diff --git a/ViCaR/AutomaticParsing.v b/ViCaR/AutomaticParsing.v deleted file mode 100644 index aa46c8c..0000000 --- a/ViCaR/AutomaticParsing.v +++ /dev/null @@ -1,216 +0,0 @@ -Require Import CategoryTypeclass. - -Open Scope Cat_scope. -Open Scope Mor_scope. -Open Scope Mon_scope. - -Inductive cat_expr {C} {cC : Category C} - : forall {A B : C}, A ~> B -> Prop := - | cat_id {A : C} : cat_expr (id_ A) - | cat_const {A B : C} (f : A ~> B) - : cat_expr f - | cat_compose {A B M : C} (f : A ~> M) (g : M ~> B) - : cat_expr (f ∘ g). - -Arguments cat_expr {_} _%Cat {_ _}%Obj _%Mor. - -Inductive moncat_expr {C} {cC : Category C} {mC : MonoidalCategory cC} - : forall {A B : C}, A ~> B -> Prop := - | moncat_id (A : C) - : moncat_expr (id_ A) - | moncat_assoc_for (A B M : C) - : moncat_expr (associator A B M).(forward) - | moncat_assoc_rev (A B M : C) - : moncat_expr (associator A B M).(reverse) - | moncat_lunit_for (A : C) - : moncat_expr (left_unitor A).(forward) - | moncat_lunit_rev (A : C) - : moncat_expr (left_unitor A).(reverse) - | moncat_runit_for (A : C) - : moncat_expr (right_unitor A).(forward) - | moncat_runit_rev (A : C) - : moncat_expr (right_unitor A).(reverse) - | moncat_compose {A B M : C} {f : A ~> B} {g : B ~> M} - : moncat_expr f -> moncat_expr g -> moncat_expr (f ∘ g) - | moncat_tensor {A1 A2 B1 B2 : C} {f : A1 ~> A2} {g : B1 ~> B2} - : moncat_expr f -> moncat_expr g -> moncat_expr (f ⊗ g) - | moncat_const {A B : C} (f : A ~> B) - : moncat_expr f. - -Arguments moncat_expr {_} {_}%Cat _%Cat {_ _}%Obj _%Mor. - -Ltac make_moncat_expr (* mC *) t := - let rec make_mon (* mC *) t := - (* lazymatch type of mC with - | @MonoidalCategory ?C ?cC => *) - match t with - | id_ ?A => constr:(moncat_id A) - | (associator ?A ?B ?C).(forward) => constr:(moncat_assoc_for A B C) - | (associator ?A ?B ?C).(forward) => constr:(moncat_assoc_for A B C) - | (left_unitor ?A).(forward) => constr:(moncat_lunit_for A) - | (left_unitor ?A).(reverse) => constr:(moncat_lunit_rev A) - | (right_unitor ?A).(forward) => constr:(moncat_runit_for A) - | (right_unitor ?A).(reverse) => constr:(moncat_runit_rev A) - | (?g ∘ ?h)%Mor => let mcg := make_mon g in let mch := make_mon h in - constr:(moncat_compose mcg mch) - | (?g ⊗ ?h)%Mon => let mcg := make_mon g in let mch := make_mon h in - constr:(moncat_tensor mcg mch) - | _ => constr:(moncat_const t) - end - (* end *) - in make_mon (* mC *) t. - -Ltac make_moncat_expr_debug (* mC *) t := - let rec make_mon (* mC *) t := - (* lazymatch type of mC with - | @MonoidalCategory ?C ?cC => *) - lazymatch t with - | id_ ?A => - - constr:(moncat_id A) - | (associator ?A ?B ?C).(forward) => - constr:(moncat_assoc_for A B C) - | (associator ?A ?B ?C).(forward) => - constr:(moncat_assoc_for A B C) - | (left_unitor ?A).(forward) => constr:(moncat_lunit_for A) - | (left_unitor ?A).(reverse) => constr:(moncat_lunit_rev A) - | (right_unitor ?A).(forward) => constr:(moncat_runit_for A) - | (right_unitor ?A).(reverse) => constr:(moncat_runit_rev A) - | (?g ∘ ?h)%Mor => let mcg := make_mon g in let mch := make_mon h in - constr:(moncat_compose mcg mch) - | (?g ⊗ ?h)%Mon => let mcg := make_mon g in let mch := make_mon h in - constr:(moncat_tensor mcg mch) - | _ => constr:(moncat_const t) - end - (* end *) - in make_mon (* mC *) t. - -Require Import String. - -(* Definition moncat_to_str : *) - -Lemma parse_test_0 : forall {C : Type} - {cC : Category C} {mC : MonoidalCategory cC} - {top botIn botMid botOut : C} - (f0 : botIn ~> botMid) (f1 : botMid ~> botOut), - (id_ top) ⊗ (f0 ∘ f1) ≃ id_ top ⊗ f0 ∘ (id_ top ⊗ f1). -Proof. - intros. - match goal with - |- (?t ≃ ?t')%Cat => let mct := make_moncat_expr_debug t in pose mct; - let mct' := make_moncat_expr_debug t' in pose mct' - end. - Admitted. - -Lemma parse_test_1 : forall {C : Type} - {cC : Category C} {mC : MonoidalCategory cC} - {top botIn botMid botOut : C} - (f0 : botIn ~> botMid) (f1 : botMid ~> botOut), - (id_ top) ⊗ (f0 ∘ f1) ∘ (λ_ _)^-1 ≃ id_ top ⊗ f0 ∘ (id_ top ⊗ f1) ∘ (λ_ _)^-1. -Proof. - intros. - match goal with - |- (?t ≃ ?t')%Cat => let mct := make_moncat_expr_debug t in pose mct; - let mct' := make_moncat_expr_debug t' in pose mct' - end. - Admitted. - - -Ltac print_moncat_expr m := - first [lazymatch m with - | moncat_id ?A => - idtac "['id', (1), (1), ['" A "'], 'id_{&1}']" - (* : moncat_expr (id_ A) *) - | moncat_assoc_for ?A ?B ?M => - idtac "['assoc_for', ((1,2),3), (1,(2,3)), ['" A "','" B "','" M "'], 'α^-1_{&1,&2,&3}']" - (* : moncat_expr (associator A B M).(forward) *) - | moncat_assoc_rev ?A ?B ?M => - idtac "['assoc_rev', ((1,2),3), (1,(2,3)), ['" A "','" B "','" M "'], 'α_{&1,&2,&3}']" - (* : moncat_expr (associator A B M).(reverse) *) - | moncat_lunit_for ?A => - idtac "['lunit_for', ('I',1), (1), ['" A "'], 'λ_{&1}']" - (* : moncat_expr (left_unitor A).(forward) *) - | moncat_lunit_rev ?A => - idtac "['lunit_rev', (1), ('I',1), ['" A "'], 'λ^-1_{&1}']" - (* : moncat_expr (left_unitor A).(reverse) *) - | moncat_runit_for ?A => - idtac "['runit_for', (1,'I'), (1), ['" A "'], 'ρ_{&1}']" - (* : moncat_expr (right_unitor A).(forward) *) - | moncat_runit_rev ?A => - idtac "['runit_rev', (1), (1,'I'), ['" A "'], 'ρ^-1_{&1}']" - (* : moncat_expr (right_unitor A).(reverse) *) - | @moncat_compose _ _ _ ?A ?B ?M _ _ ?mcf ?mcg => - idtac "['*compose', (1), (2), ['" A "','" M "'], ["; - print_moncat_expr mcf; idtac ","; print_moncat_expr mcg; idtac "]]" - (* : moncat_expr f -> moncat_expr g -> moncat_expr (f ∘ g) *) - | @moncat_tensor _ _ _ ?A1 ?A2 ?B1 ?B2 _ _ ?mcf ?mcg => - idtac "['*tensor', (1,2), (3,4), ['" A1 "','" B1 "','" A2 "','" B2 "'], ["; - print_moncat_expr mcf; idtac ","; print_moncat_expr mcg; idtac "]]" - (* : moncat_expr g -> moncat_expr h -> moncat_expr (g ⊗ h) *) - | @moncat_const _ _ _ ?A ?B ?f => - idtac "['const', ('" A "'), ('" A "'), '" f "']" - (* : moncat_expr f. *) - end | idtac "BAD EXPR" m]. - - -Ltac print_moncat_expr_JSON m := - first [lazymatch m with - | moncat_id ?A => - idtac "['id', [1], [1], ['"A"'], 'id_{&1}']" - (* : moncat_expr (id_ A) *) - | moncat_assoc_for ?A ?B ?M => - idtac "['assoc_for', [[1,2],3], [1,[2,3]], ['"A"','"B"','"M"'], 'α^-1_{&1,&2,&3}']" - (* : moncat_expr (associator A B M).(forward) *) - | moncat_assoc_rev ?A ?B ?M => - idtac "['assoc_rev', [[1,2],3], [1,[2,3]], ['"A"','"B"','"M"'], 'α_{&1,&2,&3}']" - (* : moncat_expr (associator A B M).(reverse) *) - | moncat_lunit_for ?A => - idtac "['lunit_for', ['I',1], [1], ['"A"'], 'λ_{&1}']" - (* : moncat_expr (left_unitor A).(forward) *) - | moncat_lunit_rev ?A => - idtac "['lunit_rev', [1], ['I',1], ['"A"'], 'λ^-1_{&1}']" - (* : moncat_expr (left_unitor A).(reverse) *) - | moncat_runit_for ?A => - idtac "['runit_for', [1,'I'], [1], ['"A"'], 'ρ_{&1}']" - (* : moncat_expr (right_unitor A).(forward) *) - | moncat_runit_rev ?A => - idtac "['runit_rev', [1], [1,'I'], ['"A"'], 'ρ^-1_{&1}']" - (* : moncat_expr (right_unitor A).(reverse) *) - | @moncat_compose _ _ _ ?A ?B ?M _ _ ?mcf ?mcg => - idtac "['*compose', [1], [2], ['"A"','"M"'], ["; - print_moncat_expr mcf; idtac ","; print_moncat_expr mcg; idtac "]]" - (* : moncat_expr f -> moncat_expr g -> moncat_expr (f ∘ g) *) - | @moncat_tensor _ _ _ ?A1 ?A2 ?B1 ?B2 _ _ ?mcf ?mcg => - idtac "['*tensor', [1,2], [3,4], ['"A1"','"B1"','"A2"','"B2"'], ["; - print_moncat_expr mcf; idtac ","; print_moncat_expr mcg; idtac "]]" - (* : moncat_expr g -> moncat_expr h -> moncat_expr (g ⊗ h) *) - | @moncat_const _ _ _ ?A ?B ?f => - idtac "['const', ['"A"'], ['"A"'], '"f"']" - (* : moncat_expr f. *) - end | idtac "BAD EXPR" m]. - -Lemma print_test_0 : forall {C : Type} - {cC : Category C} {mC : MonoidalCategory cC} - {top botIn botMid botOut : C} - (f0 : botIn ~> botMid) (f1 : botMid ~> botOut), - (id_ top) ⊗ (f0 ∘ f1) ≃ id_ top ⊗ f0 ∘ (id_ top ⊗ f1). -Proof. - intros. - match goal with - |- (?t ≃ ?t')%Cat => let mct := make_moncat_expr_debug t in print_moncat_expr_JSON mct - (* let mct' := make_moncat_expr_debug t' in pose mct' *) - end. - Admitted. - -Lemma print_test_1 : forall {C : Type} - {cC : Category C} {mC : MonoidalCategory cC} - {top botIn botMid botOut : C} - (f0 : botIn ~> botMid) (f1 : botMid ~> botOut), - (id_ top) ⊗ (f0 ∘ f1) ∘ (λ_ _)^-1 ≃ id_ top ⊗ f0 ∘ (id_ top ⊗ f1) ∘ (λ_ _)^-1. -Proof. - intros. - match goal with - |- (?t ≃ ?t')%Cat => let mct := make_moncat_expr_debug t in pose mct; - let mct' := make_moncat_expr_debug t' in pose mct' - end. - Admitted. diff --git a/ViCaR/CategoryAutomation.v b/ViCaR/CategoryAutomation.v index 7382cce..bc109ee 100644 --- a/ViCaR/CategoryAutomation.v +++ b/ViCaR/CategoryAutomation.v @@ -49,9 +49,9 @@ Ltac fold_Category cC := let cat_fold f := (let base := base_fn @f in let catted := catify @f in - change base with catted in *) in + change base with catted in * ) in try cat_fold @morphism; (* has issues, e.g., with ZX - - might be fixable, but likely not necessary*) + might be fixable, but likely not necessary *) cat_fold @compose; cat_fold @c_identity; let cid := base_fn @c_identity in @@ -60,13 +60,13 @@ Ltac fold_Category cC := let x' := eval unfold x in x in let cidx := eval cbn in (cid x') in pose (eq_refl : cidx = cC.(c_identity) x') as H; - erewrite H; clear x H)); + erewrite H in *; clear x H)); let eqbase := base_fn @c_equiv in let eqcat := catify @c_equiv in - change eqbase with eqcat; (* I think this is a no-op *) + change eqbase with eqcat in *; (* I think this is a no-op *) repeat (progress match goal with |- eqbase ?A ?B ?f ?g - => change (eqbase A B f g) with (eqcat A B f g) + => change (eqbase A B f g) with (eqcat A B f g) in * end); try let H' := fresh in enough (H':(@c_equiv _ _ _ _ _ _)) by (eapply H') @@ -92,22 +92,20 @@ Ltac fold_MonoidalCategory mC := let t := eval cbn in raw in constr:(t)) in let f_fold f := (let base := base_fn @f in - change base with f) in + change base with f in *) in let cat_fold f := (let base := cbase_fn @f in let catted := catify @f in - change base with catted in *) in + change base with catted in * ) in let mcat_fold f := (let base := mbase_fn @f in let catted := mcatify @f in - change base with catted in *) in - let tens := mbase_fn @tensor in - let ob_base := base_fn (@obj_bimap C C C cC cC cC tens) in - change ob_base with mC.(tensor).(obj_bimap); - let mor_base := base_fn (@morphism_bimap C C C cC cC cC tens) in - change mor_base with (@morphism_bimap C C C cC cC cC mC.(tensor)) - (* (@tensor C cC mC).(@morphism_bimap C C C cC cC cC) *); - mcat_fold @I; + change base with catted in * ) in + let tens_obj := base_fn (@obj_tensor C cC mC) in + change tens_obj with (@obj_tensor C cC mC) in *; + let tens_mor := base_fn (@mor_tensor C cC mC) in + change tens_mor with (@mor_tensor C cC mC) in *; + mcat_fold @mon_I; let lunit := mbase_fn @left_unitor in repeat progress ( let H := fresh in let x := fresh in @@ -115,7 +113,7 @@ Ltac fold_MonoidalCategory mC := let x' := eval unfold x in x in let lunitx := eval cbn in ((lunit x').(forward)) in pose (eq_refl : lunitx = (mC.(left_unitor) x').(forward)) as H; - erewrite H; clear x H); + erewrite H in *; clear x H); let runit := mbase_fn @right_unitor in repeat progress ( let H := fresh in let x := fresh in @@ -123,7 +121,7 @@ Ltac fold_MonoidalCategory mC := let x' := eval unfold x in x in let runitx := eval cbn in ((runit x').(forward)) in pose (eq_refl : runitx = (mC.(right_unitor) x').(forward)) as H; - erewrite H; clear x H) + erewrite H in *; clear x H) end. Ltac fold_all_monoidal_categories := @@ -153,36 +151,56 @@ Ltac fold_BraidedMonoidalCategory bC := let cat_fold f := (let base := cbase_fn @f in let catted := catify @f in - change base with catted in *) in + change base with catted in * ) in let mcat_fold f := (let base := mbase_fn @f in let catted := mcatify @f in - change base with catted in *) in + change base with catted in * ) in let bcat_fold f := (let base := bbase_fn @f in let catted := bcatify @f in - change base with catted in *) in + change base with catted in * ) in let braid := bbase_fn @braiding in - let braidbase := constr:(ltac:(first [exact (ltac:(eval unfold braid in braid)) | exact braid])) in - let braidforw := eval cbn in - (fun A B => (braidbase.(component_biiso) A B).(forward)) in - repeat progress (let H := fresh in let y := fresh in let x := fresh in - evar (y : C); evar (x : C); - let x' := eval unfold x in x in let y' := eval unfold y in y in - let braidforwxy := eval cbn in (braidforw x' y') in - pose (eq_refl : braidforwxy = - (bC.(braiding).(component_biiso) x' y').(forward)) as H; - erewrite H; clear x y H); + change braid with (@braiding C cC mC bC); + let braidbase := constr:( + ltac:(first [exact (ltac:(eval unfold braid in braid)) + | exact braid])) in - let braidrev := eval cbn in - (fun A B => (braidbase.(component_biiso) A B).(reverse)) in - repeat progress (let H := fresh in let y := fresh in let x := fresh in - evar (y : C); evar (x : C); - let x' := eval unfold x in x in let y' := eval unfold y in y in - let braidrevxy := eval cbn in (braidrev x' y') in - pose (eq_refl : braidrevxy = - (bC.(braiding).(component_biiso) x' y').(reverse)) as H; - erewrite H; clear x y H) + let braidforw := eval simpl in + (fun A B => (braidbase A B).(forward)) in + first [ + progress ( + match braidforw with + | fun n m => ?f n m => + change (f ?k ?l) with ((bC.(braiding) k l).(forward)) + end + ) + | repeat progress (let H := fresh in let y := fresh in let x := fresh in + evar (y : C); evar (x : C); + let x' := eval unfold x in x in let y' := eval unfold y in y in + let braidforwxy := eval simpl in (braidforw x' y') in + pose (eq_refl : braidforwxy = + (bC.(braiding) x' y').(forward)) as H; + erewrite H; clear x y H) + ]; + + let braidrev := eval simpl in + (fun A B => (braidbase A B).(reverse)) in + first [ + progress ( + match braidrev with + | fun n m => ?f n m => + change (f ?k ?l) with ((bC.(braiding) k l).(reverse)) + end + ) + | repeat progress (let H := fresh in let y := fresh in let x := fresh in + evar (y : C); evar (x : C); + let x' := eval unfold x in x in let y' := eval unfold y in y in + let braidrevxy := eval simpl in (braidrev x' y') in + pose (eq_refl : braidrevxy = + (bC.(braiding) x' y').(reverse)) as H; + erewrite H; clear x y H) + ] end. Ltac fold_all_braided_monoidal_categories := @@ -192,132 +210,197 @@ Ltac fold_all_braided_monoidal_categories := fold_BraidedMonoidalCategory bC; subst x end. -Ltac fold_CompactClosedCategory ccC := - match type of ccC with @CompactClosedCategory ?C ?cC ?mC ?bC ?sC => +Ltac to_Cat := + fold_all_categories; fold_all_monoidal_categories; + fold_all_braided_monoidal_categories. + + + +Ltac fold_Category_in cC H := + match type of cC with Category ?C => + let catify f := constr:(@f C cC) in + let base_fn f := (let raw := catify f in + let t := eval cbn in raw in constr:(t)) in + let cat_fold f := + (let base := base_fn @f in + let catted := catify @f in + change base with catted in H) in + try cat_fold @morphism; + cat_fold @compose; + cat_fold @c_identity; + let cid := base_fn @c_identity in + (repeat progress ( + let H := fresh in let x := fresh in evar (x : C); + let x' := eval unfold x in x in + let cidx := eval cbn in (cid x') in + pose (eq_refl : cidx = cC.(c_identity) x') as H; + erewrite H in *; clear x H)); + let eqbase := base_fn @c_equiv in + let eqcat := catify @c_equiv in + change eqbase with eqcat in H; (* I think this is a no-op *) + repeat (progress match goal with + |- eqbase ?A ?B ?f ?g + => change (eqbase A B f g) with (eqcat A B f g) in H + end); + try let H' := fresh in + enough (H':(@c_equiv _ _ _ _ _ _)) by (eapply H') + end. + +Ltac fold_all_categories_in H := + saturate_instances Category; + repeat match goal with + | x := ?cC : Category ?C |- _ => (* idtac x ":=" cC ": Category" C ; *) + fold_Category_in cC H; subst x + (* | x : Category ?C |- _ => idtac x; fold_Category x; subst x *) + end. + +Ltac fold_MonoidalCategory_in mC H := + match type of mC with @MonoidalCategory ?C ?cC => let catify f := constr:(@f C cC) in let mcatify f := constr:(@f C cC mC) in - let bcatify f := constr:(@f C cC mC bC) in - let cccatify f := constr:(@f C cC mC bC sC ccC) in let base_fn f := (let t := eval cbn in f in constr:(t)) in let cbase_fn f := (let raw := catify f in let t := eval cbn in raw in constr:(t)) in let mbase_fn f := (let raw := mcatify f in let t := eval cbn in raw in constr:(t)) in - let bbase_fn f := (let raw := bcatify f in - let t := eval cbn in raw in constr:(t)) in - let ccbase_fn f := (let raw := cccatify f in - let t := eval cbn in raw in constr:(t)) in let f_fold f := (let base := base_fn @f in - change base with f) in + change base with f in H) in let cat_fold f := (let base := cbase_fn @f in let catted := catify @f in - change base with catted in *) in + change base with catted in H) in let mcat_fold f := (let base := mbase_fn @f in let catted := mcatify @f in - change base with catted in *) in - let bcat_fold f := - (let base := bbase_fn @f in - let catted := bcatify @f in - change base with catted in *) in - let cccat_fold f := - (let base := ccbase_fn @f in - let catted := cccatify @f in - change base with catted in *) in - - let dua := ccbase_fn @dual in - first [ - (unify dua (@id C) (*; idtac "would loop" *) ) - | ( + change base with catted in H) in + let tens_obj := base_fn (@obj_tensor C cC mC) in + change tens_obj with (@obj_tensor C cC mC) in H; + let tens_mor := base_fn (@mor_tensor C cC mC) in + change tens_mor with (@mor_tensor C cC mC) in H; + mcat_fold @mon_I; + let lunit := mbase_fn @left_unitor in repeat progress ( - let H := fresh in let x := fresh in - evar (x : C); (* TODO: Test this - last I tried it was uncooperative *) - let x' := eval unfold x in x in - let duax := eval cbn in (dua x') in - pose (eq_refl : duax = ccC.(dual) x') as H; - erewrite H; clear x H))]; - - cccat_fold @unit; - cccat_fold @counit; - - let uni := ccbase_fn @unit in - repeat progress (let H := fresh in let x := fresh in - evar (x : C); + let Hx := fresh in let x := fresh in + evar (x : C); (* TODO: Test this - last I tried it was uncooperative *) let x' := eval unfold x in x in - let unix := eval cbn in (uni x') in - pose (eq_refl : unix = - ccC.(unit) x') as H; - erewrite H; clear x H); - - let couni := ccbase_fn @counit in - repeat progress (let H := fresh in let x := fresh in - evar (x : C); + let lunitx := eval cbn in ((lunit x').(forward)) in + pose (eq_refl : lunitx = (mC.(left_unitor) x').(forward)) as Hx; + erewrite Hx in H; clear x Hx); + repeat progress ( + let Hx := fresh in let x := fresh in + evar (x : C); (* TODO: Test this - last I tried it was uncooperative *) + let x' := eval unfold x in x in + let lunitx := eval cbn in ((lunit x').(reverse)) in + pose (eq_refl : lunitx = (mC.(left_unitor) x').(reverse)) as Hx; + erewrite Hx in H; clear x Hx); + let runit := mbase_fn @right_unitor in + repeat progress ( + let Hx := fresh in let x := fresh in + evar (x : C); (* TODO: Test this - last I tried it was uncooperative *) + let x' := eval unfold x in x in + let runitx := eval cbn in ((runit x').(forward)) in + pose (eq_refl : runitx = (mC.(right_unitor) x').(forward)) as Hx; + erewrite Hx in H; clear x Hx); + repeat progress ( + let Hx := fresh in let x := fresh in + evar (x : C); (* TODO: Test this - last I tried it was uncooperative *) let x' := eval unfold x in x in - let counix := eval cbn in (couni x') in - pose (eq_refl : counix = - ccC.(counit) x') as H; - erewrite H; clear x H) + let runitx := eval cbn in ((runit x').(reverse)) in + pose (eq_refl : runitx = (mC.(right_unitor) x').(reverse)) as Hx; + erewrite Hx in H; clear x Hx) end. -Ltac fold_all_compact_closed_categories := - saturate_instances CompactClosedCategory; +Ltac fold_all_monoidal_categories_in H := + saturate_instances MonoidalCategory; repeat match goal with - | x := ?ccC : CompactClosedCategory ?C |- _ => - fold_CompactClosedCategory ccC; subst x + | x := ?mC : MonoidalCategory ?C |- _ => (* idtac x ":=" cC ": Category" C ; *) + fold_MonoidalCategory_in mC H; subst x + (* | x : Category ?C |- _ => idtac x; fold_Category x; subst x *) end. - -Ltac fold_DaggerCategory dC := - match type of dC with @DaggerCategory ?C ?cC => +Ltac fold_BraidedMonoidalCategory_in bC H := + match type of bC with @BraidedMonoidalCategory ?C ?cC ?mC => let catify f := constr:(@f C cC) in - let dcatify f := constr:(@f C cC dC) in + let mcatify f := constr:(@f C cC mC) in + let bcatify f := constr:(@f C cC mC bC) in let base_fn f := (let t := eval cbn in f in constr:(t)) in let cbase_fn f := (let raw := catify f in let t := eval cbn in raw in constr:(t)) in - let dbase_fn f := (let raw := dcatify f in + let mbase_fn f := (let raw := mcatify f in + let t := eval cbn in raw in constr:(t)) in + let bbase_fn f := (let raw := bcatify f in let t := eval cbn in raw in constr:(t)) in let f_fold f := (let base := base_fn @f in - change base with f) in + change base with f in H) in let cat_fold f := (let base := cbase_fn @f in let catted := catify @f in - change base with catted in *) in - let dcat_fold f := - (let base := dbase_fn @f in - let catted := dcatify @f in - change base with catted in *) in - - dcat_fold @adjoint; - - let adj := dbase_fn @adjoint in - repeat progress (let H := fresh in - let x := fresh in let y := fresh in - evar (x : C); evar (y : C); - let x' := eval unfold x in x in - let y' := eval unfold y in y in - let adjxy := eval cbn in (adj x' y') in - pose (eq_refl : adjxy = - dC.(adjoint) (A:=x') (B:=y')) as H; - erewrite H; clear x y H) + change base with catted in H) in + let mcat_fold f := + (let base := mbase_fn @f in + let catted := mcatify @f in + change base with catted in H) in + let bcat_fold f := + (let base := bbase_fn @f in + let catted := bcatify @f in + change base with catted in H) in + let braid := bbase_fn @braiding in + change braid with (@braiding C cC mC bC); + let braidbase := constr:( + ltac:(first [exact (ltac:(eval unfold braid in braid)) + | exact braid])) in + + let braidforw := eval simpl in + (fun A B => (braidbase A B).(forward)) in + first [ + progress ( + match braidforw with + | fun n m => ?f n m => + change (f ?k ?l) with ((bC.(braiding) k l).(forward)) in H + end + ) + | repeat progress (let Hx := fresh in let y := fresh in let x := fresh in + evar (y : C); evar (x : C); + let x' := eval unfold x in x in let y' := eval unfold y in y in + let braidforwxy := eval simpl in (braidforw x' y') in + pose (eq_refl : braidforwxy = + (bC.(braiding) x' y').(forward)) as Hx; + erewrite Hx in H; clear x y Hx) + ]; + + let braidrev := eval simpl in + (fun A B => (braidbase A B).(reverse)) in + first [ + progress ( + match braidrev with + | fun n m => ?f n m => + change (f ?k ?l) with ((bC.(braiding) k l).(reverse)) in H + end + ) + | repeat progress (let Hx := fresh in let y := fresh in let x := fresh in + evar (y : C); evar (x : C); + let x' := eval unfold x in x in let y' := eval unfold y in y in + let braidrevxy := eval simpl in (braidrev x' y') in + pose (eq_refl : braidrevxy = + (bC.(braiding) x' y').(reverse)) as Hx; + erewrite Hx in H; clear x y Hx) + ] end. -Ltac fold_all_dagger_categories := - saturate_instances DaggerCategory; +Ltac fold_all_braided_monoidal_categories_in H := + saturate_instances BraidedMonoidalCategory; repeat match goal with - | x := ?dC : DaggerCategory ?C |- _ => - fold_DaggerCategory dC; subst x + | x := ?bC : BraidedMonoidalCategory ?C |- _ => + fold_BraidedMonoidalCategory_in bC H; subst x end. -Ltac to_Cat := - fold_all_categories; fold_all_monoidal_categories; - fold_all_braided_monoidal_categories; - fold_all_compact_closed_categories; - fold_all_dagger_categories. +Ltac to_Cat_in H := + fold_all_categories_in H; fold_all_monoidal_categories_in H; + fold_all_braided_monoidal_categories_in H. @@ -325,7 +408,7 @@ Ltac to_Cat := Ltac tensor_free f := try match f with - | context[@morphism_bimap _ _ _ _ _ _ (@tensor _ _ _)] => fail 2 + | context[@mor_tensor _ _ _ _ _ _ _] => fail 2 end. Ltac compose_free f := @@ -356,21 +439,15 @@ Ltac tensor_only f := Ltac compose_only f := first [pseudo_const f - | lazymatch f with (* TODO: Does lazy matter here? Pretty sure it doesn't hurt, but idk if it'd ever match more than once anyways*) - | (?g ∘ ?h)%Mor => (* old: @compose _ _ _ _ _ ?g ?h *) + | lazymatch f with + | (?g ∘ ?h)%Cat => compose_only g; compose_only h end]. -(* Old: -| @compose _ _ _ _ _ ?g ?h => tensor_only g; is_weak_fenced h -| @morphism_bimap _ _ _ _ _ _ (@tensor _ _ _) _ _ _ _ ?g ?h => - tensor_only g; tensor_only h -| _ => pseudo_const f - *) Ltac is_weak_fenced f := lazymatch f with - | (?g ∘ ?h)%Mor => tensor_only g; is_weak_fenced h - | (?g ⊗ ?h)%Mor => + | (?g ∘ ?h)%Cat => tensor_only g; is_weak_fenced h + | (?g ⊗ ?h)%Cat => tensor_only g; tensor_only h | _ => pseudo_const f end. @@ -382,138 +459,45 @@ Ltac is_weak_fenced f := | comp g h : tensor_only g -> is_weak_fence h -> is_weak_fence (g ∘ h). *) -(* NOTE: This old partial implementation took a very step-like approach. I'd - rather progress in big stes (i.e. strongly recurse) *) -Ltac weak_fencepost_form_debug_old f := - let rec weak_fencepost_form f := - match f with - | (?g ∘ ?h)%Mor => - let _ := match goal with _ => tensor_only g end in (* test for tensor-only while returning constr *) - let _ := match goal with _ => idtac "left composite" g "is tensor-only" end in - let Nh := weak_fencepost_form h in - (* let _ := match goal with _ => idtac h "normalizes to" Nh end in *) - constr:((g ∘ Nh)%Mor) - | (?g ∘ ?h)%Mor => - match g with - | (?g' ∘ ?h')%Mor => - (* We _can_ do this, but we can also just recurse: - let _ := tensor_only g' in - let _ := match goal with _ => idtac "associating as" g' "is tensor-only" end in - let Ng'h := weak_fencepost_form (g' ∘ h) in *) - (* Ditto for this... - let _ := match goal with _ => idtac "associating to " g' "∘ (" h' "∘" h ")" end in - let Ng' := weak_fencepost_form g' in - let Nh'h := weak_fencepost_form (h' ∘ h) in - let Nf := weak_fencepost_form (Ng' ∘ Nh'h) in - constr:(Nf)*) - let _ := match goal with _ => idtac "associating to " g' "∘ (" h' "∘" h ")" end in - let Nf := weak_fencepost_form (g' ∘ (h' ∘ h))%Mor in - constr:(Nf) - | _ => - let _ := match goal with _ => - idtac "WARNING:" g "is not tensor-only or a composition" end in - let Nh := weak_fencepost_form h in - constr:((g ∘ h)%Mor) - end - | _ => let _ := match goal with _ => tensor_only f end in - let _ := match goal with _ => - idtac f "is tensor-only" end in - constr:(f) - | _ => - let _ := match goal with _ => - idtac "INFO:" f "is const or unsupported" end in - constr:(f) - end - in weak_fencepost_form f. - - -Ltac right_associate f := +Ltac right_associate_term f := match f with - | ((?g ∘ ?h) ∘ ?i)%Mor => right_associate (g ∘ (h ∘ i))%Mor - | (?g ∘ ?h)%Mor => (* g shouldn't be a composition *) - let RAh := right_associate h in - constr:((g ∘ RAh)%Mor) + | ((?g ∘ ?h) ∘ ?i)%Cat => right_associate_term (g ∘ (h ∘ i))%Cat + | (?g ∘ ?h)%Cat => (* g shouldn't be a composition *) + let RAh := right_associate_term h in + constr:((g ∘ RAh)%Cat) | _ => constr:(f) end. (* TODO: Test this! *) -Ltac left_associate f := +Ltac left_associate_term f := match f with - | (?g ∘ (?h ∘ ?i))%Mor => left_associate ((g ∘ h) ∘ i)%Mor - | (?g ∘ ?h)%Mor => (* h shouldn't be a composition *) - let LAg := left_associate g in - constr:((LAg ∘ h)%Mor) + | (?g ∘ (?h ∘ ?i))%Cat => left_associate_term ((g ∘ h) ∘ i)%Cat + | (?g ∘ ?h)%Cat => (* h shouldn't be a composition *) + let LAg := left_associate_term g in + constr:((LAg ∘ h)%Cat) | _ => constr:(f) end. -Ltac merge_stacked_composition_old g h := - (* g ⊗ h = (g0 ∘ (g1 ∘ ...)) ⊗ (h0 ∘ (h1 ∘ ...)) - ===> (g0 ⊗ h0) ∘ (g1 ⊗ h1) ∘ ... - with id_ inserted as needed. *) - (* In gallina: - match g, h with - | ?g0 ∘ ?g1, ?h0 ∘ ?h1 => - let rest := merge_stacked_composition g1 h1 in - constr:(g0 ⊗ h0 ∘ rest) - | ?g0 ∘ ?g1, ?h0 => - match type of h0 with ?A ~> ?B => - let rest := merge_stacked_composition g1 (id_ B) in - constr:(g0 ⊗ h0 ∘ rest) - end - | ?g0, ?h0 ∘ ?h1 => - match type of g0 with ?A ~> ?B => - let rest := merge_stacked_composition (id_ B) h1 in - constr:(g0 ⊗ h0 ∘ rest) - end - | _, _ => constr:(g ⊗ h) - end.*) - (* With ltac restrictions, *) - let rec merge_stacked_composition g h := - match g with - | (?g0 ∘ ?g1)%Mor => - match h with - | (?h0 ∘ ?h1)%Mor => - let rest := merge_stacked_composition g1 h1 in - constr:((g0 ⊗ h0 ∘ rest)%Mor) - | _ => - match type of h with (?A ~> ?B)%Cat => - let rest := merge_stacked_composition g1 (id_ B)%Mor in - constr:((g0 ⊗ h ∘ rest)%Mor) - end - end - | _ => - match h with - | (?h0 ∘ ?h1)%Mor => - match type of g with (?A ~> ?B)%Cat => - let rest := merge_stacked_composition ((id_ B)%Mor) h1 in - constr:((g ⊗ h0 ∘ rest)%Mor) - end - | _ => - constr:((g ⊗ h)%Mor) - end - end - in merge_stacked_composition g h. - Ltac merge_stacked_composition_debug gh := let rec merge_stacked_composition gh := match type of gh with @morphism ?C ?cC _ _ => match gh with - @morphism_bimap C _ _ cC _ _ (@tensor C cC ?mC) ?gA ?gB ?hA ?hB ?g ?h => + @mor_tensor C cC ?mC ?gA ?gB ?hA ?hB ?g ?h => lazymatch g with - | (?g0 ∘ ?g1)%Mor => + | (?g0 ∘ ?g1)%Cat => let _ := match goal with _ => idtac "found decomp of first as" g0 "∘" g1 end in lazymatch h with - | (?h0 ∘ ?h1)%Mor => + | (?h0 ∘ ?h1)%Cat => let _ := match goal with _ => idtac "found decomp of second as" h0 "∘" h1 end in - let rest := merge_stacked_composition ((mC.(tensor) @@ g1, h1)%Mor) in + let rest := merge_stacked_composition ((mC.(mor_tensor) g1 h1)%Cat) in let _ := match goal with _ => idtac "remaining terms became" rest end in - let res := constr:(cC.(compose) (mC.(tensor) @@ g0, h0) rest) in + let res := constr:(cC.(compose) (mC.(mor_tensor) g0 h0) rest) in let _ := match goal with _ => idtac " which combined to" res end in constr:(res) @@ -523,10 +507,10 @@ Ltac merge_stacked_composition_debug gh := match type of h with @morphism _ _ ?A ?B => let _ := match goal with _ => idtac "resolved second as type" hA "~>" hB end in - let rest := merge_stacked_composition ((mC.(tensor) @@ g1, (cC.(c_identity) hB))%Mor) in + let rest := merge_stacked_composition ((mC.(mor_tensor) g1 (cC.(c_identity) hB))%Cat) in let _ := match goal with _ => idtac "remaining terms became" rest end in - let res := constr:(cC.(compose) (mC.(tensor) @@ g0, h) rest) in + let res := constr:(cC.(compose) (mC.(mor_tensor) g0 h) rest) in let _ := match goal with _ => idtac " which combined to" res end in constr:(res) @@ -536,16 +520,16 @@ Ltac merge_stacked_composition_debug gh := let _ := match goal with _ => idtac "found first to be atomic:" g end in lazymatch h with - | (?h0 ∘ ?h1)%Mor => + | (?h0 ∘ ?h1)%Cat => let _ := match goal with _ => idtac "found decomp of second as" h0 "∘" h1 end in match type of g with @morphism _ _ ?A ?B => let _ := match goal with _ => idtac "resolved second as type" gA "~>" gB end in - let rest := merge_stacked_composition ((mC.(tensor) @@ (cC.(c_identity) gB), h1)%Mor) in + let rest := merge_stacked_composition ((mC.(mor_tensor) (cC.(c_identity) gB) h1)%Cat) in let _ := match goal with _ => idtac "remaining terms became" rest end in - let res := constr:(cC.(compose) (mC.(tensor) @@ g, h0) rest) in + let res := constr:(cC.(compose) (mC.(mor_tensor) g h0) rest) in let _ := match goal with _ => idtac " which combined to" res end in constr:(res) @@ -553,7 +537,7 @@ Ltac merge_stacked_composition_debug gh := | _ => let _ := match goal with _ => idtac "found second to be atomic as well:" h end in - constr:((mC.(tensor) @@ g, h)%Mor) + constr:((mC.(mor_tensor) g h)%Cat) end end end end @@ -563,28 +547,28 @@ Ltac merge_stacked_composition gh := let rec merge_stacked_composition gh := match type of gh with @morphism ?C ?cC _ _ => match gh with - @morphism_bimap C _ _ cC _ _ (@tensor C cC ?mC) ?gA ?gB ?hA ?hB ?g ?h => + @mor_tensor C cC ?mC ?gA ?gB ?hA ?hB ?g ?h => lazymatch g with - | (?g0 ∘ ?g1)%Mor => + | (?g0 ∘ ?g1)%Cat => lazymatch h with - | (?h0 ∘ ?h1)%Mor => - let rest := merge_stacked_composition ((mC.(tensor) @@ g1, h1)%Mor) in - constr:(cC.(compose) (mC.(tensor) @@ g0, h0) rest) + | (?h0 ∘ ?h1)%Cat => + let rest := merge_stacked_composition ((mC.(mor_tensor) g1 h1)%Cat) in + constr:(cC.(compose) (mC.(mor_tensor) g0 h0) rest) | _ => let rest := merge_stacked_composition - ((mC.(tensor) @@ g1, (cC.(c_identity) hB))%Mor)in - constr:(cC.(compose) (mC.(tensor) @@ g0, h) rest) + ((mC.(mor_tensor) g1 (cC.(c_identity) hB))%Cat)in + constr:(cC.(compose) (mC.(mor_tensor) g0 h) rest) end | _ => lazymatch h with - | (?h0 ∘ ?h1)%Mor => + | (?h0 ∘ ?h1)%Cat => let rest := merge_stacked_composition - ((mC.(tensor) @@ (cC.(c_identity) gB), h1)%Mor) in - constr:(cC.(compose) (mC.(tensor) @@ g, h0) rest) + ((mC.(mor_tensor) (cC.(c_identity) gB) h1)%Cat) in + constr:(cC.(compose) (mC.(mor_tensor) g h0) rest) | _ => - constr:((mC.(tensor) @@ g, h)%Mor) + constr:((mC.(mor_tensor) g h)%Cat) end end end end @@ -600,18 +584,18 @@ Ltac weak_fencepost_form_debug f := let Nh := weak_fencepost h in let _ := match goal with _ => idtac "... getting" g "∘" h "into" end in - let res := right_associate (cC.(compose) Ng Nh) in + let res := right_associate_term (cC.(compose) Ng Nh) in let _ := match goal with _ => idtac " " res end in constr:(res) - | @morphism_bimap ?C _ _ ?cC _ _ (@tensor ?C ?cC ?mC) _ _ _ _ ?g ?h => + | @mor_tensor ?C ?cC ?mC _ _ _ _ ?g ?h => let _ := match goal with _ => idtac "splitting on ⊗ into" g "and" h "..." end in let Ng := weak_fencepost g in let Nh := weak_fencepost h in let _ := match goal with _ => idtac "... getting" g "⊗" h "into" end in - let res := merge_stacked_composition ((mC.(tensor) @@ Ng, Nh)%Mor) in + let res := merge_stacked_composition ((mC.(mor_tensor) Ng Nh)%Cat) in let _ := match goal with _ => idtac " " res end in constr:(res) @@ -628,25 +612,36 @@ Ltac weak_fencepost_form f := | @compose ?C ?cC _ _ _ ?g ?h => let Ng := weak_fencepost g in let Nh := weak_fencepost h in - right_associate (cC.(compose) Ng Nh) - | @morphism_bimap ?C _ _ ?cC _ _ (@tensor ?C ?cC ?mC) _ _ _ _ ?g ?h => + right_associate_term (cC.(compose) Ng Nh) + | @mor_tensor ?C ?cC ?mC _ _ _ _ ?g ?h => let Ng := weak_fencepost g in let Nh := weak_fencepost h in - merge_stacked_composition ((mC.(tensor) @@ Ng, Nh)%Mor) + merge_stacked_composition ((mC.(mor_tensor) Ng Nh)%Cat) | _ => (* f "is const or unsupported" *) constr:(f) end in weak_fencepost f. +Section HelperLemmas. + +Context {C} {cC : Category C} {cCh : CategoryCoherence cC}. + Local Open Scope Cat_scope. -Lemma assoc_compat_helper {C} `{Category C} {A B M N : C} : + +Lemma assoc_compat_helper {A B M N : C} : forall (f : A ~> B) (g : B ~> M) (h : M ~> N) (fgh : A ~> N), f ∘ (g ∘ h) ≃ fgh -> (f ∘ g) ∘ h ≃ fgh. Proof. intros; rewrite assoc; easy. Qed. -Lemma compose_compat_right {C} `{Category C} {A B M : C} : +Lemma assoc_compat_helper' {A B M N : C} + : forall (f : A ~> B) + (g : B ~> M) (h : M ~> N) (fgh : A ~> N), + (f ∘ g) ∘ h ≃ fgh -> f ∘ (g ∘ h) ≃ fgh. +Proof. intros; rewrite <- assoc; easy. Qed. + +Lemma compose_compat_right {A B M : C} : forall (f : A ~> B) (g g' : B ~> M), g ≃ g' -> f ∘ g ≃ f ∘ g'. Proof. @@ -654,18 +649,28 @@ Proof. apply compose_compat; easy. Qed. -Lemma stack_compose_distr_compat {C} `{MonoidalCategory C} -{A B M A' B' M': C} : +Lemma compose_compat_trans_helper {A B M : C} : forall + (f f' : A ~> B) (g g' : B ~> M) (fg : A ~> M), + f ≃ f' -> g ≃ g' -> f' ∘ g' ≃ fg -> f ∘ g ≃ fg. +Proof. + intros. + transitivity (f' ∘ g')%Cat; [|easy]. + apply compose_compat; easy. +Qed. + +Context {mC : MonoidalCategory cC} {mCh : MonoidalCategoryCoherence mC}. + +Lemma stack_compose_distr_compat {A B M A' B' M': C} : forall (f : A ~> B) (g : B ~> M) (h : A' ~> B') (i : B' ~> M') (gi : B × B' ~> M × M'), g ⊗ i ≃ gi -> (f ∘ g) ⊗ (h ∘ i) ≃ f ⊗ h ∘ gi. Proof. intros. - rewrite compose_bimap. + rewrite tensor_compose. apply compose_compat; easy. Qed. -Lemma stack_distr_pushout_r_top_compat {C} `{MonoidalCategory C} +Lemma stack_distr_pushout_r_top_compat {a b m n o} : forall (f : a ~> b) (g : m ~> n) (h : n ~> o) (ih : b × n ~> b × o), id_ b ⊗ h ≃ ih -> f ⊗ (g ∘ h) ≃ f ⊗ g ∘ ih. @@ -674,11 +679,11 @@ Proof. (* `rewrite stack_distr_pushout_r_top.` is replaced here manually to simplify dependencies *) rewrite <- (right_unit f) at 1. - rewrite compose_bimap. + rewrite tensor_compose. apply compose_compat; easy. Qed. -Lemma stack_distr_pushout_r_bot_compat {C} `{MonoidalCategory C} +Lemma stack_distr_pushout_r_bot_compat {a b c n o : C} : forall (f : a ~> b) (g : b ~> c) (h : n ~> o) (gi : b × o ~> c × o), g ⊗ id_ o ≃ gi -> (f ∘ g) ⊗ h ≃ f ⊗ h ∘ gi. @@ -687,49 +692,40 @@ Proof. (* `rewrite stack_distr_pushout_r_bot.` is replaced here manually to simplify dependencies *) rewrite <- (right_unit h) at 1. - rewrite compose_bimap. + rewrite tensor_compose. apply compose_compat; easy. Qed. -Lemma compose_compat_trans_helper {C} `{cC:Category C} {A B M : C} : forall - (f f' : A ~> B) (g g' : B ~> M) (fg : A ~> M), - f ≃ f' -> g ≃ g' -> f' ∘ g' ≃ fg -> f ∘ g ≃ fg. -Proof. - intros. - transitivity (f' ∘ g')%Mor; [|easy]. - apply compose_compat; easy. -Qed. - -Lemma stack_compat_trans_helper {C} `{mC : MonoidalCategory C} {A A' B B' : C} : +Lemma stack_compat_trans_helper {A A' B B' : C} : forall (f f' : A ~> A') (g g' : B ~> B') (fg : A × B ~> A' × B'), f ≃ f' -> g ≃ g' -> f' ⊗ g' ≃ fg -> f ⊗ g ≃ fg. Proof. intros. transitivity (f' ⊗ g'); [|easy]. - apply morphism_bicompat; easy. + apply tensor_compat; easy. Qed. -Lemma show_equiv_stack_comp_id_bot_helper {C} `{MonoidalCategory C} - {A M A' B : C} : forall (g : A ~> M) (gs : M ~> A') (gsB : M × B ~> A' × B), +Lemma show_equiv_stack_comp_id_bot_helper {A M A' B : C} : + forall (g : A ~> M) (gs : M ~> A') (gsB : M × B ~> A' × B), gs ⊗ id_ B ≃ gsB -> (g ∘ gs) ⊗ id_ B ≃ g ⊗ id_ B ∘ gsB. Proof. intros. rewrite <- (right_unit (id_ B)) at 1. - rewrite compose_bimap. + rewrite tensor_compose. apply compose_compat; easy. Qed. -Lemma show_equiv_stack_comp_id_top_helper {C} `{MonoidalCategory C} - {A B M B' : C} : forall (g : B ~> M) (gs : M ~> B') (Ags : A × M ~> A × B'), +Lemma show_equiv_stack_comp_id_top_helper {A B M B' : C} : + forall (g : B ~> M) (gs : M ~> B') (Ags : A × M ~> A × B'), id_ A ⊗ gs ≃ Ags -> id_ A ⊗ (g ∘ gs) ≃ id_ A ⊗ g ∘ Ags. Proof. intros. rewrite <- (right_unit (id_ A)) at 1. - rewrite compose_bimap. + rewrite tensor_compose. apply compose_compat; easy. Qed. -Lemma show_equiv_unfold_tensor_stack_helper {C} `{MonoidalCategory C} +Lemma show_equiv_unfold_tensor_stack_helper {fA fB gA gB : C} (f uf : fA ~> fB) (g ug : gA ~> gB) (ufs : fA × gA ~> fB × gA) (ugs : fB × gA ~> fB × gB) : f ≃ uf -> g ≃ ug -> @@ -738,31 +734,65 @@ Lemma show_equiv_unfold_tensor_stack_helper {C} `{MonoidalCategory C} Proof. intros Hf Hg Huf Hug. rewrite Hf, Hg. - rewrite <- (right_unit uf), <- (left_unit ug), compose_bimap. + rewrite <- (right_unit uf), <- (left_unit ug), tensor_compose. apply compose_compat; easy. Qed. Close Scope Cat_scope. +End HelperLemmas. + -(* Shows the goal f ≃ right_associate f by mirroring the code - path of right_associate with `apply`s. *) -Ltac show_equiv_right_associate f := - let rec show_equiv_right_associate f := +(* Shows the goal f ≃ right_associate_term f by mirroring the code + path of right_associate_term with `apply`s. *) +Ltac show_equiv_right_associate_term f := + let rec show_equiv_right_associate_term f := match f with - | ((?g ∘ ?h) ∘ ?i)%Mor => - (* RHS is `right_associate (g ∘ (h ∘ i))` *) + | ((?g ∘ ?h) ∘ ?i)%Cat => + (* RHS is `right_associate_term (g ∘ (h ∘ i))` *) apply assoc_compat_helper; - show_equiv_right_associate ((g ∘ (h ∘ i))%Mor) - | (?g ∘ ?h)%Mor => (* g shouldn't be a composition *) - (* RHS is `(g ∘ right_associate h)` *) + show_equiv_right_associate_term ((g ∘ (h ∘ i))%Cat) + | (?g ∘ ?h)%Cat => (* g shouldn't be a composition *) + (* RHS is `(g ∘ right_associate_term h)` *) apply compose_compat_right; - show_equiv_right_associate h + show_equiv_right_associate_term h + | _ => + (* RHS is `constr:(f)` *) + reflexivity + end + in show_equiv_right_associate_term f. + +(* Shows the goal f ≃ left_associate_term f by mirroring the code + path of left_associate_term with `apply`s. *) +Ltac show_equiv_left_associate_term f := + let rec show_left f := + match f with + | (?g ∘ (?h ∘ ?i))%Cat => + (* RHS is `left_associate_term ((g ∘ h) ∘ i)%Cat` *) + apply assoc_compat_helper'; + show_left (((g ∘ h) ∘ i)%Cat) + | (?g ∘ ?h)%Cat => (* h shouldn't be a composition *) + (* RHS is `(left_associate_term g ∘ h)` *) + apply compose_cancel_r; + show_left g | _ => (* RHS is `constr:(f)` *) reflexivity end - in show_equiv_right_associate f. + in show_left f. + +Ltac rassoc t := + let H := fresh in + let rt := right_associate_term t in + assert (H: (t ≃ rt)%Cat) by (show_equiv_right_associate_term t); + rewrite H; clear H. + +Ltac lassoc t := + let H := fresh in + let rt := left_associate_term t in + assert (H: (t ≃ rt)%Cat) by (show_equiv_left_associate_term t); + rewrite H; clear H. + (* Shows the goal f ≃ merge_stack_composition f by mirroring the code path of merge_stacked_composition with `apply`s. *) @@ -770,25 +800,25 @@ Ltac show_equiv_merge_stacked_composition gh := let rec show_equiv_merge_stacked_composition gh := match type of gh with @morphism ?C ?cC _ _ => match gh with - @morphism_bimap C _ _ cC _ _ (@tensor C cC ?mC) ?gA ?gB ?hA ?hB ?g ?h => + @mor_tensor C cC ?mC ?gA ?gB ?hA ?hB ?g ?h => lazymatch g with - | (?g0 ∘ ?g1)%Mor => + | (?g0 ∘ ?g1)%Cat => lazymatch h with - | (?h0 ∘ ?h1)%Mor => + | (?h0 ∘ ?h1)%Cat => (* RHS is g0 ⊗ h0 ∘ merge_stacked_composition (g1 ⊗ h1) *) apply stack_compose_distr_compat; - show_equiv_merge_stacked_composition ((mC.(tensor) @@ g1, h1)%Mor) + show_equiv_merge_stacked_composition ((mC.(mor_tensor) g1 h1)%Cat) | _ => (* RHS is g0 ⊗ h0 ∘ merge_stacked_composition (g1 ⊗ id_ hB) *) apply stack_distr_pushout_r_bot_compat; - show_equiv_merge_stacked_composition ((mC.(tensor) @@ g1, (cC.(c_identity) hB))%Mor) + show_equiv_merge_stacked_composition ((mC.(mor_tensor) g1 (cC.(c_identity) hB))%Cat) end | _ => lazymatch h with - | (?h0 ∘ ?h1)%Mor => + | (?h0 ∘ ?h1)%Cat => (* RHS is g0 ⊗ h0 ∘ merge_stacked_composition (id_ gB ⊗ h1) *) apply stack_distr_pushout_r_top_compat; - show_equiv_merge_stacked_composition ((mC.(tensor) @@ (cC.(c_identity) gB), h1)%Mor) + show_equiv_merge_stacked_composition ((mC.(mor_tensor) (cC.(c_identity) gB) h1)%Cat) | _ => (* RHS is g ⊗ h *) reflexivity @@ -806,19 +836,19 @@ Ltac show_equiv_weak_fencepost_form f := | @compose ?C ?cC _ _ _ ?g ?h => let Ng := weak_fencepost g in let Nh := weak_fencepost h in - let res := right_associate (cC.(compose) Ng Nh) in + let res := right_associate_term (cC.(compose) Ng Nh) in apply (compose_compat_trans_helper (cC:=cC) g Ng h Nh res ltac:(show_equiv_weak_fencepost_form g) ltac:(show_equiv_weak_fencepost_form h) - ltac:(show_equiv_right_associate (cC.(compose) Ng Nh))) - | @morphism_bimap ?C _ _ ?cC _ _ (@tensor ?C ?cC ?mC) _ _ _ _ ?g ?h => + ltac:(show_equiv_right_associate_term (cC.(compose) Ng Nh))) + | @mor_tensor ?C ?cC ?mC _ _ _ _ ?g ?h => let Ng := weak_fencepost g in let Nh := weak_fencepost h in - let res := merge_stacked_composition ((mC.(tensor) @@ Ng, Nh)%Mor) in + let res := merge_stacked_composition ((mC.(mor_tensor) Ng Nh)%Cat) in apply (stack_compat_trans_helper (cC:=cC) g Ng h Nh res ltac:(show_equiv_weak_fencepost_form g) ltac:(show_equiv_weak_fencepost_form h) - ltac:(show_equiv_merge_stacked_composition ((mC.(tensor) @@ Ng, Nh))%Mor)) + ltac:(show_equiv_merge_stacked_composition ((mC.(mor_tensor) Ng Nh))%Cat)) | _ => (* f "is const or unsupported" *) (* constr:(f) *) reflexivity @@ -830,13 +860,13 @@ Ltac show_equiv_weak_fencepost_form f := in the given monoidal category mC. *) Ltac stack_comp_id_bot f B mC := let base g := - constr:((mC.(tensor) @@ g, id_ B)%Mor) in + constr:((mC.(mor_tensor) g (id_ B))%Cat) in let rec stack_comp h := match h with - | (?g ∘ ?gs)%Mor => + | (?g ∘ ?gs)%Cat => let stg := base g in let stgs := stack_comp gs in - constr:((stg ∘ stgs)%Mor) + constr:((stg ∘ stgs)%Cat) | ?g => base h end @@ -846,13 +876,13 @@ Ltac stack_comp_id_bot f B mC := in the given monoidal category mC. *) Ltac stack_comp_id_top f A mC := let base g := - constr:((mC.(tensor) @@ id_ A, g)%Mor) in + constr:((mC.(mor_tensor) (id_ A) g)%Cat) in let rec stack_comp h := match h with - | (?g ∘ ?gs)%Mor => + | (?g ∘ ?gs)%Cat => let stg := base g in let stgs := stack_comp gs in - constr:((stg ∘ stgs)%Mor) + constr:((stg ∘ stgs)%Cat) | ?g => base h end @@ -866,12 +896,12 @@ Ltac stack_comp_id_top f A mC := Ltac unfold_tensor_stack f := let rec unfold_tensor_stack f := lazymatch f with - | @morphism_bimap _ _ _ _ _ _ (@tensor _ _ ?mC) ?gA ?gB ?hA ?hB ?g ?h => + | @mor_tensor _ _ ?mC ?gA ?gB ?hA ?hB ?g ?h => let ug := unfold_tensor_stack g in let uh := unfold_tensor_stack h in let sg := stack_comp_id_bot ug hA mC in let sh := stack_comp_id_top uh gB mC in - constr:((sg ∘ sh)%Mor) + constr:((sg ∘ sh)%Cat) | _ => constr:(f) end in unfold_tensor_stack f. @@ -881,26 +911,29 @@ Ltac unfold_tensor_stack_no_id f := let rec unfold_tensor_stack f := lazymatch f with (* TODO: is this case smart to have? *) - | @morphism_bimap _ _ _ _ _ _ (@tensor _ ?cC ?mC) - ?gA ?gA ?hA ?hA (id_ ?gA)%Mor (id_ ?hA)%Mor => - constr:(cC.(c_identity) (mC.(tensor) gA hA)) + (* | @mor_tensor _ ?cC ?mC + ?gA _ ?hA _ (id_ ?gA)%Cat (id_ ?hA)%Cat => + constr:(cC.(c_identity) (mC.(obj_tensor) gA hA)) *) - | @morphism_bimap _ _ _ _ _ _ (@tensor _ _ ?mC) ?gA ?gA ?hA ?hB (id_ ?gA)%Mor ?h => + | @mor_tensor _ ?cC ?mC + ?gA ?gA ?hA ?hB (id_ ?gA)%Cat ?h => let uh := unfold_tensor_stack h in let sh := stack_comp_id_top uh gA mC in constr:(sh) - | @morphism_bimap _ _ _ _ _ _ (@tensor _ _ ?mC) ?gA ?gB ?hA ?hA ?g (id_ ?hA)%Mor => + | @mor_tensor _ ?cC ?mC + ?gA ?gB ?hA ?hA ?g (id_ ?hA)%Cat => let ug := unfold_tensor_stack g in let sg := stack_comp_id_bot ug hA mC in constr:(sg) - | @morphism_bimap _ _ _ _ _ _ (@tensor _ _ ?mC) ?gA ?gB ?hA ?hB ?g ?h => + | @mor_tensor _ ?cC ?mC + ?gA ?gB ?hA ?hB ?g ?h => let ug := unfold_tensor_stack g in let uh := unfold_tensor_stack h in let sg := stack_comp_id_bot ug hA mC in let sh := stack_comp_id_top uh gB mC in - constr:((sg ∘ sh)%Mor) + constr:((sg ∘ sh)%Cat) | _ => constr:(f) end in unfold_tensor_stack f. @@ -911,10 +944,10 @@ Ltac unfold_tensor_stack_no_id f := Ltac strong_fencepost_form_of_weak f := let rec strong_fence f := lazymatch f with - | (?g ∘ ?h)%Mor => + | (?g ∘ ?h)%Cat => let ug := strong_fence g in let uh := strong_fence h in - right_associate (ug ∘ uh)%Mor + right_associate_term (ug ∘ uh)%Cat | _ => unfold_tensor_stack f end @@ -924,10 +957,10 @@ Ltac strong_fencepost_form_of_weak f := Ltac strong_fencepost_form_of_weak_no_id f := let rec strong_fence f := lazymatch f with - | (?g ∘ ?h)%Mor => + | (?g ∘ ?h)%Cat => let ug := strong_fence g in let uh := strong_fence h in - right_associate (ug ∘ uh)%Mor + right_associate_term (ug ∘ uh)%Cat | _ => unfold_tensor_stack_no_id f end @@ -936,10 +969,10 @@ Ltac strong_fencepost_form_of_weak_no_id f := Ltac show_equiv_stack_comp_id_bot f B mC := let base g := - constr:((mC.(tensor) @@ g, id_ B)%Mor) in + constr:((mC.(mor_tensor) g (id_ B))%Cat) in let rec show_stack_comp h := match h with - | (?g ∘ ?gs)%Mor => + | (?g ∘ ?gs)%Cat => (* let stg := base g in let stgs := stack_comp gs in constr:(stg ∘ stgs)%Cat *) @@ -953,10 +986,10 @@ Ltac show_equiv_stack_comp_id_bot f B mC := Ltac show_equiv_stack_comp_id_top f A mC := let base g := - constr:((mC.(tensor) @@ id_ A, g)%Mor) in + constr:((mC.(mor_tensor) (id_ A) g)%Cat) in let rec show_stack_comp h := match h with - | (?g ∘ ?gs)%Mor => + | (?g ∘ ?gs)%Cat => (* let stg := base g in let stgs := stack_comp gs in constr:(stg ∘ stgs)%Cat *) @@ -972,7 +1005,7 @@ Ltac show_equiv_stack_comp_id_top f A mC := Ltac show_equiv_unfold_tensor_stack f := let rec show_unfold f := lazymatch f with - | @morphism_bimap _ _ _ _ _ _ (@tensor _ ?cC ?mC) ?gA ?gB ?hA ?hB ?g ?h => + | @mor_tensor _ ?cC ?mC ?gA ?gB ?hA ?hB ?g ?h => let ug := unfold_tensor_stack g in let uh := unfold_tensor_stack h in let sg := stack_comp_id_bot ug hA mC in @@ -995,28 +1028,28 @@ Ltac show_equiv_unfold_tensor_stack_no_id f := let rec show_unfold f := lazymatch f with (* TODO: is this case smart to have? *) - | @morphism_bimap _ _ _ _ _ _ (@tensor _ _ ?mC) - ?gA ?gA ?hA ?hA (id_ ?gA)%Mor (id_ ?hA)%Mor => + (* | @mor_tensor _ ?cC ?mC + ?gA ?gA ?hA ?hA (id_ ?gA)%Cat (id_ ?hA)%Cat => (* constr:(cC.(c_identity) (mC.(tensor) gA hA)) *) - apply (mC.(tensor).(id_bimap) gA hA) + apply (tensor_id gA hA) *) - | @morphism_bimap _ _ _ _ _ _ (@tensor _ ?cC ?mC) ?gA ?gA ?hA ?hB (id_ ?gA)%Mor ?h => + | @mor_tensor _ ?cC ?mC ?gA ?gA ?hA ?hB (id_ ?gA)%Cat ?h => let uh := unfold_tensor_stack h in let sh := stack_comp_id_top uh gA mC in (* constr:(sh) *) apply (stack_compat_trans_helper (cC.(c_identity) gA) (cC.(c_identity) gA) h uh sh ltac:(reflexivity) ltac:(show_unfold h) - ltac:(show_equiv_stack_comp_id_top h gA mC)) + ltac:(show_equiv_stack_comp_id_top uh gA mC)) - | @morphism_bimap _ _ _ _ _ _ (@tensor _ ?cC ?mC) ?gA ?gB ?hA ?hA ?g (id_ ?hA)%Mor => + | @mor_tensor _ ?cC ?mC ?gA ?gB ?hA ?hA ?g (id_ ?hA)%Cat => let ug := unfold_tensor_stack g in let sg := stack_comp_id_bot ug hA mC in (* constr:(sg) *) apply (stack_compat_trans_helper g ug (cC.(c_identity) hA) (cC.(c_identity) hA) sg ltac:(show_unfold g) ltac:(reflexivity) - ltac:(show_equiv_stack_comp_id_bot g hA mC)) + ltac:(show_equiv_stack_comp_id_bot ug hA mC)) - | @morphism_bimap _ _ _ _ _ _ (@tensor _ _ ?mC) ?gA ?gB ?hA ?hB ?g ?h => + | @mor_tensor _ ?cC ?mC ?gA ?gB ?hA ?hB ?g ?h => let ug := unfold_tensor_stack g in let uh := unfold_tensor_stack h in let sg := stack_comp_id_bot ug hA mC in @@ -1038,30 +1071,35 @@ Ltac show_equiv_unfold_tensor_stack_no_id_debug f := let rec show_unfold f := lazymatch f with (* TODO: is this case smart to have? *) - | @morphism_bimap _ _ _ _ _ _ (@tensor _ _ ?mC) ?gA _ ?hA _ (id_ ?gA)%Mor (id_ ?hA)%Mor => + (* | @mor_tensor _ ?cC ?mC ?gA _ ?hA _ (id_ ?gA)%Cat (id_ ?hA)%Cat => idtac "id id case:"; print_state; (* constr:(cC.(c_identity) (mC.(tensor) gA hA)) *) - apply (mC.(tensor).(id_bimap) gA hA) + apply (tensor_id gA hA) + ; idtac "worked" *) - | @morphism_bimap _ _ _ _ _ _ (@tensor _ ?cC ?mC) ?gA ?gA ?hA ?hB (id_ ?gA)%Mor ?h => + | @mor_tensor _ ?cC ?mC ?gA ?gA ?hA ?hB (id_ ?gA)%Cat ?h => let uh := unfold_tensor_stack h in let sh := stack_comp_id_top uh gA mC in (* constr:(sh) *) idtac "left id case:"; print_state; apply (stack_compat_trans_helper (cC.(c_identity) gA) (cC.(c_identity) gA) h uh sh ltac:(reflexivity) ltac:(show_unfold h) - ltac:(show_equiv_stack_comp_id_top h gA mC)) + ltac:(show_equiv_stack_comp_id_top uh gA mC)) + ; idtac "worked" - | @morphism_bimap _ _ _ _ _ _ (@tensor _ ?cC ?mC) ?gA ?gB ?hA ?hA ?g (id_ ?hA)%Mor => + | @mor_tensor _ ?cC ?mC ?gA ?gB ?hA ?hA ?g (id_ ?hA)%Cat => let ug := unfold_tensor_stack g in let sg := stack_comp_id_bot ug hA mC in (* constr:(sg) *) idtac "right id case:"; print_state; apply (stack_compat_trans_helper g ug (cC.(c_identity) hA) (cC.(c_identity) hA) sg ltac:(show_unfold g) ltac:(reflexivity) - ltac:(show_equiv_stack_comp_id_bot g hA mC)) + ltac:(show_equiv_stack_comp_id_bot ug hA mC)) + (* ); idtac "applied" g hA mC; print_state; show_equiv_stack_comp_id_bot ug hA mC; + print_state + ; idtac "worked" *) - | @morphism_bimap _ _ _ _ _ _ (@tensor _ _ ?mC) ?gA ?gB ?hA ?hB ?g ?h => + | @mor_tensor _ ?cC ?mC ?gA ?gB ?hA ?hB ?g ?h => let ug := unfold_tensor_stack g in let uh := unfold_tensor_stack h in let sg := stack_comp_id_bot ug hA mC in let sh := stack_comp_id_top uh gB mC in @@ -1071,6 +1109,7 @@ Ltac show_equiv_unfold_tensor_stack_no_id_debug f := ltac:(show_equiv_stack_comp_id_bot ug hA mC) ltac:(show_equiv_stack_comp_id_top uh gB mC) ) + ; idtac "worked" | _ => (* constr:(f) *) reflexivity end in show_unfold f. @@ -1080,16 +1119,16 @@ Ltac show_equiv_strong_fencepost_form_of_weak f := let strong_fence := strong_fencepost_form_of_weak in let rec show_strong_fence f := lazymatch f with - | (?g ∘ ?h)%Mor => + | (?g ∘ ?h)%Cat => let ug := strong_fence g in let uh := strong_fence h in - let rassoc := right_associate (ug ∘ uh)%Mor in - (* right_associate (ug ∘ uh)%Cat *) + let rassoc := right_associate_term (ug ∘ uh)%Cat in + (* right_associate_term (ug ∘ uh)%Cat *) apply (compose_compat_trans_helper g ug h uh rassoc ltac:(show_strong_fence g) ltac:(show_strong_fence h) - ltac:(show_equiv_right_associate (ug ∘ uh)%Mor) + ltac:(show_equiv_right_associate_term (ug ∘ uh)%Cat) ) | _ => (* unfold_tensor_stack f *) @@ -1102,16 +1141,16 @@ Ltac show_equiv_strong_fencepost_form_of_weak_no_id f := let strong_fence := strong_fencepost_form_of_weak_no_id in let rec show_strong_fence f := lazymatch f with - | (?g ∘ ?h)%Mor => + | (?g ∘ ?h)%Cat => let ug := strong_fence g in let uh := strong_fence h in - let rassoc := right_associate (ug ∘ uh)%Mor in - (* right_associate (ug ∘ uh)%Cat *) + let rassoc := right_associate_term (ug ∘ uh)%Cat in + (* right_associate_term (ug ∘ uh)%Cat *) apply (compose_compat_trans_helper g ug h uh rassoc ltac:(show_strong_fence g) ltac:(show_strong_fence h) - ltac:(show_equiv_right_associate (ug ∘ uh)%Mor) + ltac:(show_equiv_right_associate_term (ug ∘ uh)%Cat) ) | _ => (* unfold_tensor_stack f *) @@ -1123,16 +1162,16 @@ Ltac show_equiv_strong_fencepost_form_of_weak_no_id_debug f := let strong_fence := strong_fencepost_form_of_weak_no_id in let rec show_strong_fence f := lazymatch f with - | (?g ∘ ?h)%Mor => + | (?g ∘ ?h)%Cat => let ug := strong_fence g in let uh := strong_fence h in - let rassoc := right_associate (ug ∘ uh)%Mor in - (* right_associate (ug ∘ uh)%Cat *) + let rassoc := right_associate_term (ug ∘ uh)%Cat in + (* right_associate_term (ug ∘ uh)%Cat *) apply (compose_compat_trans_helper g ug h uh rassoc ltac:(show_strong_fence g) ltac:(show_strong_fence h) - ltac:(show_equiv_right_associate (ug ∘ uh)%Mor) + ltac:(show_equiv_right_associate_term (ug ∘ uh)%Cat) ) | _ => (* unfold_tensor_stack f *) @@ -1180,54 +1219,827 @@ Ltac strong_fencepost_no_id_debug f := setoid_rewrite H; clear H. + +Ltac right_associate_term' f := + let rec rassoc f := + lazymatch f with + | ((?g ∘ ?h) ∘ ?i)%Cat => rassoc (g ∘ (h ∘ i))%Cat + | (?g ∘ ?h)%Cat => (* g shouldn't be a composition *) + let RAh := rassoc h in + constr:((g ∘ RAh)%Cat) + | @mor_tensor _ ?cC ?mC _ _ _ _ ?g ?h => + let RAg := rassoc g in let RAh := rassoc h in + constr:(mC.(mor_tensor) RAg RAh) + | _ => constr:(f) + end + in rassoc f. + +Ltac show_equiv_right_associate_term' term := + let rec show_equiv_right_associate_term f := + try easy; + lazymatch f with + | ((?g ∘ ?h) ∘ ?i)%Cat => + (* RHS is `right_associate_term (g ∘ (h ∘ i))` *) + apply assoc_compat_helper; + show_equiv_right_associate_term ((g ∘ (h ∘ i))%Cat) + | (?g ∘ ?h)%Cat => (* g shouldn't be a composition *) + (* RHS is `(g ∘ right_associate_term h)` *) + apply compose_compat_right; + show_equiv_right_associate_term h + | @mor_tensor _ ?cC ?mC _ _ _ _ ?g ?h => + apply (tensor_compat); + [ltac:(show_equiv_right_associate_term g) | + ltac:(show_equiv_right_associate_term h)] + | _ => + (* RHS is `constr:(f)` *) + reflexivity + end + in show_equiv_right_associate_term term. + +Ltac rassoc' t := + let H := fresh in + let rt := right_associate_term' t in + assert (H: (t ≃ rt)%Cat) by (show_equiv_right_associate_term' t); + rewrite H; clear H. + + +Ltac partnered_in_RA_term_nofail t s term := + let rec partnered t s term := + match term with + | (?g ∘ (?h ∘ ?i))%Cat => + let _ := lazymatch goal with _ => + unify t g; unify s h end + in constr:((g ∘ h ∘ i)%Cat) + | (?g ∘ (?h ∘ ?i))%Cat => + let subpart := constr:((h ∘ i)%Cat) in + let ptnered := partnered t s subpart in + constr:((g ∘ ptnered)%Cat) + | _ => constr:(term) + end + in partnered t s term. + +Ltac partnered_in_term_nofail t s term := + let raterm := right_associate_term' term in + partnered_in_RA_term_nofail t s raterm. + + + +Ltac partnered_in_RA_term t s term := + let rec partnered t s term := + match term with + | (?g ∘ (?h ∘ ?i))%Cat => + let _ := lazymatch goal with _ => + unify t g; unify s h end in + constr:((g ∘ h ∘ i)%Cat) + | (?g ∘ (?h ∘ ?i))%Cat => + let subpart := constr:((h ∘ i)%Cat) in + let ptnered := partnered t s subpart in + constr:((g ∘ ptnered)%Cat) + | (?g ∘ ?h)%Cat => + let _ := lazymatch goal with _ => + unify t g; unify s h end in + constr:((g ∘ h)%Cat) + | @mor_tensor _ ?cC ?mC _ _ _ _ ?f ?g => + let out := match goal with + | _ => let mf := partnered f in + constr:(mC.(mor_tensor) mf g) + | _ => let mg := partnered g in + constr:(mC.(mor_tensor) f mg) + end + in constr:(out) + end + in partnered t s term. + +Ltac partnered_in_term t s term := + let raterm := right_associate_term' term in + partnered_in_RA_term t s raterm. + +Ltac try_partnered_in_term t s term := + let out := match goal with + | |- _ => + let raterm := right_associate_term' term in + partnered_in_RA_term t s raterm + | |- _ => constr:(term) + end in constr:(out). + +Ltac show_equiv_partnered_in_RA_term t s term := + let rec show_part t s term := + match term with + | (?g ∘ (?h ∘ ?i))%Cat => + let _ := lazymatch goal with _ => + unify t g; unify s h end in + (* constr:((g ∘ h ∘ i)%Cat) *) + symmetry; apply assoc + | (?g ∘ (?h ∘ ?i))%Cat => + let subpart := constr:((h ∘ i)%Cat) in + (* constr:((g ∘ ptnered)%Cat) *) + apply compose_cancel_l; + show_part t s subpart + | (?g ∘ ?h)%Cat => + (* let _ := lazymatch goal with _ => + unify t g; unify s h end in + constr:((g ∘ h)%Cat) *) + reflexivity + | @mor_tensor _ ?cC ?mC _ _ _ _ ?f ?g => + apply (tensor_compat); + [ first [reflexivity | show_part f] | first [reflexivity | show_part g] ] + end + in show_part t s term. + +Ltac show_equiv_partnered_in_term t s term := + let raterm := right_associate_term' term in + transitivity raterm; + [ show_equiv_right_associate_term' term + | show_equiv_partnered_in_RA_term t s raterm ]. + +Ltac partner_in_term t s term := + let ptnered := partnered_in_term t s term in + let H := fresh in + assert (H : (term ≃ ptnered)%Cat) + by (show_equiv_partnered_in_term t s term); + setoid_rewrite H; + clear H. + + +Ltac show_equiv_partnered_in_RA_term_debug t s term := + let rec show_part t s term := + match term with + | (?g ∘ (?h ∘ ?i))%Cat => + let _ := lazymatch goal with _ => + unify t g; unify s h end in + (* constr:((g ∘ h ∘ i)%Cat) *) + idtac "unified; associating"; + print_state; + symmetry; apply assoc + | (?g ∘ (?h ∘ ?i))%Cat => + idtac "starting cancelling"; + let subpart := constr:((h ∘ i)%Cat) in + (* constr:((g ∘ ptnered)%Cat) *) + idtac "cancelling" g "down to" subpart; + apply compose_cancel_l; + print_state; + show_part t s subpart + | (?g ∘ ?h)%Cat => + (* let _ := lazymatch goal with _ => + unify t g; unify s h end in + constr:((g ∘ h)%Cat) *) + idtac "ended; reflexivity"; + print_state; + reflexivity + | @mor_tensor _ ?cC ?mC _ _ _ _ ?f ?g => + apply (tensor_compat); + [ first [reflexivity | show_part f] | first [reflexivity | show_part g] ] + end + in show_part t s term. + +Ltac show_equiv_partnered_in_term_debug t s term := + let raterm := right_associate_term' term in + transitivity raterm; + [ + (* idtac "rassoc:"; + print_state; *) + show_equiv_right_associate_term' term + | + idtac "partnr:"; + print_state; + show_equiv_partnered_in_RA_term_debug t s raterm ]. + +Ltac partner_in_term_debug t s term := + let ptnered := partnered_in_term t s term in + let H := fresh in + assert (H : (term ≃ ptnered)%Cat) + by (show_equiv_partnered_in_term_debug t s term); + setoid_rewrite H; + clear H. + + +Ltac gen_partnered_in_RA_term test term := + let rec partnered test term := + match term with + | (?g ∘ (?h ∘ ?i))%Cat => + let _ := lazymatch goal with _ => + test g h end in + constr:((g ∘ h ∘ i)%Cat) + | (?g ∘ (?h ∘ ?i))%Cat => + let subpart := constr:((h ∘ i)%Cat) in + let ptnered := partnered test subpart in + constr:((g ∘ ptnered)%Cat) + | (?g ∘ ?h)%Cat => + let _ := lazymatch goal with _ => + test g h end in + constr:((g ∘ h)%Cat) + | @mor_tensor _ ?cC ?mC _ _ _ _ ?f ?g => + let out := match goal with + | _ => let mf := partnered test f in + constr:(mC.(mor_tensor) mf g) + | _ => let mg := partnered test g in + constr:(mC.(mor_tensor) f mg) + end + in constr:(out) + end + in partnered test term. + +Ltac gen_partnered_in_term test term := + let raterm := right_associate_term' term in + gen_partnered_in_RA_term test raterm. + +Ltac try_gen_partnered_in_term test term := + let out := match goal with + | |- _ => + let raterm := right_associate_term' term in + gen_partnered_in_RA_term test raterm + | |- _ => constr:(term) + end in constr:(out). + +Ltac show_equiv_gen_partnered_in_RA_term test term := + let rec show_part test term := + match term with + | (?g ∘ (?h ∘ ?i))%Cat => + let _ := lazymatch goal with _ => + test g h end in + (* constr:((g ∘ h ∘ i)%Cat) *) + symmetry; apply assoc + | (?g ∘ (?h ∘ ?i))%Cat => + let subpart := constr:((h ∘ i)%Cat) in + (* constr:((g ∘ ptnered)%Cat) *) + apply compose_cancel_l; + show_part test subpart + | (?g ∘ ?h)%Cat => + (* let _ := lazymatch goal with _ => + unify t g; unify s h end in + constr:((g ∘ h)%Cat) *) + reflexivity + | @mor_tensor _ ?cC ?mC _ _ _ _ ?f ?g => + apply (tensor_compat); + [ first [reflexivity | show_part test f] | first [reflexivity | show_part test g] ] + end + in show_part test term. + +Ltac show_equiv_gen_partnered_in_term test term := + let raterm := right_associate_term' term in + transitivity raterm; + [ show_equiv_right_associate_term' term + | show_equiv_gen_partnered_in_RA_term test raterm ]. + +Ltac gen_partner_in_term test term := + let ptnered := gen_partnered_in_term test term in + let H := fresh in + assert (H : (term ≃ ptnered)%Cat) + by (show_equiv_gen_partnered_in_term test term); + setoid_rewrite H; + clear H. + +Ltac gen_partner_in_term_then test term tac := + let ptnered := gen_partnered_in_term test term in + let H := fresh in + assert (H : (term ≃ ptnered)%Cat) + by (show_equiv_gen_partnered_in_term test term); + setoid_rewrite H; + clear H; + tac ptnered. + +Ltac rep_gen_partner_in_term test term tac := + let ptnered := gen_partnered_in_term test term in + let H := fresh in + assert (H : (term ≃ ptnered)%Cat) + by (show_equiv_gen_partnered_in_term test term); + setoid_rewrite H; + clear H; + first [rep_gen_partner_in_term test term | tac term]. + +Ltac test_iso_inv_l t s := + let comp := constr:((t ∘ s)%Cat) in + lazymatch comp with + | (reverse ?f ∘ forward ?f)%Cat => idtac + end. + +Ltac test_iso_inv_r t s := + let comp := constr:((t ∘ s)%Cat) in + lazymatch comp with + | (forward ?f ∘ reverse ?f)%Cat => idtac + end. + +Ltac test_iso_inv_lr t s := + let comp := constr:((t ∘ s)%Cat) in + lazymatch comp with + | (reverse ?f ∘ forward ?f)%Cat => idtac + | (forward ?f ∘ reverse ?f)%Cat => idtac + end. + + +Ltac cancel_ids_term term := + let rec clean term := + lazymatch term with + | (compose ?cC (id_ _) (id_ ?A))%Cat => + constr:(cC.(c_identity) A) + | (id_ _ ∘ ?g)%Cat => + let r := clean g in constr:(r) + | (?g ∘ id_ _)%Cat => + let r := clean g in constr:(r) + | (?f ∘ ?g)%Cat => + let fr := clean f in let gr := clean g in + constr:((fr ∘ gr)%Cat) + | @mor_tensor ?C ?cC ?mC _ _ _ _ ?f ?g => + let fr := clean f in let gr := clean g in + constr:((mC.(mor_tensor) fr gr)%Cat) + | _ => constr:(term) + end + in clean term. + +Ltac show_equiv_cancel_ids_term term := + let rec show_clean term := + lazymatch term with + | (id_ _ ∘ id_ ?A)%Cat => + (* constr:((id_ A)%Cat) *) + apply left_unit + | (id_ _ ∘ ?g)%Cat => + (* let r := clean g in constr:(r) *) + transitivity g; + [ apply left_unit | show_clean g ] + | (?g ∘ id_ _)%Cat => + (* let r := clean g in constr:(r) *) + transitivity g; + [ apply right_unit | show_clean g ] + | (id_ _ ∘ ?g)%Cat => + (* let r := clean g in constr:(r) *) + transitivity g; + [ apply left_unit | show_clean g ] + (* constr:((id_ A)%Cat) *) + | (?f ∘ ?g)%Cat => + (* let fr := clean f in let gr := clean g in + constr:((fr ∘ gr)%Cat) *) + apply compose_compat; + [ show_clean f | show_clean g ] + | @mor_tensor ?C ?cC ?mC _ _ _ _ ?f ?g => + (* let fr := clean f in let gr := clean g in + constr:((mC.(mor_tensor) fr gr)%Cat) *) + apply tensor_compat; + [ show_clean f | show_clean g ] + | _ => reflexivity + end + in show_clean term. + +Ltac cancel_ids_in term := + let c := cancel_ids_term term in + let H := fresh in + assert (H : (term ≃ c)%Cat) by (show_equiv_cancel_ids_term term); + setoid_rewrite H; + clear H. + +Ltac cancel_all_ids_term term := + let clean := cancel_ids_term in + let rec deep_clean term := + let r := clean term in + let out := match goal with + | _ => let _ := match goal with _ => unify r term end in + constr:(term) + | _ => let rr := deep_clean r in constr:(rr) + end in constr:(out) + in deep_clean term. + + +(* This ends up being faster than the functionally-equivlant + "repeat (cancel_ids_in term)" + (found ~20% faster on small example with 2 rounds needed)*) +Ltac cancel_all_ids term := + let c := cancel_ids_term term in + tryif unify c term then idtac else + let H := fresh in + assert (H : (term ≃ c)%Cat) by (show_equiv_cancel_ids_term term); + setoid_rewrite H; + clear H; + cancel_all_ids c. + +Ltac cancel_liso term := + gen_partner_in_term test_iso_inv_l term; + rewrite ?iso_inv_l. + +Ltac cancel_lisos term := + repeat cancel_liso term. + +(* These don't work because term changes before cancel_all_ids. + This could be fixed with a highly-modified gen_partner, but + for now the cases we care about (LHS, RHS) can be done specially. *) +(* Ltac cancel_lisos' term := + cancel_lisos term; cancel_all_ids term. *) + +Ltac cancel_riso term := + gen_partner_in_term test_iso_inv_r term; + rewrite ?iso_inv_r. + +Ltac cancel_risos term := + repeat cancel_riso term. + +(* Ltac cancel_risos' term := + cancel_risos term; cancel_all_ids term. *) + +Ltac cancel_lriso term := + gen_partner_in_term test_iso_inv_lr term; + rewrite ?iso_inv_l, ?iso_inv_r. + +Ltac cancel_lrisos term := + repeat cancel_lriso term. + +(* Ltac cancel_lrisos' term := + cancel_lrisos term; cancel_all_ids term. *) + + + + +Ltac lassoc_n_term_rec n base t := + match n with + | O => constr:((base ∘ t)%Cat) + | S ?n' => match t with + | (?g ∘ ?h)%Cat => + let base' := constr:((base ∘ g)%Cat) in + (* let rest := constr:((h ∘ i)%Cat) in *) + lassoc_n_term_rec n' base' h + | _ => constr:((base ∘ t)%Cat) + end + end. + +Ltac lassoc_n_term n t := + match n with + | O => constr:(t) + | S ?n' => match t with + | (?g ∘ ?h)%Cat => + lassoc_n_term_rec n' g h + end + end. + + +Ltac lassoc_n_term_rec_debug n base t := + match n with + | O => constr:((base ∘ t)%Cat) + | S ?n' => + let _ := match goal with _ => idtac "at" n' base t end in + match t with + | (?g ∘ ?h)%Cat => + let base' := constr:((base ∘ g)%Cat) in + (* let rest := constr:((h ∘ i)%Cat) in *) + lassoc_n_term_rec n' base' h + | _ => constr:((base ∘ t)%Cat) + end + end. + +Ltac lassoc_n_term_debug n t := + match n with + | O => constr:(t) + | S ?n' => match t with + | (?g ∘ ?h)%Cat => + lassoc_n_term_rec_debug n' g h + end + end. + +(* Shows base ∘ t ≃ lassoc_n_term_rec n base t + (relies on there being enough compositions of t, of course)*) +Ltac show_equiv_lassoc_n_term_rec n base t := + match n with + | O => (* constr:((base ∘ t)%Cat) *) + reflexivity + | S ?n' => match t with + | (?g ∘ ?h)%Cat => + let base' := constr:((base ∘ g)%Cat) in + (* let rest := constr:((h ∘ i)%Cat) in *) + (* lassoc_n_term_rec n' base' h *) + transitivity ((base ∘ g ∘ h)%Cat); + [ symmetry; apply assoc | + show_equiv_lassoc_n_term_rec n' base' h] + | _ => reflexivity + end + end. + +Ltac show_equiv_lassoc_n_term n t := + match n with + | O => (* constr:(t) *) + reflexivity + | S ?n' => match t with + | (?g ∘ ?h)%Cat => + show_equiv_lassoc_n_term_rec n g h + end + end. + + + +Ltac __next_Sn_of_comp n t := + lazymatch t with + | (?g ∘ ?h)%Cat => + match n with + | O => constr:(g) + | S ?n' => let rest := __next_Sn_of_comp n' h in + constr:((g ∘ rest)%Cat) + end + | ?g => match n with O => constr:(g) end + end. + +Ltac __next_n_of_comp n t := + match n with + | S ?n' => __next_Sn_of_comp n' t + end. + +Ltac n_partnered_in_RA_term pat n term := + let rec n_partnered term := + match term with + | _ => let nextn := __next_n_of_comp n term in + let _ := match goal with _ => unify nextn pat end in + lassoc_n_term n term + | (?g ∘ ?h)%Cat => let npart := n_partnered h in + constr:((g ∘ npart)%Cat) + | mor_tensor ?mC ?f ?g => + let out := + match goal with + | _ => let mf := n_partnered f in + constr:((mC.(mor_tensor) mf g)) + | _ => let mg := n_partnered g in + constr:((mC.(mor_tensor) f mg)) + end in constr:(out) + end in + n_partnered term. + +Ltac n_partnered_in_term pat n term := + let raterm := right_associate_term' term in + n_partnered_in_RA_term pat n raterm. + +Ltac n_partnered_in_RA_term_debug pat n term := + let rec n_partnered term := + match term with + | _ => let nextn := __next_n_of_comp n term in + let _ := match goal with _ => + idtac "trying" nextn; + unify nextn pat; + idtac "unified" end in + lassoc_n_term_debug n term + | (?g ∘ ?h)%Cat => + let _ := match goal with _ => idtac "moving into" h end in + let npart := n_partnered h in + constr:((g ∘ npart)%Cat) + | mor_tensor ?mC ?f ?g => + let _ := match goal with _ => idtac "searching tensor" end in + let out := + match goal with + | _ => let mf := n_partnered f in + let _ := match goal with _ => idtac "found in top" f "=~=>" mf end in + constr:((mC.(mor_tensor) mf g)) + | _ => let mg := n_partnered g in + let _ := match goal with _ => idtac "found in bot" g "=~=>" mg end in + constr:((mC.(mor_tensor) f mg)) + end in constr:(out) + end in + n_partnered term. + +Ltac n_partnered_in_term_debug pat n term := + let raterm := right_associate_term' term in + n_partnered_in_RA_term_debug pat n raterm. + +Ltac show_equiv_n_partnered_in_RA_term pat n term := + let rec show_n_partnered term := + match term with + | _ => + show_equiv_lassoc_n_term n term + (* let nextn := __next_n_of_comp n term in + let _ := match goal with _ => unify nextn pat end in + lassoc_n_term n term *) + | (?g ∘ ?h)%Cat => + apply compose_cancel_l; + show_n_partnered h + | mor_tensor ?mC ?f ?g => + first [ + apply tensor_cancel_r; show_n_partnered f + | apply tensor_cancel_l; show_n_partnered g ] + end in + show_n_partnered term. + +Ltac show_equiv_n_partnered_in_term pat n term := + let raterm := right_associate_term' term in + transitivity raterm; + [ show_equiv_right_associate_term' term + | show_equiv_n_partnered_in_RA_term pat n raterm ]. + +Ltac n_partner_in_term pat n term := + let npart := n_partnered_in_term pat n term in + tryif unify npart term then idtac else + let H := fresh in + assert (H : (term ≃ npart)%Cat) by + (show_equiv_n_partnered_in_term pat n term); + setoid_rewrite H; + clear H. + +Ltac n_partner_in_term_debug pat n term := + let npart := n_partnered_in_term_debug pat n term in + tryif unify npart term then idtac else + let H := fresh in + assert (H : (term ≃ npart)%Cat) by + (show_equiv_n_partnered_in_term pat n term); + setoid_rewrite H; + clear H. + +Ltac __count_comps t := + lazymatch t with + | (?g ∘ ?h)%Cat => let n' := __count_comps h in constr:(S n') + | _ => constr:(O) + end. + +Ltac __count_comp_terms t := + lazymatch t with + | (?g ∘ ?h)%Cat => let n' := __count_comp_terms h in constr:(S n') + | _ => constr:(S O) + end. + +(* Section on handy versions of these tactics: *) + +Ltac apply_to_LHS tac := + lazymatch goal with |- (?LHS ≃ ?RHS)%Cat => tac LHS end. + +Ltac apply_to_RHS tac := + lazymatch goal with |- (?LHS ≃ ?RHS)%Cat => tac RHS end. + +Ltac apply_to_LRHS tac := + lazymatch goal with |- (?LHS ≃ ?RHS)%Cat => + (try tac LHS); (try tac RHS) end. + + +Ltac rassoc_LHS := apply_to_LHS rassoc. + (* match goal with |- (?LHS ≃ ?RHS)%Cat => rassoc LHS end. *) +Ltac rassoc_RHS := apply_to_RHS rassoc. + (* match goal with |- (?LHS ≃ ?RHS)%Cat => rassoc RHS end. *) +Ltac rassoc_LRHS := apply_to_LRHS rassoc. + +Ltac rassoc'_LHS := apply_to_LHS rassoc'. +Ltac rassoc'_RHS := apply_to_RHS rassoc'. +Ltac rassoc'_LRHS := apply_to_LRHS rassoc'. + +Ltac lassoc_LHS := apply_to_LHS lassoc. + (* match goal with |- (?LHS ≃ ?RHS)%Cat => lassoc LHS end. *) +Ltac lassoc_RHS := apply_to_RHS lassoc. + (* match goal with |- (?LHS ≃ ?RHS)%Cat => lassoc RHS end. *) +Ltac lassoc_LRHS := apply_to_LRHS lassoc. + + +Ltac partner_LHS t s := + let func := partner_in_term t s in + apply_to_LHS func. + +Ltac partner_RHS t s := + let func := partner_in_term t s in + apply_to_RHS func. + +Ltac partner_LRHS t s := + let func := partner_in_term t s in + apply_to_LRHS func. + +Ltac gen_partner_LHS test := + let func := gen_partner_in_term test in + apply_to_LHS func. + +Ltac gen_partner_RHS test := + let func := gen_partner_in_term test in + apply_to_RHS func. + +Ltac gen_partner_LRHS test := + let func := gen_partner_in_term test in + apply_to_LRHS func. + +Ltac cancel_id_LHS := apply_to_LHS cancel_ids_in. +Ltac cancel_id_RHS := apply_to_RHS cancel_ids_in. +Ltac cancel_id_LRHS := apply_to_LRHS cancel_ids_in. + +Ltac cancel_ids_LHS := apply_to_LHS cancel_all_ids. +Ltac cancel_ids_RHS := apply_to_RHS cancel_all_ids. +Ltac cancel_ids_LRHS := apply_to_LRHS cancel_all_ids. + +Ltac cancel_lisos_LHS := apply_to_LHS cancel_lisos; cancel_ids_LHS. +Ltac cancel_lisos_RHS := apply_to_RHS cancel_lisos; cancel_ids_RHS. +Ltac cancel_lisos_LRHS := apply_to_LRHS cancel_lisos; cancel_ids_LRHS. + +Ltac cancel_risos_LHS := apply_to_LHS cancel_risos; cancel_ids_LHS. +Ltac cancel_risos_RHS := apply_to_RHS cancel_risos; cancel_ids_RHS. +Ltac cancel_risos_LRHS := apply_to_LRHS cancel_risos; cancel_ids_LRHS. + +Ltac cancel_lrisos_LHS := apply_to_LHS cancel_lrisos; cancel_ids_LHS. +Ltac cancel_lrisos_RHS := apply_to_RHS cancel_lrisos; cancel_ids_RHS. +Ltac cancel_lrisos_LRHS := apply_to_LRHS cancel_lrisos; cancel_ids_LRHS. + +Ltac weak_fencepost_LHS := apply_to_LHS weak_fencepost. +Ltac weak_fencepost_RHS := apply_to_RHS weak_fencepost. +Ltac weak_fencepost_LRHS := apply_to_LRHS weak_fencepost. + +Ltac strong_fencepost_LHS := apply_to_LHS strong_fencepost_no_id. +Ltac strong_fencepost_RHS := apply_to_RHS strong_fencepost_no_id. +Ltac strong_fencepost_LRHS := apply_to_LRHS strong_fencepost_no_id. + + +Ltac cancel_ids := cancel_ids_LRHS. +Ltac cancel_isos := cancel_lrisos_LRHS. + +Ltac cat_cleanup := repeat (cancel_isos; cancel_ids). + +Ltac cat_easy := cat_cleanup; rassoc_LRHS; easy || rewrite !tensor_id; easy. + + + +Tactic Notation "LHS" tactic(tac) := apply_to_LHS tac. +Tactic Notation "RHS" tactic(tac) := apply_to_RHS tac. +Tactic Notation "LRHS" tactic(tac) := apply_to_LRHS tac. + +Tactic Notation "partners_rw" open_constr(lem) "within" constr(term) := + let e := fresh in let e' := fresh in + epose proof @lem as e; + repeat (rename e into e'; epose proof (e' _) as e; clear e'); + match type of e with + | (?g ≃ _)%Cat => + let rg := right_associate_term g in + let n := __count_comp_terms rg in + n_partner_in_term rg n term; + let lg := left_associate_term g in + tryif unify g lg then idtac else ( + let H := fresh in + print_state; + assert (H : (g ≃ lg)%Cat) by (show_equiv_left_associate_term g); + try setoid_rewrite H in e; + clear H); + setoid_rewrite e; + clear e + end. + +Tactic Notation "partners_rw" open_constr(lem) := + match goal with + |- (?LHS ≃ ?RHS)%Cat => + first [ + partners_rw lem within LHS + | partners_rw lem within RHS ] + end. + +Tactic Notation "partners_rw_to_Cat" open_constr(lem) "within" constr(term) := + let e := fresh in let e' := fresh in + epose proof @lem as e; + to_Cat_in e; + repeat (rename e into e'; epose proof (e' _) as e; clear e'); + match type of e with + | (?g ≃ _)%Cat => + let rg := right_associate_term g in + let n := __count_comp_terms rg in + n_partner_in_term rg n term; + let lg := left_associate_term g in + tryif unify g lg then idtac else ( + let H := fresh in + print_state; + assert (H : (g ≃ lg)%Cat) by (show_equiv_left_associate_term g); + try setoid_rewrite H in e; + clear H); + setoid_rewrite e; + clear e + end. + + +Tactic Notation "partners_rw_to_Cat" open_constr(lem) := + match goal with + |- (?LHS ≃ ?RHS)%Cat => + first [ + partners_rw_to_Cat lem within LHS + | partners_rw_to_Cat lem within RHS ] + end. + + Section Testing. Local Open Scope Cat_scope. Variables (C : Type) (cC cC' cC'' : Category C) - (mC0 mC1 : @MonoidalCategory C cC) - (mC0' mC1' : @MonoidalCategory C cC') - (mC0'' mC1'' : @MonoidalCategory C cC'') + (cCh : CategoryCoherence cC) (cC'h : CategoryCoherence cC') (cC''h : CategoryCoherence cC'') + (mC0 mC1 : @MonoidalCategory C cC) (mC0' mC1' : @MonoidalCategory C cC') (mC0'' mC1'' : @MonoidalCategory C cC'') + (mC0h : MonoidalCategoryCoherence mC0) (mC0'h : MonoidalCategoryCoherence mC0') (mC0''h : MonoidalCategoryCoherence mC0'') + (mC1h : MonoidalCategoryCoherence mC1) (mC1'h : MonoidalCategoryCoherence mC1') (mC1''h : MonoidalCategoryCoherence mC1'') (A B M N : C) - (f f0 : cC.(morphism) A B) - (g g0 : cC.(morphism) B M) - (h h0 : cC.(morphism) A M) - (i i0 : cC.(morphism) M N) - (j j0 : cC.(morphism) B M) - (k k0 : cC.(morphism) A M) - (f' f0' : cC'.(morphism) A B) - (g' g0' : cC'.(morphism) B M) - (h' h0' : cC'.(morphism) A M) - (i' i0' : cC'.(morphism) M N) - (j' j0' : cC'.(morphism) B M) - (k' k0' : cC'.(morphism) A M) - (f'' f0'' : cC''.(morphism) A B) - (g'' g0'' : cC''.(morphism) B M) - (h'' h0'' : cC''.(morphism) A M) - (i'' i0'' : cC''.(morphism) M N) - (j'' j0'' : cC''.(morphism) B M) - (k'' k0'' : cC''.(morphism) A M). + (f f0 : cC.(morphism) A B) (g g0 : cC.(morphism) B M) (h h0 : cC.(morphism) A M) (i i0 : cC.(morphism) M N) + (* (j j0 : cC.(morphism) B M) (k k0 : cC.(morphism) A M) (f' f0' : cC'.(morphism) A B) (g' g0' : cC'.(morphism) B M) + (h' h0' : cC'.(morphism) A M) (i' i0' : cC'.(morphism) M N) (j' j0' : cC'.(morphism) B M) (k' k0' : cC'.(morphism) A M) + (f'' f0'' : cC''.(morphism) A B) (g'' g0'' : cC''.(morphism) B M) (h'' h0'' : cC''.(morphism) A M) (i'' i0'' : cC''.(morphism) M N) + (j'' j0'' : cC''.(morphism) B M) (k'' k0'' : cC''.(morphism) A M) *) + . (* Goal True. *) -Existing Instance cC. -Existing Instance cC'. -Existing Instance cC''. +Existing Instance cC. Existing Instance cC'. Existing Instance cC''. Existing Instance mC0. Existing Instance mC1. Existing Instance mC0'. Existing Instance mC1'. Existing Instance mC0''. Existing Instance mC1''. -Lemma test_weak_fencepost : forall {C : Type} - `{Cat : Category C} `{MonCat : MonoidalCategory C} + + + +Lemma test_weak_fencepost : forall {a b m n o} (f : a ~> b) (g : m ~> n) (h : n ~> o), f ⊗ (g ∘ h) ≃ f ⊗ g ∘ (id_ b ⊗ h). Proof. intros. match goal with |- ?T ≃ _ => weak_fencepost T + (* let wf := weak_fencepost_form T in + let H := fresh in + assert (H : T ≃ wf) by show_equiv_weak_fencepost_form T *) + (* setoid_rewrite H *) end. easy. Qed. -Lemma test_strong_fencepost : forall {C : Type} - `{Cat : Category C} `{MonCat : MonoidalCategory C} +Lemma test_strong_fencepost : forall {a b m n o} (f : a ~> b) (g : m ~> n) (h : n ~> o), f ⊗ (g ∘ h) ≃ f ⊗ g ∘ (id_ b ⊗ h). Proof. @@ -1238,8 +2050,7 @@ Proof. easy. Qed. -Lemma test_strong_fencepost_no_id_1 : forall {C : Type} - `{Cat : Category C} `{MonCat : MonoidalCategory C} +Lemma test_strong_fencepost_no_id_1 : forall {a b m n o} (f : a ~> b) (g : m ~> n) (h : n ~> o), f ⊗ (g ∘ h) ≃ f ⊗ g ∘ (id_ b ⊗ h). Proof. @@ -1250,23 +2061,107 @@ Proof. easy. Qed. -(* Lemma test_strong_fencepost_no_id_2 : forall {C : Type} - `{Cat : Category C} `{MonCat : MonoidalCategory C} +(* Require Import Program.Tactics. *) + + + +Lemma test_strong_fencepost_no_id_2' : forall {a b m n o} (f : a ~> b) (g : m ~> n) (h : n ~> o), - f ⊗ (g ∘ h ∘ id_ _) ⊗ (id_ a ⊗ id_ b) ≃ + f ⊗ (g ∘ id_ _ ∘ h ∘ id_ _) ⊗ (id_ a ⊗ id_ b) ≃ f ⊗ g ⊗ (id_ a ⊗ id_ b) ∘ ((id_ b ⊗ h) ⊗ (id_ a ⊗ id_ b)). Proof. intros. - match goal with - |- ?T ≃ ?T' => strong_fencepost_no_id_debug T - (* ; strong_fencepost_no_id T' *) - end. + do 2 partners_rw right_unit. + LHS weak_fencepost. + rewrite tensor_id. easy. -Qed. *) +Qed. + +Lemma test_strong_fencepost_no_id_2 : forall + {a b m n o} (f : a ~> b) (g : m ~> n) (h : n ~> o), + f ⊗ (g ∘ (id_ _ ∘ h) ∘ (id_ _ ∘ id_ _)) ⊗ (id_ a ⊗ id_ b) ≃ + f ⊗ g ⊗ (id_ a ⊗ id_ b) ∘ ((id_ b ⊗ h) ⊗ (id_ a ⊗ id_ b)). +Proof. + intros. + rewrite !right_unit. + partners_rw right_unit. + LHS weak_fencepost. + cat_easy. +Qed. + +Lemma gen_partner_test : forall {C} {cC:Category C} + {cCh : CategoryCoherence cC} {A B1 M1 B2 M2 : C} + (g1 : B1 <~> M1) (f: M1 ~> B2) (g2 : B2 <~> M2), + g1 ∘ f ∘ g2 ∘ g2^-1 ≃ g1 ∘ f. +Proof. + intros. + cancel_isos. + cat_easy. +Qed. Goal True. +assert ((id_ _ ∘ id_ _ ∘ id_ _ ∘ f ∘ (g ∘ i)) ≃ f ∘ g ∘ i) by cat_easy. + + +Ltac test_show_partnered t s term := + let ptnered := partnered_in_term t s term in + let H := fresh in + assert (H : (term ≃ ptnered)%Cat) by (show_equiv_partnered_in_term t s term); + (* setoid_rewrite H; *) + clear H. +Ltac test_show_partnered_debug t s term := + let ptnered := partnered_in_term t s term in + let H := fresh in + assert (H : (term ≃ ptnered)%Cat) by (show_equiv_partnered_in_term_debug t s term); + (* setoid_rewrite H; *) + clear H. + +test_show_partnered f g (f ∘ g). +test_show_partnered g i (f ∘ (g ∘ i)). +test_show_partnered f g (f ∘ (g ∘ i)). + +test_show_partnered f g (id_ _ ∘ id_ _ ∘ id_ _ ∘ f ∘ (g ∘ i)). +test_show_partnered f g (id_ _ ∘ id_ _ ∘ id_ _ ∘ f ∘ g ∘ i). +test_show_partnered f g (id_ _ ∘ id_ _ ∘ id_ _ ∘ f ∘ g ∘ i ∘ id_ _ ∘ id_ _). + + +Ltac test_partnered t s term := + let ptnrd := partnered_in_term t s term in + (* idtac term "≈≈>" ptnrd. *) + (* For compile: *) + idtac. + +test_partnered f g (f ∘ g). +test_partnered g i (f ∘ (g ∘ i)). +Fail test_partnered f i (f ∘ (g ∘ i)). +test_partnered f g (f ∘ (g ∘ i)). + +test_partnered f g (id_ _ ∘ id_ _ ∘ id_ _ ∘ f ∘ (g ∘ i)). +test_partnered f g (id_ _ ∘ id_ _ ∘ id_ _ ∘ f ∘ g ∘ i). +test_partnered f g (id_ _ ∘ id_ _ ∘ id_ _ ∘ f ∘ g ∘ i ∘ id_ _ ∘ id_ _). + + +Ltac test_partnered_nf t s term := + let ptnrd := partnered_in_term_nofail t s term in + (* idtac term "≈≈>" ptnrd. *) + (* For compile: *) + idtac. + +test_partnered_nf f g (f ∘ g). +test_partnered_nf g i (f ∘ (g ∘ i)). +test_partnered_nf f i (f ∘ (g ∘ i)). +test_partnered_nf f g (f ∘ (g ∘ i)). + +test_partnered_nf f g (id_ _ ∘ id_ _ ∘ id_ _ ∘ f ∘ (g ∘ i)). +test_partnered_nf f g (id_ _ ∘ id_ _ ∘ id_ _ ∘ f ∘ g ∘ i). +test_partnered_nf f g (id_ _ ∘ id_ _ ∘ id_ _ ∘ f ∘ g ∘ i ∘ id_ _ ∘ id_ _). + + + + + Local Ltac test_show_unfold_no_id_of_wf f := let wf := weak_fencepost_form f in let sf := unfold_tensor_stack_no_id wf in @@ -1275,22 +2170,29 @@ Local Ltac test_show_unfold_no_id_of_wf f := assert (H : wf ≃ sf) by (show_equiv_unfold_tensor_stack_no_id wf); clear H. -test_show_unfold_no_id_of_wf (f ⊗ (f0 ∘ g0 ∘ id_ _) ⊗ (id_ A ⊗ id_ B))%Mor. - - +test_show_unfold_no_id_of_wf (f ⊗ (f0 ∘ g0 ∘ id_ _) ⊗ (id_ A ⊗ id_ B))%Cat. Local Ltac test_show_unfold_no_id f := let sf := unfold_tensor_stack_no_id f in (* idtac sf; *) let H := fresh in - assert (H : f ≃ sf) by (show_equiv_unfold_tensor_stack_no_id f; print_state); + assert (H : f ≃ sf) by + (show_equiv_unfold_tensor_stack_no_id f); + (* (show_equiv_unfold_tensor_stack_no_id f; print_state); *) clear H. -test_show_unfold_no_id ((id_ B ⊗ id_ M ⊗ id_ (A × B)))%Mor. -test_show_unfold_no_id (f ⊗ (f0 ∘ g0 ∘ id_ _) ⊗ (id_ A ⊗ id_ B))%Mor. -test_show_unfold_no_id (f ⊗ f0 ⊗ (id_ A ⊗ id_ B))%Mor. -test_show_unfold_no_id (id_ B ⊗ g0 ⊗ id_ (A × B))%Mor. +test_show_unfold_no_id ((id_ B ⊗ id_ M ⊗ id_ (A × B)))%Cat. +test_show_unfold_no_id ((f0 ∘ g0 ∘ id_ _) ⊗ (id_ A))%Cat. +test_show_unfold_no_id (f ⊗ (f0 ∘ g0 ∘ id_ _))%Cat. + + + +test_show_unfold_no_id (f ⊗ (f0 ∘ g0) ⊗ (id_ A))%Cat. +test_show_unfold_no_id (f ⊗ (f0 ∘ g0 ∘ id_ _) ⊗ (id_ A ⊗ id_ B))%Cat. +test_show_unfold_no_id (f ⊗ (f0 ∘ g0 ∘ id_ _) ⊗ (id_ (A × B)))%Cat. +test_show_unfold_no_id (f ⊗ f0 ⊗ (id_ A ⊗ id_ B))%Cat. +test_show_unfold_no_id (id_ B ⊗ g0 ⊗ id_ (A × B))%Cat. @@ -1302,11 +2204,11 @@ Local Ltac test_show_st_of_wk f := clear H. test_show_st_of_wk f. -test_show_st_of_wk (f ∘ g)%Mor. -test_show_st_of_wk (f ⊗ g)%Mor. +test_show_st_of_wk (f ∘ g)%Cat. +test_show_st_of_wk (f ⊗ g)%Cat. -test_show_st_of_wk (f ⊗ (f ∘ g))%Mor. -test_show_st_of_wk ((f ⊗ ((f ∘ g) ⊗ (f0 ∘ g0))))%Mor. +test_show_st_of_wk (f ⊗ (f ∘ g))%Cat. +test_show_st_of_wk ((f ⊗ ((f ∘ g) ⊗ (f0 ∘ g0))))%Cat. @@ -1318,9 +2220,8 @@ Local Ltac test_show_unfold f := clear H. test_show_unfold f. -test_show_unfold (f ∘ g)%Mor. -test_show_unfold (f ⊗ g)%Mor. -Local Open Scope Mor_scope. +test_show_unfold (f ∘ g)%Cat. +test_show_unfold (f ⊗ g)%Cat. test_show_unfold ((f ⊗ (f ∘ g) ⊗ (f0 ∘ g0))). test_show_unfold ((f ∘ g) ⊗ (f0 ∘ g0)). @@ -1331,6 +2232,7 @@ test_show_unfold ((f ⊗ ((f ∘ g) ⊗ (f0 ∘ g0)))). Local Ltac test_show_st_bot f B mC := let sf := stack_comp_id_bot f B mC in + (* idtac f sf; *) let H := fresh in assert (H : f ⊗ id_ B ≃ sf) by (show_equiv_stack_comp_id_bot f B mC); clear H. @@ -1350,8 +2252,11 @@ test_show_st_top_bot (f ∘ g) B mC0. test_show_st_top_bot (f ∘ g ∘ i) B mC0. test_show_st_top_bot (f ∘ g ∘ i ∘ id_ _ ∘ id_ _) B mC0. test_show_st_top_bot (f ∘ (g ∘ id_ _) ∘ i ∘ id_ _ ∘ id_ _) B mC0. +test_show_st_top_bot (f ∘ (g ∘ id_ _) ∘ (i ∘ id_ _) ∘ id_ _) B mC0. test_show_st_top_bot ((f ⊗ g) ∘ (g ⊗ i)) B mC0. +test_show_st_bot ((f ⊗ id_ A ∘ (id_ B ⊗ f0 ∘ id_ B ⊗ g0))) A mC1. + Local Ltac test_st_of_wk f := let wf := weak_fencepost_form f in @@ -1424,8 +2329,8 @@ Local Ltac test_merge gh := let Mg := merge_stacked_composition gh in idtac. -test_merge (mC0.(tensor) @@ f, g). -test_merge (mC0.(tensor) @@ (mC0.(tensor) @@ f, g), (cC.(compose) f0 g0)). +test_merge (mC0.(mor_tensor) f g). +test_merge (mC0.(mor_tensor) (mC0.(mor_tensor) f g) (cC.(compose) f0 g0)). @@ -1459,7 +2364,7 @@ Fail tensor_only (f ∘ g). tensor_only ((g⊗h) ⊗ f ⊗ (g⊗(g⊗(g⊗h)))). Fail tensor_only ((g⊗h) ⊗ f ⊗ (g⊗(g⊗(g⊗h ∘ id_ _)))). (* Note this will match any tensor products, so less useful in Rig category *) -tensor_only (mC0.(tensor) @@ f, (mC1.(tensor) @@ g, h)). +tensor_only (mC0.(mor_tensor) f (mC1.(mor_tensor) g h)). @@ -1474,8 +2379,8 @@ is_weak_fenced ((f ⊗ (g ⊗ h)) ∘ (id_ _ ⊗ id_ _)). Fail is_weak_fenced ((f ⊗ (g ⊗ h)) ∘ (id_ _ ⊗ id_ _) ∘ id_ _). is_weak_fenced ((f ⊗ (g ⊗ h)) ∘ ((id_ _ ⊗ id_ _) ∘ id_ _)). (* Note this also works over multiple tensors at once, perhaps undesirably: *) -is_weak_fenced (mC0.(tensor) @@ (mC1.(tensor) @@ g, h), f). -is_weak_fenced (mC0.(tensor) @@ f, (mC1.(tensor) @@ g, h)). +is_weak_fenced (mC0.(mor_tensor) (mC1.(mor_tensor) g h) f). +is_weak_fenced (mC0.(mor_tensor) f (mC1.(mor_tensor) g h)). exact Logic.I. @@ -1483,13 +2388,11 @@ Qed. End Testing. -Local Close Scope Mor_scope. Local Close Scope Cat_scope. Module FutureDirections. Local Open Scope Cat_scope. -Local Open Scope Mor_scope. Section CatExpr_orig. diff --git a/ViCaR/CategoryAutomationCompatibility.v b/ViCaR/CategoryAutomationCompatibility.v new file mode 100644 index 0000000..f6877fc --- /dev/null +++ b/ViCaR/CategoryAutomationCompatibility.v @@ -0,0 +1,4 @@ +Require CategoryAutomation. +Export -(notations) CategoryAutomation. + +Require Export CategoryAltNotations. \ No newline at end of file diff --git a/ViCaR/Classes/BraidedMonoidal.v b/ViCaR/Classes/BraidedMonoidal.v index cabbf7b..375f79d 100644 --- a/ViCaR/Classes/BraidedMonoidal.v +++ b/ViCaR/Classes/BraidedMonoidal.v @@ -5,103 +5,95 @@ Require Import Setoid. #[local] Set Universe Polymorphism. Local Open Scope Cat_scope. -Local Open Scope Mor_scope. -Local Open Scope Obj_scope. -Local Open Scope Mon_scope. -Declare Scope Brd_scope. -Delimit Scope Brd_scope with Brd. -Local Open Scope Brd_scope. -(* Notation CommuteBifunctor' F := ({| - obj_bimap := fun A B => F B A; - morphism_bimap := fun A1 A2 B1 B2 fAB1 fAB2 => F @@ fAB2, fAB1; - id_bimap := ltac:(intros; apply id_bimap); - compose_bimap := ltac:(intros; apply compose_bimap); - morphism_bicompat := ltac:(intros; apply morphism_bicompat; easy); -|}) (only parsing). *) - - - -(* Reserved Notation "'B_' x , y" (at level 39). *) Class BraidedMonoidalCategory {C : Type} {cC : Category C} (mC : MonoidalCategory cC) : Type := { - braiding : NaturalBiIsomorphism mC.(tensor) (CommuteBifunctor mC.(tensor)); - (* where "'B_' x , y " := (braiding x y); *) + braiding (A B : C) : A × B <~> B × A; +}. +Arguments BraidedMonoidalCategory {_} {_}%Cat (_)%Cat. +Arguments braiding {_} {_}%Cat {_}%Cat (_)%Cat. +Notation "'β_' x , y" := (braiding _%Cat x%Cat y%Cat) + (at level 20) : Cat_scope. + +Class BraidedMonoidalCategoryCoherence {C} {cC : Category C} + {mC : MonoidalCategory cC} {cCh : CategoryCoherence cC} + {mCh : MonoidalCategoryCoherence mC} (bC : BraidedMonoidalCategory mC) + : Type := { + braiding_natural {A1 B1 A2 B2} (f1 : A1 ~> B1) (f2 : A2 ~> B2) : + f1 ⊗ f2 ∘ β_ B1, B2 ≃ β_ A1, A2 ∘ f2 ⊗ f1; hexagon_1 (A B M : C) : - (braiding A B ⊗ id_ M) ∘ associator _ _ _ ∘ (id_ B ⊗ braiding A M) - ≃ associator _ _ _ ∘ (braiding A (B × M)) ∘ associator _ _ _; + (β_ A, B ⊗ id_ M) ∘ α_ _, _, _ ∘ (id_ B ⊗ β_ A, M) + ≃ α_ _, _, _ ∘ (β_ A, (B × M)) ∘ α_ _, _, _; hexagon_2 (A B M : C) : - ((braiding B A) ^-1 ⊗ id_ M) ∘ associator _ _ _ - ∘ (id_ B ⊗ (braiding M A)^-1) - ≃ associator _ _ _ ∘ (braiding (B × M) A)^-1 ∘ associator _ _ _; + ((β_ B, A) ^-1 ⊗ id_ M) ∘ α_ _, _, _ + ∘ (id_ B ⊗ (β_ M, A)^-1) + ≃ α_ _, _, _ ∘ (β_ (B × M), A)^-1 ∘ α_ _, _, _; MonoidalCategory_of_BraidedMonoidalCategory := mC; }. -Arguments BraidedMonoidalCategory {_} {_}%Cat (_)%Cat. -Arguments braiding {_} {_}%Cat {_}%Cat (_)%Cat. -Notation "'B_' x , y" := (braiding _%Cat x%Obj y%Obj) - (at level 20) : Brd_scope. - -Coercion MonoidalCategory_of_BraidedMonoidalCategory - : BraidedMonoidalCategory >-> MonoidalCategory. - -Lemma braiding_natural {C} {cC : Category C} {mC : MonoidalCategory cC} - {bC : BraidedMonoidalCategory mC} {A1 B1 A2 B2 : C} - (f1 : A1 ~> B1) (f2 : A2 ~> B2) : - f1 ⊗ f2 ∘ B_ B1, B2 ≃ B_ A1, A2 ∘ f2 ⊗ f1. -Proof. rewrite component_biiso_natural; easy. Qed. - -Lemma inv_braiding_natural {C} {cC : Category C} {mC : MonoidalCategory cC} - {bC : BraidedMonoidalCategory mC} {A1 B1 A2 B2 : C} - (f1 : A1 ~> B1) (f2 : A2 ~> B2) : - f1 ⊗ f2 ∘ (B_ B2, B1)^-1 ≃ (B_ A2, A1)^-1 ∘ f2 ⊗ f1. -Proof. symmetry. rewrite (component_biiso_natural_reverse); easy. Qed. - -Lemma braiding_inv_r {C} {cC : Category C} {mC : MonoidalCategory cC} - {bC : BraidedMonoidalCategory mC} (A B : C) : - B_ A, B ∘ (B_ A, B)^-1 ≃ id_ (A × B). + +Section BraidedCoherenceRewrites. + + +Context {C : Type} {cC : Category C} {mC : MonoidalCategory cC} + {bC : BraidedMonoidalCategory mC} {cCh : CategoryCoherence cC} + {mCh : MonoidalCategoryCoherence mC} + {bCh : BraidedMonoidalCategoryCoherence bC}. + +#[export, program] Instance BraidingNaturalBiIsomorphism : + NaturalBiIsomorphism (tensor _) (CommuteBifunctor (tensor _)) := { + component_biiso := braiding _; +}. +Next Obligation. rewrite bCh.(braiding_natural); easy. Qed. + +Lemma inv_braiding_natural {A1 B1 A2 B2 : C} (f1 : A1 ~> B1) (f2 : A2 ~> B2) : + f1 ⊗ f2 ∘ (β_ B2, B1)^-1 ≃ (β_ A2, A1)^-1 ∘ f2 ⊗ f1. +Proof. symmetry. + apply (component_biiso_natural_reverse (N:=BraidingNaturalBiIsomorphism)). +Qed. + +Lemma braiding_inv_r (A B : C) : + β_ A, B ∘ (β_ A, B)^-1 ≃ id_ (A × B). Proof. apply iso_inv_r. Qed. -Lemma braiding_inv_l {C} {cC : Category C} {mC : MonoidalCategory cC} - {bC : BraidedMonoidalCategory mC} (A B : C) : - (B_ A, B)^-1 ∘ B_ A, B ≃ id_ (B × A). +Lemma braiding_inv_l (A B : C) : + (β_ A, B)^-1 ∘ β_ A, B ≃ id_ (B × A). Proof. apply iso_inv_l. Qed. -Lemma hexagon_resultant_1 {C} {cC : Category C} {mC : MonoidalCategory cC} - {bC : BraidedMonoidalCategory mC} (A B M : C) : - id_ B ⊗ B_ M, A ∘ B_ B, (A×M) ∘ associator A M B ∘ id_ A ⊗ (B_ B, M)^-1 - ≃ associator B M A ^-1 ∘ B_ (B × M), A. +Lemma hexagon_resultant_1 (A B M : C) : + id_ B ⊗ β_ M, A ∘ β_ B, (A×M) ∘ associator A M B ∘ id_ A ⊗ (β_ B, M)^-1 + ≃ associator B M A ^-1 ∘ β_ (B × M), A. Proof. pose proof (hexagon_2 A B M) as hex2. replace (id_ A) with (IdentityIsomorphism A ^-1) by easy. rewrite <- (compose_tensor_iso_r' (associator B M A ^-1 ∘ B_ B × M, A) (IdentityIsomorphism A) (B_ B, M)). simpl. rewrite 2!compose_iso_r. - rewrite !assoc. + rewrite !(assoc). rewrite <- compose_iso_l. Check compose_tensor_iso_r. replace (id_ B) with (forward (IdentityIsomorphism B)) by easy. rewrite (compose_tensor_iso_r _ (IdentityIsomorphism _)). - rewrite assoc, compose_iso_l'. + rewrite (assoc), compose_iso_l'. symmetry in hex2. - rewrite assoc in hex2. + rewrite (assoc) in hex2. rewrite compose_iso_l in hex2. rewrite hex2. simpl. - rewrite <- 1!assoc. + rewrite <- !(assoc). apply compose_cancel_r. pose proof (hexagon_1 B A M) as hex1. rewrite <- compose_iso_l'. replace (id_ M) with (IdentityIsomorphism M ^-1) by easy. rewrite <- (compose_tensor_iso_l' _ (IdentityIsomorphism _)). simpl. - rewrite <- 3!assoc. + rewrite <- 3!(assoc). rewrite <- 2!compose_iso_r. rewrite <- hex1. easy. Qed. -Local Close Scope Cat. +Local Close Scope Cat. \ No newline at end of file diff --git a/ViCaR/Classes/CastCategory.v b/ViCaR/Classes/CastCategory.v deleted file mode 100644 index 6eabe9e..0000000 --- a/ViCaR/Classes/CastCategory.v +++ /dev/null @@ -1,634 +0,0 @@ -Require Import Setoid. - -Require Export CategoryTypeclass. -Require Logic.Eqdep_dec. (* only for UIP_dec *) -Require PeanoNat. (* only for eq_dep on nat *) - -#[local] Set Universe Polymorphism. - -Local Open Scope Cat. - -Class CastCategory {C : Type} (cC : Category C) : Type := { - cast {A B A' B' : C} (HA : A = A') (HB : B = B') : - A' ~> B' -> A ~> B; - (* After experimentation, this is actually all we need: *) - cast_refl {A B : C} (HA : A = A) (HB : B = B) f : - cast HA HB f ≃ f; - (* That's tenable! *) - - (* Old, for reference: - (* Will need some coherence conditions, probably - cast_id, cast_compose? Might be it. Oh, we may want - cast HA HB (cast (eq_sym HA) (eq_sym HB) f) ≃ f as - well, just giving bijectivity. All should be trivial - for any sensible cast. This seems sufficient on no - reflection, so let's see how far it can go: *) - cast_invertible {A B A' B' : C} - (HA : A' = A) (HB : B' = B) (HA' : A = A') (HB' : B = B') f : - cast HA' HB' (cast HA HB f) ≃ f; - (* On more reflection, this allows any involution to - go along with the cast (e.g. cast is cast of a matrix - but also scales by -1). So, we do need: *) - - (* In theory, we could specialize this to just - @eq_refl C A and @eq_refl C B, but I think - this might be better. Maybe not, though, as - it *purely conjuecturally* *seems* like this - property with equality correlates with decidable - equality, which we _might_ not want to require? *) - (* I think we need this, too: *) - (* cast_compose {A A' B M M' : C} (HA : A = A') (HM : M = M') - (f : A' ~> B) (g: B ~> M') : - cast HA HM (f ∘ g) ≃ cast HA eq_refl f ∘ cast eq_refl HM g; *) - *) -}. - -Notation "'cast' A B f" := (@cast _ _ _ A B _ _ _ _ f) - (at level 0, A at next level, B at next level, only printing). - -Add Parametric Morphism {C : Type} {cC : Category C} {castC : CastCategory cC} - {A B : C} {A' B' : C} - {prfA : A = A'} {prfB : B = B'} : (@cast C cC castC A B A' B' prfA prfB) - with signature - (@c_equiv C cC A' B') ==> (@c_equiv C cC A B) as cast_equiv_morphism. -Proof. - intros. - subst. - rewrite 2!cast_refl. - easy. -Qed. - -Lemma cast_id {C} {cC : Category C} {castC : CastCategory cC} - (A A' : C) (HA1 HA2 : A = A') : - cast HA1 HA2 (id_ A') ≃ id_ A. -Proof. - subst. - apply cast_refl. -Qed. - -Lemma cast_compose_gen {C} {cC : Category C} {castC : CastCategory cC} - {A A' B M M' : C} (HA : A = A') (HM : M = M') - (HB1 HB2 : B = B) - (f : A' ~> B) (g: B ~> M') : - cast HA HM (f ∘ g) ≃ cast HA HB1 f ∘ cast HB2 HM g. -Proof. - subst. - rewrite ?cast_refl. - easy. -Qed. - -Lemma cast_compose {C} {cC : Category C} {castC : CastCategory cC} - {A A' B M M' : C} (HA : A = A') (HM : M = M') - (f : A' ~> B) (g: B ~> M') : - cast HA HM (f ∘ g) ≃ cast HA eq_refl f ∘ cast eq_refl HM g. -Proof. - apply cast_compose_gen. -Qed. - -Lemma cast_invertible {C} {cC : Category C} {castC : CastCategory cC} - {A B A' B' : C} - (HA : A' = A) (HB : B' = B) (HA' : A = A') (HB' : B = B') f : - cast HA' HB' (cast HA HB f) ≃ f. -Proof. - subst. - rewrite ?cast_refl. - easy. -Qed. - -Lemma cast_cast_gen {C} {cC : Category C} {castC : CastCategory cC} - {A B A' B' A'' B''} (HA : A = A') (HB : B = B') - (HA' : A' = A'') (HB' : B' = B'') (HA'' : A = A'') (HB'' : B = B'') - : forall f, - cast HA HB (cast HA' HB' f) ≃ cast HA'' HB'' f. -Proof. - intros. - subst. - rewrite 3!cast_refl. - easy. -Qed. - -Lemma cast_cast {C} {cC : Category C} {castC : CastCategory cC} - {A B A' B' A'' B''} (HA : A = A') (HB : B = B') - (HA' : A' = A'') (HB' : B' = B'') : forall f, - cast HA HB (cast HA' HB' f) ≃ cast (eq_trans HA HA') (eq_trans HB HB') f. -Proof. - apply cast_cast_gen. -Qed. - -Lemma cast_irrelevance {C} {cC : Category C} {castC : CastCategory cC} - {A B A' B'} (HA1 HA2 : A = A') (HB1 HB2 : B = B') : forall f, - cast HA1 HB1 f ≃ cast HA2 HB2 f. -Proof. - intros. - subst. - rewrite 2!cast_refl. - easy. -Qed. - -Lemma cast_simplify {C} {cC : Category C} {castC : CastCategory cC} - {A B A' B'} (HA1 HA2 : A = A') (HB1 HB2 : B = B') : forall f g, - f ≃ g -> cast HA1 HB1 f ≃ cast HA2 HB2 g. -Proof. - intros f g Hfg. - rewrite Hfg. - apply cast_irrelevance. -Qed. - -Lemma cast_compatability {C} {cC : Category C} {castC : CastCategory cC} - {A B A' B'} (HA1 HA2 : A = A') (HB1 HB2 : B = B') : forall f g, - cast HA1 HB1 f ≃ cast HA2 HB2 g -> f ≃ g. -Proof. - intros. - subst. - rewrite 2!cast_refl in *. - easy. -Qed. - -Lemma cast_equiv_iff {C} {cC : Category C} {castC : CastCategory cC} - {A B A' B'} (HA1 HA2 : A = A') (HB1 HB2 : B = B') : forall f g, - cast HA1 HB1 f ≃ cast HA2 HB2 g <-> f ≃ g. -Proof. - intros. - split. - - apply cast_compatability. - - apply cast_simplify. -Qed. - -Lemma cast_symmetry {C} {cC : Category C} {castC : CastCategory cC} - {A B A' B'} (HA : A = A') (HB : B = B') - (HA' : A' = A) (HB' : B' = B) : forall f g, - cast HA HB f ≃ g <-> f ≃ cast HA' HB' g. -Proof. - intros. - subst. - rewrite 2!cast_refl. - easy. -Qed. - -Lemma cast_symmetry_swap {C} {cC : Category C} {castC : CastCategory cC} - {A B A' B'} (HA : A = A') (HB : B = B') - (HA' : A' = A) (HB' : B' = B) : forall f g, - cast HA HB f ≃ g <-> cast HA' HB' g ≃ f. -Proof. - intros. - rewrite cast_symmetry. - split; intros H; symmetry; apply H. -Qed. - -Lemma compose_cast_mid_gen {C} {cC : Category C} {castC : CastCategory cC} - {A M M' B : C} (HM1 HM2 : M = M') (HA : A = A) (HB : B = B) : - forall (f : A ~> M') (g : M' ~> B), - f ∘ g ≃ cast HA HM1 f ∘ cast HM2 HB g. -Proof. - intros. - subst. - rewrite 2!cast_refl. - easy. -Qed. - -Lemma compose_cast_mid {C} {cC : Category C} {castC : CastCategory cC} - {A M M' B : C} (HM : M = M'): - forall (f : A ~> M') (g : M' ~> B), - f ∘ g ≃ cast eq_refl HM f ∘ cast HM eq_refl g. -Proof. - apply compose_cast_mid_gen. -Qed. - -Lemma cast_compose_distribute_gen {C} {cC : Category C} {castC : CastCategory cC} - (A A' M M' B B' : C) - (HA1 HA2 : A = A') (HB1 HB2 : B = B') (HM1 HM2 : M = M') : - forall (f : A' ~> M') (g : M' ~> B'), - cast HA1 HB1 (f ∘ g) ≃ cast HA2 HM1 f ∘ cast HM2 HB2 g. -Proof. - intros. - subst. - rewrite 3!cast_refl. - easy. -Qed. - -Lemma cast_compose_distribute {C} {cC : Category C} {castC : CastCategory cC} - (A A' M M' B B' : C) (HA : A = A') (HB : B = B') (HM : M = M') : - forall (f : A' ~> M') (g : M' ~> B'), - cast HA HB (f ∘ g) ≃ cast HA HM f ∘ cast HM HB g. -Proof. - apply cast_compose_distribute_gen. -Qed. - -Lemma cast_compose_split_gen {C} {cC : Category C} {castC : CastCategory cC} - (A A' M B B' : C) - (HA1 HA2 : A = A') (HB1 HB2 : B = B') (HM1 HM2 : M = M) : - forall (f : A' ~> M) (g : M ~> B'), - cast HA1 HB1 (f ∘ g) ≃ cast HA2 HM1 f ∘ cast HM2 HB2 g. -Proof. - apply cast_compose_distribute_gen. -Qed. - -Lemma cast_compose_split {C} {cC : Category C} {castC : CastCategory cC} - (A A' M B B' : C) (HA : A = A') (HB : B = B') : - forall (f : A' ~> M) (g : M ~> B'), - cast HA HB (f ∘ g) ≃ cast HA eq_refl f ∘ cast eq_refl HB g. -Proof. - apply cast_compose_split_gen. -Qed. - - -Lemma cast_compose_l_gen {C} {cC : Category C} {castC : CastCategory cC} - (A A' M M' B : C) (HA1 HA2 : A = A') (HM : M = M') - (HM' : M' = M) (HB1 HB2 : B = B) : - forall (f : A' ~> M') (g : M ~> B), - cast HA1 HM f ∘ g ≃ cast HA2 HB1 (f ∘ cast HM' HB2 g). -Proof. - intros. - subst. - rewrite 3!cast_refl. - easy. -Qed. - -Lemma cast_compose_l {C} {cC : Category C} {castC : CastCategory cC} - (A A' M M' B : C) (HA : A = A') (HM : M = M') : - forall (f : A' ~> M') (g : M ~> B), - cast HA HM f ∘ g ≃ cast HA eq_refl (f ∘ cast (eq_sym HM) eq_refl g). -Proof. - apply cast_compose_l_gen. -Qed. - -Lemma cast_compose_r_gen {C} {cC : Category C} {castC : CastCategory cC} - (A M M' B B' : C) (HA1 HA2 : A = A) (HM : M = M') - (HM' : M' = M) (HB1 HB2 : B = B') : - forall (f : A ~> M) (g : M' ~> B'), - f ∘ cast HM HB1 g ≃ cast HA1 HB2 (cast HA2 HM' f ∘ g). -Proof. - intros. - subst. - rewrite 3!cast_refl. - easy. -Qed. - -Lemma cast_compose_r {C} {cC : Category C} {castC : CastCategory cC} - (A M M' B B' : C) (HM : M = M') (HB : B = B') : - forall (f : A ~> M) (g : M' ~> B'), - f ∘ cast HM HB g ≃ cast eq_refl HB (cast eq_refl (eq_sym HM) f ∘ g). -Proof. - apply cast_compose_r_gen. -Qed. - -Lemma cast_contract_l_gen {C} {cC : Category C} {castC : CastCategory cC} - {A A' A'' B B' B'' : C} (HA1 : A = A') (HA2 : A = A'') - (HB1 : B = B') (HB2 : B = B'') (HA3 : A'' = A') (HB3 : B'' = B'): - forall (f : A' ~> B') (g : A'' ~> B''), - cast HA1 HB1 f ≃ cast HA2 HB2 g <-> cast HA3 HB3 f ≃ g. -Proof. - intros. - subst. - rewrite 3!cast_refl. - easy. -Qed. - -Lemma cast_contract_l {C} {cC : Category C} {castC : CastCategory cC} - {A A' A'' B B' B'' : C} (HA1 : A = A') (HA2 : A = A'') - (HB1 : B = B') (HB2 : B = B'') : - forall (f : A' ~> B') (g : A'' ~> B''), - cast HA1 HB1 f ≃ cast HA2 HB2 g <-> - cast (eq_trans (eq_sym HA2) HA1) (eq_trans (eq_sym HB2) HB1) f ≃ g. -Proof. - apply cast_contract_l_gen. -Qed. - -Lemma cast_contract_r_gen {C} {cC : Category C} {castC : CastCategory cC} - {A A' A'' B B' B'' : C} (HA1 : A = A') (HA2 : A = A'') - (HB1 : B = B') (HB2 : B = B'') (HA3 : A' = A'') (HB3 : B' = B''): - forall (f : A' ~> B') (g : A'' ~> B''), - cast HA1 HB1 f ≃ cast HA2 HB2 g <-> f ≃ cast HA3 HB3 g. -Proof. - intros. - subst. - rewrite 3!cast_refl. - easy. -Qed. - -Lemma cast_contract_r {C} {cC : Category C} {castC : CastCategory cC} - {A A' A'' B B' B'' : C} (HA1 : A = A') (HA2 : A = A'') - (HB1 : B = B') (HB2 : B = B'') : - forall (f : A' ~> B') (g : A'' ~> B''), - cast HA1 HB1 f ≃ cast HA2 HB2 g <-> - f ≃ cast (eq_trans (eq_sym HA1) HA2) (eq_trans (eq_sym HB1) HB2) g. -Proof. - apply cast_contract_r_gen. -Qed. - - - - - -Lemma cast_function {C} {cC : Category C} {castC : CastCategory cC} - {A B A' B'} (HA : A = A') (HB : B = B') (f : forall a b, a ~> b) : - cast HA HB (f A' B') ≃ f A B. -Proof. - subst. - rewrite cast_refl. - easy. -Qed. - -Lemma cast_function_diag {C} {cC : Category C} {castC : CastCategory cC} - {A A'} (HA1 HA2 : A = A') (f : forall a, a ~> a) : - cast HA1 HA2 (f A') ≃ f A. -Proof. - subst. - rewrite cast_refl. - easy. -Qed. - -Lemma cast_functor_gen {C D} `{cC : Category C} `{castC : @CastCategory C cC} - `{cD : Category D} `{castD : @CastCategory D cD} - {A B A' B'} (F : Functor cC cD) (HA : A = A') (HB : B = B') - (HFA : F A = F A') (HFB : F B = F B') : - forall f, - cast HFA HFB (F @ f) ≃ F @ (cast HA HB f). -Proof. - intros f. - subst. - rewrite 2!cast_refl. - easy. -Qed. - -Definition __cast_functor_proof {C D} {A A'} (F : C -> D) (HA : A = A') : - F A = F A' := ltac:(subst; easy). - -Lemma functor_cast {C D} `{cC : Category C} `{castC : @CastCategory C cC} - `{cD : Category D} `{castD : @CastCategory D cD} - {A B A' B'} (F : Functor cC cD) (HA : A = A') (HB : B = B') : - forall f, F @ (cast HA HB f) ≃ - cast (__cast_functor_proof F HA) (__cast_functor_proof F HB) (F @ f). -Proof. - symmetry. - apply cast_functor_gen. -Qed. - -Lemma cast_contravariant_functor_gen {C D} - `{cC : Category C} `{castC : @CastCategory C cC} - `{cD : Category D} `{castD : @CastCategory D cD} - {A B A' B'} (F : ContravariantFunctor cC cD) (HA : A = A') (HB : B = B') - (HFA : F A = F A') (HFB : F B = F B') : - forall f, - cast HFB HFA (F @ f) ≃ F @ (cast HA HB f). -Proof. - intros f. - subst. - rewrite 2!cast_refl. - easy. -Qed. - -Lemma contravariant_functor_cast {C D} - `{cC : Category C} `{castC : @CastCategory C cC} - `{cD : Category D} `{castD : @CastCategory D cD} - {A B A' B'} (F : ContravariantFunctor cC cD) (HA : A = A') (HB : B = B') : - forall (f : cC.(morphism) A' B'), F @ (cast HA HB f) ≃ - cast (__cast_functor_proof F HB) (__cast_functor_proof F HA) (F @ f). -Proof. - symmetry. - apply cast_contravariant_functor_gen. -Qed. - -Lemma cast_endofunctor_gen {C} `{cC : Category C} `{castC : @CastCategory C cC} - {A B A' B'} (F : Functor cC cC) (HA : A = A') (HB : B = B') - (HFA : F A = F A') (HFB : F B = F B') : - forall f, - cast HFA HFB (F @ f) ≃ F @ (cast HA HB f). -Proof. - apply cast_functor_gen. -Qed. - -Lemma endofunctor_cast {C} `{cC : Category C} `{castC : @CastCategory C cC} - {A B A' B'} (F : Functor cC cC) (HA : A = A') (HB : B = B') : - forall f, F @ (cast HA HB f) ≃ - cast (__cast_functor_proof F HA) (__cast_functor_proof F HB) (F @ f). -Proof. - symmetry. - apply cast_endofunctor_gen. -Qed. - -Lemma cast_contravariant_endofunctor_gen - {C} `{cC : Category C} `{castC : @CastCategory C cC} - {A B A' B'} (F : ContravariantFunctor cC cC) (HA : A = A') (HB : B = B') - (HFA : F A = F A') (HFB : F B = F B') : - forall (f : cC.(morphism) A' B'), - cast HFB HFA (F @ f) ≃ F @ (cast HA HB f). -Proof. - intros. - apply cast_contravariant_functor_gen. -Qed. - -Lemma contravariant_endofunctor_cast - {C} `{cC : Category C} `{castC : @CastCategory C cC} - {A B A' B'} (F : ContravariantFunctor cC cC) (HA : A = A') (HB : B = B') - (HFA : F A = F A') (HFB : F B = F B') : - forall (f : cC.(morphism) A' B'), - F @ (cast HA HB f) ≃ - cast (__cast_functor_proof F HB) (__cast_functor_proof F HA) (F @ f). -Proof. - symmetry. - apply cast_contravariant_endofunctor_gen. -Qed. - - -Lemma cast_bifunctor_gen {C1 C2 D} - `{cC1 : Category C1} `{castC1 : @CastCategory C1 cC1} - `{cC2 : Category C2} `{castC2 : @CastCategory C2 cC2} - `{cD : Category D} `{castD : @CastCategory D cD} - {A A' B B' : C1} {M M' N N' : C2} (F : Bifunctor cC1 cC2 cD) - (HA : A = A') (HB : B = B') (HM : M = M') (HN : N = N') - (HFAM : F A M = F A' M') (HFBN : F B N = F B' N') : - forall f g, - cast HFAM HFBN (F @@ f, g) ≃ F @@ (cast HA HB f), (cast HM HN g). -Proof. - intros f g. - subst. - rewrite 3!cast_refl. - easy. -Qed. - -Definition __cast_bifunctor_proof {C1 C2 D} {A A'} {B B'} (F : C1 -> C2 -> D) - (HA : A = A') (HB : B = B') : - F A B = F A' B' := ltac:(subst; easy). - -Lemma cast_bifunctor {C1 C2 D} - `{cC1 : Category C1} `{castC1 : @CastCategory C1 cC1} - `{cC2 : Category C2} `{castC2 : @CastCategory C2 cC2} - `{cD : Category D} `{castD : @CastCategory D cD} - {A A' B B' : C1} {M M' N N' : C2} (F : Bifunctor cC1 cC2 cD) - (HA : A = A') (HB : B = B') (HM : M = M') (HN : N = N') - (HFAM : F A M = F A' M') (HFBN : F B N = F B' N') : - forall f g, - F @@ (cast HA HB f), (cast HM HN g) ≃ - cast (__cast_bifunctor_proof F HA HM) (__cast_bifunctor_proof F HB HN) (F @@ f, g). -Proof. - symmetry. - apply cast_bifunctor_gen. -Qed. - -(* Lemma cast_bifunctor_gen {C1 C2 D} - `{cC1 : Category C1} `{castC1 : @CastCategory C1 cC1} - `{cC2 : Category C2} `{castC2 : @CastCategory C2 cC2} - `{cD : Category D} `{castD : @CastCategory D cD} - {A A' B B' : C1} {M M' N N' : C2} (F : Bifunctor cC1 cC2 cD) - (HA : A = A') (HB : B = B') (HM : M = M') (HN : N = N') - (HFAM : F A M = F A' M') (HFBN : F B N = F B' N') : - forall f g, - cast HFAM HFBN (F @@ f, g) ≃ F @@ (cast HA HB f), (cast HM HN g). -Proof. - intros f g. - subst. - rewrite 3!cast_refl. - easy. -Qed. *) - - -(* Section for monoidal category casts *) -Lemma cast_tensor {C} {cC : Category C} - {monC : MonoidalCategory cC} {castC : CastCategory cC} - {A B M N A' B' M' N' : C} - (HA : A = A') (HB : B = B') (HM : M = M') (HN : N = N') - (HAB : A × B = A' × B') (HMN : M × N = M' × N') : forall f g, - cast HAB HMN (f ⊗ g) ≃ cast HA HM f ⊗ cast HB HN g. -Proof. - intros f g. - subst. - rewrite 3!cast_refl. - easy. -Qed. - -Lemma cast_tensor_r_gen {C} {cC : Category C} - {monC : MonoidalCategory cC} {castC : CastCategory cC} - {A B M N M' N' : C} (HM : M = M') (HN : N = N') - (HAM : A × M = A × M') (HBN : B × N = B × N') : forall f g, - f ⊗ cast HM HN g ≃ cast HAM HBN (f ⊗ g). -Proof. - intros f g. - subst. - rewrite 2!cast_refl. - easy. -Qed. - -Definition __cast_tensor_r_proof {C} `{monC : MonoidalCategory C} - {M M' : C} (HM : M = M') A : A × M = A × M' := ltac:(subst;easy). - -Lemma cast_tensor_r {C} {cC : Category C} - {monC : MonoidalCategory cC} {castC : CastCategory cC} - {A B M N M' N' : C} (HM : M = M') (HN : N = N') : forall (f : A ~> B) g, - f ⊗ cast HM HN g ≃ - cast (__cast_tensor_r_proof HM A) (__cast_tensor_r_proof HN B) (f ⊗ g). -Proof. - apply cast_tensor_r_gen. -Qed. - -Lemma cast_tensor_l_gen {C} {cC : Category C} - {monC : MonoidalCategory cC} {castC : CastCategory cC} - {A B A' B' M N : C} (HA : A = A') (HB : B = B') - (HAM : A × M = A' × M) (HBN : B × N = B' × N) : forall f g, - cast HA HB f ⊗ g ≃ cast HAM HBN (f ⊗ g). -Proof. - intros f g. - subst. - rewrite 2!cast_refl. - easy. -Qed. - -Definition __cast_tensor_l_proof {C} `{monC : MonoidalCategory C} - {A A' : C} (HA : A = A') M : A × M = A' × M := ltac:(subst;easy). - -Lemma cast_tensor_l {C} {cC : Category C} - {monC : MonoidalCategory cC} {castC : CastCategory cC} - {A B A' B' M N : C} (HA : A = A') (HB : B = B') : forall f (g: M ~> N), - cast HA HB f ⊗ g ≃ - cast (__cast_tensor_l_proof HA M) (__cast_tensor_l_proof HB N) (f ⊗ g). -Proof. - apply cast_tensor_l_gen. -Qed. - - -(* Section for dagger category casts *) -Lemma cast_dagger_gen {C} {cC : Category C} - {dagC : DaggerCategory cC} {castC : CastCategory cC} - {A B A' B' : C} (HA1 HA2 : A = A') (HB1 HB2 : B = B'): - forall f, (cast HA1 HB1 f) † ≃ cast HB2 HA2 (f †). -Proof. - intros. - subst. - rewrite 2!cast_refl. - easy. -Qed. - -Lemma cast_dagger {C} {cC : Category C} - {dagC : DaggerCategory cC} {castC : CastCategory cC} - {A B A' B' : C} (HA : A = A') (HB : B = B'): - forall f, (cast HA HB f) † ≃ cast HB HA (f †). -Proof. - apply cast_dagger_gen; easy. -Qed. - -(* -cast_n_compose: - forall {n n' m : nat} (zx : ZX n n) (prf : n' = n), - n_compose m ($ n', n' ::: zx $) ∝ $ n', n' ::: n_compose m zx $ -cast_n_stack1: - forall {n n' : nat} (prfn prfm : n' = n) (zx : ZX 1 1), - $ n', n' ::: n ↑ zx $ ∝ n' ↑ zx - - - - - -cast_compose_mid_contract: - forall {n m o : nat} (n' m' o' : nat) (prfn prfn' : n' = n) - (prfm prfm' : m' = m) (prfo prfo' : o' = o) (zx0 : ZX n m) - (zx1 : ZX m o), - $ n', o' ::: zx0 ⟷ zx1 $ ∝ $ n', m' ::: zx0 $ ⟷ $ m', o' ::: zx1 $ -cast_compose_partial_contract_r: - forall {n m o : nat} (n' m' o' o'' : nat) (prfn : n' = n) - (prfm : m' = m) (prfo : o' = o') (prfo' : o' = o) - (prfo'' : o' = o'') (prfo''' : o'' = o) (zx0 : ZX n m') - (zx1 : ZX m o), - $ n', o' ::: zx0 ⟷ $ m', o' ::: zx1 $ $ - ∝ $ n', o' ::: zx0 ⟷ $ m', o'' ::: zx1 $ $ -cast_compose_partial_contract_l: - forall {n m o : nat} (n' n'' m' o' : nat) (prfn : n' = n') - (prfn' : n' = n) (prfn'' : n' = n'') (prfn''' : n'' = n) - (prfm : m' = m) (prfo : o' = o) (zx0 : ZX n m) - (zx1 : ZX m' o), - $ n', o' ::: $ n', m' ::: zx0 $ ⟷ zx1 $ - ∝ $ n', o' ::: $ n'', m' ::: zx0 $ ⟷ zx1 $ - -*) - -Definition CastCategory_of_DecEq_Category {C : Type} (cC: Category C) - (dec : forall A B : C, {A = B} + {A <> B}) : - @CastCategory C cC := {| - cast := fun A B A' B' HA HB => - match HA in (_ = a) return (a ~> B' -> A ~> B) with (* Tell coq that A = A' *) - | eq_refl => - fun f => - match HB in (_ = b) return (A ~> b -> A ~> B) with - | eq_refl => fun f' => f' - end f - end; - cast_refl := ltac:(intros A B HA HB; - rewrite (@Eqdep_dec.UIP_dec C dec A _ _ (eq_refl)); - rewrite (@Eqdep_dec.UIP_dec C dec B _ _ (eq_refl)); - reflexivity); -|}. - -Definition CastCategory_of_NatCategory (cC: Category nat) : - @CastCategory nat cC := {| - cast := fun A B A' B' HA HB => - match HA in (_ = a) return (a ~> B' -> A ~> B) with (* Tell coq that A = A' *) - | eq_refl => - fun f => - match HB in (_ = b) return (A ~> b -> A ~> B) with - | eq_refl => fun f' => f' - end f - end; - cast_refl := ltac:(intros A B HA HB; - rewrite (Eqdep_dec.UIP_refl_nat _ HA); - rewrite (Eqdep_dec.UIP_refl_nat _ HB); - reflexivity); -|}. \ No newline at end of file diff --git a/ViCaR/Classes/Category.v b/ViCaR/Classes/Category.v index cc0e1ad..d82caa1 100644 --- a/ViCaR/Classes/Category.v +++ b/ViCaR/Classes/Category.v @@ -4,23 +4,11 @@ Declare Scope Cat_scope. Delimit Scope Cat_scope with Cat. Local Open Scope Cat. -Declare Scope Obj_scope. -Delimit Scope Obj_scope with Obj. -Local Open Scope Obj. - -Declare Scope Mor_scope. -Delimit Scope Mor_scope with Mor. -Local Open Scope Mor. - -Declare Scope Func_scope. -Delimit Scope Func_scope with Func. -Local Open Scope Func_scope. - #[local] Set Universe Polymorphism. -Reserved Notation "A ~> B" (at level 60). -Reserved Notation "f ≃ g" (at level 70). -Reserved Notation "A ≅ B" (at level 70). +Reserved Notation "A ~> B" (at level 38). +Reserved Notation "f ≃ g" (at level 70). (* \simeq *) +Reserved Notation "A ≅ B" (at level 70). (* \cong *) Reserved Notation "'id_' A" (at level 15). Class Category (C : Type) : Type := { @@ -30,64 +18,76 @@ Class Category (C : Type) : Type := { (* Morphism equivalence *) c_equiv {A B : C} : relation (A ~> B) where "f ≃ g" := (c_equiv f g) : Cat_scope; - c_equiv_rel {A B : C} : equivalence (A ~> B) c_equiv; compose {A B M : C} : (A ~> B) -> (B ~> M) -> (A ~> M); - compose_compat {A B M : C} : - forall (f g : A ~> B), f ≃ g -> - forall (h j : B ~> M), h ≃ j -> - compose f h ≃ compose g j; - assoc {A B M N : C} - (f : A ~> B) (g : B ~> M) (h : M ~> N) : - compose (compose f g) h ≃ compose f (compose g h); c_identity (A : C) : A ~> A; - left_unit {A B : C} (f : A ~> B) : compose (c_identity A) f ≃ f; - right_unit {A B : C} (f : A ~> B) : compose f (c_identity B) ≃ f; }. -Arguments morphism {_} (cC)%Cat (A B)%Obj : rename. -Arguments c_equiv {_} (cC)%Cat {A B}%Obj (f g)%Mor : rename. -Arguments c_equiv_rel {_} (cC)%Cat {A B}%Obj : rename. -Arguments compose {_} (cC)%Cat {_ _ _}%Obj (f g)%Mor : rename. -Arguments compose_compat {_} {cC}%Cat {_ _ _}%Obj (f g)%Mor _ (h j)%Mor : rename. -Arguments assoc {_} {cC}%Cat {_ _ _ _}%Obj (f g h)%Mor : rename. -Arguments c_identity {_} {cC}%Cat (A)%Obj : rename. -Arguments left_unit {_} {cC}%Cat {A B}%Obj (f)%Mor : rename. -Arguments right_unit {_} {cC}%Cat {A B}%Obj (f)%Mor : rename. - -Notation "'id_' A" := (c_identity A%Obj) - (at level 15, A at next level, no associativity) : Mor_scope. -Notation "A ~> B" := (morphism _%Cat A%Obj B%Obj) - (at level 60, B at next level, no associativity) : Cat_scope. -Notation "f ≃ g" := (c_equiv _%Cat f%Mor g%Mor) +Arguments morphism {_} (cC)%Cat (A B)%Cat : rename. +Arguments c_equiv {_} (cC)%Cat {A B}%Cat (f g)%Cat : rename. +Arguments compose {_} (cC)%Cat {_ _ _}%Cat (f g)%Cat : rename. +Arguments c_identity {_} {cC}%Cat (A)%Cat : rename. + +Notation "'id_' A" := (c_identity A%Cat) + (at level 15, A at next level, no associativity) : Cat_scope. +Notation "A ~> B" := (morphism _%Cat A%Cat B%Cat) + (at level 38, B at next level, no associativity) : Cat_scope. +Notation "f ≃ g" := (c_equiv _%Cat f%Cat g%Cat) (at level 70, g at next level) : Cat_scope. (* \simeq *) -Notation "f ∘ g" := (compose _%Cat f%Mor g%Mor) - (at level 65, g at next level, left associativity) : Mor_scope. (* \circ *) +Notation "f ∘ g" := (compose _%Cat f%Cat g%Cat) + (at level 40, g at next level, left associativity) : Cat_scope. (* \circ *) +Class CategoryCoherence {C} (cC : Category C) : Type := { + (* to_base_struct_cat := cC; *) + + c_equiv_rel {A B : C} : equivalence (A ~> B) cC.(c_equiv); + + compose_compat {A B M : C} : + forall (f g : A ~> B), f ≃ g -> + forall (h j : B ~> M), h ≃ j -> + f ∘ h ≃ g ∘ j; + + assoc {A B M N : C} + (f : A ~> B) (g : B ~> M) (h : M ~> N) : + (f ∘ g) ∘ h ≃ f ∘ (g ∘ h); + + left_unit {A B : C} (f : A ~> B) : id_ A ∘ f ≃ f; + right_unit {A B : C} (f : A ~> B) : f ∘ id_ B ≃ f; +}. + +Arguments c_equiv_rel {_} {cC}%Cat {cCh}%Cat {A B}%Cat : rename. +Arguments assoc {_} {cC}%Cat {cCh}%Cat {_ _ _ _}%Cat (f g h)%Cat : rename. +Arguments compose_compat {_} {cC}%Cat {cCh}%Cat {_ _ _}%Cat (f g)%Cat _ (h j)%Cat : rename. +Arguments left_unit {_} {cC}%Cat {cCh}%Cat {A B}%Cat (f)%Cat : rename. +Arguments right_unit {_} {cC}%Cat {cCh}%Cat {A B}%Cat (f)%Cat : rename. + +(* Coercion to_base_struct_cat : CategoryCoherence >-> Category. *) Add Parametric Relation {C : Type} {cC : Category C} - (A B : C) : (A ~> B) cC.(c_equiv) - reflexivity proved by (equiv_refl (A ~> B) _ cC.(c_equiv_rel)) - symmetry proved by (equiv_sym (A ~> B) _ cC.(c_equiv_rel)) - transitivity proved by (equiv_trans (A ~> B) _ cC.(c_equiv_rel)) + {cCh : CategoryCoherence cC} (A B : C) : (A ~> B) cC.(c_equiv) + reflexivity proved by (equiv_refl (A ~> B) _ (c_equiv_rel)) + symmetry proved by (equiv_sym (A ~> B) _ (c_equiv_rel)) + transitivity proved by (equiv_trans (A ~> B) _ (c_equiv_rel)) as prop_equiv_rel. -Add Parametric Morphism {C : Type} {cC : Category C} (n o m : C) : cC.(compose) - with signature (@c_equiv C cC n m) ==> (@c_equiv C cC m o) ==> cC.(c_equiv) as compose_mor. -Proof. apply compose_compat; assumption. Qed. +Add Parametric Morphism {C : Type} {cC : Category C} {cCh : CategoryCoherence cC} + (n o m : C) : cC.(compose) + with signature (@c_equiv C cC n m) ==> + (@c_equiv C cC m o) ==> cC.(c_equiv) as compose_mor. +Proof. apply compose_compat. Qed. Notation " 'is_inverse' f g" := - (f%Mor ∘ g%Mor ≃ id_ _ /\ g%Mor ∘ f%Mor ≃ id_ _) + (f%Cat ∘ g%Cat ≃ id_ _ /\ g%Cat ∘ f%Cat ≃ id_ _) (at level 10, f at next level, g at next level) : Cat_scope. Lemma inverse_comm {C} {cC : Category C} {A B : C} (f : A ~> B) (g : B ~> A) : is_inverse f g <-> is_inverse g f. Proof. easy. Qed. -Lemma inverse_unique {C} {cC : Category C} {A B : C} +Lemma inverse_unique {C} {cC : Category C} {cCh : CategoryCoherence cC} {A B : C} {f : A ~> B} {g g' : B ~> A} (Hg : is_inverse f g) (Hg' : is_inverse f g') : g ≃ g'. Proof. @@ -102,25 +102,26 @@ Qed. Definition isomorphic {C : Type} {cC : Category C} (A B : C) := exists (f : A ~> B) (g : B ~> A), is_inverse f g. -Arguments isomorphic {_} {_}%Cat (_ _)%Obj. +Arguments isomorphic {_} {_}%Cat (_ _)%Cat. Class Isomorphism {C : Type} {cC : Category C} (A B : C) := { forward : A ~> B; reverse : B ~> A; isomorphism_inverse : is_inverse forward reverse; }. -Arguments Isomorphism {_} {_}%Cat (_ _)%Obj. -Arguments forward {_} {_}%Cat {_ _}%Obj (f)%Mor : rename. -Arguments reverse {_} {_}%Cat {_ _}%Obj (f)%Mor : rename. +Arguments Isomorphism {_} {_}%Cat (_ _)%Cat. +Arguments forward {_} {_}%Cat {_ _}%Cat (f)%Cat : rename. +Arguments reverse {_} {_}%Cat {_ _}%Cat (f)%Cat : rename. (* Notation id_A I := (proj1 I.(isomorphism_inverse)). Notation id_B I := (proj2 I.(isomorphism_inverse)). *) Coercion forward : Isomorphism >-> morphism. -Notation "f '^-1'" := (reverse f%Mor) (at level 25) : Mor_scope. +Notation "f '^-1'" := (reverse f%Cat) (at level 25, only parsing) : Cat_scope. +Notation "f '⁻¹'" := (reverse f%Cat) (at level 25) : Cat_scope. (* \^- \^1 *) -Notation "A '<~>' B" := (Isomorphism A%Obj B%Obj) (at level 70) : Cat_scope. -Notation "A ≅ B" := (isomorphic A%Obj B%Obj) (at level 70) : Cat_scope. (* \cong *) +Notation "A '<~>' B" := (Isomorphism A%Cat B%Cat) (at level 70) : Cat_scope. +Notation "A ≅ B" := (isomorphic A%Cat B%Cat) (at level 70) : Cat_scope. (* \cong *) -Lemma isomorphic_iff_Isomorphism {C : Type} `{Category C} (A B : C) : +Lemma isomorphic_iff_Isomorphism {C : Type} {cC : Category C} (A B : C) : isomorphic A B <-> exists _: Isomorphism A B, True. Proof. split. @@ -132,7 +133,15 @@ Proof. exists f; exists g; auto. Qed. -Lemma isomorphic_refl {C : Type} `{Category C} (A : C) : isomorphic A A. +Lemma left_unit' {C : Type} {cC : Category C} {cCh : CategoryCoherence cC} + {A B : C} (f : A ~> B) : id_ _ ∘ f ≃ f. +Proof. + rewrite (left_unit). + easy. +Qed. + +Lemma isomorphic_refl {C : Type} {cC : Category C} {cCh : CategoryCoherence cC} + (A : C) : isomorphic A A. Proof. exists (c_identity A). exists (c_identity A). @@ -140,21 +149,23 @@ Proof. split; reflexivity. Qed. -Lemma isomorphic_trans {C : Type} `{Category C} (A B M : C) : +Lemma isomorphic_trans {C : Type} {cC : Category C} {cCh : CategoryCoherence cC} + (A B M : C) : isomorphic A B -> isomorphic B M -> isomorphic A M. Proof. intros [fAB [fBA [HfABA HfBAB]]] [fBM [fMB [HfBMB HfMBM]]]. exists (fAB ∘ fBM). exists (fMB ∘ fBA). split; - rewrite assoc. - - rewrite <- (assoc fBM), HfBMB, left_unit, HfABA. + rewrite (assoc). + - rewrite <- (assoc fBM), HfBMB, (left_unit), HfABA. reflexivity. - - rewrite <- (assoc fBA), HfBAB, left_unit, HfMBM. + - rewrite <- (assoc fBA), HfBAB, (left_unit), HfMBM. reflexivity. Qed. -Lemma isomorphic_sym {C : Type} `{Category C} (A B : C) : +Lemma isomorphic_sym {C : Type} {cC : Category C} {cCh : CategoryCoherence cC} + (A B : C) : isomorphic A B -> isomorphic B A. Proof. intros [fAB [fBA [HfABA HfBAB]]]. @@ -163,7 +174,8 @@ Proof. split; reflexivity. Qed. -Add Parametric Relation {C : Type} `{Cat : Category C} : C (@isomorphic C Cat) +Add Parametric Relation {C : Type} {cC : Category C} + {cCh : CategoryCoherence cC} : C (@isomorphic C cC) reflexivity proved by isomorphic_refl symmetry proved by isomorphic_sym transitivity proved by isomorphic_trans @@ -179,22 +191,23 @@ Class Functor {C D : Type} (cC: Category C) (cD : Category D) : Type := { morphism_compat {A B : C} (f g : A ~> B) : f ≃ g -> (morphism_map f) ≃ (morphism_map g); }. -Arguments obj_map {_ _} {_ _}%Cat (F)%Func (A)%Obj : rename. -Arguments morphism_map {_ _} {_ _}%Cat (F)%Func {A B}%Obj (f)%Mor : rename. -Arguments id_map {_ _} {_ _}%Cat (F)%Func (A)%Obj : rename. -Arguments compose_map {_ _} {_ _}%Cat (F)%Func {A B M}%Obj (f g)%Mor : rename. -Arguments morphism_compat {_ _} {_ _}%Cat (F)%Func {A B}%Obj (f g)%Mor : rename. +Arguments obj_map {_ _} {_ _}%Cat (F)%Cat (A)%Cat : rename. +Arguments morphism_map {_ _} {_ _}%Cat (F)%Cat {A B}%Cat (f)%Cat : rename. +Arguments id_map {_ _} {_ _}%Cat (F)%Cat (A)%Cat : rename. +Arguments compose_map {_ _} {_ _}%Cat (F)%Cat {A B M}%Cat (f g)%Cat : rename. +Arguments morphism_compat {_ _} {_ _}%Cat (F)%Cat {A B}%Cat (f g)%Cat : rename. Coercion obj_map : Functor >-> Funclass. -Arguments morphism_map {_ _} {_ _}%Cat (_)%Func {_ _}%Obj (_)%Mor. -Notation "F @ X" := (morphism_map F%Func X%Mor) (at level 55) : Mor_scope. +Arguments morphism_map {_ _} {_ _}%Cat (_)%Cat {_ _}%Cat (_)%Cat. +Notation "F @ f" := (morphism_map F%Cat f%Cat) (at level 38) : Cat_scope. Add Parametric Morphism {C D : Type} {cC : Category C} {cD : Category D} (F : Functor cC cD) (A B : C): F.(morphism_map) with signature (@c_equiv C cC A B) ==> (@c_equiv D cD (F A) (F B)) as functor_equiv_mor. Proof. apply morphism_compat. Qed. -Add Parametric Morphism {C D : Type} `{cC : Category C} `{cD : Category D} +Add Parametric Morphism {C D : Type} {cC : Category C} + {cCh : CategoryCoherence cC} {cD : Category D} {cDh : CategoryCoherence cD} (F : Functor cC cD) : F.(obj_map) with signature (@isomorphic C cC) ==> (@isomorphic D cD) as functor_isomorphic_mor. Proof. @@ -220,18 +233,18 @@ Class Bifunctor {C1 C2 D : Type} (cC1: Category C1) f ≃ f' -> g ≃ g' -> morphism_bimap f g ≃ morphism_bimap f' g' }. -Arguments obj_bimap {_ _ _} {_ _ _}%Cat (F)%Func (A1 A2)%Obj : rename. -Arguments morphism_bimap {_ _ _} {_ _ _}%Cat (F)%Func - {A1 B1 A2 B2}%Obj (f1 f2)%Mor : rename. -Arguments id_bimap {_ _ _} {_ _ _}%Cat (F)%Func (A1 A2)%Obj : rename. -Arguments compose_bimap {_ _ _} {_ _ _}%Cat (F)%Func - {A1 B1 M1 A2 B2 M2}%Obj (f1 g1 f2 g2)%Mor : rename. -Arguments morphism_bicompat {_ _ _} {_ _ _}%Cat (F)%Func - {A1 B1 A2 B2}%Obj (f1 f1' f2 f2')%Mor : rename. +Arguments obj_bimap {_ _ _} {_ _ _}%Cat (F)%Cat (A1 A2)%Cat : rename. +Arguments morphism_bimap {_ _ _} {_ _ _}%Cat (F)%Cat + {A1 B1 A2 B2}%Cat (f1 f2)%Cat : rename. +Arguments id_bimap {_ _ _} {_ _ _}%Cat (F)%Cat (A1 A2)%Cat : rename. +Arguments compose_bimap {_ _ _} {_ _ _}%Cat (F)%Cat + {A1 B1 M1 A2 B2 M2}%Cat (f1 g1 f2 g2)%Cat : rename. +Arguments morphism_bicompat {_ _ _} {_ _ _}%Cat (F)%Cat + {A1 B1 A2 B2}%Cat (f1 f1' f2 f2')%Cat : rename. Coercion obj_bimap : Bifunctor >-> Funclass. -Notation " F '@@' X , Y " := (morphism_bimap F%Func X%Mor Y%Mor) - (at level 55) : Mor_scope. +Notation " F '@@' X , Y " := (morphism_bimap F%Cat X%Cat Y%Cat) + (at level 38) : Cat_scope. Add Parametric Morphism {C1 C2 D : Type} {cC1 : Category C1} {cC2 : Category C2} {cD : Category D} @@ -240,8 +253,10 @@ Add Parametric Morphism {C1 C2 D : Type} ==> (@c_equiv D cD (F A1 A2) (F B1 B2)) as bifunctor_equiv_mor. Proof. intros. apply morphism_bicompat; easy. Qed. -Add Parametric Morphism {C1 C2 D : Type} `{cC1 : Category C1} `{cC2 : Category C2} - `{cD : Category D} (F : Bifunctor cC1 cC2 cD) : F.(obj_bimap) +Add Parametric Morphism {C1 C2 D : Type} {cC1 : Category C1} + {cC1h : CategoryCoherence cC1} {cC2 : Category C2} + {cC2h : CategoryCoherence cC2} {cD : Category D} {cDh : CategoryCoherence cD} + (F : Bifunctor cC1 cC2 cD) : F.(obj_bimap) with signature (@isomorphic C1 cC1) ==> (@isomorphic C2 cC2) ==> (@isomorphic D cD) as bifunctor_isometric_mor. Proof. @@ -261,7 +276,7 @@ Definition CommuteBifunctor {C1 C2 D : Type} `{cC1 : Category C1} morphism_bicompat := ltac:(intros; apply morphism_bicompat; easy); |}. -Arguments CommuteBifunctor {_ _ _} {_ _ _}%Cat (_)%Func /. +Arguments CommuteBifunctor {_ _ _} {_ _ _}%Cat (_)%Cat /. #[export] Typeclasses Transparent CommuteBifunctor. @@ -269,15 +284,13 @@ Arguments CommuteBifunctor {_ _ _} {_ _ _}%Cat (_)%Func /. Class NaturalTransformation {C D : Type} `{cC: Category C} `{cD : Category D} (F G : Functor cC cD) := { component_map (A : C) : F A ~> G A; - (* where "'α_' X" := (component_map X); *) component_map_natural {A B : C} (f : A ~> B) : F @ f ∘ component_map B ≃ component_map A ∘ G @ f; }. -Arguments NaturalTransformation {_ _} {_ _}%Cat (_ _)%Func. -Arguments component_map {_ _} {_ _}%Cat {_ _}%Func (N)%Func (_)%Obj : rename. -Arguments component_map_natural {_ _} {_ _}%Cat {_ _}%Func - {N}%Func {_ _}%Obj (f)%Mor : rename. -Notation "'α_' X" := (component_map _%Func X%Obj) (at level 20) : Mor_scope. +Arguments NaturalTransformation {_ _} {_ _}%Cat (_ _)%Cat. +Arguments component_map {_ _} {_ _}%Cat {_ _}%Cat (N)%Cat (_)%Cat : rename. +Arguments component_map_natural {_ _} {_ _}%Cat {_ _}%Cat + {N}%Cat {_ _}%Cat (f)%Cat : rename. Class NaturalIsomorphism {C D : Type} `{cC: Category C} `{cD : Category D} (F G : Functor cC cD) := { @@ -285,12 +298,12 @@ Class NaturalIsomorphism {C D : Type} `{cC: Category C} `{cD : Category D} component_iso_natural (A B : C) (f : A ~> B) : F@f ∘ component_iso B ≃ component_iso A ∘ G@f; }. -Arguments NaturalIsomorphism {_ _} {_ _}%Cat (_ _)%Func. -Arguments component_iso {_ _} {_ _}%Cat {_ _}%Func (N)%Func (_)%Obj : rename. -Arguments component_iso_natural {_ _} {_ _}%Cat {_ _}%Func - {N}%Func {_ _}%Obj (f)%Mor : rename. +Arguments NaturalIsomorphism {_ _} {_ _}%Cat (_ _)%Cat. +Arguments component_iso {_ _} {_ _}%Cat {_ _}%Cat (N)%Cat (_)%Cat : rename. +Arguments component_iso_natural {_ _} {_ _}%Cat {_ _}%Cat + {N}%Cat {_ _}%Cat (f)%Cat : rename. -Coercion component_iso : NaturalIsomorphism >-> Funclass. (* TODO: is this sensible? I think not *) +Coercion component_iso : NaturalIsomorphism >-> Funclass. Definition NaturalTransformation_of_NaturalIsomorphism {C D : Type} `{cC : Category C} `{cD : Category D} {F G : Functor cC cD} @@ -298,28 +311,22 @@ Definition NaturalTransformation_of_NaturalIsomorphism {C D : Type} component_map := component_iso N; component_map_natural := ltac:(intros; apply component_iso_natural); |}. -(* Just to simplify notation, I'm removing this for now... *) -(* Coercion NaturalTransformation_of_NaturalIsomorphism : - NaturalIsomorphism >-> NaturalTransformation. *) Class NaturalBiTransformation {C1 C2 D : Type} `{cC1 : Category C1} `{cC2 : Category C2} `{cD : Category D} (F G : Bifunctor cC1 cC2 cD) := { component_bimap (A1 : C1) (A2 : C2) : F A1 A2 ~> G A1 A2; - (* where "'β_' X , Y" := (component_bimap X Y); *) component_bimap_natural {A1 B1 : C1} {A2 B2 : C2} (f1 : A1 ~> B1) (f2 : A2 ~> B2) : (F @@ f1, f2) ∘ (component_bimap B1 B2) ≃ (component_bimap A1 A2) ∘ (G @@ f1, f2) }. -Arguments NaturalBiTransformation {_ _ _} {_ _ _}%Cat (_ _)%Func. -Arguments component_bimap {_ _ _} {_ _ _}%Cat {_ _}%Func - (N)%Func (_ _)%Obj : rename. -Arguments component_bimap_natural {_ _ _} {_ _ _}%Cat {_ _}%Func - {N}%Func {_ _ _ _}%Obj (f1 f2)%Mor : rename. -Notation "'β_' X , Y" := (component_bimap _%Func X%Obj Y%Obj) - (at level 20) : Mor_scope. +Arguments NaturalBiTransformation {_ _ _} {_ _ _}%Cat (_ _)%Cat. +Arguments component_bimap {_ _ _} {_ _ _}%Cat {_ _}%Cat + (N)%Cat (_ _)%Cat : rename. +Arguments component_bimap_natural {_ _ _} {_ _ _}%Cat {_ _}%Cat + {N}%Cat {_ _ _ _}%Cat (f1 f2)%Cat : rename. Class NaturalBiIsomorphism {C1 C2 D : Type} `{cC1 : Category C1} `{cC2 : Category C2} `{cD : Category D} (F G : Bifunctor cC1 cC2 cD) := { @@ -327,11 +334,11 @@ Class NaturalBiIsomorphism {C1 C2 D : Type} `{cC1 : Category C1} component_biiso_natural (A1 B1 : C1) (A2 B2 : C2) (f1 : A1 ~> B1) (f2 : A2 ~> B2) : (F @@ f1, f2) ∘ (component_biiso B1 B2) ≃ (component_biiso A1 A2) ∘ (G @@ f1, f2) }. -Arguments NaturalBiIsomorphism {_ _ _} {_ _ _}%Cat (_ _)%Func. -Arguments component_biiso {_ _ _} {_ _ _}%Cat {_ _}%Func - (N)%Func (_ _)%Obj : rename. -Arguments component_bimap_natural {_ _ _} {_ _ _}%Cat {_ _}%Func - {N}%Func {_ _ _ _}%Obj (f1 f2)%Mor : rename. +Arguments NaturalBiIsomorphism {_ _ _} {_ _ _}%Cat (_ _)%Cat. +Arguments component_biiso {_ _ _} {_ _ _}%Cat {_ _}%Cat + (N)%Cat (_ _)%Cat : rename. +Arguments component_bimap_natural {_ _ _} {_ _ _}%Cat {_ _}%Cat + {N}%Cat {_ _ _ _}%Cat (f1 f2)%Cat : rename. Coercion component_biiso : NaturalBiIsomorphism >-> Funclass. Definition NaturalBiTransformation_of_NaturalBiIsomorphism {C1 C2 D : Type} @@ -352,21 +359,27 @@ Definition NaturalBiTransformation_of_NaturalBiIsomorphism {C1 C2 D : Type} Definition OppositeCategory {C : Type} (cC : Category C) : Category C := {| morphism := fun A B => cC.(morphism) B A; c_equiv := fun A B => @c_equiv C cC B A; - c_equiv_rel := fun A B => @c_equiv_rel C cC B A; compose := fun A B M fAB fBM => fBM ∘ fAB; - compose_compat := fun A B M f g Hfg h j Hhj => - @compose_compat C cC M B A _ _ Hhj _ _ Hfg; - assoc := fun A B M N f g h => - equiv_sym _ _ cC.(c_equiv_rel) _ _ (@assoc C cC N M B A h g f); c_identity := c_identity; - left_unit := fun A B => @right_unit C cC B A; - right_unit := fun A B => @left_unit C cC B A; |}. Notation "C '^op'" := (OppositeCategory C%Cat) - (at level 40, left associativity) : Cat_scope. + (at level 34, left associativity, only parsing) : Cat_scope. +Notation "C 'ᵒᵖ'" := (OppositeCategory C%Cat) + (at level 34, left associativity) : Cat_scope. (* \^o \^p *) #[export] Typeclasses Transparent OppositeCategory. +#[program] Definition OppositeCategoryCoherence {C : Type} {cC : Category C} + (cCh : CategoryCoherence cC) : CategoryCoherence (cC^op) := {| + c_equiv_rel := fun A B => @c_equiv_rel C cC cCh B A; + compose_compat := fun A B M f g Hfg h j Hhj => + @compose_compat C cC cCh M B A _ _ Hhj _ _ Hfg; + assoc := fun A B M N f g h => + equiv_sym _ _ (c_equiv_rel) _ _ (@assoc C cC cCh N M B A h g f); + left_unit := fun A B => @right_unit C cC cCh B A; + right_unit := fun A B => @left_unit C cC cCh B A; +|}. + Class ContraFunctor {C D : Type} (cC: Category C) (cD : Category D) : Type := { obj_contramap : C -> D; morphism_contramap {A B : C} : (A ~> B) -> @@ -378,17 +391,17 @@ Class ContraFunctor {C D : Type} (cC: Category C) (cD : Category D) : Type := { (morphism_contramap f) ≃ (morphism_contramap g); }. Arguments ContraFunctor {_ _} (_ _)%Cat. -Arguments obj_contramap {_ _} {_ _}%Cat (F)%Func (A)%Obj : rename. -Arguments morphism_contramap {_ _} {_ _}%Cat (F)%Func - {A B}%Obj (f)%Mor : rename. -Arguments id_contramap {_ _} {_ _}%Cat (F)%Func (A)%Obj : rename. -Arguments compose_contramap {_ _} {_ _}%Cat (F)%Func - {A B M}%Obj (f g)%Mor : rename. -Arguments morphism_contracompat {_ _} {_ _}%Cat (F)%Func - {A B}%Obj (f g)%Mor : rename. +Arguments obj_contramap {_ _} {_ _}%Cat (F)%Cat (A)%Cat : rename. +Arguments morphism_contramap {_ _} {_ _}%Cat (F)%Cat + {A B}%Cat (f)%Cat : rename. +Arguments id_contramap {_ _} {_ _}%Cat (F)%Cat (A)%Cat : rename. +Arguments compose_contramap {_ _} {_ _}%Cat (F)%Cat + {A B M}%Cat (f g)%Cat : rename. +Arguments morphism_contracompat {_ _} {_ _}%Cat (F)%Cat + {A B}%Cat (f g)%Cat : rename. Coercion obj_contramap : ContraFunctor >-> Funclass. Notation "F @' X" := (F.(morphism_contramap) X) - (at level 55, no associativity) : Mor_scope. + (at level 55, no associativity) : Cat_scope. Add Parametric Morphism {C D : Type} {cC : Category C} {cD : Category D} (F : ContraFunctor cC cD) (A B : C): F.(morphism_contramap) @@ -396,7 +409,8 @@ Add Parametric Morphism {C D : Type} {cC : Category C} {cD : Category D} (@c_equiv C cC A B) ==> (@c_equiv D cD (F B) (F A)) as contrafunctor_equiv_mor. Proof. apply morphism_contracompat. Qed. -Add Parametric Morphism {C D : Type} `{cC : Category C} `{cD : Category D} +Add Parametric Morphism {C D : Type} {cC : Category C} {cCh : CategoryCoherence cC} + {cD : Category D} {cDh : CategoryCoherence cD} (F : ContraFunctor cC cD) : F.(obj_contramap) with signature (@isomorphic C cC) ==> (@isomorphic D cD) as contrafunctor_isomorphic_mor. @@ -483,7 +497,7 @@ Definition product_relation_equivalence {T U : Type} {simT : relation T} {simU : |} end end. -#[local, program] Instance ProductCategory {C D : Type} +#[local] Instance ProductCategory {C D : Type} (cC : Category C) (cD : Category D) : Category (C*D) := { morphism := fun AB MN => prod (cC.(morphism) (fst AB) (fst MN)) (cD.(morphism) (snd AB) (snd MN)); @@ -493,11 +507,6 @@ end end. match A'B' with (A', B') => product_relation (@c_equiv C cC A A') (@c_equiv D cD B B') end end; - c_equiv_rel := fun AB A'B' => - match AB with (A, B) => - match A'B' with (A', B') => - product_relation_equivalence (@c_equiv_rel C cC A A') (@c_equiv_rel D cD B B') - end end; compose := fun AB MN EF => match AB with (A, B) => match MN with (M, N) => @@ -508,6 +517,18 @@ end end. end end end end end; c_identity := fun AB => match AB with (A, B) => (c_identity A, c_identity B) end; }. + +#[program] Instance ProductCategoryCoherence {C D : Type} + {cC : Category C} {cD : Category D} + (cCh : CategoryCoherence cC) (cDh : CategoryCoherence cD) + : CategoryCoherence (ProductCategory cC cD) := { + c_equiv_rel := fun AB A'B' => + match AB with (A, B) => + match A'B' with (A', B') => + product_relation_equivalence (@c_equiv_rel C cC cCh A A') + (@c_equiv_rel D cD cDh B B') + end end; +}. Next Obligation. split; apply compose_compat; easy. Qed. @@ -536,7 +557,8 @@ Definition ProductCategoryFunctor_of_Bifunctor {C1 C2 D : Type} `{cC1 : Category |}. Definition Bifunctor_of_ProductCategoryFunctor {C1 C2 D : Type} `{cC1 : Category C1} - `{cC2 : Category C2} `{cD : Category D} (F : Functor (ProductCategory cC1 cC2) cD) : + `{cC2 : Category C2} {cD : Category D} {cDh : CategoryCoherence cD} + (F : Functor (ProductCategory cC1 cC2) cD) : Bifunctor cC1 cC2 cD := {| obj_bimap := fun A B => F (A, B); morphism_bimap := fun A1 B1 A2 B2 fA1B1 fA2B2 => @@ -548,76 +570,85 @@ Definition Bifunctor_of_ProductCategoryFunctor {C1 C2 D : Type} `{cC1 : Category End ProductCategories. -(* Lemma equiv_of_iso_compose_l {C : Type} `{cC : Category C} {A A' B : C} - (I : Isomorphism A A') (f g : A' ~> B) (H : I ∘ f ≃ I ∘ g) : - f ≃ g. -Proof. - rewrite <- (left_unit (f:=f)), <- (left_unit (f:=g)). - rewrite <- (id_B I), 2!assoc, H. - easy. -Qed. - -Lemma equiv_of_iso_compose_r {C : Type} `{cC : Category C} {A B' B : C} - (I : Isomorphism B' B) (f g : A ~> B') (H : f ∘ I ≃ g ∘ I) : - f ≃ g. -Proof. - rewrite <- (right_unit (f:=f)), <- (right_unit (f:=g)). - rewrite <- (id_A I), <- 2!assoc, H. - easy. -Qed. *) - - - (* Some useful little lemmas *) -Lemma compose_cancel_r : forall {C} {cC:Category C} {A B M : C} +Lemma compose_cancel_r : forall {C} {cC : Category C} + {cCh : CategoryCoherence cC} {A B M : C} (f g : A ~> B) (h : B ~> M), f ≃ g -> f ∘ h ≃ g ∘ h. Proof. intros. apply compose_compat; easy. Qed. -Lemma compose_cancel_l : forall {C} {cC:Category C} {A B M : C} +Lemma compose_cancel_l : forall {C} {cC:Category C} + {cCh : CategoryCoherence cC} {A B M : C} (f : A ~> B) (g h : B ~> M), g ≃ h -> f ∘ g ≃ f ∘ h. Proof. intros. apply compose_compat; easy. Qed. -Lemma iso_inv_r : forall {C} {cC:Category C} {A B : C} +Lemma iso_inv_r : forall {C} {cC:Category C} + {cCh : CategoryCoherence cC} {A B : C} (f : A <~> B), f ∘ f^-1 ≃ id_ A. Proof. intros; apply isomorphism_inverse. Qed. -Lemma iso_inv_l : forall {C} {cC:Category C} {A B : C} +Lemma iso_inv_l : forall {C} {cC:Category C} + {cCh : CategoryCoherence cC} {A B : C} (f : A <~> B), f^-1 ∘ f ≃ id_ B. Proof. intros; apply isomorphism_inverse. Qed. -Lemma compose_iso_r : forall {C} {cC:Category C} {A B M : C} +Lemma equiv_of_iso_compose_l {C : Type} `{cC : Category C} + {cCh : CategoryCoherence cC} {A A' B : C} + (I : Isomorphism A A') (f g : A' ~> B) (H : I ∘ f ≃ I ∘ g) : + f ≃ g. +Proof. + rewrite <- (left_unit f), <- (left_unit g). + rewrite <- (iso_inv_l I), 2!assoc, H. + easy. +Qed. + +Lemma equiv_of_iso_compose_r {C : Type} `{cC : Category C} + {cCh : CategoryCoherence cC} {A B' B : C} + (I : Isomorphism B' B) (f g : A ~> B') (H : f ∘ I ≃ g ∘ I) : + f ≃ g. +Proof. + rewrite <- (right_unit f), <- (right_unit g). + rewrite <- (iso_inv_r I), <- 2!assoc, H. + easy. +Qed. + + +Lemma compose_iso_r : forall {C} {cC:Category C} + {cCh : CategoryCoherence cC} {A B M : C} (f : A ~> B) (g : B <~> M) (h : A ~> M), f ∘ g ≃ h <-> f ≃ h ∘ g^-1. Proof. intros; split; intro Heq. - - rewrite <- Heq, assoc, iso_inv_r, right_unit; easy. - - rewrite Heq, assoc, iso_inv_l, right_unit; easy. + - rewrite <- Heq, (assoc), iso_inv_r, right_unit; easy. + - rewrite Heq, (assoc), iso_inv_l, right_unit; easy. Qed. -Lemma compose_iso_l : forall {C} {cC:Category C} {A B M : C} +Lemma compose_iso_l : forall {C} {cC:Category C} + {cCh : CategoryCoherence cC} {A B M : C} (f : A <~> B) (g : B ~> M) (h : A ~> M), f ∘ g ≃ h <-> g ≃ f^-1 ∘ h. Proof. intros; split; intro Heq. - - rewrite <- Heq, <- assoc, iso_inv_l, left_unit; easy. - - rewrite Heq, <- assoc, iso_inv_r, left_unit; easy. + - rewrite <- Heq, <- (assoc), iso_inv_l, left_unit; easy. + - rewrite Heq, <- (assoc), iso_inv_r, left_unit; easy. Qed. -Lemma compose_iso_r' : forall {C} {cC:Category C} {A B M : C} +Lemma compose_iso_r' : forall {C} {cC:Category C} + {cCh : CategoryCoherence cC} {A B M : C} (f : A ~> B) (g : B <~> M) (h : A ~> M), h ≃ f ∘ g <-> h ∘ g^-1 ≃ f. Proof. intros; split; symmetry; apply compose_iso_r; easy. Qed. -Lemma compose_iso_l' : forall {C} {cC:Category C} {A B M : C} +Lemma compose_iso_l' : forall {C} {cC:Category C} + {cCh : CategoryCoherence cC} {A B M : C} (f : A <~> B) (g : B ~> M) (h : A ~> M), h ≃ f ∘ g <-> f^-1 ∘ h ≃ g. Proof. @@ -625,7 +656,8 @@ Proof. Qed. -#[program] Definition FunctorIsomorphism {C D} {cC : Category C} {cD : Category D} +#[program] Definition FunctorIsomorphism {C D} {cC : Category C} + {cCh : CategoryCoherence cC} {cD : Category D} {cDh : CategoryCoherence cD} {A B : C} (F : Functor cC cD) (f : A <~> B) : F A <~> F B := {| forward := F @ f; reverse := F @ (f ^-1); @@ -635,8 +667,9 @@ Next Obligation. (proj2 isomorphism_inverse), 2!id_map; easy. Qed. -#[program] Definition BifunctorIsomorphism {C1 C2 D} - {cC1 : Category C1} {cC2 : Category C2} {cD : Category D} +#[program] Definition BifunctorIsomorphism {C1 C2 D} {cC1 : Category C1} + {cC1h : CategoryCoherence cC1} {cC2 : Category C2} + {cC2h : CategoryCoherence cC2} {cD : Category D} {cDh : CategoryCoherence cD} {A1 B1 : C1} {A2 B2 : C2} (F : Bifunctor cC1 cC2 cD) (f1 : A1 <~> B1) (f2 : A2 <~> B2) : F A1 A2 <~> F B1 B2 := {| forward := F @@ f1, f2; @@ -647,7 +680,8 @@ Next Obligation. 2!(proj2 isomorphism_inverse), 2!id_bimap; easy. Qed. -#[export] Instance IdentityIsomorphism {C} `{cC : Category C} (A : C) +#[export] Instance IdentityIsomorphism {C} {cC : Category C} + {cCh : CategoryCoherence cC} (A : C) : Isomorphism A A := {| forward := id_ A; reverse := id_ A; @@ -655,7 +689,8 @@ Qed. |}. Lemma component_iso_natural_reverse : forall {C D} - {cC : Category C} {cD : Category D} + {cC : Category C} {cCh : CategoryCoherence cC} + {cD : Category D} {cDh : CategoryCoherence cD} {F G : Functor cC cD} {N : NaturalIsomorphism F G} (A B : C) (f : A ~> B), N A ^-1 ∘ F.(morphism_map) f ≃ G.(morphism_map) f ∘ N B ^-1. @@ -669,9 +704,10 @@ Proof. Qed. Lemma component_biiso_natural_reverse : forall {C1 C2 D : Type} - {cC1 : Category C1} {cC2 : Category C2} - {cD : Category D} {F G : Bifunctor cC1 cC2 cD} - {N : NaturalBiIsomorphism F G} + {cC1 : Category C1} {cC1h : CategoryCoherence cC1} + {cC2 : Category C2} {cC2h : CategoryCoherence cC2} + {cD : Category D} {cDh : CategoryCoherence cD} + {F G : Bifunctor cC1 cC2 cD} {N : NaturalBiIsomorphism F G} (A1 B1 : C1) (A2 B2 : C2) (f1 : A1 ~> B1) (f2 : A2 ~> B2), (N A1 A2) ^-1 ∘ F.(morphism_bimap) f1 f2 ≃ G.(morphism_bimap) f1 f2 ∘ N B1 B2 ^-1. diff --git a/ViCaR/Classes/CategoryAltNotations.v b/ViCaR/Classes/CategoryAltNotations.v new file mode 100644 index 0000000..e596d98 --- /dev/null +++ b/ViCaR/Classes/CategoryAltNotations.v @@ -0,0 +1,30 @@ +Require CategoryTypeclass. +Import -(notations) CategoryTypeclass. +Delimit Scope Cat_scope with Cat. + +Notation "'id_' A" := (c_identity A%Cat) + (at level 15, A at next level, no associativity) : Cat_scope. +Notation "A ~> B" := (morphism _%Cat A%Cat B%Cat) + (at level 70, B at next level, no associativity) : Cat_scope. +Notation "f ≃ g" := (c_equiv _%Cat f%Cat g%Cat) + (at level 70, g at next level) : Cat_scope. (* \simeq *) +Notation "f ∘ g" := (compose _%Cat f%Cat g%Cat) + (at level 40, g at next level, left associativity) : Cat_scope. (* \circ *) + +Notation "'catI'" := (mon_I _%Cat) (at level 0) : Cat_scope. +Notation "A '∗' B" := (obj_tensor _%Cat A%Cat B%Cat) + (at level 34, left associativity) : Cat_scope. (* \ast *) +Notation "f '⧆' g" := (mor_tensor _%Cat f%Cat g%Cat) + (at level 34, left associativity) : Cat_scope . (* \otimes *) +Notation "'α_' A , B , M" := + (associator A%Cat B%Cat M%Cat) + (at level 20, no associativity) : Cat_scope. (* \alpha *) +Notation "'λ_' x" := (left_unitor x) + (at level 20, no associativity) : Cat_scope. (* \lambda \^- \^1 *) +Notation "'ρ_' x" := (right_unitor x) + (at level 20, no associativity) : Cat_scope. (* \rho \^- \^1 *) +Notation "'β_' x , y" := (braiding _%Cat x%Cat y%Cat) + (at level 20) : Cat_scope. +Notation "f '⁻¹'" := (reverse f%Cat) (at level 25) : Cat_scope. (* \^- \^1 *) + +Notation "A '<~>' B" := (Isomorphism A%Cat B%Cat) (at level 70) : Cat_scope. diff --git a/ViCaR/Classes/CategoryTypeclass.v b/ViCaR/Classes/CategoryTypeclass.v index 51197f3..541eb8a 100644 --- a/ViCaR/Classes/CategoryTypeclass.v +++ b/ViCaR/Classes/CategoryTypeclass.v @@ -1,9 +1,4 @@ Require Export Category. Require Export Monoidal. Require Export BraidedMonoidal. -Require Export SymmetricMonoidal. -Require Export Dagger. -Require Export DaggerMonoidal. -Require Export CompactClosed. -Require Export DaggerBraidedMonoidal. -Require Export DaggerSymmetricMonoidal. \ No newline at end of file +Require Export SymmetricMonoidal. \ No newline at end of file diff --git a/ViCaR/Classes/CategoryTypeclassCompatibility.v b/ViCaR/Classes/CategoryTypeclassCompatibility.v new file mode 100644 index 0000000..8a7e738 --- /dev/null +++ b/ViCaR/Classes/CategoryTypeclassCompatibility.v @@ -0,0 +1,4 @@ +Require CategoryTypeclass. +Export -(notations) CategoryTypeclass. + +Require Export CategoryAltNotations. \ No newline at end of file diff --git a/ViCaR/Classes/CompactClosed.v b/ViCaR/Classes/CompactClosed.v deleted file mode 100644 index d9cf4a0..0000000 --- a/ViCaR/Classes/CompactClosed.v +++ /dev/null @@ -1,32 +0,0 @@ -Require Import Category. -Require Import Monoidal. -Require Import BraidedMonoidal. -Require Import SymmetricMonoidal. - -#[local] Set Universe Polymorphism. - -Local Open Scope Cat. -Local Open Scope Mon. -(* Reserved Notation "A ★" (at level 0). (* \bigstar *) *) - -(* A compact closed category is a right autonomous symmetric monoidal category *) -Class CompactClosedCategory {C : Type} {cC : Category C} {mC : MonoidalCategory cC} - {bC : BraidedMonoidalCategory mC} (sC : SymmetricMonoidalCategory bC) : Type := { - dual (A : C) : C; - (* where "A ★" := (dual A); *) - unit (A : C) : I ~> dual A × A; - counit (A : C) : A × dual A ~> I; - - triangle_1 (A : C) : - (right_unitor A)^-1 ∘ (id_ A ⊗ unit A) ∘ (associator A (dual A) A)^-1 - ∘ (counit A ⊗ id_ A) ∘ left_unitor A ≃ id_ A; - - triangle_2 {A : C} : - (left_unitor (dual A))^-1 ∘ (unit A ⊗ id_ (dual A)) - ∘ (associator (dual A) A (dual A)) - ∘ (id_ (dual A) ⊗ counit A) ∘ right_unitor (dual A) ≃ id_ (dual A); -}. - -Notation "A ★" := (dual A%Obj) (at level 0) : Obj_scope. - -Local Close Scope Cat. diff --git a/ViCaR/Classes/Dagger.v b/ViCaR/Classes/Dagger.v deleted file mode 100644 index 43515d1..0000000 --- a/ViCaR/Classes/Dagger.v +++ /dev/null @@ -1,100 +0,0 @@ -Require Import Setoid. -Require Import Category. - -#[local] Set Universe Polymorphism. - -Local Open Scope Cat. - -(* Reserved Notation "f †" (at level 0). *) - -Class DaggerCategory {C : Type} (cC : Category C) : Type := { - adjoint {A B : C} (f : A ~> B) : B ~> A; - (* where "f †" := (adjoint f); *) - - adjoint_involutive {A B : C} (f : A ~> B) : - adjoint (adjoint f) ≃ f; - - adjoint_id (A : C) : adjoint (id_ A) ≃ id_ A; - - adjoint_contravariant {A B M : C} (f : A ~> B) (g : B ~> M) : - adjoint (f ∘ g) ≃ adjoint g ∘ adjoint f; - - adjoint_compat {A B : C} (f g : A ~> B) : - f ≃ g -> adjoint f ≃ adjoint g; -}. - -Notation "f †" := (adjoint f) (at level 0) : Mor_scope. (* \dag *) - -#[export] Instance AdjointContraFunctor {C} {cC : Category C} - {dagC : DaggerCategory cC} : ContraFunctor cC cC := {| - obj_contramap := fun A => A; - morphism_contramap := @adjoint C cC dagC; - id_contramap := @adjoint_id C cC dagC; - compose_contramap := @adjoint_contravariant C cC dagC; - morphism_contracompat := @adjoint_compat C cC dagC; -|}. - -Add Parametric Morphism {C} {cC : Category C} - {dagC : DaggerCategory cC} {A B : C} : - (@adjoint C cC dagC A B) - with signature - (@c_equiv C cC A B) ==> (@c_equiv C cC B A) - as adjoint_equiv_morph. -Proof. intros. apply adjoint_compat. easy. Qed. - - -Notation "'unitary' f" := (is_inverse f%Mor f%Mor †) - (at level 10, f at next level) : Cat_scope. - - - -Add Parametric Morphism {C : Type} {cC : Category C} {dagC : DaggerCategory cC} -(A B : C) : (fun f => is_inverse f f†) - with signature (@c_equiv C cC A B) ==> iff as unitary_mor. -Proof. - intros f g Hequiv. - rewrite Hequiv. - easy. -Qed. - -Lemma unitary_compose {C} {cC : Category C} {dagC : DaggerCategory cC} - {A B M : C} {f : A ~> B} {g : B ~> M} (Hf : unitary f) (Hg : unitary g) : - unitary (f ∘ g). -Proof. - split. - - rewrite adjoint_contravariant, assoc, <- (assoc g), (proj1 Hg), - left_unit; easy. - - rewrite adjoint_contravariant, assoc, <- (assoc f†), (proj2 Hf), - left_unit; easy. -Qed. - -Lemma unitary_adjoint {C} {cC : Category C} {dagC : DaggerCategory cC} - {A B : C} {f : A ~> B} (Hf : unitary f) : unitary (f†). -Proof. - rewrite adjoint_involutive. - split; easy. -Qed. - -Lemma unitary_inverse {C} {cC : Category C} {dagC : DaggerCategory cC} - {A B : C} {f : A ~> B} {g : B ~> A} (Hf : unitary f) (Hinv : is_inverse f g) : - unitary g. -Proof. - rewrite (inverse_unique Hinv Hf). - apply unitary_adjoint; easy. -Qed. - -Class UnitaryIsomorphism {C : Type} {cC : Category C} {dagC : DaggerCategory cC} - (A B : C) := { - unitary_isomorphism : A <~> B; - is_unitary_isomorphism : unitary unitary_isomorphism.(forward) -}. -Coercion unitary_isomorphism : UnitaryIsomorphism >-> Isomorphism. - -Lemma is_unitary_isomorphism_rev {C} {cC : Category C} {dagC : DaggerCategory cC} - {A B : C} (f : UnitaryIsomorphism A B) : unitary f.(reverse). -Proof. - apply (unitary_inverse is_unitary_isomorphism isomorphism_inverse). -Qed. - - -Local Close Scope Cat. diff --git a/ViCaR/Classes/DaggerBraidedMonoidal.v b/ViCaR/Classes/DaggerBraidedMonoidal.v deleted file mode 100644 index b328c04..0000000 --- a/ViCaR/Classes/DaggerBraidedMonoidal.v +++ /dev/null @@ -1,17 +0,0 @@ -Require Import Category. -Require Import Dagger. -Require Import Monoidal. -Require Import BraidedMonoidal. -Require Import DaggerMonoidal. - -#[local] Set Universe Polymorphism. - -Local Open Scope Cat. - -Class DaggerBraidedMonoidalCategory {C : Type} {cC : Category C} - {dagC : DaggerCategory cC} {mC : MonoidalCategory cC} - (mdagC : DaggerMonoidalCategory dagC mC) (bC : BraidedMonoidalCategory mC) - : Type := {}. - - -Local Close Scope Cat. \ No newline at end of file diff --git a/ViCaR/Classes/DaggerMonoidal.v b/ViCaR/Classes/DaggerMonoidal.v deleted file mode 100644 index 91f40a9..0000000 --- a/ViCaR/Classes/DaggerMonoidal.v +++ /dev/null @@ -1,40 +0,0 @@ -Require Import Category. -Require Import Dagger. -Require Import Monoidal. - -#[local] Set Universe Polymorphism. - -Local Open Scope Cat. -Local Open Scope Mon. - -Class DaggerMonoidalCategory {C : Type} {cC : Category C} - (dagC : DaggerCategory cC) (mC : MonoidalCategory cC) : Type := { - dagger_tensor_compat {A B M N : C} (f : A ~> M) (g : B ~> N) : - f † ⊗ g † ≃ (f ⊗ g) †; - - associator_unitary {A B M : C} : - unitary (associator A B M); - left_unitor_unitary {A : C} : - unitary (left_unitor A); - - right_unitor_unitary {A : C} : - unitary (right_unitor A); - - - (* associator_unitary_r {A B M : C} : - associator ∘ associator † ≃ c_identity (A × B × M); - associator_unitary_l {A B M : C} : - associator † ∘ associator ≃ c_identity (A × (B × M)); - - left_unitor_unitary_r {A : C} : - left_unitor ∘ left_unitor † ≃ c_identity (I × A); - left_unitor_unitary_l {A : C} : - left_unitor † ∘ left_unitor ≃ c_identity A; - - right_unitor_unitary_r {A : C} : - right_unitor ∘ right_unitor † ≃ c_identity (A × I); - right_unitor_unitary_l {A : C} : - right_unitor † ∘ right_unitor ≃ c_identity A; *) -}. - -Local Close Scope Cat. diff --git a/ViCaR/Classes/DaggerSymmetricMonoidal.v b/ViCaR/Classes/DaggerSymmetricMonoidal.v deleted file mode 100644 index 0091908..0000000 --- a/ViCaR/Classes/DaggerSymmetricMonoidal.v +++ /dev/null @@ -1,21 +0,0 @@ -Require Import Category. -Require Import Dagger. -Require Import Monoidal. -Require Import BraidedMonoidal. -Require Import DaggerMonoidal. -Require Import SymmetricMonoidal. -Require Import DaggerBraidedMonoidal. - -#[local] Set Universe Polymorphism. - -Local Open Scope Cat. - -Class DaggerSymmetricMonoidalCategory {C : Type} {cC : Category C} - {dagC : DaggerCategory cC} {mC : MonoidalCategory cC} - {bC : BraidedMonoidalCategory mC} - {mdagC : DaggerMonoidalCategory dagC mC} - (bdagC : DaggerBraidedMonoidalCategory mdagC bC) - (sC : SymmetricMonoidalCategory bC) : Type := {}. - - -Local Close Scope Cat. \ No newline at end of file diff --git a/ViCaR/Classes/Monoidal.v b/ViCaR/Classes/Monoidal.v index 4b617a3..d1ff429 100644 --- a/ViCaR/Classes/Monoidal.v +++ b/ViCaR/Classes/Monoidal.v @@ -4,214 +4,183 @@ Require Import Setoid. #[local] Set Universe Polymorphism. Local Open Scope Cat_scope. -Local Open Scope Mor_scope. -Local Open Scope Obj_scope. -Declare Scope Mon_scope. -Delimit Scope Mon_scope with Mon. -Open Scope Mon_scope. +Reserved Notation "x × y" (at level 34, left associativity). (* \times *) +Reserved Notation "x ⊗ y" (at level 34, left associativity). (* \otimes *) +Reserved Notation "'λ_' x" (at level 20, no associativity). (* \lambda *) +Reserved Notation "'ρ_' x" (at level 20, no associativity). (* \rho *) +Reserved Notation "'α_' A , B , M" (at level 20, no associativity). (* \alpha *) - -(* -Removing the reservation to allow confining these to a particular scope. - -Reserved Notation "x × y" (at level 40, left associativity). (* \times *) -Reserved Notation "x ⊗ y" (at level 40, left associativity). (* \otimes *) -Reserved Notation "'λ_' x" (at level 30, no associativity). (* \lambda *) -Reserved Notation "'ρ_' x" (at level 30, no associativity). (* \rho *) -*) Class MonoidalCategory {C : Type} (cC : Category C) : Type := { - tensor : Bifunctor cC cC cC; - (* where "x × y" := (tensor x y); *) + obj_tensor : C -> C -> C + where "x × y" := (obj_tensor x y); + mor_tensor {A1 B1 A2 B2 : C} (f : A1 ~> B1) (g : A2 ~> B2) : + A1 × A2 ~> B1 × B2; mon_I : C; - (* where "x ⊗ y" := (tensor @@ x, y) ;*) associator (A B M : C) : - tensor (tensor A B) M <~> tensor A (tensor B M); + (A × B) × M <~> A × (B × M); - left_unitor (A : C) : tensor mon_I A <~> A; - (* where "'λ_' x" := (left_unitor x); *) + left_unitor (A : C) : mon_I × A <~> A; - right_unitor (A : C) : tensor A mon_I <~> A; - (* where "'ρ_' x" := (right_unitor x); *) + right_unitor (A : C) : A × mon_I <~> A; +}. +Arguments obj_tensor {_} {_}%Cat (mC)%Cat (A B)%Cat : rename. +Arguments mor_tensor {_} {_}%Cat (mC)%Cat {_ _ _ _}%Cat (_ _)%Cat : rename. +Arguments mon_I {_} {_}%Cat (mC)%Cat : rename. +Arguments associator {_} {_}%Cat {mC}%Cat (_ _ _)%Cat : rename. +Arguments left_unitor {_} {_}%Cat {mC}%Cat (_)%Cat : rename. +Arguments right_unitor {_} {_}%Cat {mC}%Cat (_)%Cat : rename. + +Notation "'I'" := (mon_I _%Cat) (at level 0) : Cat_scope. +Notation "A '×' B" := (obj_tensor _%Cat A%Cat B%Cat) + (at level 34, left associativity) : Cat_scope. (* \times *) +Notation "f '⊗' g" := + (mor_tensor _%Cat f%Cat g%Cat) + (at level 34, left associativity) : Cat_scope . (* \otimes *) +Notation "'α_' A , B , M" := + (associator A%Cat B%Cat M%Cat) + (at level 20, no associativity) : Cat_scope. (* \alpha *) +Notation "'λ_' x" := (left_unitor x) + (at level 20, no associativity) : Cat_scope. (* \lambda \^- \^1 *) +Notation "'ρ_' x" := (right_unitor x) + (at level 20, no associativity) : Cat_scope. (* \rho \^- \^1 *) +(* Notation "'α⁻¹_' A , B , M" := + ((associator A%Cat B%Cat M%Cat).(reverse)) + (at level 20, no associativity) : Cat_scope. (* \alpha \^- \^1 *) +Notation "'λ⁻¹_' x" := ((left_unitor x).(reverse)) + (at level 20, no associativity) : Cat_scope. (* \lambda \^- \^1 *) +Notation "'ρ⁻¹_' x" := ((right_unitor x).(reverse)) + (at level 20, no associativity) : Cat_scope. (* \rho \^- \^1 *) + *) + + +Class MonoidalCategoryCoherence {C : Type} {cC : Category C} + {cCh : CategoryCoherence cC} (mC : MonoidalCategory cC) : Type := { + tensor_id (A1 A2 : C) : (id_ A1) ⊗ (id_ A2) ≃ id_ (A1 × A2); + tensor_compose {A1 B1 M1 A2 B2 M2 : C} + (f1 : A1 ~> B1) (g1 : B1 ~> M1) + (f2 : A2 ~> B2) (g2 : B2 ~> M2) : + (f1 ∘ g1) ⊗ (f2 ∘ g2) ≃ f1 ⊗ f2 ∘ g1 ⊗ g2; + tensor_compat {A1 B1 A2 B2 : C} + (f f' : A1 ~> B1) (g g' : A2 ~> B2) : + f ≃ f' -> g ≃ g' -> f ⊗ g ≃ f' ⊗ g'; (* Coherence conditions for α, λ, ρ *) associator_cohere {A B M N P Q : C} (f : A ~> B) (g : M ~> N) (h : P ~> Q) : - associator A M P ∘ (tensor @@ f, (tensor @@ g, h)) - ≃ (tensor @@ (tensor @@ f, g), h) ∘ associator B N Q; + α_ A, M, P ∘ (f ⊗ (g ⊗ h)) + ≃ ((f ⊗ g) ⊗ h) ∘ α_ B, N, Q; left_unitor_cohere {A B : C} (f : A ~> B) : - left_unitor A ∘ f ≃ (tensor @@ id_ mon_I, f) ∘ left_unitor B; + λ_ A ∘ f ≃ (id_ I ⊗ f) ∘ λ_ B; right_unitor_cohere {A B : C} (f : A ~> B) : - right_unitor A ∘ f ≃ (tensor @@ f, id_ mon_I) ∘ right_unitor B; + ρ_ A ∘ f ≃ (f ⊗ id_ I) ∘ ρ_ B; (* Commutative diagrams *) triangle (A B : C) : - associator A mon_I B ∘ (tensor @@ id_ A, left_unitor B) - ≃ tensor @@ right_unitor A, id_ B; + α_ A, I, B ∘ (id_ A ⊗ λ_ B) + ≃ ρ_ A ⊗ id_ B; pentagon (A B M N : C) : - (tensor @@ associator A B M, id_ N) - ∘ associator A (tensor B M) N - ∘ (tensor @@ id_ A, associator B M N) - ≃ associator (tensor A B) M N ∘ associator A B (tensor M N); + (α_ A, B, M ⊗ id_ N) + ∘ α_ A, (B × M), N + ∘ (id_ A ⊗ α_ B, M, N) + ≃ α_ (A × B), M, N ∘ α_ A, B, (M × N); }. -Arguments MonoidalCategory {_} (_)%Cat. -Arguments tensor {_} {_}%Cat (mC)%Cat : rename. -Arguments mon_I {_} {_}%Cat (mC)%Cat : rename. -Arguments associator {_} {_}%Cat {mC}%Cat (_ _ _)%Obj : rename. -Arguments left_unitor {_} {_}%Cat {mC}%Cat (_)%Obj : rename. -Arguments right_unitor {_} {_}%Cat {mC}%Cat (_)%Obj : rename. -Arguments associator_cohere {_} {_}%Cat {mC}%Cat - {_ _ _ _ _ _}%Obj (_ _ _)%Mor : rename. -Arguments left_unitor_cohere {_} {_}%Cat {mC}%Cat {_ _}%Obj (_)%Mor : rename. -Arguments right_unitor_cohere {_} {_}%Cat {mC}%Cat {_ _}%Obj (_)%Mor : rename. -Arguments triangle {_} {_}%Cat {mC}%Cat (_ _)%Obj: rename. -Arguments pentagon {_} {_}%Cat {mC}%Cat (_ _ _ _)%Obj : rename. - -Notation "'I'" := (mon_I _%Cat) (at level 0) : Mon_scope. -Notation "A '×' B" := (tensor _%Cat A%Obj B%Obj) - (at level 40, left associativity) : Mon_scope. (* \times *) -Notation "f '⊗' g" := - (morphism_bimap (tensor _%Cat) f%Obj g%Obj) - (at level 40, left associativity) : Mon_scope . (* \otimes *) -Notation "'λ_' x" := (left_unitor x) - (at level 20, no associativity) : Mon_scope. (* \lambda *) -Notation "'ρ_' x" := (right_unitor x) - (at level 20, no associativity) : Mon_scope. (* \rho *) -Add Parametric Morphism {C : Type} - {cC : Category C} {mC: MonoidalCategory cC} - (A1 B1 A2 B2 : C) : mC.(tensor).(morphism_bimap) +Arguments associator_cohere {_} {_ _ _}%Cat {mCh}%Cat + {_ _ _ _ _ _}%Cat (_ _ _)%Cat : rename. +Arguments left_unitor_cohere {_} {_ _ _}%Cat {mCh}%Cat {_ _}%Cat (_)%Cat : rename. +Arguments right_unitor_cohere {_} {_ _ _}%Cat {mCh}%Cat {_ _}%Cat (_)%Cat : rename. +Arguments triangle {_} {_ _ _}%Cat {mCh}%Cat (_ _)%Cat: rename. +Arguments pentagon {_} {_ _ _}%Cat {mCh}%Cat (_ _ _ _)%Cat : rename. + +Add Parametric Morphism {C : Type} {cC : Category C} + {mC : MonoidalCategory cC} {cCh : CategoryCoherence cC} + {mCh : MonoidalCategoryCoherence mC} {A1 B1 A2 B2 : C} : (mor_tensor mC) with signature (cC.(c_equiv) (A:=A1) (B:=B1)) ==> (cC.(c_equiv) (A:=A2) (B:=B2)) ==> - cC.(c_equiv) as stack_equiv_mor. + cC.(c_equiv) as mor_tensor_equiv_mor. +Proof. intros. apply tensor_compat; assumption. Qed. + +Section TensorBifunctor. + +Context {C : Type} {cC : Category C} (mC : MonoidalCategory cC) + {cCh : CategoryCoherence cC} {mCh : MonoidalCategoryCoherence mC}. + +#[export, program] Instance TensorIsomorphism {A1 B1 A2 B2} + (f1 : A1 <~> B1) (f2 : A2 <~> B2) : A1 × A2 <~> B1 × B2 := { + forward := f1 ⊗ f2; + reverse := f1^-1 ⊗ f2^-1; +}. +Next Obligation. + rewrite <- 2!(tensor_compose), 2!iso_inv_r, 2!iso_inv_l, 2!tensor_id. + easy. +Qed. + +#[export] Instance tensor : Bifunctor cC cC cC := { + obj_bimap := mC.(obj_tensor); + morphism_bimap := fun A1 B1 B1 B2 => mC.(mor_tensor); + id_bimap := tensor_id; + compose_bimap := fun A1 B1 M1 A2 B2 M2 => tensor_compose; + morphism_bicompat := fun A1 B1 A2 B2 => tensor_compat; +}. + +Add Parametric Morphism (A1 B1 A2 B2 : C) : tensor.(morphism_bimap) + with signature + (cC.(c_equiv) (A:=A1) (B:=B1)) ==> + (cC.(c_equiv) (A:=A2) (B:=B2)) ==> + cC.(c_equiv) as tensor_equiv_mor. Proof. intros. apply morphism_bicompat; assumption. Qed. -Add Parametric Morphism {C : Type} - {cC : Category C} {mC : MonoidalCategory cC} : mC.(tensor).(obj_bimap) +Add Parametric Morphism : tensor.(obj_bimap) with signature (@isomorphic C cC) ==> (@isomorphic C cC) ==> @isomorphic C cC as stack_isomorphic_mor. Proof. intros A B [fAB [fBA [HfAB HfBA]]] M N [fMN [fNM [HfMN HfNM]]]. exists (fAB ⊗ fMN); exists (fBA ⊗ fNM). - rewrite <- 2!compose_bimap, HfAB, HfBA, HfMN, HfNM. - rewrite 2!id_bimap; easy. + simpl. + rewrite <- 2!(tensor_compose), HfAB, HfBA, HfMN, HfNM. + rewrite 2!(tensor_id); easy. Qed. -Fixpoint n_times_r {C} {cC : Category C} {mC : MonoidalCategory cC} - (n : nat) (A : C) : C := - match n with - | O => I - | S n' => A × (n_times_r n' A) - end. - -Fixpoint n_times_l {C} {cC : Category C} {mC : MonoidalCategory cC} - (n : nat) (A : C) : C := - match n with - | O => I - | S n' => (n_times_l n' A) × A - end. - -Fixpoint n_tensor_r {C} {cC : Category C} {mC : MonoidalCategory cC} - {A B : C} (n : nat) (f : A ~> B) - : (n_times_r n A ~> n_times_r n B) := - match n with - | O => id_ I - | S n' => f ⊗ (n_tensor_r n' f) - end. - -Fixpoint n_tensor_l {C} {cC : Category C} {mC : MonoidalCategory cC} - {A B : C} (n : nat) (f : A ~> B) - : (n_times_l n A ~> n_times_l n B) := - match n with - | O => id_ I - | S n' => (n_tensor_l n' f) ⊗ f - end. - -Arguments n_times_r {_} {_ _}%Cat _ _%Obj. -Arguments n_times_l {_} {_ _}%Cat _ _%Obj. -Arguments n_tensor_r {_} {_ _}%Cat {_ _}%Obj _ _%Mor. -Arguments n_tensor_l {_} {_ _}%Cat {_ _}%Obj _ _%Mor. - -Add Parametric Morphism {C : Type} - {cC : Category C} {mC : MonoidalCategory cC} : (@n_times_r C cC mC) - with signature - (@eq nat) ==> (@isomorphic C cC) ==> (@isomorphic C cC) - as n_times_r_isomorphic_mor. -Proof. - intros n A B [fAB [fBA HAB]]. - exists (n_tensor_r n fAB); exists (n_tensor_r n fBA). - induction n; simpl. - - rewrite left_unit; split; easy. - - rewrite <- 2!compose_bimap, (proj1 HAB), (proj2 HAB), - (proj1 IHn), (proj2 IHn), 2!id_bimap. - split; easy. -Qed. +End TensorBifunctor. -Add Parametric Morphism {C : Type} - {cC : Category C} {mC : MonoidalCategory cC} : (@n_times_l C cC mC) - with signature - (@eq nat) ==> (@isomorphic C cC) ==> (@isomorphic C cC) - as n_times_l_isomorphic_mor. -Proof. - intros n A B [fAB [fBA HAB]]. - exists (n_tensor_l n fAB); exists (n_tensor_l n fBA). - induction n; simpl. - - rewrite left_unit; split; easy. - - rewrite <- 2!compose_bimap, (proj1 HAB), (proj2 HAB), - (proj1 IHn), (proj2 IHn), 2!id_bimap. - split; easy. -Qed. +Arguments tensor {_} {_}%Cat (mC)%Cat {_ _}%Cat. +Arguments TensorIsomorphism {_} {_ mC cCh mCh}%Cat {_ _ _ _}%Cat (_ _)%Cat. -Add Parametric Morphism {C : Type} - {cC : Category C} {mC : MonoidalCategory cC} {A B : C} {n : nat} : - (@n_tensor_r C cC mC A B n) - with signature - (@c_equiv C cC A B) - ==> (@c_equiv C cC (n_times_r n A) (n_times_r n B)) - as n_tensor_r_equiv_mor. +Section TensorIsomorphismRewrites. + +Context {C : Type} {cC : Category C} {mC : MonoidalCategory cC} + {cCh : CategoryCoherence cC} {mCh : MonoidalCategoryCoherence mC}. + +Lemma tensor_cancel_l : forall {A1 B1 A2 B2} (f : A1 ~> B1) (g g' : A2 ~> B2), + g ≃ g' -> f ⊗ g ≃ f ⊗ g'. Proof. - intros f g Hfg. - induction n. - - easy. - - apply morphism_bicompat; assumption. + intros; apply tensor_compat; easy. Qed. -Add Parametric Morphism {C : Type} - {cC : Category C} {mC : MonoidalCategory cC} {A B : C} {n : nat} : - (@n_tensor_l C cC mC A B n) - with signature - (@c_equiv C cC A B) - ==> (@c_equiv C cC (n_times_l n A) (n_times_l n B)) - as n_tensor_l_equiv_mor. +Lemma tensor_cancel_r : forall {A1 B1 A2 B2} (f f' : A1 ~> B1) (g : A2 ~> B2), + f ≃ f' -> f ⊗ g ≃ f' ⊗ g. Proof. - intros f g Hfg. - induction n. - - easy. - - apply morphism_bicompat; assumption. + intros; apply tensor_compat; easy. Qed. - -Lemma compose_tensor_iso_r : forall {C} {cC : Category C} - {mC : MonoidalCategory cC} - {A B1 M1 B2 M2 : C} (f : A ~> B1 × B2) +Lemma compose_tensor_iso_r : forall {A B1 M1 B2 M2 : C} (f : A ~> B1 × B2) (g1 : B1 <~> M1) (g2 : B2 <~> M2) (h : A ~> M1 × M2), f ∘ g1⊗g2 ≃ h <-> f ≃ h ∘ (g1^-1 ⊗ g2^-1). Proof. intros; split; intro Heq. - - rewrite <- Heq, assoc, <- compose_bimap, - 2!iso_inv_r, id_bimap, right_unit; easy. - - rewrite Heq, assoc, <- compose_bimap, - 2!iso_inv_l, id_bimap, right_unit; easy. + - rewrite <- Heq, (assoc), <- (tensor_compose), + 2!iso_inv_r, (tensor_id), (right_unit); easy. + - rewrite Heq, (assoc), <- (tensor_compose), + 2!iso_inv_l, (tensor_id), (right_unit); easy. Qed. -Lemma compose_tensor_iso_r' : forall {C} {cC : Category C} - {mC : MonoidalCategory cC} - {A B1 M1 B2 M2 : C} (f : A ~> B1 × B2) +Lemma compose_tensor_iso_r' : forall {A B1 M1 B2 M2 : C} (f : A ~> B1 × B2) (g1 : B1 <~> M1) (g2 : B2 <~> M2) (h : A ~> M1 × M2), h ≃ f ∘ g1⊗g2 <-> h ∘ (g1^-1 ⊗ g2^-1) ≃ f. Proof. @@ -219,30 +188,26 @@ Proof. split; symmetry; apply compose_tensor_iso_r; easy. Qed. -Lemma compose_tensor_iso_l : forall {C} {cC : Category C} - {mC : MonoidalCategory cC} - {A1 B1 M A2 B2 : C} (f1 : A1 <~> B1) (f2 : A2 <~> B2) +Lemma compose_tensor_iso_l : forall {A1 B1 M A2 B2 : C} + (f1 : A1 <~> B1) (f2 : A2 <~> B2) (g : B1 × B2 ~> M) (h : A1 × A2 ~> M), f1⊗f2 ∘ g ≃ h <-> g ≃ (f1^-1 ⊗ f2^-1) ∘ h. Proof. intros; split; intro Heq. - - rewrite <- Heq, <- assoc, <- compose_bimap, - 2!iso_inv_l, id_bimap, left_unit; easy. - - rewrite Heq, <- assoc, <- compose_bimap, - 2!iso_inv_r, id_bimap, left_unit; easy. + - rewrite <- Heq, <- (assoc), <- (tensor_compose), + 2!iso_inv_l, (tensor_id), (left_unit); easy. + - rewrite Heq, <- (assoc), <- (tensor_compose), + 2!iso_inv_r, (tensor_id), (left_unit); easy. Qed. -Lemma compose_tensor_iso_l' : forall {C} {cC : Category C} - {mC : MonoidalCategory cC} - {A1 B1 M A2 B2 : C} (f1 : A1 <~> B1) (f2 : A2 <~> B2) +Lemma compose_tensor_iso_l' : forall {A1 B1 M A2 B2 : C} + (f1 : A1 <~> B1) (f2 : A2 <~> B2) (g : B1 × B2 ~> M) (h : A1 × A2 ~> M), h ≃ f1⊗f2 ∘ g <-> (f1^-1 ⊗ f2^-1) ∘ h ≃ g. Proof. intros; split; symmetry; apply compose_tensor_iso_l; easy. Qed. +End TensorIsomorphismRewrites. -Local Close Scope Mon_scope. -Local Close Scope Obj_scope. -Local Close Scope Mor_scope. Local Close Scope Cat_scope. diff --git a/ViCaR/Classes/RigCategory.v b/ViCaR/Classes/RigCategory.v deleted file mode 100644 index 8c89a62..0000000 --- a/ViCaR/Classes/RigCategory.v +++ /dev/null @@ -1,972 +0,0 @@ -Require Import Setoid. - -Require Import Category. -Require Import Monoidal. -Require Import BraidedMonoidal. -Require Import SymmetricMonoidal. - -#[local] Set Universe Polymorphism. - -Require Import CategoryAutomation. -Local Close Scope Mon_scope. - -Local Open Scope Cat_scope. - -Declare Scope Rig_scope. -Delimit Scope Rig_scope with Rig. -Open Scope Rig_scope. - -Class PreDistributiveBimonoidalCategory {C : Type} {cC : Category C} - {mAddC : MonoidalCategory cC} {bAddC : BraidedMonoidalCategory mAddC} - (AddC : SymmetricMonoidalCategory bAddC) - (MulC : MonoidalCategory cC) := { - addC := AddC; - mulC := MulC; - (* left_distr_iso (A B M : C) : - A × (B + M) <~> A×B + A×M; - left_distr_natural {A B M A' B' M' : C} : forall - (f : A ~> A') (g : B ~> B') (h : M ~> M'), - left_distr_iso A B M ∘ (f ⊠ g ⊞ f ⊠ h) - ≃ (f ⊠ (g ⊞ h)) ∘ left_distr_iso A' B' M'; *) - - (* NOTE: By not defining these as actual natural transformations, we - *should* only be 'losing' the declaration that the underlying - transformations are functorial... but as compositions of functors, - they are — plus, it's awkward to get the definitions working nicely - in functor land. So for now, we do it directly. *) - left_distr_iso (A B M : C) : - MulC.(tensor) A (AddC.(tensor) B M) - <~> AddC.(tensor) (MulC.(tensor) A B) (MulC.(tensor) A M); - left_distr_natural {A B M A' B' M' : C} : forall - (f : A ~> A') (g : B ~> B') (h : M ~> M'), - left_distr_iso A B M ∘ - (AddC.(tensor) @@ (MulC.(tensor) @@ f, g), (MulC.(tensor) @@ f, h)) - ≃ MulC.(tensor) @@ f, (AddC.(tensor) @@ g, h) - ∘ left_distr_iso A' B' M'; - - right_distr_iso (A B M : C) : - MulC.(tensor) (AddC.(tensor) A B) M - <~> AddC.(tensor) (MulC.(tensor) A M) (MulC.(tensor) B M); - right_distr_natural {A B M A' B' M' : C} : forall - (f : A ~> A') (g : B ~> B') (h : M ~> M'), - right_distr_iso A B M ∘ - (AddC.(tensor) @@ (MulC.(tensor) @@ f, h), (MulC.(tensor) @@ g, h)) - ≃ MulC.(tensor) @@ (AddC.(tensor) @@ f, g), h - ∘ right_distr_iso A' B' M'; -}. - -Arguments PreDistributiveBimonoidalCategory {_} {_ _ _}%Cat (_ _)%Cat. -Arguments left_distr_iso {_} {_ _ _ _ _ _}%Cat (_ _ _)%Rig. -Arguments right_distr_iso {_} {_ _ _ _ _ _}%Cat (_ _ _)%Rig. - -Close Scope Mon_scope. -Close Scope Brd_scope. -Open Scope Mor_scope. - - -Notation "'addC_m'" := (addC -.(BraidedMonoidalCategory_of_SymmetricMonoidalCategory) -.(MonoidalCategory_of_BraidedMonoidalCategory)). - -Notation "'addC_b'" := (addC -.(BraidedMonoidalCategory_of_SymmetricMonoidalCategory)). - - -Notation "'cplus'" := (tensor addC_m). - -Notation "'cmul'" := (tensor mulC). - -Notation "A + B" := (cplus A B) - (at level 50, left associativity) : Rig_scope. -Notation "A × B" := (cmul A B) - (at level 40, left associativity) : Rig_scope. - -Notation "f ⊞ g" := (addC - .(BraidedMonoidalCategory_of_SymmetricMonoidalCategory) - .(MonoidalCategory_of_BraidedMonoidalCategory) - .(tensor) @@ f, g) - (at level 50, left associativity) : Rig_scope. -Notation "f ⊠ g" := (mulC.(tensor) @@ f, g) - (at level 40, left associativity) : Rig_scope. - - -Notation "'c1'" := (mon_I mulC) - (at level 0, no associativity) : Rig_scope. -Notation "'c0'" := (mon_I addC_m) - (at level 0, no associativity) : Rig_scope. - -Notation "'α_' A , B , M" := - (reverse (mulC.(associator) A B M)) - (at level 20, no associativity) : Rig_scope. -Notation "'α'_' A , B , M" := - (reverse (addC_m.(associator) A B M)) - (at level 20, no associativity) : Rig_scope. -Notation "'α^-1_' A , B , M" := - (mulC.(associator) A B M).(forward) - (at level 20, no associativity) : Rig_scope. -Notation "'α'^-1_' A , B , M" := - (addC_m.(associator) A B M).(forward) - (at level 20, no associativity) : Rig_scope. - - -Notation "'δ_' A , B , M" := - (left_distr_iso A B M) - (at level 20, no associativity) : Rig_scope. -Notation "'δ#_' A , B , M" := - (right_distr_iso A B M) - (at level 20, no associativity) : Rig_scope. - -Notation "'γ'_' A , B" := - (braiding (addC.(BraidedMonoidalCategory_of_SymmetricMonoidalCategory)) A B) - (at level 35, A at next level, B at next level, no associativity) : Rig_scope. -(* NOTE! This one won't work unless in a Braided Rig: *) -Notation "'γ_' A , B" := - (braiding (mC:=mulC) _ A B) - (at level 35, A at next level, B at next level, no associativity) : Rig_scope. - -Notation "'λ_' A" := - (mulC.(left_unitor) A) - (at level 20, no associativity) : Rig_scope. -Notation "'λ'_' A" := - (addC_m.(left_unitor) A) - (at level 20, no associativity) : Rig_scope. - -Notation "'ρ_' A" := - (mulC.(right_unitor) A) - (at level 20, no associativity) : Rig_scope. -Notation "'ρ'_' A" := - (addC_m.(right_unitor) A) - (at level 20, no associativity) : Rig_scope. - -Class DistributiveBimonoidalCategory {C : Type} {cC : Category C} - {mAddC : MonoidalCategory cC} {bAddC : BraidedMonoidalCategory mAddC} - {AddC : SymmetricMonoidalCategory bAddC} - {MulC : MonoidalCategory cC} - (pdC : PreDistributiveBimonoidalCategory AddC MulC) := { - predistr_cat := pdC; - - left_absorbtion_iso (A : C) : - c0 × A <~> c0; - left_absorbtion_natural {A A' : C} : forall (f : A ~> A'), - left_absorbtion_iso A ∘ (id_ c0) - ≃ (id_ c0) ⊠ f ∘ left_absorbtion_iso A'; - - right_absorbtion_iso (A : C) : - A × c0 <~> c0; - right_absorbtion_natural {A A' : C} : forall (f : A ~> A'), - right_absorbtion_iso A ∘ (id_ c0) - ≃ f ⊠ (id_ c0) ∘ right_absorbtion_iso A'; -}. - -Coercion predistr_cat : - DistributiveBimonoidalCategory >-> PreDistributiveBimonoidalCategory. - -Notation "'λ*_' A" := - (left_absorbtion_iso A) - (at level 20, no associativity) : Rig_scope. -Notation "'ρ*_' A" := - (right_absorbtion_iso A) - (at level 20, no associativity) : Rig_scope. - - -(* Notation "'[+]'" := (addC) (at level 0) : Rig_scope. -Notation "'[×]'" := (addC) (at level 0) : Rig_scope. *) -(* -Section CoherenceConditions. - -Context {DD : Type} {cC : Category DD} - {mAddC : MonoidalCategory cC} {bAddC : BraidedMonoidalCategory mAddC} - {AddC : SymmetricMonoidalCategory bAddC} - {MulC : MonoidalCategory cC} - (pdC : PreDistributiveBimonoidalCategory AddC MulC) - (rC : DistributiveBimonoidalCategory pdC). - - -Definition condition_I := forall (A B C : DD), - δ_ A,B,C ∘ γ'_ (A×B), (A×C) ≃ - id_ A ⊠ γ'_ B, C ∘ δ_ A, C, B. - -Definition condition_III := forall (A B C : DD), - δ#_ A,B,C ∘ γ'_ (A×C),(B×C) - ≃ (γ'_ A,B ⊠ id_ C) ∘ δ#_ B,A,C. - -Definition condition_IV := forall (A B C D : DD), - (* c_equiv (A:=(A+(B+C))×D) (B:=(A×D + B×D + C×D)) *) - δ#_ A,(B+C),D ∘ (id_ (A×D) ⊞ δ#_ B,C,D) ∘ α'_ A×D, B×D, (C×D) - ≃ (α'_ A,B,C ⊠ (id_ D)) ∘ δ#_ A+B,C,D ∘ (δ#_ A,B,D ⊞ id_ (C×D)). - -Definition condition_V := forall (A B C D : DD), - δ_ A,B,(C+D) ∘ (id_(A×B) ⊞ δ_ A,C,D) ∘ α'_ A×B,A×C,(A×D) - ≃ id_ A ⊠ α'_ B,C,D ∘ δ_ A,B+C,D ∘ (δ_ A,B,C ⊞ id_(A×D)). - -Definition condition_VI := forall (A B C D : DD), - id_ A ⊠ δ_ B,C,D ∘ δ_ A,B×C,(B×D) ∘ (α_ A,B,C ⊞ α_ A,B,D) - ≃ α_ A,B,(C+D) ∘ δ_ (A×B),C,D. - -Definition condition_VII := forall (A B C D : DD), - δ#_ A,B,(C×D) ∘ (α_ A,C,D ⊞ α_ B,C,D) - ≃ α_ A+B,C,D ∘ (δ#_A,B,C ⊠ id_ D) ∘ (δ#_ A×C,B×C,D). - -Definition condition_VIII := forall (A B C D : DD), - (id_ A ⊠ δ#_ B,C,D) ∘ δ_ A,B×D,(C×D) ∘ (α_ A,B,D ⊞ α_ A,C,D) - ≃ α_ A,B+C,D ∘ (δ_ A,B,C ⊠ id_ D) ∘ δ#_ A×B,A×C,D. - -Definition condition_IX := forall (A B C D : DD), - δ#_ A,B,(C+D) ∘ (δ_ A,C,D ⊞ δ_ B,C,D) ∘ (α'_ (A×C+A×D),B×C,(B×D)) - ∘ ((α'^-1_ A×C,A×D,(B×C)) ⊞ id_ (B×D)) - ∘ ((id_ (A×C) ⊞ γ'_ A×D, (B×C)) ⊞ id_ (B×D)) - ∘ (α'_ A×C,B×C,(A×D) ⊞ id_ (B×D)) - ≃ δ_ A+B,C,D ∘ (δ#_ A,B,C ⊞ δ#_ A,B,D) ∘ α'_ (A×C+B×C),A×D,(B×D). - -Definition condition_X := - λ*_ c0 ≃ ρ*_ c0. - -Definition condition_XI := forall (A B : DD), - δ_ c0,A,B ∘ (λ*_ A ⊞ λ*_ B) ∘ λ'_ c0 - ≃ λ*_ (A + B). - -Definition condition_XII := forall (A B : DD), - δ#_ A,B,c0 ∘ (ρ*_ A ⊞ ρ*_ B) ∘ (λ'_ c0) - ≃ ρ*_ (A+B). - -Definition condition_XIII := - ρ_ c0 ≃ λ*_ c1. - -Definition condition_XIV := - λ_ c0 ≃ ρ*_ c1. - -Definition condition_XVI := forall (A B : DD), - α_ c0,A,B ∘ (λ*_ A ⊠ id_ B) ∘ λ*_B - ≃ λ*_ (A×B). - -Definition condition_XVII := forall (A B : DD), - α_ A,c0,B ∘ (ρ*_ A ⊠ id_ B) ∘ λ*_ B - ≃ id_ A ⊠ λ*_ B ∘ ρ*_ A. - -Definition condition_XVIII := forall (A B : DD), - α_ A,B,c0 ∘ ρ*_ (A×B) - ≃ (id_ A ⊠ ρ*_ B) ∘ ρ*_ A. - -Definition condition_XIX := forall (A B : DD), - δ_ A,c0,B ∘ (ρ*_ A ⊞ id_ (A×B)) ∘ λ'_ (A×B) - ≃ id_ A ⊠ λ'_ B. - -Definition condition_XX := forall (A B : DD), - δ#_ c0,B,A ∘ (λ*_ A ⊞ id_ (B×A)) ∘ λ'_ (B×A) - ≃ λ'_ B ⊠ id_ A. - -Definition condition_XXI := forall (A B : DD), - δ_ A,B,c0 ∘ (id_ (A×B) ⊞ ρ*_ A) ∘ (ρ'_ (A×B)) - ≃ id_ A ⊠ ρ'_ B. - -Definition condition_XXII := forall (A B : DD), - δ#_ A,c0,B ∘ (id_ (A×B) ⊞ λ*_ B) ∘ (ρ'_ (A×B)) - ≃ ρ'_ A ⊠ id_ B. - -Definition condition_XXIII := forall (A B : DD), - δ_ c1,A,B ∘ (λ_ A ⊞ λ_ B) - ≃ λ_ (A+B). - -Definition condition_XXIV := forall (A B : DD), - δ#_ A,B,c1 ∘ (ρ_ A ⊞ ρ_ B) - ≃ ρ_ (A+B). - -(* Conditions specific to braided rigs: *) -Definition condition_II `{@BraidedMonoidalCategory DD cC MulC} := - forall (A B C : DD), - (δ#_ A, B, C) ∘ (γ_ A,C ⊞ γ_ B,C) - ≃ γ_ A+B, C ∘ δ_ C,A,B. - -Definition condition_XV `{@BraidedMonoidalCategory DD cC MulC} := - forall (A : DD), - ρ*_ A ≃ γ_ A,c0 ∘ λ*_A. - - - - - - - -Lemma equivalence_1 `{BMul: @BraidedMonoidalCategory DD cC MulC} : - condition_II -> (condition_I <-> condition_III). -Proof. - unfold condition_II, condition_I, condition_III. - intros cond_2; split. - - intros cond_1 A B C. - symmetry. - rewrite <- right_unit; fold addC. - rewrite <- (addC.(tensor).(id_bimap)). - rewrite <- 2!(proj1 (BMul.(braiding) _ _).(isomorphism_inverse)); fold mulC. - rewrite compose_bimap. - rewrite <- assoc, (assoc (_ ⊠ _)). - rewrite (cond_2 B A C). - rewrite <- assoc. - rewrite braiding_natural. - symmetry. - rewrite <- (right_unit (δ#_ A,B, C)). - rewrite <- (addC.(tensor).(id_bimap)). - rewrite <- 2!braiding_inv_r. - rewrite compose_bimap. - rewrite <- assoc. - rewrite cond_2. - rewrite (assoc _ (_ ⊞ _)). - rewrite braiding_natural. - rewrite (assoc (γ_ (A+B), C)). - simpl. - rewrite <- (assoc (δ_ C,A,B)). - rewrite (cond_1 C A B). - rewrite <- 2!assoc. - easy. - - intros cond_3 A B C. - rewrite <- (left_unit (δ_ A,B,C)). - fold addC. - rewrite <- braiding_inv_l. - rewrite (assoc _ _ (δ_ _, _, _)). - rewrite <- cond_2. - rewrite 2!assoc, braiding_natural. - rewrite <- (assoc (δ#_ B, C, A)). - rewrite cond_3. - symmetry. - rewrite <- (left_unit (δ_ A,C,B)). - rewrite <- braiding_inv_l. - rewrite (assoc _ _ (δ_ _, _, _)). - rewrite <- cond_2. - rewrite <- (assoc (id_ A ⊠ _)). - rewrite inv_braiding_natural. - rewrite 2!assoc. - easy. -Qed. - - -Local Open Scope Mon_scope. - -Lemma stack_distr_pushout_r_bot : forall {C : Type} - {cC : Category C} {mC : MonoidalCategory cC} - {a b d m n} (f : a ~> b) (g : b ~> d) (h : m ~> n), - (f ∘ g) ⊗ h ≃ f ⊗ h ∘ (g ⊗ (id_ n)). -Proof. - intros. - rewrite <- compose_bimap, right_unit. - easy. -Qed. - -Lemma stack_distr_pushout_r_top : forall {C : Type} - {cC : Category C} {mC : MonoidalCategory cC} - {a b m n o} (f : a ~> b) (g : m ~> n) (h : n ~> o), - f ⊗ (g ∘ h) ≃ f ⊗ g ∘ (id_ b ⊗ h). -Proof. - intros. - rewrite <- compose_bimap, right_unit. - easy. -Qed. - -Lemma stack_distr_pushout_l_bot : forall {C : Type} - {cC : Category C} {mC : MonoidalCategory cC} - {a b d m n} (f : a ~> b) (g : b ~> d) (h : m ~> n), - (f ∘ g) ⊗ h ≃ f ⊗ (id_ m) ∘ (g ⊗ h). -Proof. - intros. - rewrite <- compose_bimap, left_unit. - easy. -Qed. - -Lemma stack_distr_pushout_l_top : forall {C : Type} - {cC : Category C} {mC : MonoidalCategory cC} - {a b m n o} (f : a ~> b) (g : m ~> n) (h : n ~> o), - f ⊗ (g ∘ h) ≃ id_ a ⊗ g ∘ (f ⊗ h). -Proof. - intros. - rewrite <- compose_bimap, left_unit. - easy. -Qed. - -Lemma stack_id_compose_split_top : forall {C : Type} - {cC : Category C} {mC : MonoidalCategory cC} - {topIn topMid topOut bot : C} - (f0 : topIn ~> topMid) (f1 : topMid ~> topOut), - (f0 ∘ f1) ⊗ (id_ bot) ≃ f0 ⊗ id_ bot ∘ (f1 ⊗ id_ bot). -Proof. - intros. - rewrite <- compose_bimap, left_unit. - easy. -Qed. - -Lemma stack_id_compose_split_bot : forall {C : Type} - {cC : Category C} {mC : MonoidalCategory cC} - {top botIn botMid botOut : C} - (f0 : botIn ~> botMid) (f1 : botMid ~> botOut), - (id_ top) ⊗ (f0 ∘ f1) ≃ id_ top ⊗ f0 ∘ (id_ top ⊗ f1). -Proof. - intros. - rewrite <- compose_bimap, left_unit. - easy. -Qed. - -Lemma id_stack_compose_topleft_general : forall {C : Type} - {cC : Category C} {mC : MonoidalCategory cC} - {topIn botIn topOut botOut : C} - (f0 : botIn ~> botOut) (f1 : topIn ~> topOut), - ((c_identity topIn) ⊗ f0) ∘ (f1 ⊗ (c_identity botOut)) ≃ (f1 ⊗ f0). -Proof. - intros. - rewrite <- compose_bimap. - rewrite left_unit; rewrite right_unit. - easy. -Qed. - -Lemma id_stackcompose_topright_general : forall {C : Type} - {cC : Category C} {mC : MonoidalCategory cC} - {topIn botIn topOut botOut : C} - (f0 : topIn ~> topOut) (f1 : botIn ~> botOut), - (f0 ⊗ (c_identity botIn)) ∘ ((c_identity topOut) ⊗ f1) ≃ (f0 ⊗ f1). -Proof. - intros. - rewrite <- compose_bimap. - rewrite right_unit, left_unit. - easy. -Qed. - -Lemma associator_cohere_inv : forall {C } - {cC : Category C} {mC : MonoidalCategory cC} - {A B M N P Q : C} - (f : A ~> B) (g : M ~> N) (h : P ~> Q), - (associator A M P)^-1 ∘ (f ⊗ g ⊗ h) ≃ (f ⊗ (g ⊗ h)) ∘ (associator B N Q)^-1. -Proof. - intros. - symmetry. - rewrite <- compose_iso_l, <- assoc, <- compose_iso_r'. - apply associator_cohere. -Qed. - -Local Close Scope Mon_scope. - - -Open Scope Cat_scope. -Open Scope Rig_scope. - - -Ltac rassoc t := - let H := fresh in - let rt := right_associate t in - assert (H: (t ≃ rt)%Cat) by (show_equiv_right_associate t); - rewrite H; clear H. - -Ltac rassoc_LHS := - match goal with |- (?LHS ≃ ?RHS) => rassoc LHS end. - -Ltac rassoc_RHS := - match goal with |- (?LHS ≃ ?RHS) => rassoc RHS end. - -Ltac right_associate f := - match f with - | ((?g ∘ ?h) ∘ ?i)%Cat => right_associate (g ∘ (h ∘ i))%Cat - | (?g ∘ ?h)%Cat => (* g shouldn't be a composition *) - let RAh := right_associate h in - constr:((g ∘ RAh)%Cat) - | _ => constr:(f) - end. - -(* TODO: Test this! *) -Ltac left_associate f := - match f with - | (?g ∘ (?h ∘ ?i))%Cat => left_associate ((g ∘ h) ∘ i)%Cat - | (?g ∘ ?h)%Cat => (* h shouldn't be a composition *) - let LAg := left_associate g in - constr:((LAg ∘ h)%Cat) - | _ => constr:(f) - end. - -Lemma assoc_compat_helper' : forall {C} - {cC : Category C} {A B M N : C} (f : A ~> B) - (g : B ~> M) (h : M ~> N) (fgh : A ~> N), - (f ∘ g) ∘ h ≃ fgh -> f ∘ (g ∘ h) ≃ fgh. -Proof. intros; rewrite <- assoc; easy. Qed. - -Ltac show_equiv_left_associate f := - let rec show_left f := - match f with - | (?g ∘ (?h ∘ ?i))%Cat => - (* RHS is `left_associate ((g ∘ h) ∘ i)%Cat` *) - apply assoc_compat_helper'; - show_left (((g ∘ h) ∘ i)%Cat) - | (?g ∘ ?h)%Cat => (* h shouldn't be a composition *) - (* RHS is `(left_associate g ∘ h)` *) - apply compose_cancel_r; - show_left g - | _ => - (* RHS is `constr:(f)` *) - reflexivity - end - in show_left f. - -Ltac lassoc t := - let H := fresh in - let rt := left_associate t in - assert (H: (t ≃ rt)%Cat) by (show_equiv_left_associate t); - rewrite H; clear H. - -Ltac lassoc_LHS := - match goal with |- (?LHS ≃ ?RHS) => lassoc LHS end. - -Ltac lassoc_RHS := - match goal with |- (?LHS ≃ ?RHS) => lassoc RHS end. - - - -Lemma equivalence_2 `{BMul: @BraidedMonoidalCategory DD cC MulC} : - condition_II -> (condition_IV <-> condition_V). -Proof. - unfold condition_II, condition_IV, condition_V. - intros cond_2. - split. - - intros cond_4 A B C D. - symmetry in cond_2. - setoid_rewrite compose_iso_l in cond_2. - symmetry. - rewrite assoc. - rewrite <- (compose_iso_l' - (BifunctorIsomorphism cmul (IdentityIsomorphism _) (associator _ _ _))). - simpl. - rewrite 2!cond_2. - rewrite 2!assoc. - rewrite <- compose_iso_l'. - rewrite 2!stack_id_compose_split_top. - rewrite assoc, <- 3!(assoc (_ ⊞ _)), <- (assoc (_ ⊠ _)). - rewrite <- compose_bimap, iso_inv_r, right_unit. - rewrite <- compose_bimap, right_unit, left_unit. - rewrite <- assoc. - rewrite <- id_stackcompose_topright_general, <- assoc. - symmetry in cond_4. - setoid_rewrite assoc in cond_4. - setoid_rewrite <- (compose_iso_l' - (BifunctorIsomorphism cmul (associator _ _ _) (IdentityIsomorphism _))) - in cond_4. - rewrite cond_4. - simpl. - rewrite 2!(assoc (_ ⊠ _)). - (* rewrite ?assoc. *) - rewrite (compose_iso_l - (BifunctorIsomorphism cmul (associator _ _ _) (IdentityIsomorphism A))). - simpl. - symmetry. - rewrite <-(assoc (associator _ _ _ ^-1 ⊠ id_ A)). - rewrite braiding_natural. - simpl in cond_4. - rewrite assoc. - simpl. - rewrite <- 2!(assoc (id_ A ⊠ α'_ _, _, _)). - rewrite <- compose_bimap, iso_inv_l, left_unit, id_bimap, left_unit. - rewrite 2!cond_2. - rewrite <- !assoc. - rewrite iso_inv_r, left_unit. - (* rewrite stack_id_compose_split_bot. *) - rewrite 2!assoc. - rewrite <- 1!(assoc (γ_ _, _ ⊞ γ_ _, _)). - rewrite <- compose_bimap, right_unit. - rewrite <- 2!(assoc ((γ_ _, A))). - rewrite iso_inv_r, left_unit. - rewrite ?assoc. - apply compose_cancel_l. - rewrite <- compose_bimap, left_unit, right_unit. - rewrite associator_cohere_inv. - rewrite <- assoc. - apply compose_cancel_r. - rewrite stack_distr_pushout_l_top. - easy. - - intros cond_5 A B C D. - setoid_rewrite (compose_iso_r _ (BifunctorIsomorphism cplus _ _)) in cond_2. - symmetry. - rewrite assoc. - simpl in cond_2. - rewrite <- (compose_iso_l' - (BifunctorIsomorphism cmul (associator _ _ _) (IdentityIsomorphism D))). - rewrite cond_2. - rewrite 2!(assoc (γ_ _, _)). - rewrite compose_iso_l. - rewrite assoc. - simpl. - rewrite <- compose_bimap. - rewrite right_unit. - setoid_rewrite assoc in cond_2. - setoid_rewrite compose_iso_l' in cond_2. - rewrite (cond_2 A B D). - rewrite stack_distr_pushout_l_bot, <- assoc. - setoid_rewrite assoc in cond_5. - setoid_rewrite <- (compose_iso_l - (BifunctorIsomorphism cmul - (IdentityIsomorphism _) (associator _ _ _))) in cond_5. - simpl in cond_5. - rewrite <- cond_5. - rewrite <- 2!(assoc _ _ (associator _ _ _ ^-1)). - rewrite (assoc _ (associator _ _ _^-1)). - rewrite associator_cohere_inv. - - (* rewrite <- 3!assoc. *) - rewrite <- compose_iso_l. - rewrite <- 2!(assoc (γ_ _, _)). - rewrite <- braiding_natural. - rewrite 2!(assoc (associator _ _ _ ⊠ id_ D)). - apply compose_cancel_l. - rewrite <- assoc. - apply compose_cancel_r. - rewrite (assoc (γ_ _, _)). - rewrite compose_iso_l. - rewrite <- (assoc ((γ_ _, _) ^-1)). - simpl. - rewrite (cond_2 A (B + C) D). - rassoc_LHS; rassoc_RHS. - apply compose_cancel_l. - symmetry. - simpl. - rewrite <- compose_bimap, right_unit. - rewrite (cond_2 B C D). - apply stack_distr_pushout_l_top. -Qed. - -Lemma distributive_hexagon_1 {BMulC : BraidedMonoidalCategory MulC} : - forall (A B C D : DD), - id_ A ⊠ γ_ B, C ⊞ id_ A ⊠ γ_ B, D ∘ γ_ A,(C×B) ⊞ γ_ A, (D×B) - ∘ α^-1_ C,B,A ⊞ α^-1_ D,B,A ∘ id_ C ⊠ (γ_ A, B)^-1 ⊞ id_ D ⊠ (γ_ A, B)^-1 - ≃ α_ A,B,C ⊞ α_ A,B,D ∘ γ_ (A×B),C ⊞ γ_ (A×B), D. -Proof. - intros A B C D. - rewrite <- 4!(compose_bimap cplus). - apply morphism_bicompat; apply hexagon_resultant_1. -Qed. - -Lemma distributive_hexagon_2 {BMulC : BraidedMonoidalCategory MulC} : - forall (A B C D : DD), - id_ A ⊠ γ_ B,(C+D) ∘ γ_ A,((C+D)×B) - ∘ α^-1_ (C+D), B, A ∘ id_(C+D) ⊠ (γ_ A,B)^-1 - ≃ α_ A, B, (C+D) ∘ γ_ (A×B),(C+D). -Proof. - intros. - apply hexagon_resultant_1. -Qed. - -Lemma equivalence_3_helper_1 {BMul : BraidedMonoidalCategory MulC} : - forall (A B C D : DD), - α^-1_ (C+D), B, A - ≃ (γ_ A,((C+D)×B)) ^-1 ∘ id_ A ⊠ (γ_ B, (C+D)) ^-1 - ∘ α_ A, B, (C+D) ∘ γ_ (A×B),(C+D) ∘ id_(C+D) ⊠ γ_ A,B. -Proof. - intros A B C D. - simpl. - epose proof (compose_tensor_iso_r' (mC:=mulC) _ - (IdentityIsomorphism (C+D)) (γ_ A,B)) - as e; simpl in e; rewrite e; clear e. - rassoc_RHS. - rewrite <- compose_iso_l. - rewrite <- (compose_tensor_iso_l (IdentityIsomorphism _)). - simpl. - lassoc_LHS. - apply distributive_hexagon_2. -Qed. - -(* We don't need laplaza's coherence requirement, as that's baked-in for us. *) -Lemma equivalence_3 `{BMul: @BraidedMonoidalCategory DD cC MulC} : - condition_II -> (condition_VI <-> condition_VII). -Proof. Abort. - (* unfold condition_II, condition_VI, condition_VII. - intros cond_2; simpl in *; split. - - intros cond_6 A B C D. - symmetry. - setoid_rewrite (compose_iso_r _ - (BifunctorIsomorphism tensor - (BMul.(braiding) _ _) (BMul.(braiding) _ _))) in cond_2. - simpl in cond_2. - epose proof (braiding.(component_biiso_natural)) as Hbr. - setoid_rewrite compose_iso_r in Hbr. - simpl in Hbr. - rewrite (Hbr _ _ _ _ (δ#_ A, B, C) (id_ D)). - rassoc_LHS. - rewrite <- compose_iso_l'. - rewrite 2!cond_2. - rassoc_LHS. - - rewrite <- (assoc (g:= (B_ _, _))). - rewrite iso_inv_l, left_unit. - rewrite stack_id_compose_split_bot. - rewrite (assoc (g:= id_ _ ⊗ _)). - rewrite <- (assoc (f:= id_ _ ⊗ (_ ⊗ _))). - rewrite <- left_distr_natural. - lassoc_LHS. - rewrite (assoc (g:= id_ _ ⊗ (_ ∘ _))). - rewrite stack_id_compose_split_bot. - rewrite (assoc (f:= id_ _ ⊗ B_ _, _)). - setoid_rewrite <- (compose_iso_r' _ - (BifunctorIsomorphism tensor associator associator)) in cond_6. - Local Open Scope Cat_scope. - simpl in cond_6. - rewrite (cond_6 D C A B). - rewrite cond_2. - - rassoc_LHS. - rewrite <- 2!compose_bimap. - - pose proof BMul.(@hexagon_2 _ _ _) as hex2. - setoid_rewrite assoc in hex2. - - pose proof BMul.(@hexagon_1 _ _ _) as hex1. - setoid_rewrite (compose_iso_r _ - (BifunctorIsomorphism tensor (IdentityIsomorphism _) (BMul.(braiding) _ _))) - in hex1. - setoid_rewrite <- (compose_iso_r _ - (BifunctorIsomorphism tensor (IdentityIsomorphism _) - (BMul.(braiding) _ _))) in hex1. - setoid_rewrite assoc in hex1. - setoid_rewrite (compose_iso_l - (BifunctorIsomorphism tensor - (BMul.(braiding) _ _) (IdentityIsomorphism _))) - in hex1. - simpl in hex1. - setoid_rewrite compose_iso_l in hex1. - rewrite compose_iso_l. - rewrite hex1. - rassoc_LHS. - rewrite <- compose_iso_l'. - lassoc_RHS. - setoid_rewrite <- assoc in hex2. - rewrite <- hex2. - rassoc_RHS. - do 2 apply compose_cancel_l. - rewrite compose_iso_l. - lassoc_RHS. - epose proof (component_biiso_natural_reverse - (N:=BMul.(braiding))) as e; simpl in e; - rewrite e; clear e. - rewrite (assoc (g:= (B_ _, _)^-1)). - rewrite iso_inv_l, right_unit. - lassoc_LHS. - rewrite iso_inv_r, left_unit. - rewrite <- id_bimap. - rewrite <- left_distr_natural. - rassoc_RHS. - apply compose_cancel_l. - rewrite <- 2!compose_bimap. - - specialize hex2 as hex2'. - setoid_rewrite <- (compose_iso_r' _ - (BifunctorIsomorphism tensor - (IdentityIsomorphism _) - (BMul.(braiding) _ _))) in hex2'. - setoid_rewrite compose_iso_r in hex2'. - unfold CommuteBifunctor in hex2'. - simpl in hex2'. - - pose proof BMul.(@hexagon_1 _ _ _) as hex1'. - setoid_rewrite (compose_iso_r _ - (BifunctorIsomorphism tensor - (IdentityIsomorphism _) (BMul.(braiding) _ _))) - in hex1'. - setoid_rewrite (compose_iso_r _ associator) in hex1'. - do 3 setoid_rewrite assoc in hex1'. - do 3 setoid_rewrite compose_iso_l' in hex1'. - do 2 setoid_rewrite <- assoc in hex1'. - setoid_rewrite (compose_iso_r _ - (BifunctorIsomorphism tensor - (BMul.(braiding) _ _) (IdentityIsomorphism _))) - in hex1'. - simpl in hex1'. - - apply morphism_bicompat; - - rewrite (hex2' D C _); - rassoc_RHS; - apply compose_cancel_l; - - rewrite <- (assoc (f:=associator ^-1)); - rewrite (hex1' _ C D); - rewrite <- 2!(assoc (f:= id_ C ⊗ B_ _, _)), - <- compose_bimap, iso_inv_r, left_unit, id_bimap, left_unit; - rewrite <- (assoc (f:=associator)), iso_inv_r, left_unit; - - epose proof (component_biiso_natural_reverse - (N:=BMul.(braiding)) _ _ _ _ _) as e; - simpl in e; rewrite <- e; - easy. - - intros cond_7 A B C D. - symmetry in cond_2. - setoid_rewrite (compose_iso_l) in cond_2. - simpl in cond_2. - epose proof (braiding.(component_biiso_natural)) as Hbr. - epose proof (braiding.(component_biiso_natural)) as Hbr'. - epose proof (braiding.(component_biiso_natural)) as Hbr''. - setoid_rewrite compose_iso_r in Hbr'. - setoid_rewrite compose_iso_l' in Hbr''. - simpl in Hbr, Hbr', Hbr''. - rewrite <- (Hbr'' _ _ _ _ (δ_ B, C, D) (id_ A)). - rassoc_LHS; rewrite <- compose_iso_l'. - - - rewrite 2!cond_2. - rewrite <- 2!(assoc (f:=B_ _, A)). - rewrite iso_inv_r, left_unit. - symmetry. - rewrite cond_2. - - - (* symmetry. *) - rewrite 2!stack_id_compose_split_top. - rassoc_RHS. - epose proof (compose_iso_l (BifunctorIsomorphism tensor - (BMul.(braiding) _ _) (IdentityIsomorphism _))) as e; - simpl in e; rewrite <- e; clear e. - lassoc_LHS. - rewrite Hbr. - - - pose proof BMul.(@hexagon_2 _ _ _) as hex2. - setoid_rewrite assoc in hex2. - - pose proof BMul.(@hexagon_1 _ _ _) as hex1. - setoid_rewrite (compose_iso_r _ - (BifunctorIsomorphism tensor (IdentityIsomorphism _) (BMul.(braiding) _ _))) - in hex1. - setoid_rewrite <- (compose_iso_r _ - (BifunctorIsomorphism tensor (IdentityIsomorphism _) - (BMul.(braiding) _ _))) in hex1. - setoid_rewrite assoc in hex1. - setoid_rewrite (compose_iso_l - (BifunctorIsomorphism tensor - (BMul.(braiding) _ _) (IdentityIsomorphism _))) - in hex1. - simpl in hex1. - lassoc_RHS. - rewrite <- (compose_iso_r _ - (BifunctorIsomorphism tensor associator associator)). - unfold CommuteBifunctor. - simpl. - symmetry. - - epose proof (compose_iso_r _ - (BifunctorIsomorphism AddC.(tensor) - (BMul.(braiding) _ _) (BMul.(braiding) _ _))) as e; - simpl in e. - rewrite e; clear e. - Admitted. *) - - -End CoherenceConditions. - -Class SemiCoherent_DistributiveBimonoidalCategory {DD : Type} {cC : Category DD} - {mAddC : MonoidalCategory cC} {bAddC : BraidedMonoidalCategory mAddC} - {AddC : SymmetricMonoidalCategory bAddC} - {MulC : MonoidalCategory cC} - {pdC : PreDistributiveBimonoidalCategory AddC MulC} - (rC : DistributiveBimonoidalCategory pdC) := { - cond_I : condition_I _; - cond_III : condition_III _; - cond_IV : condition_IV _; - cond_V : condition_V _; - cond_VI : condition_VI _; - cond_VII : condition_VII _; - cond_VIII : condition_VIII _; - cond_IX : condition_IX _; - cond_X : condition_X _ _; - cond_XI : condition_XI _ _; - cond_XII : condition_XII _ _; - cond_XIII : condition_XIII _ _; - cond_XIV : condition_XIV _ _; - cond_XVI : condition_XVI _ _; - cond_XVII : condition_XVII _ _; - cond_XVIII : condition_XVIII _ _; - cond_XIX : condition_XIX _ _; - cond_XX : condition_XX _ _; - cond_XXI : condition_XXI _ _; - cond_XXII : condition_XXII _ _; - cond_XXIII : condition_XXIII _; - cond_XXIV : condition_XXIV _; - -(* condition_I (A B C : DD) : - δ_ A,B,C ∘ γ'_ (A×B), (A×C) ≃ - id_ A ⊠ γ'_ B, C ∘ δ_ A, C, B; - condition_III (A B C : DD) : - δ#_ A,B,C ∘ γ'_ (A×C),(B×C) - ≃ (γ'_ A,B ⊠ id_ C) ∘ δ#_ B,A,C; - condition_IV (A B C D : DD) : - (* c_equiv (A:=(A+(B+C))×D) (B:=(A×D + B×D + C×D)) *) - δ#_ A,(B+C),D ∘ (id_ (A×D) ⊞ δ#_ B,C,D) ∘ α'_ A×D, B×D, (C×D) - ≃ (α'_ A,B,C ⊠ (id_ D)) ∘ δ#_ A+B,C,D ∘ (δ#_ A,B,D ⊞ id_ (C×D)); - condition_V (A B C D : DD) : - δ_ A,B,(C+D) ∘ (id_(A×B) ⊞ δ_ A,C,D) ∘ α'_ A×B,A×C,(A×D) - ≃ id_ A ⊠ α'_ B,C,D ∘ δ_ A,B+C,D ∘ (δ_ A,B,C ⊞ id_(A×D)); - condition_VI (A B C D : DD) : - id_ A ⊠ δ_ B,C,D ∘ δ_ A,B×C,(B×D) ∘ (α_ A,B,C ⊞ α_ A,B,D) - ≃ α_ A,B,(C+D) ∘ δ_ (A×B),C,D; - condition_VII (A B C D : DD) : - δ#_ A,B,(C×D) ∘ (α_ A,C,D ⊞ α_ B,C,D) - ≃ α_ A+B,C,D ∘ (δ#_A,B,C ⊠ id_ D) ∘ (δ#_ A×C,B×C,D); - condition_VIII (A B C D : DD) : - (id_ A ⊠ δ#_ B,C,D) ∘ δ_ A,B×D,(C×D) ∘ (α_ A,B,D ⊞ α_ A,C,D) - ≃ α_ A,B+C,D ∘ (δ_ A,B,C ⊠ id_ D) ∘ δ#_ A×B,A×C,D; - condition_IX (A B C D : DD) : - δ#_ A,B,(C+D) ∘ (δ_ A,C,D ⊞ δ_ B,C,D) ∘ (α'_ (A×C+A×D),B×C,(B×D)) - ∘ ((α'^-1_ A×C,A×D,(B×C)) ⊞ id_ (B×D)) - ∘ ((id_ (A×C) ⊞ γ'_ A×D, (B×C)) ⊞ id_ (B×D)) - ∘ (α'_ A×C,B×C,(A×D) ⊞ id_ (B×D)) - ≃ δ_ A+B,C,D ∘ (δ#_ A,B,C ⊞ δ#_ A,B,D) ∘ α'_ (A×C+B×C),A×D,(B×D); - condition_X : - λ*_ c0 ≃ ρ*_ c0; - condition_XI (A B : DD) : - δ_ c0,A,B ∘ (λ*_ A ⊞ λ*_ B) ∘ λ'_ c0 - ≃ λ*_ (A + B); - condition_XII (A B : DD) : - δ#_ A,B,c0 ∘ (ρ*_ A ⊞ ρ*_ B) ∘ (λ'_ c0) - ≃ ρ*_ (A+B); - condition_XIII : - ρ_ c0 ≃ λ*_ c1; - condition_XIV : - λ_ c0 ≃ ρ*_ c1; - condition_XVI (A B : DD) : - α_ c0,A,B ∘ (λ*_ A ⊠ id_ B) ∘ λ*_B - ≃ λ*_ (A×B); - condition_XVII (A B : DD) : - α_ A,c0,B ∘ (ρ*_ A ⊠ id_ B) ∘ λ*_ B - ≃ id_ A ⊠ λ*_ B ∘ ρ*_ A; - condition_XVIII (A B : DD) : - α_ A,B,c0 ∘ ρ*_ (A×B) - ≃ (id_ A ⊠ ρ*_ B) ∘ ρ*_ A; - condition_XIX (A B : DD) : - δ_ A,c0,B ∘ (ρ*_ A ⊞ id_ (A×B)) ∘ λ'_ (A×B) - ≃ id_ A ⊠ λ'_ B; - condition_XX (A B : DD) : - δ#_ c0,B,A ∘ (λ*_ A ⊞ id_ (B×A)) ∘ λ'_ (B×A) - ≃ λ'_ B ⊠ id_ A; - condition_XXI (A B : DD) : - δ_ A,B,c0 ∘ (id_ (A×B) ⊞ ρ*_ A) ∘ (ρ'_ (A×B)) - ≃ id_ A ⊠ ρ'_ B; - condition_XXII (A B : DD) : - δ#_ A,c0,B ∘ (id_ (A×B) ⊞ λ*_ B) ∘ (ρ'_ (A×B)) - ≃ ρ'_ A ⊠ id_ B; - condition_XXIII (A B : DD) : - δ_ c1,A,B ∘ (λ_ A ⊞ λ_ B) - ≃ λ_ (A+B); - condition_XXIV (A B : DD) : - δ#_ A,B,c1 ∘ (ρ_ A ⊞ ρ_ B) - ≃ ρ_ (A+B); -*) -}. - - - - -Class SemiCoherent_BraidedDistributiveBimonoidalCategory {DD : Type} {cC : Category DD} - {mAddC : MonoidalCategory cC} {bAddC : BraidedMonoidalCategory mAddC} - {AddC : SymmetricMonoidalCategory bAddC} - {MulC : MonoidalCategory cC} - {pdC : PreDistributiveBimonoidalCategory AddC MulC} - (rC : DistributiveBimonoidalCategory pdC) - (bMulC : BraidedMonoidalCategory MulC) := { - cond_II : condition_II _; - cond_XV : condition_XV _ _; -(* - condition_II (A B C : DD) : - (δ#_ A, B, C) ∘ (γ_ A,C ⊞ γ_ B,C) - ≃ γ_ A+B, C ∘ δ_ C,A,B; - condition_XV (A : DD) : - ρ*_ A ≃ γ_ A,c0 ∘ λ*_A; -*) -}. *) - -Close Scope Rig_scope. diff --git a/ViCaR/Classes/SymmetricMonoidal.v b/ViCaR/Classes/SymmetricMonoidal.v index 5a0ac86..104a9d6 100644 --- a/ViCaR/Classes/SymmetricMonoidal.v +++ b/ViCaR/Classes/SymmetricMonoidal.v @@ -5,19 +5,17 @@ Require Import BraidedMonoidal. #[local] Set Universe Polymorphism. Local Open Scope Cat. -Local Open Scope Brd. Class SymmetricMonoidalCategory {C : Type} {cC : Category C} {mC : MonoidalCategory cC} (bC : BraidedMonoidalCategory mC) : Type := { - symmetry (A B : C) : B_ A,B ≃ (B_ B,A)^-1; + symmetry (A B : C) : β_ A,B ≃ (β_ B,A)^-1; BraidedMonoidalCategory_of_SymmetricMonoidalCategory := bC; }. Arguments SymmetricMonoidalCategory {_} {_ _}%Cat (_)%Cat. -Arguments symmetry {_} {_ _ _ _}%Cat (_ _)%Obj. +Arguments symmetry {_} {_ _ _ _}%Cat (_ _)%Cat. Coercion BraidedMonoidalCategory_of_SymmetricMonoidalCategory : SymmetricMonoidalCategory >-> BraidedMonoidalCategory. -Local Close Scope Brd. Local Close Scope Cat. diff --git a/ViCaR/GeneralLemmas.v b/ViCaR/GeneralLemmas.v index 3bb68cd..0931429 100644 --- a/ViCaR/GeneralLemmas.v +++ b/ViCaR/GeneralLemmas.v @@ -2,9 +2,14 @@ Require Import Category. Require Import Monoidal. Require Import Setoid. +Section GeneralLemmas. + Local Open Scope Cat. -Lemma compose_simplify_general : forall {C : Type} {cC : Category C} +Context {C : Type} {cC : Category C} {cCh : CategoryCoherence cC} + {mC : MonoidalCategory cC} {mCh : MonoidalCategoryCoherence mC}. + +Lemma compose_simplify_general : forall {A B M : C} (f1 f3 : A ~> B) (f2 f4 : B ~> M), f1 ≃ f3 -> f2 ≃ f4 -> (f1 ∘ f2) ≃ (f3 ∘ f4). Proof. @@ -13,8 +18,7 @@ Proof. reflexivity. Qed. -Lemma stack_simplify_general : forall {C : Type} - {cC : Category C} {mC : MonoidalCategory cC} +Lemma stack_simplify_general : forall {A B M N : C} (f1 f3 : A ~> B) (f2 f4 : M ~> N), f1 ≃ f3 -> f2 ≃ f4 -> (f1 ⊗ f2) ≃ (f3 ⊗ f4). Proof. @@ -23,137 +27,89 @@ Proof. reflexivity. Qed. -Lemma stack_compose_distr_test : forall {C : Type} - {cC : Category C} {mC : MonoidalCategory cC} - {A B D M N P : C} (f : A ~> B) (g : B ~> D) - (h : M ~> N) (i : N ~> P), - (f ∘ g) ⊗ (h ∘ i) ≃ (f ⊗ h) ∘ (g ⊗ i). +Lemma stack_distr_pushout_r_bot : forall + {a b d m n} (f : a ~> b) (g : b ~> d) (h : m ~> n), + (f ∘ g) ⊗ h ≃ f ⊗ h ∘ (g ⊗ (id_ n)). Proof. intros. - rewrite compose_bimap. + rewrite <- tensor_compose, right_unit. easy. Qed. -(* Local Notation "A ⨂ B" := (morphism_bimap (Bifunctor:=tensor) A B) (only printing). *) - +Lemma stack_distr_pushout_r_top : forall + {a b m n o} (f : a ~> b) (g : m ~> n) (h : n ~> o), + f ⊗ (g ∘ h) ≃ f ⊗ g ∘ (id_ b ⊗ h). +Proof. + intros. + rewrite <- tensor_compose, right_unit. + easy. +Qed. -Lemma stack_distr_pushout_r_bot : forall {C : Type} - `{cC : Category C} `{mC : MonoidalCategory C} +Lemma stack_distr_pushout_l_bot : forall {a b d m n} (f : a ~> b) (g : b ~> d) (h : m ~> n), - (f ∘ g) ⊗ h ≃ f ⊗ h ∘ (g ⊗ (id_ n)). + (f ∘ g) ⊗ h ≃ f ⊗ (id_ m) ∘ (g ⊗ h). Proof. intros. - rewrite <- compose_bimap, right_unit. + rewrite <- tensor_compose, left_unit. easy. Qed. -(* TODO: the other two; _l_bot and _l_top *) - -Lemma stack_distr_pushout_r_top : forall {C : Type} - {cC : Category C} {mC : MonoidalCategory cC} +Lemma stack_distr_pushout_l_top : forall {a b m n o} (f : a ~> b) (g : m ~> n) (h : n ~> o), - f ⊗ (g ∘ h) ≃ f ⊗ g ∘ (id_ b ⊗ h). + f ⊗ (g ∘ h) ≃ id_ a ⊗ g ∘ (f ⊗ h). Proof. intros. - rewrite <- compose_bimap, right_unit. + rewrite <- tensor_compose, left_unit. easy. Qed. -(* Ltac fencestep := - let test_simple t := match t with - | context[(_ ⊗ _) ∘ _] => fail 2 - | context[_ ∘ (_ ⊗ _)] => fail 2 - | context[(_ ∘ _) ⊗ _] => fail 2 - | context[_ ⊗ (_ ∘ _)] => fail 2 - | _ => idtac - end - in - first [ match goal with - |- context[(?f ∘ ?g) ⊗ (?h ∘ ?i)] => - test_simple f; test_simple g; test_simple h; test_simple i; - rewrite (compose_bimap f g h i) - end | match goal with - |- context[(?f) ⊗ (?g ∘ ?h)] => - test_simple f; test_simple g; test_simple h; - rewrite (stack_distr_pushout_r_top f g h) - end | match goal with - |- context[(?f ∘ ?g) ⊗ (?h)] => - test_simple f; test_simple g; test_simple h; - rewrite (stack_distr_pushout_r_bot f g h) - end]. *) - - - -(* Ltac fencepost := - - let rec process_term term := - match term with - | ?A ⊗ ?B => match A, B with - | ?A' ⊗ B', _ => process_term A - | ?f ∘ ?g (* TODO: test if these are "simple", in some sense. - Also see if I can come up with an (even informal) algorithm, or - even as much as an explicit spec (e.g., strict fenceposing may - be required to have functions go in descending order, so - f id id i id f - id g id id g id - id id h id , etc. is good, but not id id ... - id id id id id id - ). I suspect the strict spec will make reasoning - much easier, i.e. to process f⊗g, we *must* push it out - to f⊗id ∘ id⊗g. *) -*) - - - - - - - -Lemma nwire_stack_compose_topleft_general : forall {C : Type} - {cC : Category C} {mC : MonoidalCategory cC} + + + + + +Lemma nwire_stack_compose_topleft_general : forall {topIn botIn topOut botOut : C} (f0 : botIn ~> botOut) (f1 : topIn ~> topOut), ((c_identity topIn) ⊗ f0) ∘ (f1 ⊗ (c_identity botOut)) ≃ (f1 ⊗ f0). Proof. intros. - rewrite <- compose_bimap. + rewrite <- tensor_compose. rewrite left_unit; rewrite right_unit. easy. Qed. -Lemma nwire_stackcompose_topright_general : forall {C : Type} - {cC : Category C} {mC : MonoidalCategory cC} +Lemma nwire_stackcompose_topright_general : forall {topIn botIn topOut botOut : C} (f0 : topIn ~> topOut) (f1 : botIn ~> botOut), (f0 ⊗ (c_identity botIn)) ∘ ((c_identity topOut) ⊗ f1) ≃ (f0 ⊗ f1). Proof. intros. - rewrite <- compose_bimap. + rewrite <- tensor_compose. rewrite right_unit, left_unit. easy. Qed. -Lemma stack_id_compose_split_top : forall {C : Type} - {cC : Category C} {mC : MonoidalCategory cC} +Lemma stack_id_compose_split_top : forall {topIn topMid topOut bot : C} (f0 : topIn ~> topMid) (f1 : topMid ~> topOut), (f0 ∘ f1) ⊗ (id_ bot) ≃ f0 ⊗ id_ bot ∘ (f1 ⊗ id_ bot). Proof. intros. - rewrite <- compose_bimap, left_unit. + rewrite <- tensor_compose, left_unit. easy. Qed. -Lemma stack_id_compose_split_bot : forall {C : Type} - {cC : Category C} {mC : MonoidalCategory cC} +Lemma stack_id_compose_split_bot : forall {top botIn botMid botOut : C} (f0 : botIn ~> botMid) (f1 : botMid ~> botOut), (id_ top) ⊗ (f0 ∘ f1) ≃ id_ top ⊗ f0 ∘ (id_ top ⊗ f1). Proof. intros. - rewrite <- compose_bimap, left_unit. + rewrite <- tensor_compose, left_unit. easy. Qed. - Local Close Scope Cat. + +End GeneralLemmas. diff --git a/ViCaR/MonoidalCoherence.v b/ViCaR/MonoidalCoherence.v deleted file mode 100644 index 638cdfa..0000000 --- a/ViCaR/MonoidalCoherence.v +++ /dev/null @@ -1,1015 +0,0 @@ -Require Import String. -Require Import Psatz. -Require Import CategoryTypeclass. -Require Import Arith. - -Section BWord. - -Inductive W := - | e0 : W - | var : W - | tens : W -> W -> W. - -Inductive W0 := - | var0 : W0 - | tens0 : W0 -> W0 -> W0. - -Fixpoint W_eqb (b b' : W) : bool := - match b, b' with - | e0, e0 => true - | var, var => true - | tens b0 b1, tens b0' b1' => W_eqb b0 b0' && W_eqb b1 b1' - | _, _ => false - end. - -Fixpoint W0_eqb (b b' : W0) : bool := - match b, b' with - | var0, var0 => true - | tens0 b0 b1, tens0 b0' b1' => W0_eqb b0 b0' && W0_eqb b1 b1' - | _, _ => false - end. - - - -Require Import Bool. - -Lemma W_eqbP : forall b b', b = b' <-> W_eqb b b' = true. -Proof. - intros b b'. - split. - - intros H; subst b'; induction b; - simpl; try apply andb_true_intro; easy. - - revert b'. - induction b; - intros b' H; destruct b'; try easy. - simpl in H. - rewrite andb_true_iff in H. - erewrite IHb1, IHb2 by apply H. - easy. -Qed. - -Lemma W0_eqbP : forall b b', b = b' <-> W0_eqb b b' = true. -Proof. - intros b b'. - split. - - intros H; subst b'; induction b; - simpl; try apply andb_true_intro; easy. - - revert b'. - induction b; - intros b' H; destruct b'; try easy. - simpl in H. - rewrite andb_true_iff in H. - erewrite IHb1, IHb2 by apply H. - easy. -Qed. - -Lemma bw_eq_dec : forall b b' : W, {b = b'} + {b <> b'}. -Proof. - intros b b'. - destruct (W_eqb b b') eqn:e. - - rewrite <- W_eqbP in e. - left; exact e. - - right. - intros Heq. - rewrite (proj1 (W_eqbP b b') Heq) in e. - discriminate e. -Qed. - - -Fixpoint eval_W {A} (a : A) (zero : A) - (bind : A -> A -> A) (b : W) : A := - match b with - | e0 => zero - | var => a - | tens b0 b1 => - bind (eval_W a zero bind b0) (eval_W a zero bind b0) - end. - -Fixpoint print_word (b : W) : string := - match b with - | e0 => "e"%string - | var => "—"%string - | tens b0 b1 => - ("(" ++ print_word b0 ++ ")⊗(" ++ print_word b1 ++ ")")%string - end. - - - -Fixpoint interp_W0 {C} {cC : Category C} {mC : MonoidalCategory cC} - (c : C) (b : W0) := - match b with - | var0 => c - | tens0 b0 b1 => interp_W0 c b0 × interp_W0 c b1 - end. - -End BWord. - -Section Associativity. - -Inductive aS := - | as_W0 : W0 -> aS - | as_assoc : W0 -> W0 -> W0 -> aS - | as_invassoc : W0 -> W0 -> W0 -> aS - | as_tensor_l : aS -> W0 -> aS - | as_tensor_r : aS -> W0 -> aS. - -Fixpoint a_dom (s : aS) : W0 := - match s with - | as_W0 b => b - | as_assoc u v w => tens0 u (tens0 v w) - | as_invassoc u v w => tens0 (tens0 u v) w - | as_tensor_l a b => - tens0 b (a_dom a) - | as_tensor_r a b => - tens0 (a_dom a) b - end. - -Fixpoint a_codom (s : aS) : W0 := - match s with - | as_W0 b => b - | as_assoc u v w => tens0 (tens0 u v) w - | as_invassoc u v w => tens0 u (tens0 v w) - | as_tensor_l a b => - tens0 b (a_codom a) - | as_tensor_r a b => - tens0 (a_codom a) b - end. - -Fixpoint w_Sn (n : nat) : W0 := - match n with - | 0 => var0 - | S n' => tens0 (w_Sn n') var0 - end. - -Fixpoint len0 (b : W0) : nat := - match b with - | var0 => 1 - | tens0 b0 b1 => len0 b0 + len0 b1 - end. - -Lemma len0_pos : forall b, len0 b > 0. -Proof. induction b; simpl; lia. Qed. - -Lemma len_w_Sn : forall n, len0 (w_Sn n) = S n. -Proof. - induction n; simpl; lia. -Qed. - -Fixpoint rho0 (b : W0) : nat := - match b with - | var0 => 0 - | tens0 v w => rho0 v + rho0 w + len0 w - 1 - end. - -Lemma rho0_w_Sn : forall n, - rho0 (w_Sn n) = 0. -Proof. - induction n; simpl; lia. -Qed. - -Lemma rho_eq0_iff : forall b : W0, - rho0 b = 0 <-> b = w_Sn (len0 b - 1). -Proof. - intros b. - split. - - induction b; [easy|]. - destruct b2. - + simpl in *. - intros. - assert (rho0 b1 = 0) by lia. - rewrite IHb1 at 1 by lia. - assert (H':len0 b1 > 0) by (apply len0_pos); revert H'. - generalize (len0 b1) as n. - intros n. - induction n; [easy|]. - intros _. - simpl; rewrite 2!Nat.sub_0_r, Nat.add_1_r. - easy. - + intros H. - contradict H. - simpl. - pose (len0_pos b2_1); pose (len0_pos b2_2). - lia. - - intros H; rewrite H. - apply rho0_w_Sn. -Qed. - -Inductive a_dir : aS -> Prop := - | a_dir_assoc u v w : a_dir (as_assoc u v w) - | a_dir_tens_l s b : a_dir s -> a_dir (as_tensor_l s b) - | a_dir_tens_r s b : a_dir s -> a_dir (as_tensor_r s b). - -Inductive a_antidir : aS -> Prop := - | a_antidir_invassoc u v w : a_antidir (as_invassoc u v w) - | a_antidir_tens_l s b : a_antidir s -> a_antidir (as_tensor_l s b) - | a_antidir_tens_r s b : a_antidir s -> a_antidir (as_tensor_r s b). - -Lemma ne_tens0_l : forall b b' : W0, - b <> tens0 b b'. -Proof. - intros b. - induction b; [easy|]. - intros b' H. - inversion H. - eapply IHb1; eassumption. -Qed. - -Lemma ne_tens0_r : forall b b' : W0, - b <> tens0 b' b. -Proof. - intros b. - induction b; [easy|]. - intros b' H. - inversion H. - eapply IHb2; eassumption. -Qed. - -Lemma ne_tens0_assoc : forall b b' b'' : W0, - tens0 (tens0 b b') b'' <> tens0 b (tens0 b' b''). -Proof. - intros b. - induction b; [easy|]. - intros b' b'' H; inversion H. - symmetry in H1. - eapply ne_tens0_l, H1. -Qed. - -Ltac W0cong := - simpl in *; subst; intros; try easy; try lia; first [ - exfalso; eapply ne_tens0_r; solve [eauto] | - exfalso; eapply ne_tens0_l; solve [eauto] | - exfalso; eapply ne_tens0_assoc; solve [eauto] | - match goal with - | H : ?f _ _ = ?g _ _ |- _ => inversion H; clear H; W0cong - end | - match goal with - | H : forall _, _ |- _ => apply H; clear H; W0cong - end | - (progress f_equal); solve [eauto] | - (progress hnf); W0cong - ]. - -Ltac W0cong_debug := - simpl in *; subst; intros; try easy; try lia; first [ - exfalso; eapply ne_tens0_r; idtac "ner"; solve [eauto] | - exfalso; eapply ne_tens0_l; idtac "nel"; solve [eauto] | - exfalso; eapply ne_tens0_assoc; idtac "nea"; solve [eauto] | - match goal with - | H : ?f _ _ = ?g _ _ |- _ => inversion H; - idtac "invers" H; clear H; W0cong_debug - end | - match goal with - | H : forall _, _ |- _ => apply H; idtac "apply" H; clear H; W0cong_debug - end | - (progress f_equal); idtac "feq"; solve [eauto] | - (progress hnf); idtac "hnf"; W0cong_debug - ]. - -Hint Resolve ne_tens0_r ne_tens0_l ne_tens0_assoc : tens0_db. - -Lemma a_dir_not_id : forall s : aS, a_dir s -> a_dom s <> a_codom s. -Proof. - intros s Hs. - induction Hs; W0cong. -Qed. - -Lemma a_dir_len_dom_codom : forall s : aS, - a_dir s -> - len0 (a_dom s) = len0 (a_codom s). -Proof. - intros s Hs; induction Hs; simpl; auto with arith. -Qed. - -Lemma a_no_parallel_dir : forall s s' : aS, - a_dir s -> a_dir s' -> - a_dom s = a_dom s' -> a_codom s = a_codom s' -> s = s'. -Proof. - intros s s' Hs Hs'; revert s' Hs'. - induction Hs. - - simpl; intros s' Hs'; inversion Hs'; W0cong. - - simpl. intros s' Hs'. - inversion Hs'; try W0cong. - pose proof a_dir_not_id. - simpl. - intros. - subst. - inversion H2; inversion H3. - exfalso; eapply H1; eauto; subst; easy. - - simpl. intros s' Hs'. - inversion Hs'; try W0cong. - pose proof a_dir_not_id. - simpl. - intros. - subst. - inversion H2; inversion H3. - exfalso; eapply H1; eauto; subst; easy. -Qed. - -Lemma rho0_dir_decreasing : forall s, - a_dir s -> - rho0 (a_codom s) < rho0 (a_dom s). -Proof. - intros s Hs. - induction Hs. - - simpl. - pose (len0_pos u); - pose (len0_pos v); - pose (len0_pos w). - lia. - - simpl. - pose (len0_pos (a_dom s)). - rewrite <- a_dir_len_dom_codom by easy. - lia. - - simpl. - pose (len0_pos b). - lia. -Qed. - -Fixpoint a_inv (s : aS) : aS := - match s with - | as_W0 b => as_W0 b - | as_assoc u v w => as_invassoc u v w - | as_invassoc u v w => as_assoc u v w - | as_tensor_l a b => as_tensor_l (a_inv a) b - | as_tensor_r a b => as_tensor_r (a_inv a) b - end. - -Lemma a_inv_involutive : forall s, - a_inv (a_inv s) = s. -Proof. induction s; simpl; rewrite ?IHs; easy. Qed. - -Lemma a_inv_dom : forall s, - a_dom (a_inv s) = a_codom s. -Proof. induction s; simpl; rewrite ?IHs; easy. Qed. - -Lemma a_inv_codom : forall s, - a_codom (a_inv s) = a_dom s. -Proof. induction s; simpl; rewrite ?IHs; easy. Qed. - -Lemma a_inv_dir : forall s, - a_dir s -> a_antidir (a_inv s). -Proof. - intros s Hs. - induction Hs; constructor; assumption. -Qed. - -Lemma a_inv_antidir : forall s, - a_antidir s -> a_dir (a_inv s). -Proof. - intros s Hs. - induction Hs; constructor; assumption. -Qed. - - -Inductive a_lst := - | a_lst_init : aS -> a_lst - | a_lst_cons : aS -> a_lst -> a_lst. - -Definition a_lst_dom (ls : a_lst) : W0 := - match ls with - | a_lst_init s - | a_lst_cons s _ => a_dom s - end. - -Fixpoint a_lst_codom (ls : a_lst) : W0 := - match ls with - | a_lst_init s => a_codom s - | a_lst_cons _ l => a_lst_codom l - end. - -Fixpoint a_path (ls : a_lst) : bool := - match ls with - | a_lst_init s => true - | a_lst_cons s l => W0_eqb (a_codom s) (a_lst_dom l) && a_path l - end. - -Fixpoint a_path_prop (ls : a_lst) : Prop := - match ls with - | a_lst_init s => True - | a_lst_cons s l => (a_codom s) = (a_lst_dom l) /\ a_path_prop l - end. - -Fixpoint assoc_free (b : W0) : bool := - match b with - | var0 => true - | tens0 b' b'' => match b'' with - | var0 => assoc_free b' - | tens0 _ _ => false - end - end. - -Lemma assoc_free_iff_rho0_eq_0 : forall b, - assoc_free b = true <-> rho0 b = 0. -Proof. - intros b. - rewrite rho_eq0_iff. - split. - - intros H. - induction b; [easy|]. - simpl in H. - destruct b2; [|easy]. - simpl. - rewrite Nat.add_sub. - rewrite IHb1 at 1 by easy. - pose proof (len0_pos b1) as e ; revert e. - generalize (len0 b1) as n; induction n; [easy|]. - intros _. - simpl; rewrite Nat.sub_0_r; easy. - - generalize (len0 b). - intros n Hb; subst b. - destruct n; [easy|]. - simpl; rewrite Nat.sub_0_r. - induction n; easy. -Qed. - -Fixpoint a_can_next_step (b : W0) : aS := - match b with - | var0 => as_W0 var0 - | tens0 c d => match d with - | var0 => as_tensor_r (a_can_next_step c) var0 - | tens0 d' d'' => as_assoc c d' d'' - end - end. - -Lemma a_can_next_step_dom : forall b, - a_dom (a_can_next_step b) = b. -Proof. - intros b; induction b; [easy|]; simpl. - induction b2; [|easy]. - simpl; rewrite IHb1; easy. -Qed. - -Lemma a_can_next_step_dir : forall b, - rho0 b <> 0 -> - a_dir (a_can_next_step b). -Proof. - intros b H. - induction b as [|b0 b1]; [easy|]. - intros. simpl. - destruct b2; constructor. - apply b1. - simpl in H. - lia. -Qed. - -Lemma rho0_a_can_next_step_codom : forall b, - assoc_free b = false -> - rho0 (a_codom (a_can_next_step b)) < rho0 b. -Proof. - intros. - rewrite <- (a_can_next_step_dom b) at 2. - apply rho0_dir_decreasing, a_can_next_step_dir. - rewrite <- assoc_free_iff_rho0_eq_0. - rewrite H; easy. -Qed. - -Require Coq.Program.Wf. -Require Import FunInd Recdef. - -Function a_can_path (b : W0) {measure rho0 b} : a_lst := - match assoc_free b as H with - | true => a_lst_init (as_W0 b) - | false => - let cns := a_can_next_step b in - a_lst_cons (cns) (a_can_path (a_codom cns)) - end. -apply rho0_a_can_next_step_codom. -Qed. - - -(* Function a_can_path (b : W0) {measure (rho0 b)} : a_lst := - match assoc_free b as H with - | true => a_lst_init (as_W0 b) - | false => - let cns := a_can_next_step b in - a_lst_cons (cns) (a_can_path (a_codom cns)) - end. -Next Obligation. - apply rho0_a_can_next_step_codom; easy. -Qed. *) - -Lemma a_can_path_dom : forall b, a_lst_dom (a_can_path b) = b. -Proof. - intros b. - pattern (a_can_path b). - revert b. - apply a_can_path_ind; [easy|]. - intros b Hb. - simpl. - intros _. - apply a_can_next_step_dom. -Qed. - -Lemma a_can_path_codom_rho0 : forall b, - rho0 (a_lst_codom (a_can_path b)) = 0. -Proof. - intros b; - pattern (a_can_path b); - revert b. - apply a_can_path_ind; [|easy]. - intros b Hb. - rewrite assoc_free_iff_rho0_eq_0 in Hb. - easy. -Qed. - -Lemma a_can_next_step_codom_len : forall b, - len0 (a_codom (a_can_next_step b)) = len0 b. -Proof. - induction b; [easy|]. - simpl. - destruct b2; simpl; lia. -Qed. - -Lemma a_can_path_codom_len : forall b, - len0 (a_lst_codom (a_can_path b)) = len0 b. -Proof. - intros b; - pattern (a_can_path b); - revert b. - apply a_can_path_ind; [easy|]. - simpl; intros. - rewrite <- (a_can_next_step_codom_len b). - easy. -Qed. - -Lemma a_can_path_codom : forall b, - a_lst_codom (a_can_path b) = w_Sn (len0 b - 1). -Proof. - intros b. - pose proof (a_can_path_codom_rho0 b) as H. - rewrite rho_eq0_iff in H. - rewrite H, a_can_path_codom_len. - easy. -Qed. - -Lemma a_can_path_path : forall b, a_path_prop (a_can_path b). -Proof. - apply a_can_path_ind. - - easy. - - intros b Hb. - simpl. - intros Hpath; split; [|easy]. - rewrite a_can_path_dom. - easy. -Qed. - - - -Fixpoint interp_aS {C} {cC : Category C} {mC : MonoidalCategory cC} - (c : C) (s : aS) - : cC.(morphism) (interp_W0 c (a_dom s)) (interp_W0 c (a_codom s)) := - match s with - | as_W0 b => id_ _ - | as_assoc u v w => (associator _ _ _)^-1 - | as_invassoc u v w => (associator _ _ _) - | as_tensor_l a b => (id_ _) ⊗ (interp_aS c a) - | as_tensor_r a b => (interp_aS c a) ⊗ (id_ _) - end%Mor. - -Fixpoint a_dir_lst (l : a_lst) : Prop := - match l with - | a_lst_init s => a_dir s - | a_lst_cons s ls => a_dir s /\ a_dir_lst ls - end. - -Definition a_dir_path (l : a_lst) := a_path_prop l /\ a_dir_lst l. - -Fixpoint interp_a_path_prop {C} {cC : Category C} - {mC : MonoidalCategory cC} (c : C) (l : a_lst) - (H : a_path_prop l) {struct l} : - (interp_W0 c (a_lst_dom l) ~> interp_W0 c (a_lst_codom l))%Cat. -destruct l. -- apply interp_aS. -- simpl. - eapply compose. - + apply interp_aS. - + destruct H as [Ha Hl]. - rewrite Ha. - apply interp_a_path_prop, Hl. -Defined. - -Definition unique_directed_path_image_prop : Prop. - refine ( forall C (cC : Category C) - (mC : MonoidalCategory cC) (c : C) (v : W0) (l l' : a_lst) - (Hl : a_dir_path l) (Hl' : a_dir_path l') - (Hdoml : a_lst_dom l = v) (Hdoml' : a_lst_dom l' = v) - (Hcodoml : a_lst_codom l = w_Sn (len0 v - 1)) - (Hcodoml' : a_lst_codom l' = w_Sn (len0 v - 1)), - (interp_a_path_prop c l (proj1 Hl) ≃ _)%Cat). - rewrite Hdoml, <- Hdoml', Hcodoml, <- Hcodoml'. - exact (interp_a_path_prop c l' (proj1 Hl')). -Defined. - -Lemma generalized_induction {T} {P : T -> Prop} - (f : T -> nat) : - (forall t, f t = 0 -> P t) -> - (forall t, (forall s, f s < f t -> P s) -> P t) -> - forall t, P t. -Proof. - intros Hbase Hrec. - enough (Hen: forall n, forall t, f t <= n -> P t) - by (intros t; apply (Hen (f t) t (le_n _))). - intros n. - induction n. - - intros t Hft. - apply Hbase; lia. - - intros t Hft. - destruct (Nat.lt_trichotomy (f t) (S n)) - as [Hlt | [Heq | Hfalse]]; [| | lia]. - + apply IHn; lia. - + apply Hrec. - intros s Hs. - apply IHn; lia. -Qed. - -Lemma generalized_induction_base {T} {P : T -> Prop} - (f : T -> nat) (base : nat) : - (forall t, f t <= base -> P t) -> - (forall t, (forall s, f s < f t -> P s) -> P t) -> - forall t, P t. -Proof. - intros Hbase Hrec. - apply (generalized_induction (fun t => f t - base)). - - intros t Ht; apply Hbase; lia. - - intros t Hless. - apply Hrec. - intros s Hs. - destruct (Nat.lt_trichotomy (f s) base) as [Hlt | [Heq | Hgt]]; - [apply Hbase; lia | apply Hbase; lia |]. - apply Hless; lia. -Qed. - -Require Import Coq.Init.Wf. - -Lemma generalized_double_induction {T} {P : T -> Prop} - (f g : T -> nat) : - (forall t, f t = 0 -> P t) -> - (forall t, g t = 0 -> P t) -> - (forall t, - (forall s, f s = f t -> g s < g t -> P s) -> - (forall s, f s < f t -> P s) -> P t) -> - forall t, P t. -Proof. - - intros Hbasef Hbaseg Hrec. - enough (Hpair:forall s t : T, P s /\ P t) by - (intros t; apply (Hpair t t)). - set (lexico_rel (* : nat * nat -> nat * nat -> Prop *) := - fun ab cd :nat*nat=> let (a,b) := ab in let (c,d):=cd in - a < c \/ (a = c /\ b < d)). - intros s t; pattern t; revert t; pattern s; revert s. - apply well_founded_induction_type_2 with lexico_rel. - - intros (a, b). - induction a; induction b. - + constructor. - intros (c, d) Hy; unfold lexico_rel in *. - lia. - + constructor. - intros (c, d). - intros [Hfalse | [Hc Hd]]; [unfold lexico_rel in *; lia|]. subst c. - destruct (Nat.lt_trichotomy d b) as [Hlt | [Heq | Hf]]; - [ | | lia]. - * apply Acc_inv with (0, b). - apply IHb. - hnf; lia. - * subst d; easy. - + constructor. - intros (a', b'). - unfold lexico_rel. - intros [Ha' | Hf]; [|lia]. - destruct (Nat.lt_trichotomy a' a) as [Hlt | [Heq | Hf]]; - [ | | lia]. - * apply Acc_inv with (a, 0); easy || lia. - * subst a'. - induction b'; [easy|]. - constructor. - intros (a'', b''). - intros [Has | [Haeq Hb]]; - [ apply Acc_inv with (a, b'); easy || lia | ]. - subst a''. - destruct (Nat.lt_trichotomy b'' b') as [Hlt | [Heq | Hf]]; - [ | | lia]. - --apply Acc_inv with (a, b'); easy || lia. - --subst b''; easy. - + specialize (IHb ltac:(apply Acc_inv with (a, S b); hnf; easy || lia)). - constructor. - intros (c, d). - intros [Halt | [Haeq Hblt]]. - * apply Acc_inv with (S a, b); hnf; easy || lia. - * subst c. - induction d. - --destruct b; [easy|]. - apply Acc_inv with (S a, S b); hnf; easy || lia. - --constructor. - intros (c', d'). - intros [Hclt | [Hceq Hdlt]]. - ++apply Acc_inv with (S a, b); hnf; easy || lia. - ++destruct (Nat.lt_trichotomy d' d) as [Hlt | [Heq | Hf]]; - [ | | lia]. - **apply Acc_inv with (S a, d); hnf; apply IHd || lia; lia. - **subst; apply IHd; lia. - - intros x x' Hind. - Abort. - - -Lemma unique_directed_path_image : unique_directed_path_image_prop. -Proof. - unfold unique_directed_path_image_prop. - intros C cC mC c v. - pattern v. - match goal with |- ?f ?v => set (gP := f) end. - assert (HwSn: forall w, (exists n, w = w_Sn n) -> gP w). 1:{ - subst gP. - intros w [n Hn]. - subst w. - intros. - unfold a_dir_path in Hl. - destruct l. - + simpl in *. - pose proof (rho0_dir_decreasing a (proj2 Hl)) as H. - replace (a_dom a) in H. - rewrite rho0_w_Sn in H. - easy. - + simpl in *. - pose proof (rho0_dir_decreasing a (proj1 (proj2 Hl))) as H. - replace (a_dom a) in H. - rewrite rho0_w_Sn in H; easy. - } - apply (generalized_induction rho0). - - intros t. - rewrite rho_eq0_iff. - intros Ht. - apply HwSn. - exists (len0 t - 1); easy. - - intros t; - pattern t; revert t. - apply (generalized_induction_base len0 1). - + intros t Ht. - destruct t; try (simpl in Ht; lia). - * intros _. - apply HwSn. - exists 0; easy. - * simpl in Ht. - pose (len0_pos t1); pose (len0_pos t2); lia. - Abort. - - - - - - - - - - - - - - - - - - - - - - - - - -Inductive arrow_S' := - | arr_W (b : W) - | arr_assoc (u v w : W) - | arr_invassoc (u v w : W) - | arr_lunit (b : W) - | arr_invlunit (b : W) - | arr_runit (b : W) - | arr_invrunit (b : W) - | arr_tensor_l : arrow_S' -> W -> arrow_S' - | arr_tensor_r : arrow_S' -> W -> arrow_S'. - -Class Quiver := { - q_edges : Type; - q_verts : Type; - dom : q_edges -> q_verts; - codom : q_edges -> q_verts; -}. - -Fixpoint arrow_dom (a : arrow_S') : W := - match a with - | arr_W b => b - | arr_assoc u v w => tens (tens u v) w - | arr_invassoc u v w => tens u (tens v w) - | arr_lunit b => tens e0 b - | arr_invlunit b => b - | arr_runit b => tens b e0 - | arr_invrunit b => b - | arr_tensor_l a b => - tens b (arrow_dom a) - | arr_tensor_r a b => - tens (arrow_dom a) b - end. - -Fixpoint arrow_codom (a : arrow_S') : W := - match a with - | arr_W b => b - | arr_assoc u v w => tens u (tens v w) - | arr_invassoc u v w => tens (tens u v) w - | arr_lunit b => b - | arr_invlunit b => tens e0 b - | arr_runit b => b - | arr_invrunit b => tens b e0 - | arr_tensor_l s b => - tens b (arrow_codom s) - | arr_tensor_r s b => - tens (arrow_codom s) b - end. - -#[export] Instance quiver_S' : Quiver := { - q_edges := arrow_S'; - q_verts := W; - dom := arrow_dom; - codom := arrow_codom; -}. - -Fixpoint len (b : W) : nat := - match b with - | e0 => 0 - | var => 1 - | tens b0 b1 => len b0 + len b1 - end. - -Lemma len_dom_codom : forall s, - len (arrow_dom s) = len (arrow_codom s). -Proof. - induction s; simpl; lia. -Qed. - - -Fixpoint word_to_mc {C} {cC : Category C} (mC : MonoidalCategory cC) - (b : W) : C -> C := - fun c => - match b with - | e0 => mC.(mon_I) - | var => c - | tens b0 b1 => (word_to_mc mC b0 c) × (word_to_mc mC b1 c) - end. - - -(* Local Open Scope Mor_scope. *) -Fixpoint arr_to_mc {C} {cC : Category C} (mC : MonoidalCategory cC) - (s : arrow_S') (c : C) : - ((word_to_mc mC (arrow_dom s) c) ~> - (word_to_mc mC (arrow_codom s) c))%Cat := - match s with - | arr_W b => id_ _ - | arr_assoc u v w => - associator (word_to_mc mC u c) (word_to_mc mC v c) (word_to_mc mC w c) - | arr_invassoc u v w => - (associator (word_to_mc mC u c) (word_to_mc mC v c) (word_to_mc mC w c))^-1 - | arr_lunit b => - λ_ _ - | arr_invlunit b => (λ_ _)^-1 - | arr_runit b => (ρ_ _) - | arr_invrunit b => (ρ_ _)^-1 - | arr_tensor_l s' b => - id_ (word_to_mc mC b c) ⊗ (arr_to_mc mC s' c) - | arr_tensor_r s' b => - (arr_to_mc mC s' c) ⊗ id_ (word_to_mc mC b c) - end%Mor. - -Inductive arr_lst := - | lst_init : arrow_S' -> arr_lst - | lst_cons : arrow_S' -> arr_lst -> arr_lst. - -Definition arr_lst_dom (ls : arr_lst) : W := - match ls with - | lst_init s - | lst_cons s _ => dom s - end. - -Fixpoint arr_lst_codom (ls : arr_lst) : W := - match ls with - | lst_init s => codom s - | lst_cons _ l => arr_lst_codom l - end. - -Fixpoint arr_path (ls : arr_lst) : bool := - match ls with - | lst_init s => true - | lst_cons s l => W_eqb (codom s) (arr_lst_dom l) && arr_path l - end. - -Fixpoint arr_path_prop (ls : arr_lst) : Prop := - match ls with - | lst_init s => True - | lst_cons s l => (codom s) = (arr_lst_dom l) /\ arr_path_prop l - end. - -Lemma arr_lst_pathP (ls : arr_lst) : - arr_path_prop ls <-> arr_path ls = true. -Proof. - induction ls; [easy|]. - - simpl. - (* rewrite andb_true_iff, W_eqbP. - split; intros []; split; easy || apply IHls; easy. *) -Admitted. - -Definition morphism_of_arr_path {C} {cC : Category C} (mC : MonoidalCategory cC) - (c : C) (ls : arr_lst) (H : arr_path_prop ls) : - (word_to_mc mC (arr_lst_dom ls) c ~> word_to_mc mC (arr_lst_codom ls) c)%Cat. -induction ls. -- apply arr_to_mc. -- eapply compose. - + apply arr_to_mc. - + simpl in H. rewrite (proj1 H). - apply IHls, H. -Defined. - - - -Definition coherence_theorem_v1 : Prop. -refine ( - forall {C} {cC : Category C} (mC : MonoidalCategory cC) (c : C) - (v w : W) (Hlen : len v = len w) (ls ls' : arr_lst) - (Hls : arr_path_prop ls) (Hls' : arr_path_prop ls') - (Hdom : arr_lst_dom ls = arr_lst_dom ls') - (Hcodom : arr_lst_codom ls = arr_lst_codom ls'), - _ : Prop). -pose (morphism_of_arr_path mC c ls Hls) as morls; -pose (morphism_of_arr_path mC c ls' Hls') as morls'. -rewrite Hdom, Hcodom in morls. -exact (morls ≃ morls')%Cat. -Defined. - -Section VersionTwo. - -Require Import Arith. - -#[export, program] Instance WordCat : Category W := { - morphism := fun b b' => if len b =? len b' then True else False; - c_equiv := fun b b' f f' => True; - compose := fun b b' b'' f g => _; -}. -Next Obligation. - split; destruct (len A =? len B); easy. -Qed. -Next Obligation. - destruct (len b =? len b') eqn:e. - - rewrite Nat.eqb_eq in e. - rewrite e. - apply g. - - destruct f. -Defined. -Next Obligation. - rewrite Nat.eqb_refl. - exact Logic.I. -Defined. - -Local Open Scope Cat_scope. - - -Context {C : Type} {cC : Category C} {D : Type} {cD : Category D}. - -Class StrictMonoidalFunctor - (mC : MonoidalCategory cC) (mD : MonoidalCategory cD) : Type := { - strict_functor : Functor cC cD; - mon_I_eq : strict_functor I = I; - tensor_eq (c c' : C) : - strict_functor (c × c') = strict_functor c × strict_functor c'; - mon_mu_ij (i j : C) : - strict_functor i × strict_functor j <~> strict_functor (i × j) := - eq_rect_r (fun d : D => strict_functor i × strict_functor j <~> d) - (IdentityIsomorphism (strict_functor i × strict_functor j)) - (tensor_eq i j); - mon_eps : I <~> strict_functor I := - eq_rect_r (fun d : D => I <~> d) - (IdentityIsomorphism I) - mon_I_eq; -}. - -Coercion strict_functor : StrictMonoidalFunctor >-> Functor. - -Context {mC : MonoidalCategory cC} {mD : MonoidalCategory cD}. - -(* Definition mu_ij (F : StrictMonoidalFunctor mC mD) : forall (i j : C), -F i × F j <~> F (i × j). -intros. -rewrite tensor_eq. -apply IdentityIsomorphism. -Defined. *) - -Lemma strict_monoidal_functor_assoc (F : StrictMonoidalFunctor mC mD) - : forall (x y z : C), - associator (F x) (F y) (F z) ∘ id_ (F x) ⊗ mon_mu_ij y z ∘ mon_mu_ij x (y × z) - ≃ mon_mu_ij x y ⊗ id_ (F z) ∘ mon_mu_ij (x × y) z ∘ F @ associator x y z. -Proof. - intros. - Abort. - - -End VersionTwo. - -End Associativity. - - diff --git a/examples/CatExample.v b/examples/CatExample.v index 5bc44a8..dac0d5d 100644 --- a/examples/CatExample.v +++ b/examples/CatExample.v @@ -5,35 +5,21 @@ Open Scope Cat_scope. Set Universe Polymorphism. -Generalizable Variables A B C D E cA cB cC cD cE. +(* Generalizable Variables A B C D E cA cB cC cD cE. *) -Record Cat := { base_type : Type; cat_inst : Category base_type }. +Section IsomorphismInstances. -Program Definition ComposeFunctors - `{cC : Category C, cD : Category D, cE : Category E} - (F : Functor cC cD) (G : Functor cD cE) : Functor cC cE := {| - obj_map := fun A => G (F A); - morphism_map := fun A B f => G @ (F @ f); - |}. -Next Obligation. - rewrite 2!id_map; easy. Qed. -Next Obligation. - rewrite 2!compose_map; easy. Qed. -Next Obligation. - apply morphism_compat, morphism_compat; easy. -Qed. -Unset Printing Universes. +Context {C:Type} {cC : Category C} {cCh : CategoryCoherence cC}. -Notation "F '○' G" := (ComposeFunctors F G) (at level 59, left associativity). - -#[export] Instance IdentityIsomorphism `{cC : Category C} (A : C) +#[export] Instance IdentityIsomorphism (A : C) : Isomorphism A A := {| forward := id_ A; reverse := id_ A; isomorphism_inverse := ltac:(rewrite left_unit; easy) |}. -Definition ReverseIsomorphism `{cC : Category C} {A B : C} +Definition ReverseIsomorphism {C : Type} {cC : Category C} + {cCh : CategoryCoherence cC} {A B : C} (f : A <~> B) : B <~> A := {| forward := f.(reverse); reverse := f.(forward); @@ -41,39 +27,61 @@ Definition ReverseIsomorphism `{cC : Category C} {A B : C} let (hA, hB) := f.(isomorphism_inverse) in conj hB hA |}. -Program Definition ComposeIsomorphisms `{cC : Category C} {A B M : C} +Program Definition ComposeIsomorphisms {A B M : C} (f : A <~> B) (g : B <~> M) := {| forward := f ∘ g; reverse := g^-1 ∘ f^-1; |}. Next Obligation. split. - - rewrite <- assoc, (assoc (f:=f)), (id_A g), right_unit, (id_A f); easy. - - rewrite <- assoc, (assoc (f:=g^-1)), (id_B f), right_unit, (id_B g); easy. + - rewrite <- assoc, (assoc (f)), (iso_inv_r g), right_unit, (iso_inv_r f); easy. + - rewrite <- assoc, (assoc (g^-1)), (iso_inv_l f), right_unit, (iso_inv_l g); easy. +Qed. + +Program Definition Isomorphism_by_Functor + {D} {cD : Category D} {cDh : CategoryCoherence cD} + {A B : C} (F : Functor cC cD) (H : A <~> B) : F A <~> F B := + {| + forward := F @ H; + reverse := F @ H^-1 + |}. +Next Obligation. + split; rewrite <- compose_map. + - rewrite (iso_inv_r H), id_map; easy. + - rewrite (iso_inv_l H), id_map; easy. Qed. - -Program Definition ReverseNaturalIsomorphism `{cC : Category C, cD : Category D} +End IsomorphismInstances. + +Section NaturalTransformationInstances. + +Context {C D E : Type} + {cC : Category C} {cD : Category D} {cE : Category E} + {cCh : CategoryCoherence cC} + {cDh : CategoryCoherence cD} + {cEh : CategoryCoherence cE}. + +Program Definition ReverseNaturalIsomorphism {F G : Functor cC cD} (N : NaturalIsomorphism F G) : NaturalIsomorphism G F := {| - component_iso := fun A => ReverseIsomorphism (N A) + component_iso := fun A => (ReverseIsomorphism (N A)); |}. Next Obligation. match goal with - |- ?g ≃ ?h => rewrite <- (right_unit (f:=h)) + |- ?g ≃ ?h => rewrite <- (right_unit h) end. rewrite <- (proj1 (N B).(isomorphism_inverse)). - rewrite <- assoc, !(assoc (h:= N B)), N.(component_iso_natural). + rewrite <- assoc, !(assoc _ _ (N B)), N.(component_iso_natural). apply compose_compat; [|easy]. rewrite <- assoc, (proj2 (N A).(isomorphism_inverse)). rewrite left_unit. easy. Qed. -Program Definition ComposeNaturalTransformations `{cC : Category C, cD : Category D} +Program Definition ComposeNaturalTransformations {F G H : Functor cC cD} (N : NaturalTransformation F G) (M : NaturalTransformation G H) : NaturalTransformation F H := {| - component_map := fun A => (α_ A) ∘ (α_ A) + component_map := fun A => (N.(component_map) A) ∘ (M.(component_map) A) |}. Next Obligation. rewrite <- assoc, N.(component_map_natural), @@ -84,33 +92,49 @@ Qed. Notation "N '⨀' M" := (ComposeNaturalTransformations N M) (at level 59, left associativity) : Cat_scope. -Program Definition ComposeNaturalIsomorphisms `{cC : Category C, cD : Category D} +Program Definition ComposeNaturalIsomorphisms {F G H : Functor cC cD} (N : NaturalIsomorphism F G) (M : NaturalIsomorphism G H) : NaturalIsomorphism F H := {| component_iso := fun A => ComposeIsomorphisms (N A) (M A) |}. Next Obligation. - rewrite <- assoc, N.(component_map_natural), - assoc, M.(component_map_natural), <- assoc. + rewrite <- assoc, N.(component_iso_natural), + assoc, M.(component_iso_natural), <- assoc. easy. Qed. -Program Definition Isomorphism_by_Functor `{cC : Category C, cD : Category D} - {A B : C} (F : Functor cC cD) (H : A <~> B) : F A <~> F B := - {| - forward := F @ H; - reverse := F @ H^-1 +Program Definition ComposeFunctors {C D E : Type} + {cC : Category C} {cCh : CategoryCoherence cC} + {cD : Category D} {cDh : CategoryCoherence cD} + {cE : Category E} {cEh : CategoryCoherence cE} + (F : Functor cC cD) (G : Functor cD cE) : Functor cC cE := {| + obj_map := fun A => G (F A); + morphism_map := fun A B f => G @ (F @ f); |}. Next Obligation. - split; rewrite <- compose_map. - - rewrite (id_A H), id_map; easy. - - rewrite (id_B H), id_map; easy. + rewrite 2!id_map; easy. Qed. +Next Obligation. + rewrite 2!compose_map; easy. Qed. +Next Obligation. + apply morphism_compat, morphism_compat; easy. Qed. +Notation "F '○' G" := (ComposeFunctors F G) + (at level 59, left associativity) : Cat_scope. + +Program Definition ComposeFunctorsAssociator {A B C D} + {cA : Category A} {cB : Category B} {cC : Category C} {cD : Category D} + {cAh : CategoryCoherence cA} {cBh : CategoryCoherence cB} + {cCh : CategoryCoherence cC} {cDh : CategoryCoherence cD} + (F : Functor cA cB) (G : Functor cB cC) (H : Functor cC cD) : + NaturalIsomorphism (F ○ G ○ H) (F ○ (G ○ H)) := + {| component_iso := fun A => IdentityIsomorphism _ |}. +Next Obligation. + rewrite left_unit, right_unit; easy. +Qed. Program Definition NaturalIsomorphism_of_Compose - `{cC : Category C, cD : Category D, cE : Category E} {F F' : Functor cC cD} {G G' : Functor cD cE} (N : NaturalIsomorphism F F') (M : NaturalIsomorphism G G') : NaturalIsomorphism (F ○ G) (F' ○ G') := @@ -120,7 +144,7 @@ Program Definition NaturalIsomorphism_of_Compose |}. Next Obligation. rewrite 2!component_iso_natural, <- assoc. - rewrite component_iso_natural, 2!(assoc (f:=M (F A))). + rewrite component_iso_natural, 2!(assoc (M (F A))). apply compose_compat; [easy|]. rewrite <- 2!compose_map. rewrite component_iso_natural. @@ -128,35 +152,48 @@ Next Obligation. Qed. -Program Definition IdentityFunctor `(cC : Category C) : Functor cC cC := +Program Definition IdentityFunctor {C} (cC : Category C) + {cCh : CategoryCoherence cC} : Functor cC cC := {| obj_map := id; morphism_map := fun _ _ => id |}. Solve All Obligations with (rewrite ?left_unit, ?right_unit; easy). -Program Definition IdentityNaturalIsomorphism - `{cC : Category C, cD : Category D} (F : Functor cC cD) : +Program Definition IdentityNaturalIsomorphism {C D} + {cC : Category C} {cD : Category D} + {cCh : CategoryCoherence cC} {cDh : CategoryCoherence cD} + (F : Functor cC cD) : NaturalIsomorphism F F := {| component_iso := fun A => IdentityIsomorphism _ |}. Next Obligation. rewrite ?left_unit, ?right_unit; easy. Qed. +End NaturalTransformationInstances. + +Notation "N '⨀' M" := (ComposeNaturalTransformations N M) + (at level 59, left associativity) : Cat_scope. +Notation "F '○' G" := (ComposeFunctors F G) + (at level 59, left associativity) : Cat_scope. + +Section NaturallyIsomorphic. -Definition naturally_isomorphic `{cC : Category C, cD : Category D} +Context {C D E : Type} + {cC : Category C} {cCh : CategoryCoherence cC} + {cD : Category D} {cDh : CategoryCoherence cD}. + +Definition naturally_isomorphic {C D} {cC : Category C} {cD : Category D} : relation (Functor cC cD) := fun F G => exists (comp_map : forall A : C, F A <~> G A), forall (A B : C) (f : A ~> B), F @ f ∘ comp_map B ≃ comp_map A ∘ G @ f. -Lemma naturally_isomorphic_of_NaturalIsomorphism - `{cC : Category C, cD : Category D} {F G : Functor cC cD} : +Lemma naturally_isomorphic_of_NaturalIsomorphism {F G : Functor cC cD} : NaturalIsomorphism F G -> naturally_isomorphic F G. Proof. intros [c Hc]. exists c; exact Hc. Qed. -Lemma naturally_isomorphic_iff_ex_NaturalIsomorphism - `{cC : Category C, cD : Category D} (F G : Functor cC cD) : +Lemma naturally_isomorphic_iff_ex_NaturalIsomorphism (F G : Functor cC cD) : naturally_isomorphic F G <-> (exists _ : NaturalIsomorphism F G, True). Proof. split. @@ -166,16 +203,16 @@ Proof. - intros [[c Hc] _]; exists c; exact Hc. Qed. -Lemma naturally_isomorphic_refl `{cC : Category C, cD : Category D} : - reflexive _ (@naturally_isomorphic C cC D cD). +Lemma naturally_isomorphic_refl : + reflexive _ (@naturally_isomorphic C D cC cD). Proof. intros F. apply naturally_isomorphic_of_NaturalIsomorphism, IdentityNaturalIsomorphism. Qed. -Lemma naturally_isomorphic_sym `{cC : Category C, cD : Category D} : - symmetric _ (@naturally_isomorphic C cC D cD). +Lemma naturally_isomorphic_sym : + symmetric _ (@naturally_isomorphic C D cC cD). Proof. intros F G. rewrite naturally_isomorphic_iff_ex_NaturalIsomorphism. @@ -184,8 +221,8 @@ Proof. exact (ReverseNaturalIsomorphism N). Qed. -Lemma naturally_isomorphic_trans `{cC : Category C, cD : Category D} : - transitive _ (@naturally_isomorphic C cC D cD). +Lemma naturally_isomorphic_trans : + transitive _ (@naturally_isomorphic C D cC cD). Proof. intros F G H. rewrite 2!naturally_isomorphic_iff_ex_NaturalIsomorphism. @@ -194,31 +231,36 @@ Proof. exact (ComposeNaturalIsomorphisms N M). Qed. -Definition in_cat_of_cat {U} {cU : Category U} : Cat := - {| base_type := U |}. - -Definition cat_of_in_cat {U : Cat} := U.(cat_inst). -Coercion cat_of_in_cat : Cat >-> Category. +End NaturallyIsomorphic. +Section CategoryCat. -Program Definition ComposeFunctorsAssociator - `{cA : Category A, cB : Category B, cC : Category C, cD : Category D} - (F : Functor cA cB) (G : Functor cB cC) (H : Functor cC cD) : - NaturalIsomorphism (F ○ G ○ H) (F ○ (G ○ H)) := - {| component_iso := fun A => IdentityIsomorphism _ |}. -Next Obligation. - rewrite left_unit, right_unit; easy. -Qed. - -Lemma compose_functors_assoc - `{cA : Category A, cB : Category B, cC : Category C, cD : Category D} +Lemma compose_functors_assoc {A B C D} + {cA : Category A} {cB : Category B} {cC : Category C} {cD : Category D} + {cAh : CategoryCoherence cA} {cBh : CategoryCoherence cB} + {cCh : CategoryCoherence cC} {cDh : CategoryCoherence cD} (F : Functor cA cB) (G : Functor cB cC) (H : Functor cC cD) : naturally_isomorphic (F ○ G ○ H) (F ○ (G ○ H)). Proof. exact (naturally_isomorphic_of_NaturalIsomorphism (ComposeFunctorsAssociator F G H)). Qed. +Record Cat := { + base_type : Type; + cat_inst : Category base_type; + cat_coh :> CategoryCoherence cat_inst +}. + +Existing Instance cat_coh. + +Definition in_cat_of_category {U} (cU : Category U) + {cUh : CategoryCoherence cU} : Cat := + {| base_type := U |}. + +Coercion cat_inst : Cat >-> Category. + + #[export] Program Instance CatCategory : Category Cat := {| @@ -227,6 +269,8 @@ Qed. compose := fun A B C => ComposeFunctors; c_identity := fun A => IdentityFunctor A |}. + +#[export] Program Instance CatCategoryC : CategoryCoherence CatCategory. Next Obligation. split. - apply naturally_isomorphic_refl. @@ -236,13 +280,12 @@ Qed. Next Obligation. destruct H as [cfg Hcfg]. destruct H0 as [chj Hchj]. - eexists (fun A => ComposeIsomorphisms (Isomorphism_by_Functor h (cfg A)) (chj (g A))). intros x y fxy. simpl. rewrite 2!Hchj, <- assoc. - rewrite Hchj, 2!(assoc (f:=chj (f x))). + rewrite Hchj, 2!(assoc (chj (f x))). apply compose_compat; [easy|]. rewrite <- 2!compose_map. rewrite Hcfg. @@ -268,7 +311,17 @@ Next Obligation. easy. Qed. -Class CategoryProduct `{cC : Category C} (A B : C) := { +End CategoryCat. + +Section CategoryCatProperties. + + +Context {C D E : Type} + {cC : Category C} {cCh : CategoryCoherence cC} + {cD : Category D} {cDh : CategoryCoherence cD} + {cE : Category E} {cEh : CategoryCoherence cE}. + +Class CategoryProduct (A B : C) := { prod_AB : C; p_A : prod_AB ~> A; p_B : prod_AB ~> B; @@ -287,7 +340,7 @@ Class CategoryProduct `{cC : Category C} (A B : C) := { Local Notation "'@' AB" := (AB.(prod_AB)) (at level 8) : Cat_scope. -Lemma prod_mor_unique' `{cC : Category C} {A B : C} +Lemma prod_mor_unique' {A B : C} (AB : CategoryProduct A B) {Q} (fA : Q ~> A) (fB : Q ~> B) : forall (fAB fAB' : Q ~> AB.(prod_AB)), fAB ∘ p_A ≃ fA -> fAB ∘ p_B ≃ fB -> @@ -301,7 +354,7 @@ Proof. easy. Qed. -Program Definition category_product_unique `{cC : Category C} (A B : C) : +Program Definition category_product_unique (A B : C) : forall (AB AB' : CategoryProduct A B), Isomorphism @AB @AB' := fun AB AB' => {| @@ -315,25 +368,25 @@ Next Obligation. all: rewrite left_unit; easy. Qed. -Class CategoryBigProd `{cC : Category C} {I : Type} - (obj : I -> C) (prod_I : C) := { - p_i : forall i, prod_I ~> obj i; +Class CategoryBigProd {J : Type} + (obj : J -> C) (prod_J : C) := { + p_i : forall i, prod_J ~> obj i; big_cat_prod_mor : - forall (Q : C) (fQ_ : forall i, Q ~> obj i), Q ~> prod_I; + forall (Q : C) (fQ_ : forall i, Q ~> obj i), Q ~> prod_J; big_cat_prod_mor_prop: forall (Q : C) (fQ_ : forall i, Q ~> obj i), forall i, (big_cat_prod_mor Q fQ_) ∘ p_i i ≃ fQ_ i; big_cat_prod_mor_unique : forall (Q : C) (fQ_ : forall i, Q ~> obj i) - (fAB' : Q ~> prod_I), + (fAB' : Q ~> prod_J), (forall i, fQ_ i ≃ fAB' ∘ p_i i) -> fAB' ≃ big_cat_prod_mor Q fQ_ }. -Lemma big_cat_prod_mor_unique' `{cC : Category C} {I} {obj : I -> C} {pI : C} - (HI : CategoryBigProd obj pI) {Q} (fQ_ : forall i, Q ~> obj i) : - forall (fAB fAB' : Q ~> pI), +Lemma big_cat_prod_mor_unique' {J} {obj : J -> C} {pJ : C} + (HJ : CategoryBigProd obj pJ) {Q} (fQ_ : forall i, Q ~> obj i) : + forall (fAB fAB' : Q ~> pJ), (forall i, fAB ∘ p_i i ≃ fQ_ i) -> (forall i, fAB' ∘ p_i i ≃ fQ_ i) -> fAB ≃ fAB'. @@ -345,13 +398,13 @@ Proof. easy. Qed. -Program Definition category_big_cat_prod_unique `{cC : Category C} {I} {obj : I->C} : - forall {pI pI'} (HI : CategoryBigProd obj pI) (HI' : CategoryBigProd obj pI'), - Isomorphism pI pI' := - fun pI pI' HI HI' => +Program Definition category_big_cat_prod_unique {J} {obj : J -> C} : + forall {pJ pJ'} (HJ : CategoryBigProd obj pJ) (HI' : CategoryBigProd obj pJ'), + Isomorphism pJ pJ' := + fun pJ pJ' HJ HJ' => {| - forward := HI'.(big_cat_prod_mor) pI p_i; - reverse := HI.(big_cat_prod_mor) pI' p_i; + forward := HJ'.(big_cat_prod_mor) pJ p_i; + reverse := HJ.(big_cat_prod_mor) pJ' p_i; |}. Obligations. Next Obligation. @@ -369,15 +422,20 @@ Definition nat_trans_equiv `{cC : Category C, cD : Category D} (* #[program] Definition nat_trans_equiv_equivalence *) +Coercion NaturalTransformation_of_NaturalIsomorphism + : NaturalIsomorphism >-> NaturalTransformation. (* I'm fine exporting, seeing as we're unlikely to be asking for any other category on Functor cC cD... or any at all! *) -#[export, program] Instance FunctorCategory `(cC : Category C, cD : Category D) : +#[export, program] Instance FunctorCategory : Category (Functor cC cD) := {| morphism := @NaturalTransformation C D cC cD; c_equiv := nat_trans_equiv; compose := fun F G H => ComposeNaturalTransformations; c_identity := fun F => IdentityNaturalIsomorphism F; |}. + +#[export, program] Instance FunctorCategoryC : + CategoryCoherence FunctorCategory. Next Obligation. split; try easy. - intros F G H h1 h2 ?. @@ -402,15 +460,20 @@ Next Obligation. rewrite right_unit; easy. Qed. -Definition big_prod {I : Type} (CI : I -> Type) := forall i, CI i. +Definition big_prod {J : Type} (CJ : J -> Type) := forall i, CJ i. -#[export, program] Instance BigProductCategory {I : Type} (CI : I -> Type) - (cI : forall i : I, Category (CI i)) : Category (big_prod CI) := {| +#[export, program] Instance BigProductCategory {J : Type} (CJ : J -> Type) + (cJ : forall i : J, Category (CJ i)) : Category (big_prod CJ) := {| morphism := fun is js => forall i, is i ~> js i; c_equiv := fun is js fis fjs => forall i, fis i ≃ fjs i; compose := fun is js ks fijs fjks => fun i => fijs i ∘ fjks i; c_identity := fun is => fun i => id_ (is i); |}. + +#[export, program] Instance BigProductCategoryC {J : Type} (CJ : J -> Type) + (cJ : forall i : J, Category (CJ i)) + (cJh : forall i : J, CategoryCoherence (cJ i)) : + CategoryCoherence (BigProductCategory CJ cJ). Next Obligation. split; intros ? **; try easy. - etransitivity; easy. @@ -428,18 +491,20 @@ Next Obligation. rewrite right_unit; easy. Qed. -#[program, export] Instance BigProductComponentFunctor {I} (CI : I -> Type) - (cI : forall i : I, Category (CI i)) (i : I) : - Functor (BigProductCategory CI cI) (cI i) := { +#[program, export] Instance BigProductComponentFunctor {J} (CJ : J -> Type) + (cJ : forall i : J, Category (CJ i)) + (cJh : forall i : J, CategoryCoherence (cJ i)) + (i : J) : + Functor (BigProductCategory CJ cJ) (cJ i) := { obj_map := fun is => is i; morphism_map := fun is js fijs => fijs i; }. Solve All Obligations with easy. -#[export] Instance BigProductFunctor_of_Components {I} - (CI : I -> Type) (cI : forall i : I, Category (CI i)) (Q : Cat) - (fQ_ : forall i, Functor Q.(cat_inst) (cI i)) : - Functor Q (BigProductCategory CI cI) := { +#[export] Instance BigProductFunctor_of_Components {J} + (CJ : J -> Type) (cJ : forall i : J, Category (CJ i)) (Q : Cat) + (fQ_ : forall i, Functor Q.(cat_inst) (cJ i)) : + Functor Q (BigProductCategory CJ cJ) := { obj_map := fun q => fun i => fQ_ i q; morphism_map := fun q r fqr => fun i => (fQ_ i).(morphism_map) fqr; id_map := fun a => fun i => (fQ_ i).(id_map) a; @@ -447,8 +512,8 @@ Solve All Obligations with easy. morphism_compat := fun a b f g H => fun i => (fQ_ i).(morphism_compat) f g H; }. -#[export, program] Instance BigProductIsomorphism_of_Components {I} - (CI : I -> Type) (cI : forall i : I, Category (CI i)) (As Bs : big_prod CI) +#[export, program] Instance BigProductIsomorphism_of_Components {J} + (CJ : J -> Type) (cJ : forall i : J, Category (CJ i)) (As Bs : big_prod CJ) (Hiso : forall i, As i <~> Bs i) : As <~> Bs := { forward := fun i => Hiso i; reverse := fun i => Hiso i ^-1; @@ -457,248 +522,4 @@ Next Obligation. split; intros i; apply (Hiso i).(isomorphism_inverse). Qed. -(* -Lemma exists_big_product_isomorphism_iff_forall_exists {I} - (CI : I -> Type) (cI : forall i : I, Category (CI i)) (As Ba : big_prod CI) P : - exists f : As <~> Bs, P f <-> forall i, exists fi : As i <~> Bs i *) - -(* #[program, export] Instance big_product_category_is_product {I} - (CI : I -> Type) (cI : forall i : I, Category (CI i)) : - CategoryBigProd (cC := CatCategory) - (fun i => {|base_type:=CI i; cat_inst:=cI i|}) - {| base_type := big_prod CI; cat_inst := BigProductCategory CI cI |} := { - p_i := fun i => BigProductComponentFunctor CI cI i; - big_cat_prod_mor := BigProductFunctor_of_Components CI cI; -}. -Next Obligation. - exists (fun a => IdentityIsomorphism _). - intros; rewrite right_unit, left_unit; easy. -Qed. -Next Obligation. - (* edestruct H as [Hiso Hnat]. *) - (* unfold naturally_isomorphic in H. *) - - - assert (Hex: exists f : forall i, NaturalIsomorphism (fQ_ i) (fAB' ○ BigProductComponentFunctor CI cI i), True) - by admit. - destruct Hex as [Hiso _]. - eexists. - unshelve (instantiate (1:= (fun a => _))). - - eapply Build_Isomorphism. - unshelve (instantiate (1:= (fun i => _))). - exact (Hiso a).(reverse). - - - pose (ReverseIsomorphism (Hiso a)). - simpl in *. - - eapply Build_Isomorphism. - unshelve (instantiate (1:= (fun i => _))). - exact (X i a).(reverse). - pose (X a).(reverse) as m. - exact m. - simpl in m. - simpl. - destruct (H i) as [N _]. - - - refine (ex_intro _ (fun a => _), _). - - -Definition big_sum {I : Type} (CI : I -> Type) := {i & CI i}. - -#[program] Definition BigSumCategory {I : Type} - (UIPI : forall (i:I) (p q : i = i), p = q) - (UIUIPI : forall (i:I) (p q : i = i) (H H' : p = q), H = H') - (CI : I -> Type) - (cI : forall i : I, Category (CI i)) : Category (big_sum CI) := {| - morphism := fun somei somej => let (i, ai) := somei in let (j, aj) := somej in - {H : i = j & _}; - c_equiv := fun somei somej somefi somefj => _; - (* c_identity := fun somei => let (i, ai) := somei in existT _ eq_refl (id_ ai); *) - (* compose := fun is js ks fijs fjks => fun i => fijs i ∘ fjks i; - c_identity := fun is => fun i => id_ (is i); *) -|}. -Next Obligation. - exact (ai ~> aj). -Defined. -Next Obligation. - destruct somei as [i ai]; destruct somej as [j aj]. - destruct somefi as [Hij mori], somefj as [? morj]. - subst. - rewrite (UIPI _ _ eq_refl) in morj. - exact (mori ≃ morj). -Defined. -Next Obligation. - destruct A as [i ai], B as [j aj]. - split. - - intros [Hij mor]. - subst; simpl in *. - rewrite (UIUIPI _ _ _ _ eq_refl). - reflexivity. - - intros [Hij mor] [Hij' mor'] [Hij'' mor'']. - subst. - pose (UIPI _ Hij'' eq_refl); - pose (UIPI _ Hij' eq_refl); subst. - simpl. - rewrite (UIUIPI _ _ _ _ eq_refl). - simpl; intros. etransitivity; eauto. - - intros [Hij mor] [Hij' mor']. - subst. - pose (UIPI _ Hij' eq_refl); subst. - simpl. - rewrite (UIUIPI _ _ _ _ eq_refl). - easy. -Qed. -Next Obligation. - destruct A as [i ai], B as [j aj], M as [k ak]. - destruct X as [Hij morij], X0 as [Hjk morjk]. - eapply (existT _ (eq_trans Hij Hjk)). - subst. - exact (morij ∘ morjk). -Defined. -Next Obligation. - rename j into h'. - destruct A as [i ai], B as [j aj], M as [k ak]. - destruct f, g, h, h'. - subst. - pose (UIPI _ x0 eq_refl); - pose (UIPI _ x2 eq_refl); subst. - simpl in *. - rewrite (UIUIPI _ _ _ _ eq_refl) in *. - apply compose_compat; easy. -Qed. -Next Obligation. - destruct A as [i ai], B as [j aj], M as [k ak], N as [l al]. - destruct f as [Hij morij], g as [Hjk morjk], h as [Hkl morkl]. - subst. - simpl. - rewrite (UIUIPI _ _ _ _ eq_refl). - rewrite assoc; easy. -Qed. -Next Obligation. - destruct A as [i ai]. - apply (existT _ eq_refl (id_ ai)). -Defined. -Next Obligation. - destruct A as [i ai], B as [j aj]. - destruct f as [Hij morij]. - subst. - simpl. - rewrite (UIUIPI _ _ _ _ eq_refl). - apply left_unit. -Qed. -Next Obligation. - destruct A as [i ai], B as [j aj]. - destruct f as [Hij morij]. - subst. - simpl. - rewrite (UIUIPI _ _ _ _ eq_refl). - apply right_unit. -Qed. - - - - - -Require Import Logic.Eqdep_dec. - -#[universes(polymorphic=yes), program] Definition BigSumCategory {I : Type} - (eqdec_I : forall i j : I, {i = j} + {i <> j}) - (CI : I -> Type) - (cI : forall i : I, Category (CI i)) : Category (big_sum CI) := {| - morphism := fun somei somej => let (i, ai) := somei in let (j, aj) := somej in _; - (* match eqdec_I i j with - | left H => _ - | _ => False - end; *) - c_equiv := fun somei somej somefi somefj => _; - (* compose := fun is js ks fijs fjks => fun i => fijs i ∘ fjks i; - c_identity := fun is => fun i => id_ (is i); *) -|}. -Next Obligation. - destruct (eqdec_I i j). - - subst. - exact (ai ~> aj). - - exact False. -Defined. -Next Obligation. - destruct somei as [i ai]; destruct somej as [j aj]. - (* destruct somefi as mori, somefj as morj. *) - unfold BigSumCategory_obligation_1 in *. - destruct (eqdec_I i j). - - subst. exact (somefi ≃ somefj). - - destruct somefi. -Defined. -Next Obligation. - destruct A as [i ai], B as [j aj]. - unfold BigSumCategory_obligation_1 in *. - simpl in *. - destruct (eqdec_I i j); [ | easy]. - subst; simpl. - split. - - intros somefi; easy. - - intros ? **. - etransitivity; eauto. - - intros ?. - easy. -Qed. -Next Obligation. - destruct A as [i ai], B as [j aj], M as [k ak]. - unfold BigSumCategory_obligation_1 in *. - simpl in *. - destruct (eqdec_I i j); [ | easy]. - destruct (eqdec_I j k); [ | easy]. - subst. - destruct (eqdec_I k k); [ | easy]. - rewrite (UIP_dec eqdec_I _ eq_refl). - simpl. - - unfold eq_rect_r in *. - simpl in *. - eapply compose; eassumption. -Defined. -Next Obligation. - rename j into j'. - destruct A as [i ai], B as [j aj], M as [k ak]. - unfold BigSumCategory_obligation_1 in *. - (* simpl in H, H0, f, g, h, j'. *) - simpl in *. - destruct (eqdec_I i j); [ | easy]. - destruct (eqdec_I j k); [ | easy]. - subst. - unfold EqdepFacts.internal_eq_rew_r_dep, eq_rect_r in *. - simpl in *. - destruct (eqdec_I k k) as [e | ]; [ | easy]. - - rewrite (UIP_dec eqdec_I _ eq_refl) in *. - simpl. - - - simpl in *. - unfold EqdepFacts.internal_eq_sym_internal, - EqdepFacts.internal_eq_sym_internal0. - apply compose_compat. - - rewrite (UIP_dec eqdec_I _ eq_refl). - - - - - - - - - rewrite assoc; easy. -Qed. -Next Obligation. - rewrite left_unit; easy. -Qed. -Next Obligation. - rewrite right_unit; easy. -Qed. - - - - -Class CartesianMonoidalCategory *) +End CategoryCatProperties. \ No newline at end of file diff --git a/examples/CategoryConstructions.v b/examples/CategoryConstructions.v new file mode 100644 index 0000000..2e89010 --- /dev/null +++ b/examples/CategoryConstructions.v @@ -0,0 +1,340 @@ +From ViCaR Require Import CategoryTypeclass. + +Section Constructions. + +Local Open Scope Cat_scope. +Generalizable Variable C. + +Context {C : Type} {cC : Category C}. + +Set Universe Polymorphism. + + +Class CategoryProduct `{cC : Category C} (A B : C) (AB : C) := { + p_A : AB ~> A; + p_B : AB ~> B; + prod_mor : + forall (Q : C) (fA : Q ~> A) (fB : Q ~> B), Q ~> AB; + prod_mor_prop: + forall (Q : C) (fA : Q ~> A) (fB : Q ~> B), + (prod_mor Q fA fB) ∘ p_A ≃ fA /\ + (prod_mor Q fA fB) ∘ p_B ≃ fB; + prod_mor_unique : + forall (Q : C) (fA : Q ~> A) (fB : Q ~> B) + (fAB' : Q ~> AB), + fA ≃ fAB' ∘ p_A -> fB ≃ fAB' ∘ p_B -> + fAB' ≃ prod_mor Q fA fB +}. + +(* Local Notation "'@' AB" := (AB.(prod_AB)) (at level 8) : Cat_scope. *) + +Context {cCh : CategoryCoherence cC}. + +Lemma prod_mor_unique' {A B AB : C} + (HAB : CategoryProduct A B AB) {Q} (fA : Q ~> A) (fB : Q ~> B) : + forall (fAB fAB' : Q ~> AB), + fAB ∘ p_A ≃ fA -> fAB ∘ p_B ≃ fB -> + fAB' ∘ p_A ≃ fA -> fAB' ∘ p_B ≃ fB -> + fAB ≃ fAB'. +Proof. + intros. + rewrite (prod_mor_unique Q fA fB) by easy. + symmetry; + rewrite (prod_mor_unique Q fA fB) by easy. + easy. +Qed. + +Definition category_product_unique (A B : C) : + forall {AB AB'} (HAB : CategoryProduct A B AB) + (HAB' : CategoryProduct A B AB'), Isomorphism AB AB'. + refine (fun AB AB' HAB HAB' => + {| + forward := HAB'.(prod_mor) AB p_A p_B; + reverse := HAB.(prod_mor) AB' p_A p_B; + |}). + split; eapply (prod_mor_unique' _ p_A p_B); rewrite 1?cCh.(assoc). + 1,5: rewrite 2!(proj1 (prod_mor_prop _ _ _)); easy. + 1,4: rewrite 2!(proj2 (prod_mor_prop _ _ _)); easy. + all: rewrite cCh.(left_unit); easy. +Qed. + +Class CartesianMonoidalCategory (mC : MonoidalCategory cC) := { + tensor_cartesian : forall (A B : C), + CategoryProduct A B (A × B); +}. + + +Class CategoryBigProd `{cC : Category C} {J : Type} + (obj : J -> C) (prod_J : C) := { + p_i : forall i, prod_J ~> obj i; + big_prod_mor : + forall (Q : C) (fQ_ : forall i, Q ~> obj i), Q ~> prod_J; + big_prod_mor_prop: + forall (Q : C) (fQ_ : forall i, Q ~> obj i), + forall i, + (big_prod_mor Q fQ_) ∘ p_i i ≃ fQ_ i; + big_prod_mor_unique : + forall (Q : C) (fQ_ : forall i, Q ~> obj i) + (fAB' : Q ~> prod_J), + (forall i, fQ_ i ≃ fAB' ∘ p_i i) -> + fAB' ≃ big_prod_mor Q fQ_ +}. + +Lemma big_prod_mor_unique' {J} {obj : J -> C} {pJ : C} + (HI : CategoryBigProd obj pJ) {Q} (fQ_ : forall i, Q ~> obj i) : + forall (fAB fAB' : Q ~> pJ), + (forall i, fAB ∘ p_i i ≃ fQ_ i) -> + (forall i, fAB' ∘ p_i i ≃ fQ_ i) -> + fAB ≃ fAB'. +Proof. + intros. + rewrite (big_prod_mor_unique Q fQ_) by easy. + symmetry; + rewrite (big_prod_mor_unique Q fQ_) by easy. + easy. +Qed. + +Obligation Tactic := idtac. + +Program Definition category_big_prod_unique {J} {obj : J->C} : + forall {pJ pJ'} (HI : CategoryBigProd obj pJ) (HI' : CategoryBigProd obj pJ'), + Isomorphism pJ pJ' := + fun pJ pJ' HI HI' => + {| + forward := HI'.(big_prod_mor) pJ p_i; + reverse := HI.(big_prod_mor) pJ' p_i; + |}. +Next Obligation. + split; eapply (big_prod_mor_unique' _ p_i); rewrite 1?assoc. + 1,3: intros i; rewrite cCh.(assoc), 2!(big_prod_mor_prop _ _); easy. + all: intros; rewrite cCh.(left_unit); easy. +Qed. + +(* Require Import FunctionalExtensionality. *) + + +Class CategoryCoproduct (A B : C) (A_B : C) := { + coprod_obj := A_B; + iota_A : A ~> A_B; + iota_B : B ~> A_B; + coprod_mor : + forall (Q : C) (fA : A ~> Q) (fB : B ~> Q), A_B ~> Q; + coprod_mor_prop: + forall (Q : C) (fA : A ~> Q) (fB : B ~> Q), + iota_A ∘ (coprod_mor Q fA fB) ≃ fA /\ + iota_B ∘ (coprod_mor Q fA fB) ≃ fB; + coprod_mor_unique : + forall (Q : C) (fA : A ~> Q) (fB : B ~> Q) + (fAB' : A_B ~> Q), + fA ≃ iota_A ∘ fAB' -> fB ≃ iota_B ∘ fAB' -> + fAB' ≃ coprod_mor Q fA fB +}. + +Lemma coprod_mor_unique' {A B A_B : C} + (HAB : CategoryCoproduct A B A_B) {Q} (fA : A ~> Q) (fB : B ~> Q) : + forall (fAB fAB' : A_B ~> Q), + iota_A ∘ fAB ≃ fA -> iota_B ∘ fAB ≃ fB -> + iota_A ∘ fAB' ≃ fA -> iota_B ∘ fAB' ≃ fB -> + fAB ≃ fAB'. +Proof. + intros. + rewrite (coprod_mor_unique Q fA fB) by easy. + symmetry; + rewrite (coprod_mor_unique Q fA fB) by easy. + easy. +Qed. + +Program Definition category_coproduct_unique (A B : C) : + forall {A_B A_B'} (HA_B : CategoryCoproduct A B A_B) + (HA_B' : CategoryCoproduct A B A_B'), Isomorphism A_B A_B' := + fun A_B A_B' HA_B HA_B' => + {| + forward := HA_B.(coprod_mor) A_B' iota_A iota_B; + reverse := HA_B'.(coprod_mor) A_B iota_A iota_B; + |}. +Next Obligation. + split; eapply (coprod_mor_unique' _ iota_A iota_B); rewrite <- 1?cCh.(assoc). + 1,5: rewrite 2!(proj1 (coprod_mor_prop _ _ _)); easy. + 1,4: rewrite 2!(proj2 (coprod_mor_prop _ _ _)); easy. + all: rewrite cCh.(right_unit); easy. +Qed. + +Class CategoryBigCoprod {J : Type} + (obj : J -> C) (coprod_J : C) := { + big_coprod_obj := coprod_J; + iota_i : forall i, obj i ~> coprod_J; + big_coprod_mor : + forall (Q : C) (fQ_ : forall i, obj i ~> Q), coprod_J ~> Q; + big_coprod_mor_prop: + forall (Q : C) (fQ_ : forall i, obj i ~> Q), + forall i, + iota_i i ∘ (big_coprod_mor Q fQ_) ≃ fQ_ i; + big_coprod_mor_unique : + forall (Q : C) (fQ_ : forall i, obj i ~> Q) + (fAB' : coprod_J ~> Q), + (forall i, fQ_ i ≃ iota_i i ∘ fAB') -> + fAB' ≃ big_coprod_mor Q fQ_ +}. + + +Reserved Notation "A ∏ B" (at level 34). +Reserved Notation "f @∏ g" (at level 34). + +Class Category_with_Products (cC : Category C) := { + cat_prod : C -> C -> C + where "A ∏ B" := (cat_prod A B); + cat_prod_product : forall {A B}, CategoryProduct A B (A ∏ B); + cat_mor_prod : + forall {A A' B B'} (fA : A ~> B) (fB : A' ~> B'), + A ∏ A' ~> B ∏ B' := fun A A' B B' fA fB => + cat_prod_product.(prod_mor) + (A ∏ A') (p_A ∘ fA) (p_B ∘ fB) + where "f @∏ g" := (cat_mor_prod f g); +}. + +Notation "A ∏ B" := (cat_prod A B) (at level 34). +Notation "f @∏ g" := (cat_mor_prod f g) (at level 34). + +#[program, export] Instance Category_with_Products_of_CartesianMonoidalCategory + {mC : MonoidalCategory cC} + (ccc : @CartesianMonoidalCategory mC) : Category_with_Products cC := {| + cat_prod := mC.(obj_tensor); + cat_prod_product := tensor_cartesian; +|}. + +Coercion Category_with_Products_of_CartesianMonoidalCategory : + CartesianMonoidalCategory >-> Category_with_Products. + + +Reserved Notation "'λ' g" (at level 20). +Class CategoryExponential `{cC : Category C} {pC : Category_with_Products cC} + (A B A_B : C) := { + exp_obj := A_B; + eval_mor : A_B ∏ B ~> A; + eval_mor_transpose : forall {M : C} (f : M ∏ B ~> A), + M ~> A_B + where "'λ' g" := (eval_mor_transpose g); + eval_mor_transpose_prop : forall (M : C) (f : M ∏ B ~> A), + f ≃ (λ f @∏ id_ B) ∘ eval_mor; + eval_mor_transpose_unique : forall (M : C) (f : M ∏ B ~> A) + (lf' : M ~> A_B), f ≃ (lf' @∏ id_ B) ∘ eval_mor -> λ f ≃ lf'; +}. + +Class Cone {C D} {cC : Category C} {cD : Category D} + (F : Functor cD cC) := { + cone_obj : C; + cone_mor : forall (d : D), cone_obj ~> F d; + cone_mor_prop : forall {d d' : D} (f : d ~> d'), + cone_mor d ∘ F @ f ≃ cone_mor d'; +}. + +Coercion cone_mor : Cone >-> Funclass. + +Reserved Notation "'lim' F" (at level 30). + +Class Limit {C D} {cC : Category C} {cD : Category D} + (F : Functor cD cC) := { + limit_cone : Cone F; + limit_cone_factor : + forall (N : Cone F), N.(cone_obj) ~> limit_cone.(cone_obj); + limit_cone_factor_prop : forall (N : Cone F) (d : D), + limit_cone_factor N ∘ limit_cone d ≃ N d; + limit_cone_factor_unique : forall (N : Cone F) + (f' : N.(cone_obj) ~> limit_cone.(cone_obj)), + (forall (d : D), f' ∘ limit_cone d ≃ N d) -> f' ≃ limit_cone_factor N; +}. + +Lemma limit_cone_factor_unique' {D} {cD : Category D} + {F : Functor cD cC} (L : Limit F) : forall (N : Cone F) + (f f' : N.(cone_obj) ~> L.(limit_cone).(cone_obj)), + (forall (d : D), f ∘ limit_cone d ≃ N d) -> + (forall (d : D), f' ∘ limit_cone d ≃ N d) -> + f ≃ f'. +Proof. + intros. + rewrite limit_cone_factor_unique by easy. + symmetry. + rewrite limit_cone_factor_unique by easy. + easy. +Qed. + +Coercion limit_cone : Limit >-> Cone. + +Notation "'lim' F" := (limit_cone (F:=F)).(cone_obj) (at level 30). + +#[program] Definition limit_unique {D} {cD : Category D} + {F : Functor cD cC} (L L' : Limit F) : L.(cone_obj) <~> L'.(cone_obj) := {| + forward := L'.(limit_cone_factor) L; + reverse := L.(limit_cone_factor) L' + |}. +Next Obligation. + split; (apply limit_cone_factor_unique'; + [| intros; rewrite cCh.(left_unit); easy]); + intros; rewrite cCh.(assoc), 2!limit_cone_factor_prop; easy. +Qed. + + +Class Cocone {C D} {cC : Category C} {cD : Category D} + (F : Functor cD cC) := { + cocone_obj : C; + cocone_mor : forall (d : D), F d ~> cocone_obj; + cocone_mor_prop : forall {d d' : D} (f : d ~> d'), + F @ f ∘ cocone_mor d' ≃ cocone_mor d; +}. + +Coercion cocone_mor : Cocone >-> Funclass. + +Reserved Notation "'colim' F" (at level 30). + +Class Colimit {C D} {cC : Category C} {cD : Category D} + (F : Functor cD cC) := { + colimit_cocone : Cocone F; + colimit_cocone_factor : forall (N : Cocone F), + colimit_cocone.(cocone_obj) ~> N.(cocone_obj); + colimit_cocone_factor_prop : forall (N : Cocone F) (d : D), + colimit_cocone d ∘ colimit_cocone_factor N ≃ N d; + colimit_cocone_factor_unique : forall (N : Cocone F) + (f' : colimit_cocone.(cocone_obj) ~> N.(cocone_obj)), + (forall (d : D), colimit_cocone d ∘ f' ≃ N d) -> + f' ≃ colimit_cocone_factor N; +}. + +Lemma colimit_cocone_factor_unique' {D} {cD : Category D} + {F : Functor cD cC} (L : Colimit F) : forall (N : Cocone F) + (f f' : L.(colimit_cocone).(cocone_obj) ~> N.(cocone_obj)), + (forall (d : D), colimit_cocone d ∘ f ≃ N d) -> + (forall (d : D), colimit_cocone d ∘ f' ≃ N d) -> + f ≃ f'. +Proof. + intros. + rewrite colimit_cocone_factor_unique by easy. + symmetry. + rewrite colimit_cocone_factor_unique by easy. + easy. +Qed. + +Coercion colimit_cocone : Colimit >-> Cocone. + +Notation "'colim' F" := (colimit_cocone (F:=F)).(cocone_obj) (at level 30). + +#[program] Definition colimit_unique {D} {cD : Category D} + {F : Functor cD cC} (L L' : Colimit F) : L.(cocone_obj) <~> L'.(cocone_obj) := {| + forward := L.(colimit_cocone_factor) L'; + reverse := L'.(colimit_cocone_factor) L + |}. +Next Obligation. + split; (apply colimit_cocone_factor_unique'; + [| intros; rewrite cCh.(right_unit); easy]); + intros; rewrite <- cCh.(assoc), 2!colimit_cocone_factor_prop; easy. +Qed. + + +Class Category_with_Small_Limits {C} (cC : Category C) := { + take_set_limit : forall {D : Set} (cD : Category D) (F : Functor cD cC), Limit F; +}. + +Class Category_with_Limits {C} (cC : Category C) := { + take_limit : forall {D : Type} (cD : Category D) (F : Functor cD cC), Limit F; +}. + +End Constructions. \ No newline at end of file diff --git a/examples/DirectSumMatrixExample.v b/examples/DirectSumMatrixExample.v deleted file mode 100644 index 9d807d1..0000000 --- a/examples/DirectSumMatrixExample.v +++ /dev/null @@ -1,1147 +0,0 @@ -Require Import MatrixExampleBase. -Require Import MatrixPermBase. -Require Import ExamplesAutomation. - -#[export] Instance MxCategory : Category nat := { - morphism := Matrix; - - c_equiv := @mat_equiv; (* fun m n => @eq (Matrix m n); *) - c_equiv_rel := @mat_equiv_equivalence; - - compose := @Mmult; - compose_compat := fun n m o f g Hfg h i Hhi => - @Mmult_simplify_mat_equiv n m o f g h i Hfg Hhi; - assoc := @Mmult_assoc_mat_equiv; - - c_identity n := I n; - left_unit := Mmult_1_l_mat_eq; - right_unit := Mmult_1_r_mat_eq; -}. - -Lemma direct_sum'_id_mat_equiv : forall n m, - direct_sum' (I n) (I m) ≡ I (n + m). -Proof. - intros n m. - intros i j Hi Hj. - unfold direct_sum', I. - bdestructΩ'simp. -Qed. - -Lemma direct_sum'_id : forall n m, - direct_sum' (I n) (I m) = I (n + m). -Proof. - intros n m. - rewrite <- mat_equiv_eq_iff by auto with wf_db. - apply direct_sum'_id_mat_equiv. -Qed. - -Lemma big_sum_split : forall {G : Type} {H : Monoid G} (n k : nat) (f : nat -> G), - big_sum f (n + k) = - (big_sum f n + big_sum (fun x : nat => f (n + x)%nat) k)%G. -Proof. - intros G H n. - induction k; intros f. - - simpl; rewrite Nat.add_0_r, Gplus_0_r; easy. - - rewrite Nat.add_succ_r, <- big_sum_extend_r, IHk, <- Gplus_assoc. - reflexivity. -Qed. - -Lemma direct_sum'_Mmult : forall {n m o p q r} - (A : Matrix n m) (B : Matrix m o) (C : Matrix p q) (D : Matrix q r), - direct_sum' (A × B) (C × D) ≡ direct_sum' A C × direct_sum' B D. -Proof. - intros n m o p q r A B C D. - intros i j Hi Hj. - symmetry. - unfold direct_sum', Mmult. - bdestruct (i {| - forward := (I (n + m + o) : Matrix (n + m + o) (n + (m +o))); - reverse := (I (n + m + o) : Matrix (n + (m +o)) (n + m + o)); - isomorphism_inverse := ltac:(simpl; rewrite Nat.add_assoc, Mmult_1_r_mat_eq; easy); - (* id_A := ltac:(simpl; rewrite Nat.add_assoc, Mmult_1_r_mat_eq; easy); - id_B := ltac:(simpl; rewrite Nat.add_assoc, Mmult_1_r_mat_eq; easy); *) - |}; - - left_unitor := fun n => {| - forward := (I n : Matrix (0 + n) n); - reverse := (I n : Matrix n (0 + n)); - isomorphism_inverse := ltac:(split; simpl; rewrite Mmult_1_r_mat_eq; easy); - (* id_A := ltac:(simpl; rewrite Mmult_1_r_mat_eq; easy); - id_B := ltac:(simpl; rewrite Mmult_1_r_mat_eq; easy); *) - |}; - - right_unitor := fun n => {| - forward := (I n : Matrix (n + 0) n); - reverse := (I n : Matrix n (n + 0)); - isomorphism_inverse := ltac:(split; rewrite Nat.add_0_r, Mmult_1_r_mat_eq; easy); - (* id_A := ltac:(rewrite Nat.add_0_r, Mmult_1_r_mat_eq; easy); - id_B := ltac:(rewrite Nat.add_0_r, Mmult_1_r_mat_eq; easy); *) - |}; - - associator_cohere := ltac:(intros; simpl in *; - rewrite 2!Nat.add_assoc, Mmult_1_l_mat_eq, - Mmult_1_r_mat_eq, <-2!Nat.add_assoc; rewrite (direct_sum'_assoc f g h); easy); - left_unitor_cohere := ltac:(intros; simpl; - rewrite direct_sum'_0_l, Mmult_1_l_mat_eq, Mmult_1_r_mat_eq; easy); - right_unitor_cohere := ltac:(intros; simpl; rewrite direct_sum'_0_r; - rewrite 2!Nat.add_0_r, Mmult_1_r_mat_eq, Mmult_1_l_mat_eq; easy); - - pentagon := ltac:(intros; simpl in *; restore_dims; - rewrite 4!Nat.add_assoc, Mmult_1_r_mat_eq, - 2!direct_sum'_id, 2!Mmult_1_l_mat_eq, 2!Nat.add_assoc; easy); - triangle := ltac:(intros; simpl in *; - rewrite Nat.add_0_r, direct_sum'_id, Mmult_1_r_mat_eq; easy); -}. - - - - -Notation mx_braiding := (fun n m => (perm_mat (n+m) (rotr (n+m) n) : Matrix (n+m) (m+n))%nat). - -Lemma mx_braiding_compose_inv : forall n m, - (mx_braiding n m) × (mx_braiding m n) ≡ I (n + m). -Proof. - intros n m. - simpl. - rewrite (Nat.add_comm m n). - rewrite perm_mat_Mmult by auto with perm_db. - cleanup_perm. - rewrite perm_mat_I by easy. - easy. -Qed. - - - - - -Lemma mx_braiding_braids_eq : forall n m o p (A : Matrix n m) (B : Matrix o p), - (direct_sum' A B × perm_mat (m + p) (rotr (m + p) m) - = perm_mat (n + o) (rotr (n + o) n) × direct_sum' B A). -Proof. - intros n m o p A B. - apply equal_on_basis_vectors_implies_equal; - [|rewrite Nat.add_comm, (Nat.add_comm m p) |]; auto with wf_db. - intros k Hk. - rewrite <- mat_equiv_eq_iff; - [| | apply WF_mult; [apply WF_mult|]]; auto with wf_db; - [|rewrite (Nat.add_comm m p), (Nat.add_comm n o); auto with wf_db]. - rewrite Mmult_assoc. - rewrite perm_mat_permutes_basis_vectors_r, basis_vector_eq_e_i by easy. - rewrite <- (matrix_by_basis _ _ Hk). - rewrite <- matrix_by_basis. - 2: { (* TODO: replace with 'apply permutation_is_bounded; auto with perm_db' *) - assert (Hp: (permutation (m+p) (rotr (m + p) m))%nat) by auto with perm_db. - destruct Hp as [finv Hfinv]. - destruct (Hfinv k Hk); easy. - } - intros x y Hx Hy; replace y with O by lia; clear y Hy. - unfold get_vec. - rewrite Nat.eqb_refl. - rewrite perm_mat_permutes_matrix_r by auto with perm_db. - rewrite perm_inv_of_rotr by easy. - rewrite rotl_eq_rotr_sub. - bdestruct (o =? 0); - [subst o; - rewrite Nat.add_0_r, Nat.mod_same, Nat.sub_0_r, rotr_n by lia| - rewrite Nat.mod_small by lia; - replace (n + o - n)%nat with o by lia]; - unfold direct_sum', rotr, rotl; simpl; - rewrite 1?Nat.add_0_r; - - [ replace_bool_lia (x Build_Isomorphism nat MxCategory (n+m)%nat (m+n)%nat - (mx_braiding n m) (mx_braiding m n) - (conj (mx_braiding_compose_inv n m) (mx_braiding_compose_inv m n)). - -#[export] Instance MxBraidingBiIsomorphism : - NaturalBiIsomorphism MxDirectSumBiFunctor (CommuteBifunctor MxDirectSumBiFunctor) := {| - component_biiso := MxBraidingIsomorphism; - component_biiso_natural := ltac:(intros; simpl in *; - restore_dims; rewrite mx_braiding_braids_eq; easy); -|}. - -Lemma if_mult_dist_r (b : bool) (z : C) : - (if b then C1 else C0) * z = - if b then z else C0. -Proof. - destruct b; lca. -Qed. - -Lemma if_mult_dist_l (b : bool) (z : C) : - z * (if b then C1 else C0) = - if b then z else C0. -Proof. - destruct b; lca. -Qed. - -Lemma if_mult_dist_r_gen (b : bool) (z x y : C) : - (if b then x else y) * z = - if b then x*z else y*z. -Proof. - destruct b; lca. -Qed. - -Lemma if_mult_dist_l_gen (b : bool) (z x y : C) : - z * (if b then x else y) = - if b then z*x else z*y. -Proof. - destruct b; lca. -Qed. - -Lemma if_mult_and (b c : bool) : - (if b then C1 else C0) * (if c then C1 else C0) = - if (b && c) then C1 else C0. -Proof. - destruct b; destruct c; lca. -Qed. - -Lemma if_if_if_combine {A : Type} : forall (x y : A) (b c d:bool), - (if b then if c then x else y - else if d then x else y) = - if (b&&c)||((¬b) && d) then x else y. -Proof. - intros. - bdestructΩ'. -Qed. - -(* Definition unshift {A} (f : nat -> A) (k : nat) (x : A) : nat -> A := - fun i => if i if (O if (i permutation m g -> - direct_sum' (perm_mat n f) (perm_mat m g) ≡ perm_mat (n+m) (stack_perms n m f g). -Proof. - intros n m f g Hf Hg. - apply mat_equiv_of_equiv_on_ei. - intros k Hk. - rewrite perm_mat_permutes_ei_r, direct_sum'_mul_vec_r by lia. - unfold make_WF, Mplus. - intros i j Hi Hj. - replace j with O by lia; clear j Hj. - simpl_bools. - unfold make_WF, perm_mat, stack_perms, unshift_mx, shift, e_i, Mmult. - replace_bool_lia (k apply big_sum_unique - | |- _ = C0 => rewrite big_sum_0_bounded; [easy|]; intros; bdestructΩ'simp - end. - 4: { - rewrite Nat.add_sub in *; - lia. - } - - exists k; repeat split; try lia; intros; bdestructΩ'simp. - - destruct Hf as [? Hf]. - specialize (Hf k). - lia. - - exists (k - n)%nat; repeat split; intros; bdestructΩ'simp. -Qed. - -Lemma perm_mat_idn : forall n, - perm_mat n idn = I n. -Proof. - intros n. - apply perm_mat_I; easy. -Qed. - -Lemma direct_sum'_stack_perm_I_r : forall n m f, - permutation n f -> - direct_sum' (perm_mat n f) (I m) ≡ perm_mat (n+m) (stack_perms n m f idn). -Proof. - intros n m f Hf. - rewrite <- perm_mat_idn. - rewrite direct_sum'_stack_perms; (auto with perm_db). - easy. -Qed. - -Lemma direct_sum'_stack_perm_I_l : forall n m f, - permutation m f -> - direct_sum' (I n) (perm_mat m f) ≡ perm_mat (n+m) (stack_perms n m idn f). -Proof. - intros n m f Hf. - rewrite <- perm_mat_idn. - rewrite direct_sum'_stack_perms; (auto with perm_db). - easy. -Qed. - -Lemma mx_braiding_hexagon1: forall n m o (* M B A *), - ((direct_sum' (mx_braiding n m) (I o) × I (m + n + o) - × direct_sum' (I m) (mx_braiding n o))) - ≡ (I (n + m + o) × (mx_braiding n (m + o)%nat) × I (m + o + n)). -Proof. - intros n m o. - (* replace (n + (o+m))%nat with (m+o+n)%nat by lia. *) - rewrite 2!Mmult_1_r_mat_eq, Mmult_1_l_mat_eq. - rewrite (Nat.add_comm m n), (Nat.add_comm o n). - rewrite direct_sum'_stack_perm_I_r by auto with perm_db. - replace (n+m+o)%nat with (m+(n+o))%nat by lia. - replace (m+o+n)%nat with (m+(n+o))%nat by lia. - rewrite direct_sum'_stack_perm_I_l by auto with perm_db. - rewrite perm_mat_Mmult by auto with perm_db. - apply perm_mat_equiv_of_perm_eq'; [lia|]. - intros k Hk. - unfold stack_perms, rotr, Basics.compose; - replace_bool_lia (k if (k =? j) then _ else _); shelve. - bdestruct (k =? j). - - replace j with k by easy. - rewrite Nat.eqb_refl. - unshelve (instantiate (1:=_)). - refine (if (k if (k subst x - end. - -Ltac evarify_1func_once f := - match goal with - |- context[@f ?a ?A] => - not_evar' a; let x:= fresh in evar (x : nat); - replace (@f a A) with (@f x A) by - ( replace (@f a) with (@f x) by (replace a with x by shelve; reflexivity); - reflexivity) - end. - -Ltac evarify_1func f := repeat (evarify_1func_once f). -Ltac evarify_1func' f := repeat (evarify_1func_once f); subst_evars. - -Ltac evarify_2func_once f := - match goal with - |- context[@f ?a ?b ?A] => - not_evar' a; not_evar' b; - let x:= fresh in let y := fresh in - evar (x : nat); evar (y : nat); - replace (@f a b A) with (@f x y A) by - ( replace (@f a) with (@f x) by (replace a with x by shelve; reflexivity); - replace (@f x b) with (@f x y) by (replace b with y by shelve; reflexivity); - reflexivity) - end. - -Ltac evarify_2func f := repeat (evarify_2func_once f). -Ltac evarify_2func' f := repeat (evarify_2func_once f); subst_evars. - -Ltac evarify_3func_once f := - match goal with - |- context[@f ?a ?b ?c ?A] => - not_evar' a; not_evar' b; not_evar' c; - let x:= fresh in let y := fresh in let z := fresh in - evar (x : nat); evar (y : nat); evar (z : nat); - replace (@f a b c A) with (@f x y z A) by - ( replace (@f a) with (@f x) by (replace a with x by shelve; reflexivity); - replace (@f x b) with (@f x y) by (replace b with y by shelve; reflexivity); - replace (@f x y c) with (@f x y z) by (replace c with z by shelve; reflexivity); - reflexivity) - end. - -Ltac evarify_3func f := repeat (evarify_3func_once f). -Ltac evarify_3func' f := repeat (evarify_3func_once f); subst_evars. - -Ltac evarify_4func_once f := - match goal with - |- context[@f ?a ?b ?c ?d ?A] => - not_evar' a; not_evar' b; not_evar' c; not_evar' d; - let x:= fresh in let y := fresh in let z := fresh in let zz := fresh in - evar (x : nat); evar (y : nat); evar (z : nat); evar (zz : nat); - replace (@f a b c d A) with (@f x y z zz A) by - ( replace (@f a) with (@f x) by (replace a with x by shelve; reflexivity); - replace (@f x b) with (@f x y) by (replace b with y by shelve; reflexivity); - replace (@f x y c) with (@f x y z) by (replace c with z by shelve; reflexivity); - replace (@f x y z c) with (@f x y z zz) by (replace d with zz by shelve; reflexivity); - reflexivity) - end. - -Ltac evarify_4func f := repeat (evarify_4func_once f). -Ltac evarify_4func' f := repeat (evarify_4func_once f); subst_evars. - -(* Ltac lem_conclusion lem := - match type of lem with - | forall _ _ _ _ _ _ _ _ _ _ _ _ _ _, ?P => P - | forall _ _ _ _ _ _ _ _ _ _ _ _ _, ?P => P - | forall _ _ _ _ _ _ _ _ _ _ _ _, ?P => P - | forall _ _ _ _ _ _ _ _ _ _ _, ?P => P - | forall _ _ _ _ _ _ _ _ _ _, ?P => P - | forall _ _ _ _ _ _ _ _ _, ?P => P - | forall _ _ _ _ _ _ _ _, ?P => P - | forall _ _ _ _ _ _ _, ?P => P - | forall _ _ _ _ _ _, ?P => P - | forall _ _ _ _ _, ?P => P - | forall _ _ _ _, ?P => P - | forall _ _ _, ?P => P - | forall _ _, ?P => P - | forall _, ?P => P - | ?P => P - end. *) - -(* Ltac lem_conclusion lem := - lazymatch type of lem with - | forall (a : _) (b : _) (c : _) (d : _) (e : _) (f : _) (g : _) (h : _) (i : _) (j : _) (k : _) (l : _) (m : _) (n : _), ?P => open_constr:(P) - | forall (a : _) (b : _) (c : _) (d : _) (e : _) (f : _) (g : _) (h : _) (i : _) (j : _) (k : _) (l : _) (m : _), ?P => open_constr:(P) - | forall (a : _) (b : _) (c : _) (d : _) (e : _) (f : _) (g : _) (h : _) (i : _) (j : _) (k : _) (l : _), ?P => open_constr:(P) - | forall (a : _) (b : _) (c : _) (d : _) (e : _) (f : _) (g : _) (h : _) (i : _) (j : _) (k : _), ?P => open_constr:(P) - | forall (a : _) (b : _) (c : _) (d : _) (e : _) (f : _) (g : _) (h : _) (i : _) (j : _), ?P => open_constr:(P) - | forall (a : _) (b : _) (c : _) (d : _) (e : _) (f : _) (g : _) (h : _) (i : _), ?P => open_constr:(P) - | forall (a : _) (b : _) (c : _) (d : _) (e : _) (f : _) (g : _) (h : _), ?P => open_constr:(P) - | forall (a : _) (b : _) (c : _) (d : _) (e : _) (f : _) (g : _), ?P => open_constr:(P) - | forall (a : _) (b : _) (c : _) (d : _) (e : _) (f : _), ?P => open_constr:(P) - | forall (a : _) (b : _) (c : _) (d : _) (e : _), ?P => open_constr:(P) - | forall (a : _) (b : _) (c : _) (d : _), ?P => open_constr:(P) - | forall (a : _) (b : _) (c : _), ?P => open_constr:(P) - | forall (a : _) (b : _), ?P => open_constr:(P) - | forall (a : _), ?P => open_constr:(P) - | ?P => open_constr:(P) - end. - -Ltac eq_lem_lhs lem := - let t:= lem_conclusion lem in - match open_constr:(t) with - | ?x = ?y => uconstr:(x) - end. - -Ltac eq_lem_lhs_head' lem := - let t:= eq_lem_lhs lem in - let s := type of t in idtac s; - match type of t with - | ?f ?x => idtac f - end. - -Ltac eq_lem_lhs_head lem := - let hd ty := - match type of ty with - | ?f _ _ _ _ _ _ _ _ _ _ _ => f - | ?f _ _ _ _ _ _ _ _ _ _ => f - | ?f _ _ _ _ _ _ _ _ _ => f - | ?f _ _ _ _ _ _ _ _ => f - | ?f _ _ _ _ _ _ _ => f - | ?f _ _ _ _ _ _ => f - | ?f _ _ _ _ _ => f - | ?f _ _ _ _ => f - | ?f _ _ _ => f - | ?f _ _ => f - | ?f _ => f - end in - lazymatch type of lem with - | forall (a : _) (b : _) (c : _) (d : _) (e : _) (f : _) (g : _) (h : _) (i : _) (j : _) (k : _) (l : _) (m : _) (n : _), ?P => hd P - | forall (a : _) (b : _) (c : _) (d : _) (e : _) (f : _) (g : _) (h : _) (i : _) (j : _) (k : _) (l : _) (m : _), ?P => hd P - | forall (a : _) (b : _) (c : _) (d : _) (e : _) (f : _) (g : _) (h : _) (i : _) (j : _) (k : _) (l : _), ?P => hd P - | forall (a : _) (b : _) (c : _) (d : _) (e : _) (f : _) (g : _) (h : _) (i : _) (j : _) (k : _), ?P => hd P - | forall (a : _) (b : _) (c : _) (d : _) (e : _) (f : _) (g : _) (h : _) (i : _) (j : _), ?P => hd P - | forall (a : _) (b : _) (c : _) (d : _) (e : _) (f : _) (g : _) (h : _) (i : _), ?P => hd P - | forall (a : _) (b : _) (c : _) (d : _) (e : _) (f : _) (g : _) (h : _), ?P => hd P - | forall (a : _) (b : _) (c : _) (d : _) (e : _) (f : _) (g : _), ?P => hd P - | forall (a : _) (b : _) (c : _) (d : _) (e : _) (f : _), ?P => hd P - | forall (a : _) (b : _) (c : _) (d : _) (e : _), ?P => hd P - | forall (a : _) (b : _) (c : _) (d : _), ?P => hd P - | forall (a : _) (b : _) (c : _), ?P => hd P - | forall (a : _) (b : _), ?P => hd P - | forall (a : _), ?P => hd P - | ?P => hd P - end. *) - -(* NOTE: modified from https://coq.discourse.group/t/ltac-that-goes-under-binders-to-return-a-non-constr-value/257/5 -Ltac head_application th := - match th with - | ?f _ _ _ _ _ _ _ _ _ _ _ => f - | ?f _ _ _ _ _ _ _ _ _ _ => f - | ?f _ _ _ _ _ _ _ _ _ => f - | ?f _ _ _ _ _ _ _ _ => f - | ?f _ _ _ _ _ _ _ => f - | ?f _ _ _ _ _ _ => f - | ?f _ _ _ _ _ => f - | ?f _ _ _ _ => f - | ?f _ _ _ => f - | ?f _ _ => f - | ?f _ => f - end. - -(* go under binder and rebuild a term with a good name inside, - catchable by a match context. *) -Ltac get_head_name' lem lemty := - lazymatch lemty with - | forall z:?A , ?B => - let name := - constr:( - forall z:A, - ltac:( - let h' := constr:(lem z) in - let th' := type of h' in - get_head_name' h' th')) in - (* remove the let *) - eval lazy zeta in name - | _ => - (* let nme := fallback_rename_hyp h th in *) - head_application lemty - end. - -Ltac get_head_name'' lemty := - lazymatch lemty with - | forall z:?A , ?B => - let t := type of B in - get_head_name' t - (* let name := get_head_name' t in - (* remove the let *) - eval lazy zeta in name *) - | _ => - (* let nme := fallback_rename_hyp h th in *) - head_application lemty - end. *) - -(*endcomment*) - -Definition DUMMY {T : Type} := fun (x : T) => x. -(* Extract the head of an application *) -Ltac head_application th := - let ty := type of th in - match th with - | ?f _ _ _ _ _ _ _ _ _ _ _ => constr:(Some f) - | ?f _ _ _ _ _ _ _ _ _ _ => constr:(Some f) - | ?f _ _ _ _ _ _ _ _ _ => constr:(Some f) - | ?f _ _ _ _ _ _ _ _ => constr:(Some f) - | ?f _ _ _ _ _ _ _ => constr:(Some f) - | ?f _ _ _ _ _ _ => constr:(Some f) - | ?f _ _ _ _ _ => constr:(Some f) - | ?f _ _ _ _ => constr:(Some f) - | ?f _ _ _ => constr:(Some f) - | ?f _ _ => constr:(Some f) - | ?f _ => constr:(Some f) - | _ => constr:(None : option ty) - end. - -(* go under binder and rebuild a term with dummy labels, - catchable by a match context. *) -Ltac build_dummy_quantified h th := - lazymatch th with - | forall z:?A , ?B => - let tA := type of A in let A' := build_dummy_quantified A tA in - let x := - constr:( - forall z:A', - let h' := (h z) in - ltac:( - let th' := type of h' in - let res := build_dummy_quantified h' th' in - exact res)) in - (* remove the let *) - eval lazy zeta in x - | _ => - (* let nme := fallback_rename_hyp h th in *) - let hd := head_application th in - match hd with - | Some ?f => constr:(DUMMY f) - | None => constr:(h) - end - (* let frshnme := fresh hd in - (* Build something catchable with mathc context *) - constr:(forall frshnme:Prop, DUMMY frshnme) *) - end. - -Lemma xxx: - (forall x y z:nat, x y< z -> x < z)%nat -> True. -Proof. - intros h. - let x := build_dummy_quantified 4 nat in idtac x. - let typ := type of h in - let x := build_dummy_quantified h typ in - idtac x. (* Se the type built by the tactic *) - (* use it to rename h: *) - match x with - | context [forall (Xnam : _), DUMMY Xnam] => let nme := fresh "h_" Xnam in rename h into nme - end. -Qed. - -Lemma test : forall (A B M : nat), -(and - (@mat_equiv (Nat.add (Nat.add A B) M) (Nat.add (Nat.add A B) M) - (@Mmult (Nat.add (Nat.add A B) M) (Nat.add A (Nat.add B M)) - (Nat.add (Nat.add A B) M) (I (Init.Nat.add (Init.Nat.add A B) M)) - (@adjoint (Nat.add (Nat.add A B) M) (Nat.add A (Nat.add B M)) - (I (Init.Nat.add (Init.Nat.add A B) M)))) - (I (Nat.add (Nat.add A B) M))) - (@mat_equiv (Nat.add A (Nat.add B M)) (Nat.add A (Nat.add B M)) - (@Mmult (Nat.add A (Nat.add B M)) (Nat.add (Nat.add A B) M) - (Nat.add A (Nat.add B M)) - (@adjoint (Nat.add (Nat.add A B) M) (Nat.add A (Nat.add B M)) - (I (Init.Nat.add (Init.Nat.add A B) M))) - (I (Init.Nat.add (Init.Nat.add A B) M))) - (I (Nat.add A (Nat.add B M))))). -intros A B M. - -Ltac get_head term := - match term with - | ?f _ _ _ _ _ _ _ _ _=> constr:(f) - | ?f _ _ _ _ _ _ _ _ => constr:(f) - | ?f _ _ _ _ _ _ _ => constr:(f) - | ?f _ _ _ _ _ _ => constr:(f) - | ?f _ _ _ _ _ => constr:(f) - | ?f _ _ _ _ => constr:(f) - | ?f _ _ _ => constr:(f) - | ?f _ _ => constr:(f) - | ?f _ => constr:(f) - | ?f => constr:(f) - end. - - (*stdpp: *) -Ltac mk_evar T := - let T := constr:(T : Type) in - let e := fresh in - let _ := match goal with _ => evar (e:T) end in - let e' := eval unfold e in e in - let _ := match goal with _ => clear e end in - e'. - -Ltac head_fn term := - lazymatch type of term with - | forall (a : ?A), _ => (* idtac "forall" a; *) - let H := fresh in - let x := mk_evar A in - let t := constr:(term x) in - let _ := match goal with _ => specialize (term x) as H end in (* idtac H; *) - let h := head_fn H in - let _ := match goal with _ => clear H end in - (* let _ := match goal with _ => idtac "propogate" h end in *) constr:(h) - (* | ?A -> ?B => idtac A B; constr:(Type) - | _ -> ?a = ?b => idtac a b; constr:(Type) *) - | ?a = ?b => let h := get_head a in - let t := type of h in - let _ := match goal with _ => idtac "finish at" a "=" b "with" t end in - constr:(h) - | ?T => let h := get_head T in - let _ := match goal with _ => idtac "finish at" T "with" h end in - constr:(h) - end. - (* lazymatch ty with - | forall (a : ?A), ?P a => idtac "dep"; - (* let t := type of (P a) in idtac t "(dep)"; *) - let h := head_fn (P a) in constr:(h) - | forall (a : ?A), ?P => idtac "fun"; - (* let t := type of P in *) - (* let __ := match goal with _ => idtac t "(fun)" end in *) - let h := head_fn P in constr:(h) - | ?a = ?b => - (* let t := type of a in idtac t "(eq)"; *) - let h := head_fn a in idtac h; constr:(h) - | ?P => - let t := type of P in constr:(t) - end. *) -(* Ltac test := constr:(Matrix). -let t := test in idtac t. *) - -let h := (head_fn @Mmult_1_l) in idtac h. -Inductive ltac_tuple := - | ltac_nil - | ltac_cons (T : Type) (t : T) (rest : ltac_tuple). - -Fixpoint ltac_app (tu1 tu2 : ltac_tuple) : ltac_tuple := - match tu1 with - | ltac_nil => tu2 - | ltac_cons T t rest => ltac_app rest (ltac_cons T t tu2) - end. - -Local Notation "'#' x" := (ltac_cons _ x ltac_nil) (at level 100). - -Ltac get_heads term := - match term with - | ?f _ _ _ _ _ _ _ _ _=> constr:(# f) - | ?f _ _ _ _ _ _ _ _ => constr:(# f) - | ?f _ _ _ _ _ _ _ => constr:(# f) - | ?f _ _ _ _ _ _ => constr:(# f) - | ?f _ _ _ _ _ => constr:(# f) - | ?f _ _ _ _ => constr:(# f) - | ?f _ _ _ => constr:(# f) - | ?f _ _ => constr:(# f) - | ?f _ => constr:(# f) - | ?f => constr:(ltac_nil) - end. -Ltac head_fns term := - lazymatch type of term with - | forall (a : ?A), _ => (* idtac "forall" a; *) - let H := fresh in - let x := mk_evar A in - let t := constr:(term x) in - let _ := match goal with _ => specialize (term x) as H end in (* idtac H; *) - let rest := head_fns H in - let this := head_fns x in - (* let out := constr:(app h) *) - let _ := match goal with _ => clear H end in - (* let _ := match goal with _ => idtac "propogate" h end in *) - let out := constr:(ltac_app this rest) in - let out := eval cbn in out in constr:(out) - (* | ?A -> ?B => idtac A B; constr:(Type) - | _ -> ?a = ?b => idtac a b; constr:(Type) *) - | ?a = ?b => let h := get_heads a in constr:(h) - | ?T => let h := get_heads T in - (* let t := type of h in - let _ := match goal with _ => idtac t end in *) - (* let __ := match goal with _ => idtac "finish at" T "with" h end in *) - constr:(h) - end. -let h := (head_fns @direct_sum'_id) in idtac h. -let h := (head_fn @Mmult_1_l) in idtac h. - - -(* Require Import List. *) - -Ltac evarify_idxs f idxs evars := - lazymatch idxs with - | [] => constr:(f) - | O :: ?ls => - let _ := match goal with _ => idtac "hit" end in - lazymatch type of f with - | forall (a : ?A), ?P => - let _:=match goal with _=>idtac "indep:" f ": forall (a :" A ")," P end in - (* let a' := mk_evar A in *) - lazymatch evars with | ?a :: ?vars => - let f' := evarify_idxs (f a) ls vars in - let _:= match goal with _ => idtac " rec to" f' end in - constr:(fun (b:A) => f') - | _ => - let _ := match goal with - _ => - idtac "ERROR: Not enough evars! For evarify_idxs" f idxs evars "(returning first arg)" - end in constr:(f) - end - (* | forall (a : ?A), ?P a => - idtac "dep:" f ": forall (a :" A ")," P "a"; - let a' := mk_evar A in - constr:(forall (a:A), - ltac:( - exact (f a')) - ) *) - end - | (S ?n') :: ?ls => - let _ := match goal with _ => idtac "dec" end in - lazymatch type of f with - | forall (a : ?A), ?P => - constr:(fun (a : A) => - ltac:( - let f' := evarify_idxs (f a) constr:(n' :: ls) evars in - let _ := match goal with _ => idtac "subcall gave" f' end in - exact f')) - | _ => - let _ := match goal with _ => idtac "no more args" end in - constr:(f) - end - end. - - -let a := mk_evar nat in -let b := mk_evar nat in -let t := (evarify_idxs @Mmult [S O; O] [a; b]) in idtac t. - -specialize id_adjoint_eq as H; -let a := mk_evar nat in -let b := mk_evar nat in -let t := (evarify_idxs @adjoint [O; O] [a; b]) in -match H with -| context[@adjoint ?n ?m] => - replace (@adjoint n m) with (t n m) in H -end. -replace (@adjoint ?n ?m) with (t ?n ?m) in H. - - - - -evarify_2func' @adjoint. - -rewrite id_adjoint_eq. -evarify_3func' @Mmult. -evarify_2func' @mat_equiv. -rewrite Mmult_1_r_mat_eq. -split; f_equiv; lia. -Unshelve. -all : lia. -Qed. - - - - - - - - - -#[export] Instance MxDaggerMonoidalCategory : DaggerMonoidalCategory nat := { - dagger_tensor_compat := ltac: (intros; apply direct_sum'_adjoint); - - associator_unitary := ltac:(intros; unfold unitary; simpl in *; - rewrite Nat.add_assoc, id_adjoint_eq, Mmult_1_r by auto with wf_db; - easy); - (* associator_unitary_l := ltac:(intros; simpl in *; - rewrite Nat.add_assoc, id_adjoint_eq, Mmult_1_r by auto with wf_db; - easy); *) - left_unitor_unitary := ltac:(intros; unfold unitary; simpl in *; - rewrite id_adjoint_eq, Mmult_1_r by auto with wf_db; - easy); - (* left_unitor_unitary_l := ltac:(intros; simpl in *; - rewrite id_adjoint_eq, Mmult_1_r by auto with wf_db; - easy); *) - right_unitor_unitary := ltac:(intros; unfold unitary; simpl in *; - rewrite Nat.add_0_r, id_adjoint_eq, Mmult_1_r by auto with wf_db; - easy); - (* right_unitor_unitary_l := ltac:(intros; simpl in *; - rewrite Nat.add_0_r, id_adjoint_eq, Mmult_1_r by auto with wf_db; - easy); *) -}. - -#[export] Instance MxDaggerBraidedMonoidalCategory : DaggerBraidedMonoidalCategory nat := {}. - -#[export] Instance MxDaggerSymmetricMonoidalCategory : DaggerSymmetricMonoidalCategory nat := {}. \ No newline at end of file diff --git a/examples/DiscreteCategories.v b/examples/DiscreteCategories.v deleted file mode 100644 index b7d9661..0000000 --- a/examples/DiscreteCategories.v +++ /dev/null @@ -1,27 +0,0 @@ -From ViCaR Require Import CategoryTypeclass. - -Close Scope Cat_scope. - -From mathcomp Require Import fintype eqtype ssrnat. - -Open Scope Cat_scope. -Require Import TypesExample. -Open Scope bool_scope. - -#[program] Definition DiscreteCategory {n} : Category ('I_ n) := {| - morphism := fun m k => if m == k then True else False; - c_equiv := fun n m f g => - if (n == m) then True else False; - compose := fun n m k f g => - if (n==m) && (m == k) then _ else _ -|}. -Next Obligation. - split. - - intros ?. - destruct (A =P B) as [H|H]; easy. - - intros f g h. - destruct ( A =P B) as [H | H]; easy. - - intros f. - destruct (A =P B) as [H'|H']; easy. -Qed. -Admit Obligations. \ No newline at end of file diff --git a/examples/KronComm.v b/examples/KronComm.v index 30714fd..f653b4f 100644 --- a/examples/KronComm.v +++ b/examples/KronComm.v @@ -5,7 +5,7 @@ From VyZX Require Import CoreRules. From VyZX Require Import PermutationRules. Require Import MatrixExampleBase. Require Import MatrixPermBase. -From ViCaR Require Export CategoryTypeclass. +From ViCaR Require Import CategoryTypeclassCompatibility. Require Import ExamplesAutomation. Local Open Scope matrix_scope. @@ -2419,157 +2419,6 @@ Proof. apply div_divides; easy. Qed. - - -(* Lemma gcd_prod_helper : forall n m k, - n <> O -> m <> O -> k <> O -> - (Nat.gcd (Nat.gcd n m) k * Nat.gcd (m*n) k = (Nat.gcd n k) * (Nat.gcd m k))%nat. - (* (Nat.gcd (Nat.gcd (S n) (S m)) (S k) * Nat.gcd (S n * S m) (S k) = (Nat.gcd (S n) (S k)) * (Nat.gcd (S m) (S k)))%nat. *) -Proof. - intros n m k Hn Hm Hk. - Admitted. -Lemma gcd_prod : forall n m k, - n <> O -> m <> O -> k <> O -> - (Nat.gcd (m*n) k = (Nat.gcd n k) * (Nat.gcd m k) / Nat.gcd (Nat.gcd n m) k)%nat. -Proof. - intros n m k Hn Hm Hk. - rewrite <- gcd_prod_helper by easy. - rewrite (Nat.mul_comm (Nat.gcd _ _) _), Nat.div_mul; [easy | - rewrite ?Nat.gcd_eq_0; lia]. -Qed. - -Lemma sum_if_prod_eq_one_plus_gcd : forall m n, - Σ (fun j => Σ (fun i => if j * n =? i * m then C1 else 0) (S n)) (S m) = - INR (1 + Nat.gcd m n). -Proof. - intros m n. - destruct m, n; - try replace j with O by lia; - rewrite 1?Nat.mul_0_r, 1?Nat.mul_0_l, 1?Nat.mod_0_l, 1?Nat.eqb_refl by easy. - 1-3: (erewrite big_sum_eq_bounded; [ | - intros j Hj; - erewrite big_sum_eq_bounded; [ | - intros i Hi; - rewrite 2?Nat.mul_0_r; - reflexivity - ]; - reflexivity - ]). - 1: lca. - 1: rewrite Nat.gcd_0_l; simpl big_sum. - 2: erewrite Nat.gcd_0_r, big_sum_eq_bounded; [| intros j Hj; simpl; rewrite Cplus_0_l; - reflexivity]. - 1,2: rewrite big_sum_constant, Nat.add_comm, times_n_C1, - 1?plus_INR, ?S_INR; Csimpl; rewrite 2!RtoC_plus; lca. - rewrite (big_sum_eq_bounded _ (fun j => if (j*(S n) mod (S m) =? 0) then C1 else C0)). - 2: { - intros j Hj. - bdestruct_one. - - rewrite Nat.mod_divide in H by easy. - destruct H as [k Hk]. - apply big_sum_unique. - exists k; split; [nia | split; intros; bdestructΩ'simp]. - nia. - - rewrite big_sum_0_bounded; [easy|]. - intros i Hi. - bdestructΩ'simp. - rewrite H0, Nat.mod_mul in * by easy. - easy. - } - rewrite <- big_sum_extend_l. - rewrite plus_INR, RtoC_plus. - unfold C_is_monoid, Gplus. f_equal. - 1: rewrite Nat.mul_0_l, Nat.mod_0_l; easy. - pose proof (Nat.gcd_divide (S m) (S n)) as Hgcddiv. - (* Search (Nat.divide _ (_*_)). *) - assert (Hgcdnz: Nat.gcd (S m) (S n) <> O) by (rewrite Nat.gcd_eq_0; lia). - rewrite (big_sum_eq_bounded _ - (fun j => if ((S j) mod ((S m) / (Nat.gcd (S m) (S n))) =? 0) then C1 else C0)). - 2: { - intros j Hj. - enough (H: ((S j * S n) mod S m =? 0) = (S j mod (S m / Nat.gcd (S m) (S n)) =? 0)) - by (rewrite H; easy). - rewrite eq_iff_eq_true, 2!Nat.eqb_eq. - rewrite 2!Nat.mod_divide by (try easy; - rewrite Nat.div_small_iff by easy; - pose proof (gcd_le m n); lia). - split. - - intros [k Hk]. - exists (k * S m * Nat.gcd (S m) (S n) / (S m * S n))%nat. - rewrite <- Hk. - rewrite div_mul_combine; try easy. - 1: replace (S j * S n * Nat.gcd (S m) (S n) * S m)%nat with - (S j * (S m * S n * Nat.gcd (S m) (S n)))%nat by lia. - 1: rewrite Nat.div_mul by lia; easy. - rewrite Hk. - assert (Hgcd: Nat.gcd (S j * S n) (S n) = Nat.gcd (k * S m) (S n)) by - (rewrite Hk; easy). - assert (Hmul: (forall j', Nat.gcd (S j' * S n) (S n) = S n)%nat) by - (induction j'; [rewrite Nat.mul_1_l; apply Nat.gcd_diag| - rewrite Nat.mul_succ_l, Nat.gcd_comm, Nat.gcd_add_diag_r, - Nat.gcd_comm; apply IHj']). - rewrite Hmul in Hgcd. - rewrite (Nat.mul_comm k), <- Nat.mul_assoc. - rewrite Nat.mul_divide_cancel_l by easy. - rewrite Hgcd at 1. - exists (Nat.gcd (Nat.gcd (S m) k ) (S n) * (k/Nat.gcd k (S n)))%nat. - symmetry. - rewrite Nat.mul_comm, Nat.mul_assoc, (Nat.mul_comm (Nat.gcd (_*_) _)), gcd_prod_helper by lia. - rewrite <-Nat.mul_assoc, Nat.mul_comm. - f_equal. - rewrite <- Nat.divide_div_mul_exact, Nat.mul_comm, Nat.div_mul; try easy; - try apply Nat.gcd_divide; - rewrite Nat.gcd_eq_0; lia. - - intros [k Hk]. - rewrite Hk. - exists (k * (S n / Nat.gcd (S m) (S n)))%nat. - rewrite Nat.mul_comm, Nat.mul_assoc. - rewrite <- Nat.divide_div_mul_exact by easy. - symmetry. - rewrite Nat.mul_comm, Nat.mul_assoc. - rewrite <- Nat.divide_div_mul_exact by easy. - f_equal. - lia. - } - rewrite <- (Cplus_0_l (big_sum _ _)). - set (f:= fun j => if j mod (S m / Nat.gcd (S m) (S n)) =? 0 then C1 else 0). - - erewrite <- (Cminus_diag (f O) _ ltac:(reflexivity)) at 1. - unfold Cminus. - rewrite (Cplus_comm (_) (- _)), <- Cplus_assoc. - assert (Hdiv: (S m / Nat.gcd (S m) (S n) <> 0)%nat)by (rewrite Nat.div_small_iff by easy; - enough (Nat.gcd (S m) (S n) <= S m)%nat by lia; apply gcd_le). - pose proof (big_sum_extend_l (S m) f) as H. - unfold f at 2 in H. - unfold C_is_monoid, Gplus in H. - rewrite H. - set (g := fun (i j : nat) => if j =? 0 then C1 else 0). - rewrite (big_sum_eq_bounded f - (fun j => g (j / (S m / Nat.gcd (S m) (S n))) (j mod (S m / Nat.gcd (S m) (S n)))))%nat - by easy. - rewrite <- big_sum_extend_r. - replace (S m) with (S m / Nat.gcd (S m) (S n) * Nat.gcd (S m) (S n))%nat at 1. - rewrite <- big_sum_double_sum. - rewrite (big_sum_eq_bounded _ (fun _ => C1)). - 2: { - intros j Hj. - apply big_sum_unique. - exists O; split; [lia|unfold g; split; intros; bdestructΩ'simp]. - } - rewrite big_sum_constant, times_n_C1. - rewrite Cplus_assoc, (Cplus_comm (- _)), <- Cplus_assoc. - symmetry. - rewrite <- (Cplus_0_r (INR _)). - f_equal; try lca. - rewrite div_div by easy. - pose proof (div_divides (S m) (Nat.gcd (S m) (S n)) ltac:(easy) ltac:(easy)). - erewrite (proj2 (Nat.mod_divide _ _ _)); try easy. - unfold f, g; rewrite Nat.mod_0_l, Nat.eqb_refl; try lca; try easy. - Unshelve. - 2: easy. - rewrite Nat.mul_comm, <- Nat.divide_div_mul_exact, Nat.mul_comm, Nat.div_mul; easy. -Qed. *) - Lemma f_to_vec_split : forall (f : nat -> bool) (m n : nat), f_to_vec (m + n) f = f_to_vec m f ⊗ f_to_vec n (VectorStates.shift f m). Proof. diff --git a/examples/KronComm_orig.v b/examples/KronComm_orig.v deleted file mode 100644 index 4059149..0000000 --- a/examples/KronComm_orig.v +++ /dev/null @@ -1,1215 +0,0 @@ -Require Import Setoid. - -From VyZX Require Import CoreData. -From VyZX Require Import CoreRules. -From VyZX Require Import PermutationRules. -From ViCaR Require Export CategoryTypeclass. - -Local Open Scope matrix_scope. - -Lemma Msum_transpose : forall n m p f, - (big_sum (G:=Matrix n m) f p) ⊤ = - big_sum (G:=Matrix n m) (fun i => (f i) ⊤) p. -Proof. - intros. - rewrite (big_sum_func_distr f transpose); easy. -Qed. - -Ltac print_state := - try (match goal with | H : ?p |- _ => idtac H ":" p; fail end); - idtac "---------------------------------------------------------"; - match goal with |- ?P => idtac P -end. - - -Ltac is_C0 x := - assert (x = C0) by lca. - -Ltac is_C1 x := - assert (x = C1) by lca. - -Ltac print_C x := - tryif is_C0 x then idtac "0" else - tryif is_C1 x then idtac "1" else idtac "X". - -Ltac print_LHS_matU := - intros; - (let i := fresh "i" in - let j := fresh "j" in - let Hi := fresh "Hi" in - let Hj := fresh "Hj" in - intros i j Hi Hj; try solve_end; - repeat - (destruct i as [| i]; [ | apply <- Nat.succ_lt_mono in Hi ]; - try solve_end); clear Hi; - repeat - (destruct j as [| j]; [ | apply <- Nat.succ_lt_mono in Hj ]; - try solve_end); clear Hj); - match goal with |- ?x = ?y ?i ?j => autounfold with U_db; simpl; - match goal with - | |- ?x = _ => idtac i; idtac j; print_C x; idtac "" - end -end. - -Definition kron_comm p q : Matrix (p*q) (p*q):= - @make_WF (p*q) (p*q) (fun s t => - (* have blocks H_ij, p by q of them, and each is q by p *) - let i := (s / q)%nat in let j := (t / p)%nat in - let k := (s mod q)%nat in let l := (t mod p) in - (* let k := (s - q * i)%nat in let l := (t - p * t)%nat in *) - if (i =? l) && (j =? k) then C1 else C0 - (* s/q =? t mod p /\ t/p =? s mod q *) -). - -Lemma WF_kron_comm p q : WF_Matrix (kron_comm p q). -Proof. unfold kron_comm; auto with wf_db. Qed. -#[export] Hint Resolve WF_kron_comm : wf_db. - -(* Lemma test_kron : kron_comm 2 3 = Matrix.Zero. -Proof. - apply mat_equiv_eq; unfold kron_comm; auto with wf_db. - print_LHS_matU. -*) - -Lemma kron_comm_transpose : forall p q, - (kron_comm p q) ⊤ = kron_comm q p. -Proof. - intros p q. - apply mat_equiv_eq; auto with wf_db. - 1: rewrite Nat.mul_comm; apply WF_kron_comm. - intros i j Hi Hj. - unfold kron_comm, transpose, make_WF. - rewrite andb_comm, Nat.mul_comm. - rewrite (andb_comm (_ =? _)). - easy. -Qed. - -Lemma kron_comm_1_r : forall p, - (kron_comm p 1) = Matrix.I p. -Proof. - intros p. - apply mat_equiv_eq; [|rewrite 1?Nat.mul_1_r|]; auto with wf_db. - intros s t Hs Ht. - unfold kron_comm. - unfold make_WF. - unfold Matrix.I. - rewrite Nat.mul_1_r, Nat.div_1_r, Nat.mod_1_r, Nat.div_small, Nat.mod_small by lia. - bdestructΩ'. -Qed. - -Lemma kron_comm_1_l : forall p, - (kron_comm 1 p) = Matrix.I p. -Proof. - intros p. - apply mat_equiv_eq; [|rewrite 1?Nat.mul_1_l|]; auto with wf_db. - intros s t Hs Ht. - unfold kron_comm. - unfold make_WF. - unfold Matrix.I. - rewrite Nat.mul_1_l, Nat.div_1_r, Nat.mod_1_r, Nat.div_small, Nat.mod_small by lia. - bdestructΩ'. -Qed. - -Definition mx_to_vec {n m} (A : Matrix n m) : Vector (n*m) := - make_WF (fun i j => A (i mod n)%nat (i / n)%nat - (* Note: goes columnwise. Rowwise would be: - make_WF (fun i j => A (i / m)%nat (i mod n)%nat - *) -). - -Lemma WF_mx_to_vec {n m} (A : Matrix n m) : WF_Matrix (mx_to_vec A). -Proof. unfold mx_to_vec; auto with wf_db. Qed. -#[export] Hint Resolve WF_mx_to_vec : wf_db. - -(* Compute vec_to_list (mx_to_vec (Matrix.I 2)). *) -From Coq Require Import ZArith. -Ltac Zify.zify_post_hook ::= PreOmega.Z.div_mod_to_equations. - -Lemma kron_comm_vec_helper : forall i p q, (i < p * q)%nat -> - (p * (i mod q) + i / q < p * q)%nat. -Proof. - intros i p q. - intros Hi. - assert (i / q < p)%nat by (apply Nat.div_lt_upper_bound; lia). - destruct p; [lia|]; - destruct q; [lia|]. - enough (S p * (i mod S q) <= S p * q)%nat by lia. - apply Nat.mul_le_mono; [lia | ]. - pose proof (Nat.mod_upper_bound i (S q) ltac:(easy)). - lia. -Qed. - -Lemma mx_to_vec_additive {n m} (A B : Matrix n m) : - mx_to_vec (A .+ B) = mx_to_vec A .+ mx_to_vec B. -Proof. - apply mat_equiv_eq; auto with wf_db. - intros i j Hi Hj. - replace j with O by lia; clear dependent j. - unfold mx_to_vec, make_WF, Mplus. - bdestructΩ'. -Qed. - -Lemma if_mult_dist_r (b : bool) (z : C) : - (if b then C1 else C0) * z = - if b then z else C0. -Proof. - destruct b; lca. -Qed. - -Lemma if_mult_dist_l (b : bool) (z : C) : - z * (if b then C1 else C0) = - if b then z else C0. -Proof. - destruct b; lca. -Qed. - -Lemma if_mult_and (b c : bool) : - (if b then C1 else C0) * (if c then C1 else C0) = - if (b && c) then C1 else C0. -Proof. - destruct b; destruct c; lca. -Qed. - -Lemma kron_comm_vec : forall p q (A : Matrix p q), - kron_comm p q × mx_to_vec A = mx_to_vec (A ⊤). -Proof. - intros p q A. - apply mat_equiv_eq; [|rewrite Nat.mul_comm|]; auto with wf_db. - intros i j Hi Hj. - replace j with O by lia; clear dependent j. - unfold transpose, mx_to_vec, kron_comm, make_WF, Mmult. - rewrite (Nat.mul_comm q p). - replace_bool_lia (i . - destruct p; [lia|]. - destruct q; [lia|]. - split. - + rewrite Nat.add_comm, Nat.mul_comm. - rewrite Nat.mod_add by easy. - rewrite Nat.mod_small; [lia|]. - apply Nat.div_lt_upper_bound; lia. - + rewrite Nat.mul_comm, Nat.div_add_l by easy. - rewrite Nat.div_small; [lia|]. - apply Nat.div_lt_upper_bound; lia. - - intros [Hmodp Hdivp]. - rewrite (Nat.div_mod_eq k p). - lia. - } - apply big_sum_unique. - exists (p * (i mod q) + i / q)%nat; repeat split; - [apply kron_comm_vec_helper; easy | rewrite Nat.eqb_refl | intros; bdestructΩ']. - destruct p; [lia|]; - destruct q; [lia|]. - f_equal. - - rewrite Nat.add_comm, Nat.mul_comm, Nat.mod_add, Nat.mod_small; try easy. - apply Nat.div_lt_upper_bound; lia. - - rewrite Nat.mul_comm, Nat.div_add_l by easy. - rewrite Nat.div_small; [lia|]. - apply Nat.div_lt_upper_bound; lia. -Qed. - -Lemma kron_comm_ei_kron_ei_sum : forall p q, - kron_comm p q = - big_sum (G:=Square (p*q)) (fun i => big_sum (G:=Square (p*q)) (fun j => - (@e_i p i ⊗ @e_i q j) × ((@e_i q j ⊗ @e_i p i) ⊤)) - q) p. -Proof. - intros p q. - apply mat_equiv_eq; auto with wf_db. - 1: apply WF_Msum; intros; apply WF_Msum; intros; - rewrite Nat.mul_comm; apply WF_mult; - auto with wf_db; rewrite Nat.mul_comm; - auto with wf_db. - intros i j Hi Hj. - rewrite Msum_Csum. - erewrite big_sum_eq_bounded. - 2: { - intros k Hk. - rewrite Msum_Csum. - erewrite big_sum_eq_bounded. - 2: { - intros l Hl. - unfold Mmult, kron, transpose, e_i. - erewrite big_sum_eq_bounded. - 2: { - intros m Hm. - (* replace m with O by lia. *) - rewrite Nat.div_1_r, Nat.mod_1_r. - replace_bool_lia (m =? 0) true; rewrite 4!andb_true_r. - rewrite 3!if_mult_and. - match goal with - |- context[if ?b then _ else _] => - replace b with ((i =? k * q + l) && (j =? l * p + k)) - end. - 1: reflexivity. (* set our new function *) - clear dependent m. - rewrite eq_iff_eq_true, 8!andb_true_iff, - 6!Nat.eqb_eq, 4!Nat.ltb_lt. - split. - - intros [Hieq Hjeq]. - subst i j. - rewrite 2!Nat.div_add_l, Nat.div_small, Nat.add_0_r by lia. - rewrite Nat.add_comm, Nat.mod_add, Nat.mod_small, - Nat.div_small, Nat.add_0_r by lia. - rewrite Nat.add_comm, Nat.mod_add, Nat.mod_small by lia. - easy. - - intros [[[] []] [[] []]]. - split. - + rewrite (Nat.div_mod_eq i q) by lia; lia. - + rewrite (Nat.div_mod_eq j p) by lia; lia. - } - simpl; rewrite Cplus_0_l. - reflexivity. - } - apply big_sum_unique. - exists (i mod q). - split; [|split]. - - apply Nat.mod_upper_bound; lia. - - reflexivity. - - intros l Hl Hnmod. - bdestructΩ'. - exfalso; apply Hnmod. - rewrite Nat.add_comm, Nat.mod_add, Nat.mod_small by lia; lia. - } - symmetry. - apply big_sum_unique. - exists (j mod p). - repeat split. - - apply Nat.mod_upper_bound; lia. - - unfold kron_comm, make_WF. - replace_bool_lia (i - enough (b = c) by bdestructΩ' - end. - rewrite eq_iff_eq_true, 2!andb_true_iff, 4!Nat.eqb_eq. - split. - + intros [Hieq Hjeq]. - split; [rewrite Hieq | rewrite Hjeq]; - rewrite Hieq, Nat.div_add_l by lia; - (rewrite Nat.div_small; [lia|]); - apply Nat.mod_upper_bound; lia. - + intros [Hidiv Hjdiv]. - rewrite (Nat.div_mod_eq i q) at 1 by lia. - rewrite (Nat.div_mod_eq j p) at 2 by lia. - lia. - - intros k Hk Hkmod. - bdestructΩ'. - exfalso; apply Hkmod. - rewrite Nat.add_comm, Nat.mod_add, Nat.mod_small by lia; lia. -Qed. - -Lemma kron_comm_ei_kron_ei_sum' : forall p q, - kron_comm p q = - big_sum (G:=Square (p*q)) (fun ij => - let i := (ij / q)%nat in let j := (ij mod q) in - ((@e_i p i ⊗ @e_i q j) × ((@e_i q j ⊗ @e_i p i) ⊤))) (p*q). -Proof. - intros p q. - rewrite kron_comm_ei_kron_ei_sum, big_sum_double_sum, Nat.mul_comm. - reflexivity. -Qed. - -Lemma div_eq_iff : forall a b c, b <> O -> - (a / b)%nat = c <-> (b * c <= a /\ a < b * (S c))%nat. -Proof. - intros a b c Hb. - split. - intros Hadivb. - split; - subst c. - etransitivity; [ - apply Nat.div_mul_le, Hb |]. - rewrite Nat.mul_comm, Nat.div_mul; easy. - apply Nat.mul_succ_div_gt, Hb. - intros [Hge Hlt]. - symmetry. - apply (Nat.div_unique _ _ _ (a - b*c)); [lia|]. - lia. -Qed. - -Lemma kron_e_i_transpose_l : forall k n m o (A : Matrix m o), (k < n)%nat -> - (o <> O) -> (m <> O) -> - (@e_i n k)⊤ ⊗ A = (fun i j => - if (i - (o <> O) -> (m <> O) -> - (@e_i n k) ⊗ A = (fun i j => - if (j - (o <> O) -> (m <> O) -> - (@e_i n k)⊤ ⊗ A = (fun i j => - if (i ((j/o)%nat=k)) by lia; - rewrite Hrw; clear Hrw. - symmetry. - rewrite div_eq_iff by lia. - lia. - - replace (i / m =? 0) with false. - rewrite andb_false_r; easy. - symmetry. - rewrite Nat.eqb_neq. - rewrite Nat.div_small_iff; lia. -Qed. - -Lemma kron_e_i_l' : forall k n m o (A : Matrix m o), (k < n)%nat -> - (o <> O) -> (m <> O) -> - (@e_i n k) ⊗ A = (fun i j => - if (j ((i/m)%nat=k)) by lia; - rewrite Hrw; clear Hrw. - symmetry. - rewrite div_eq_iff by lia. - lia. - - replace (j / o =? 0) with false. - rewrite andb_false_r; easy. - symmetry. - rewrite Nat.eqb_neq. - rewrite Nat.div_small_iff; lia. -Qed. - -Lemma kron_e_i_r : forall k n m o (A : Matrix m o), (k < n)%nat -> - (o <> O) -> (m <> O) -> - A ⊗ (@e_i n k) = (fun i j => - if (i mod n =? k) then A (i / n)%nat j else 0). -Proof. - intros k n m o A Hk Ho Hm. - apply functional_extensionality; intros i; - apply functional_extensionality; intros j. - unfold kron, e_i. - rewrite if_mult_dist_l, Nat.div_1_r. - rewrite Nat.mod_1_r, Nat.eqb_refl, andb_true_r. - replace (i mod n - (o <> O) -> (m <> O) -> - A ⊗ (@e_i n k) ⊤ = (fun i j => - if (j mod n =? k) then A i (j / n)%nat else 0). -Proof. - intros k n m o A Hk Ho Hm. - apply functional_extensionality; intros i; - apply functional_extensionality; intros j. - unfold kron, transpose, e_i. - rewrite if_mult_dist_l, Nat.div_1_r. - rewrite Nat.mod_1_r, Nat.eqb_refl, andb_true_r. - replace (j mod n m <> O -> - (@e_i n k) ⊤ ⊗ (Matrix.I m) ⊗ (@e_i n k) = - (fun i j => if (i mod n =? k) && (j / m =? k)%nat - && (i / n =? j - k * m) && (i / n - (@e_i n j) ⊤ ⊗ (Matrix.I m) ⊗ (@e_i n j)) n. -Proof. - intros m n. - apply mat_equiv_eq; auto with wf_db. - 1: apply WF_Msum; intros; apply WF_kron; auto with wf_db arith. - intros i j Hi Hj. - rewrite Msum_Csum. - erewrite big_sum_eq_bounded. - 2: { - intros ij Hij. - rewrite ei_kron_I_kron_ei by lia. - reflexivity. - } - unfold kron_comm, make_WF. - replace_bool_lia (i - (@e_i m i) ⊗ (Matrix.I n) ⊗ (@e_i m i)⊤) m. -Proof. - intros. - rewrite <- (kron_comm_transpose n m). - rewrite (kron_comm_kron_form_sum n m). - rewrite Msum_transpose. - apply big_sum_eq_bounded. - intros k Hk. - rewrite Nat.mul_1_l. - pose proof (kron_transpose _ _ _ _ ((@e_i m k) ⊤ ⊗ Matrix.I n) (@e_i m k)) as H; - rewrite Nat.mul_1_l, Nat.mul_1_r in H; - rewrite (Nat.mul_comm n m), H in *; clear H. - pose proof (kron_transpose _ _ _ _ ((@e_i m k) ⊤) (Matrix.I n)) as H; - rewrite Nat.mul_1_l in H; - rewrite H; clear H. - rewrite transpose_involutive, id_transpose_eq; easy. -Qed. - -Lemma e_i_dot_is_component : forall p k (x : Vector p), - (k < p)%nat -> WF_Matrix x -> - (@e_i p k) ⊤ × x = x k O .* Matrix.I 1. -Proof. - intros p k x Hk HWF. - apply mat_equiv_eq; auto with wf_db. - intros i j Hi Hj; - replace i with O by lia; - replace j with O by lia; - clear i Hi; - clear j Hj. - unfold Mmult, transpose, scale, e_i, Matrix.I. - bdestructΩ'. - rewrite Cmult_1_r. - apply big_sum_unique. - exists k. - split; [easy|]. - bdestructΩ'. - rewrite Cmult_1_l. - split; [easy|]. - intros l Hl Hkl. - bdestructΩ'; lca. -Qed. - -Lemma kron_e_i_e_i : forall p q k l, - (k < p)%nat -> (l < q)%nat -> - @e_i q l ⊗ @e_i p k = @e_i (p*q) (l*p + k). -Proof. - intros p q k l Hk Hl. - apply functional_extensionality; intro i. - apply functional_extensionality; intro j. - unfold kron, e_i. - rewrite Nat.mod_1_r, Nat.div_1_r. - rewrite if_mult_and. - lazymatch goal with - |- (if ?b then _ else _) = (if ?c then _ else _) => - enough (H : b = c) by (rewrite H; easy) - end. - rewrite Nat.eqb_refl, andb_true_r. - destruct (j =? 0); [|rewrite 2!andb_false_r; easy]. - rewrite 2!andb_true_r. - rewrite eq_iff_eq_true, 4!andb_true_iff, 3!Nat.eqb_eq, 3!Nat.ltb_lt. - split. - - intros [[] []]. - rewrite (Nat.div_mod_eq i p). - split; nia. - - intros []. - subst i. - rewrite Nat.div_add_l, Nat.div_small, Nat.add_0_r, - Nat.add_comm, Nat.mod_add, Nat.mod_small by lia. - easy. -Qed. - -Lemma kron_eq_sum : forall p q (x : Vector q) (y : Vector p), - WF_Matrix x -> WF_Matrix y -> - y ⊗ x = big_sum (fun ij => - let i := (ij / q)%nat in let j := ij mod q in - (x j O * y i O) .* (@e_i p i ⊗ @e_i q j)) (p * q). -Proof. - intros p q x y Hwfx Hwfy. - - erewrite big_sum_eq_bounded. - 2: { - intros ij Hij. - simpl. - rewrite (@kron_e_i_e_i q p) by - (try apply Nat.mod_upper_bound; try apply Nat.div_lt_upper_bound; lia). - rewrite (Nat.mul_comm (ij / q) q). - rewrite <- (Nat.div_mod_eq ij q). - reflexivity. - } - apply mat_equiv_eq; [|rewrite Nat.mul_comm|]; auto with wf_db. - intros i j Hi Hj. - replace j with O by lia; clear j Hj. - simpl. - rewrite Msum_Csum. - symmetry. - apply big_sum_unique. - exists i. - split; [lia|]. - unfold e_i; split. - - unfold scale, kron; bdestructΩ'. - rewrite Cmult_1_r, Cmult_comm. - easy. - - intros j Hj Hij. - unfold scale, kron; bdestructΩ'. - rewrite Cmult_0_r. - easy. -Qed. - -Lemma kron_comm_commutes_vectors_l : forall p q (x : Vector q) (y : Vector p), - WF_Matrix x -> WF_Matrix y -> - kron_comm p q × (x ⊗ y) = (y ⊗ x). -Proof. - intros p q x y Hwfx Hwfy. - rewrite kron_comm_ei_kron_ei_sum', Mmult_Msum_distr_r. - erewrite big_sum_eq_bounded. - 2: { - intros k Hk. - simpl. - match goal with - |- (?A × ?B) × ?C = _ => - assert (Hassoc: (A × B) × C = A × (B × C)) by apply Mmult_assoc - end. - simpl in Hassoc. - rewrite (Nat.mul_comm q p) in *. - rewrite Hassoc. clear Hassoc. - pose proof (kron_transpose _ _ _ _ (@e_i q (k mod q)) (@e_i p (k / q))) as Hrw; - rewrite (Nat.mul_comm q p) in Hrw; - simpl in Hrw; rewrite Hrw; clear Hrw. - pose proof (kron_mixed_product ((e_i (k mod q)) ⊤) ((e_i (k / q)) ⊤) x y) as Hrw; - rewrite (Nat.mul_comm q p) in Hrw; - simpl in Hrw; rewrite Hrw; clear Hrw. - rewrite 2!e_i_dot_is_component; [| - apply Nat.div_lt_upper_bound; lia | - easy | - apply Nat.mod_upper_bound; lia | - easy]. - rewrite Mscale_kron_dist_l, Mscale_kron_dist_r, Mscale_assoc. - rewrite kron_1_l, Mscale_mult_dist_r, Mmult_1_r by auto with wf_db. - reflexivity. - } - rewrite <- kron_eq_sum; easy. -Qed. - -Lemma kron_basis_vector_basis_vector : forall p q k l, - (k < p)%nat -> (l < q)%nat -> - basis_vector q l ⊗ basis_vector p k = basis_vector (p*q) (l*p + k). -Proof. - intros p q k l Hk Hl. - apply functional_extensionality; intros i. - apply functional_extensionality; intros j. - unfold kron, basis_vector. - rewrite Nat.mod_1_r, Nat.div_1_r, Nat.eqb_refl, andb_true_r, if_mult_and. - bdestructΩ'; - try pose proof (Nat.div_mod_eq i p); - try nia. - rewrite Nat.div_add_l, Nat.div_small in * by lia. - lia. -Qed. - -Lemma kron_extensionality : forall n m s t (A B : Matrix (n*m) (s*t)), - WF_Matrix A -> WF_Matrix B -> - (forall (x : Vector s) (y :Vector t), - WF_Matrix x -> WF_Matrix y -> - A × (x ⊗ y) = B × (x ⊗ y)) -> - A = B. -Proof. - intros b n s t A B HwfA HwfB Hext. - apply equal_on_basis_vectors_implies_equal; try easy. - intros i Hi. - - pose proof (Nat.div_lt_upper_bound i t s ltac:(lia) ltac:(lia)). - pose proof (Nat.mod_upper_bound i s ltac:(lia)). - pose proof (Nat.mod_upper_bound i t ltac:(lia)). - - specialize (Hext (basis_vector s (i / t)) (basis_vector t (i mod t)) - ltac:(apply basis_vector_WF; easy) - ltac:(apply basis_vector_WF; easy) - ). - rewrite (kron_basis_vector_basis_vector t s) in Hext by lia. - - simpl in Hext. - rewrite (Nat.mul_comm (i/t) t), <- (Nat.div_mod_eq i t) in Hext. - rewrite (Nat.mul_comm t s) in Hext. easy. -Qed. - -Lemma kron_comm_commutes : forall n s m t - (A : Matrix n s) (B : Matrix m t), - WF_Matrix A -> WF_Matrix B -> - kron_comm m n × (A ⊗ B) × (kron_comm s t) = (B ⊗ A). -Proof. - intros n s m t A B HwfA HwfB. - apply (kron_extensionality _ _ t s); [| - apply WF_kron; try easy; lia |]. - rewrite (Nat.mul_comm t s); apply WF_mult; auto with wf_db; - apply WF_mult; auto with wf_db; - rewrite Nat.mul_comm; auto with wf_db. - (* rewrite Nat.mul_comm; apply WF_mult; [rewrite Nat.mul_comm|auto with wf_db]; - apply WF_mult; auto with wf_db; rewrite Nat.mul_comm; auto with wf_db. *) - intros x y Hwfx Hwfy. - (* simpl. *) - (* Search "assoc" in Matrix. *) - rewrite (Nat.mul_comm s t). - rewrite (Mmult_assoc (_ × _)). - rewrite (Nat.mul_comm t s). - rewrite kron_comm_commutes_vectors_l by easy. - rewrite Mmult_assoc, (Nat.mul_comm m n). - rewrite kron_mixed_product. - rewrite (Nat.mul_comm n m), kron_comm_commutes_vectors_l by (auto with wf_db). - rewrite <- kron_mixed_product. - f_equal; lia. -Qed. - -Lemma commute_kron : forall n s m t - (A : Matrix n s) (B : Matrix m t), - WF_Matrix A -> WF_Matrix B -> - (A ⊗ B) = kron_comm n m × (B ⊗ A) × (kron_comm t s). -Proof. - intros n s m t A B HA HB. - rewrite (kron_comm_commutes m t n s B A HB HA); easy. -Qed. - -#[export] Hint Extern 4 (WF_Matrix (@Mmult ?m ?n ?o ?A ?B)) => ( - match type of A with Matrix ?m' ?n' => - match type of B with Matrix ?n'' ?o'' => - let Hm' := fresh "Hm'" in let Hn' := fresh "Hn'" in - let Hn'' := fresh "Hn''" in let Ho'' := fresh "Hoo'" in - assert (Hm' : m = m') by lia; - assert (Hn' : n = n') by lia; - assert (Hn'' : n = n'') by lia; - assert (Ho' : o = o'') by lia; - replace (@Mmult m n o A B) with (@Mmult m' n' o A B) - by (first [try (rewrite Hm' at 1); try (rewrite Hn' at 1); reflexivity | f_equal; lia]); - apply WF_mult; [ - auto with wf_db | - apply WF_Matrix_dim_change; - auto with wf_db - ] - end end) : wf_db. - -Lemma kron_comm_mul_inv : forall p q, - kron_comm p q × kron_comm q p = Matrix.I _. -Proof. - intros p q. - apply mat_equiv_eq; auto with wf_db. - intros i j Hi Hj. - unfold Mmult, kron_comm, make_WF. - erewrite big_sum_eq_bounded. - 2: { - intros k Hk. - rewrite <- 2!andb_if, if_mult_and. - replace_bool_lia (k - replace b with ((i =? j) && (k =? (i mod q) * p + (j/q))) - end; - [reflexivity|]. - rewrite eq_iff_eq_true, 4!andb_true_iff, 6!Nat.eqb_eq. - split. - - intros [? ?]; subst. - destruct p; [easy|destruct q;[lia|]]. - assert (j / S q < S p)%nat by (apply Nat.div_lt_upper_bound; lia). - rewrite Nat.div_add_l, (Nat.div_small (j / (S q))), Nat.add_0_r by easy. - rewrite Nat.add_comm, Nat.mod_add, Nat.mod_small by easy. - easy. - - intros [[Hiqkp Hkpiq] [Hkpjq Hjqkp]]. - split. - + rewrite (Nat.div_mod_eq i q), (Nat.div_mod_eq j q). - lia. - + rewrite (Nat.div_mod_eq k p). - lia. - } - bdestruct (i =? j). - - subst. - apply big_sum_unique. - exists ((j mod q) * p + (j/q))%nat. - split; [|split]. - + rewrite Nat.mul_comm. apply kron_comm_vec_helper; easy. - + unfold Matrix.I. - rewrite Nat.eqb_refl; bdestructΩ'. - + intros; bdestructΩ'. - - unfold Matrix.I. - replace_bool_lia (i =? j) false. - rewrite andb_false_l. - rewrite big_sum_0; [easy|]. - intros; rewrite andb_false_l; easy. -Qed. - -Lemma kron_comm_mul_transpose_r : forall p q, - kron_comm p q × (kron_comm p q) ⊤ = Matrix.I _. -Proof. - intros p q. - rewrite (kron_comm_transpose p q). - apply kron_comm_mul_inv. -Qed. - -Lemma kron_comm_mul_transpose_l : forall p q, - (kron_comm p q) ⊤ × kron_comm p q = Matrix.I _. -Proof. - intros p q. - rewrite <- (kron_comm_transpose q p). - rewrite (Nat.mul_comm p q). - rewrite (transpose_involutive _ _ (kron_comm q p)). - apply kron_comm_mul_transpose_r. -Qed. - -Lemma kron_comm_commutes_l : forall n s m t - (A : Matrix n s) (B : Matrix m t), - WF_Matrix A -> WF_Matrix B -> - kron_comm m n × (A ⊗ B) = (B ⊗ A) × (kron_comm t s). -Proof. - intros n s m t A B HwfA HwfB. - match goal with |- ?A = ?B => - rewrite <- (Mmult_1_r _ _ A), <- (Mmult_1_r _ _ B) ; auto with wf_db - end. - rewrite (Nat.mul_comm t s). - rewrite <- (kron_comm_mul_transpose_r), <- 2!Mmult_assoc. - rewrite (kron_comm_commutes n s m t) by easy. - apply Mmult_simplify; [|easy]. - rewrite Mmult_assoc. - rewrite (Nat.mul_comm s t), (kron_comm_mul_inv t s), Mmult_1_r by auto with wf_db. - easy. -Qed. - -Lemma kron_comm_commutes_r : forall n s m t - (A : Matrix n s) (B : Matrix m t), - WF_Matrix A -> WF_Matrix B -> - (A ⊗ B) × kron_comm s t = (kron_comm n m) × (B ⊗ A). -Proof. - intros n s m t A B HA HB. - rewrite kron_comm_commutes_l; easy. -Qed. - - - -(* Lemma kron_comm_commutes_r : forall n s m t - (A : Matrix n s) (B : Matrix m t), - WF_Matrix A -> WF_Matrix B -> - kron_comm m n × (A ⊗ B) = (B ⊗ A) × (kron_comm t s). -Proof. - intros n s m t A B HwfA HwfB. - match goal with |- ?A = ?B => - rewrite <- (Mmult_1_r _ _ A), <- (Mmult_1_r _ _ B) ; auto with wf_db - end. - rewrite (Nat.mul_comm t s). - rewrite <- (kron_comm_mul_transpose_r), <- 2!Mmult_assoc. - rewrite (kron_comm_commutes n s m t) by easy. - apply Mmult_simplify; [|easy]. - rewrite Mmult_assoc. - rewrite (Nat.mul_comm s t), (kron_comm_mul_inv t s), Mmult_1_r by auto with wf_db. - easy. -Qed. *) - - - - -Lemma vector_eq_basis_comb : forall n (y : Vector n), - WF_Matrix y -> - y = big_sum (G:=Vector n) (fun i => y i O .* @e_i n i) n. -Proof. - intros n y Hwfy. - apply mat_equiv_eq; auto with wf_db. - intros i j Hi Hj. - replace j with O by lia; clear j Hj. - symmetry. - rewrite Msum_Csum. - apply big_sum_unique. - exists i. - repeat split; try easy. - - unfold ".*", e_i; bdestructΩ'; lca. - - intros l Hl Hnk. - unfold ".*", e_i; bdestructΩ'; lca. -Qed. - -Lemma kron_vecT_matrix_vec : forall m n o p - (P : Matrix m o) (y : Vector n) (z : Vector p), - WF_Matrix y -> WF_Matrix z -> WF_Matrix P -> - (z⊤) ⊗ P ⊗ y = @Mmult (m*n) (m*n) (o*p) (kron_comm m n) ((y × (z⊤)) ⊗ P). -Proof. - intros m n o p P y z Hwfy Hwfz HwfP. - match goal with |- ?A = ?B => - rewrite <- (Mmult_1_l _ _ A) ; auto with wf_db - end. - rewrite Nat.mul_1_l. - rewrite <- (kron_comm_mul_transpose_r), Mmult_assoc at 1. - rewrite Nat.mul_1_r, (Nat.mul_comm o p). - apply Mmult_simplify; [easy|]. - rewrite kron_comm_kron_form_sum. - rewrite Msum_transpose. - rewrite Mmult_Msum_distr_r. - erewrite big_sum_eq_bounded. - 2: { - intros k Hk. - pose proof (kron_transpose _ _ _ _ ((@e_i n k) ⊤ ⊗ Matrix.I m) (@e_i n k)) as H; - rewrite Nat.mul_1_l, Nat.mul_1_r, (Nat.mul_comm m n) in *; - rewrite H; clear H. - pose proof (kron_transpose _ _ _ _ ((@e_i n k) ⊤) (Matrix.I m)) as H; - rewrite Nat.mul_1_l in *; - rewrite H; clear H. - restore_dims. - rewrite 2!kron_mixed_product. - rewrite id_transpose_eq, Mmult_1_l by easy. - rewrite e_i_dot_is_component, transpose_involutive by easy. - (* rewrite <- Mmult_transpose. *) - rewrite Mscale_kron_dist_r, <- 2!Mscale_kron_dist_l. - rewrite kron_1_r. - rewrite <- Mscale_mult_dist_l. - reflexivity. - } - rewrite <- (kron_Msum_distr_r n _ P). - rewrite <- (Mmult_Msum_distr_r). - rewrite <- vector_eq_basis_comb by easy. - easy. -Qed. - -Lemma kron_vec_matrix_vecT : forall m n o p - (Q : Matrix n o) (x : Vector m) (z : Vector p), - WF_Matrix x -> WF_Matrix z -> WF_Matrix Q -> - x ⊗ Q ⊗ (z⊤) = @Mmult (m*n) (m*n) (o*p) (kron_comm m n) (Q ⊗ (x × z⊤)). -Proof. - intros m n o p Q x z Hwfx Hwfz HwfQ. - match goal with |- ?A = ?B => - rewrite <- (Mmult_1_l _ _ A) ; auto with wf_db - end. - rewrite Nat.mul_1_r. - rewrite <- (kron_comm_mul_transpose_r), Mmult_assoc at 1. - rewrite Nat.mul_1_l. - apply Mmult_simplify; [easy|]. - rewrite kron_comm_kron_form_sum'. - rewrite Msum_transpose. - rewrite Mmult_Msum_distr_r. - erewrite big_sum_eq_bounded. - 2: { - intros k Hk. - pose proof (kron_transpose _ _ _ _ ((@e_i m k) ⊗ Matrix.I n) ((@e_i m k) ⊤)) as H; - rewrite Nat.mul_1_l, Nat.mul_1_r, (Nat.mul_comm m n) in *; - rewrite H; clear H. - pose proof (kron_transpose _ _ _ _ ((@e_i m k)) (Matrix.I n)) as H; - rewrite Nat.mul_1_l, (Nat.mul_comm m n) in *; - rewrite H; clear H. - restore_dims. - rewrite 2!kron_mixed_product. - rewrite id_transpose_eq, Mmult_1_l by easy. - rewrite e_i_dot_is_component, transpose_involutive by easy. - (* rewrite <- Mmult_transpose. *) - rewrite 2!Mscale_kron_dist_l, kron_1_l, <-Mscale_kron_dist_r by easy. - rewrite <- Mscale_mult_dist_l. - restore_dims. - reflexivity. - } - rewrite <- (kron_Msum_distr_l m _ Q). - rewrite <- (Mmult_Msum_distr_r). - rewrite <- vector_eq_basis_comb by easy. - easy. -Qed. - -Lemma kron_comm_triple_cycle : forall m n s t p q (A : Matrix m n) - (B : Matrix s t) (C : Matrix p q), WF_Matrix A -> WF_Matrix B -> WF_Matrix C -> - A ⊗ B ⊗ C = (kron_comm (m*s) p) × (C ⊗ A ⊗ B) × (kron_comm q (t*n)). -Proof. - intros m n s t p q A B C HA HB HC. - rewrite (commute_kron _ _ _ _ (A ⊗ B) C) by auto with wf_db. - rewrite kron_assoc by easy. - f_equal; try lia; f_equal; lia. -Qed. - -Lemma kron_comm_triple_cycle2 : forall m n s t p q (A : Matrix m n) - (B : Matrix s t) (C : Matrix p q), WF_Matrix A -> WF_Matrix B -> WF_Matrix C -> - A ⊗ B ⊗ C = (kron_comm m (s*p)) × (B ⊗ C ⊗ A) × (kron_comm (q*t) n). -Proof. - intros m n s t p q A B C HA HB HC. - rewrite kron_assoc by easy. - rewrite (commute_kron _ _ _ _ A (B ⊗ C)) by auto with wf_db. - f_equal; try lia; f_equal; lia. -Qed. - -Lemma id_eq_sum_kron_e_is : forall n, - Matrix.I n = big_sum (G:=Square n) (fun i => @e_i n i ⊗ (@e_i n i) ⊤) n. -Proof. - intros n. - symmetry. - apply mat_equiv_eq; auto with wf_db. - intros i j Hi Hj. - rewrite Msum_Csum. - erewrite big_sum_eq_bounded. - 2: { - intros k Hk. - rewrite kron_e_i_l by lia. - unfold transpose, e_i. - rewrite <- andb_if. - replace_bool_lia (j rewrite (commute_kron _ _ _ _ A B) by auto with wf_db - end. - (* restore_dims. *) - reflexivity. - } - (* rewrite ?Nat.mul_1_r, ?Nat.mul_1_l. *) - (* rewrite <- Mmult_Msum_distr_r. *) - - rewrite <- (Mmult_Msum_distr_r n _ (kron_comm (t*1) (n*s))). - rewrite <- Mmult_Msum_distr_l. - erewrite big_sum_eq_bounded. - 2: { - intros j Hj. - rewrite <- kron_assoc, (kron_assoc (Matrix.I t)) by auto with wf_db. - restore_dims. - reflexivity. - } - (* rewrite Nat.mul_1_l *) - rewrite <- (kron_Msum_distr_r n _ (Matrix.I s)). - rewrite <- (kron_Msum_distr_l n _ (Matrix.I t)). - rewrite 2!Nat.mul_1_r, 2!Nat.mul_1_l. - rewrite <- (id_eq_sum_kron_e_is n). - rewrite 2!id_kron. - restore_dims. - rewrite Mmult_1_r by auto with wf_db. - rewrite (Nat.mul_comm t n), (Nat.mul_comm n s). - easy. -Qed. - -Lemma kron_comm_cycle_indices_rev : forall t s n, - @Mmult (s*(n*t)) (s*(n*t)) (t*(s*n)) (kron_comm s (n*t)) (kron_comm t (s*n)) = kron_comm (t*s) n. -Proof. - intros. - rewrite <- kron_comm_cycle_indices. - easy. -Qed. - - -Lemma kron_comm_triple_id : forall t s n, - (kron_comm (t*s) n) × (kron_comm (s*n) t) × (kron_comm (n*t) s) = Matrix.I (t*s*n). -Proof. - intros t s n. - rewrite kron_comm_cycle_indices. - restore_dims. - rewrite (Mmult_assoc (kron_comm s (n*t))). - restore_dims. - rewrite (Nat.mul_comm (s*n) t). (* TODO: Fix kron_comm_mul_inv to have the - right dimensions by default (or, better yet, to be ambivalent) *) - rewrite (kron_comm_mul_inv t (s*n)). - restore_dims. - rewrite Mmult_1_r by auto with wf_db. - rewrite (Nat.mul_comm (n*t) s). - rewrite (kron_comm_mul_inv). - f_equal; lia. -Qed. - -Lemma kron_comm_triple_id' : forall t s n, - (kron_comm n (t*s)) × (kron_comm t (s*n)) × (kron_comm s (n*t)) = Matrix.I (t*s*n). -Proof. - intros t s n. - apply transpose_matrices. - rewrite 2!Mmult_transpose. - (* restore_dims. *) - rewrite (kron_comm_transpose s (n*t)). - restore_dims. - rewrite (kron_comm_transpose n (t*s)). - restore_dims. - replace (n*(t*s))%nat with (t*(s*n))%nat by lia. - replace (s*(n*t))%nat with (t*(s*n))%nat by lia. - rewrite (kron_comm_transpose t (s*n)). - restore_dims. - rewrite Nat.mul_assoc, id_transpose_eq. - restore_dims. - (* rewrite (kron_comm_triple_id n t s). *) - Admitted. - -Lemma kron_comm_triple_indices_commute : forall t s n, - @Mmult (s*t*n) (s*t*n) (t*(s*n)) (kron_comm (s*t) n) (kron_comm t (s*n)) = - @Mmult (t*(s*n)) (t*(s*n)) (s*t*n) (kron_comm t (s*n)) (kron_comm (s*t) n). -Proof. - intros t s n. - rewrite <- (kron_comm_cycle_indices_rev s t n). - Admitted. - (* rewrite kron_comm_triple_id. -rewrite <- (kron_comm_cycle_indices t s n). *) - - -Lemma f_to_vec_split : forall (f : nat -> bool) (m n : nat), - f_to_vec (m + n) f = f_to_vec m f ⊗ f_to_vec n (VectorStates.shift f m). -Proof. - intros f m n. - rewrite f_to_vec_merge. - apply f_to_vec_eq. - intros i Hi. - unfold VectorStates.shift. - bdestructΩ'. - f_equal; lia. -Qed. - -Lemma n_top_to_bottom_semantics_eq_kron_comm : forall n o, - ⟦ n_top_to_bottom n o ⟧ = kron_comm (2^o) (2^n). -Proof. - intros n o. - rewrite zxperm_permutation_semantics by auto with zxperm_db. - unfold zxperm_to_matrix. - rewrite perm_of_n_top_to_bottom. - apply equal_on_basis_states_implies_equal; auto with wf_db. - 1: { - rewrite Nat.add_comm, Nat.pow_add_r. - auto with wf_db. - } - intros f. - pose proof (perm_to_matrix_permutes_qubits (n + o) (rotr (n+o) n) f). - unfold perm_to_matrix in H. - rewrite H by auto with perm_db. - rewrite (f_to_vec_split f). - pose proof (kron_comm_commutes_vectors_l (2^o) (2^n) - (f_to_vec n f) (f_to_vec o (@VectorStates.shift bool f n)) - ltac:(auto with wf_db) ltac:(auto with wf_db)). - replace (2^(n+o))%nat with (2^o *2^n)%nat by (rewrite Nat.pow_add_r; lia). - simpl in H0. - rewrite H0. - rewrite Nat.add_comm, f_to_vec_split. - f_equal. - - apply f_to_vec_eq. - intros i Hi. - unfold VectorStates.shift. - f_equal; unfold rotr. - rewrite Nat.mod_small by lia. - bdestructΩ'. - - apply f_to_vec_eq. - intros i Hi. - unfold VectorStates.shift, rotr. - rewrite <- Nat.add_assoc, mod_add_n_r, Nat.mod_small by lia. - bdestructΩ'. -Qed. - - -Lemma compose_semantics' : -forall {n m o : nat} (zx0 : ZX n m) (zx1 : ZX m o), -@eq (Matrix (Nat.pow 2 o) (Nat.pow 2 n)) - (@ZX_semantics n o (@Compose n m o zx0 zx1)) - (@Mmult (Nat.pow 2 o) (Nat.pow 2 m) (Nat.pow 2 n) - (@ZX_semantics m o zx1) (@ZX_semantics n m zx0)). -Proof. - intros. - rewrite (@compose_semantics n m o). - easy. -Qed. diff --git a/examples/KronMatrixExample.v b/examples/KronMatrixExample.v index e6ae3ea..ef60cec 100644 --- a/examples/KronMatrixExample.v +++ b/examples/KronMatrixExample.v @@ -2,20 +2,23 @@ Require Import MatrixPermBase. Require Import KronComm. Require Export MatrixExampleBase. Require Import ExamplesAutomation. +From ViCaR Require Import CategoryTypeclassCompatibility. #[export] Instance MxCategory : Category nat := { morphism := Matrix; - c_equiv := @mat_equiv; (* fun m n => @eq (Matrix m n); *) + compose := @Mmult; + c_identity n := I n; +}. + +#[export] Instance MxCategoryC : CategoryCoherence MxCategory := { c_equiv_rel := @mat_equiv_equivalence; - compose := @Mmult; compose_compat := fun n m o f g Hfg h i Hhi => @Mmult_simplify_mat_equiv n m o f g h i Hfg Hhi; assoc := @Mmult_assoc_mat_equiv; - c_identity n := I n; left_unit := Mmult_1_l_mat_eq; right_unit := Mmult_1_r_mat_eq; }. @@ -29,16 +32,17 @@ Definition MxKronBiFunctor : Bifunctor MxCategory MxCategory MxCategory := {| morphism_bicompat := ltac:(intros; apply kron_mat_equiv_morph; easy); |}. +Obligation Tactic := idtac. - -#[export] Instance MxKronMonoidalCategory : MonoidalCategory nat := { - tensor := MxKronBiFunctor; - I := 1; +#[export, program] Instance MxKronMonoidalCategory + : MonoidalCategory MxCategory := { + obj_tensor := Nat.mul; + mor_tensor := @kron; + mon_I := 1; associator := fun n m o => {| forward := (I (n * m * o) : Matrix (n * m * o) (n * (m * o))); reverse := (I (n * m * o) : Matrix (n * (m * o)) (n * m * o)); - isomorphism_inverse := ltac:(split; simpl; rewrite Nat.mul_assoc, Mmult_1_r_mat_eq; easy); (* id_A := ltac:(simpl; rewrite Nat.mul_assoc, Mmult_1_r_mat_eq; easy); id_B := ltac:(simpl; rewrite Nat.mul_assoc, Mmult_1_r_mat_eq; easy); *) |}; @@ -46,8 +50,6 @@ Definition MxKronBiFunctor : Bifunctor MxCategory MxCategory MxCategory := {| left_unitor := fun n => {| forward := (I n : Matrix (1 * n) n); reverse := (I n : Matrix n (1 * n)); - isomorphism_inverse - := ltac:(abstract(split; rewrite Nat.mul_1_l, Mmult_1_r_mat_eq; easy)); (* id_A := ltac:(rewrite Nat.mul_1_l, Mmult_1_r_mat_eq; easy); id_B := ltac:(rewrite Nat.mul_1_l, Mmult_1_r_mat_eq; easy); *) |}; @@ -55,36 +57,74 @@ Definition MxKronBiFunctor : Bifunctor MxCategory MxCategory MxCategory := {| right_unitor := fun n => {| forward := (I n : Matrix (n * 1) n); reverse := (I n : Matrix n (n * 1)); - isomorphism_inverse - := ltac:(abstract(split; rewrite Nat.mul_1_r, Mmult_1_r_mat_eq; easy)); (* id_A := ltac:(rewrite Nat.mul_1_r, Mmult_1_r_mat_eq; easy); id_B := ltac:(rewrite Nat.mul_1_r, Mmult_1_r_mat_eq; easy); *) |}; +}. +Next Obligation. + split; + simpl; + rewrite Nat.mul_assoc, Mmult_1_r_mat_eq; + easy. +Qed. +Next Obligation. + split; + rewrite Nat.mul_1_l, Mmult_1_r_mat_eq; + easy. +Qed. +Next Obligation. + split; + rewrite Nat.mul_1_r, Mmult_1_r_mat_eq; + easy. +Qed. - associator_cohere := ltac:(intros; simpl in *; - rewrite kron_assoc_mat_equiv; - rewrite 2!Nat.mul_assoc, Mmult_1_r_mat_eq, Mmult_1_l_mat_eq; - easy - ); - left_unitor_cohere := ltac:(intros; cbn; - rewrite kron_1_l_mat_equiv, 2!Nat.add_0_r, - Mmult_1_l_mat_eq, Mmult_1_r_mat_eq; easy); - right_unitor_cohere := ltac:(intros; cbn; +#[export, program] Instance MxKronMonoidalCategoryC : + MonoidalCategoryCoherence MxKronMonoidalCategory := {}. +Next Obligation. + intros; simpl in *. + rewrite id_kron; easy. +Qed. +Next Obligation. + simpl; intros. + rewrite kron_mixed_product; easy. +Qed. +Next Obligation. + simpl; intros. + apply kron_mat_equiv_morph; easy. +Qed. +Next Obligation. + intros; simpl in *. + rewrite kron_assoc_mat_equiv. + rewrite 2!Nat.mul_assoc, Mmult_1_r_mat_eq, Mmult_1_l_mat_eq. + easy. +Qed. +Next Obligation. + intros; simpl. + rewrite kron_1_l_mat_equiv, 2!Nat.add_0_r, + Mmult_1_l_mat_eq, Mmult_1_r_mat_eq. + easy. +Qed. +Next Obligation. + intros; simpl. rewrite kron_1_r_mat_equiv, 2!Nat.mul_1_r, - Mmult_1_l_mat_eq, Mmult_1_r_mat_eq; easy); - - pentagon := ltac:(intros; simpl in *; - rewrite ?Nat.mul_assoc, 2!id_kron, Mmult_1_l_mat_eq; - rewrite ?Nat.mul_assoc, Mmult_1_l_mat_eq; - easy - ); - triangle := ltac:(intros; - cbn; - rewrite Nat.mul_1_r, Nat.add_0_r in *; - rewrite Mmult_1_l_mat_eq; - easy - ); -}. + Mmult_1_l_mat_eq, Mmult_1_r_mat_eq; easy. +Qed. +Next Obligation. + intros; simpl in *. + (* change (B+0)%nat with (1*B)%nat. *) + rewrite Nat.add_0_r, ?Nat.mul_assoc. + rewrite Nat.mul_1_r. + rewrite id_kron. + rewrite Mmult_1_r_mat_eq. + easy. +Qed. +Next Obligation. + intros; simpl in *. + rewrite ?Nat.mul_assoc. + rewrite 2!id_kron, 2!Mmult_1_l_mat_eq. + rewrite ?Nat.mul_assoc. + easy. +Qed. Definition MxKronBraidingIsomorphism : forall n m, Isomorphism (MxKronBiFunctor n m) ((CommuteBifunctor MxKronBiFunctor) n m) := @@ -282,16 +322,30 @@ Proof. Qed. -#[export, program] Instance MxBraidedMonoidal : BraidedMonoidalCategory nat := { - braiding := MxKronBraidingBiIsomorphism +#[export] Instance MxBraidedMonoidal + : BraidedMonoidalCategory MxKronMonoidalCategory := { + braiding := MxKronBraidingIsomorphism +}. + +#[export, program] Instance MxBraidedMonoidalC + : BraidedMonoidalCategoryCoherence MxBraidedMonoidal := { + (* braiding_natural *) }. Next Obligation. + intros; simpl in *. + rewrite (Nat.mul_comm B2 B1), (Nat.mul_comm A2 A1). + rewrite (kron_comm_commutes_r_mat_equiv). + easy. +Qed. +Next Obligation. + intros; simpl in *. rewrite 3!Nat.mul_assoc. rewrite Mmult_1_r_mat_eq, Mmult_1_r_mat_eq. rewrite Mmult_1_l_mat_eq. apply mx_hexagon_helper. Qed. Next Obligation. + intros; simpl in *. rewrite 3!Nat.mul_assoc. rewrite Mmult_1_r_mat_eq, Mmult_1_r_mat_eq. rewrite Mmult_1_l_mat_eq. @@ -299,255 +353,14 @@ Next Obligation. Qed. #[export, program] Instance MxSymmetricMonoidal : - SymmetricMonoidalCategory nat. + SymmetricMonoidalCategory MxBraidedMonoidal. Next Obligation. easy. Qed. -Definition mat_unit (n : nat) : Matrix 1 (n*n) := - fun i j => if (i =? 0) && (0 @e_i n i ⊗ @e_i n i) n. -Proof. - intros i j Hi Hj; replace j with O by lia; clear j Hj. - unfold mat_unit, e_i, transpose, kron. - replace_bool_lia (0 A; - unit := mat_unit; - counit := fun n => (mat_unit n) ⊤; -}. -Next Obligation. - rewrite Nat.add_0_r, Nat.mul_1_r. - rewrite Mmult_1_r_mat_eq, Mmult_1_l_mat_eq. - rewrite <- Nat.mul_assoc. - rewrite Mmult_1_r_mat_eq. - apply mat_equiv_of_equiv_on_ei. - intros k Hk. - rewrite Mmult_1_l_mat_eq. - rewrite Mmult_assoc. - pose (kron_e_i_e_i A 1 k 0 Hk (le_n 1)) as Hek. - rewrite !Nat.mul_1_r, Nat.mul_0_l, Nat.add_0_l in Hek. - rewrite <- Hek at 1. - rewrite Nat.mul_assoc. - pose (@kron_mixed_product (A*A)%nat 1 1 A A 1) as Hkron; - rewrite !Nat.mul_1_l in Hkron. - rewrite Hkron. - assert (H: @e_i 1 0 = I 1) by (unfold e_i; solve_matrix; bdestructΩ'simp). - rewrite H. - rewrite Mmult_1_r_mat_eq, Mmult_1_l_mat_eq. - intros i j Hi Hj; replace j with O by lia; clear j Hj. - unfold I, e_i, kron, mat_unit, Mmult, transpose. - rewrite !Nat.div_1_r, !Nat.mod_1_r. - replace_bool_lia (0 0)%nat by lia. - assert (A*A <> 0)%nat by lia. - exists ((A*A)*k + k*(1 + A))%nat. - pose proof (Nat.mod_small k A Hi) as Hmodk. - assert (k < A * A)%nat by nia. - pose proof (Nat.mod_small k (A*A) ltac:(nia)). - split; [nia|]. - rewrite Nat.mul_comm, Nat.div_add_l, Nat.div_small, - Nat.add_0_r, Nat.eqb_refl, andb_true_r, Cmult_1_l by (easy || nia). - rewrite Nat.add_comm, Nat.mod_add, Nat.mod_small by (easy || nia). - rewrite Nat.mul_add_distr_l, Nat.mul_1_r, Nat.div_add, Nat.mod_add, - Nat.div_small, Nat.mod_small, Nat.eqb_refl, Cmult_1_l by (easy || nia). - rewrite Nat.div_div, Nat.div_add, Nat.div_small, Nat.mul_assoc, - 2!Nat.div_add, Nat.div_small, Nat.mod_add, Nat.mod_small, - Nat.eqb_refl, Cmult_1_l by (easy || nia). - split; [rewrite 2!Nat.mod_add, Hmodk; bdestructΩ'simp|]. - intros k' Hk' Hfalse. - pose proof (Nat.div_mod k' (A*A) ltac:(nia)) as Hk'eq. - rewrite Hk'eq. - pose proof (Nat.mod_upper_bound k' (A*A) ltac:(lia)) as Hk'modAA. - rewrite Nat.mul_comm, Nat.div_add_l, (Nat.div_small _ _ Hk'modAA), - Nat.add_0_r, andb_true_r by (easy || nia). - bdestruct (k =? k' / (A*A)); [|lca]. - replace <- (k' / (A * A))%nat. - rewrite Nat.add_comm, Nat.mod_add, Cmult_1_l by (easy || nia). - rewrite Nat.mod_mod, mod_product, Nat.div_div, Nat.div_add by (easy || nia). - rewrite Nat.mul_assoc, Nat.div_add, Nat.mod_add, <- mod_of_div by easy. - rewrite Nat.mod_mod, Nat.mod_add, mod_product by easy. - bdestruct (k' mod A =? k); [|lca]. - replace (k' mod A). - replace_bool_lia (k 0)%nat by lia. - assert (A*A <> 0)%nat by lia. - exists ((A*A)*k + k*(1 + A))%nat. - pose proof (Nat.mod_small k A Hi) as Hmodk. - assert (k < A * A)%nat by nia. - pose proof (Nat.mod_small k (A*A) ltac:(nia)). - split; [nia|]. - rewrite Nat.div_small, Nat.mul_comm, Nat.div_div, Nat.div_add_l, - Nat.div_small, Nat.add_0_r, Nat.add_comm, Nat.mul_assoc, - Nat.mul_add_distr_l, Nat.mul_1_r, 2!Nat.div_add, Nat.mod_add, - Nat.div_small, Nat.mod_small, Nat.add_0_l, 2!Nat.eqb_refl, - andb_true_r, Cmult_1_l, Hmodk by (easy || nia). - rewrite 2!Nat.mod_add, Hmodk, Nat.eqb_refl, <- Nat.mul_assoc, Nat.mod_add, - mod_product, Nat.mod_add, Nat.mod_small, Nat.div_add, Nat.div_small, - Hmodk, Nat.eqb_refl by (easy || nia). - replace_bool_lia (k j < m -> A i j <-> B i j. - -Definition compose_relns {n m o : nat} (A : reln n m) (B : reln m o) : reln n o := - fun i j => exists k, k < m /\ A i k /\ B k j. - -Local Notation "A '~' B" := (reln_equiv A B) (at level 80). -Local Notation "A '\o' B" := (compose_relns A B) (at level 40, left associativity). - -Definition idn {n} : reln n n := - fun i j => i = j. - -Lemma compose_relns_assoc {n m o p : nat} - (A : reln n m) (B : reln m o) (C : reln o p) : - A \o B \o C ~ A \o (B \o C). -Proof. - intros i j Hi Hj. - unfold compose_relns. - split; - [ intros [? [? [[? [? [? ?]]] ?]]] - | intros [? [? [? [? [? [? ?]]]]]]]; - eexists; repeat split; try eexists; - repeat split; eassumption. -Qed. - -Lemma compose_relns_id_l {n m} (A : reln n m) : - idn \o A ~ A. -Proof. - intros i j Hi Hj. - unfold compose_relns, idn. - split; - [intros [? [? []]]; subst; easy | - intros; eexists; repeat split; eassumption]. -Qed. - -Lemma compose_relns_id_r {n m} (A : reln n m) : - A \o idn ~ A. -Proof. - intros i j Hi Hj. - unfold compose_relns, idn. - split; - [intros [? [? []]]; subst; easy | - intros; eexists; repeat split; eassumption]. -Qed. - -Lemma reln_equiv_equivalence {n m} : - equivalence (reln n m) (@reln_equiv n m). -Proof. - split; unfold reln_equiv; try easy. - - intros A B C HAB HBC i j Hi Hj. - rewrite HAB by easy. - apply HBC; easy. - - intros A B HAB. symmetry; apply HAB; easy. -Qed. - -Lemma compose_relns_compat {n m o} : forall (A A' : reln n m), - A ~ A' -> forall (B B' : reln m o), - B ~ B' -> A \o B ~ A' \o B'. -Proof. - unfold reln_equiv, compose_relns. - intros A A' HA B B' HB i j Hi Hj. - split; - intros [k [Hk [HAk HBk]]]; - [ rewrite HA in HAk by easy; rewrite HB in HBk by easy - | rewrite <-HA in HAk by easy; rewrite <-HB in HBk by easy]; - eexists; - repeat split; try eassumption. -Qed. - -#[export] Instance RelationCategory : Category nat := { - morphism := @reln; - c_equiv := @reln_equiv; - c_equiv_rel := @reln_equiv_equivalence; - compose := @compose_relns; - compose_compat := @compose_relns_compat; - assoc := @compose_relns_assoc; - c_identity := @idn; - left_unit := @compose_relns_id_l; - right_unit := @compose_relns_id_r; -}. - -Definition stack_relns n0 m0 n1 m1 (A : reln n0 m0) (B : reln n1 m1) - : reln (n0+n1) (m0+m1) := - fun i j => - (i < n0) /\ (j < m0) /\ A i j - \/ (n0 <= i < n0 + n1) /\ (m0 <= j < m0 + m1) /\ B (i - n0) (j - m0). - - - - - -Ltac solve_relations_fast := - repeat (repeat (intros ?); subst; try lia; match goal with - | H : (exists _, _) |- _ => destruct H - | H : _ /\ _ |- _ => destruct H - | |- _ <-> _ => split - | |- _ /\ _ => split - | |- ?f ?x ?y => match type of f with reln ?n ?m => eassumption end - | |- ?x = _ => match type of x with nat => reflexivity || lia end - | |- _ \/ _ => (left; solve [solve_relations_fast; lia]) - || (right; solve [solve_relations_fast; lia]) || (* print_state; *) fail 2 - - | |- exists _, _ => - (* idtac "trying vars"; print_state; *) - match goal with - | x : nat |- _ => match goal with - | _ : reln x _ |- _ => (* idtac "not" x; *) fail 1 - | _ : reln _ x |- _ => (* idtac "not" x; *) fail 1 - | _ => exists x; (* idtac "yes" x; *) (solve [solve_relations_fast; lia]) - end - end - | H : ?A ?i ?j |- ?A ?i' ?j' => - match type of A with - | reln ?n ?m => - (* idtac "matched " H " :"; print_state; *) - (apply H || ( (* idtac "needed rewrite"; *) - replace i' with i by lia; replace j' with j by lia; apply H) - (* || idtac "failed" *) - ) - end - | H : ?A ~ ?A' |- ?A ?i ?j => - eassumption || rewrite (H i j) by lia; eassumption - | H : ?A ~ ?A' |- ?A' ?i ?j => - eassumption || rewrite <- (H i j) by lia; eassumption - - - end); - repeat ( - try eassumption; - try reflexivity; - try easy; - try lia); - try lia. - - -Ltac __solve_relations_end := - try eassumption; - try reflexivity; - try easy; - (* try lia); *) - try lia; - try match goal with - | |- ?f ?x ?y => match type of f with reln ?n ?m => eassumption end - | |- ?x = _ => match type of x with nat => reflexivity || lia end - | H : ?A ?i ?j |- ?A ?i' ?j' => - match type of A with - | reln ?n ?m => - (* idtac "matched " H " :"; print_state; *) - (apply H || ( (* idtac "needed rewrite"; *) - replace i' with i by lia; replace j' with j by lia; apply H) - (* || idtac "failed" *) - ) - end - | H : ?A ~ ?A' |- ?A ?i ?j => - eassumption || rewrite (H i j) by lia; eassumption - | H : ?A ~ ?A' |- ?A' ?i ?j => - eassumption || rewrite <- (H i j) by lia; eassumption - (* - match type of H with - | _ < _ => fail 1 - | _ <= _ => fail 1 - | _ => - (* idtac "matched " H " :"; print_state; *) - (apply H || ( (* idtac "needed rewrite"; *) - replace i' with i by lia; replace j' with j by lia; apply H) - (* || idtac "failed" *) - ) - end - *) - end. - - -Ltac __solve_relations_l := idtac. -Ltac __solve_relations_r := idtac. - -Ltac __solve_relations_l ::= - repeat (repeat (intros ?); subst; auto; try lia; match goal with - | H : (exists _, _) |- _ => destruct H - | H : _ /\ _ |- _ => destruct H - | |- _ <-> _ => split - | |- _ /\ _ => split - | H : _ \/ _ |- _ => destruct H; [__solve_relations_l | __solve_relations_r] - | |- _ \/ _ => - (left; solve [__solve_relations_l; lia]) - || (right; solve [__solve_relations_r; lia]) - || (* print_state; *) fail 2 - | |- exists _, _ => - (* idtac "trying vars"; print_state; *) - match goal with - | x : nat |- _ => match goal with - | _ : reln x _ |- _ => (* idtac "not" x; *) fail 1 - | _ : reln _ x |- _ => (* idtac "not" x; *) fail 1 - | _ => exists x; (* idtac "yes" x; *) (solve [__solve_relations_l; lia]) - end - end; match goal with - | |- _ => (* idtac "none worked"; *) eexists - end - end); - __solve_relations_end. - -Ltac __solve_relations_r ::= - repeat (repeat (intros ?); subst; auto; try lia; match goal with - | H : (exists _, _) |- _ => destruct H - | H : _ /\ _ |- _ => destruct H - | |- _ <-> _ => split - | |- _ /\ _ => split - | H : _ \/ _ |- _ => destruct H; [__solve_relations_l | __solve_relations_r] - | |- _ \/ _ => - (right; solve [__solve_relations_r; lia]) - || (left; solve [__solve_relations_l; lia]) - || (* print_state; *) fail 2 - | |- exists _, _ => - (* idtac "trying vars"; print_state; *) - match goal with - | x : nat |- _ => match goal with - | _ : reln x _ |- _ => (* idtac "not" x; *) fail 1 - | _ : reln _ x |- _ => (* idtac "not" x; *) fail 1 - | _ => exists x; (* idtac "yes" x; *) (solve [__solve_relations_r; lia]) - end - end; match goal with - | |- _ => (* idtac "none worked"; *) eexists - end - end); - __solve_relations_end. - -Ltac solve_relations := - simpl in *; - unfold idn, compose_relns, stack_relns, reln_equiv in *; - __solve_relations_l. - -Lemma stack_relns_idn : forall n m, - stack_relns n n m m idn idn ~ idn. -Proof. - intros n m i j Hi Hj; - unfold stack_relns, "id_"; simpl; - unfold idn; - lia. -Qed. - -Lemma stack_relns_compose : forall n0 m0 o0 n1 m1 o1 - (A1 : reln n0 m0) (B1 : reln m0 o0) (A2 : reln n1 m1) (B2 : reln m1 o1), - stack_relns n0 o0 n1 o1 (A1 \o B1) (A2 \o B2) ~ - stack_relns n0 m0 n1 m1 A1 A2 \o stack_relns m0 o0 m1 o1 B1 B2. -Proof. - intros n0 m0 o0 n1 m1 o1 A1 B1 A2 B2. - simpl. - unfold stack_relns, compose_relns. - intros i j Hi Hj. - split. - + intros [[Hi' [Hj' [k [Hk [HAk HBk]]]]] | [Hi' [Hj' [k [Hk [HAk HBk]]]]]]. - * exists k; split; [lia|]. - split; left; repeat split; lia || easy. - * exists (m0 + k); split; [lia|]. - split; right; repeat split; lia || (try easy); - replace (m0+k-m0) with k by lia; easy. - + intros [k [Hk [[[? [? ?]] | [? [? ?]]] [[? [? ?]] | [[? ?] [? ?]]]]]]; - try lia; [left | right]; - repeat (split; try easy); eexists; repeat split; try eassumption; lia. -Qed. - -Lemma stack_relns_compat : forall n0 m0 n1 m1 - (A A' : reln n0 m0) (B B' : reln n1 m1), A ~ A' -> B ~ B' -> - stack_relns n0 m0 n1 m1 A B ~ stack_relns n0 m0 n1 m1 A' B'. -Proof. - simpl. - intros n0 m0 n1 m1 A A' B B' HA HB. - revert n0 m0 n1 m1 A A' B B' HA HB. - intros i j Hi Hj. - simpl in *. - unfold stack_relns, reln_equiv in *. - split; - (intros [[? [? ?]] | [? [? ?]]]; [left | right]); - (split; [lia | split; [lia | ]]); - (rewrite 1?HA, 1?HB by lia; easy) || - (rewrite <- 1?HA, <- 1?HB by lia; easy). -Qed. - -#[export] Instance StackRelnsBifunctor : - Bifunctor RelationCategory RelationCategory RelationCategory. - refine {| - obj_bimap := Nat.add; - morphism_bimap := stack_relns; - |}. -Proof. - - apply stack_relns_idn. - - apply stack_relns_compose. - - apply stack_relns_compat. -Defined. - -#[export] Instance RelationMonoidalCategory : MonoidalCategory nat := { - tensor := StackRelnsBifunctor; - I := O; - - associator := fun n m o => {| - forward := (@idn (n + m + o) : reln (n + m + o) (n + (m +o))); - reverse := (@idn (n + m + o) : reln (n + (m +o)) (n + m + o)); - isomorphism_inverse := ltac:(abstract(solve_relations)); - |}; - - left_unitor := fun n => {| - forward := (@idn n : reln (0 + n) n); - reverse := (@idn n : reln n (0 + n)); - isomorphism_inverse := ltac:(abstract(solve_relations)); - |}; - - right_unitor := fun n => {| - forward := (@idn n : reln (n + 0) n); - reverse := (@idn n : reln n (n + 0)); - isomorphism_inverse := ltac:(abstract(solve_relations)); - |}; - - associator_cohere := ltac:(abstract(solve_relations)); - left_unitor_cohere := ltac:(abstract(solve_relations)); - right_unitor_cohere := ltac:(abstract(solve_relations)); - - pentagon := ltac:(abstract(solve_relations)); - triangle := ltac:(abstract(solve_relations)); -}. - -Definition braid_relation n m : reln (n + m) (m + n) := - fun i j => - (i < n) /\ (m <= j <= m + n) /\ (i = j - m) \/ - (n <= i < n + m) /\ (j < m) /\ (i - n = j). - -Definition BraidRelationIsomorphism n m : Isomorphism (n + m) (m + n). - refine ({| - forward := braid_relation n m; - reverse := braid_relation m n; |}). -Proof. - unfold braid_relation. - solve_relations; intros; subst j. - - bdestruct (i A j i. - -Lemma flipn_idn {n} : - flipn (@idn n) ~ @idn n. -Proof. - easy. -Qed. - -Lemma flipn_compose {n m o} (A : reln n m) (B : reln m o) : - flipn (A \o B) ~ flipn B \o flipn A. -Proof. - unfold flipn. - solve_relations. -Qed. - -Lemma flipn_compat {n m} (A B : reln n m) : - A ~ B -> flipn A ~ flipn B. -Proof. - unfold flipn. - solve_relations. -Qed. - -#[export] Instance RelationDaggerCategory : - DaggerCategory nat. - apply Build_DaggerCategory with (fun A B => flipn). - - reflexivity. - - apply @flipn_idn. - - apply @flipn_compose. - - apply @flipn_compat. -Defined. - -#[export] Instance RelationDaggerMonoidalCategory : - DaggerMonoidalCategory nat. -Proof. - apply Build_DaggerMonoidalCategory. - - simpl ; unfold flipn. - solve_relations. - - unfold unitary; simpl; unfold flipn. - solve_relations. - - unfold unitary; simpl; unfold flipn. - solve_relations. - - unfold unitary; simpl; unfold flipn. - solve_relations. -Qed. - - diff --git a/examples/RelationsExample.v b/examples/RelationsExample.v index 4ab462d..1556b95 100644 --- a/examples/RelationsExample.v +++ b/examples/RelationsExample.v @@ -2,7 +2,6 @@ Require Import Setoid. Require Import Logic. Require Import Basics (flip). From ViCaR Require Export CategoryTypeclass. -From ViCaR Require Import RigCategory. Definition reln (A B : Type) : Type := A -> B -> Prop. @@ -255,11 +254,14 @@ Qed. #[export] Instance Rel : Category Type := { morphism := reln; c_equiv := @reln_equiv; - c_equiv_rel := reln_equiv_equivalence; compose := @compose_relns; + c_identity := @idn; +}. + +#[export] Instance RelC : CategoryCoherence Rel := { + c_equiv_rel := reln_equiv_equivalence; compose_compat := @compose_relns_compat; assoc := @compose_relns_assoc; - c_identity := @idn; left_unit := @compose_idn_l; right_unit := @compose_idn_r; }. @@ -322,13 +324,18 @@ Qed. }. #[export, program] Instance RelMonoidal : MonoidalCategory Rel | 0 := { - tensor := ProductRelation; + obj_tensor := prod; + mor_tensor := @reln_prod; mon_I := ⊤; associator := @RelAssociator; left_unitor := @RelLeftUnitor; right_unitor := @RelRightUnitor; }. +#[export, program] Instance RelMonoidalC + : MonoidalCategoryCoherence RelMonoidal := {}. + + #[export, program] Instance RelBraidingIsomorphism {A B} : Isomorphism (A * B) (B * A) := { forward := fun ab ba' => match ab with (a, b) => match ba' with (b', a') => @@ -353,22 +360,6 @@ Qed. SymmetricMonoidalCategory RelBraidedMonoidal | 0 := { }. -#[export, program] Instance RelCompactClosed - : CompactClosedCategory RelSymmetricMonoidal := { - dual := fun A => A; - unit := @reln_unit; - counit := fun A => flip reln_unit; -}. - -#[export, program] Instance RelDagger : DaggerCategory Rel := { - adjoint := fun A B => @flip A B Prop; -}. - -#[export, program] Instance RelDaggerMonoidal - : DaggerMonoidalCategory RelDagger RelMonoidal := { -}. - - @@ -432,13 +423,18 @@ Qed. }. #[export, program] Instance RelSumMonoidal : MonoidalCategory Rel | 10 := { - tensor := SumRelation; + obj_tensor := sum; + mor_tensor := @reln_sum; mon_I := ⊥; associator := @RelSumAssociator; left_unitor := @RelSumLeftUnitor; right_unitor := @RelSumRightUnitor; }. +#[export, program] Instance RelSumMonoidalC + : MonoidalCategoryCoherence RelSumMonoidal := {}. + + #[export, program] Instance RelSumBraidingIsomorphism {A B} : Isomorphism (A + B) (B + A) := { forward := fun ab ba' => @@ -469,27 +465,6 @@ Qed. SymmetricMonoidalCategory RelSumBraidedMonoidal | 10 := { }. -Lemma not_RelSumCompactClosed : - @CompactClosedCategory Type Rel RelSumMonoidal - RelSumBraidedMonoidal RelSumSymmetricMonoidal -> False. -Proof. - intros []. - pose proof (triangle_1 ⊤) as Hfalse. - simpl in Hfalse. - specialize (Hfalse tt tt). - specialize (proj2 Hfalse eq_refl). - clear Hfalse. - solve_relations. -Qed. - -#[export, program] Instance RelSumDagger : DaggerCategory Rel | 10 := { - adjoint := fun A B => @flip A B Prop; -}. - -#[export, program] Instance RelSumDaggerMonoidal - : DaggerMonoidalCategory RelSumDagger RelSumMonoidal | 10 := { -}. - #[export, program] Instance RelLeftDistributivityIsomorphism (A B M : Type) : @Isomorphism Type Rel (A * (B + M)) ((A * B) + (A * M)) := { @@ -552,142 +527,15 @@ Qed. reverse := fun bot => match bot with end; }. -#[export, program] Instance RelPreDistr : - PreDistributiveBimonoidalCategory RelSumSymmetricMonoidal - RelMonoidal := { - left_distr_iso := RelLeftDistributivityIsomorphism; - right_distr_iso := RelRightDistributivityIsomorphism; - left_distr_natural := @rel_left_distributivity_isomorphism_natural; - right_distr_natural := @rel_right_distributivity_isomorphism_natural; -}. - -#[export, program] Instance RelDistrBimonoidal : - DistributiveBimonoidalCategory RelPreDistr := { - left_absorbtion_iso := RelLeftAbsorbtion; - right_absorbtion_iso := RelRightAbsorbtion; -}. -Obligation Tactic := idtac. - -#[export, program] Instance RelSemiCoherent : - SemiCoherent_DistributiveBimonoidalCategory RelDistrBimonoidal := { - (* condition_I (A B C : Type) := ltac:(abstract(solve_relations)); - condition_III (A B C : Type) := ltac:(abstract(solve_relations)); - condition_IV (A B C D : Type) := ltac:(abstract(solve_relations)); - condition_V (A B C D : Type) := ltac:(abstract(solve_relations)); - condition_VI (A B C D : Type) := ltac:(abstract(solve_relations)); - condition_VII (A B C D : Type) := ltac:(abstract(solve_relations)); - condition_VIII (A B C D : Type) := ltac:(abstract(solve_relations)); - condition_IX (A B C D : Type) := ltac:(abstract(solve_relations)); - condition_X := ltac:(abstract(solve_relations)); - condition_XI (A B : Type) := ltac:(abstract(solve_relations)); - condition_XII (A B : Type) := ltac:(abstract(solve_relations)); - condition_XIII := ltac:(abstract(solve_relations)); - condition_XIV := ltac:(abstract(solve_relations)); - condition_XVI (A B : Type) := ltac:(abstract(solve_relations)); - condition_XVII (A B : Type) := ltac:(abstract(solve_relations)); - condition_XVIII (A B : Type) := ltac:(abstract(solve_relations)); - condition_XIX (A B : Type) := ltac:(abstract(solve_relations)); - condition_XX (A B : Type) := ltac:(abstract(solve_relations)); - condition_XXI (A B : Type) := ltac:(abstract(solve_relations)); - condition_XXII (A B : Type) := ltac:(abstract(solve_relations)); - condition_XXIII (A B : Type) := ltac:(abstract(solve_relations)); - condition_XXIV (A B : Type) := ltac:(abstract(solve_relations)); *) -}. -Next Obligation. hnf; solve_relations. Qed. -Next Obligation. hnf; solve_relations. Qed. -Next Obligation. hnf; solve_relations. Qed. -Next Obligation. hnf; solve_relations. Qed. -Next Obligation. hnf; solve_relations. Qed. -Next Obligation. hnf; solve_relations. Qed. -Next Obligation. hnf; solve_relations. Qed. -Next Obligation. hnf; solve_relations. Qed. -Next Obligation. hnf; solve_relations. Qed. -Next Obligation. hnf; solve_relations. Qed. -Next Obligation. hnf; solve_relations. Qed. -Next Obligation. hnf; solve_relations. Qed. -Next Obligation. hnf; solve_relations. Qed. -Next Obligation. hnf; solve_relations. Qed. -Next Obligation. hnf; solve_relations. Qed. -Next Obligation. hnf; solve_relations. Qed. -Next Obligation. hnf; solve_relations. Qed. -Next Obligation. hnf; solve_relations. Qed. -Next Obligation. hnf; solve_relations. Qed. -Next Obligation. hnf; solve_relations. Qed. -Next Obligation. hnf; solve_relations. Qed. -Next Obligation. hnf; solve_relations. Qed. +Obligation Tactic := try (hnf; solve_relations). -Obligation Tactic := (hnf; solve_relations). +Require Import CategoryConstructions. -#[export, program] Instance RelSemiCoherentBraided : - SemiCoherent_BraidedDistributiveBimonoidalCategory RelDistrBimonoidal RelBraidedMonoidal := { - (* condition_II (A B C : Type) := ltac:(abstract(solve_relations)); - condition_XV (A : Type) := ltac:(abstract(solve_relations)); *) -}. +Section RelProperties. Local Open Scope Cat_scope. -Local Open Scope Obj_scope. -Generalizable Variable C. - -Set Universe Polymorphism. - - -Class CategoryProduct `{cC : Category C} (A B : C) (AB : C) := { - p_A : AB ~> A; - p_B : AB ~> B; - prod_mor : - forall (Q : C) (fA : Q ~> A) (fB : Q ~> B), Q ~> AB; - prod_mor_prop: - forall (Q : C) (fA : Q ~> A) (fB : Q ~> B), - (prod_mor Q fA fB) ∘ p_A ≃ fA /\ - (prod_mor Q fA fB) ∘ p_B ≃ fB; - prod_mor_unique : - forall (Q : C) (fA : Q ~> A) (fB : Q ~> B) - (fAB' : Q ~> AB), - fA ≃ fAB' ∘ p_A -> fB ≃ fAB' ∘ p_B -> - fAB' ≃ prod_mor Q fA fB -}. - -Arguments CategoryProduct {_} {_}%Cat (_ _ _)%Obj. -(* Local Notation "'@' AB" := (AB.(prod_AB)) (at level 8) : Cat_scope. *) - - - -Lemma prod_mor_unique' `{cC : Category C} {A B AB : C} - (HAB : CategoryProduct A B AB) {Q} (fA : Q ~> A) (fB : Q ~> B) : - forall (fAB fAB' : Q ~> AB), - fAB ∘ p_A ≃ fA -> fAB ∘ p_B ≃ fB -> - fAB' ∘ p_A ≃ fA -> fAB' ∘ p_B ≃ fB -> - fAB ≃ fAB'. -Proof. - intros. - rewrite (prod_mor_unique Q fA fB) by easy. - symmetry; - rewrite (prod_mor_unique Q fA fB) by easy. - easy. -Qed. -Obligation Tactic := idtac. -Program Definition category_product_unique `{cC : Category C} (A B : C) : - forall {AB AB'} (HAB : CategoryProduct A B AB) - (HAB' : CategoryProduct A B AB'), Isomorphism AB AB' := - fun AB AB' HAB HAB' => - {| - forward := HAB'.(prod_mor) AB p_A p_B; - reverse := HAB.(prod_mor) AB' p_A p_B; - |}. -Obligations. -Next Obligation. - split; eapply (prod_mor_unique' _ p_A p_B); rewrite 1?assoc. - 1,5: rewrite 2!(proj1 (prod_mor_prop _ _ _)); easy. - 1,4: rewrite 2!(proj2 (prod_mor_prop _ _ _)); easy. - all: rewrite left_unit; easy. -Qed. - -Class CartesianMonoidalCategory `(mC : MonoidalCategory C) := { - tensor_cartesian : forall (A B : C), - CategoryProduct A B (A × B)%Mon; -}. #[export, program] Instance RelCartesianMonoidalCategory : CartesianMonoidalCategory RelSumMonoidal := { @@ -701,10 +549,6 @@ Class CartesianMonoidalCategory `(mC : MonoidalCategory C) := { end |} }. -Obligations. -Next Obligation. - solve_relations. -Qed. Next Obligation. __setup_solve_relations. intros A B Q fA fB fAB'. @@ -714,98 +558,6 @@ Next Obligation. solve_relations. Qed. -Module BigPredProd. -Class CategoryBigProd `{cC : Category C} (J : C -> Prop) (prod_J : C) := { - p_i : forall i, J i -> prod_J ~> i; - big_prod_mor : - forall (Q : C) (fQ_ : forall i, J i -> Q ~> i), Q ~> prod_J; - big_prod_mor_prop: - forall (Q : C) (fQ_ : forall i, J i -> Q ~> i), - forall i (Hi : J i), - (big_prod_mor Q fQ_) ∘ p_i i Hi ≃ fQ_ i Hi; - big_prod_mor_unique : - forall (Q : C) (fQ_ : forall i, J i -> Q ~> i) - (fAB' : Q ~> prod_J), - (forall i (Hi : J i), fQ_ i Hi ≃ fAB' ∘ p_i i Hi) -> - fAB' ≃ big_prod_mor Q fQ_ -}. - -Lemma big_prod_mor_unique' `{cC : Category C} {J} {pJ : C} - (HJ : CategoryBigProd J pJ) {Q} (fQ_ : forall i, J i -> Q ~> i) : - forall (fAB fAB' : Q ~> pJ), - (forall i (Hi : J i), fAB ∘ p_i i Hi ≃ fQ_ i Hi) -> - (forall i (Hi : J i), fAB' ∘ p_i i Hi ≃ fQ_ i Hi) -> - fAB ≃ fAB'. -Proof. - intros. - rewrite (big_prod_mor_unique Q fQ_) by easy. - symmetry; - rewrite (big_prod_mor_unique Q fQ_) by easy. - easy. -Qed. - -Program Definition category_big_prod_unique `{cC : Category C} J : - forall {pJ pJ'} (HJ : CategoryBigProd J pJ) (HJ' : CategoryBigProd J pJ'), - Isomorphism pJ pJ' := - fun pJ pJ' HJ HJ' => - {| - forward := HJ'.(big_prod_mor) pJ p_i; - reverse := HJ.(big_prod_mor) pJ' p_i; - |}. -Obligations. -Next Obligation. - split; eapply (big_prod_mor_unique' _ p_i); rewrite 1?assoc. - 1,3: intros i Hi; rewrite assoc, 2!(big_prod_mor_prop _ _); easy. - all: intros; rewrite left_unit; easy. -Qed. -End BigPredProd. - - -Class CategoryBigProd `{cC : Category C} {J : Type} - (obj : J -> C) (prod_J : C) := { - p_i : forall i, prod_J ~> obj i; - big_prod_mor : - forall (Q : C) (fQ_ : forall i, Q ~> obj i), Q ~> prod_J; - big_prod_mor_prop: - forall (Q : C) (fQ_ : forall i, Q ~> obj i), - forall i, - (big_prod_mor Q fQ_) ∘ p_i i ≃ fQ_ i; - big_prod_mor_unique : - forall (Q : C) (fQ_ : forall i, Q ~> obj i) - (fAB' : Q ~> prod_J), - (forall i, fQ_ i ≃ fAB' ∘ p_i i) -> - fAB' ≃ big_prod_mor Q fQ_ -}. - -Lemma big_prod_mor_unique' `{cC : Category C} {J} {obj : J -> C} {pJ : C} - (HJ : CategoryBigProd obj pJ) {Q} (fQ_ : forall i, Q ~> obj i) : - forall (fAB fAB' : Q ~> pJ), - (forall i, fAB ∘ p_i i ≃ fQ_ i) -> - (forall i, fAB' ∘ p_i i ≃ fQ_ i) -> - fAB ≃ fAB'. -Proof. - intros. - rewrite (big_prod_mor_unique Q fQ_) by easy. - symmetry; - rewrite (big_prod_mor_unique Q fQ_) by easy. - easy. -Qed. - -Program Definition category_big_prod_unique `{cC : Category C} {J} {obj : J->C} : - forall {pJ pJ'} (HJ : CategoryBigProd obj pJ) (HJ' : CategoryBigProd obj pJ'), - Isomorphism pJ pJ' := - fun pJ pJ' HJ HJ' => - {| - forward := HJ'.(big_prod_mor) pJ p_i; - reverse := HJ.(big_prod_mor) pJ' p_i; - |}. -Obligations. -Next Obligation. - split; eapply (big_prod_mor_unique' _ p_i); rewrite 1?assoc. - 1,3: intros i; rewrite assoc, 2!(big_prod_mor_prop _ _); easy. - all: intros; rewrite left_unit; easy. -Qed. - Inductive big_sum (J : Type -> Type) := | in_i i (a : J i) : big_sum J. @@ -813,23 +565,16 @@ Inductive big_sum (J : Type -> Type) := Arguments in_i {_} _ _. (* Set Printing Universes. *) -Definition my_f_equal [A B : Type] (f : A -> B) [x y : A] (H : x = y) : - f x = f y := - match H in (_ = y') return (f x = f y') with eq_refl => eq_refl end. - Obligation Tactic := Tactics.program_simpl. #[export, program] Instance RelBigProd (J : Type -> Type) : CategoryBigProd J (big_sum J) := { p_i := fun i => fun somei a => match somei with | in_i i' a' => exists (H: i = i'), _ - (* let Hi := my_f_equal J H in - match Hi in (_ = Ji') return Ji' with eq_refl => a end = a' *) end; big_prod_mor := fun Q fQ_ => fun q somei => match somei with | in_i i a => fQ_ i q a end - }. Next Obligation. exact (a = a'). @@ -850,7 +595,6 @@ Next Obligation. Qed. Next Obligation. __setup_solve_relations. - intros q [i a]. rewrite H. split. @@ -863,362 +607,4 @@ Next Obligation. easy. Qed. - -(* Section TransitiveClosure. - -Local Open Scope Rig_scope. - -(* Definition encup {A A' B : Type} (R : reln (A + B) (A' + B)) : reln (A + ⊤) _ (* (A' + ⊤) *) := - (id_ A ⊞ reln_unit) ∘ (α'_ A, B, B ∘ (R ⊞ id_ B) ∘ associator ∘ (id_ A ⊞ unit †). *) - - - -(* Definition dec_and_R {A : Type} (R : reln A A) : reln (A * nat) (A * nat) := - fun an bm => - let (a, n) := an in let (b, m) := bm in - R a b /\ (m + 1 = n)%nat. *) - -(* Definition dec : reln nat nat := - fun i j => (i = j + 1)%nat. - - -Definition if_0 : reln nat (⊤ + nat) := - fun i j => match i, j with - | O, inl _ => True - | S n', inr m => m = *) - -Definition dec_if0 : reln nat (⊤ + nat) := - fun i j => match i, j with - | O, inl _ => True - | S n', inr m => n' = m - | _, _ => False - end. - -(* TODO: generalize? *) -Definition REPNZ {A : Type} - : reln (A * nat) (A + (A * nat)) := - (id_ A ⊠ dec_if0) ∘ δ_ A, ⊤, nat ∘ (ρ_ A ⊞ id_ (A * nat)). - - -(* Note: this is in fact a property of coproducts *) -Definition join_plus {B : Type} : reln (B + B) B := - fun b_or_b c => match b_or_b with - | inl b | inr b => b = c - end. - -(* Note: I _believe_ we can get this with (categorical) sums and products? *) -Definition prod_to_sum {A B : Type} : reln (A*B) (A + B) := - fun ab a'_or_b' => match ab, a'_or_b' with - | (a, _), inl a' => a = a' - | (_, b), inr b' => b = b' - end. - -Definition repn_top {A : Type} (R : reln A A) - : reln (A * nat + A * nat) (A + (A * nat)) := - join_plus ∘ REPNZ ∘ (id_ A ⊞ (R ⊠ id_ _)). - -(* Actually a property of cartesian category: *) -Definition proj_l {A B : Type} : reln (A * B) A := - fun ab a' => let (a, _) := ab in a = a'. - -(* Note: quite suspect! Maybe we get it bc we have all trivial morphisms? *) -Definition only_l {A B : Type} : reln (A + B) A := - fun a_or_b a' => match a_or_b with - | inl a => a = a' - | inr b => False - end. - -(* Q: Given R : A + B ~> A' + B, can we form Rconn : A ~> A'? *) - -Definition REPN {A : Type} (R : reln A A) : reln (A * nat) _ := - ρ_ _ ^-1 ∘ (id_ _ ⊠ reln_unit) ∘ α_ _, _, _ - ∘ ((prod_to_sum ∘ repn_top R) ⊠ id_ _) ∘ δ#_ _, _, _ - ∘ (proj_l ⊞ flip reln_unit) ∘ only_l. - -Fixpoint compose_n {C} {cC : Category C} (n : nat) - {A : C} (f : A ~> A) : A ~> A := - match n with - | O => id_ A - | S n' => compose_n n' f ∘ f - end. - -Theorem REPN_correct {A : Type} (R : reln A A) : - REPN R ≃ fun an a' => let (a, n) := an in (compose_n n R) a a'. -Proof. - intros [a n]; revert a. - induction n. - - intros a b. simpl. - unfold REPN. - rewrite ?compose_relns_assoc. - (* repn_top, only_l, proj_l, - prod_to_sum, join_plus, REPNZ, dec_if0. *) - - split. - + __setup_solve_relations. - __process_solve_relations. - unfold prod_to_sum, repn_top, proj_l, join_plus, REPNZ, only_l in *. - revert H2; - __setup_solve_relations. - solve_relations. - - solve_relations. - solve_relations. - __solve_relations_mid_step. - __process_solve_relations_step; __process_solve_relations_cleanup_vars; [|easy]. - __process_solve_relations_step. - __process_solve_relations_step; __process_solve_relations_cleanup_vars; [|easy]. - - - - solve_relations. - - - - - -Definition anynat : reln ⊤ nat := - fun t n => True. - - - -Definition anynat_alt : reln ⊤ nat := - reln_unit ∘ proj_l. - -Lemma anynat_alt_correct : anynat ≃ anynat_alt. -Proof. - unfold anynat, anynat_alt. - solve_relations. -Qed. - - -Definition if_0 (A : Type) : reln (A * nat) (A + (A * nat)) := - fun an b_or_bm => match an, b_or_bm with - | (a, O), inl b => a = b - | (a, S n'), inr (b, m) => a = b /\ S n' = m - | _, _ => False - end. - -Definition either (B : Type) : reln (B * B) B := - fun bb' c => let (b, b') := bb' in - b = c \/ b' = c. - -Definition transitive_closure_setup {A : Type} (R : reln A A) - : reln A ((A * nat) * (A * nat)) := - right_unitor ^-1 ∘ (id_ A ⊗ anynat) ∘ right_unitor^-1 - ∘ (id_ _ ⊗ reln_unit) ∘ associator ^-1 - ∘ (either _ ⊗ id_ _). - -Definition transitive_closure {A : Type} (R : reln A A) : reln A _ := - transitive_closure_setup R ∘ (R ⊗ dec ⊗ id_ _) ∘ - (if_0 A ⊗ id_ _) ∘ RelRightDistributivityIsomorphism A (A*nat) (A*nat) - ∘ (proj_l ⊞ (flip reln_unit ∘ @reln_unit ⊥ ∘ proj_l))%Rig - ∘ RelSumMonoidal.(right_unitor).(forward). - -Require Import Lia. - -Ltac __solve_relations_mid_step ::= - repeat (intros ?); intros; subst; auto; try lia. - -Lemma test_trans_clo : - (transitive_closure (fun i j => (i + 1 = j)%nat)) 2 3. -Proof. - unfold transitive_closure, transitive_closure_setup, either, - anynat, dec, proj_l, if_0. - __setup_solve_relations. - (* straight through, so... *) - exists (inl 3); split; [|easy]. - eexists (inl (_, (_, _))); split; [|reflexivity]. - eexists ((inl 3, (_, _))); split; [|split; reflexivity]. - eexists ((_, 0), (_, _)); split; [|split]. - eexists ((_, _), (_, _)); split; [|split; [split|]; reflexivity]. - eexists ((_, _), (_, _), (_, _)). - split; [|split; [|reflexivity]]. - eexists ((_, _), ((_, _), (_, _))). - split; [|split; [|split]; reflexivity]. - eexists ((_, _), tt). - split; [|split; reflexivity]. - eexists (_, _). - split; [|reflexivity]. - eexists (_, tt). - split; [|split]; reflexivity. - left. - all: reflexivity. - Unshelve. all: constructor. -Qed. - -Lemma test_trans_clo_2 : - (transitive_closure (fun i j => (i + 1 = j)%nat)) 2 1 -> False. -Proof. - unfold transitive_closure, transitive_closure_setup, either, - anynat, dec, proj_l, if_0. - __setup_solve_relations. - intros Hbad. - __process_solve_relations. - destruct H1. - destruct n7; [|easy]. - inversion H1. - lia. - destruct n7; [|easy]. - - __process_solve_relations_cleanup_vars. - (* straight through, so... *) - eexists (inl _); split; [|easy]. - eexists (inl (_, (_, _))); split; [|reflexivity]. - eexists ((inl 3, (_, _))); split; [|split; reflexivity]. - eexists ((_, 0), (_, _)); split; [|split]. - eexists ((_, _), (_, _)); split; [|split; [split|]; reflexivity]. - eexists ((_, _), (_, _), (_, _)). - split; [|split; [|reflexivity]]. - eexists ((_, _), ((_, _), (_, _))). - split; [|split; [|split]; reflexivity]. - eexists ((_, _), tt). - split; [|split; reflexivity]. - eexists (_, _). - split; [|reflexivity]. - eexists (_, tt). - split; [|split]; reflexivity. - left. - all: reflexivity. - Unshelve. all: constructor. -Qed. - -Lemma test_trans_clo_2 : - (transitive_closure (fun i j => (i + 1 = j)%nat)) 2 4. -Proof. - unfold transitive_closure, transitive_closure_setup, either, anynat, dec. - __setup_solve_relations. - do 2 (eexists ((_, _), (_, _)); split). - 2: split; [split|]; reflexivity. - eexists ((_, _), (_, _), (_, _)). - split; [|split; [|reflexivity]]. - eexists ((_, _), ((_, _), (_, _))). - split; [|split; [|split]; reflexivity]. - eexists ((_, _), tt). - split; [|split; reflexivity]. - eexists (_, _). - split; [|reflexivity]. - eexists (_, tt). - split; [|split]; reflexivity. - 2: { - eexists ((_, _), ((_, _), (_, _))). - split; [|split; [|split]; reflexivity]. - eexists ((_, _), tt). - split; [|split; reflexivity]. - eexists (_, _). - split; [|reflexivity]. - eexists (_, tt). - split; [|split]; reflexivity. - } - left; reflexivity. - split; [right|]; reflexivity. - Unshelve. - all: constructor. -Qed. - -Lemma test_trans_clo_3 : - (transitive_closure (fun i j => (i + 1 = j)%nat)) 2 1. -Proof. - unfold transitive_closure, transitive_closure_setup, either, anynat, dec. - __setup_solve_relations. - do 2 (eexists ((_, _), (_, _)); split). - 2: split; [split|]; reflexivity. - eexists ((_, _), (_, _), (_, _)). - split; [|split; [|reflexivity]]. - eexists ((_, _), ((_, _), (_, _))). - split; [|split; [|split]; reflexivity]. - eexists ((_, _), tt). - split; [|split; reflexivity]. - eexists (_, _). - split; [|reflexivity]. - eexists (_, tt). - split; [|split]; reflexivity. - 2: { - eexists ((_, _), ((_, _), (_, _))). - split; [|split; [|split]; reflexivity]. - eexists ((_, _), tt). - split; [|split; reflexivity]. - eexists (_, _). - split; [|reflexivity]. - eexists (_, tt). - split; [|split]; reflexivity. - } - left; reflexivity. - split; [right|]; reflexivity. - Unshelve. - all: constructor. -Qed. - - -Definition trans_join (A : Type) : reln (A*A) A := - fun ii' j => let (i, i') := ii' in - i = j \/ i' = j. - -Definition trans_split (A : Type) : reln A (A*A) := - fun i jj' => let (j, j') := jj' in i = j /\ i = j'. - -Definition trans_clo (A : Type) (R : reln A A) := - (@RelRightUnitor A) ^-1 ∘ (RelMonoidal.(tensor) @@ (id_ A), reln_unit) - ∘ (associator ^-1) ∘ (RelMonoidal.(tensor) @@ trans_join A, id_ A) - ∘ (RelMonoidal.(tensor) @@ R, id_ A) ∘ (RelMonoidal.(tensor) @@ trans_split A, id_ A) - ∘ associator ∘ (RelMonoidal.(tensor) @@ id_ A, flip reln_unit) - ∘ RelMonoidal.(right_unitor).(forward). - -Lemma trans_clo_test : (trans_clo nat (fun i j => (i + 1 = j)%nat)) 2 3. -Proof. - unfold trans_clo, trans_join, trans_split. - simpl. - __setup_solve_relations. - exists (3, tt). - split; [|easy]. - eexists (_, (_, _)); split; [|split; reflexivity]. - eexists ((_, _), _); split; [|split; [|split]; reflexivity]. - eexists (_, _); split; [|split; [split|]; reflexivity]. - eexists (2, 3); split; [|easy]. - eexists ((_, _), _); split; [|split; [left|]; reflexivity]. - eexists (_, (_, _)); split; [|split; [|split]; reflexivity]. - eexists (_, tt). - easy. -Qed. - -Lemma trans_clo_test_2 : (trans_clo nat (fun i j => (i + 1 = j)%nat)) 2 4. -Proof. - unfold trans_clo, trans_join, trans_split. - simpl. - unfold compose_relns. - __setup_solve_relations. - exists (4, tt). - split; [|easy]. - eexists (_, (_, _)); split; [|split; reflexivity]. - eexists ((_, _), _); split; [|split; [|split]; reflexivity]. - eexists (_, _); split; [|split; [split|]; reflexivity]. - eexists (3, 4); split; [|easy]. - eexists ((_, _), _); split; [|split; [left|]; reflexivity]. - eexists (_, (_, _)); split; [|split; [|split]; reflexivity]. - eexists (_, tt). - easy. -Qed. - -End TransitiveClosure. - - - -From ViCaR Require Import CategoryAutomation. - - -Lemma test {A} : forall (f : reln A (A★)%Cat) (g : reln (A★)%Cat A), - ((unit (A:=A) \o (braiding (H:=RelMonoidal) _ A).(forward) \o (f \* g) \o counit) † ~ - (unit (A:=A) \o g \* f \o (braiding (H:=RelMonoidal) A _).(forward) \o counit) †)%Cat. -Proof. - simpl. - intros f g. - to_Cat. - Admitted. - - -Lemma test2 {A B} : forall (f g : reln A B) (h : reln ⊤ A) (i : reln (⊤ * (A*A)) _), - (RelMonoidal.(left_unitor).(forward) \o (idn \o f \* g) ~ i \o f \* g)%Cat. -Proof. - intros f g h i. - simpl. - to_Cat. - Admitted. *) +End RelProperties. \ No newline at end of file diff --git a/examples/TypesExample.v b/examples/TypesExample.v index 5f7d40f..ecf9753 100644 --- a/examples/TypesExample.v +++ b/examples/TypesExample.v @@ -1,7 +1,6 @@ Require Import Setoid. Require Import Logic. From ViCaR Require Export CategoryTypeclass. -From ViCaR Require Import RigCategory. Local Notation "f '\o' g" := (fun A => g (f A)) (at level 45, left associativity). @@ -180,6 +179,9 @@ Ltac __process_solve_functions_step := | |- _ => (* idtac "none worked"; *) eexists end end + | |- (_, _) = (_, _) => f_equal + | |- inl _ = inl _ => f_equal + | |- inr _ = inr _ => f_equal end. Ltac __process_solve_functions := @@ -246,11 +248,14 @@ Qed. #[export] Instance Typ : Category Type := { morphism := fun A B => A -> B; c_equiv := @fun_equiv; - c_equiv_rel := fun_equiv_equivalence; compose := fun A B C f g => f \o g; + c_identity := @id; +}. + +#[export] Instance TypC : CategoryCoherence Typ := { + c_equiv_rel := fun_equiv_equivalence; compose_compat := @compose_funs_compat; assoc := @compose_funs_assoc; - c_identity := @id; left_unit := @compose_id_l; right_unit := @compose_id_r; }. @@ -312,23 +317,23 @@ Qed. }. #[export] Instance TypMonoidal : MonoidalCategory Typ := { - tensor := ProductFunction; - mon_I := ⊤; + obj_tensor := prod; + mor_tensor := @fun_prod; + mon_I := ⊤; associator := @TypAssociator; left_unitor := @TypLeftUnitor; right_unitor := @TypRightUnitor; - associator_cohere := ltac:(abstract(solve_functions)); - left_unitor_cohere := ltac:(abstract(solve_functions)); - right_unitor_cohere := ltac:(abstract(solve_functions)); - triangle := ltac:(abstract(solve_functions)); - pentagon := ltac:(abstract(solve_functions)); }. -#[export] Instance TypBraidingIsomorphism {A B} : +Obligation Tactic := try (hnf; solve_functions). + +#[export, program] Instance TypMonoidalC + : MonoidalCategoryCoherence TypMonoidal := {}. + +#[export, program] Instance TypBraidingIsomorphism {A B} : Isomorphism (A * B) (B * A) := { forward := fun ab => let (a, b) := ab in (b, a); - reverse := fun ba => let (b, a) := ba in (a, b); - isomorphism_inverse := ltac:(abstract(solve_functions)) + reverse := fun ba => let (b, a) := ba in (a, b) }. #[export] Instance TypBraiding : @@ -338,38 +343,12 @@ Qed. }. #[export] Instance TypBraidedMonoidal : BraidedMonoidalCategory TypMonoidal := { - braiding := TypBraiding; - hexagon_1 := ltac:(abstract(solve_functions)); - hexagon_2 := ltac:(abstract(solve_functions)); -}. - -#[export] Instance TypSymmetricMonoidal : - SymmetricMonoidalCategory TypBraidedMonoidal := { - symmetry := ltac:(abstract(solve_functions)); -}. - -(* #[export] Instance TypCompactClosed : CompactClosedCategory Type := { - dual := fun A => A; - unit := @fun_unit; - counit := fun A => flip fun_unit; - triangle_1 := ltac:(abstract(solve_functions)); - triangle_2 := ltac:(abstract(solve_functions)); -}. *) - -(* #[export] Instance TypDagger : DaggerCategory Type := { - adjoint := fun A B => @flip A B Prop; - adjoint_involutive := ltac:(easy); - adjoint_id := ltac:(easy); - adjoint_contravariant := ltac:(abstract(solve_functions)); - adjoint_compat := ltac:(abstract(solve_functions)); -}. + braiding := TypBraiding; }. +#[export, program] Instance TypBraidedMonoidalC : + BraidedMonoidalCategoryCoherence TypBraidedMonoidal := {}. -#[export] Instance TypDaggerMonoidal : DaggerMonoidalCategory Type := { - dagger_tensor_compat := ltac:(abstract(solve_functions)); - associator_unitary := ltac:(abstract(solve_functions)); - left_unitor_unitary := ltac:(abstract(solve_functions)); - right_unitor_unitary := ltac:(abstract(solve_functions)); -}. *) +#[export, program] Instance TypSymmetricMonoidal : + SymmetricMonoidalCategory TypBraidedMonoidal := {}. @@ -392,9 +371,6 @@ Qed. id_bimap := @sum_id; compose_bimap := @sum_compose; }. -Next Obligation. - intros [a | b]; simpl; [rewrite H | rewrite H0]; easy. -Qed. #[export] Instance TypSumAssociator {A B C : Type} : Isomorphism (A + B + C) (A + (B + C)) := { @@ -435,18 +411,17 @@ Qed. }. #[export] Instance TypSumMonoidal : MonoidalCategory Typ | 10 := { - tensor := SumFunction; - mon_I := ⊥; + obj_tensor := sum; + mor_tensor := @fun_sum; + mon_I := ⊥; associator := @TypSumAssociator; left_unitor := @TypSumLeftUnitor; right_unitor := @TypSumRightUnitor; - associator_cohere := ltac:(abstract(solve_functions)); - left_unitor_cohere := ltac:(abstract(solve_functions)); - right_unitor_cohere := ltac:(abstract(solve_functions)); - triangle := ltac:(abstract(solve_functions)); - pentagon := ltac:(abstract(solve_functions)); }. +#[export, program] Instance TypSumMonoidalC + : MonoidalCategoryCoherence TypSumMonoidal := {}. + #[export] Instance TypSumBraidingIsomorphism {A B} : Isomorphism (A + B) (B + A) := { forward := fun a_b => @@ -471,22 +446,13 @@ Qed. #[export] Instance TypSumBraidedMonoidal : BraidedMonoidalCategory TypSumMonoidal | 10 := { braiding := TypSumBraiding; - hexagon_1 := ltac:(abstract(solve_functions)); - hexagon_2 := ltac:(abstract(solve_functions)); }. -#[export] Instance TypSumSymmetricMonoidal : - SymmetricMonoidalCategory TypSumBraidedMonoidal | 10 := { - symmetry := ltac:(abstract(solve_functions)); -}. +#[export, program] Instance TypSumBraidedMonoidalC + : BraidedMonoidalCategoryCoherence TypSumBraidedMonoidal := {}. -Lemma not_TypSumCompactClosed : - @CompactClosedCategory Type Typ TypSumMonoidal - TypSumBraidedMonoidal TypSumSymmetricMonoidal -> False. -Proof. - intros []. - destruct (counit ⊤ (inl tt)). -Qed. +#[export, program] Instance TypSumSymmetricMonoidal : + SymmetricMonoidalCategory TypSumBraidedMonoidal | 10 := {}. #[export] Instance TypLeftDistributivityIsomorphism (A B M : Type) : @Isomorphism Type Typ (A * (B + M)) ((A * B) + (A * M)) := { @@ -519,7 +485,7 @@ Lemma rel_left_distributivity_isomorphism_natural {A B M A' B' M' : Type} (TypLeftDistributivityIsomorphism A B M ∘ (fun_sum (fun_prod f g) (fun_prod f h)) ≃ fun_prod f (fun_sum g h) - ∘ TypLeftDistributivityIsomorphism A' B' M')%Cat. + ∘ TypLeftDistributivityIsomorphism A' B' M')%Cat. Proof. solve_functions. Qed. @@ -529,7 +495,7 @@ Lemma rel_right_distributivity_isomorphism_natural {A B M A' B' M' : Type} (TypRightDistributivityIsomorphism A B M ∘ (fun_sum (fun_prod f h) (fun_prod g h)) ≃ fun_prod (fun_sum f g) h - ∘ TypRightDistributivityIsomorphism A' B' M')%Cat. + ∘ TypRightDistributivityIsomorphism A' B' M')%Cat. Proof. solve_functions. Qed. @@ -548,215 +514,13 @@ Qed. isomorphism_inverse := ltac:(abstract(solve_functions)); }. -#[export] Instance TypPreDistr : - PreDistributiveBimonoidalCategory TypSumSymmetricMonoidal - TypMonoidal := { - left_distr_iso := TypLeftDistributivityIsomorphism; - right_distr_iso := TypRightDistributivityIsomorphism; - left_distr_natural := @rel_left_distributivity_isomorphism_natural; - right_distr_natural := @rel_right_distributivity_isomorphism_natural; -}. - -#[export] Instance TypDistrBimonoidal : - DistributiveBimonoidalCategory TypPreDistr := { - left_absorbtion_iso := TypLeftAbsorbtion; - right_absorbtion_iso := TypRightAbsorbtion; - left_absorbtion_natural := ltac:(abstract(solve_functions)); - right_absorbtion_natural := ltac:(abstract(solve_functions)); -}. - -#[export] Instance TypSemiCoherent : - SemiCoherent_DistributiveBimonoidalCategory TypDistrBimonoidal := { - cond_I (A B C : Type) := ltac:(abstract(solve_functions)); - cond_III (A B C : Type) := ltac:(abstract(solve_functions)); - cond_IV (A B C D : Type) := ltac:(abstract(solve_functions)); - cond_V (A B C D : Type) := ltac:(abstract(solve_functions)); - cond_VI (A B C D : Type) := ltac:(abstract(solve_functions)); - cond_VII (A B C D : Type) := ltac:(abstract(solve_functions)); - cond_VIII (A B C D : Type) := ltac:(abstract(solve_functions)); - cond_IX (A B C D : Type) := ltac:(abstract(solve_functions)); - cond_X := ltac:(abstract(solve_functions)); - cond_XI (A B : Type) := ltac:(abstract(solve_functions)); - cond_XII (A B : Type) := ltac:(abstract(solve_functions)); - cond_XIII := ltac:(abstract(solve_functions)); - cond_XIV := ltac:(abstract(solve_functions)); - cond_XVI (A B : Type) := ltac:(abstract(solve_functions)); - cond_XVII (A B : Type) := ltac:(abstract(solve_functions)); - cond_XVIII (A B : Type) := ltac:(abstract(solve_functions)); - cond_XIX (A B : Type) := ltac:(abstract(solve_functions)); - cond_XX (A B : Type) := ltac:(abstract(solve_functions)); - cond_XXI (A B : Type) := ltac:(abstract(solve_functions)); - cond_XXII (A B : Type) := ltac:(abstract(solve_functions)); - cond_XXIII (A B : Type) := ltac:(abstract(solve_functions)); - cond_XXIV (A B : Type) := ltac:(abstract(solve_functions)); -(* - condition_I (A B C : Type) := ltac:(abstract(solve_functions)); - condition_III (A B C : Type) := ltac:(abstract(solve_functions)); - condition_IV (A B C D : Type) := ltac:(abstract(solve_functions)); - condition_V (A B C D : Type) := ltac:(abstract(solve_functions)); - condition_VI (A B C D : Type) := ltac:(abstract(solve_functions)); - condition_VII (A B C D : Type) := ltac:(abstract(solve_functions)); - condition_VIII (A B C D : Type) := ltac:(abstract(solve_functions)); - condition_IX (A B C D : Type) := ltac:(abstract(solve_functions)); - condition_X := ltac:(abstract(solve_functions)); - condition_XI (A B : Type) := ltac:(abstract(solve_functions)); - condition_XII (A B : Type) := ltac:(abstract(solve_functions)); - condition_XIII := ltac:(abstract(solve_functions)); - condition_XIV := ltac:(abstract(solve_functions)); - condition_XVI (A B : Type) := ltac:(abstract(solve_functions)); - condition_XVII (A B : Type) := ltac:(abstract(solve_functions)); - condition_XVIII (A B : Type) := ltac:(abstract(solve_functions)); - condition_XIX (A B : Type) := ltac:(abstract(solve_functions)); - condition_XX (A B : Type) := ltac:(abstract(solve_functions)); - condition_XXI (A B : Type) := ltac:(abstract(solve_functions)); - condition_XXII (A B : Type) := ltac:(abstract(solve_functions)); - condition_XXIII (A B : Type) := ltac:(abstract(solve_functions)); - condition_XXIV (A B : Type) := ltac:(abstract(solve_functions)); -*) -}. - -#[export] Instance TypSemiCoherentBraided : - SemiCoherent_BraidedDistributiveBimonoidalCategory TypDistrBimonoidal TypBraidedMonoidal := { - cond_II (A B C : Type) := ltac:(abstract(solve_functions)); - cond_XV (A : Type) := ltac:(abstract(solve_functions)); -}. - - - - - - +Require Import CategoryConstructions. +Require Import FunctionalExtensionality. +Require Import Eqdep. +Section TypProperties. Local Open Scope Cat_scope. -Generalizable Variable C. - -Set Universe Polymorphism. - - -Class CategoryProduct `{cC : Category C} (A B : C) (AB : C) := { - p_A : AB ~> A; - p_B : AB ~> B; - prod_mor : - forall (Q : C) (fA : Q ~> A) (fB : Q ~> B), Q ~> AB; - prod_mor_prop: - forall (Q : C) (fA : Q ~> A) (fB : Q ~> B), - (prod_mor Q fA fB) ∘ p_A ≃ fA /\ - (prod_mor Q fA fB) ∘ p_B ≃ fB; - prod_mor_unique : - forall (Q : C) (fA : Q ~> A) (fB : Q ~> B) - (fAB' : Q ~> AB), - fA ≃ fAB' ∘ p_A -> fB ≃ fAB' ∘ p_B -> - fAB' ≃ prod_mor Q fA fB -}. - -Arguments CategoryProduct {_} {_}%Cat (_ _ _)%Obj. - -(* Local Notation "'@' AB" := (AB.(prod_AB)) (at level 8) : Cat_scope. *) - - - -Lemma prod_mor_unique' `{cC : Category C} {A B AB : C} - (HAB : CategoryProduct A B AB) {Q} (fA : Q ~> A) (fB : Q ~> B) : - forall (fAB fAB' : Q ~> AB), - fAB ∘ p_A ≃ fA -> fAB ∘ p_B ≃ fB -> - fAB' ∘ p_A ≃ fA -> fAB' ∘ p_B ≃ fB -> - fAB ≃ fAB'. -Proof. - intros. - rewrite (prod_mor_unique Q fA fB) by easy. - symmetry; - rewrite (prod_mor_unique Q fA fB) by easy. - easy. -Qed. - -Program Definition category_product_unique `{cC : Category C} (A B : C) : - forall {AB AB'} (HAB : CategoryProduct A B AB) - (HAB' : CategoryProduct A B AB'), Isomorphism AB AB' := - fun AB AB' HAB HAB' => - {| - forward := HAB'.(prod_mor) AB p_A p_B; - reverse := HAB.(prod_mor) AB' p_A p_B; - |}. -Obligations. -Next Obligation. - split; eapply (prod_mor_unique' _ p_A p_B); rewrite 1?assoc. - 1,5: rewrite 2!(proj1 (prod_mor_prop _ _ _)); easy. - 1,4: rewrite 2!(proj2 (prod_mor_prop _ _ _)); easy. - all: rewrite left_unit; easy. -Qed. - - - -Class CartesianMonoidalCategory `(mC : MonoidalCategory C) := { - tensor_cartesian : forall (A B : C), - CategoryProduct A B (A × B)%Mon; -}. - -#[export, program] Instance TypCartesianMonoidalCategory : - CartesianMonoidalCategory TypMonoidal := { - tensor_cartesian := fun A B => {| - p_A := fun ab => let (a, _) := ab in a; - p_B := fun ab => let (_, b) := ab in b; - prod_mor := fun Q fA fB => - fun q => (fA q, fB q) - |} - }. -Next Obligation. - easy. -Qed. -Next Obligation. - intros q. - rewrite H, H0. - destruct (fAB' q). - easy. -Qed. - -Class CategoryBigProd `{cC : Category C} {J : Type} - (obj : J -> C) (prod_J : C) := { - p_i : forall i, prod_J ~> obj i; - big_prod_mor : - forall (Q : C) (fQ_ : forall i, Q ~> obj i), Q ~> prod_J; - big_prod_mor_prop: - forall (Q : C) (fQ_ : forall i, Q ~> obj i), - forall i, - (big_prod_mor Q fQ_) ∘ p_i i ≃ fQ_ i; - big_prod_mor_unique : - forall (Q : C) (fQ_ : forall i, Q ~> obj i) - (fAB' : Q ~> prod_J), - (forall i, fQ_ i ≃ fAB' ∘ p_i i) -> - fAB' ≃ big_prod_mor Q fQ_ -}. - -Lemma big_prod_mor_unique' `{cC : Category C} {J} {obj : J -> C} {pJ : C} - (HI : CategoryBigProd obj pJ) {Q} (fQ_ : forall i, Q ~> obj i) : - forall (fAB fAB' : Q ~> pJ), - (forall i, fAB ∘ p_i i ≃ fQ_ i) -> - (forall i, fAB' ∘ p_i i ≃ fQ_ i) -> - fAB ≃ fAB'. -Proof. - intros. - rewrite (big_prod_mor_unique Q fQ_) by easy. - symmetry; - rewrite (big_prod_mor_unique Q fQ_) by easy. - easy. -Qed. - -Program Definition category_big_prod_unique `{cC : Category C} {J} {obj : J->C} : - forall {pJ pJ'} (HI : CategoryBigProd obj pJ) (HI' : CategoryBigProd obj pJ'), - Isomorphism pJ pJ' := - fun pJ pJ' HI HI' => - {| - forward := HI'.(big_prod_mor) pJ p_i; - reverse := HI.(big_prod_mor) pJ' p_i; - |}. -Next Obligation. - split; eapply (big_prod_mor_unique' _ p_i); rewrite 1?assoc. - 1,3: intros i; rewrite assoc, 2!(big_prod_mor_prop _ _); easy. - all: intros; rewrite left_unit; easy. -Qed. - -Require Import FunctionalExtensionality. Definition big_prod {J} (obj : J -> Type) := forall (i : J), obj i. @@ -766,82 +530,15 @@ Definition big_prod {J} (obj : J -> Type) := forall (i : J), obj i. big_prod_mor := fun Q fQ_ => fun q => fun i => fQ_ i q; }. Next Obligation. - easy. -Qed. -Next Obligation. - intros q. - unfold big_prod in *. + unfold big_prod. + __setup_solve_functions. + __solve_functions_mid_step. apply functional_extensionality_dep. intros i. symmetry. apply H. Qed. - - -Class CategoryCoproduct `{cC : Category C} (A B : C) (A_B : C) := { - coprod_obj := A_B; - iota_A : A ~> A_B; - iota_B : B ~> A_B; - coprod_mor : - forall (Q : C) (fA : A ~> Q) (fB : B ~> Q), A_B ~> Q; - coprod_mor_prop: - forall (Q : C) (fA : A ~> Q) (fB : B ~> Q), - iota_A ∘ (coprod_mor Q fA fB) ≃ fA /\ - iota_B ∘ (coprod_mor Q fA fB) ≃ fB; - coprod_mor_unique : - forall (Q : C) (fA : A ~> Q) (fB : B ~> Q) - (fAB' : A_B ~> Q), - fA ≃ iota_A ∘ fAB' -> fB ≃ iota_B ∘ fAB' -> - fAB' ≃ coprod_mor Q fA fB -}. - -Lemma coprod_mor_unique' `{cC : Category C} {A B A_B : C} - (HAB : CategoryCoproduct A B A_B) {Q} (fA : A ~> Q) (fB : B ~> Q) : - forall (fAB fAB' : A_B ~> Q), - iota_A ∘ fAB ≃ fA -> iota_B ∘ fAB ≃ fB -> - iota_A ∘ fAB' ≃ fA -> iota_B ∘ fAB' ≃ fB -> - fAB ≃ fAB'. -Proof. - intros. - rewrite (coprod_mor_unique Q fA fB) by easy. - symmetry; - rewrite (coprod_mor_unique Q fA fB) by easy. - easy. -Qed. - -Program Definition category_coproduct_unique `{cC : Category C} (A B : C) : - forall {A_B A_B'} (HA_B : CategoryCoproduct A B A_B) - (HA_B' : CategoryCoproduct A B A_B'), Isomorphism A_B A_B' := - fun A_B A_B' HA_B HA_B' => - {| - forward := HA_B.(coprod_mor) A_B' iota_A iota_B; - reverse := HA_B'.(coprod_mor) A_B iota_A iota_B; - |}. -Next Obligation. - split; eapply (coprod_mor_unique' _ iota_A iota_B); rewrite <- 1?assoc. - 1,5: rewrite 2!(proj1 (coprod_mor_prop _ _ _)); easy. - 1,4: rewrite 2!(proj2 (coprod_mor_prop _ _ _)); easy. - all: rewrite right_unit; easy. -Qed. - -Class CategoryBigCoprod `{cC : Category C} {J : Type} - (obj : J -> C) (coprod_J : C) := { - big_coprod_obj := coprod_J; - iota_i : forall i, obj i ~> coprod_J; - big_coprod_mor : - forall (Q : C) (fQ_ : forall i, obj i ~> Q), coprod_J ~> Q; - big_coprod_mor_prop: - forall (Q : C) (fQ_ : forall i, obj i ~> Q), - forall i, - iota_i i ∘ (big_coprod_mor Q fQ_) ≃ fQ_ i; - big_coprod_mor_unique : - forall (Q : C) (fQ_ : forall i, obj i ~> Q) - (fAB' : coprod_J ~> Q), - (forall i, fQ_ i ≃ iota_i i ∘ fAB') -> - fAB' ≃ big_coprod_mor Q fQ_ -}. - Inductive big_sum {J} (obj : J -> Type) := | in_i i (a : obj i) : big_sum obj. @@ -855,56 +552,28 @@ Arguments in_i {_ _} _ _. end; }. Next Obligation. - easy. -Qed. -Next Obligation. - intros [i ai]. - rewrite H. - easy. + intros J obj Q fQ_ fAB' H [i ai]. + symmetry. + apply H. Qed. -Reserved Notation "A ∏ B" (at level 40). -Reserved Notation "f @∏ g" (at level 40). - -Class Category_with_Products {C} (cC : Category C) := { - cat_prod : C -> C -> C - where "A ∏ B" := (cat_prod A B); - cat_prod_product : forall {A B}, CategoryProduct A B (A ∏ B); - cat_mor_prod : - forall {A A' B B'} (fA : A ~> B) (fB : A' ~> B'), - A ∏ A' ~> B ∏ B' := fun A A' B B' fA fB => - cat_prod_product.(prod_mor) - (A ∏ A') (p_A ∘ fA) (p_B ∘ fB) - where "f @∏ g" := (cat_mor_prod f g); -}. - -Notation "A ∏ B" := (cat_prod A B) (at level 40). -Notation "f @∏ g" := (cat_mor_prod f g) (at level 40). - -#[program, export] Instance Category_with_Products_of_CartesianMonoidalCategory {C cC mC} - (ccc : @CartesianMonoidalCategory C cC mC) : Category_with_Products cC := {| - cat_prod := mC.(tensor).(obj_bimap); - cat_prod_product := tensor_cartesian; -|}. - -Coercion Category_with_Products_of_CartesianMonoidalCategory : - CartesianMonoidalCategory >-> Category_with_Products. - - -Reserved Notation "'λ' g" (at level 20). -Class CategoryExponential `{cC : Category C} {pC : Category_with_Products cC} - (A B A_B : C) := { - exp_obj := A_B; - eval_mor : A_B ∏ B ~> A; - eval_mor_transpose : forall {M : C} (f : M ∏ B ~> A), - M ~> A_B - where "'λ' g" := (eval_mor_transpose g); - eval_mor_transpose_prop : forall (M : C) (f : M ∏ B ~> A), - f ≃ (λ f @∏ id_ B) ∘ eval_mor; - eval_mor_transpose_unique : forall (M : C) (f : M ∏ B ~> A) - (lf' : M ~> A_B), f ≃ (lf' @∏ id_ B) ∘ eval_mor -> λ f ≃ lf'; -}. - +#[export, program] Instance TypCartesianMonoidalCategory : + CartesianMonoidalCategory TypMonoidal := { + tensor_cartesian := fun A B => {| + p_A := fun ab => let (a, _) := ab in a; + p_B := fun ab => let (_, b) := ab in b; + prod_mor := fun Q fA fB => + fun q => (fA q, fB q) + |} + }. +Next Obligation. + intros ** q. + simpl. + rewrite H, H0. + simpl. + destruct (fAB' q). + easy. +Qed. #[program, export] Instance TypExponential {A B : Type} : CategoryExponential A B (B -> A) := { @@ -912,134 +581,13 @@ Class CategoryExponential `{cC : Category C} {pC : Category_with_Products cC} eval_mor_transpose := fun M f => fun m b => f (m, b); }. Next Obligation. - solve_functions. -Qed. -Next Obligation. - intros i. + intros ** i. apply functional_extensionality. intros b. rewrite H. easy. Qed. - -Class Cone {C D} {cC : Category C} {cD : Category D} - (F : Functor cD cC) := { - cone_obj : C; - cone_mor : forall (d : D), cone_obj ~> F d; - cone_mor_prop : forall {d d' : D} (f : d ~> d'), - cone_mor d ∘ F @ f ≃ cone_mor d'; -}. - -Coercion cone_mor : Cone >-> Funclass. - -Reserved Notation "'lim' F" (at level 30). - -Class Limit {C D} {cC : Category C} {cD : Category D} - (F : Functor cD cC) := { - limit_cone : Cone F; - limit_cone_factor : - forall (N : Cone F), N.(cone_obj) ~> limit_cone.(cone_obj); - limit_cone_factor_prop : forall (N : Cone F) (d : D), - limit_cone_factor N ∘ limit_cone d ≃ N d; - limit_cone_factor_unique : forall (N : Cone F) - (f' : N.(cone_obj) ~> limit_cone.(cone_obj)), - (forall (d : D), f' ∘ limit_cone d ≃ N d) -> f' ≃ limit_cone_factor N; -}. - -Lemma limit_cone_factor_unique' {C D} {cC : Category C} {cD : Category D} - {F : Functor cD cC} (L : Limit F) : forall (N : Cone F) - (f f' : N.(cone_obj) ~> L.(limit_cone).(cone_obj)), - (forall (d : D), f ∘ limit_cone d ≃ N d) -> - (forall (d : D), f' ∘ limit_cone d ≃ N d) -> - f ≃ f'. -Proof. - intros. - rewrite limit_cone_factor_unique by easy. - symmetry. - rewrite limit_cone_factor_unique by easy. - easy. -Qed. - -Coercion limit_cone : Limit >-> Cone. - -Notation "'lim' F" := (limit_cone (F:=F)).(cone_obj) (at level 30). - -#[program] Definition limit_unique {C D} {cC : Category C} {cD : Category D} - {F : Functor cD cC} (L L' : Limit F) : L.(cone_obj) <~> L'.(cone_obj) := {| - forward := L'.(limit_cone_factor) L; - reverse := L.(limit_cone_factor) L' - |}. -Next Obligation. - split; (apply limit_cone_factor_unique'; - [| intros; rewrite left_unit; easy]); - intros; rewrite assoc, 2!limit_cone_factor_prop; easy. -Qed. - - -Class Cocone {C D} {cC : Category C} {cD : Category D} - (F : Functor cD cC) := { - cocone_obj : C; - cocone_mor : forall (d : D), F d ~> cocone_obj; - cocone_mor_prop : forall {d d' : D} (f : d ~> d'), - F @ f ∘ cocone_mor d' ≃ cocone_mor d; -}. - -Coercion cocone_mor : Cocone >-> Funclass. - -Reserved Notation "'colim' F" (at level 30). - -Class Colimit {C D} {cC : Category C} {cD : Category D} - (F : Functor cD cC) := { - colimit_cocone : Cocone F; - colimit_cocone_factor : forall (N : Cocone F), - colimit_cocone.(cocone_obj) ~> N.(cocone_obj); - colimit_cocone_factor_prop : forall (N : Cocone F) (d : D), - colimit_cocone d ∘ colimit_cocone_factor N ≃ N d; - colimit_cocone_factor_unique : forall (N : Cocone F) - (f' : colimit_cocone.(cocone_obj) ~> N.(cocone_obj)), - (forall (d : D), colimit_cocone d ∘ f' ≃ N d) -> - f' ≃ colimit_cocone_factor N; -}. - -Lemma colimit_cocone_factor_unique' {C D} {cC : Category C} {cD : Category D} - {F : Functor cD cC} (L : Colimit F) : forall (N : Cocone F) - (f f' : L.(colimit_cocone).(cocone_obj) ~> N.(cocone_obj)), - (forall (d : D), colimit_cocone d ∘ f ≃ N d) -> - (forall (d : D), colimit_cocone d ∘ f' ≃ N d) -> - f ≃ f'. -Proof. - intros. - rewrite colimit_cocone_factor_unique by easy. - symmetry. - rewrite colimit_cocone_factor_unique by easy. - easy. -Qed. - -Coercion colimit_cocone : Colimit >-> Cocone. - -Notation "'colim' F" := (colimit_cocone (F:=F)).(cocone_obj) (at level 30). - -#[program] Definition colimit_unique {C D} {cC : Category C} {cD : Category D} - {F : Functor cD cC} (L L' : Colimit F) : L.(cocone_obj) <~> L'.(cocone_obj) := {| - forward := L.(colimit_cocone_factor) L'; - reverse := L'.(colimit_cocone_factor) L - |}. -Next Obligation. - split; (apply colimit_cocone_factor_unique'; - [| intros; rewrite right_unit; easy]); - intros; rewrite <- assoc, 2!colimit_cocone_factor_prop; easy. -Qed. - - -Class Category_with_Small_Limits {C} (cC : Category C) := { - take_set_limit : forall {D : Set} (cD : Category D) (F : Functor cD cC), Limit F; -}. - -Class Category_with_Limits {C} (cC : Category C) := { - take_limit : forall {D : Type} (cD : Category D) (F : Functor cD cC), Limit F; -}. - Definition typ_limit_obj {D} {cD : Category D} (F : Functor cD Typ) : Type := { s : big_prod F.(obj_map) | forall (d d' : D) (f : d ~> d'), (F @ f) (s d) = s d' }. @@ -1064,23 +612,20 @@ Qed. cone_mor_prop := fun d d' => typ_limit_obj_cone_mor_prop F; }. -Require Import Eqdep. - #[export, program] Instance TypLimit {D} {cD : Category D} (F : Functor cD Typ) : Limit F := { limit_cone := TypLimitCone F; }. Next Obligation. + intros D cD F N X. exists (fun d => N.(cone_mor) d X). intros d d' f. apply N. Defined. Next Obligation. - easy. -Qed. -Next Obligation. - intros c. + intros ** c. unfold TypLimit_obligation_1. + simpl. destruct (f' c) as [s' P'] eqn:e. match goal with |- exist _ ?s ?P = exist _ ?s' ?P' => assert (Hs: s = s') @@ -1089,6 +634,7 @@ Next Obligation. intros d. specialize (H d c). rewrite <- H. + simpl. unfold typ_limit_obj_cone_mor. rewrite e. easy. @@ -1102,6 +648,8 @@ Next Obligation. apply UIP. Qed. +End TypProperties. + diff --git a/examples/TypoidExample.v b/examples/TypoidExample.v index 625e382..dd137ed 100644 --- a/examples/TypoidExample.v +++ b/examples/TypoidExample.v @@ -1,7 +1,7 @@ Require Import Setoid Morphisms. (* Morphisms_Prop may be added? *) Require Import Logic. From ViCaR Require Export CategoryTypeclass. -From ViCaR Require Import RigCategory. + @@ -201,6 +201,28 @@ Ltac __saturate_funs Hf := end end. +Ltac __saturate_s_equiv Hf := + match type of Hf with s_equiv ?x ?y => + let Hsym := constr:(s_equiv_sym x y Hf) in + extend (Hsym); + repeat match goal with + | H : s_equiv ?z x |- _ => + let Hg := constr:(s_equiv_trans z x y H Hf) in + extend Hg; __saturate_s_equiv Hg + | H : s_equiv x ?z |- _ => + let Hg := constr:(s_equiv_trans y x z Hsym H) in + extend Hg; __saturate_s_equiv Hg + | H : s_equiv ?z y |- _ => + let Hg := constr:(s_equiv_trans z y x H Hsym) in + extend Hg; __saturate_s_equiv Hg + | H : s_equiv y ?z |- _ => + let Hg := constr:(s_equiv_trans x y z Hf H) in + extend Hg; __saturate_s_equiv Hg + end +end. + + + Ltac __solve_morphisms_end_forward := repeat match goal with | H : ?f ?i ≡ ?g ?j |- _ => (* match type of f with ?A → ?B => *) @@ -279,12 +301,15 @@ Ltac __solve_morphisms_end_no_constr := try (__solve_morphisms_end_match __solve_morphisms_end_no_constr || idtac (* " no match" *)). +Ltac __solve_morphisms_end_extra := idtac. + Ltac __solve_morphisms_end ::= try eassumption; try (etransitivity; eauto); try reflexivity; try easy; auto; + __solve_morphisms_end_extra; try (__solve_morphisms_end_match __solve_morphisms_end || constructor; solve [__solve_morphisms_end_no_constr] || econstructor; solve [__solve_morphisms_end_no_constr]). @@ -348,7 +373,8 @@ Ltac __process_solve_morphisms_step := | H : _ /\ _ |- _ => destruct H | |- _ <-> _ => split | |- _ /\ _ => split - | H : _ \/ _ |- _ => destruct H; [solve_morphisms | solve_morphisms] + | H : _ \/ _ |- _ => destruct H; [try (left; solve_morphisms) | try (right; solve_morphisms)] + | H : context[exists _, _] |- _ => edestruct H; try clear H; solve_morphisms | |- _ \/ _ => (left; solve [solve_morphisms; auto]) || (right; solve [solve_morphisms; auto]) @@ -369,10 +395,13 @@ Ltac __process_solve_morphisms_step := end end. +Ltac __process_solve_morphisms_extra := idtac. + Ltac __process_solve_morphisms := repeat ( __solve_morphisms_mid_step; - try __process_solve_morphisms_step); + try __process_solve_morphisms_step; + try __process_solve_morphisms_extra); repeat (try __solve_morphisms_mid_step; try __process_solve_morphisms_cleanup_vars). @@ -421,11 +450,14 @@ Qed. #[export] Instance Std : Category Setoid := { morphism := Morphism; c_equiv := @mor_equiv; - c_equiv_rel := @mor_equiv_equivalence; compose := @compose_morphisms; + c_identity := @IdentityMorphism; +}. + +#[export] Instance StdC : CategoryCoherence Std := { + c_equiv_rel := @mor_equiv_equivalence; compose_compat := @compose_morphisms_compat; assoc := @compose_morphisms_assoc; - c_identity := @IdentityMorphism; left_unit := @compose_id_l; right_unit := @compose_id_r; }. @@ -486,7 +518,8 @@ Obligation Tactic := (hnf; solve_morphisms). }. #[export, program] Instance StdMonoidal : MonoidalCategory Std := { - tensor := ProductMorphism; + obj_tensor := ProdSetoid; + mor_tensor := fun _ _ _ _ => mor_prod; mon_I := ⊤; associator := @StdAssociator; left_unitor := @StdLeftUnitor; @@ -498,6 +531,9 @@ Obligation Tactic := (hnf; solve_morphisms). pentagon := ltac:(abstract(solve_morphisms)); *) }. +#[export, program] Instance StdMonoidalC + : MonoidalCategoryCoherence StdMonoidal := {}. + #[export, program] Instance StdBraidingIsomorphism (A B : Setoid) : Isomorphism (A .* B) (B .* A) := { @@ -513,8 +549,10 @@ Obligation Tactic := (hnf; solve_morphisms). #[export, program] Instance StdBraidedMonoidal : BraidedMonoidalCategory StdMonoidal := { braiding := StdBraiding; - (* hexagon_1 := ltac:(abstract(solve_morphisms)); - hexagon_2 := ltac:(abstract(solve_morphisms)); *) +}. + +#[export, program] Instance StdBraidedMonoidalC + : BraidedMonoidalCategoryCoherence StdBraidedMonoidal := { }. #[export, program] Instance StdSymmetricMonoidal : @@ -605,7 +643,9 @@ Qed. }. #[export, program] Instance StdSumMonoidal : MonoidalCategory Std := { - tensor := SumMorphism; + obj_tensor := SumSetoid; + mor_tensor := fun _ _ _ _ => mor_sum; + (* tensor := SumMorphism; *) mon_I := ⊥; associator := @StdSumAssociator; left_unitor := @StdSumLeftUnitor; @@ -617,6 +657,9 @@ Qed. pentagon := ltac:(abstract(solve_morphisms)); *) }. +#[export, program] Instance StdSumMonoidalC : + MonoidalCategoryCoherence StdSumMonoidal := {}. + #[export, program] Instance StdSumBraidingIsomorphism {A B} : Isomorphism (A .+ B) (B .+ A) := { forward := {| base_fn := fun a_b : A.+B => @@ -649,14 +692,6 @@ Qed. (* symmetry := ltac:(abstract(solve_morphisms)); *) }. -Lemma not_StdSumCompactClosed : - @CompactClosedCategory Setoid Std StdSumMonoidal - StdSumBraidedMonoidal StdSumSymmetricMonoidal -> False. -Proof. - intros []. - destruct (counit ⊤ (inl tt)). -Qed. - #[export, program] Instance StdLeftDistributivityIsomorphism (A B M : Setoid) : Isomorphism (A .* (B .+ M)) ((A .* B) .+ (A .* M)) := { forward := {| base_fn := fun abm : A .* (B .+ M) => match abm with @@ -721,97 +756,11 @@ Qed. match bot with end : A .* ⊥ |}; }. -#[export, program] Instance StdPreDistr : - PreDistributiveBimonoidalCategory StdSumSymmetricMonoidal - StdMonoidal := { - left_distr_iso := StdLeftDistributivityIsomorphism; - right_distr_iso := StdRightDistributivityIsomorphism; - left_distr_natural := @rel_left_distributivity_isomorphism_natural; - right_distr_natural := @rel_right_distributivity_isomorphism_natural; -}. - -#[export, program] Instance StdDistrBimonoidal : - DistributiveBimonoidalCategory StdPreDistr := { - left_absorbtion_iso := StdLeftAbsorbtion; - right_absorbtion_iso := StdRightAbsorbtion; -}. - - - -#[export, program] Instance StdSemiCoherent : - SemiCoherent_DistributiveBimonoidalCategory StdDistrBimonoidal. - - - - -#[export, program] Instance StdSemiCoherentBraided : - SemiCoherent_BraidedDistributiveBimonoidalCategory StdDistrBimonoidal StdBraidedMonoidal := { -}. - - - - -Local Open Scope Cat_scope. -Generalizable Variable C. -Set Universe Polymorphism. - - -Class CategoryProduct `{cC : Category C} (A B : C) (AB : C) := { - p_A : AB ~> A; - p_B : AB ~> B; - prod_mor : - forall (Q : C) (fA : Q ~> A) (fB : Q ~> B), Q ~> AB; - prod_mor_prop: - forall (Q : C) (fA : Q ~> A) (fB : Q ~> B), - (prod_mor Q fA fB) ∘ p_A ≃ fA /\ - (prod_mor Q fA fB) ∘ p_B ≃ fB; - prod_mor_unique : - forall (Q : C) (fA : Q ~> A) (fB : Q ~> B) - (fAB' : Q ~> AB), - fA ≃ fAB' ∘ p_A -> fB ≃ fAB' ∘ p_B -> - fAB' ≃ prod_mor Q fA fB -}. - -(* Local Notation "'@' AB" := (AB.(prod_AB)) (at level 8) : Cat_scope. *) - - - -Lemma prod_mor_unique' `{cC : Category C} {A B AB : C} - (HAB : CategoryProduct A B AB) {Q} (fA : Q ~> A) (fB : Q ~> B) : - forall (fAB fAB' : Q ~> AB), - fAB ∘ p_A ≃ fA -> fAB ∘ p_B ≃ fB -> - fAB' ∘ p_A ≃ fA -> fAB' ∘ p_B ≃ fB -> - fAB ≃ fAB'. -Proof. - intros. - rewrite (prod_mor_unique Q fA fB) by easy. - symmetry; - rewrite (prod_mor_unique Q fA fB) by easy. - easy. -Qed. - -Definition category_product_unique `{cC : Category C} (A B : C) : - forall {AB AB'} (HAB : CategoryProduct A B AB) - (HAB' : CategoryProduct A B AB'), Isomorphism AB AB'. - refine (fun AB AB' HAB HAB' => - {| - forward := HAB'.(prod_mor) AB p_A p_B; - reverse := HAB.(prod_mor) AB' p_A p_B; - |}). - split; eapply (prod_mor_unique' _ p_A p_B); rewrite 1?assoc. - 1,5: rewrite 2!(proj1 (prod_mor_prop _ _ _)); easy. - 1,4: rewrite 2!(proj2 (prod_mor_prop _ _ _)); easy. - all: rewrite left_unit; easy. -Qed. - -Class CartesianMonoidalCategory `(mC : MonoidalCategory C) := { - tensor_cartesian : forall (A B : C), - CategoryProduct A B (A × B)%Mon; -}. +Require Import CategoryConstructions. Obligation Tactic := (try (hnf; solve_morphisms)). @@ -832,51 +781,7 @@ Next Obligation. easy. Qed. -Class CategoryBigProd `{cC : Category C} {J : Type} - (obj : J -> C) (prod_J : C) := { - p_i : forall i, prod_J ~> obj i; - big_prod_mor : - forall (Q : C) (fQ_ : forall i, Q ~> obj i), Q ~> prod_J; - big_prod_mor_prop: - forall (Q : C) (fQ_ : forall i, Q ~> obj i), - forall i, - (big_prod_mor Q fQ_) ∘ p_i i ≃ fQ_ i; - big_prod_mor_unique : - forall (Q : C) (fQ_ : forall i, Q ~> obj i) - (fAB' : Q ~> prod_J), - (forall i, fQ_ i ≃ fAB' ∘ p_i i) -> - fAB' ≃ big_prod_mor Q fQ_ -}. - -Lemma big_prod_mor_unique' `{cC : Category C} {J} {obj : J -> C} {pJ : C} - (HI : CategoryBigProd obj pJ) {Q} (fQ_ : forall i, Q ~> obj i) : - forall (fAB fAB' : Q ~> pJ), - (forall i, fAB ∘ p_i i ≃ fQ_ i) -> - (forall i, fAB' ∘ p_i i ≃ fQ_ i) -> - fAB ≃ fAB'. -Proof. - intros. - rewrite (big_prod_mor_unique Q fQ_) by easy. - symmetry; - rewrite (big_prod_mor_unique Q fQ_) by easy. - easy. -Qed. - -Program Definition category_big_prod_unique `{cC : Category C} {J} {obj : J->C} : - forall {pJ pJ'} (HI : CategoryBigProd obj pJ) (HI' : CategoryBigProd obj pJ'), - Isomorphism pJ pJ' := - fun pJ pJ' HI HI' => - {| - forward := HI'.(big_prod_mor) pJ p_i; - reverse := HI.(big_prod_mor) pJ' p_i; - |}. -Next Obligation. - split; eapply (big_prod_mor_unique' _ p_i); rewrite 1?assoc. - 1,3: intros i; rewrite assoc, 2!(big_prod_mor_prop _ _); easy. - all: intros; rewrite left_unit; easy. -Qed. -(* Require Import FunctionalExtensionality. *) #[program] Definition big_setoid_prod {J} (obj : J -> Setoid) := {| base_type := forall (i : J), obj i; @@ -890,70 +795,6 @@ Qed. {| base_fn := fun q => fun i => fQ_ i q |}; }. - -Class CategoryCoproduct `{cC : Category C} (A B : C) (A_B : C) := { - coprod_obj := A_B; - iota_A : A ~> A_B; - iota_B : B ~> A_B; - coprod_mor : - forall (Q : C) (fA : A ~> Q) (fB : B ~> Q), A_B ~> Q; - coprod_mor_prop: - forall (Q : C) (fA : A ~> Q) (fB : B ~> Q), - iota_A ∘ (coprod_mor Q fA fB) ≃ fA /\ - iota_B ∘ (coprod_mor Q fA fB) ≃ fB; - coprod_mor_unique : - forall (Q : C) (fA : A ~> Q) (fB : B ~> Q) - (fAB' : A_B ~> Q), - fA ≃ iota_A ∘ fAB' -> fB ≃ iota_B ∘ fAB' -> - fAB' ≃ coprod_mor Q fA fB -}. - -Lemma coprod_mor_unique' `{cC : Category C} {A B A_B : C} - (HAB : CategoryCoproduct A B A_B) {Q} (fA : A ~> Q) (fB : B ~> Q) : - forall (fAB fAB' : A_B ~> Q), - iota_A ∘ fAB ≃ fA -> iota_B ∘ fAB ≃ fB -> - iota_A ∘ fAB' ≃ fA -> iota_B ∘ fAB' ≃ fB -> - fAB ≃ fAB'. -Proof. - intros. - rewrite (coprod_mor_unique Q fA fB) by easy. - symmetry; - rewrite (coprod_mor_unique Q fA fB) by easy. - easy. -Qed. - -Program Definition category_coproduct_unique `{cC : Category C} (A B : C) : - forall {A_B A_B'} (HA_B : CategoryCoproduct A B A_B) - (HA_B' : CategoryCoproduct A B A_B'), Isomorphism A_B A_B' := - fun A_B A_B' HA_B HA_B' => - {| - forward := HA_B.(coprod_mor) A_B' iota_A iota_B; - reverse := HA_B'.(coprod_mor) A_B iota_A iota_B; - |}. -Next Obligation. - split; eapply (coprod_mor_unique' _ iota_A iota_B); rewrite <- 1?assoc. - 1,5: rewrite 2!(proj1 (coprod_mor_prop _ _ _)); easy. - 1,4: rewrite 2!(proj2 (coprod_mor_prop _ _ _)); easy. - all: rewrite right_unit; easy. -Qed. - -Class CategoryBigCoprod `{cC : Category C} {J : Type} - (obj : J -> C) (coprod_J : C) := { - big_coprod_obj := coprod_J; - iota_i : forall i, obj i ~> coprod_J; - big_coprod_mor : - forall (Q : C) (fQ_ : forall i, obj i ~> Q), coprod_J ~> Q; - big_coprod_mor_prop: - forall (Q : C) (fQ_ : forall i, obj i ~> Q), - forall i, - iota_i i ∘ (big_coprod_mor Q fQ_) ≃ fQ_ i; - big_coprod_mor_unique : - forall (Q : C) (fQ_ : forall i, obj i ~> Q) - (fAB' : coprod_J ~> Q), - (forall i, fQ_ i ≃ iota_i i ∘ fAB') -> - fAB' ≃ big_coprod_mor Q fQ_ -}. - Inductive big_sum {J} (obj : J -> Setoid) := | in_i i (a : obj i) : big_sum obj. @@ -1013,47 +854,6 @@ Next Obligation. easy. Qed. -Reserved Notation "A ∏ B" (at level 40). -Reserved Notation "f @∏ g" (at level 40). - -Class Category_with_Products {C} (cC : Category C) := { - cat_prod : C -> C -> C - where "A ∏ B" := (cat_prod A B); - cat_prod_product : forall {A B}, CategoryProduct A B (A ∏ B); - cat_mor_prod : - forall {A A' B B'} (fA : A ~> B) (fB : A' ~> B'), - A ∏ A' ~> B ∏ B' := fun A A' B B' fA fB => - cat_prod_product.(prod_mor) - (A ∏ A') (p_A ∘ fA) (p_B ∘ fB) - where "f @∏ g" := (cat_mor_prod f g); -}. - -Notation "A ∏ B" := (cat_prod A B) (at level 40). -Notation "f @∏ g" := (cat_mor_prod f g) (at level 40). - -#[program, export] Instance Category_with_Products_of_CartesianMonoidalCategory {C cC mC} - (ccc : @CartesianMonoidalCategory C cC mC) : Category_with_Products cC := {| - cat_prod := mC.(tensor).(obj_bimap); - cat_prod_product := tensor_cartesian; -|}. - -Coercion Category_with_Products_of_CartesianMonoidalCategory : - CartesianMonoidalCategory >-> Category_with_Products. - - -Reserved Notation "'λ' g" (at level 20). -Class CategoryExponential `{cC : Category C} {pC : Category_with_Products cC} - (A B A_B : C) := { - exp_obj := A_B; - eval_mor : A_B ∏ B ~> A; - eval_mor_transpose : forall {M : C} (f : M ∏ B ~> A), - M ~> A_B - where "'λ' g" := (eval_mor_transpose g); - eval_mor_transpose_prop : forall (M : C) (f : M ∏ B ~> A), - f ≃ (λ f @∏ id_ B) ∘ eval_mor; - eval_mor_transpose_unique : forall (M : C) (f : M ∏ B ~> A) - (lf' : M ~> A_B), f ≃ (lf' @∏ id_ B) ∘ eval_mor -> λ f ≃ lf'; -}. #[program, export] Instance MorphismSetoid (A B : Setoid) : Setoid := { base_type := A → B; @@ -1067,125 +867,9 @@ Class CategoryExponential `{cC : Category C} {pC : Category_with_Products cC} {| base_fn := fun b => f (m, b) |} |}; }. +Open Scope Cat_scope. -Class Cone {C D} {cC : Category C} {cD : Category D} - (F : Functor cD cC) := { - cone_obj : C; - cone_mor : forall (d : D), cone_obj ~> F d; - cone_mor_prop : forall {d d' : D} (f : d ~> d'), - cone_mor d ∘ F @ f ≃ cone_mor d'; -}. - -Coercion cone_mor : Cone >-> Funclass. - -Reserved Notation "'lim' F" (at level 30). - -Class Limit {C D} {cC : Category C} {cD : Category D} - (F : Functor cD cC) := { - limit_cone : Cone F; - limit_cone_factor : - forall (N : Cone F), N.(cone_obj) ~> limit_cone.(cone_obj); - limit_cone_factor_prop : forall (N : Cone F) (d : D), - limit_cone_factor N ∘ limit_cone d ≃ N d; - limit_cone_factor_unique : forall (N : Cone F) - (f' : N.(cone_obj) ~> limit_cone.(cone_obj)), - (forall (d : D), f' ∘ limit_cone d ≃ N d) -> f' ≃ limit_cone_factor N; -}. - -Lemma limit_cone_factor_unique' {C D} {cC : Category C} {cD : Category D} - {F : Functor cD cC} (L : Limit F) : forall (N : Cone F) - (f f' : N.(cone_obj) ~> L.(limit_cone).(cone_obj)), - (forall (d : D), f ∘ limit_cone d ≃ N d) -> - (forall (d : D), f' ∘ limit_cone d ≃ N d) -> - f ≃ f'. -Proof. - intros. - rewrite limit_cone_factor_unique by easy. - symmetry. - rewrite limit_cone_factor_unique by easy. - easy. -Qed. - -Coercion limit_cone : Limit >-> Cone. - -Notation "'lim' F" := (limit_cone (F:=F)).(cone_obj) (at level 30). - -#[program] Definition limit_unique {C D} {cC : Category C} {cD : Category D} - {F : Functor cD cC} (L L' : Limit F) : L.(cone_obj) <~> L'.(cone_obj) := {| - forward := L'.(limit_cone_factor) L; - reverse := L.(limit_cone_factor) L' - |}. -Next Obligation. - split; (apply limit_cone_factor_unique'; - [| intros; rewrite left_unit; easy]); - intros; rewrite assoc, 2!limit_cone_factor_prop; easy. -Qed. - - -Class Cocone {C D} {cC : Category C} {cD : Category D} - (F : Functor cD cC) := { - cocone_obj : C; - cocone_mor : forall (d : D), F d ~> cocone_obj; - cocone_mor_prop : forall {d d' : D} (f : d ~> d'), - F @ f ∘ cocone_mor d' ≃ cocone_mor d; -}. - -Coercion cocone_mor : Cocone >-> Funclass. - -Reserved Notation "'colim' F" (at level 30). - -Class Colimit {C D} {cC : Category C} {cD : Category D} - (F : Functor cD cC) := { - colimit_cocone : Cocone F; - colimit_cocone_factor : forall (N : Cocone F), - colimit_cocone.(cocone_obj) ~> N.(cocone_obj); - colimit_cocone_factor_prop : forall (N : Cocone F) (d : D), - colimit_cocone d ∘ colimit_cocone_factor N ≃ N d; - colimit_cocone_factor_unique : forall (N : Cocone F) - (f' : colimit_cocone.(cocone_obj) ~> N.(cocone_obj)), - (forall (d : D), colimit_cocone d ∘ f' ≃ N d) -> - f' ≃ colimit_cocone_factor N; -}. - -Lemma colimit_cocone_factor_unique' {C D} {cC : Category C} {cD : Category D} - {F : Functor cD cC} (L : Colimit F) : forall (N : Cocone F) - (f f' : L.(colimit_cocone).(cocone_obj) ~> N.(cocone_obj)), - (forall (d : D), colimit_cocone d ∘ f ≃ N d) -> - (forall (d : D), colimit_cocone d ∘ f' ≃ N d) -> - f ≃ f'. -Proof. - intros. - rewrite colimit_cocone_factor_unique by easy. - symmetry. - rewrite colimit_cocone_factor_unique by easy. - easy. -Qed. - -Coercion colimit_cocone : Colimit >-> Cocone. - -Notation "'colim' F" := (colimit_cocone (F:=F)).(cocone_obj) (at level 30). - -#[program] Definition colimit_unique {C D} {cC : Category C} {cD : Category D} - {F : Functor cD cC} (L L' : Colimit F) : L.(cocone_obj) <~> L'.(cocone_obj) := {| - forward := L.(colimit_cocone_factor) L'; - reverse := L'.(colimit_cocone_factor) L - |}. -Next Obligation. - split; (apply colimit_cocone_factor_unique'; - [| intros; rewrite right_unit; easy]); - intros; rewrite <- assoc, 2!colimit_cocone_factor_prop; easy. -Qed. - - -Class Category_with_Small_Limits {C} (cC : Category C) := { - take_set_limit : forall {D : Set} (cD : Category D) (F : Functor cD cC), Limit F; -}. - -Class Category_with_Limits {C} (cC : Category C) := { - take_limit : forall {D : Type} (cD : Category D) (F : Functor cD cC), Limit F; -}. - #[export, program] Instance std_limit_setoid {D} {cD : Category D} (F : Functor cD Std) : Setoid := { base_type := { s : big_setoid_prod F.(obj_map) | diff --git a/examples/ZXExample.v b/examples/ZXExample.v index a7c3c93..e5827f3 100644 --- a/examples/ZXExample.v +++ b/examples/ZXExample.v @@ -4,20 +4,7 @@ From VyZX Require Import CoreData. From VyZX Require Import CoreRules. From VyZX Require Import PermutationRules. -(* From ViCaR Require Import Category (Category, Bifunctor, NaturalBiIsomorphism, Isomorphism). -From ViCaR Require Import Monoidal (MonoidalCategory). -From ViCaR Require Import BraidedMonoidal (BraidedMonoidalCategory). -From ViCaR Require Import SymmetricMonoidal (SymmetricMonoidalCategory). -From ViCaR Require Import Dagger (DaggerCategory). -From ViCaR Require Import DaggerMonoidal (DaggerMonoidalCategory). *) - -From ViCaR Require CategoryTypeclass. -Import -(notations) CategoryTypeclass. - -(* From ViCaR Require Import CategoryTypeclass (Category, MonoidalCategory, -BraidedMonoidalCategory, SymmetricMonoidalCategory, DaggerBraidedMonoidalCategory, -DaggerCategory, DaggerSymmetricMonoidalCategory). *) - +From ViCaR Require Import CategoryTypeclassCompatibility. Lemma proportional_equiv {n m : nat} : equivalence (ZX n m) proportional. Proof. @@ -37,15 +24,17 @@ Qed. #[export] Instance ZXCategory : Category nat := { morphism := ZX; - c_equiv := @proportional; + compose := @Compose; + c_identity n := n_wire n; +}. + +#[export] Instance ZXCategoryC : CategoryCoherence ZXCategory := { c_equiv_rel := @proportional_equiv; - compose := @Compose; compose_compat := @Proportional.compose_compat; assoc := @ComposeRules.compose_assoc; - c_identity n := n_wire n; left_unit _ _ := nwire_removal_l; right_unit _ _ := nwire_removal_r; }. @@ -205,62 +194,37 @@ Definition ZXTensorBiFunctor : Bifunctor ZXCategory ZXCategory ZXCategory := {| |}. #[export] Instance ZXMonoidalCategory : MonoidalCategory ZXCategory := { - tensor := ZXTensorBiFunctor; + obj_tensor := Nat.add; + mor_tensor := @Stack; + mon_I := O; associator := fun n m o => {| forward := @zx_associator n m o; reverse := @zx_inv_associator n m o; - isomorphism_inverse := conj (@zx_associator_inv_compose n m o) (@zx_inv_associator_compose n m o) - (* id_A := @zx_associator_inv_compose n m o; - id_B := @zx_inv_associator_compose n m o; *) + isomorphism_inverse := + conj (@zx_associator_inv_compose n m o) (@zx_inv_associator_compose n m o) |}; left_unitor := fun n => {| forward := @zx_left_unitor n; reverse := @zx_inv_left_unitor n; - isomorphism_inverse := conj (@zx_left_inv_compose n) (@zx_inv_left_compose n) - (* id_A := @zx_left_inv_compose n; - id_B := @zx_inv_left_compose n; *) + isomorphism_inverse := + conj (@zx_left_inv_compose n) (@zx_inv_left_compose n) |}; right_unitor := fun n => {| forward := @zx_right_unitor n; reverse := @zx_inv_right_unitor n; - isomorphism_inverse := conj (@zx_right_inv_compose n) (@zx_inv_right_compose n) - (* id_A := @zx_right_inv_compose n; - id_B := @zx_inv_right_compose n; *) + isomorphism_inverse := + conj (@zx_right_inv_compose n) (@zx_inv_right_compose n) |}; +}. - associator_cohere := @zx_associator_cohere; - left_unitor_cohere := @zx_left_unitor_cohere; - right_unitor_cohere := @zx_right_unitor_cohere; - - triangle := @zx_triangle_lemma; - pentagon := @zx_pentagon_lemma; -(* - tensor := Nat.add; - I := 0; - - tensor_morph _ _ _ _ := Stack; - tensor_morph_compat := stack_compat; - - associator := @zx_associator; - inv_associator := @zx_inv_associator; - associator_inv_compose := @zx_associator_inv_compose; - inv_associator_compose := @zx_inv_associator_compose; - - left_unitor := @zx_left_unitor; - inv_left_unitor := @zx_inv_left_unitor; - left_inv_compose := @zx_left_inv_compose; - inv_left_compose := @zx_inv_left_compose; - - right_unitor := @zx_right_unitor; - inv_right_unitor := @zx_inv_right_unitor; - right_inv_compose := @zx_right_inv_compose; - inv_right_compose := @zx_inv_right_compose; - - bifunctor_id := n_wire_stack; - bifunctor_comp := @stack_compose_distr; +#[export] Instance ZXMonoidalCategoryC : + MonoidalCategoryCoherence ZXMonoidalCategory := { + tensor_id := n_wire_stack; + tensor_compose := @stack_compose_distr; + tensor_compat := @stack_simplify; associator_cohere := @zx_associator_cohere; left_unitor_cohere := @zx_left_unitor_cohere; @@ -268,7 +232,6 @@ Definition ZXTensorBiFunctor : Bifunctor ZXCategory ZXCategory ZXCategory := {| triangle := @zx_triangle_lemma; pentagon := @zx_pentagon_lemma; -*) }. Definition zx_braiding {n m} := @@ -300,18 +263,6 @@ Proof. prop_perm_eq. rewrite Nat.add_comm. cleanup_perm_of_zx; easy. - (* intros. - unfold zx_braiding. unfold zx_inv_braiding. - unfold n_top_to_bottom. - unfold n_bottom_to_top. - rewrite cast_compose_mid. - simpl_casts. - fold (n_compose_bot n (m + n)). - rewrite cast_fn_eq_dim. - rewrite n_compose_top_compose_bottom. - reflexivity. - Unshelve. - all: rewrite (Nat.add_comm n m); easy. *) Qed. Lemma zx_inv_braiding_compose : forall {n m}, @@ -321,18 +272,6 @@ Proof. prop_perm_eq. rewrite Nat.add_comm. cleanup_perm_of_zx; easy. - (* intros. - unfold zx_braiding. unfold zx_inv_braiding. - unfold n_top_to_bottom. - unfold n_bottom_to_top. - rewrite cast_compose_mid. - simpl_casts. - fold (n_compose_top n (n + m)). - rewrite cast_fn_eq_dim. - rewrite n_compose_bottom_compose_top. - reflexivity. - Unshelve. - all: rewrite (Nat.add_comm n m); easy. *) Qed. Lemma n_top_to_bottom_split : forall {n m o o'} prf1 prf2 prf3 prf4, @@ -352,20 +291,6 @@ Proof. intros. prop_perm_eq. solve_modular_permutation_equalities. - (* intros. - unfold zx_braiding. unfold zx_associator. - simpl_casts. - rewrite cast_compose_l. simpl_casts. - rewrite compose_assoc. - rewrite cast_compose_l. simpl_casts. - cleanup_zx. simpl_casts. - rewrite cast_compose_l. - simpl_casts. cleanup_zx. - rewrite cast_compose_l. simpl_casts. - rewrite (cast_compose_r _ _ _ (n_wire (m + o + n))). - cleanup_zx. simpl_casts. - rewrite n_top_to_bottom_split. - reflexivity. *) Qed. Lemma n_bottom_to_top_split : forall {n m o o'} prf1 prf2 prf3 prf4, @@ -383,27 +308,9 @@ Lemma hexagon_lemma_2 : forall {n m o}, Proof. prop_perm_eq. solve_modular_permutation_equalities. - (* intros. - unfold zx_inv_braiding. unfold zx_associator. - simpl_casts. - rewrite cast_compose_l. simpl_casts. - rewrite compose_assoc. - rewrite cast_compose_l. simpl_casts. - cleanup_zx. simpl_casts. - rewrite cast_compose_l. - simpl_casts. cleanup_zx. - rewrite cast_compose_l. simpl_casts. - rewrite (cast_compose_r _ _ _ (n_wire (m + o + n))). - cleanup_zx. simpl_casts. - rewrite n_bottom_to_top_split. - reflexivity. *) Qed. -Require Export KronComm_orig. -(* Search "cast_" -ZXperm -Box -Z -X -n_box. *) - - - +Require Export KronComm. @@ -468,20 +375,16 @@ Definition ZXBraidingIsomorphism : forall n m, |}. -#[export] Instance ZXBraidedMonoidalCategory : BraidedMonoidalCategory nat := { +#[export] Instance ZXBraidedMonoidalCategory : + BraidedMonoidalCategory ZXMonoidalCategory := { braiding := ZXBraidingBiIsomorphism; +}. +#[export] Instance ZXBraidedMonoidalCategoryC : + BraidedMonoidalCategoryCoherence ZXBraidedMonoidalCategory := { + braiding_natural := zx_braiding_iso_natural; hexagon_1 := @hexagon_lemma_1; hexagon_2 := @hexagon_lemma_2; -(* - braiding := @zx_braiding; - inv_braiding := @zx_inv_braiding; - braiding_inv_compose := @zx_braiding_inv_compose; - inv_braiding_compose := @zx_inv_braiding_compose; - - hexagon_1 := @hexagon_lemma_1; - hexagon_2 := @hexagon_lemma_2; -*) }. Lemma n_top_to_bottom_is_bottom_to_top : forall {n m}, @@ -489,37 +392,6 @@ Lemma n_top_to_bottom_is_bottom_to_top : forall {n m}, Proof. prop_perm_eq. solve_modular_permutation_equalities. -(* - intros. - unfold n_bottom_to_top. - unfold bottom_to_top. - unfold n_top_to_bottom. - induction n. - - intros. - rewrite n_compose_0. - simpl. - rewrite <- n_compose_transpose. - rewrite n_compose_n_top_to_bottom. - rewrite n_wire_transpose. - reflexivity. - - intros. - rewrite n_compose_grow_l. - assert ((S n + m)%nat = (n + S m)%nat) by lia. - assert (top_to_bottom (S n + m) - ∝ cast (S n + m) (S n + m) H H (top_to_bottom (n + S m))) - by (rewrite cast_fn_eq_dim; reflexivity). - rewrite H0. - rewrite cast_n_compose. - rewrite IHn. - rewrite <- cast_n_compose. - rewrite <- H0. - rewrite n_compose_grow_l. - rewrite <- cast_transpose. - rewrite <- H0. - fold (bottom_to_top (S n + m)). - rewrite <- compose_assoc. - rewrite top_to_bottom_to_top. cleanup_zx. - reflexivity. *) Qed. Lemma braiding_symmetry : forall n m, @@ -527,14 +399,10 @@ Lemma braiding_symmetry : forall n m, Proof. prop_perm_eq. solve_modular_permutation_equalities. -(* intros. - unfold zx_braiding. unfold zx_inv_braiding. - apply cast_compat. - rewrite n_top_to_bottom_is_bottom_to_top. - reflexivity. *) Qed. -#[export] Instance ZXSymmetricMonoidalCategory : SymmetricMonoidalCategory nat := { +#[export] Instance ZXSymmetricMonoidalCategory + : SymmetricMonoidalCategory ZXBraidedMonoidalCategory := { symmetry := braiding_symmetry; }. @@ -567,1152 +435,73 @@ Proof. easy. Qed. -#[export] Instance ZXDaggerCategory : DaggerCategory nat := { - adjoint := @ZXCore.adjoint; - adjoint_involutive := @Proportional.adjoint_involutive; - adjoint_id := nwire_adjoint; - adjoint_contravariant := @compose_adjoint; - adjoint_compat := Proportional.adjoint_compat; -}. - -Lemma zx_dagger_tensor_compat : forall {n n' m m'} - (zx0: ZX n m) (zx1 : ZX n' m'), - zx0 †%ZX ↕ zx1 †%ZX ∝ (zx0 ↕ zx1) †%ZX. -Proof. - intros. - rewrite stack_adjoint. - easy. -Qed. - -Lemma zx_associator_unitary_r : forall {n m o}, - zx_associator ⟷ zx_associator †%ZX ∝ n_wire (n + m + o). -Proof. - intros. - unfold zx_associator. - rewrite cast_adj. - rewrite nwire_adjoint. - simpl_permlike_zx. - reflexivity. -Qed. - -Lemma zx_associator_unitary_l : forall {n m o}, - zx_associator †%ZX ⟷ zx_associator ∝ n_wire (n + (m + o)). -Proof. - intros. - unfold zx_associator. - rewrite cast_adj. - rewrite nwire_adjoint. - simpl_permlike_zx. - rewrite cast_n_wire. - reflexivity. -Qed. - -Lemma zx_left_unitor_unitary_r : forall {n}, - zx_left_unitor ⟷ zx_left_unitor †%ZX ∝ n_wire (0 + n). -Proof. - intros. unfold zx_left_unitor. - simpl_permlike_zx. - rewrite nwire_adjoint. - reflexivity. -Qed. - -Lemma zx_left_unitor_unitary_l : forall {n}, - zx_left_unitor †%ZX ⟷ zx_left_unitor ∝ n_wire n. -Proof. - intros. unfold zx_left_unitor. - simpl_permlike_zx. - rewrite nwire_adjoint. - reflexivity. -Qed. - -Lemma zx_right_unitor_unitary_r : forall {n}, - zx_right_unitor ⟷ zx_right_unitor †%ZX ∝ n_wire (n + 0). -Proof. - intros. unfold zx_right_unitor. - simpl_permlike_zx. - rewrite cast_adj, nwire_adjoint. - simpl_permlike_zx. - rewrite cast_n_wire. - reflexivity. -Qed. - -Lemma zx_right_unitor_unitary_l : forall {n}, - zx_right_unitor †%ZX ⟷ zx_right_unitor ∝ n_wire n. -Proof. - intros. unfold zx_right_unitor. - simpl_permlike_zx. - rewrite cast_adj, nwire_adjoint. - simpl_permlike_zx. - reflexivity. -Qed. - -Lemma helper_eq: forall n m (prf: n = m), (n + n = m + m)%nat. -Proof. intros. subst. reflexivity. Qed. - -Lemma cast_n_cup_unswapped : forall n m prf1 prf2, - cast (n + n) 0 (helper_eq _ _ prf1) prf2 (n_cup_unswapped m) ∝ n_cup_unswapped n. -Proof. - intros. - subst. - rewrite cast_id_eq. - easy. -Qed. - -Lemma wire_stack_distr_compose_l : forall n m o (zx0 : ZX n m) (zx1 : ZX m o), - — ↕ (zx0 ⟷ zx1) ∝ (— ↕ zx0) ⟷ (— ↕ zx1). -Proof. - intros. - rewrite <- stack_compose_distr. - cleanup_zx. - easy. -Qed. - -Lemma wire_stack_distr_compose_r : forall n m o (zx0 : ZX n m) (zx1 : ZX m o), - (zx0 ⟷ zx1) ↕ — ∝ (zx0 ↕ —) ⟷ (zx1 ↕ —). -Proof. - intros. - rewrite <- stack_compose_distr. - cleanup_zx. - easy. -Qed. - -Lemma n_cup_unswapped_grow_r : forall n prf1 prf2, - n_cup_unswapped (S n) ∝ - cast _ _ prf1 prf2 (— ↕ n_cup_unswapped n ↕ —) ⟷ ⊃. -Proof. - intros. - induction n. - - simpl. cleanup_zx. - apply compose_simplify; [|easy]. - prop_perm_eq. - - rewrite n_cup_unswapped_grow_l. - rewrite IHn at 1. - rewrite n_cup_unswapped_grow_l. - bundle_wires. - rewrite <- compose_assoc. - apply compose_simplify; [|easy]. - rewrite wire_stack_distr_compose_l, wire_stack_distr_compose_r. - rewrite (prop_iff_double_cast _ (1 + 0 + 1) _ _ eq_refl). - simpl_casts. - rewrite (cast_compose_mid_contract _ (1 + (n + n) + 1)). - rewrite 2!cast_contract, cast_id. - apply compose_simplify; [|easy]. - simpl_casts. - simpl (n_wire S n). - rewrite 4!stack_assoc. - rewrite (stack_assoc — (n_wire n) _). - simpl_casts. - repeat (rewrite cast_stack_distribute; apply stack_simplify; try prop_perm_eq). - rewrite cast_id; easy. - Unshelve. - all: lia. -Qed. -Lemma nwire_stack_distr_compose_l : forall k n m o (zx0 : ZX n m) (zx1 : ZX m o), - n_wire k ↕ (zx0 ⟷ zx1) ∝ (n_wire k ↕ zx0) ⟷ (n_wire k ↕ zx1). -Proof. - intros. - rewrite <- stack_compose_distr. - cleanup_zx. - easy. -Qed. -Lemma nwire_stack_distr_compose_r : forall k n m o (zx0 : ZX n m) (zx1 : ZX m o), - (zx0 ⟷ zx1) ↕ n_wire k ∝ (zx0 ↕ n_wire k) ⟷ (zx1 ↕ n_wire k). -Proof. - intros. - rewrite <- stack_compose_distr. - cleanup_zx. - easy. -Qed. -Lemma n_cup_unswapped_comm_1 : forall k prf1 prf2 prf3 prf4, - cast (S k + (S k)) _ prf1 prf2 (n_wire k ↕ ⊃ ↕ n_wire k) ⟷ (n_cup_unswapped k) - ∝ cast _ _ prf3 prf4 (— ↕ n_cup_unswapped k ↕ —) ⟷ ⊃. -Proof. - intros. - rewrite <- n_cup_unswapped_grow_l. - rewrite <- n_cup_unswapped_grow_r. - easy. -Qed. -Lemma n_wire_add_stack : forall n k, - n_wire (n + k) ∝ n_wire n ↕ n_wire k. -Proof. prop_perm_eq. Qed. -Lemma n_wire_add_stack_rev : forall n k prf1 prf2, - n_wire (n + k) ∝ cast _ _ prf1 prf2 (n_wire k ↕ n_wire n). -Proof. prop_perm_eq. Qed. +From ViCaR Require Import CategoryAutomationCompatibility. -Lemma stack_assoc' : forall {n0 n1 n2 m0 m1 m2} (zx0 : ZX n0 m0) - (zx1 : ZX n1 m1) (zx2 : ZX n2 m2) prfn prfm, - zx0 ↕ (zx1 ↕ zx2) ∝ cast _ _ prfn prfm ((zx0 ↕ zx1) ↕ zx2). -Proof. - intros. - rewrite stack_assoc. - rewrite cast_cast_eq, cast_id. - easy. - Unshelve. - all: lia. -Qed. +Section Testing. -Lemma n_cup_unswapped_comm : forall n k prf1 prf2 prf3 prf4, - cast (S n + k + (S n + k)) _ prf1 prf2 (n_wire (n + k) ↕ ⊃ ↕ n_wire (n + k)) ⟷ (n_wire n ↕ n_cup_unswapped k ↕ n_wire n) - ∝ cast _ _ prf3 prf4 (n_wire (S n) ↕ n_cup_unswapped k ↕ n_wire (S n)) ⟷ (n_wire n ↕ ⊃ ↕ n_wire n). -Proof. - intros. - rewrite n_wire_add_stack at 1. - rewrite n_wire_add_stack_rev at 1. - rewrite (n_wire_add_stack_rev 1 n) at 1. - rewrite (n_wire_add_stack 1 n) at 1. - rewrite 5!stack_assoc. - repeat rewrite cast_cast_eq. - simpl_casts. - rewrite stack_assoc. - rewrite (prop_iff_double_cast (n + (k + 2 + k + n)) (n + (0 + n))). - rewrite (cast_compose_mid_contract _ (n + (k + 0 + k + n))). - repeat rewrite cast_cast_eq. - rewrite (cast_compose_mid_contract _ (n + (2 + n))). - rewrite 2!cast_cast_eq. - erewrite 3!(cast_stack_distribute _ _ _ _ _ _ (n_wire n)). - rewrite 4!cast_id_eq. - rewrite <- 2!nwire_stack_distr_compose_l. - apply stack_simplify; [easy|]. - rewrite <- wire_to_n_wire. - rewrite 3!stack_assoc'. - rewrite (stack_assoc' (_ ↕ _) (—) (n_wire n)). - repeat rewrite cast_cast_eq. - rewrite 3!(cast_stack_distribute (o':=n)). - rewrite (cast_id (n:=n)). - rewrite <- 2!nwire_stack_distr_compose_r. - apply stack_simplify; [|easy]. - rewrite (prop_iff_double_cast ((S k) + (S k)) (0)). - simpl_permlike_zx. - symmetry. - rewrite (cast_compose_mid_contract _ 2). - symmetry. - rewrite cast_id_eq. - rewrite cast_cast_eq. - rewrite <- n_cup_unswapped_comm_1. - rewrite (prop_iff_double_cast ((S k) + (S k)) (0)). - rewrite (cast_compose_mid_contract _ (k + k)). - simpl_casts. - easy. - Unshelve. - all: lia. -Qed. +(* Import (notations) CategoryAutomation. *) +Delimit Scope Cat_scope with Cat. -Lemma n_cup_unswapped_grow_k_l : forall n k prf1 prf2, - n_cup_unswapped (n + k) ∝ - cast _ _ prf1 prf2 (n_wire n ↕ (n_cup_unswapped k) ↕ n_wire n) ⟷ n_cup_unswapped n. -Proof. - intros. - induction n. - - rewrite stack_empty_l, stack_empty_r, cast_cast_eq, - cast_id_eq, compose_empty_r. - easy. - - rewrite (prop_iff_double_cast (S (n + k) + S (n + k)) 0). - rewrite cast_n_cup_unswapped. - rewrite n_cup_unswapped_grow_l, IHn. - rewrite n_cup_unswapped_grow_l. - simpl_permlike_zx. - (* simpl_casts. *) - repeat rewrite <- compose_assoc. - apply compose_simplify; [|easy]. - symmetry. - rewrite (cast_compose_mid (n + 2 + n)). - erewrite (prop_iff_double_cast ((S n + k) + (S n + k)) (n + 0 + n) _ _ eq_refl). - rewrite 2!cast_contract. - rewrite cast_compose_mid_contract, 2!cast_contract, cast_id. - rewrite cast_compose_mid_contract, 2!cast_contract, cast_id. - rewrite n_cup_unswapped_comm. - easy. - Unshelve. - all: try easy; auto with arith; lia. -Qed. - -Lemma n_cup_unswapped_add_comm : forall n k prf1 prf2, - n_cup_unswapped (n + k) ∝ cast _ _ prf1 prf2 (n_cup_unswapped (k + n)). -Proof. - intros. - assert (H: (k + n = n + k)%nat) by lia. - generalize dependent (k + n)%nat. - generalize dependent (n + k)%nat. - intros; subst. - rewrite cast_id_eq. - easy. -Qed. - -Lemma n_cup_unswapped_grow_k_r : forall n k prf1 prf2, - n_cup_unswapped (n + k) ∝ - cast _ _ prf1 prf2 (n_wire k ↕ (n_cup_unswapped n) ↕ n_wire k) ⟷ n_cup_unswapped k. -Proof. - intros. - rewrite n_cup_unswapped_add_comm. - rewrite n_cup_unswapped_grow_k_l. - rewrite (cast_compose_mid_contract _ (k + k)%nat). - simpl_casts. - apply compose_simplify; [|easy]. - erewrite cast_proof_independence. - reflexivity. - Unshelve. all:lia. -Qed. -Lemma stack_ncup_unswapped_split : forall {n0t n0b n1t n1b} n m (zx0top : ZX n0t m) (zx0bot : ZX n0b m) - (zx1top : ZX n1t n) (zx1bot : ZX n1b n) prf1 prf2 prf3 prf4 prf5 prf6, - (zx1top ↕ zx0top) ↕ (zx0bot ↕ zx1bot) - ⟷ cast ((n + m) + (m + n)) 0 prf1 prf2 (n_cup_unswapped (n + m)) - ∝ cast _ _ prf5 prf6 (zx1top ↕ ((zx0top ↕ zx0bot) ⟷ n_cup_unswapped m) ↕ zx1bot - ⟷ cast (n + 0 + n) 0 prf3 prf4 (n_cup_unswapped n)). +Lemma test_part_rw : forall {nIn nOut} (zx : ZX nIn 0) (zx1 : ZX 0 nOut), + zx ⟷ (⦰ ⟷ zx1) ∝ zx ⟷ zx1. Proof. intros. - rewrite cast_compose_r. - simpl_permlike_zx. - rewrite n_cup_unswapped_grow_k_l, <- compose_assoc. - rewrite (cast_compose_mid_contract _ (n + n)). - simpl_permlike_zx. - apply compose_simplify; [|easy]. - rewrite stack_assoc, (stack_assoc' zx0top). - simpl_casts. - rewrite stack_assoc'. - rewrite cast_cast_eq. - rewrite (prop_iff_double_cast (n1t + (n0t + n0b) + n1b) (n + 0 + n)). - rewrite (cast_compose_mid_contract _ (n + (m + m) + n)). - simpl_permlike_zx. - rewrite <- 2!stack_compose_distr, 2!nwire_removal_r. - easy. - Unshelve. - all: lia. -Qed. - -(* Local Open Scope matrix_scope. *) - - - -Lemma sem_n_cup_unswapped_2 : - ⟦ n_cup_unswapped 2 ⟧ = - fun x y => if (x=?0) && ((y=?0) || (y=?6) || (y=?9) || (y=?15)) then C1 else C0. -Proof. - unfold n_cup_unswapped. - repeat (simpl; - rewrite cast_semantics; simpl). - rewrite 2!id_kron. - replace (list2D_to_matrix [[C1; C0; C0; C1]]) with - (fun x y => if (x =? 0)&&((y =? 0) || (y =? 3)) then C1 else C0) by solve_matrix. - solve_matrix. -Qed. - - - - - -Lemma swap_2cup_transport : - ⟦ n_cup_unswapped 2 ⟧ × (kron swap (Matrix.I (2^2))) - = ⟦ n_cup_unswapped 2 ⟧ × (kron (Matrix.I (2^2)) swap). -Proof. - apply mat_equiv_eq; auto with wf_db. - rewrite sem_n_cup_unswapped_2. - by_cell; try lca. -Qed. - -Lemma swap_2cup_flip : - ⨉ ↕ n_wire 2 ⟷ n_cup_unswapped 2 ∝ n_wire 2 ↕ ⨉ ⟷ n_cup_unswapped 2. -Proof. - prop_exists_nonzero 1. - rewrite (compose_semantics (⨉ ↕ n_wire 2) (n_cup_unswapped 2)). (* For some reason, it needs *) - rewrite (compose_semantics (n_wire 2 ↕ ⨉) (n_cup_unswapped 2)). (* its arguments explicitly *) - rewrite Mscale_1_l. - rewrite 2!stack_semantics, n_wire_semantics. - simpl (⟦ ⨉ ⟧). - apply swap_2cup_transport. -Qed. - -Tactic Notation "preplace" constr(zx0) "with" constr(zx1) := - (let H := fresh "H" in - enough (H: zx0 ∝ zx1); - [rewrite H; clear H| ]). - -Tactic Notation "preplace" constr(zx0) "with" constr(zx1) "in" hyp(target) := - (let H := fresh "H" in - enough (H: zx0 ∝ zx1); - [rewrite H in target; clear H| ]). - -Tactic Notation "preplace" constr(zx0) "with" constr(zx1) "in" "*" := - (let H := fresh "H" in - enough (H: zx0 ∝ zx1); - [rewrite H in *; clear H| ]). - -Tactic Notation "preplace" constr(zx0) "with" constr(zx1) "by" tactic(slv) := - (let H := fresh "H" in - assert (H: zx0 ∝ zx1) by slv; - rewrite H; clear H). - -Tactic Notation "preplace" constr(zx0) "with" constr(zx1) "in" hyp(target) "by" tactic(slv) := - (let H := fresh "H" in - assert (H: zx0 ∝ zx1) by slv; - rewrite H in target; clear H). - -Tactic Notation "preplace" constr(zx0) "with" constr(zx1) "in" "*" "by" tactic(slv) := - (let H := fresh "H" in - assert (H: zx0 ∝ zx1) by slv; - rewrite H in *; clear H). - - - -Lemma n_cup_unswapped_split_stack : forall {n00 m0 n01 m1 n10 n11 nbot} - (zx00 : ZX n00 m0) (zx01 : ZX n01 m1) (zx10 : ZX n10 m1) (zx11 : ZX n11 m0) - prf1 prf2 prf3 prf4 prf5 prf6, - (zx00 ↕ zx01) ↕ (cast nbot (m0 + m1) prf1 prf2 (zx10 ↕ zx11)) ⟷ n_cup_unswapped (m0 + m1) - ∝ - cast _ 0 prf5 prf6 (n_wire n00 - ↕ (zx01 ↕ zx10 ⟷ n_cup_unswapped m1) ↕ n_wire n11 ⟷ - cast _ 0 prf3 prf4 (zx00 ↕ zx11 ⟷ n_cup_unswapped m0)). -Proof. - intros. - rewrite cast_stack_r. - rewrite n_cup_unswapped_grow_k_l. - rewrite <- compose_assoc. - rewrite stack_assoc, (stack_assoc' zx01). - simpl_casts. - rewrite stack_assoc', cast_cast_eq. - rewrite <- cast_compose_mid_contract. - rewrite <- 2!stack_compose_distr, 2!nwire_removal_r. - rewrite pull_out_bot, <- (nwire_removal_l zx11), stack_compose_distr. - rewrite cast_compose_distribute. - rewrite (cast_compose_mid_contract _ (n00 + 0 + n11)). - rewrite compose_assoc. - apply compose_simplify; [easy | ]. - simpl_casts. - rewrite cast_compose_mid_contract, cast_id_eq. - apply compose_simplify; [|easy]. - rewrite nwire_removal_l. - easy. - Unshelve. - all: lia. -Qed. - -Lemma n_cup_unswapped_split_stack' : forall {n00 m0 n01 m1 n10 n11 ntop} - (zx00 : ZX n00 m0) (zx01 : ZX n01 m1) (zx10 : ZX n10 m1) (zx11 : ZX n11 m0) - prf1 prf2 prf3 prf4 prf5 prf6, - (cast ntop (m1 + m0) prf1 prf2 (zx00 ↕ zx01)) ↕ (zx10 ↕ zx11) ⟷ n_cup_unswapped (m1 + m0) - ∝ - cast _ 0 prf5 prf6 (n_wire n00 - ↕ (zx01 ↕ zx10 ⟷ n_cup_unswapped m1) ↕ n_wire n11 ⟷ - cast _ 0 prf3 prf4 (zx00 ↕ zx11 ⟷ n_cup_unswapped m0)). -Proof. - intros. - rewrite (prop_iff_double_cast (ntop + (n10 + n11)) 0). - rewrite (cast_compose_mid_contract _ (m0 + m1 + (m0 + m1))). - rewrite cast_n_cup_unswapped by lia. - subst ntop. - rewrite cast_stack_distribute, cast_cast_eq, cast_id_eq. - rewrite n_cup_unswapped_split_stack. - simpl_casts. - easy. - Unshelve. - all: lia. -Qed. - -Lemma n_cup_unswapped_split_stack_cast : forall {n00 m0 n01 m1 n10 n11 ntop nbot} - (zx00 : ZX n00 m0) (zx01 : ZX n01 m1) (zx10 : ZX n10 m1) (zx11 : ZX n11 m0) - prf1 prf2 prf3 prf4 prf5 prf6 prf7 prf8, - (cast ntop (m0 + m1) prf1 prf2 (zx00 ↕ zx01)) - ↕ (cast nbot (m0 + m1) prf3 prf4 (zx10 ↕ zx11)) - ⟷ n_cup_unswapped (m0 + m1) ∝ - cast _ 0 prf5 prf6 ( - n_wire n00 ↕ (zx01 ↕ zx10 ⟷ n_cup_unswapped m1) ↕ n_wire n11 - ⟷ cast _ 0 prf7 prf8 (zx00 ↕ zx11 ⟷ n_cup_unswapped m0)). -Proof. - intros. - subst ntop. - rewrite cast_id_eq. - apply n_cup_unswapped_split_stack. -Qed. - - -Lemma n_cup_unswapped_split_stack_n_wire_bot : forall {n0 n1} - (zx0 : ZX n0 n0) (zx1 : ZX n1 n1) prf1 prf2 prf3 prf4, - zx0 ↕ zx1 ↕ n_wire (n0 + n1) ⟷ n_cup_unswapped (n0 + n1) - ∝ cast _ 0 prf1 prf2 ( - n_wire n0 ↕ (zx1 ↕ n_wire n1 ⟷ n_cup_unswapped n1) ↕ n_wire n0 ⟷ - cast _ 0 prf3 prf4 (zx0 ↕ n_wire n0 ⟷ n_cup_unswapped n0)). -Proof. - intros. - rewrite n_wire_add_stack_rev. - rewrite n_cup_unswapped_split_stack. - easy. - Unshelve. - all: lia. -Qed. - - - -Lemma n_cup_unswapped_split_stack_n_wire_top : forall {n0 n1} - (zx0 : ZX n0 n0) (zx1 : ZX n1 n1) prf1 prf2 prf3 prf4, - n_wire (n0 + n1) ↕ (zx0 ↕ zx1) ⟷ n_cup_unswapped (n0 + n1) - ∝ cast _ 0 prf1 prf2 ( - n_wire n1 ↕ (n_wire n0 ↕ zx0 ⟷ n_cup_unswapped n0) ↕ n_wire n1 ⟷ - cast _ 0 prf3 prf4 (n_wire n1 ↕ zx1 ⟷ n_cup_unswapped n1)). -Proof. - intros. - rewrite n_wire_add_stack_rev. - rewrite n_cup_unswapped_split_stack'. - easy. - Unshelve. - all: lia. -Qed. - -Lemma n_cup_unswapped_split_stack_n_wire_bot' : forall {n0 n1 ntop} - (zx0 : ZX n0 n0) (zx1 : ZX n1 n1) prf1 prf2 prf3 prf4 prf5 prf6, - cast ntop _ prf1 prf2 (zx0 ↕ zx1) ↕ n_wire (n1 + n0) ⟷ n_cup_unswapped (n1 + n0) - ∝ cast _ 0 prf3 prf4 ( - n_wire n0 ↕ (zx1 ↕ n_wire n1 ⟷ n_cup_unswapped n1) ↕ n_wire n0 ⟷ - cast _ 0 prf5 prf6 (zx0 ↕ n_wire n0 ⟷ n_cup_unswapped n0)). -Proof. - intros. - rewrite n_wire_add_stack. - rewrite n_cup_unswapped_split_stack'. - easy. - Unshelve. - all: lia. -Qed. - -Lemma n_cup_unswapped_split_stack_n_wire_top' : forall {n0 n1 ntop} - (zx0 : ZX n0 n0) (zx1 : ZX n1 n1) prf1 prf2 prf3 prf4 prf5 prf6, - n_wire (n1 + n0) ↕ cast ntop _ prf1 prf2 (zx0 ↕ zx1) ⟷ n_cup_unswapped (n1 + n0) - ∝ cast _ 0 prf3 prf4 ( - n_wire n1 ↕ (n_wire n0 ↕ zx0 ⟷ n_cup_unswapped n0) ↕ n_wire n1 ⟷ - cast _ 0 prf5 prf6 (n_wire n1 ↕ zx1 ⟷ n_cup_unswapped n1)). -Proof. - intros. - rewrite n_wire_add_stack. - rewrite n_cup_unswapped_split_stack. - easy. - Unshelve. - all: lia. -Qed. - - -Lemma n_cup_unswapped_grow_r_back : forall n prf1 prf2, - (— ↕ n_cup_unswapped (n) ↕ — ⟷ ⊃) - ∝ cast _ _ prf1 prf2 (n_cup_unswapped (S n)). -Proof. - intros. - rewrite (n_cup_unswapped_grow_r n). - rewrite cast_compose_l. - rewrite cast_cast_eq, 2!cast_id_eq. - easy. - Unshelve. - all: lia. -Qed. - -Lemma n_cup_unswapped_grow_k_r_back : forall n k prf1 prf2 prf3 prf4, - (n_wire k ↕ n_cup_unswapped n ↕ n_wire k) - ⟷ cast (k + 0 + k) 0 prf1 prf2 (n_cup_unswapped k) - ∝ cast _ 0 prf3 prf4 (n_cup_unswapped (n + k)). -Proof. - intros. - rewrite n_cup_unswapped_grow_k_r. - rewrite (cast_compose_mid_contract _ (k + 0 + k)). - rewrite cast_cast_eq, cast_id_eq. + to_Cat. + partners_rw_to_Cat compose_empty_r. easy. - Unshelve. - all: lia. -Qed. - -Lemma compose_n_wire_comm : forall {n m} (zx : ZX n m), - n_wire n ⟷ zx ∝ zx ⟷ n_wire m. -Proof. - intros. - cleanup_zx; easy. Qed. -Lemma compose_stack_n_wire_comm : forall {n0 m0 n1 m1} (zx0 : ZX n0 m0) (zx1 : ZX n1 m1), - zx0 ↕ n_wire n1 ⟷ (n_wire m0 ↕ zx1) ∝ n_wire n0 ↕ zx1 ⟷ (zx0 ↕ n_wire m1). +Lemma test_part_rw_long : forall {nIn nOut} + (zx0 zx2 zx3 zx4 zx5 zx6 zx7: ZX nIn nIn) + (zx : ZX nIn 0) (zx1 : ZX 0 nOut), + zx0 ⟷ (zx2 ⟷ (zx3 ⟷ (zx4 ⟷ (zx5 ⟷ + (zx6 ⟷ (zx7 ⟷ (zx ⟷ (⦰ ⟷ zx1)))))))) ∝ zx ⟷ zx1. Proof. intros. - cleanup_zx. - easy. -Qed. - -Lemma top_to_bottom_cup_flip : forall k, - top_to_bottom k ↕ n_wire k ⟷ n_cup_unswapped k - ∝ n_wire k ↕ top_to_bottom k ⟷ n_cup_unswapped k. -Proof. - destruct k; [prop_perm_eq|]. - induction k; - [ apply compose_simplify; [prop_perm_eq | easy] | ]. - rewrite top_to_bottom_grow_r at 1. - rewrite nwire_stack_distr_compose_r. - rewrite compose_assoc. - rewrite (n_wire_add_stack 2 k) at 2. - rewrite (n_cup_unswapped_split_stack' (n_wire k) ⨉ (n_wire 2) (n_wire k)). - rewrite swap_2cup_flip. - bundle_wires; cleanup_zx. - rewrite nwire_stack_distr_compose_l, nwire_stack_distr_compose_r. - rewrite compose_assoc. - rewrite n_cup_unswapped_grow_k_r_back. - rewrite (stack_assoc' (n_wire k)). - rewrite 2!cast_stack_l. - rewrite (stack_assoc _ ⨉), cast_cast_eq. - rewrite (cast_compose_mid_contract _ (S (S k) + S (S k))). - simpl_casts. - rewrite 2!cast_stack_distribute. - simpl_permlike_zx. - bundle_wires; rewrite cast_n_wire. - rewrite <- compose_assoc, compose_stack_n_wire_comm. - rewrite (n_wire_add_stack 1 (S k)) at 2. - rewrite compose_assoc. - rewrite (n_cup_unswapped_split_stack' (top_to_bottom (S k)) — (n_wire 1) (n_wire S k)). - rewrite IHk. - rewrite <- n_cup_unswapped_split_stack'. - rewrite <- compose_assoc. - rewrite <- stack_compose_distr. - rewrite <- wire_to_n_wire. - rewrite <- (top_to_bottom_grow_l k). - apply compose_simplify; [|easy]. - apply stack_simplify; [|easy]. - prop_perm_eq. - Unshelve. - all: lia. -Qed. - -Lemma bottom_to_top_cup_flip : forall k, - bottom_to_top k ↕ n_wire k ⟷ n_cup_unswapped k - ∝ n_wire k ↕ bottom_to_top k ⟷ n_cup_unswapped k. -Proof. - intros k. - destruct k; [prop_perm_eq | ]. - induction k; - [ apply compose_simplify; [prop_perm_eq | easy] | ]. - rewrite bottom_to_top_grow_r at 1. - rewrite nwire_stack_distr_compose_r. - rewrite compose_assoc. - rewrite (n_wire_add_stack_rev 2 k) at 2. - rewrite (n_cup_unswapped_split_stack ⨉ (n_wire k) (n_wire k) (n_wire 2)). - rewrite swap_2cup_flip. - rewrite <- n_cup_unswapped_split_stack. - rewrite <- compose_assoc, <- n_wire_add_stack, compose_stack_n_wire_comm. - preplace (n_wire (2 + k)) with (n_wire (1 + (S k))) by easy. - rewrite (n_wire_add_stack_rev 1 (S k)). - rewrite compose_assoc. - rewrite (n_cup_unswapped_split_stack — (bottom_to_top (S k)) (n_wire S k) (n_wire 1)). - rewrite IHk. - rewrite <- n_cup_unswapped_split_stack. - rewrite <- compose_assoc. - apply compose_simplify; [|easy]. - rewrite <- stack_compose_distr. - apply stack_simplify; [prop_perm_eq | ]. - rewrite <- wire_to_n_wire. - rewrite cast_compose_r, cast_cast_eq. - rewrite (cast_compose_partial_contract_l _ (S k + 1)). - rewrite <- (bottom_to_top_grow_l k). - easy. - Unshelve. - all: lia. -Qed. - -Lemma a_swap_cup_flip : forall k, - a_swap k ↕ n_wire k ⟷ n_cup_unswapped k - ∝ n_wire k ↕ a_swap k ⟷ n_cup_unswapped k. -Proof. - intros k. - destruct k; [prop_perm_eq|]. - simpl a_swap. - rewrite nwire_stack_distr_compose_r, compose_assoc. - rewrite (n_cup_unswapped_split_stack_n_wire_bot —). - rewrite top_to_bottom_cup_flip. - rewrite <- (cast_cast_eq _ _ (k + 1 + (k + 1)) 0 (1 + k + (1 + k)) 0). - preplace (— ↕ n_wire 1) with (n_wire 1 ↕ —) by prop_perm_eq. - rewrite <- (n_cup_unswapped_split_stack_n_wire_top (top_to_bottom k) —). - rewrite (cast_compose_mid_contract _ (1 + k + (1 + k))). - rewrite cast_n_cup_unswapped, cast_stack_distribute, cast_n_wire by lia. - rewrite <- compose_assoc, compose_stack_n_wire_comm, compose_assoc. - rewrite bottom_to_top_cup_flip. - rewrite <- compose_assoc. - apply compose_simplify; [|easy]. - rewrite <- stack_compose_distr. - rewrite nwire_removal_r. - apply stack_simplify; prop_perm_eq. - solve_modular_permutation_equalities. - Unshelve. - all: lia. -Qed. - -Lemma n_swap_zxperm : forall n, - ZXperm n (n_swap n). -Proof. - induction n; simpl; auto with zxperm_db. -Qed. - -#[export] Hint Resolve n_swap_zxperm : zxperm_db. - -Lemma perm_of_n_swap : forall n, - perm_of_zx (n_swap n) = fun k => if n <=? k then k else (n - S k)%nat. -Proof. - (* destruct n; [simpl; solve_modular_permutation_equalities|]. *) - induction n; simpl perm_of_zx; cleanup_perm_of_zx; - [|rewrite IHn]; solve_modular_permutation_equalities. -Qed. - -#[export] Hint Rewrite perm_of_n_swap : perm_of_zx_cleanup_db. - -Lemma n_swap_cup_flip : forall k, - n_swap k ↕ n_wire k ⟷ n_cup_unswapped k - ∝ n_wire k ↕ n_swap k ⟷ n_cup_unswapped k. -Proof. - intros k. - (* destruct k; [prop_perm_eq|]. *) - induction k; - [prop_perm_eq | ]. - (* [apply compose_simplify; [prop_perm_eq | easy] | ]. *) - simpl (n_swap (S k)). - rewrite nwire_stack_distr_compose_r, compose_assoc. - rewrite (n_cup_unswapped_split_stack_n_wire_bot —), IHk. - preplace (— ↕ n_wire 1) with (n_wire 1 ↕ —) by prop_perm_eq. - rewrite <- (n_cup_unswapped_split_stack_n_wire_top' (n_swap k) —). - rewrite <- compose_assoc, compose_stack_n_wire_comm, compose_assoc. - rewrite bottom_to_top_cup_flip, <- compose_assoc. - apply compose_simplify; [|easy]. - rewrite <- stack_compose_distr, nwire_removal_l. - apply stack_simplify; [easy|]. - prop_perm_eq. - solve_modular_permutation_equalities. - Unshelve. - all: try lia. -Qed. - - - - - - -Local Open Scope ZX_scope. - - - -Lemma n_yank_1_l_helper_helper : forall n, - (⊃) ⊤ ↕ n_wire n ⟷ (— ↕ n_wire (1 + n)) ∝ n_wire n ⟷ (⊂ ↕ n_wire n). -Proof. - intros n. - simpl. - rewrite nwire_removal_l, nwire_removal_r. - easy. -Qed. + to_Cat. + partners_rw_to_Cat compose_empty_r. + Abort. -Lemma n_yank_1_l_helper : forall n prf1 prf2 prf3 prf4 prf5 prf6 prf7 prf8, - cast (S n + (n + n)) (1 + n + S n + S n) prf1 prf2 (n_wire (1 + n) ↕ (n_wire n ↕ ⊃ ↕ n_wire n) ⊤) - ⟷ (((cast (S n + S n) 2 prf3 prf4 (— ↕ n_cup_unswapped n ↕ —)) ⟷ ⊃) ↕ (— ↕ n_wire n)) - ∝ cast _ _ prf7 prf8 ( - (— ↕ (n_wire n ↕ n_wire n) ↕ n_wire n) ⟷ (cast (1 + (n + n) + n) (1 + 0 + n) prf5 prf6 (— ↕ (n_cup_unswapped n) ↕ n_wire n) ) - ⟷ (— ↕ ⊂ ↕ n_wire n) ⟷ (⊃ ↕ — ↕ n_wire n) - ). -Proof. - intros. - rewrite nwire_stack_distr_compose_r. - rewrite <- compose_assoc. - rewrite (cast_compose_mid_contract _ (2 + 1 + n)). - apply compose_simplify; [| - rewrite stack_assoc, cast_cast_eq, cast_id; easy]. - rewrite 2!stack_transpose, n_wire_transpose. - rewrite 2!stack_assoc', (stack_assoc — (n_wire n) (n_wire n)). - simpl_permlike_zx. - rewrite cast_stack_l, cast_cast_eq. - rewrite (stack_assoc _ ((⊃) ⊤) (n_wire n)). - rewrite (prop_iff_double_cast (S (n + n) + (0 + n)) (1 + (1 + (1 + n)))). - rewrite (cast_compose_mid_contract _ ((n + n) + 1 + (1 + (1 + n)))). - - rewrite 2!cast_cast_eq. - rewrite cast_stack_distribute. - rewrite (cast_stack_l (mTop':=2)). - rewrite (stack_assoc _ —). - rewrite 2!cast_cast_eq. - rewrite (cast_stack_distribute (o':= 1 + (1 + n))). - rewrite <- stack_compose_distr. - - simpl_permlike_zx. - rewrite n_yank_1_l_helper_helper. - rewrite <- (nwire_removal_r (— ↕ n_cup_unswapped n)). - rewrite stack_compose_distr. - rewrite compose_assoc. - apply compose_simplify; [ - bundle_wires; prop_perm_eq|]. - cleanup_zx. - enough (Hrw : — ↕ n_cup_unswapped n ↕ n_wire n ⟷ (— ↕ ⊂ ↕ n_wire n) - ∝ @cast (1 + (n + n) + n) (1 + 2 + n) (1 + (n + n) + 0 + n) (1 + 0 + 2 + n) - (ltac:(lia)) (ltac:(lia)) ( - — ↕ n_cup_unswapped n ↕ ⦰ ↕ n_wire n ⟷ (— ↕ ⦰ ↕ ⊂ ↕ n_wire n))). - - rewrite Hrw. - repeat rewrite <- stack_compose_distr. - rewrite 2!nwire_removal_l. - rewrite (stack_assoc' _ ⊂ (n_wire n)). - simpl_casts. - do 2 (apply stack_simplify; [|easy]). - apply stack_simplify; cleanup_zx; easy. - - - cleanup_zx. simpl_permlike_zx. - rewrite (cast_compose_mid_contract _ (1 + 0 + n)), cast_id_eq. - apply compose_simplify; [|easy]. - rewrite cast_stack_l, cast_cast_eq, cast_id_eq. - easy. - - Unshelve. all: lia. -Qed. - - -Lemma n_yank_1_l : forall n, - (— ↕ n_wire n) ↕ ((n_cup_unswapped (S n)) ⊤) ⟷ zx_inv_associator ⟷ ((n_cup_unswapped (S n)) ↕ (— ↕ n_wire n)) - ∝ — ↕ (n_wire n ↕ ((n_cup_unswapped n) ⊤) ⟷ zx_inv_associator ⟷ ((n_cup_unswapped n) ↕ n_wire n)). +Lemma helper_tester_2 : forall {n m} (zx0 : ZX n m) (zx1 : ZX m O), + zx0 ⟷ zx1 ⟷ ⦰ ∝ zx0 ⟷ zx1. Proof. - intros n. - rewrite n_cup_unswapped_grow_l at 1. - rewrite compose_transpose. - rewrite n_cup_unswapped_grow_r. - unfold zx_inv_associator. - simpl_permlike_zx. - (* bundle_wires. *) - rewrite nwire_stack_distr_compose_l. - rewrite (cast_compose_mid_contract _ ((S n + (n + n)))). - rewrite (cast_stack_distribute). - simpl_permlike_zx. - rewrite compose_assoc. - simpl_casts. - rewrite n_yank_1_l_helper, cast_id_eq. - rewrite compose_assoc. - rewrite <- 3!stack_compose_distr. - rewrite yank_l. - simpl (n_wire (1 + n)). - rewrite (stack_assoc —). - rewrite cast_id_eq. - rewrite wire_to_n_wire, nwire_removal_r, <- wire_to_n_wire. - rewrite (stack_assoc —). - rewrite cast_id_eq. - rewrite <- n_wire_add_stack. - rewrite 2!nwire_removal_l. - (* rewrite <- (wire_stack_distr_compose_l _ _ _ (n_cup_unswapped n ↕ n_wire n) (n_wire n)). *) - replace (S n + (n + n))%nat with (1 + (n + (n + n)))%nat by easy. - rewrite (wire_stack_distr_compose_l (n + 0)). - simpl_casts. - rewrite wire_to_n_wire at 5. - rewrite <- n_wire_add_stack, nwire_removal_r. - rewrite (prop_iff_double_cast (1 + (n + 0)) (0 + (1 + n))). - rewrite 2!(cast_compose_mid_contract _ (1 + (n + n + n))). - rewrite 2!cast_cast_eq, 2!cast_id_eq. - easy. - Unshelve. all: try easy; auto with arith; lia. -Qed. - - -Lemma n_yank_l_unswapped : forall n {prf1 prf2}, - (n_wire n) ↕ ((n_cup_unswapped n) ⊤) ⟷ zx_inv_associator ⟷ ((n_cup_unswapped n) ↕ (n_wire n)) - ∝ cast _ _ prf1 prf2 (n_wire n). -Proof. intros. - induction n. - - unfold zx_inv_associator, n_cup_unswapped. - prop_perm_eq. - - simpl (n_wire S n). - rewrite n_yank_1_l, IHn. - prop_perm_eq. - Unshelve. - all: auto with arith. + apply compose_empty_r. Qed. -Lemma compose_zx_inv_associator_r : forall {n0 n m o} (zx : ZX n0 (n + (m + o))) prf1 prf2, - zx ⟷ zx_inv_associator ∝ cast n0 (n + m + o) prf1 prf2 zx. -Proof. - intros. - rewrite (prop_iff_double_cast n0 (n + (m + o))). - rewrite (cast_compose_mid_contract _ (n + (m + o))). - rewrite cast_cast_eq, 2!cast_id_eq. - rewrite <- (nwire_removal_r zx) at 2. - apply compose_simplify; [easy | prop_perm_eq]. - Unshelve. - all: auto with arith. -Qed. - -Lemma n_yank_l_unswapped': forall n {prf1 prf2}, - cast n (n+n+n) prf1 prf2 (n_wire n ↕ (n_cup_unswapped n) ⊤) ⟷ ((n_cup_unswapped n) ↕ (n_wire n)) - ∝ n_wire n. -Proof. - intros. - rewrite (prop_iff_double_cast (n + 0) (0 + n)). - rewrite <- n_yank_l_unswapped. - rewrite compose_zx_inv_associator_r. - rewrite (cast_compose_mid_contract _ (n + n + n)), cast_id_eq, cast_cast_eq. - easy. - Unshelve. - all: auto with arith. -Qed. - -Lemma n_swap_grow_r' : forall n prf1 prf2, - n_swap (S n) ∝ top_to_bottom (S n) ⟷ cast (S n) (S n) prf1 prf2 (n_swap n ↕ —). -Proof. - intros. - prop_perm_eq. - solve_modular_permutation_equalities. -Qed. - -Lemma n_swap_n_swap : forall n, - n_swap n ⟷ n_swap n ∝ n_wire n. -Proof. - prop_perm_eq. - solve_modular_permutation_equalities. -Qed. - -#[export] Hint Rewrite n_swap_n_swap : perm_inv_db. - - -Lemma n_cup_unswapped_n_swap_n_swap : forall n, - n_swap n ↕ n_swap n ⟷ n_cup_unswapped n ∝ n_cup_unswapped n. -Proof. - intros n. - rewrite <- (nwire_removal_l (n_swap n)) at 1. - rewrite <- (nwire_removal_r (n_swap n)) at 2. - rewrite stack_compose_distr. - rewrite compose_assoc, n_swap_cup_flip, <- compose_assoc. - rewrite <- stack_compose_distr, n_swap_n_swap. - rewrite nwire_removal_l, <- n_wire_add_stack, nwire_removal_l. - easy. -Qed. - -Lemma n_cup_inv_n_swap_n_wire' : forall n, n_cup n ∝ n_wire n ↕ n_swap n ⟷ n_cup_unswapped n. -Proof. - intros n. - unfold n_cup. - apply n_swap_cup_flip. -Qed. - -Lemma n_swap_transpose : forall n, - (n_swap n) ⊤ ∝ n_swap n. -Proof. - intros n. - prop_perm_eq. - perm_eq_by_WF_inv_inj (perm_of_zx (n_swap n)) n; - [apply zxperm_permutation | apply perm_of_zx_WF - | rewrite perm_of_transpose_is_linv | ]; - auto with zxperm_db. - cleanup_perm_of_zx. - solve_modular_permutation_equalities. -Qed. - -Lemma n_yank_l : forall n {prf1 prf2}, - cast n (n + n + n) prf1 prf2 (n_wire n ↕ n_cap n) - ⟷ (n_cup n ↕ n_wire n) ∝ n_wire n. -Proof. - intros. - unfold n_cap. - rewrite n_cup_inv_n_swap_n_wire' at 2. - unfold n_cup. - (* rewrite n_cup_inv_n_swap_n_wire. *) - simpl. - rewrite n_wire_transpose, n_swap_transpose. - rewrite nwire_stack_distr_compose_l. - rewrite stack_assoc'. - rewrite (cast_compose_mid_contract n (n + n + n) (n + n + n)). - rewrite cast_cast_eq, cast_id_eq. - rewrite nwire_stack_distr_compose_r. - rewrite <- compose_assoc, (compose_assoc _ _ (_ ↕ n_swap n ↕ _)). - rewrite <- 2!stack_compose_distr, n_swap_n_swap, nwire_removal_l. - rewrite <- 2!n_wire_add_stack, nwire_removal_r. - apply n_yank_l_unswapped'. - Unshelve. - all: lia. -Qed. - -Lemma n_yank_r : forall n {prf1 prf2 prf3 prf4}, - cast n n prf3 prf4 (cast n (n + (n + n)) prf1 prf2 (n_cap n ↕ n_wire n) - ⟷ (n_wire n ↕ n_cup n)) ∝ n_wire n. -Proof. - intros. - apply transpose_diagrams. - rewrite cast_transpose, compose_transpose. - rewrite (cast_compose_mid_contract _ (n+n+n)). - rewrite cast_transpose, cast_cast_eq, cast_id_eq. - rewrite 2!stack_transpose. - rewrite n_wire_transpose. - unfold n_cap. - rewrite Proportional.transpose_involutive. - apply n_yank_l. - Unshelve. - all: lia. -Qed. - -Lemma zx_triangle_1 : forall n, - zx_inv_right_unitor ⟷ (n_wire n ↕ n_cap n) ⟷ zx_inv_associator - ⟷ (n_cup n ↕ n_wire n) ⟷ zx_left_unitor ∝ n_wire n. -Proof. - intros. - unfold zx_inv_right_unitor. - unfold zx_inv_associator. - unfold zx_left_unitor. - simpl_casts. cleanup_zx. - rewrite cast_compose_l. - simpl_casts. cleanup_zx. - rewrite cast_compose_r. - cleanup_zx. simpl_casts. - rewrite n_yank_l. - reflexivity. -Qed. - -Lemma zx_triangle_2 : forall n, - zx_inv_left_unitor ⟷ (n_cap n ↕ n_wire n) ⟷ zx_associator - ⟷ (n_wire n ↕ n_cup n) ⟷ zx_right_unitor ∝ n_wire n. -Proof. - intros. - unfold zx_inv_left_unitor. - unfold zx_associator. - unfold zx_right_unitor. - simpl_casts. cleanup_zx. - rewrite cast_compose_r. - simpl_casts. cleanup_zx. - rewrite cast_compose_r. - simpl_casts. cleanup_zx. - rewrite n_yank_r. - reflexivity. -Qed. - -#[export] Instance ZXCompactClosedCategory : CompactClosedCategory nat := { - dual n := n; - unit n := n_cap n; - counit n := n_cup n; - triangle_1 := zx_triangle_1; - triangle_2 := zx_triangle_2; -}. - - - -#[export] Instance ZXDaggerMonoidalCategory : DaggerMonoidalCategory nat := { - dagger_tensor_compat := @zx_dagger_tensor_compat; - - associator_unitary := fun A B M => - conj (@zx_associator_unitary_r A B M) (@zx_associator_unitary_l A B M); - (* associator_unitary_l := @zx_associator_unitary_l; *) - left_unitor_unitary := fun A => - conj (@zx_left_unitor_unitary_r A) (@zx_left_unitor_unitary_l A); - (* left_unitor_unitary_l := @zx_left_unitor_unitary_l; *) - right_unitor_unitary := fun A => - conj (@zx_right_unitor_unitary_r A) (@zx_right_unitor_unitary_l A); - (* right_unitor_unitary_l := @zx_right_unitor_unitary_l; *) -}. - -#[export] Instance ZXDaggerBraidedMonoidalCategory : DaggerBraidedMonoidalCategory nat := {}. - -#[export] Instance ZXDaggerSymmetricMonoidalCategory : DaggerSymmetricMonoidalCategory nat := {}. - - - - - -Ltac not_evar' v := - not_evar v; try (let v := eval unfold v in v in - tryif not_evar v then idtac else fail 1). - -Ltac subst_evars := - repeat match goal with - | x := ?y : nat |- _ => subst x - end. - -Ltac evarify_1func_once f := - match goal with - |- context[@f ?a ?A] => - not_evar' a; let x:= fresh in evar (x : nat); - replace (@f a A) with (@f x A) by - ( replace (@f a) with (@f x) by (replace a with x by shelve; reflexivity); - reflexivity) - end. - -Ltac evarify_1func f := repeat (evarify_1func_once f). -Ltac evarify_1func' f := repeat (evarify_1func_once f); subst_evars. - -Ltac evarify_2func_once f := - match goal with - |- context[@f ?a ?b ?A] => - not_evar' a; not_evar' b; - let x:= fresh in let y := fresh in - evar (x : nat); evar (y : nat); - replace (@f a b A) with (@f x y A) by - ( replace (@f a) with (@f x) by (replace a with x by shelve; reflexivity); - replace (@f x b) with (@f x y) by (replace b with y by shelve; reflexivity); - reflexivity) - end. - -Ltac evarify_2func f := repeat (evarify_2func_once f). -Ltac evarify_2func' f := repeat (evarify_2func_once f); subst_evars. - - -Lemma nat_equality_proof_equality : forall (n m : nat) (H0 H1 : n = m), - H0 = H1. +Lemma test_part_rw_long : forall {nIn nOut} + (zx0 zx2 zx3 zx4 zx5 zx6 zx7: ZX nIn nIn) + (zx : ZX nIn 0) (zx1 : ZX 0 nOut), + zx0 ⟷ (zx2 ⟷ (zx3 ⟷ (zx4 ⟷ (zx5 ⟷ + (zx6 ⟷ (zx7 ⟷ (zx ⟷ (⦰ ⟷ zx1)))))))) ∝ zx ⟷ zx1. Proof. - intros n. intros. subst n. - now rewrite Logic.Eqdep_dec.UIP_refl_nat. -Qed. - -Ltac evarify_cast_once := - match goal with - |- context[@cast ?n ?m ?n' ?m' ?prfn ?prfm ?zx] => - not_evar' prfm; not_evar' prfm; - let x:= fresh in let y := fresh in - evar (x : n = n'); evar (y : m = m'); - replace (@cast n m n' m' prfn prfm) with (@cast n m n' m' x y) by - ( replace (@cast n m n' m' prfn) with (@cast n m n' m' x) by - (replace prfn with x by (apply nat_equality_proof_equality); reflexivity); - replace (@cast n m n' m' x prfm) with (@cast n m n' m' x y) by - (replace prfm with y by (apply nat_equality_proof_equality); reflexivity); - reflexivity) - end. - -Ltac subst_eqs_evars := - repeat match goal with - | x := ?y : _ = _ |- _ => subst x - end. - -Ltac evarify_cast := repeat (evarify_cast_once). -Ltac evarify_cast' := evarify_cast; subst_eqs_evars. - - -(* For example: *) -Lemma temp_id {n} : forall (zx : ZX n n), - cast n n eq_refl eq_refl zx ∝ zx. -Admitted. - -Lemma gen_id {n} : forall (zx : ZX n n) prfn prfm, - cast n n prfn prfm zx ∝ zx. intros. - evarify_cast'. (* Replace all proofs with evars *) - rewrite temp_id. - reflexivity. -Qed. - - - - - - -From ViCaR Require Import CategoryAutomation. - -(* Ltac varify_nats f := - match ty *) + to_Cat. + partners_rw_to_Cat compose_empty_r. + Abort. Lemma test : forall {n m o} (zx0 : ZX n m) (zx1 : ZX m o), - (zx0 ⟷ zx1) ↕ — ∝ (zx0 ↕ —) ⟷ (zx1 ↕ —). + (zx0 ⟷ zx1) ↕ — ⟷ zx_braiding ∝ (zx0 ↕ —) ⟷ (zx1 ↕ —) ⟷ zx_braiding. Proof. (* setoid_rewrite wire_to_n_wire. *) intros. rewrite wire_to_n_wire. - fold_all_categories; - fold_all_monoidal_categories. - match goal with - |- (?T ≃ ?T')%Cat => - strong_fencepost_no_id T - end. - easy. + to_Cat. + strong_fencepost_LHS. + Fail easy. + cat_easy. Qed. Lemma test2 : forall {n m o p} (zx0 : ZX n m) (zx1 : ZX o p), @@ -1721,11 +510,8 @@ Proof. (* setoid_rewrite wire_to_n_wire. *) intros. to_Cat. - (* fold_all_categories; - fold_all_monoidal_categories. *) - match goal with - |- (?T ≃ ?T')%Cat => - strong_fencepost_no_id T - end. + strong_fencepost_LHS. easy. -Qed. \ No newline at end of file +Qed. + +End Testing. \ No newline at end of file