diff --git a/ViCaR/CategoryAutomation.v b/ViCaR/CategoryAutomation.v index bc109ee..bc5cffa 100644 --- a/ViCaR/CategoryAutomation.v +++ b/ViCaR/CategoryAutomation.v @@ -404,7 +404,7 @@ Ltac to_Cat_in H := -(* Section on Fenceposting *) +(* Section on foliateing *) Ltac tensor_free f := try match f with @@ -574,14 +574,14 @@ Ltac merge_stacked_composition gh := end end in merge_stacked_composition gh. -Ltac weak_fencepost_form_debug f := - let rec weak_fencepost f := +Ltac weak_foliate_form_debug f := + let rec weak_foliate f := match f with | @compose ?C ?cC _ _ _ ?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 Ng := weak_foliate g in + let Nh := weak_foliate h in let _ := match goal with _ => idtac "... getting" g "∘" h "into" end in let res := right_associate_term (cC.(compose) Ng Nh) in @@ -591,8 +591,8 @@ Ltac weak_fencepost_form_debug f := | @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 Ng := weak_foliate g in + let Nh := weak_foliate h in let _ := match goal with _ => idtac "... getting" g "⊗" h "into" end in let res := merge_stacked_composition ((mC.(mor_tensor) Ng Nh)%Cat) in @@ -604,23 +604,23 @@ Ltac weak_fencepost_form_debug f := idtac "INFO:" f "is const or unsupported" end in constr:(f) end - in weak_fencepost f. + in weak_foliate f. -Ltac weak_fencepost_form f := - let rec weak_fencepost f := +Ltac weak_foliate_form f := + let rec weak_foliate f := match f with | @compose ?C ?cC _ _ _ ?g ?h => - let Ng := weak_fencepost g in - let Nh := weak_fencepost h in + let Ng := weak_foliate g in + let Nh := weak_foliate h in 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 Ng := weak_foliate g in + let Nh := weak_foliate h in merge_stacked_composition ((mC.(mor_tensor) Ng Nh)%Cat) | _ => (* f "is const or unsupported" *) constr:(f) end - in weak_fencepost f. + in weak_foliate f. Section HelperLemmas. @@ -827,33 +827,33 @@ Ltac show_equiv_merge_stacked_composition gh := end end in show_equiv_merge_stacked_composition gh. -(* Shows the goal f ≃ weak_fencepost_form f by mirroring the code - path of weak_fencepost_form with `apply`s. *) -Ltac show_equiv_weak_fencepost_form f := - let weak_fencepost := weak_fencepost_form in - let rec show_equiv_weak_fencepost_form f := +(* Shows the goal f ≃ weak_foliate_form f by mirroring the code + path of weak_foliate_form with `apply`s. *) +Ltac show_equiv_weak_foliate_form f := + let weak_foliate := weak_foliate_form in + let rec show_equiv_weak_foliate_form f := match f with | @compose ?C ?cC _ _ _ ?g ?h => - let Ng := weak_fencepost g in - let Nh := weak_fencepost h in + let Ng := weak_foliate g in + let Nh := weak_foliate h 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_weak_foliate_form g) + ltac:(show_equiv_weak_foliate_form 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 Ng := weak_foliate g in + let Nh := weak_foliate h 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_weak_foliate_form g) + ltac:(show_equiv_weak_foliate_form h) ltac:(show_equiv_merge_stacked_composition ((mC.(mor_tensor) Ng Nh))%Cat)) | _ => (* f "is const or unsupported" *) (* constr:(f) *) reflexivity end - in show_equiv_weak_fencepost_form f. + in show_equiv_weak_foliate_form f. (* TODO: Generalize these to fold_compose base *) (* If f = f0 ∘ (f1 ∘ (...)), this gives f0 ⊗ id_ B ∘ (f1 ⊗ id_ B ∘ (...)) @@ -938,10 +938,10 @@ Ltac unfold_tensor_stack_no_id f := end in unfold_tensor_stack f. -(* Returns the strong fencepost term of a weakly fenceposted term +(* Returns the strong foliate term of a weakly foliateed term (in fact, not even requiring the term be right-associated, though - the resulting fencepost will be. )*) -Ltac strong_fencepost_form_of_weak f := + the resulting foliate will be. )*) +Ltac strong_foliate_form_of_weak f := let rec strong_fence f := lazymatch f with | (?g ∘ ?h)%Cat => @@ -954,7 +954,7 @@ Ltac strong_fencepost_form_of_weak f := in strong_fence f. (* Additionally avoids taking id ⊗ id to id ⊗ id ∘ id ⊗ id and similar *) -Ltac strong_fencepost_form_of_weak_no_id f := +Ltac strong_foliate_form_of_weak_no_id f := let rec strong_fence f := lazymatch f with | (?g ∘ ?h)%Cat => @@ -1115,8 +1115,8 @@ Ltac show_equiv_unfold_tensor_stack_no_id_debug f := in show_unfold f. -Ltac show_equiv_strong_fencepost_form_of_weak f := - let strong_fence := strong_fencepost_form_of_weak in +Ltac show_equiv_strong_foliate_form_of_weak f := + let strong_fence := strong_foliate_form_of_weak in let rec show_strong_fence f := lazymatch f with | (?g ∘ ?h)%Cat => @@ -1137,8 +1137,8 @@ Ltac show_equiv_strong_fencepost_form_of_weak f := in show_strong_fence f. -Ltac show_equiv_strong_fencepost_form_of_weak_no_id f := - let strong_fence := strong_fencepost_form_of_weak_no_id in +Ltac show_equiv_strong_foliate_form_of_weak_no_id f := + let strong_fence := strong_foliate_form_of_weak_no_id in let rec show_strong_fence f := lazymatch f with | (?g ∘ ?h)%Cat => @@ -1158,8 +1158,8 @@ Ltac show_equiv_strong_fencepost_form_of_weak_no_id f := end in show_strong_fence f. -Ltac show_equiv_strong_fencepost_form_of_weak_no_id_debug f := - let strong_fence := strong_fencepost_form_of_weak_no_id in +Ltac show_equiv_strong_foliate_form_of_weak_no_id_debug f := + let strong_fence := strong_foliate_form_of_weak_no_id in let rec show_strong_fence f := lazymatch f with | (?g ∘ ?h)%Cat => @@ -1179,43 +1179,43 @@ Ltac show_equiv_strong_fencepost_form_of_weak_no_id_debug f := end in show_strong_fence f. -Ltac weak_fencepost f := - let wf := weak_fencepost_form f in +Ltac weak_foliate f := + let wf := weak_foliate_form f in let H := fresh in - assert (H: (f ≃ wf)%Cat) by (show_equiv_weak_fencepost_form f); + assert (H: (f ≃ wf)%Cat) by (show_equiv_weak_foliate_form f); setoid_rewrite H; clear H. -Ltac strong_fencepost f := - let wf := weak_fencepost_form f in - let sf := strong_fencepost_form_of_weak wf in +Ltac strong_foliate f := + let wf := weak_foliate_form f in + let sf := strong_foliate_form_of_weak wf in let H := fresh in assert (H: (f ≃ sf)%Cat) by ( transitivity wf; - [ show_equiv_weak_fencepost_form f - | show_equiv_strong_fencepost_form_of_weak wf]); + [ show_equiv_weak_foliate_form f + | show_equiv_strong_foliate_form_of_weak wf]); setoid_rewrite H; clear H. -Ltac strong_fencepost_no_id f := - let wf := weak_fencepost_form f in - let sf := strong_fencepost_form_of_weak_no_id wf in +Ltac strong_foliate_no_id f := + let wf := weak_foliate_form f in + let sf := strong_foliate_form_of_weak_no_id wf in let H := fresh in assert (H: (f ≃ sf)%Cat) by ( transitivity wf; - [ show_equiv_weak_fencepost_form f - | show_equiv_strong_fencepost_form_of_weak_no_id wf]); + [ show_equiv_weak_foliate_form f + | show_equiv_strong_foliate_form_of_weak_no_id wf]); setoid_rewrite H; clear H. -Ltac strong_fencepost_no_id_debug f := - let wf := weak_fencepost_form f in - let sf := strong_fencepost_form_of_weak_no_id wf in +Ltac strong_foliate_no_id_debug f := + let wf := weak_foliate_form f in + let sf := strong_foliate_form_of_weak_no_id wf in let H := fresh in assert (H: (f ≃ sf)%Cat) by ( transitivity wf; - [ show_equiv_weak_fencepost_form f - | show_equiv_strong_fencepost_form_of_weak_no_id_debug wf]); + [ show_equiv_weak_foliate_form f + | show_equiv_strong_foliate_form_of_weak_no_id_debug wf]); setoid_rewrite H; clear H. @@ -1887,6 +1887,11 @@ Ltac partner_LRHS t s := let func := partner_in_term t s in apply_to_LRHS func. +Ltac partner t s := + first + [ partner_LHS t s + | partner_RHS t s]. + Ltac gen_partner_LHS test := let func := gen_partner_in_term test in apply_to_LHS func. @@ -1919,21 +1924,29 @@ 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 weak_foliate_LHS := apply_to_LHS weak_foliate. +Ltac weak_foliate_RHS := apply_to_RHS weak_foliate. +Ltac weak_foliate_LRHS := apply_to_LRHS weak_foliate. -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 strong_foliate_LHS := apply_to_LHS strong_foliate_no_id. +Ltac strong_foliate_RHS := apply_to_RHS strong_foliate_no_id. +Ltac strong_foliate_LRHS := apply_to_LRHS strong_foliate_no_id. Ltac cancel_ids := cancel_ids_LRHS. Ltac cancel_isos := cancel_lrisos_LRHS. +Ltac weak_foliate_goal := weak_foliate_LRHS. +Ltac strong_foliate_goal := strong_foliate_LRHS. + +Ltac foliate := (fun t => ltac:(strong_foliate t; rewrite ?tensor_id; cancel_ids)). +Ltac foliate_LHS := strong_foliate_LHS. +Ltac foliate_RHS := strong_foliate_RHS. +Ltac foliate_LRHS := strong_foliate_LRHS. Ltac cat_cleanup := repeat (cancel_isos; cancel_ids). -Ltac cat_easy := cat_cleanup; rassoc_LRHS; easy || rewrite !tensor_id; easy. +Ltac cat_easy := cat_cleanup; try easy; rassoc_LRHS; try easy; + rewrite ?tensor_id; try easy; weak_foliate_LRHS; easy. @@ -1941,7 +1954,7 @@ 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) := +Tactic Notation "assoc_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'); @@ -1961,15 +1974,15 @@ Tactic Notation "partners_rw" open_constr(lem) "within" constr(term) := clear e end. -Tactic Notation "partners_rw" open_constr(lem) := +Tactic Notation "assoc_rw" open_constr(lem) := match goal with |- (?LHS ≃ ?RHS)%Cat => first [ - partners_rw lem within LHS - | partners_rw lem within RHS ] + assoc_rw lem within LHS + | assoc_rw lem within RHS ] end. -Tactic Notation "partners_rw_to_Cat" open_constr(lem) "within" constr(term) := +Tactic Notation "assoc_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; @@ -1991,17 +2004,34 @@ Tactic Notation "partners_rw_to_Cat" open_constr(lem) "within" constr(term) := end. -Tactic Notation "partners_rw_to_Cat" open_constr(lem) := +Tactic Notation "assoc_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 ] + assoc_rw_to_Cat lem within LHS + | assoc_rw_to_Cat lem within RHS ] end. + Section Testing. Local Open Scope Cat_scope. + +Lemma TODO_fix_assoc_rw + {C : Type} {cC : Category C} {cCh : CategoryCoherence cC} + {mC : MonoidalCategory cC} {mCh : MonoidalCategoryCoherence mC} + {bC : BraidedMonoidalCategory mC} {bCh : BraidedMonoidalCategoryCoherence bC} + {A B M N P Q R : C} + (h : A ~> B) (f : B ~> M) (g : B ~> N) + (m : Q ~> R) (k : R × N ~> Q) (j : M ~> A) n : + α_ _, _, _ ∘ m ⊗ ((h ∘ f) ⊗ (h ∘ g) ∘ β_ _, _) ∘ (α_ _, _, _)^-1 ∘ k ⊗ j + ≃ n. +Proof. + LHS rassoc. + + Fail assoc_rw braiding_natural. (* ??? *) + Abort. + Variables (C : Type) (cC cC' cC'' : Category C) (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'') @@ -2022,41 +2052,56 @@ Existing Instance mC0'. Existing Instance mC1'. Existing Instance mC0''. Existing Instance mC1''. +(* Goal (f ⊗ g ⊗ h) ≃ (f ⊗ g ⊗ h). *) +(* easy. *) + + +(* Goal (f ⊗ (g ⊗ h)) ≃ (f ⊗ (g ⊗ h)). *) +(* easy. *) -Lemma test_weak_fencepost : forall +(* Goal (f ⊗ (g ⊗ h)) ≃ (f ⊗ (g ⊗ h)). *) +(* easy. *) + +(* Goal (f ∘ g) ⊗ h ≃ f ⊗ h ∘ g ⊗ id_ M. *) +(* Admitted. *) + +(* Goal (f ∘ g) ⊗ h ≃ f ⊗ id_ A ∘ (id_ B ⊗ h ∘ g ⊗ id_ M). *) +(* Admitted. *) + +Lemma test_weak_foliate : 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 + |- ?T ≃ _ => weak_foliate T + (* let wf := weak_foliate_form T in let H := fresh in - assert (H : T ≃ wf) by show_equiv_weak_fencepost_form T *) + assert (H : T ≃ wf) by show_equiv_weak_foliate_form T *) (* setoid_rewrite H *) end. easy. Qed. -Lemma test_strong_fencepost : forall +Lemma test_strong_foliate : 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 ≃ ?T' => strong_fencepost T; strong_fencepost T' + |- ?T ≃ ?T' => strong_foliate T; strong_foliate T' end. easy. Qed. -Lemma test_strong_fencepost_no_id_1 : forall +Lemma test_strong_foliate_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. intros. match goal with - |- ?T ≃ ?T' => strong_fencepost_no_id T; strong_fencepost_no_id T' + |- ?T ≃ ?T' => strong_foliate_no_id T; strong_foliate_no_id T' end. easy. Qed. @@ -2065,27 +2110,27 @@ Qed. -Lemma test_strong_fencepost_no_id_2' : forall +Lemma test_strong_foliate_no_id_2' : forall {a b m n o} (f : a ~> b) (g : m ~> n) (h : n ~> o), f ⊗ (g ∘ id_ _ ∘ h ∘ id_ _) ⊗ (id_ a ⊗ id_ b) ≃ f ⊗ g ⊗ (id_ a ⊗ id_ b) ∘ ((id_ b ⊗ h) ⊗ (id_ a ⊗ id_ b)). Proof. intros. - do 2 partners_rw right_unit. - LHS weak_fencepost. + do 2 assoc_rw right_unit. + LHS weak_foliate. rewrite tensor_id. easy. Qed. -Lemma test_strong_fencepost_no_id_2 : forall +Lemma test_strong_foliate_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. + assoc_rw right_unit. + LHS weak_foliate. cat_easy. Qed. @@ -2163,7 +2208,7 @@ test_partnered_nf f g (id_ _ ∘ id_ _ ∘ id_ _ ∘ f ∘ g ∘ i ∘ id_ _ ∘ Local Ltac test_show_unfold_no_id_of_wf f := - let wf := weak_fencepost_form f in + let wf := weak_foliate_form f in let sf := unfold_tensor_stack_no_id wf in (* idtac sf; *) let H := fresh in @@ -2197,10 +2242,10 @@ test_show_unfold_no_id (id_ B ⊗ g0 ⊗ id_ (A × B))%Cat. Local Ltac test_show_st_of_wk f := - let wf := weak_fencepost_form f in - let sf := strong_fencepost_form_of_weak wf in + let wf := weak_foliate_form f in + let sf := strong_foliate_form_of_weak wf in let H := fresh in - assert (H: (wf ≃ sf)%Cat) by (show_equiv_strong_fencepost_form_of_weak wf); + assert (H: (wf ≃ sf)%Cat) by (show_equiv_strong_foliate_form_of_weak wf); clear H. test_show_st_of_wk f. @@ -2259,8 +2304,8 @@ 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 - let sf := strong_fencepost_form_of_weak wf in + let wf := weak_foliate_form f in + let sf := strong_foliate_form_of_weak wf in (* idtac wf "=~=>" sf. *) (* For compile: *) idtac. @@ -2290,8 +2335,8 @@ test_ust ((f ⊗ (f ∘ g) ⊗ (f0 ∘ g0))). Local Ltac test_show_wf f := let H := fresh in - let wf := weak_fencepost_form f in - assert (H: f ≃ wf) by (show_equiv_weak_fencepost_form f); + let wf := weak_foliate_form f in + assert (H: f ≃ wf) by (show_equiv_weak_foliate_form f); clear H. test_show_wf f. @@ -2335,10 +2380,10 @@ test_merge (mC0.(mor_tensor) (mC0.(mor_tensor) f g) (cC.(compose) f0 g0)). Local Ltac test_fence f := - (* let Nf := weak_fencepost_form_debug f in - idtac "fenceposted:" Nf. *) + (* let Nf := weak_foliate_form_debug f in + idtac "foliateed:" Nf. *) (* For compile: *) - let Nf := weak_fencepost_form f in + let Nf := weak_foliate_form f in idtac. test_fence ((f ⊗ g ⊗ h) ∘ id_ _). @@ -2389,68 +2434,3 @@ Qed. End Testing. Local Close Scope Cat_scope. - -Module FutureDirections. - -Local Open Scope Cat_scope. - -Section CatExpr_orig. - -Inductive cat_expr_o {C} `{cC : Category C} - : forall {A B : C}, A ~> B -> Prop := - (* This _might_ want to be to Type instead? *) - | expr_const {A B : C} (f : A ~> B) - : cat_expr_o f - | expr_compose {A B M : C} (f : A ~> M) (g : M ~> B) - : cat_expr_o (f ∘ g) - | expr_tensor {A1 A2 B1 B2 : C} {mC : @MonoidalCategory C cC} - (g : A1 ~> A2) (h : B1 ~> B2) - : cat_expr_o (g ⊗ h). - (* Optionally, - | expr_cast `{ccC: CastCategory C} {A B A' B'} - (HA : A = A') (HB : B = B') (f : A' ~> B') - : cat_expr (cast A B HA HB f) *) - -End CatExpr_orig. - -(* Dependency cycle: - -Require Import RigCategory. - -Section CatExpr_hierarchy. - -Local Open Scope Rig_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). - -Inductive moncat_expr {C} `{mC : MonoidalCategory C} - : forall {A B : C}, A ~> B -> Prop := - | moncat_cat {A B} {f : A ~> B} - : cat_expr f -> moncat_expr f - | moncat_tensor {A1 A2 B1 B2 : C} (g : A1 ~> A2) (h : B1 ~> B2) - : moncat_expr (g ⊗ h) - | moncat_assoc_for {A B M : C} - : moncat_expr (associator (A:=A) (B:=B) (M:=M)).(forward) - | moncat_assoc_rev {A B M : C} - : moncat_expr (associator (A:=A) (B:=B) (M:=M)).(reverse) - (* ... and so on with left_unit, etc. *). - -Inductive rigcat_expr {C} `{mC : PreDistributiveBimonoidalCategory C} - : forall {A B : C}, A ~> B -> Prop := - | rigcat_cat {A B} {f : A ~> B} - : cat_expr f -> rigcat_expr f - | rigcat_plus {A1 A2 B1 B2 : C} (g : A1 ~> A2) (h : B1 ~> B2) - : rigcat_expr (g ⊞ h) - | rigcat_times {A1 A2 B1 B2 : C} (g : A1 ~> A2) (h : B1 ~> B2) - : rigcat_expr (g ⊠ h). - -End CatExpr_hierarchy. *) - -End FutureDirections. - diff --git a/ViCaR/Classes/BraidedMonoidal.v b/ViCaR/Classes/BraidedMonoidal.v index 392d155..46d1d37 100644 --- a/ViCaR/Classes/BraidedMonoidal.v +++ b/ViCaR/Classes/BraidedMonoidal.v @@ -67,15 +67,13 @@ 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. + (* rewrite <- compose_iso_l. *) 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 × M, A) (IdentityIsomorphism A) (β_ B, M)). + rewrite <- (compose_tensor_iso_r' _ (IdentityIsomorphism _)). simpl. rewrite 2!compose_iso_r. 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'. symmetry in hex2. @@ -87,7 +85,6 @@ Proof. apply compose_cancel_r. pose proof (hexagon_1 B A M) as hex1. rewrite assoc, <- compose_iso_l'. - replace (id_ M) with (IdentityIsomorphism M ^-1) by easy. rewrite <- (compose_tensor_iso_l' _ (IdentityIsomorphism _)). simpl. rewrite <- 3!(assoc). @@ -98,4 +95,5 @@ Qed. End BraidedCoherenceRewrites. -Local Close Scope Cat. + +Local Close Scope Cat. \ No newline at end of file diff --git a/ViCaR/Classes/Category.v b/ViCaR/Classes/Category.v index d82caa1..f396044 100644 --- a/ViCaR/Classes/Category.v +++ b/ViCaR/Classes/Category.v @@ -11,6 +11,7 @@ Reserved Notation "f ≃ g" (at level 70). (* \simeq *) Reserved Notation "A ≅ B" (at level 70). (* \cong *) Reserved Notation "'id_' A" (at level 15). +#[universes(polymorphic=yes,cumulative=yes)] Class Category (C : Type) : Type := { morphism : C -> C -> Type where "A ~> B" := (morphism A B); @@ -39,6 +40,7 @@ Notation "f ≃ g" := (c_equiv _%Cat f%Cat g%Cat) Notation "f ∘ g" := (compose _%Cat f%Cat g%Cat) (at level 40, g at next level, left associativity) : Cat_scope. (* \circ *) +#[universes(polymorphic=yes,cumulative=yes)] Class CategoryCoherence {C} (cC : Category C) : Type := { (* to_base_struct_cat := cC; *) @@ -99,11 +101,13 @@ Qed. (** Isomorphism of objects in a category, and equivalent typeclass, with parametric equivalence *) +#[universes(polymorphic=yes)] Definition isomorphic {C : Type} {cC : Category C} (A B : C) := exists (f : A ~> B) (g : B ~> A), is_inverse f g. Arguments isomorphic {_} {_}%Cat (_ _)%Cat. +#[universes(polymorphic=yes,cumulative=yes)] Class Isomorphism {C : Type} {cC : Category C} (A B : C) := { forward : A ~> B; reverse : B ~> A; @@ -182,6 +186,7 @@ Add Parametric Relation {C : Type} {cC : Category C} as isomorphic_equiv_rel. (** Functors, including instances as equivalence & isomorphism parametric morphisms *) +#[universes(polymorphic=yes,cumulative=yes)] Class Functor {C D : Type} (cC: Category C) (cD : Category D) : Type := { obj_map : C -> D; morphism_map {A B : C} : (A ~> B) -> (obj_map A ~> obj_map B); @@ -218,7 +223,7 @@ Proof. Qed. - +#[universes(polymorphic=yes,cumulative=yes)] Class Bifunctor {C1 C2 D : Type} (cC1: Category C1) (cC2 : Category C2) (cD : Category D) := { obj_bimap : C1 -> C2 -> D; @@ -266,6 +271,7 @@ Proof. split; apply id_bimap. Qed. +#[universes(polymorphic=yes)] Definition CommuteBifunctor {C1 C2 D : Type} `{cC1 : Category C1} `{cC2 : Category C2} `{cD : Category D} (F : Bifunctor cC1 cC2 cD) : Bifunctor cC2 cC1 cD := {| @@ -281,6 +287,7 @@ Arguments CommuteBifunctor {_ _ _} {_ _ _}%Cat (_)%Cat /. (** Natural Transformations & Isomorphisms (and the equivalents for Bifunctors) *) +#[universes(polymorphic=yes,cumulative=yes)] Class NaturalTransformation {C D : Type} `{cC: Category C} `{cD : Category D} (F G : Functor cC cD) := { component_map (A : C) : F A ~> G A; @@ -292,6 +299,7 @@ Arguments component_map {_ _} {_ _}%Cat {_ _}%Cat (N)%Cat (_)%Cat : rename. Arguments component_map_natural {_ _} {_ _}%Cat {_ _}%Cat {N}%Cat {_ _}%Cat (f)%Cat : rename. +#[universes(polymorphic=yes,cumulative=yes)] Class NaturalIsomorphism {C D : Type} `{cC: Category C} `{cD : Category D} (F G : Functor cC cD) := { component_iso (A : C) : F A <~> G A; @@ -313,7 +321,7 @@ Definition NaturalTransformation_of_NaturalIsomorphism {C D : Type} |}. - +#[universes(polymorphic=yes,cumulative=yes)] 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; @@ -328,6 +336,7 @@ Arguments component_bimap {_ _ _} {_ _ _}%Cat {_ _}%Cat Arguments component_bimap_natural {_ _ _} {_ _ _}%Cat {_ _}%Cat {N}%Cat {_ _ _ _}%Cat (f1 f2)%Cat : rename. +#[universes(polymorphic=yes,cumulative=yes)] Class NaturalBiIsomorphism {C1 C2 D : Type} `{cC1 : Category C1} `{cC2 : Category C2} `{cD : Category D} (F G : Bifunctor cC1 cC2 cD) := { component_biiso (A1 : C1) (A2 : C2) : F A1 A2 <~> G A1 A2; @@ -655,8 +664,8 @@ Proof. intros; split; symmetry; apply compose_iso_l; easy. Qed. - -#[program] Definition FunctorIsomorphism {C D} {cC : Category C} +#[universes(polymorphic=yes), + 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; @@ -667,7 +676,8 @@ Next Obligation. (proj2 isomorphism_inverse), 2!id_map; easy. Qed. -#[program] Definition BifunctorIsomorphism {C1 C2 D} {cC1 : Category C1} +#[universes(polymorphic=yes), + 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) diff --git a/ViCaR/Classes/Monoidal.v b/ViCaR/Classes/Monoidal.v index d1ff429..d3debea 100644 --- a/ViCaR/Classes/Monoidal.v +++ b/ViCaR/Classes/Monoidal.v @@ -11,6 +11,7 @@ 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 *) +#[universes(polymorphic=yes,cumulative=yes)] Class MonoidalCategory {C : Type} (cC : Category C) : Type := { obj_tensor : C -> C -> C where "x × y" := (obj_tensor x y); @@ -54,7 +55,7 @@ Notation "'ρ⁻¹_' x" := ((right_unitor x).(reverse)) (at level 20, no associativity) : Cat_scope. (* \rho \^- \^1 *) *) - +#[universes(polymorphic=yes,cumulative=yes)] 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); @@ -149,6 +150,62 @@ Qed. End TensorBifunctor. +Section InverseCoherences. + +Context {C : Type} {cC : Category C} {cCh : CategoryCoherence cC} + {mC : MonoidalCategory cC} {mCh : MonoidalCategoryCoherence mC}. + +Lemma invassociator_cohere {A B M N P Q : C} (f : A ~> B) + (g : M ~> N) (h : P ~> Q) : + α_ A, M, P ⁻¹ ∘ ((f ⊗ g) ⊗ h) + ≃ (f ⊗ (g ⊗ h)) ∘ α_ B, N, Q ⁻¹. +Proof. + rewrite <- compose_iso_l', <- assoc, <- compose_iso_r. + symmetry. + apply associator_cohere. +Qed. + +Lemma invleft_unitor_cohere {A B : C} (f : A ~> B) : + λ_ A ⁻¹ ∘ (id_ I ⊗ f) ≃ f ∘ λ_ B⁻¹. +Proof. + rewrite <- compose_iso_l', <- assoc, <- compose_iso_r. + symmetry. + apply left_unitor_cohere. +Qed. + +Lemma invright_unitor_cohere {A B : C} (f : A ~> B) : + ρ_ A ⁻¹ ∘ (f ⊗ id_ I) ≃ f ∘ ρ_ B ⁻¹. +Proof. + rewrite <- compose_iso_l', <- assoc, <- compose_iso_r. + symmetry. + apply right_unitor_cohere. +Qed. + +Lemma inv_triangle' (A B : C) : + (id_ A ⊗ λ_ B) + ≃ α_ A, I, B ^-1 ∘ ρ_ A ⊗ id_ B. +Proof. + rewrite <- compose_iso_l. + apply triangle. +Qed. + +Lemma invpentagon (A B M N : C) : + (id_ A ⊗ α_ B, M, N ^-1) ∘ + α_ A, (B × M), N ^-1 ∘ (α_ A, B, M ^-1 ⊗ id_ N) + ≃ α_ A, B, (M × N)^-1 ∘ α_ (A × B), M, N ^-1. +Proof. + symmetry. rewrite <- left_unit, <- assoc. + rewrite <- 2!compose_iso_r'. + rewrite assoc, <- pentagon. + rewrite assoc, <- 2!(assoc (_^-1 ⊗ id_ N)), <- tensor_compose. + rewrite iso_inv_l, left_unit, tensor_id, left_unit. + rewrite assoc, <- (assoc (_^-1)), iso_inv_l, left_unit. + now rewrite <- tensor_compose, iso_inv_l, left_unit, tensor_id. +Qed. + + +End InverseCoherences. + Arguments tensor {_} {_}%Cat (mC)%Cat {_ _}%Cat. Arguments TensorIsomorphism {_} {_ mC cCh mCh}%Cat {_ _ _ _}%Cat (_ _)%Cat. diff --git a/ViCaR/MonoidalCoherence/CatCategory.v b/ViCaR/MonoidalCoherence/CatCategory.v new file mode 100644 index 0000000..2aeb0b6 --- /dev/null +++ b/ViCaR/MonoidalCoherence/CatCategory.v @@ -0,0 +1,518 @@ +Require Import Setoid. +From ViCaR Require Import CategoryTypeclass. + +Open Scope Cat_scope. + +Set Universe Polymorphism. + +Section IsomorphismInstances. + +Context {C:Type} {cC : Category C} {cCh : CategoryCoherence cC}. + +#[export] Instance IdentityIsomorphism (A : C) + : Isomorphism A A := {| + forward := id_ A; + reverse := id_ A; + isomorphism_inverse := ltac:(rewrite left_unit; easy) +|}. + +Definition ReverseIsomorphism {C : Type} {cC : Category C} + {cCh : CategoryCoherence cC} {A B : C} + (f : A <~> B) : B <~> A := {| + forward := f.(reverse); + reverse := f.(forward); + isomorphism_inverse := + let (hA, hB) := f.(isomorphism_inverse) in conj hB hA +|}. + +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)), (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. + +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)); +|}. +Next Obligation. + match goal with + |- ?g ≃ ?h => rewrite <- (right_unit h) + end. + rewrite <- (proj1 (N B).(isomorphism_inverse)). + 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 + {F G H : Functor cC cD} (N : NaturalTransformation F G) + (M : NaturalTransformation G H) : NaturalTransformation F H := {| + component_map := fun A => (N.(component_map) A) ∘ (M.(component_map) A) + |}. +Next Obligation. + rewrite <- assoc, N.(component_map_natural), + assoc, M.(component_map_natural), <- assoc. + easy. +Qed. + +Notation "N '⨀' M" := (ComposeNaturalTransformations N M) + (at level 59, left associativity) : Cat_scope. + +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_iso_natural), + assoc, M.(component_iso_natural), <- assoc. + easy. +Qed. + +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. + 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 + {F F' : Functor cC cD} {G G' : Functor cD cE} + (N : NaturalIsomorphism F F') (M : NaturalIsomorphism G G') + : NaturalIsomorphism (F ○ G) (F' ○ G') := + {| + component_iso := fun A => ComposeIsomorphisms + (Isomorphism_by_Functor G (N A)) (M (F' A)) +|}. +Next Obligation. + rewrite 2!component_iso_natural, <- assoc. + rewrite component_iso_natural, 2!(assoc (M (F A))). + apply compose_compat; [easy|]. + rewrite <- 2!compose_map. + rewrite component_iso_natural. + easy. +Qed. + + +Program Definition IdentityFunctor {C} (cC : Category C) + {cCh : CategoryCoherence cC} : Functor cC cC := + {| obj_map := fun A => A; morphism_map := fun _ _ f => f |}. +Solve All Obligations with (rewrite ?left_unit, ?right_unit; easy). + +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. + +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 {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 (F G : Functor cC cD) : + naturally_isomorphic F G <-> (exists _ : NaturalIsomorphism F G, True). +Proof. + split. + - intros [c Hc]. + exists {| component_iso := c; component_iso_natural := Hc |}; + easy. + - intros [[c Hc] _]; exists c; exact Hc. +Qed. + +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 : + symmetric _ (@naturally_isomorphic C D cC cD). +Proof. + intros F G. + rewrite naturally_isomorphic_iff_ex_NaturalIsomorphism. + intros [N _]. + apply naturally_isomorphic_of_NaturalIsomorphism. + exact (ReverseNaturalIsomorphism N). +Qed. + +Lemma naturally_isomorphic_trans : + transitive _ (@naturally_isomorphic C D cC cD). +Proof. + intros F G H. + rewrite 2!naturally_isomorphic_iff_ex_NaturalIsomorphism. + intros [N _] [M _]. + apply naturally_isomorphic_of_NaturalIsomorphism. + exact (ComposeNaturalIsomorphisms N M). +Qed. + +End NaturallyIsomorphic. + + +Section CategoryCat. + +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 := + {| + morphism := fun U V => Functor U V; + c_equiv := fun A B => naturally_isomorphic; + 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. + - apply naturally_isomorphic_trans. + - apply naturally_isomorphic_sym. +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 (chj (f x))). + apply compose_compat; [easy|]. + rewrite <- 2!compose_map. + rewrite Hcfg. + easy. +Qed. +Next Obligation. + apply naturally_isomorphic_of_NaturalIsomorphism. + apply ComposeFunctorsAssociator. +Qed. +Next Obligation. + exists (fun x => IdentityIsomorphism (f x)). + intros. + simpl. + unfold id. + rewrite left_unit, right_unit. + easy. +Qed. +Next Obligation. + exists (fun x => IdentityIsomorphism (f x)). + intros. + simpl. unfold id. + rewrite left_unit, right_unit. + easy. +Qed. + +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; + prod_mor : + forall (Q : C) (fA : Q ~> A) (fB : Q ~> B), Q ~> prod_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 ~> prod_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' {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 -> + 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 (A B : C) : + forall (AB AB' : CategoryProduct A B), Isomorphism @AB @AB' := + fun AB AB' => + {| + forward := AB'.(prod_mor) @AB p_A p_B; + reverse := AB.(prod_mor) @AB' p_A p_B; + |}. +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 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_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_J), + (forall i, fQ_ i ≃ fAB' ∘ p_i i) -> + fAB' ≃ big_cat_prod_mor Q fQ_ +}. + +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'. +Proof. + intros. + rewrite (big_cat_prod_mor_unique Q fQ_) by easy. + symmetry; + rewrite (big_cat_prod_mor_unique Q fQ_) by easy. + easy. +Qed. + +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 := HJ'.(big_cat_prod_mor) pJ p_i; + reverse := HJ.(big_cat_prod_mor) pJ' p_i; + |}. +Next Obligation. + split; eapply (big_cat_prod_mor_unique' _ p_i); rewrite 1?assoc. + 1,3: intros i; rewrite assoc, 2!(big_cat_prod_mor_prop _ _); easy. + all: intros; rewrite left_unit; easy. +Qed. + + + +Definition nat_trans_equiv `{cC : Category C, cD : Category D} + (F G : Functor cC cD) : relation (NaturalTransformation F G) := + fun N M => forall (A:C), N.(component_map) A ≃ M.(component_map) A. + +Coercion NaturalTransformation_of_NaturalIsomorphism + : NaturalIsomorphism >-> NaturalTransformation. + +#[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 ?. + etransitivity; [apply h1 | apply h2]. +Qed. +Next Obligation. + intros a. + simpl. + apply compose_compat; easy. +Qed. +Next Obligation. + intros a. + simpl. + rewrite assoc; easy. +Qed. +Next Obligation. + intros a; simpl. + rewrite left_unit; easy. +Qed. +Next Obligation. + intros a; simpl. + rewrite right_unit; easy. +Qed. + +Definition big_prod {J : Type} (CJ : J -> Type) := forall i, CJ i. + +#[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. +Qed. +Next Obligation. + apply compose_compat; easy. +Qed. +Next Obligation. + rewrite assoc; easy. +Qed. +Next Obligation. + rewrite left_unit; easy. +Qed. +Next Obligation. + rewrite right_unit; easy. +Qed. + +#[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 {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; + compose_map := fun a b m f g => fun i => (fQ_ i).(compose_map) f g; + morphism_compat := fun a b f g H => fun i => (fQ_ i).(morphism_compat) f g H; +}. + +#[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; +}. +Next Obligation. + split; intros i; apply (Hiso i).(isomorphism_inverse). +Qed. + +End CategoryCatProperties. \ No newline at end of file diff --git a/ViCaR/MonoidalCoherence/MCDefinitions.v b/ViCaR/MonoidalCoherence/MCDefinitions.v new file mode 100644 index 0000000..433cd8e --- /dev/null +++ b/ViCaR/MonoidalCoherence/MCDefinitions.v @@ -0,0 +1,562 @@ +Require Export Setoid. +From ViCaR Require CategoryTypeclass. +Require PeanoNat. +Require EqdepFacts Eqdep_dec. + +Create HintDb bwdb. +Create HintDb bwarrdb. +Create HintDb monarrdb. + +Declare Scope bw_scope. +Delimit Scope bw_scope with bw. + +Section Definitions. + +#[local] Set Universe Polymorphism. + +#[universes(polymorphic=yes)] +Context {X : Type}. + +#[universes(polymorphic=yes,cumulative=yes)] +Inductive bw : Type := + | e : bw + | var (x : X) : bw + | tens (a b : bw) : bw. + +Local Notation "a '⨂' b" := (tens a b) + (at level 40, left associativity). (* \bigotimes *) + +#[universes(polymorphic=yes,cumulative=yes)] +Inductive bweq : bw -> bw -> Prop := + | bw_leftid (a : bw) : bweq (tens e a) a + | bw_rightid (a : bw) : bweq (tens a e) a + | bw_assoc (a b c : bw) : bweq (tens (tens a b) c) (tens a (tens b c)) + | bw_tens (a a' b b' : bw) : bweq a a' -> bweq b b' + -> bweq (tens a b) (tens a' b') + | bw_refl (a : bw) : bweq a a + | bw_trans (a b c : bw) : bweq a b -> bweq b c -> bweq a c + | bw_symm (a b : bw) : bweq a b -> bweq b a. + +#[universes(polymorphic=yes,cumulative=yes)] +Inductive bwnorm := + | norm_e : bwnorm + | norm_rtens (n : bwnorm) (x : X) : bwnorm. + +Fixpoint bwnorm_to_bw (a : bwnorm): bw := + match a with + | norm_e => e + | norm_rtens n x => tens (bwnorm_to_bw n) (var x) + end. + +Coercion bwnorm_to_bw : bwnorm >-> bw. + +Fixpoint bwbrac (a : bw) : bwnorm -> bwnorm := + match a with + | e => fun n => n + | var x => fun n => norm_rtens n x + | tens a b => fun n => bwbrac b (bwbrac a n) + end. + +Local Notation "'⟦' a '⟧'" := (bwbrac a). + +Definition Nf (a : bw) : bwnorm := ⟦a⟧ norm_e. + +Fixpoint bw_to_varlist (b : bw) : list X := + match b with + | e => nil + | var x => x::nil + | tens a b => bw_to_varlist b ++ bw_to_varlist a + end. + +Fixpoint bwnorm_to_varlist (n : bwnorm) : list X := + match n with + | norm_e => nil + | norm_rtens m a => a :: bwnorm_to_varlist m + end. + +Fixpoint varlist_to_bwnorm (l : list X) : bwnorm := + match l with + | nil => norm_e + | cons a l' => norm_rtens (varlist_to_bwnorm l') a + end. + +Fixpoint bwnormapp (n m : bwnorm) {struct m} : bwnorm := + match m with + | norm_e => n + | norm_rtens m' a => norm_rtens (bwnormapp n m') a + end. + +#[universes(polymorphic=yes,cumulative=yes)] +Inductive bwarr : bw -> bw -> Type := + | arrcomp {a b c : bw} (f : bwarr a b) (g : bwarr b c) : bwarr a c + | arrtens {a a' b b'} (f : bwarr a a') (g : bwarr b b') : bwarr (a ⨂ b) (a' ⨂ b') + | arrid (a : bw) : bwarr a a + | arrassoc (a b c : bw) : bwarr (a ⨂ (b ⨂ c)) (a ⨂ b ⨂ c) + | arrinvassoc (a b c : bw) : bwarr (a ⨂ b ⨂ c) (a ⨂ (b ⨂ c)) + | arrlunitor (a : bw) : bwarr (e ⨂ a) a + | arrinvlunitor (a : bw) : bwarr a (e ⨂ a) + | arrrunitor (a : bw) : bwarr (a ⨂ e) a + | arrinvrunitor (a : bw) : bwarr a (a ⨂ e). + +Section bwarr. + +Local Notation "a '⟶' b" := (bwarr a b) + (at level 60) : type_scope. (* \longrightarrow *) + +Local Notation "f '○' g" := (arrcomp f g) + (at level 59, left associativity). (* \bigcirc *) +Local Notation "f '⊠' g" := (arrtens f g) + (at level 40, left associativity). (* \boxtimes *) + +Reserved Notation "f '≅' g" (at level 70). (* \cong *) + +#[universes(polymorphic=yes,cumulative=yes)] +Inductive bwarrequiv : forall a b, relation (a ⟶ b) := + | bwarr_comp {a b c : bw} (f f' : a ⟶ b) (g g' : b ⟶ c) : + f ≅ f' -> g ≅ g' -> f ○ g ≅ f' ○ g' + | bwarr_assoc {a a' b' b : bw} (f : a ⟶ a') (g : a' ⟶ b') (h : b' ⟶ b) : + f ○ g ○ h ≅ f ○ (g ○ h) + | bwarr_lunit {a b} (f : a ⟶ b) : (arrid a) ○ f ≅ f + | bwarr_runit {a b} (f : a ⟶ b) : f ○ (arrid b) ≅ f + + | bwarr_tens {a a' b b' : bw} (f f' : a ⟶ a') (g g' : b ⟶ b') : + f ≅ f' -> g ≅ g' -> arrtens f g ≅ arrtens f' g' + | bwarr_tens_id {a b : bw} : + arrid a ⊠ arrid b ≅ arrid (a ⨂ b) + | bwarr_tens_comp {a b c a' b' c'} + (f : a ⟶ b) (g : b ⟶ c) (f' : a' ⟶ b') (g' : b' ⟶ c') : + (f ○ g) ⊠ (f' ○ g') ≅ + f ⊠ f' ○ g ⊠ g' + + | bwarr_assoc_rinv (a b c : bw) : + arrassoc a b c ○ arrinvassoc a b c ≅ arrid (a ⨂ (b ⨂ c)) + | bwarr_assoc_linv (a b c : bw) : + arrinvassoc a b c ○ arrassoc a b c ≅ arrid (a ⨂ b ⨂ c) + + | bwarr_lunitor_rinv (a : bw) : + arrlunitor a ○ arrinvlunitor a ≅ arrid (e ⨂ a) + | bwarr_lunitor_linv (a : bw) : + arrinvlunitor a ○ arrlunitor a ≅ arrid a + + | bwarr_runitor_rinv (a : bw) : + arrrunitor a ○ arrinvrunitor a ≅ arrid (a ⨂ e) + | bwarr_runitor_linv (a : bw) : + arrinvrunitor a ○ arrrunitor a ≅ arrid a + + | bwarr_assoc_nat {a b c a' b' c' : bw} + (f : a ⟶ a') (g : b ⟶ b') (h : c ⟶ c') : + arrassoc a b c ○ f ⊠ g ⊠ h + ≅ f ⊠ (g ⊠ h) ○ arrassoc a' b' c' + + | bwarr_lunitor_nat {a b : bw} (f : a ⟶ b) : + arrlunitor a ○ f ≅ (arrid e) ⊠ f ○ arrlunitor b + + | bwarr_runitor_nat {a b : bw} (f : a ⟶ b) : + arrrunitor a ○ f ≅ f ⊠ (arrid e) ○ arrrunitor b + + | bwarr_pentagon (a b c d : bw) : + arrassoc a b (c⨂d) ○ arrassoc (a⨂b) c d + ≅ arrid a ⊠ arrassoc b c d ○ arrassoc a (b⨂c) d + ○ arrassoc a b c ⊠ arrid d + + | bwarr_triangle (a b : bw) : + arrassoc a e b ○ arrrunitor a ⊠ arrid b + ≅ arrid a ⊠ arrlunitor b + | bwarr_refl {a b} (f : a ⟶ b) : f ≅ f + | bwarr_trans {a b} (f g h : a ⟶ b) : f ≅ g -> g ≅ h -> f ≅ h + | bwarr_symm {a b} (f g : a ⟶ b) : f ≅ g -> g ≅ f + where "f '≅' g" := (bwarrequiv _ _ f g). + +Local Notation "f '≅' g" := (bwarrequiv _ _ f g) (at level 70). (* \cong *) + +Fixpoint bwarrinv {A B} (h : A ⟶ B) : B ⟶ A := + match h with + | arrid a => arrid a + | arrassoc a b c => arrinvassoc a b c + | arrinvassoc a b c => arrassoc a b c + | arrlunitor a => arrinvlunitor a + | arrinvlunitor a => arrlunitor a + | arrrunitor a => arrinvrunitor a + | arrinvrunitor a => arrrunitor a + | arrcomp f g => arrcomp (bwarrinv g) (bwarrinv f) + | arrtens f g => arrtens (bwarrinv f) (bwarrinv g) + end. + +Fixpoint xi_comp_map (n : bwnorm) (A : bw) {struct A} : + n ⨂ A ⟶ ⟦A⟧ n := + match A with + | e => arrrunitor n + | var x => arrid (n ⨂ var x) + | tens a b => + arrassoc n a b ○ xi_comp_map n a ⊠ arrid b + ○ xi_comp_map (⟦a⟧ n) b + end. + +Definition to_Nf_arr (a : bw) : bwarr a (Nf a) := + (arrinvlunitor a) ○ (xi_comp_map norm_e a). + +Definition from_Nf_arr (a : bw) : bwarr (Nf a) a := + (bwarrinv (xi_comp_map norm_e a)) ○ (arrlunitor a). + +Definition cast_bwarr {n n' m m'} + (Hn : n = n') (Hm : m = m') (f : n ⟶ m) : n' ⟶ m'. +rewrite <- Hn, <- Hm. +apply f. +Defined. + +Definition bwarr_of_Nf_eq {a b : bw} (H : Nf a = Nf b) : bwarr a b := + to_Nf_arr a ○ + cast_bwarr eq_refl (f_equal _ H) (arrid (Nf a)) + ○ from_Nf_arr b. + +End bwarr. + +Section MonoidalCoherence. + +Import CategoryTypeclass. + +Context {cC : Category X} {cCh : CategoryCoherence cC} + {mC : MonoidalCategory cC} {mCh : MonoidalCategoryCoherence mC}. + +Local Open Scope Cat_scope. + +Fixpoint realize_bw (a : bw) : X := + match a with + | e => mC.(mon_I) + | var x => x + | tens a' b' => mC.(obj_tensor) (realize_bw a') (realize_bw b') + end. + +Coercion realize_bw : bw >-> X. + +Existing Instance cC | 0. +Existing Instance mC | 0. + + +Fixpoint realize_bwarr {A B} (h : bwarr A B) : (realize_bw A ~> realize_bw B) := + match h with + | arrid a => cC.(c_identity) a + | arrassoc a b c => (mC.(associator) a b c)^-1 + | arrinvassoc a b c => (mC.(associator) a b c) + | arrlunitor a => mC.(left_unitor) a + | arrinvlunitor a => (mC.(left_unitor) a)^-1 + | arrrunitor a => mC.(right_unitor) a + | arrinvrunitor a => (mC.(right_unitor) a)^-1 + | arrcomp f g => (realize_bwarr f) ∘ (realize_bwarr g) + | arrtens f g => (realize_bwarr f) ⊗ (realize_bwarr g) + end. + + +#[universes(polymorphic=yes,cumulative=yes)] +Inductive monarr : bw -> bw -> Type := + | monarrcomp {a b c : bw} (f : monarr a b) (g : monarr b c) : + monarr a c + | monarrtens {a a' b b'} (f : monarr a a') (g : monarr b b') : + monarr (a ⨂ b) (a' ⨂ b') + | mongeneric {a b : bw} (f : cC.(morphism) a b) : + monarr a b + | monarrstruct {a b : bw} (f : bwarr a b) : + monarr a b. + +Coercion monarrstruct : bwarr >-> monarr. +Local Notation "a '⟶' b" := (monarr a b) (at level 60) : type_scope. + +Fixpoint realize_monarr {A B} (h : A ⟶ B) : (realize_bw A ~> realize_bw B) := + match h with + | monarrcomp f g => realize_monarr f ∘ realize_monarr g + | monarrtens f g => realize_monarr f ⊗ realize_monarr g + | mongeneric f => f + | monarrstruct f => realize_bwarr f + end. + +Reserved Notation "f '≊' g" (at level 70). +#[universes(polymorphic=yes,cumulative=yes)] +Inductive monarrequiv : forall a b, relation (a ⟶ b) := + | monarr_comp {a b c : bw} (f f' : a ⟶ b) (g g' : b ⟶ c) : + f ≊ f' -> g ≊ g' -> monarrcomp f g ≊ monarrcomp f' g' + | monarr_assoc {a a' b' b : bw} (f : a ⟶ a') (g : a' ⟶ b') (h : b' ⟶ b) : + monarrcomp f (monarrcomp g h) ≊ monarrcomp (monarrcomp f g) h + | monarr_lunit {a b} (f : a ⟶ b) : monarrcomp (arrid a) f ≊ f + | monarr_runit {a b} (f : a ⟶ b) : monarrcomp f (arrid b) ≊ f + + | monarr_tens {a a' b b' : bw} (f f' : a ⟶ a') (g g' : b ⟶ b') : + f ≊ f' -> g ≊ g' -> monarrtens f g ≊ monarrtens f' g' + | monarr_tens_comp {a b c a' b' c'} + (f : a ⟶ b) (g : b ⟶ c) (f' : a' ⟶ b') (g' : b' ⟶ c') : + monarrtens (monarrcomp f g) (monarrcomp f' g') + ≊ monarrcomp (monarrtens f f') (monarrtens g g') + | monarr_struct {a b} (f g : bwarr a b) : + (* bwarrequiv a b f g -> *) + (* NOTE: this predicate is given by monoidal coherence! *) + f ≊ g + | monarr_arrcomp {a b c} (f : bwarr a b) (g : bwarr b c) : + monarrcomp f g ≊ arrcomp f g + | monarr_arrtens {a a' b b'} (f : bwarr a a') (g : bwarr b b') : + monarrtens f g ≊ arrtens f g + + | monarr_assoc_nat {a b c a' b' c' : bw} + (f : a ⟶ a') (g : b ⟶ b') (h : c ⟶ c') : + monarrcomp (arrassoc a b c) (monarrtens (monarrtens f g) h) + ≊ monarrcomp (monarrtens f (monarrtens g h)) (arrassoc a' b' c') + | monarr_lunitor_nat {a b : bw} (f : a ⟶ b) : + monarrcomp (arrlunitor a) f ≊ monarrcomp (monarrtens (arrid e) f) (arrlunitor b) + | monarr_runitor_nat {a b : bw} (f : a ⟶ b) : + monarrcomp (arrrunitor a) f ≊ monarrcomp (monarrtens f (arrid e)) (arrrunitor b) + + | monarr_refl {a b} (f : a ⟶ b) : f ≊ f + | monarr_symm {a b} (f g : a ⟶ b) : f ≊ g -> g ≊ f + | monarr_trans {a b} (f g h : a ⟶ b) : f ≊ g -> g ≊ h -> f ≊ h + where "f '≊' g" := (monarrequiv _ _ f g). + +End MonoidalCoherence. + +End Definitions. + +Arguments bw _ : clear implicits. +Arguments bwnorm _ : clear implicits. +#[global] Arguments cast_bwarr {_ _ _ _ _} !_ !_ _ /. +(* This will simplify only if both casts are constructors, i.e. eq_refl *) + +Section ExtraDefinitions. + +Fixpoint map_bw {A B} (f : A -> B) (a : bw A) : bw B := + match a with + | e => e + | var a' => var (f a') + | tens a' b' => tens (map_bw f a') (map_bw f b') + end. + +Fixpoint map_bwarr {A B} (f : A -> B) {a b} (g : bwarr a b) : + bwarr (map_bw f a) (map_bw f b) := + match g with + | arrid a => arrid (map_bw f a) + | arrassoc a b c => arrassoc (map_bw f a) (map_bw f b) (map_bw f c) + | arrinvassoc a b c => arrinvassoc (map_bw f a) (map_bw f b) (map_bw f c) + | arrlunitor a => arrlunitor (map_bw f a) + | arrinvlunitor a => arrinvlunitor (map_bw f a) + | arrrunitor a => arrrunitor (map_bw f a) + | arrinvrunitor a => arrinvrunitor (map_bw f a) + | arrcomp f' g' => arrcomp (map_bwarr f f') (map_bwarr f g') + | arrtens f' g' => arrtens (map_bwarr f f') (map_bwarr f g') + end. + +End ExtraDefinitions. + +#[export] Hint Constructors bweq : bwdb. +#[export] Hint Constructors bwarr : bwdb. +#[export] Hint Constructors bwarrequiv : bwarrdb. +#[export] Hint Constructors monarrequiv : monarrdb. + +Module MC_setoids. + +Import CategoryTypeclass. + +Definition true_rel {A : Type} : relation A := + fun _ _ => True. + +#[global] Add Parametric Relation (A : Type) : A true_rel + reflexivity proved by ltac:(easy) + symmetry proved by ltac:(easy) + transitivity proved by ltac:(easy) + as true_rel_relation. + +#[program, export] Instance subrel_true_rel {A : Type} (R : relation A) : + subrelation R true_rel. +Next Obligation. easy. Qed. + +Add Parametric Relation {X} : (@bw X) bweq + reflexivity proved by bw_refl + symmetry proved by bw_symm + transitivity proved by bw_trans as bweq_setoid. + +Add Parametric Relation {X} (a b : @bw X) : (bwarr a b) (bwarrequiv a b) + reflexivity proved by bwarr_refl + symmetry proved by bwarr_symm + transitivity proved by bwarr_trans as bwarrequiv_setoid. + +Add Parametric Morphism {X} (a b c : @bw X) : (@arrcomp X a b c) + with signature + (bwarrequiv a b) ==> (bwarrequiv b c) ==> (bwarrequiv a c) + as arrcomp_mor. +Proof. eauto with bwarrdb. Qed. + +Add Parametric Morphism {X} (a a' b b' : @bw X) : (@arrtens X a a' b b') + with signature + (bwarrequiv a a') ==> (bwarrequiv b b') + ==> (bwarrequiv (tens a b) (tens a' b')) + as arrtens_mor. +Proof. eauto with bwarrdb. Qed. + +Add Parametric Morphism {A B} (f : A -> B) {a b} : + (@map_bwarr A B f a b) with signature + bwarrequiv a b ==> bwarrequiv (map_bw f a) (map_bw f b) + as map_bwarr_mor. +Proof. + intros g h H. + induction H; simpl; try constructor; + eauto 3 with bwarrdb. +Qed. + + +Add Parametric Relation {X} {cC : Category X} {mC : MonoidalCategory cC} + (a b : @bw X) : (monarr a b) (monarrequiv a b) + reflexivity proved by monarr_refl + symmetry proved by monarr_symm + transitivity proved by monarr_trans as monarrequiv_setoid. + +Add Parametric Morphism {X} {cC : Category X} {mC : MonoidalCategory cC} + (a b c : @bw X) : (@monarrcomp _ _ _ a b c) + with signature + (monarrequiv a b) ==> (monarrequiv b c) ==> (monarrequiv a c) + as monarrcomp_mor. +Proof. eauto with monarrdb. Qed. + +Add Parametric Morphism {X} {cC : Category X} {mC : MonoidalCategory cC} + (a a' b b' : @bw X) : (@monarrtens _ _ _ a a' b b') + with signature + (monarrequiv a a') ==> (monarrequiv b b') + ==> (monarrequiv (tens a b) (tens a' b')) + as monarrtens_mor. +Proof. eauto with monarrdb. Qed. + +Add Parametric Morphism {X} {cC : Category X} {mC : MonoidalCategory cC} + (a b : @bw X) : (@monarrstruct X cC mC a b) + with signature true_rel ==> monarrequiv a b + as monarrstruct_mor. +Proof. eauto with monarrdb. Qed. + +End MC_setoids. + +Module MC_notations. + +Notation "a '⨂' b" := (tens a b) + (at level 40, left associativity) : bw_scope. (* \bigotimes *) +Notation "a '~' b" := (bweq a b) (at level 70) : bw_scope. +Notation "'⟦' a '⟧'" := (bwbrac a) : bw_scope. (* \llbracket , \rrbracket *) + +Notation "f '≅' g" := (bwarrequiv _ _ f g) + (at level 70) : bw_scope. (* \cong *) + +Notation "f '○' g" := (arrcomp f g) + (at level 59, left associativity) : bw_scope. (* \bigcirc *) +Notation "f '⊠' g" := (arrtens f g) + (at level 40, left associativity) : bw_scope. (* \boxtimes *) + +Notation "f '^-'" := (bwarrinv f) (at level 9) : bw_scope. + +Module bwarr_notation. +Notation "a '⟶' b" := (bwarr a b) + (at level 60) : type_scope. (* \longrightarrow *) +End bwarr_notation. + + + +Module monarr_notation. +Notation "a '⟶' b" := (monarr a b) + (at level 60) : type_scope. (* \longrightarrow *) +End monarr_notation. + + +Notation "f '≊' g" := (monarrequiv _ _ f g) + (at level 70) : bw_scope. + +Notation "f ⧆ g" := (monarrtens f g) + (at level 40, left associativity) : bw_scope. (* \boxast *) +Notation "f ◌ g" := (monarrcomp f g) + (at level 50, left associativity) : bw_scope. (* \dottedcircle *) + + +End MC_notations. + +Module MCClasses. + +(* Computing forms of proj1 and proj2, for UIP proofs *) +Definition proj1' {A B} (H : A /\ B) : A := + match H with + | conj PA PB => PA + end. + +Definition proj2' {A B} (H : A /\ B) : B := + match H with + | conj PA PB => PB + end. + +#[universes(polymorphic=yes,cumulative=yes)] +Class DECEQ (X : Type) := { + deceq : forall x y : X, {x = y} + {x <> y} +}. + +#[universes(polymorphic=yes,cumulative=yes)] +Class UIP (X : Type) := { + uip : forall {x y : X} (H H' : x = y), H = H' +}. + +#[universes(polymorphic=yes,cumulative=yes)] +Class UIP_refl (X : Type) := { + uip_refl : forall {x : X} (H : x = x), H = eq_refl +}. + +Lemma UIP_iff_UIP_refl (X : Type) : + UIP X <-> UIP_refl X. +Proof. + split. + - intros H; constructor; intros; apply uip. + - intros H; constructor. + intros x y e0. + case e0. + symmetry. + apply uip_refl. +Defined. + +#[universes(polymorphic=yes,cumulative=yes)] +Class EQ_RECT_EQ (X : Type) := { + eq_rect_eq : forall (x : X) (P : X -> Type) (a : P x) (h : x = x), + eq_rect x P a x h = a +}. + +#[universes(polymorphic=yes,cumulative=yes)] +Class EQ_REC_EQ (X : Type) := { + eq_rec_eq : forall (x : X) (P : X -> Set) (a : P x) (h : x = x), + eq_rec x P a x h = a +}. + +End MCClasses. + +Section MCClasses_instances. + +Import MCClasses. + + +#[universes(polymorphic=yes), + export] Instance DECEQ_UIP {X : Type} (H : DECEQ X) : UIP X := { + uip := Eqdep_dec.UIP_dec deceq +}. + +#[universes(polymorphic=yes), + export, program] Instance UIP_EQ_RECT_EQ {X : Type} (H : UIP X) + : EQ_RECT_EQ X := {}. +Next Obligation. + symmetry. + apply EqdepFacts.Streicher_K__eq_rect_eq. + apply EqdepFacts.UIP_refl__Streicher_K. + intros ? ?. + apply uip. +Qed. + +#[universes(polymorphic=yes), + export] Instance EQ_RECT_EQ_EQ_REC_EQ {X} (H : EQ_RECT_EQ X) + : EQ_REC_EQ X := { + eq_rec_eq := eq_rect_eq +}. + +#[universes(polymorphic=yes), + export] Instance DECEQ_nat : DECEQ nat := { + deceq := PeanoNat.Nat.eq_dec +}. + +#[universes(polymorphic=yes), + export] Instance DECEQ_bool : DECEQ bool := { + deceq := Bool.bool_dec +}. + +End MCClasses_instances. \ No newline at end of file diff --git a/ViCaR/MonoidalCoherence/MCTheory.v b/ViCaR/MonoidalCoherence/MCTheory.v new file mode 100644 index 0000000..87b0630 --- /dev/null +++ b/ViCaR/MonoidalCoherence/MCTheory.v @@ -0,0 +1,86 @@ +From ViCaR Require CategoryTypeclass. +Require CatCategory. (* FunctorCategory *) + + + +Section DiscreteCategory. + +Set Universe Polymorphism. + +Import CategoryTypeclass. + +#[universes(polymorphic=yes), + program] Definition DiscreteCategory (N : Type) : Category N := {| + morphism := @eq N; + c_equiv := fun _ _ _ _ => True; + c_identity := @eq_refl N +|}. + +#[universes(polymorphic=yes), + export, program] Instance DiscreteCategoryCoherence (N : Type) + : CategoryCoherence (DiscreteCategory N) := { +}. +Solve All Obligations with easy. + +End DiscreteCategory. + +Section Bifunctor_of_FunctorCategoryFunctor. + +Import CategoryTypeclass CatCategory (FunctorCategory). + +Set Universe Polymorphism. + +#[export, program] + Instance Bifunctor_of_FunctorCategoryFunctor + {C D E : Type} {cC : Category C} {cD : Category D} {cE : Category E} + (* {cCh : CategoryCoherence cC} *) + {cDh : CategoryCoherence cD} {cEh : CategoryCoherence cE} + (F : Functor cC (FunctorCategory (cC:=cD) (cD:=cE))) : + Bifunctor cC cD cE := { + obj_bimap := F.(obj_map); + morphism_bimap := fun A1 B1 A2 B2 f1 f2 => + (F A1 @ f2 ∘ component_map (F.(morphism_map) f1) B2)%Cat +}. +Next Obligation. + (* rewrite (F A1).(id_map), left_unit. *) + rewrite component_map_natural. + rewrite (F A1).(id_map), right_unit. + apply F. +Qed. +Next Obligation. + symmetry. + rewrite assoc, <- (assoc _ (F B1 @ g2)%Cat). + rewrite <- component_map_natural. + rewrite compose_map. + rewrite 2!assoc. + apply compose_cancel_l, compose_cancel_l. + symmetry. + apply F. +Qed. +Next Obligation. + rewrite H0. + pose proof (F.(morphism_compat) f f' ltac:(assumption)) as e. + simpl in e. + hnf in e. + rewrite e. + easy. +Qed. + +Unset Universe Polymorphism. + +End Bifunctor_of_FunctorCategoryFunctor. + +Section Groupoid. + +(* Set Universe Polymorphism. *) + +Import CategoryTypeclass. + +#[universes(polymorphic=yes,cumulative=yes)] +Class IsGroupoid {C} (cC : Category C) := { + groupoid_inv {A B : C} (f : (A ~> B)%Cat) : (B ~> A)%Cat; + groupoid_inv_is_inv {A B : C} (f : (A ~> B)%Cat) : + (is_inverse f (groupoid_inv f))%Cat +}. + +End Groupoid. \ No newline at end of file diff --git a/ViCaR/MonoidalCoherence/MCautomation.v b/ViCaR/MonoidalCoherence/MCautomation.v new file mode 100644 index 0000000..7488c7d --- /dev/null +++ b/ViCaR/MonoidalCoherence/MCautomation.v @@ -0,0 +1,1167 @@ +Require Import Setoid. +Require MCprocessing. + +Section monarr_under_over. + +Open Scope bw_scope. + +Set Universe Polymorphism. + +Import MCDefinitions MCClasses MC_setoids + MC_notations UIP_facts MCbw + MCconsequences MCmonarrlist + MCproperop MCprocessing. +Import CategoryTypeclass. +Import List ListNotations. + +(* Polymorphic *)Context {X : Type} {UIPX : UIP X}. +(* Polymorphic *)Context {cC : Category X} {cCh : CategoryCoherence cC} + {mC : MonoidalCategory cC} {mCh : MonoidalCategoryCoherence mC}. + +Local Notation bw := (bw X). +Local Notation "a ⟶ b" := (@monarr X cC mC a b) (at level 60). + +Definition realize_equiv (a b : bw) (f g : a ⟶ b) := + (realize_monarr f ≃ realize_monarr g)%Cat. + +Arguments realize_equiv _ _ _ _/. + +Definition realize_equiv' := @realize_equiv. + +Global Arguments realize_equiv' _ _ _ _ : simpl never. + +Local Notation "f '≡' g" := (realize_equiv' _ _ f g) (at level 70) : bw_scope. + +#[global] Add Parametric Relation (a b : bw) : (monarr a b) (realize_equiv' a b) + reflexivity proved by + ltac:(intros ?; unfold realize_equiv'; simpl; reflexivity) + symmetry proved by + ltac:(unfold realize_equiv'; intros ?; simpl; symmetry; easy) + transitivity proved by + ltac:(unfold realize_equiv'; intros ?; simpl; etransitivity; eauto) + as realize_equiv_setoid. + +#[program] Instance realize_equiv_subrel {x y} : + subrelation (monarrequiv x y) (realize_equiv' x y). +Next Obligation. + unfold realize_equiv'; simpl. + rewrite H. + reflexivity. +Qed. + +#[global] Add Parametric Morphism (a b c : bw) : (@monarrcomp X cC mC a b c) + with signature + (realize_equiv' a b) ==> (realize_equiv' b c) ==> (realize_equiv' a c) + as monarrcomp_mor_real. +Proof. intros; apply compose_compat; easy. Qed. + +#[global] Add Parametric Morphism (a a' b b' : bw) : (@monarrtens X cC mC a a' b b') + with signature + (realize_equiv' a a') ==> (realize_equiv' b b') ==> (realize_equiv' (a⨂b) (a'⨂b')) + as monarrtens_mor_real. +Proof. intros; apply tensor_compat; easy. Qed. + +#[global] Add Parametric Morphism (a b : bw) : (realize_equiv' a b) + with signature + monarrequiv a b ==> monarrequiv a b ==> iff + as realize_equiv_mor. +Proof. + intros * H * H0; + rewrite H, H0. + easy. +Qed. + +Lemma mon_struct_l {a a' b} (f : bwarr a a') (g : a' ⟶ b) (h : a ⟶ b) : + f ◌ g ≡ h <-> g ≡ monarrstruct (a:=a') (b:=a) f ^- ◌ h. +Proof. + now split; [intros <-|intros ->]; + rewrite monarr_assoc, monarr_struct_inv, monarr_id_l. +Qed. + +Lemma mon_struct_r {a a' b} (f : bwarr a' b) (g : a ⟶ a') (h : a ⟶ b) : + g ◌ f ≡ h <-> g ≡ h ◌ monarrstruct (a:=b) (b:=a') f ^-. +Proof. + now split; [intros <-|intros ->]; + rewrite <- monarr_assoc, monarr_struct_inv, monarr_id_r. +Qed. + +Lemma mon_struct_l' {a a' b} (f : bwarr a a') (g : a' ⟶ b) (h : a ⟶ b) : + h ≡ f ◌ g <-> monarrstruct (a:=a') (b:=a) f ^- ◌ h ≡ g. +Proof. + now split; [intros ->|intros <-]; + rewrite monarr_assoc, monarr_struct_inv, monarr_id_l. +Qed. + +Lemma mon_struct_r' {a a' b} (f : bwarr a' b) (g : a ⟶ a') (h : a ⟶ b) : + h ≡ g ◌ f <-> h ◌ monarrstruct (a:=b) (b:=a') f ^- ≡ g. +Proof. + now split; [intros ->|intros <-]; + rewrite <- monarr_assoc, monarr_struct_inv, monarr_id_r. +Qed. + +Lemma by_all_equiv_foliation {a b} (f g : a ⟶ b) : + all_monarrlist_list_equiv (foliate_monarr f) (foliate_monarr g) -> + f ≡ g. +Proof. + intros Hequiv. + apply realize_equiv_subrel. + apply monarrequiv_iff_monarr_norm_equiv. + apply (monarr_norm_equiv_trans' (foliate_correct' f) (foliate_correct' g)). + apply (compose_composable_monarrlist_list_equiv Hequiv). +Qed. + +Lemma by_reflexive_foliation {a b} (f g : a ⟶ b) : + foliate_monarr f = foliate_monarr g -> + f ≡ g. +Proof. + intros H; apply by_all_equiv_foliation. + rewrite H. + easy. +Qed. + +Lemma by_reflexive_trim_foliation {a b} (f g : a ⟶ b) : + trim_foliate_monarr f = trim_foliate_monarr g -> + f ≡ g. +Proof. + intro H. + apply realize_equiv_subrel. + apply monarrequiv_iff_monarr_norm_equiv. + apply (monarr_norm_equiv_trans' + (trim_foliate_correct' f) (trim_foliate_correct' g)). + (* revert H. *) + lazymatch goal with + |- monarr_norm_equiv (compose_totally_composable ?termf ?pf) _ => + pose proof (H : termf = _) as H'; + clear H; revert H'; + generalize termf pf (* idtac term + change term with + (trim_foliate_monarr f) *) + end. + intros; subst. + erewrite compose_totally_composable_indep. easy. +Qed. + +Lemma by_reflexive_trim_foliation_no_empty {a b} (f g : a ⟶ b) : + trim_foliate_monarr_no_empty f = trim_foliate_monarr_no_empty g -> + f ≡ g. +Proof. + intro H. + apply realize_equiv_subrel. + apply monarrequiv_iff_monarr_norm_equiv. + apply (monarr_norm_equiv_trans' + (trim_foliate_no_empty_correct' f) (trim_foliate_no_empty_correct' g)). + lazymatch goal with + |- monarr_norm_equiv (compose_totally_composable ?termf ?pf) _ => + pose proof (H : termf = _) as H'; + clear H; revert H'; + generalize termf pf + end. + intros; subst. + erewrite compose_totally_composable_indep. + easy. +Qed. + +(* Lemma realize_monarr_mor_under {a b c d : bw} (g g' : a ⟶ b) + (f f' : c ⟶ d) : + (* (f ≊ g <-> f' ≊ g') -> *) + (g ≡ g' -> f ≡ f') -> + g ≡ g' -> f ≡ f'. +Proof. + auto. +Qed. + +Definition monarr_under_rel {x y z w} (g g' : z ⟶ w) (f f' : x ⟶ y) := + g ≡ g' -> f ≡ f'. + +Arguments monarr_under_rel : simpl never. + +#[global] Add Parametric Morphism {x y z w} : (@monarr_under_rel x y z w) + with signature + monarrequiv _ _ ==> monarrequiv _ _ ==> + monarrequiv _ _ ==> monarrequiv _ _ ==> iff + as monarr_under_rel_mor. +Proof. + unfold monarr_under_rel. + intros * H * H' * H'' * H'''. + rewrite H, H', H'', H'''. + easy. +Qed. + +#[global] Add Parametric Relation {x y z w} g g' : + (x ⟶ y) (@monarr_under_rel x y z w g g') + reflexivity proved by ltac:(intros []; easy) + symmetry proved by ltac:(unfold monarr_under_rel in *; intros **; + (* match goal with |- ?g => idtac g end; *) + try symmetry; eauto; intros ? **; symmetry; auto) + transitivity proved by + ltac:(unfold monarr_under_rel in *; intros **; + try etransitivity; eauto; + intros ? **; etransitivity; eauto) + as monarr_under_rel_relation. + +#[program] Instance monarr_under_subrel {x y} : + subrelation (monarrequiv x y) (realize_equiv x y). +Next Obligation. + rewrite H. + reflexivity. +Qed. + +Lemma monarr_under_struct_r {x x' y z w} (g g' : z ⟶ w) + f (f' : bwarr x' y) (h : x ⟶ y): + monarr_under_rel g g' (f ◌ f') h <-> + monarr_under_rel g g' f (h ◌ f' ^-). +Proof. + unfold monarr_under_rel. + split; intros H Hg; specialize (H Hg); + [rewrite <- H | rewrite H]; + rewrite <- monarr_assoc, monarr_arrcomp, + ?monarr_id_l, ?monarr_id_r; easy. +Qed. + +Lemma monarr_under_struct_l {x x' y z w} (g g' : z ⟶ w) + (f : bwarr x x') f' (h : x ⟶ y): + monarr_under_rel g g' (f ◌ f') h <-> + monarr_under_rel g g' f' (f ^- ◌ h). +Proof. + unfold monarr_under_rel. + split; intros H Hg; specialize (H Hg); + [rewrite <- H | rewrite H]; + rewrite monarr_assoc, monarr_arrcomp, + ?monarr_id_l, ?monarr_id_r; easy. +Qed. + +Lemma monarr_under_struct_r' {x x' y z w} (g g' : z ⟶ w) + f (f' : bwarr x' y) (h : x ⟶ y): + monarr_under_rel g g' h (f ◌ f') <-> + monarr_under_rel g g' (h ◌ f' ^-) f. +Proof. + split; symmetry; now apply monarr_under_struct_r. +Qed. + +Lemma monarr_under_struct_l' {x x' y z w} (g g' : z ⟶ w) + (f : bwarr x x') f' (h : x ⟶ y): + monarr_under_rel g g' h (f ◌ f') <-> + monarr_under_rel g g' (f ^- ◌ h) f'. +Proof. + split; symmetry; now apply monarr_under_struct_l. +Qed. + +Lemma monarr_under_cancel_struct_l {x x' y z w} {g g' : z ⟶ w} + (h : bwarr x x') (f f' : x' ⟶ y) : + monarr_under_rel g g' (h ◌ f) (h ◌ f') <-> + monarr_under_rel g g' f f'. +Proof. + now rewrite monarr_under_struct_l, + monarr_assoc, monarr_struct_inv, monarr_id_l. +Qed. + +Lemma monarr_under_cancel_struct_r {x x' y z w} {g g' : z ⟶ w} + (h : bwarr x' y) (f f' : x ⟶ x') : + monarr_under_rel g g' (f ◌ h) (f' ◌ h) <-> + monarr_under_rel g g' f f'. +Proof. + now rewrite monarr_under_struct_r, + <- monarr_assoc, monarr_struct_inv, monarr_id_r. +Qed. + +Lemma monarr_under_cancel_l {x x' y z w} {g g' : z ⟶ w} + (h : x ⟶ x') (f f' : x' ⟶ y) : + monarr_under_rel g g' f f' -> + monarr_under_rel g g' (h ◌ f) (h ◌ f'). +Proof. + now intros H Hg; rewrite (H Hg). +Qed. + +Lemma monarr_under_cancel_r {x x' y z w} {g g' : z ⟶ w} + (h : x' ⟶ y) (f f' : x ⟶ x') : + monarr_under_rel g g' f f' -> + monarr_under_rel g g' (f ◌ h) (f' ◌ h). +Proof. + now intros H Hg; rewrite (H Hg). +Qed. + +Lemma monarr_under_rel_trans {s t u v x y z w} (g g' : u ⟶ v) + (h h' : x ⟶ y) (i i' : s ⟶ t) (fmid f f' : z ⟶ w) : + monarr_under_rel h h' f fmid -> + monarr_under_rel i i' fmid f' -> + h ≡ h' -> i ≡ i' -> + monarr_under_rel g g' f f'. +Proof. + unfold monarr_under_rel. + intros; etransitivity; eauto. +Qed. + +Lemma monarr_under_rel_trans_no_evars {u v x y z w} + {h h' : u ⟶ v} {i i' : x ⟶ y} (fmid f f' : z ⟶ w) : + monarr_under_rel h h' f fmid -> + monarr_under_rel i i' fmid f' -> + h ≡ h' -> i ≡ i' -> + monarr_under_rel f f f f'. +Proof. + apply monarr_under_rel_trans. +Qed. *) + +(* Lemma under_by_norm_equiv *) + +Unset Universe Polymorphism. + +End monarr_under_over. + +(* Arguments monarrstruct {_ _ _} _ _. + +Notation "'α_' a ',' b ',' c" := + (monarrstruct (a%monarr_scope ⨂ b%monarr_scope ⨂ c%monarr_scope) + (a%monarr_scope ⨂ (b%monarr_scope ⨂ c%monarr_scope)) _) + (in custom mon at level 20, + a custom mon_bw, b custom mon_bw, c custom mon_bw) : monarr_scope. +Notation "'α_' a ',' b ',' c ⁻¹" := + (monarrstruct (a%monarr_scope ⨂ (b%monarr_scope ⨂ c%monarr_scope)) + (a%monarr_scope ⨂ b%monarr_scope ⨂ c%monarr_scope) _) + (in custom mon at level 20, + a custom mon_bw, b custom mon_bw, c custom mon_bw) : monarr_scope. +Notation "'λ_' a" := (monarrstruct (e ⨂ a%monarr_scope) a%monarr_scope _) + (in custom mon at level 20, a custom mon_bw) : monarr_scope. +Notation "'λ_' a ⁻¹" := (monarrstruct a%monarr_scope (e ⨂ a%monarr_scope) _) + (in custom mon at level 20, a custom mon_bw) : monarr_scope. +Notation "'ρ_' a" := (monarrstruct (a%monarr_scope ⨂ e) a%monarr_scope _) + (in custom mon at level 20, a custom mon_bw) : monarr_scope. +Notation "'ρ_' a ⁻¹" := (monarrstruct a%monarr_scope (a%monarr_scope ⨂ e) _) + (in custom mon at level 20, a custom mon_bw) : monarr_scope. +Notation "'{' a '⟶' b '}'" := + (monarrstruct a%monarr_scope b%monarr_scope _) + (in custom mon at level 5, + a custom mon_bw (* at level 99 *), b custom mon_bw (* at level 99 *), + only printing) : monarr_scope. +Notation "'id_' a" := (monarrstruct a%monarr_scope a%monarr_scope _) + (in custom mon at level 20, a custom mon_bw, only printing) : monarr_scope. +Notation "a" := (var a%monarr_scope) + (in custom mon_bw at level 10, a constr at level 9, + only printing) : monarr_scope. +Notation "a × b" := (tens a%monarr_scope b%monarr_scope) + (in custom mon_bw at level 20, + a custom mon_bw, b custom mon_bw, + left associativity, + only printing) : monarr_scope. +Notation "'(' a ')'" := a%monarr_scope + (in custom mon_bw at level 0, + a custom mon_bw at level 200) : monarr_scope. +Notation "'(' a ')'" := a%monarr_scope + (in custom mon at level 0, + a custom mon at level 200) : monarr_scope. +Notation "f ⊗ g" := (monarrtens f%monarr_scope g%monarr_scope) + (in custom mon at level 34, + f custom mon, g custom mon, + left associativity) : monarr_scope. +Notation "{ a }" := (mongeneric a%monarr_scope) + (in custom mon at level 10) : monarr_scope. +Notation "f ∘ g" := (monarrcomp f%monarr_scope g%monarr_scope) + (in custom mon at level 40, + f custom mon, g custom mon (* at level 40 *), + left associativity) : monarr_scope. +Notation "f ≊ g" := (monarrequiv _ _ f%monarr_scope g%monarr_scope) + (in custom mon at level 70, f custom mon, g custom mon) : monarr_scope. +Notation "a" := a (in custom mon at level 0, a constr) : monarr_scope. + +Notation "''Monarr[' f '≊' g ']'" := + (_ ≡ _ /\ monarr_under_rel _ _ f%monarr_scope g%monarr_scope) + (f custom mon, g custom mon, only printing). *) + +Import MCDefinitions. + +Declare Custom Entry mon. +Declare Custom Entry mon_bw. + +Declare Scope monarr_scope. +Delimit Scope monarr_scope with monarr_scope. +Close Scope monarr_scope. + +Import MC_notations. + +(* Arguments monarrstruct {_ _ _} _ _. *) + +Notation "'α_' a ',' b ',' c" := + (@monarrstruct _ _ _ + (tens (tens + a%monarr_scope b%monarr_scope) c%monarr_scope) + (tens a%monarr_scope + (tens b%monarr_scope c%monarr_scope)) _) + (in custom mon at level 20, + a custom mon_bw, b custom mon_bw, c custom mon_bw) : monarr_scope. +Notation "'α_' a ',' b ',' c ⁻¹" := + (@monarrstruct _ _ _ + (tens a%monarr_scope (tens b%monarr_scope c%monarr_scope)) + (tens (tens a%monarr_scope b%monarr_scope) c%monarr_scope) _) + (in custom mon at level 20, + a custom mon_bw, b custom mon_bw, c custom mon_bw) : monarr_scope. +Notation "'λ_' a" := (@monarrstruct _ _ _ (tens e a%monarr_scope) a%monarr_scope _) + (in custom mon at level 20, a custom mon_bw) : monarr_scope. +Notation "'λ_' a ⁻¹" := (@monarrstruct _ _ _ a%monarr_scope (tens e a%monarr_scope) _) + (in custom mon at level 20, a custom mon_bw) : monarr_scope. +Notation "'ρ_' a" := (@monarrstruct _ _ _ (tens a%monarr_scope e) a%monarr_scope _) + (in custom mon at level 20, a custom mon_bw) : monarr_scope. +Notation "'ρ_' a ⁻¹" := (@monarrstruct _ _ _ a%monarr_scope (tens a%monarr_scope e) _) + (in custom mon at level 20, a custom mon_bw) : monarr_scope. +Notation "'{' a '⟶' b '}'" := + (@monarrstruct _ _ _ a%monarr_scope b%monarr_scope _) + (in custom mon at level 5, + a custom mon_bw (* at level 99 *), b custom mon_bw (* at level 99 *), + only printing) : monarr_scope. +Notation "'id_' a" := (@monarrstruct _ _ _ a%monarr_scope a%monarr_scope _) + (in custom mon at level 20, a custom mon_bw, only printing) : monarr_scope. +Notation "a" := (var a%monarr_scope) + (in custom mon_bw at level 10, a constr at level 9, + only printing) : monarr_scope. +Notation "'e'" := (@e _) + (in custom mon_bw at level 0, only printing) : monarr_scope. +Notation "a × b" := (tens a%monarr_scope b%monarr_scope) + (in custom mon_bw at level 20, + a custom mon_bw, b custom mon_bw, + left associativity, + only printing) : monarr_scope. +Notation "'(' a ')'" := a%monarr_scope + (in custom mon_bw at level 0, + a custom mon_bw at level 200) : monarr_scope. +Notation "'(' a ')'" := a%monarr_scope + (in custom mon at level 0, + a custom mon at level 200) : monarr_scope. +Notation "f ⊗ g" := (monarrtens f%monarr_scope g%monarr_scope) + (in custom mon at level 34, + f custom mon, g custom mon, + left associativity) : monarr_scope. +Notation "{ a }" := (mongeneric a%monarr_scope) + (in custom mon at level 10) : monarr_scope. +Notation "f ∘ g" := (monarrcomp f%monarr_scope g%monarr_scope) + (in custom mon at level 40, + f custom mon, g custom mon (* at level 40 *), + left associativity) : monarr_scope. +Notation "f ≊ g" := (monarrequiv _ _ f%monarr_scope g%monarr_scope) + (in custom mon at level 70, f custom mon, g custom mon) : monarr_scope. +Notation "a" := a (in custom mon at level 0, a constr) : monarr_scope. + +(* Notation "''Monarr[' f '≊' g ']'" := + (realize_equiv' f%monarr_scope g%monarr_scope) + (f custom mon, g custom mon, only printing). *) + +Notation "''Cat[' f '≃' g ']'" := + (realize_equiv' _ _ f%monarr_scope g%monarr_scope) + (f custom mon, g custom mon, only printing). + +Section Testing. + +Import MCClasses CategoryTypeclass. + +Context {X : Type} {UIPX : UIP X}. +Context {cC : Category X} {cCh : CategoryCoherence cC} + {mC : MonoidalCategory cC} {mCh : MonoidalCategoryCoherence mC}. + + +(* Print Grammar tactic. *) +(* Ltac monarr_etransitivity := + match goal with + | |- monarr_under_rel _ _ _ _ => eapply monarr_under_rel_trans_no_evars + | _ => fail 1 "Not a monarr goal (of the form 'Monarr[ _ ≊ _ ])" + end. + +Ltac monarr_over := + match goal with + |- monarr_under_rel _ _ _ _ => exact (fun x => x) + end. + +Ltac monarr_under := + refine (realize_monarr_mor_under _ _ _ _ (_ : monarr_under_rel _ _ _ _) _). *) + +Import MCprocessing MCconsequences MCmonarrlist List. + +Ltac full_process t := + let proc := eval compute in (full_process_monarr t) in + match type of proc with + | @monarr ?X ?cC ?mC ?from_proc ?to_proc => + match type of t with + | monarr ?A ?B => + let RHS := constr:(monarrcomp (monarrcomp + (@monarrstruct X cC mC A from_proc (bwarr_of_Nf_eq (full_process_in_pf t))) + proc) (@monarrstruct X cC mC to_proc A (bwarr_of_Nf_eq (full_process_out_pf t)))) in + let rw := constr:(MCprocessing.full_process_monarr_equiv t : monarrequiv _ _ t RHS) in + rewrite rw + end + end. + + +(* Goal forall (A B M N P Q : X) (f : (A×A ~> B×B)%Cat) (g : (M ~> N)%Cat) +(h : (P ~> Q)%Cat), (α_ (A×A), M, P ∘ f ⊗ (g ⊗ h) ≃ f ⊗ g ⊗ h ∘ α_ (B×B), N, Q)%Cat. +intros. +change + (realize_equiv' _ _ (arrinvassoc _ _ _ ◌ + mongeneric (a := tens (var A) (var A)) (b := tens (var B) (var B)) f + ⧆ (mongeneric (a:=var M) (b:=var N) g + ⧆ mongeneric (a:=var P) (b:=var Q) h)) + ( + mongeneric (a := tens (var A) (var A)) (b := tens (var B) (var B)) f + ⧆ mongeneric (a:=var M) (b:=var N) g + ⧆ mongeneric (a:=var P) (b:=var Q) h ◌ arrinvassoc _ _ _))%bw. +rewrite monarr_invassoc_nat_l. + *) + +(*Ltac full_process t := + let pf := constr:(MCprocessing.full_process_monarr_equiv t) in + (* let T := type of pf in idtac T; *) + lazymatch type of pf with + | monarrequiv ?a ?b ?LHS ?RHS => + match RHS with + | monarrcomp (monarrcomp + (@monarrstruct ?X ?cC ?mC ?inL ?outL ?pfL) + ?center) + (@monarrstruct _ _ _ ?inR ?outR ?pfR) => + let cinL := eval vm_compute in inL in + let coutL := eval vm_compute in outL in + let cinR := eval vm_compute in inR in + let coutR := eval vm_compute in outR in + let ccenter := eval vm_compute in center in + let rw := constr:(pf : monarrequiv a b LHS + (monarrcomp (monarrcomp + (@monarrstruct X cC mC cinL coutL pfL) + ccenter) + (@monarrstruct X cC mC cinR coutR pfR))) in + rewrite rw + end + end. + +Ltac full_process' t := + let pf := constr:(MCprocessing.full_process_monarr_equiv t) in + (* let T := type of pf in idtac T; *) + lazymatch type of pf with + | monarrequiv ?a ?b ?LHS ?RHS => + let cRHS := eval vm_compute in RHS in + let rw := constr:(pf : monarrequiv a b LHS cRHS) in + rewrite rw + end. + + +Ltac full_process'' t := + (* match type of t with + | monarr ?A ?B => + *) + let pf := constr:(MCprocessing.full_process_monarr_equiv t) in + let T := type of pf in + let cbnT := eval vm_compute in T in + let rw := constr:(pf : cbnT) in + rewrite rw. *) + +(* Require Import CategoryAutomation. *) + +(* Import MCprocessing MCconsequences MCmonarrlist List. + +Goal forall (A B M N P Q : X) (f : (A×A ~> B×B)%Cat) (g : (M ~> N)%Cat) +(h : (P ~> Q)%Cat), (α_ (A×A), M, P ∘ f ⊗ (g ⊗ h) ≃ f ⊗ g ⊗ h ∘ α_ (B×B), N, Q)%Cat. +intros. +change + (realize_equiv' _ _ (arrinvassoc _ _ _ ◌ + mongeneric (a := tens (var A) (var A)) (b := tens (var B) (var B)) f + ⧆ (mongeneric (a:=var M) (b:=var N) g + ⧆ mongeneric (a:=var P) (b:=var Q) h)) + ( + mongeneric (a := tens (var A) (var A)) (b := tens (var B) (var B)) f + ⧆ mongeneric (a:=var M) (b:=var N) g + ⧆ mongeneric (a:=var P) (b:=var Q) h ◌ arrinvassoc _ _ _))%bw. +(* change + (realize_monarr (arrinvassoc _ _ _ ◌ + mongeneric (a := tens (var A) (var A)) (b := tens (var B) (var B)) f + ⧆ (mongeneric (a:=var M) (b:=var N) g + ⧆ mongeneric (a:=var P) (b:=var Q) h)) + ≃ realize_monarr ( + mongeneric (a := tens (var A) (var A)) (b := tens (var B) (var B)) f + ⧆ mongeneric (a:=var M) (b:=var N) g + ⧆ mongeneric (a:=var P) (b:=var Q) h ◌ arrinvassoc _ _ _))%Cat%bw. *) + +(* rewrite mon_struct_l, monarr_assoc_nat_r. *) + +time (match goal with +|- realize_equiv' _ _ ?t ?g => +let proc := eval compute in (full_process_monarr t) in +match type of proc with +| @monarr ?X ?cC ?mC ?from_proc ?to_proc => +match type of t with +| monarr ?A ?B => + let RHS := constr:(monarrcomp (monarrcomp + (@monarrstruct X cC mC A from_proc (bwarr_of_Nf_eq (full_process_in_pf t))) + proc) (@monarrstruct X cC mC to_proc B (bwarr_of_Nf_eq (full_process_out_pf t)))) in + let rw := constr:(MCprocessing.full_process_monarr_equiv t : monarrequiv _ _ t RHS) in + time "rw" (rewrite rw) +end +end +end). + +(* time *) (do 10 (try (match goal with +|- realize_equiv' _ _ ?f ?g => + full_process f + (* let t := eval compute in (full_process_monarr f) in + idtac *) +end; fail))). + +(* time *) (do 10 (try (match goal with +|- realize_equiv' _ _ ?f ?g => + full_process_new f + (* let t := eval compute in (full_process_monarr f) in + idtac *) +end; fail))). + +time (do 10 (try (match goal with +|- realize_equiv' _ _ ?f ?g => + let t := eval compute in (full_process_monarr f) in + idtac +end; fail))). + + +(* match goal with +|- realize_equiv' _ _ ?f ?g => + let t := eval compute in (foliate_monarr f) in + idtac t +end. *) + +time (do 10 (try (match goal with +|- realize_equiv' _ _ ?f ?g => + let t := eval compute in (foliate_monarr f) in + idtac +end; fail))). + +let f := match goal with + |- realize_equiv' _ _ ?f ?g => constr:(f) + end in +let fol := eval compute in (foliate_monarr f) +in +time (do 10 ( + let t := eval compute in (map structify_monarrlist fol) in + idtac)). + +let f := match goal with + |- realize_equiv' _ _ ?f ?g => constr:(f) + end in +let fol := eval compute in (map structify_monarrlist (foliate_monarr f)) +in +time (do 10 ( + let t := eval compute in (map remove_structs fol) in + idtac)). + + +let f := match goal with + |- realize_equiv' _ _ ?f ?g => constr:(f) + end in +let fol := eval compute in + (map remove_structs (map structify_monarrlist (foliate_monarr f))) +in +time (do 10 ( + let t := eval compute in (map split_ids fol) in + idtac)). + + +let f := match goal with + |- realize_equiv' _ _ ?f ?g => constr:(f) + end in +let fol := eval compute in + (map split_ids (map remove_structs (map structify_monarrlist (foliate_monarr f)))) +in +time (do 10 ( + let t := eval compute in (full_right_shift fol) in + idtac)). + + + + +apply MCconsequences.monarrcomp_struct_l. +monarr_under. + +Time (do 10 (try (match goal with +|- monarr_under_rel _ _ ?f ?g => + full_process f + (* full_process f; full_process g *) +end; fail))). + +(* epose MCprocessing.full_process_monarr as T. +unfold MCprocessing.full_process_monarr in T. *) + +(* rewrite <- MCmonarrlist.monarrequiv_iff_monarr_norm_equiv. *) + +match goal with +|- monarr_under_rel _ _ ?f ?g => + full_process f + (* full_process f; full_process g *) +end. + +Time (do 10 (try (match goal with +|- monarr_under_rel _ _ ?f ?g => + full_process f + (* full_process f; full_process g *) +end; fail))). + + +rewrite monarr_under_struct_l. +rewrite monarr_invassoc_nat_r. +rewrite monarr_assoc. +rewrite monarr_arrcomp. +rewrite monarr_id_l. +monarr_over. +easy. +Qed. + + +Goal forall (A B M N : X), +(α_ A, B, M ⊗ id_ N ∘ α_ A, B × M, N ∘ id_ A ⊗ α_ B, M, N + ≃ α_ A × B, M, N ∘ α_ A, B, (M × N))%Cat. + intros. + change + (realize_monarr (monarrcomp (arrid (var M)) + (mongeneric (a:=var M) (b:=var N) g)) + ≃ realize_monarr (mongeneric (a:=var M) (b:=var N) g))%Cat. + monarr_under. + + +Goal forall (A B M N P Q : X) (f : (A×A ~> B×B)%Cat) (g : (M ~> N)%Cat) + (h : (P ~> Q)%Cat), (id_ M ∘ g ≃ g)%Cat. +intros. +change + (realize_monarr (monarrcomp (arrid (var M)) + (mongeneric (a:=var M) (b:=var N) g)) + ≃ realize_monarr (mongeneric (a:=var M) (b:=var N) g))%Cat. +monarr_under. +pentagon +full_process (@monarrcomp X cC mC (@var X M) (@var X M) (@var X N) +(@arrid X (@var X M)) (@mongeneric X cC mC (@var X M) (@var X N) g)). + +cbn. + +repeat match goal with +|- context[MCmonarrlist.monarrlist_list_source ?l] => + let simpled := eval cbn in (MCmonarrlist.monarrlist_list_source l) in + change (MCmonarrlist.monarrlist_list_source l) with simpled + +|- context[MCmonarrlist.monarrlist_list_source ?l] => +let simpled := eval cbn in (MCmonarrlist.monarrlist_list_source l) in +change (MCmonarrlist.monarrlist_list_source l) with simpled +end. +cbn [MCmonarrlist.monarrlist_list_source + MCprocessing.full_process_monarrlist_list + MCprocessing.foliate_monarr + (* MCmonarrlist.monarrlist_arr *)]. +simpl. +rewrite MCprocessing.full_process_monarr_equiv. +cbn [MCprocessing.full_process_monarr]. +monarr_etransitivity. +rewrite monarr_id_l. +monarr_over. +easy. +Qed. + + +Goal forall (A B M N P Q : X) (f : (A×A ~> B×B)%Cat) (g : (M ~> N)%Cat) +(h : (P ~> Q)%Cat), (α_ (A×A), M, P ∘ f ⊗ (g ⊗ h) ≃ f ⊗ g ⊗ h ∘ α_ (B×B), N, Q)%Cat. +intros. +change + (realize_monarr (arrinvassoc _ _ _ ◌ + mongeneric (a := tens (var A) (var A)) (b := tens (var B) (var B)) f + ⧆ (mongeneric (a:=var M) (b:=var N) g + ⧆ mongeneric (a:=var P) (b:=var Q) h)) + ≃ realize_monarr ( + mongeneric (a := tens (var A) (var A)) (b := tens (var B) (var B)) f + ⧆ mongeneric (a:=var M) (b:=var N) g + ⧆ mongeneric (a:=var P) (b:=var Q) h ◌ arrinvassoc _ _ _))%Cat. +monarr_under. +rewrite monarr_under_struct_l. +rewrite monarr_invassoc_nat_r. +rewrite monarr_assoc. +rewrite monarr_arrcomp. +rewrite monarr_id_l. +monarr_over. +easy. +Qed. + + + +(* Arguments monarrstruct {_ _ _ _ _}. *) + +Section Testing. +(* +Local Ltac monarr_over := + match goal with + |- monarr_under_rel _ _ _ _ => exact (fun x => x); try reflexivity + end. + +Local Ltac monarr_under := + refine (realize_monarr_mor_under _ _ _ _ (_ : monarr_under_rel _ _ _ _) _). + + +Goal forall (A B M N P Q : X) (f : (A×A ~> B×B)%Cat) (g : (M ~> N)%Cat) + (h : (P ~> Q)%Cat), (id_ M ∘ g ≃ g)%Cat. +intros. +change + (realize_monarr (arrid (var M) ◌ + mongeneric (a:=var M) (b:=var N) g) + ≃ realize_monarr (mongeneric (a:=var M) (b:=var N) g))%Cat. +monarr_under. +rewrite monarr_id_l. +monarr_over. +easy. +Qed. + + +Goal forall (A B M N P Q : X) (f : (A×A ~> B×B)%Cat) (g : (M ~> N)%Cat) +(h : (P ~> Q)%Cat), (α_ (A×A), M, P ∘ f ⊗ (g ⊗ h) ≃ f ⊗ g ⊗ h ∘ α_ (B×B), N, Q)%Cat. +intros. +change + (realize_monarr (arrinvassoc _ _ _ ◌ + mongeneric (a := tens (var A) (var A)) (b := tens (var B) (var B)) f + ⧆ (mongeneric (a:=var M) (b:=var N) g + ⧆ mongeneric (a:=var P) (b:=var Q) h)) + ≃ realize_monarr ( + mongeneric (a := tens (var A) (var A)) (b := tens (var B) (var B)) f + ⧆ mongeneric (a:=var M) (b:=var N) g + ⧆ mongeneric (a:=var P) (b:=var Q) h ◌ arrinvassoc _ _ _))%Cat. +monarr_under. +rewrite monarr_under_struct_l. +rewrite monarr_invassoc_nat_r. +rewrite monarr_assoc. +rewrite monarr_arrcomp. +rewrite monarr_id_l. +monarr_over. +easy. +Qed. +*) *) +End Testing. + +Set Universe Polymorphism. + +Import -(notations) Category Monoidal. + +Ltac evarT f T := + let _ := match goal with + |- _ => + evar (f : T) + end in + let x' := eval unfold f in f in + uconstr:(x'). + +Ltac quote_to_bw term C cC mC := + let rec quote term := + match term with + | mC.(mon_I) => constr:(@e C) + | mC.(obj_tensor) ?a ?b => + let qa := quote a in + let qb := quote b in + constr:(@tens C qa qb) + | ?t => + let a := fresh in let b := fresh in + let _ := match goal with + |- _ => evar (a : C); evar (b : C) + end in + let a' := eval unfold a in a in + let b' := eval unfold b in b in + let _ := match goal with + | |- _ => unify (mC.(obj_tensor) a' b') t + (* | |- _ => clear a b; fail 1 *) + (* ; idtac "unified as tensor of" a' "and" b' *) + end in + let qa' := quote a' in + let qb' := quote b' in + let _ := match goal with + |- _ => clear a b + end in + constr:(@tens C qa' qb') + | ?t => constr:(@var C t) + end in + quote term. + +Ltac quote_to_bwarr term C cC mC := + let bw_quote A := quote_to_bw A C cC mC in + let rec quote term := + match term with + | cC.(compose) ?f ?g => + let qf := quote f in + let qg := quote g in + constr:(@arrcomp C _ _ _ qf qg) + | mC.(mor_tensor) ?f ?g => + let qf := quote f in + let qg := quote g in + constr:(@arrtens C _ _ _ qf qg) + | cC.(c_identity) ?A => + let qA := bw_quote A in + constr:(arrid (* C *) qA) + | forward (mC.(associator) ?A ?B ?M) => + let _ := match goal with |- _ => idtac "found assoc directly" term end in + let qA := bw_quote A in + let qB := bw_quote B in + let qM := bw_quote M in + let _ := match goal with |- _ => idtac "quoted successfully" end in + constr:(arrinvassoc (* C *) qA qB qM) + | reverse (mC.(associator) ?A ?B ?M) => + let _ := match goal with |- _ => idtac "found reverse assoc directly" term end in + let qA := bw_quote A in + let qB := bw_quote B in + let qM := bw_quote M in + let _ := match goal with |- _ => idtac "quoted successfully" end in + constr:(arrassoc (* C *) qA qB qM) + | forward (mC.(left_unitor) ?A) => + let qA := bw_quote A in + constr:(arrlunitor (* C *) qA) + | reverse (mC.(left_unitor) ?A) => + let qA := bw_quote A in + constr:(arrinvlunitor (* C *) qA) + | forward (mC.(right_unitor) ?A) => + let qA := bw_quote A in + constr:(arrrunitor (* C *) qA) + | reverse (mC.(right_unitor) ?A) => + let qA := bw_quote A in + constr:(arrinvrunitor (* C *) qA) + | ?t => + let A' := fresh in let B' := fresh in + let A := evarT A' C in + let B := evarT B' C in + let _ := match goal with + |- _ => + let T := type of t in + unify (cC.(morphism) A B) T + end in + let qA := bw_quote A in + let qB := bw_quote B in + let _ := match goal with |- _ => clear A' B' end in + match t with + | ?t => + let _ := match goal with + |- _ => unify (cC.(c_identity) A) t + (* ; clear A' B' *) + end in + constr:(arrid (* C *) qA) + | ?t => + let _ := match goal with + |- _ => unify (forward (mC.(left_unitor) B)) t + (* ; clear A' B' *) + end in + constr:(arrlunitor (* C *) qB) + | ?t => + let _ := match goal with + |- _ => unify (reverse (mC.(left_unitor) A)) t + (* ; clear A' B' *) + end in + constr:(arrinvlunitor (* C *) qA) + | ?t => + let _ := match goal with + |- _ => unify (forward (mC.(right_unitor) B)) t + (* ; clear A' B' *) + end in + constr:(arrrunitor (* C *) qB) + | ?t => + let _ := match goal with + |- _ => unify (reverse (mC.(right_unitor) A)) t + (* ; clear A' B' *) + end in + constr:(arrinvrunitor (* C *) qA) + | ?t => + let M' := fresh in let f' := fresh in let g' := fresh in + let M := evarT M' C in + let f := evarT f' (cC.(morphism) A M) in + let g := evarT g' (cC.(morphism) M B) in + let _ := match goal with + | |- _ => unify (cC.(compose) f g) t + (* | |- _ => clear M' f' g'; fail 1 *) + end in + let qf := quote f in + let qg := quote g in + let _ := match goal with + |- _ => clear M' f' g' + end in + constr:(arrcomp (* C *) qf qg) + | ?t => + let A1' := fresh in let B1' := fresh in let A2' := fresh in + let B2' := fresh in let f' := fresh in let g' := fresh in + let A1 := evarT A1' C in + let B1 := evarT B1' C in + let A2 := evarT A2' C in + let B2 := evarT B2' C in + let f := evarT f' (cC.(morphism) A1 B1) in + let g := evarT g' (cC.(morphism) A2 B2) in + let _ := match goal with + | |- _ => unify (mC.(mor_tensor) f g) t + (* | |- _ => clear A1' B1' A2' B2' f' g' *) + end in + let qf := quote f in + let qg := quote g in + let _ := match goal with + |- _ => clear A1' B1' A2' B2' f' g' + end in + constr:(arrtens (* C *) qf qg) + | ?t => + let A'' := fresh in let B'' := fresh in let M'' := fresh in + let A' := evarT A'' C in + let B' := evarT B'' C in + let M' := evarT M'' C in + let _ := match goal with + | |- _ => unify (forward (mC.(associator) A' B' M')) t + (* ; idtac "found associator" t *) + (* | |- _ => clear A'' B'' M'' *) + end in + let qA' := bw_quote A' in + let qB' := bw_quote B' in + let qM' := bw_quote M' in + let _ := match goal with + |- _ => clear A'' B'' M'' + end in + constr:(arrinvassoc (* C *) qA' qB' qM') + | ?t => + let A'' := fresh in let B'' := fresh in let M'' := fresh in + let A' := evarT C in + let B' := evarT C in + let M' := evarT C in + let _ := match goal with + | |- _ => unify (reverse (mC.(associator) A' B' M')) t + (* | |- _ => clear A'' B'' M'' *) + end in + let qA' := bw_quote A' in + let qB' := bw_quote B' in + let qM' := bw_quote M' in + let _ := match goal with + |- _ => clear A'' B'' M'' + end in + constr:(arrassoc (* C *) qA' qB' qM') + end + end + in quote term. + +Ltac quote_to_monarr term C cC mC := + let quote ter := quote_to_monarr ter C cC mC in + let bw_quote A := quote_to_bw A C cC mC in + let bwarr_quote f := quote_to_bwarr f C cC mC in + (* let _ := match goal with |- _ => idtac "quoting" term "to monarr in" C cC mC end in *) + let t' := constr:(term) in + (* let _ := match goal with |- _ => idtac "quoting" t' "to monarr in" C cC mC end in *) + match t' with + | ?t => + (* let _ := match goal with |- _ => idtac "try bwarr..." end in *) + let qt := bwarr_quote t in + let _ := match goal with |- _ => idtac "found bwarr" qt end in + constr:(@monarrstruct C cC mC _ _ qt) + | cC.(compose) ?f ?g => + (* let _ := match goal with |- _ => idtac "compose match for monarr" end in *) + let qf := quote f in + let qg := quote g in + constr:(@monarrcomp C cC mC _ _ _ qf qg) + | mC.(mor_tensor) ?f ?g => + (* let _ := match goal with |- _ => idtac "tensor match for monarr" end in *) + let qf := quote f in + let qg := quote g in + constr:(@monarrtens C cC mC _ _ _ qf qg) + | ?t => + (* let _ := match goal with |- _ => idtac "nonsimple match for monarr" end in *) + let A' := fresh in + let B' := fresh in + let _ := match goal with + |- _ => evar (A' : C); evar (B' : C) + (* ; idtac "made A' and B'" *) + end in + let A := eval unfold A' in A' in + let B := eval unfold B' in B' in + let _ := match goal with + | |- _ => + (* idtac A B A' B'; *) + let T := type of t in + unify (cC.(morphism) A B) T + (* | |- _ => clear A' B'; fail 2 t "not a morphism" *) + (* ;idtac "unified as morphism" A "to" B *) + end in + (* let _ := match goal with |- _ => idtac A B A' B' end in *) + let qA := bw_quote A in + let qB := bw_quote B in + let _ := match goal with |- _ => clear A' B' end in + (* let _ := match goal with |- _ => idtac "here" end in *) + match t with + | ?t => + let M' := fresh in let f' := fresh in let g' := fresh in + let _ := match goal with + |- _ => evar (M' : C) + end in + let M := eval unfold M' in M' in + let _ := match goal with + |- _ => evar (f' : (cC.(morphism) A M)); + evar (g' : cC.(morphism) M B) + end in + let f := eval unfold f' in f' in + let g := eval unfold g' in g' in + let _ := match goal with + | |- _ => unify (cC.(compose) f g) t + (* | |- _ => clear M' f' g'; fail 1 *) + (* ; idtac "unified as composition of" f "and" g *) + end in + let qf := quote f in + let qg := quote g in + let _ := match goal with |- _ => clear M' f' g' end in + constr:(@monarrcomp C cC mC _ _ _ qf qg) + | ?t => + let A1' := fresh in let B1' := fresh in + let A2' := fresh in let B2' := fresh in + let f' := fresh in let g' := fresh in + let A1 := evarT A1' C in + let B1 := evarT B1' C in + let A2 := evarT A2' C in + let B2 := evarT B2' C in + let f := evarT f' (cC.(morphism) A1 B1) in + let g := evarT g' (cC.(morphism) A2 B2) in + let _ := match goal with + | |- _ => unify (mC.(mor_tensor) f g) t + (* | |- _ => clear A1' A2' B1' B2' f' g'; fail 1 *) + (* ; idtac "unified as tensor of" f "and" g *) + end in + (* let _ := match goal with + |- _ => idtac f g + end in *) + let qf := quote f in + (* let _ := match goal with + |- _ => idtac qf + end in *) + let qg := quote g in + (* let _ := match goal with + |- _ => idtac qf qg + end in *) + let _ := match goal with |- _ => clear A1' A2' B1' B2' f' g' end in + constr:(@monarrtens C cC mC _ _ _ _ qf qg) + | ?t => + (* let _ := match goal with |- _ => idtac qA qB end in *) + constr:(@mongeneric C cC mC qA qB t) + end + end. + +Ltac quote_equiv_to_monarr mC := + lazymatch type of mC with + | @MonoidalCategory ?X ?cC => + let monarr_quote ter := quote_to_monarr ter X cC mC in + let bw_quote A := quote_to_bw A X cC mC in + let bwarr_quote f := quote_to_bwarr f X cC mC in + let A := fresh "A" in + let B := fresh "B" in + let f := fresh "f" in + let g := fresh "g" in + evar (A : X); evar (B : X); + let A' := eval unfold A in A in + let B' := eval unfold B in B in + evar (f : cC.(morphism) A' B'); + evar (g : cC.(morphism) A' B'); + let f' := eval unfold f in f in + let g' := eval unfold g in g in + let GOAL := lazymatch goal with |- ?G => constr:(G) end in + unify (@c_equiv X cC A' B' f' g') GOAL; + let qA := bw_quote A' in + let qB := bw_quote B' in + let qf := monarr_quote f' in + let qg := monarr_quote g' in + clear f g A B; + change (@realize_equiv' X cC mC qA qB qf qg) + | ?T => fail "couldn't see as MonoidalCategory instance:" mC "(with type" T ")" + end. + +Tactic Notation "monoidal" constr(mC) := + quote_equiv_to_monarr mC; + apply by_reflexive_trim_foliation_no_empty; easy. + +Tactic Notation "monoidal" := + let x := fresh in + let o := open_constr:(MonoidalCategory _) in + unshelve evar (x:o); + [typeclasses eauto|..]; + let x' := eval unfold x in x in + monoidal x'. \ No newline at end of file diff --git a/ViCaR/MonoidalCoherence/MCbw.v b/ViCaR/MonoidalCoherence/MCbw.v new file mode 100644 index 0000000..952e7c0 --- /dev/null +++ b/ViCaR/MonoidalCoherence/MCbw.v @@ -0,0 +1,1170 @@ +Require Import MCDefinitions. +Require UIP_facts. +Require MCTheory. +Require CatCategory. + + +Section bw_theory. + +Set Universe Polymorphism. + +Import MC_notations MC_setoids. + +Open Scope bw_scope. + +Polymorphic Context {X : Type}. +#[local] Set Universe Polymorphism. +Local Notation bw := (bw X). +Local Notation bwnorm := (bwnorm X). +Local Notation cast_bwarr := (@cast_bwarr X). + +Section bwnorm_theory. + +Lemma varlist_to_bwnorm_app (l l' : list X) : + varlist_to_bwnorm (l ++ l') = + bwnormapp (varlist_to_bwnorm l') (varlist_to_bwnorm l). +Proof. + induction l; [easy|]. + now simpl; f_equal. +Qed. + +Lemma Nf_tens (b c : bw) : + Nf (b ⨂ c) = ⟦c⟧ (Nf b). +Proof. + easy. +Qed. + +Lemma Nf_tens_assoc (b c d : bw) : + Nf (b ⨂ (c ⨂ d)) = Nf (b ⨂ c ⨂ d). +Proof. + easy. +Qed. + +Lemma bwnormapp_lnorme (n : bwnorm) : + bwnormapp norm_e n = n. +Proof. + now induction n; auto; simpl; f_equal. +Qed. + +Lemma bwnormapp_assoc (n m o : bwnorm) : + bwnormapp n (bwnormapp m o) = bwnormapp (bwnormapp n m) o. +Proof. + induction o; [easy|]. + now simpl; f_equal. +Qed. + +Lemma bwnorm_to_varlist_to_bwnorm (a : bwnorm) : + varlist_to_bwnorm (bwnorm_to_varlist a) = a. +Proof. + induction a; [easy|]. + now simpl; f_equal. +Qed. + +Lemma varlist_to_bwnorm_to_varlist (a : list X) : +bwnorm_to_varlist (varlist_to_bwnorm a) = a. +Proof. + induction a; [easy|]. + now simpl; f_equal. +Qed. + +Lemma varlist_to_bwnorm_inj (a b : list X) : + varlist_to_bwnorm a = varlist_to_bwnorm b -> a = b. +Proof. + intros H; + apply (f_equal bwnorm_to_varlist) in H. + rewrite !varlist_to_bwnorm_to_varlist in H. + easy. +Qed. + +Lemma bwnorm_to_varlist_inj (a b : bwnorm) : + bwnorm_to_varlist a = bwnorm_to_varlist b -> a = b. +Proof. + intros H; + apply (f_equal varlist_to_bwnorm) in H. + rewrite !bwnorm_to_varlist_to_bwnorm in H. + easy. +Qed. + + +Lemma varlist_to_bwnorm_eq_iff (a b : list X) : + varlist_to_bwnorm a = varlist_to_bwnorm b <-> a = b. +Proof. + split. + - apply varlist_to_bwnorm_inj. + - intros; subst; easy. +Qed. + +Lemma bwnorm_eq_iff_varlist_eq {a b : bwnorm} : + a = b <-> bwnorm_to_varlist a = bwnorm_to_varlist b. +Proof. + rewrite <- varlist_to_bwnorm_eq_iff. + rewrite !bwnorm_to_varlist_to_bwnorm. + easy. +Qed. + +Lemma Nf_tens_bwnormapp (b c : bw) : + Nf (b ⨂ c) = bwnormapp (Nf b) (Nf c). +Proof. + revert b; induction c; intros b; [easy..|]. + rewrite Nf_tens_assoc. + rewrite Nf_tens, IHc1, IHc2. + rewrite bwnormapp_assoc. + rewrite <- IHc1, <- IHc2. + easy. +Qed. + +Lemma Nf_tens_eq {a b c d : bw} (H : Nf a = Nf b) (H' : Nf c = Nf d) : + Nf (a ⨂ c) = Nf (b ⨂ d). +Proof. + rewrite 2!Nf_tens_bwnormapp, H, H'. + easy. +Qed. + +Lemma Nf_eq_bwnorm_of_varlist (b : bw) : + Nf b = varlist_to_bwnorm (bw_to_varlist b). +Proof. + induction b as [|a|b IHb c IHc]; [easy.. |]. + rewrite Nf_tens_bwnormapp. + simpl. + rewrite varlist_to_bwnorm_app. + now f_equal. +Qed. + +Lemma bwnorm_to_varlist_Nf (b : bw) : + bwnorm_to_varlist (Nf b) = bw_to_varlist b. +Proof. + rewrite (f_equal bwnorm_to_varlist (Nf_eq_bwnorm_of_varlist b)). + now rewrite varlist_to_bwnorm_to_varlist. +Qed. + + +Lemma bwbrac_bwnorm (a b : bwnorm) : ⟦a⟧ b = bwnormapp b a. +Proof. + induction a; [easy|]. + simpl; f_equal; easy. +Qed. + +Lemma Nf_bwnorm (a : bwnorm) : Nf a = a. +Proof. + unfold Nf. + now rewrite bwbrac_bwnorm, bwnormapp_lnorme. +Qed. + +Lemma bw_of_bwnorm_inj (n m : bwnorm) : @eq bw n m -> n = m. +Proof. + intros H. + apply (f_equal Nf) in H. + rewrite !Nf_bwnorm in H. + easy. +Qed. + +Lemma Nf_eq_iff_varlist_eq (a b : bw) : + Nf a = Nf b <-> bw_to_varlist a = bw_to_varlist b. +Proof. + rewrite 2!Nf_eq_bwnorm_of_varlist. + apply varlist_to_bwnorm_eq_iff. +Qed. + +Import List ListNotations. + +Lemma bw_to_varlist_fold_right (fs : list bw) : + bw_to_varlist (fold_right (fun a n => n ⨂ a) e fs) = + concat (map bw_to_varlist fs). +Proof. + induction fs; [easy|]. + simpl. + f_equal. + easy. +Qed. + +Lemma varlist_fold_right_tens hd tl : + bw_to_varlist (X:=X) (fold_right (fun a n => n ⨂ a) e (hd ++ tl)) = + bw_to_varlist (fold_right (fun a n => n ⨂ a) e hd) + ++ bw_to_varlist (fold_right (fun a n => n ⨂ a) e tl). +Proof. + induction hd; [easy|]. + simpl. + rewrite IHhd. + apply app_assoc. +Qed. + +Lemma Nf_fold_right_tens hd tl : + Nf (fold_right (fun a n => n ⨂ a) e (hd ++ tl)) = + bwnormapp (X:=X) + (Nf (fold_right (fun a n => n ⨂ a) e tl)) + (Nf (fold_right (fun a n => n ⨂ a) e hd)). +Proof. + rewrite 3!Nf_eq_bwnorm_of_varlist. + rewrite 3!bw_to_varlist_fold_right. + rewrite map_app, concat_app. + rewrite varlist_to_bwnorm_app. + easy. +Qed. + +(* Lemma bwnormapp_tens_l a b x : + bwnormapp (X:=X) (norm_rtens a x) b = + bwnormapp a (bwnormapp (norm_rtens norm_e x) b). +Proof. + now rewrite bwnormapp_assoc. +Qed. + +Lemma bwnormapp_cancel_var_l {x} {b c : bwnorm} : + bwnormapp (bwnormapp (norm_rtens norm_e x) b) + = bwnormapp (bwnormapp (norm_rtens norm_e x) c) -> b = c. +Proof. + *) + +Lemma bwnorm_to_varlist_app (a b : bwnorm) : + bwnorm_to_varlist (bwnormapp a b) = + bwnorm_to_varlist b ++ bwnorm_to_varlist a. +Proof. + apply varlist_to_bwnorm_inj. + rewrite varlist_to_bwnorm_app. + rewrite !bwnorm_to_varlist_to_bwnorm. + easy. +Qed. + +Lemma bwnormapp_cancel_l {a b c : bwnorm} : + bwnormapp a b = bwnormapp a c -> b = c. +Proof. + rewrite !bwnorm_eq_iff_varlist_eq. + rewrite !bwnorm_to_varlist_app. + apply app_inv_tail. +Qed. + +Lemma bwnormapp_cancel_r {a b c : bwnorm} : + bwnormapp b a = bwnormapp c a -> b = c. +Proof. + rewrite !bwnorm_eq_iff_varlist_eq. + rewrite !bwnorm_to_varlist_app. + apply app_inv_head. +Qed. + +Lemma Nf_cancel_tens_l {a b c : bw} : + Nf (a ⨂ b) = Nf (a ⨂ c) -> Nf b = Nf c. +Proof. + rewrite !Nf_tens_bwnormapp. + apply bwnormapp_cancel_l. +Qed. + +Lemma Nf_cancel_tens_r {a b c : bw} : + Nf (b ⨂ a) = Nf (c ⨂ a) -> Nf b = Nf c. +Proof. + rewrite !Nf_tens_bwnormapp. + apply bwnormapp_cancel_r. +Qed. + +End bwnorm_theory. + + +Lemma bwbrac_of_bweq (a b : bw) : a ~ b -> + forall n : bwnorm, ⟦a⟧ n = ⟦b⟧ n. +Proof. + intros H. + induction H; + auto; + intros n. + - simpl. + rewrite IHbweq1. + apply IHbweq2. + - etransitivity; eauto. +Qed. + +Lemma bwnorm_ltens_bw_eq (a : bw) : forall n : bwnorm, + (tens n a) ~ (bwbrac a n). +Proof. + induction a; [repeat constructor..|]. + eauto using IHa1, IHa2 with bwdb. +Qed. + +Lemma bweq_Nf (a : bw) : a ~ (Nf a). +Proof. + transitivity (tens e a); + [do 2 constructor |]. + apply (bwnorm_ltens_bw_eq a (norm_e)). +Qed. + +Lemma Nf_eq_iff_bweq (a b : bw) : + Nf a = Nf b <-> a ~ b. +Proof. + split. + - intros Heq. + transitivity (Nf a). + 1: apply bweq_Nf. + rewrite Heq. + symmetry. + apply bweq_Nf. + - intros H; induction H; try easy + (etransitivity; eauto). + rewrite 2!Nf_tens_bwnormapp. + f_equal; easy. +Qed. + + +Section bwarr_theory. + +Local Notation "a '⟶' b" := (@bwarr X a b) + (at level 60) : type_scope. (* \longrightarrow *) + +Lemma bweq_of_arr {a b : bw} : a ⟶ b -> a ~ b. +Proof. + intros f. + induction f; eauto with bwdb. +Qed. + +Lemma ex_arr_of_bweq {a b : bw} : a ~ b -> exists (f:a ⟶ b), True. +Proof. + intros H. + induction H; try (destruct IHbweq); try (destruct IHbweq1, IHbweq2); + eexists; eauto 2 using bwarrinv with bwdb. +Qed. + +Lemma bweq_iff_ex_arr {a b : bw} : a ~ b <-> exists (f:a ⟶ b), True. +Proof. + split; [apply ex_arr_of_bweq|]. + intros [f _]; apply bweq_of_arr; easy. +Qed. + + +Lemma bwarrinv_invol {A B} (h : A ⟶ B) : + bwarrinv (bwarrinv h) = h. +Proof. + induction h; try easy; simpl; rewrite IHh1, IHh2; easy. +Qed. + +Lemma bwarrinv_linv {A B} (h : A ⟶ B) : arrcomp (bwarrinv h) h ≅ arrid B. +Proof. + induction h; [eauto 3 with bwarrdb .. ]; simpl. + - rewrite bwarr_assoc, <- (bwarr_assoc (bwarrinv h1)), IHh1; eauto with bwarrdb. + - rewrite <- bwarr_tens_comp, <- bwarr_tens_id. + apply bwarr_tens; auto. +Qed. + +Lemma bwarrinv_rinv {A B} (h : A ⟶ B) : arrcomp h (bwarrinv h) ≅ arrid A. +Proof. + induction h; [eauto 3 with bwarrdb .. ]; simpl. + - rewrite bwarr_assoc, <- (bwarr_assoc h2), IHh2; eauto with bwarrdb. + - rewrite <- bwarr_tens_comp, <- bwarr_tens_id. + apply bwarr_tens; auto. +Qed. + +Lemma bwinv_unique {a b} (f : a ⟶ b) (g g' : b ⟶ a) : + arrcomp f g ≅ arrid a -> arrcomp g' f ≅ arrid b -> + g ≅ g'. +Proof. + intros Hfg Hg'f. + rewrite <- (bwarr_lunit g), <- Hg'f. + rewrite bwarr_assoc, Hfg. + eauto 3 with bwarrdb. +Qed. + +Lemma bwinv_unique' {a b} (f : a ⟶ b) (g g' : b ⟶ a) : + arrcomp g f ≅ arrid b -> arrcomp f g' ≅ arrid a -> + g ≅ g'. +Proof. + intros Hgf Hfg'. + symmetry. + eapply bwinv_unique; eauto. +Qed. + +#[global] Add Parametric Morphism {A B} : (@bwarrinv X A B) + with signature + (bwarrequiv A B) ==> (bwarrequiv B A) + as bwarrinv_mor. +Proof. + intros x y Hxy. + apply bwinv_unique with x. + - apply bwarrinv_rinv. + - rewrite Hxy. + apply bwarrinv_linv. +Qed. + +Lemma by_bwarrinv {a b : bw} (f f' : a ⟶ b) : + bwarrinv f ≅ bwarrinv f' -> f ≅ f'. +Proof. + intros H. + rewrite <- (bwarrinv_invol f), <- (bwarrinv_invol f'). + rewrite H. + easy. +Qed. + + + +Lemma bwarr_invassoc_natl {a b c a' b' c'} (f : a' ⟶ a) + (g : b' ⟶ b) (h : c' ⟶ c) : + arrcomp (arrtens (arrtens f g) h) (arrinvassoc a b c) + ≅ arrcomp (arrinvassoc a' b' c') (arrtens f (arrtens g h)). +Proof. + apply by_bwarrinv, bwarr_assoc_nat. +Qed. + +Lemma bwarr_invlunitor_natl {a b} (f : b ⟶ a) : + arrcomp f (arrinvlunitor a) + ≅ arrcomp (arrinvlunitor b) (arrtens (arrid e) f). +Proof. + apply by_bwarrinv, bwarr_lunitor_nat. +Qed. + +Lemma bwarr_invrunitor_natl {a b} (f : b ⟶ a) : + arrcomp f (arrinvrunitor a) + ≅ arrcomp (arrinvrunitor b) (arrtens f (arrid e)). +Proof. + apply by_bwarrinv, bwarr_runitor_nat. +Qed. + +Lemma bwarr_invpentagonl (a b c d : bw) : + arrcomp (arrinvassoc (a ⨂ b) c d) (arrinvassoc a b (c ⨂ d)) + ≅ arrcomp (arrtens (arrinvassoc a b c) (arrid d)) + (arrcomp (arrinvassoc a (b ⨂ c) d) + (arrtens (arrid a) (arrinvassoc b c d))). +Proof. + apply by_bwarrinv, bwarr_pentagon. +Qed. + +Lemma bwarr_invtrianglel' (a b : bw) : + arrcomp (arrinvassoc a e b) (arrtens (arrid a) (arrlunitor b)) + ≅ arrtens (arrrunitor a) (arrid b). +Proof. + rewrite <- (bwarr_triangle a b). + rewrite <- bwarr_assoc, bwarr_assoc_linv, bwarr_lunit. + easy. +Qed. + +Lemma arrtens_pushout_top {a b c d e : bw} (f : a ⟶ b) (g : b ⟶ c) (h : d ⟶ e) : + arrtens (arrcomp f g) h + ≅ arrcomp (arrtens f h) (arrtens g (arrid e)). +Proof. + rewrite <- bwarr_tens_comp, bwarr_runit. + easy. +Qed. + +Lemma arrtens_pushin_top {a b c d e : bw} (f : a ⟶ b) (g : b ⟶ c) (h : d ⟶ e) : + arrtens (arrcomp f g) h + ≅ arrcomp (arrtens f (arrid d)) (arrtens g h). +Proof. + rewrite <- bwarr_tens_comp, bwarr_lunit. + easy. +Qed. + + Lemma arrtens_pushout_bot {a b c d e : bw} (f : a ⟶ b) (g : c ⟶ d) (h : d ⟶ e) : + arrtens f (arrcomp g h) + ≅ arrcomp (arrtens f g) (arrtens (arrid b) h). +Proof. + rewrite <- bwarr_tens_comp, bwarr_runit. + easy. +Qed. + +Lemma arrtens_pushin_bot {a b c d e : bw} (f : a ⟶ b) (g : c ⟶ d) (h : d ⟶ e) : + arrtens f (arrcomp g h) + ≅ arrcomp (arrtens (arrid a) g) (arrtens f h). +Proof. + rewrite <- bwarr_tens_comp, bwarr_lunit. + easy. +Qed. + +Lemma arrtens_split_diag {a b a' b'} (f : a ⟶ a') (g : b ⟶ b') : + f ⊠ g ≅ f ⊠ arrid b ○ arrid a' ⊠ g. +Proof. + rewrite <- bwarr_tens_comp, bwarr_lunit, bwarr_runit. + easy. +Qed. + + + +Lemma bwarr_trianglel' (a b : bw) : + arrassoc a e b ≅ arrid a ⊠ arrlunitor b ○ arrinvrunitor a ⊠ arrid b. +Proof. + rewrite <- (bwarr_runit (arrassoc a e b)), <- bwarr_tens_id, + <- (bwarr_runitor_rinv), arrtens_pushout_top, <- bwarr_assoc, + bwarr_triangle. + easy. +Qed. + +Lemma bwarr_compose_l {a b c} (f : a ⟶ b) (g : b ⟶ c) (h : a ⟶ c) : + f ○ g ≅ h <-> g ≅ bwarrinv f ○ h. +Proof. + split; intros H; [rewrite <- H | rewrite H]; + rewrite <- bwarr_assoc, ?bwarrinv_linv, ?bwarrinv_rinv, bwarr_lunit; + easy. +Qed. + +Lemma bwarr_compose_l' {a b c} (f : a ⟶ b) (g : b ⟶ c) (h : a ⟶ c) : + h ≅ f ○ g <-> bwarrinv f ○ h ≅ g. +Proof. + split; symmetry; apply bwarr_compose_l; symmetry; assumption. +Qed. + +Lemma bwarr_compose_r {a b c} (f : a ⟶ b) (g : b ⟶ c) (h : a ⟶ c) : + f ○ g ≅ h <-> f ≅ h ○ bwarrinv g. +Proof. + split; intros H; [rewrite <- H | rewrite H]; + rewrite bwarr_assoc, ?bwarrinv_linv, ?bwarrinv_rinv, bwarr_runit; + easy. +Qed. + +Lemma bwarr_compose_r' {a b c} (f : a ⟶ b) (g : b ⟶ c) (h : a ⟶ c) : + h ≅ f ○ g <-> h ○ bwarrinv g ≅ f. +Proof. + split; symmetry; apply bwarr_compose_r; symmetry; assumption. +Qed. + +Lemma bwarr_compose_cancel_l {a b c} (f : a ⟶ b) (g h : b ⟶ c) : + f ○ g ≅ f ○ h -> g ≅ h. +Proof. + intros H. + rewrite <- (bwarr_lunit g), <- (bwarr_lunit h), + <- (bwarrinv_linv f), bwarr_assoc, H. + eauto with bwarrdb. +Qed. + +Lemma bwarr_compose_cancel_r {a b c} (f g : a ⟶ b) (h : b ⟶ c) : + f ○ h ≅ g ○ h -> f ≅ g. +Proof. + intros H. + rewrite <- (bwarr_runit f), <- (bwarr_runit g), <- (bwarrinv_rinv h), + <- bwarr_assoc, H. + eauto with bwarrdb. +Qed. + +Lemma bwarr_compose_cancel_l_iff {a b c} (f : a ⟶ b) (g h : b ⟶ c) : + f ○ g ≅ f ○ h <-> g ≅ h. +Proof. + split; [apply bwarr_compose_cancel_l|now intros ->]. +Qed. + +Lemma bwarr_compose_cancel_r_iff {a b c} (f g : a ⟶ b) (h : b ⟶ c) : + f ○ h ≅ g ○ h <-> f ≅ g. +Proof. + split; [apply bwarr_compose_cancel_r|now intros ->]. +Qed. + +Lemma bwarr_tensor_cancel_e_top {a b} (f g : a ⟶ b) (h : e ⟶ e) : + h ⊠ f ≅ h ⊠ g -> f ≅ g. +Proof. + intros H. + apply bwinv_unique with (f^-); + [now rewrite bwarrinv_linv|]. + rewrite <- bwarr_lunit, bwarr_compose_r, <- bwarr_tens_id in H. + simpl in H. + rewrite <- bwarr_tens_comp, bwarrinv_rinv in H. + rewrite <- (bwarr_compose_cancel_r_iff _ _ (arrlunitor _)) in H. + rewrite <- 2!bwarr_lunitor_nat in H. + rewrite bwarr_compose_cancel_l_iff in H. + easy. +Qed. + +Lemma bwarr_tensor_cancel_e_bot {a b} (f g : a ⟶ b) (h : e ⟶ e) : + f ⊠ h ≅ g ⊠ h -> f ≅ g. +Proof. + intros H. + apply bwinv_unique with (f^-); + [now rewrite bwarrinv_linv|]. + rewrite <- bwarr_lunit, bwarr_compose_r, <- bwarr_tens_id in H. + simpl in H. + rewrite <- bwarr_tens_comp, bwarrinv_rinv in H. + rewrite <- (bwarr_compose_cancel_r_iff _ _ (arrrunitor _)) in H. + rewrite <- 2!bwarr_runitor_nat in H. + rewrite bwarr_compose_cancel_l_iff in H. + easy. +Qed. + +Lemma bwarr_tensor_cancel_e_top_iff {a b} (f g : a ⟶ b) (h : e ⟶ e) : + h ⊠ f ≅ h ⊠ g <-> f ≅ g. +Proof. + split; [apply bwarr_tensor_cancel_e_top|now intros ->]. +Qed. + +Lemma bwarr_tensor_cancel_e_bot_iff {a b} (f g : a ⟶ b) (h : e ⟶ e) : + f ⊠ h ≅ g ⊠ h <-> f ≅ g. +Proof. + split; [apply bwarr_tensor_cancel_e_bot|now intros ->]. +Qed. + + +Lemma bwarr_assoc_nat_alt {a b c a' b' c' : bw} + (f : a ⟶ a') (g : b ⟶ b') (h : c ⟶ c') : + arrassoc a b c ≅ f ⊠ (g ⊠ h) ○ arrassoc a' b' c' ○ ((f ⊠ g) ⊠ h)^-. +Proof. + rewrite bwarr_compose_r'. + simpl. + rewrite 3!bwarrinv_invol. + apply bwarr_assoc_nat. +Qed. + +Lemma bwarr_assoc_nat_alt' {a b c a' b' c' : bw} + (f : a ⟶ a') (g : b ⟶ b') (h : c ⟶ c') : + arrassoc a b c ≅ f ⊠ (g ⊠ h) ○ arrassoc a' b' c' ○ (f^- ⊠ g^-) ⊠ h^-. +Proof. + rewrite bwarr_compose_r'. + simpl. + rewrite 3!bwarrinv_invol. + apply bwarr_assoc_nat. +Qed. + +Lemma bwarr_invassoc_nat_alt {a b c a' b' c' : bw} + (f : a ⟶ a') (g : b ⟶ b') (h : c ⟶ c') : + arrinvassoc a b c ≅ (f ⊠ g) ⊠ h ○ arrinvassoc a' b' c' ○ (f ⊠ (g ⊠ h))^-. +Proof. + apply by_bwarrinv. + simpl. + rewrite !bwarrinv_invol, <- bwarr_assoc. + apply bwarr_assoc_nat_alt. +Qed. + +Lemma bwarr_invassoc_nat_alt' {a b c a' b' c' : bw} + (f : a ⟶ a') (g : b ⟶ b') (h : c ⟶ c') : + arrinvassoc a b c ≅ (f ⊠ g) ⊠ h ○ arrinvassoc a' b' c' ○ f^- ⊠ (g^- ⊠ h^-). +Proof. + apply by_bwarrinv. + simpl. + rewrite !bwarrinv_invol, <- bwarr_assoc. + apply bwarr_assoc_nat_alt. +Qed. + +Lemma bwarr_triangle_alt (a b : bw) : + arrassoc a e b ≅ + arrid a ⊠ arrlunitor b ○ arrinvrunitor a ⊠ arrid b. +Proof. + rewrite bwarr_compose_r'. + apply bwarr_triangle. +Qed. + +Lemma bwarr_invtriangle_alt (a b : bw) : + arrinvassoc a e b ≅ + arrrunitor a ⊠ arrid b ○ arrid a ⊠ arrinvlunitor b. +Proof. + apply by_bwarrinv. + simpl. + apply bwarr_triangle_alt. +Qed. + + +Lemma bwarr_lunitor_tri (b c : bw) : + arrassoc e b c ○ arrlunitor b ⊠ arrid c ≅ arrlunitor (b ⨂ c). +Proof. + pose proof (bwarr_pentagon e e b c) as p. + rewrite (bwarr_triangle_alt e b) in p. + rewrite (bwarr_triangle_alt e (b ⨂ c)) in p. + rewrite (bwarr_assoc_nat_alt (arrrunitor e) (arrid b) (arrid c)) in p. + rewrite (bwarr_assoc_nat_alt (arrid e) (arrlunitor b) (arrid c)) in p. + rewrite !arrtens_pushout_top, <- !bwarr_assoc, bwarr_compose_cancel_r_iff in p. + rewrite !bwarr_assoc, bwarrinv_linv, bwarr_runit in p. + rewrite <- !bwarr_assoc, bwarr_compose_cancel_r_iff in p. + rewrite bwarr_assoc, <- bwarr_tens_comp, bwarr_runitor_linv in p. + rewrite bwarr_lunit, 2!bwarr_tens_id, bwarr_runit in p. + rewrite <- bwarr_tens_comp, bwarr_lunit in p. + rewrite bwarr_tensor_cancel_e_top_iff in p. + symmetry. + exact p. +Qed. + + +Lemma bwarr_runitor_tri (b c : bw) : + arrid b ⊠ arrrunitor c ≅ arrassoc b c e ○ arrrunitor (b ⨂ c). +Proof. + pose proof (bwarr_invpentagonl b c e e) as p. + rewrite (bwarr_invtriangle_alt c e) in p. + rewrite (bwarr_invtriangle_alt (b ⨂ c) e) in p. + rewrite (bwarr_invassoc_nat_alt (arrid b) (arrid c) (arrlunitor e)) in p. (* sure about runitor? *) + rewrite (bwarr_invassoc_nat_alt (arrid b) (arrrunitor c) (arrid e)) in p. (* runitor?? *) + simpl in p. + rewrite !arrtens_pushout_bot, <- !bwarr_assoc, bwarr_compose_cancel_r_iff in p. + rewrite !bwarr_assoc, <- 2!bwarr_tens_comp, bwarr_runitor_linv in p. + rewrite !bwarr_lunit, !bwarr_tens_id, bwarr_runit, <- !bwarr_assoc in p. + rewrite bwarr_compose_cancel_r_iff in p. + rewrite bwarr_assoc, <- bwarr_tens_comp, bwarr_lunitor_linv in p. + rewrite bwarr_lunit, bwarr_tens_id, bwarr_runit in p. + rewrite <- bwarr_tens_comp, bwarr_lunit in p. + rewrite bwarr_tensor_cancel_e_bot_iff in p. + rewrite p, <- bwarr_assoc, bwarr_assoc_rinv, bwarr_lunit. + easy. +Qed. + + +Section bwcast. + +Import UIP_facts MCClasses. + +Context {UIPX : UIP X}. + +Lemma bw_cast_id {n m} (Hn : n = n) (Hm : m = m) f : + cast_bwarr Hn Hm f = f. +Proof. + unfold cast_bwarr. + rewrite 2!eq_rect_eq. + easy. +Qed. + +Lemma cast_bwarr_indep {n n' m m'} + (Hn Hn' : n = n') (Hm Hm' : m = m') (f : n ⟶ m) : + cast_bwarr Hn Hm f = cast_bwarr Hn' Hm' f. +Proof. + f_equal; apply uip. +Qed. + +#[global] Add Parametric Morphism {n n' m m'} : + (@cast_bwarr n n' m m') with signature + true_rel ==> true_rel ==> bwarrequiv n m ==> bwarrequiv n' m' + as cast_bwarr_mor. +Proof. + intros; subst; rewrite 2!bw_cast_id. + easy. +Qed. + +Lemma bw_compose_cast_r {n n' m m' m'' p p'} + (Hn : n = n') (Hm : m = m') + (Hm' : m'' = m') (Hp : p = p') f g : + cast_bwarr Hn Hm f ○ cast_bwarr Hm' Hp g + ≅ cast_bwarr Hn eq_refl + (f ○ cast_bwarr (eq_trans Hm' (eq_sym Hm)) Hp g). +Proof. + subst. + easy. +Qed. + +Lemma bw_cast_cast {n n' n'' m m' m''} (f : n ⟶ m) Hn Hm + (Hn' : n' = n'') (Hm' : m' = m'') : + cast_bwarr Hn' Hm' (cast_bwarr Hn Hm f) = + cast_bwarr (eq_trans Hn Hn') (eq_trans Hm Hm') f. +Proof. + subst; easy. +Qed. + +Lemma bw_compose_cast_l {n n' m m' m'' p p'} + (Hn : n = n') (Hm : m = m') + (Hm' : m'' = m') (Hp : p = p') f g : + cast_bwarr Hn Hm f ○ cast_bwarr Hm' Hp g + ≅ cast_bwarr eq_refl Hp + (cast_bwarr Hn (eq_trans Hm (eq_sym Hm')) f ○ g). +Proof. + subst. + easy. +Qed. + +Lemma bw_cast_compose_split {n n' m p p'} (f : n ⟶ m) (g : m ⟶ p) + (Hn : n = n') (Hp : p = p') : + cast_bwarr Hn Hp (f ○ g) = + cast_bwarr Hn eq_refl f ○ cast_bwarr eq_refl Hp g. +Proof. + intros; subst; reflexivity. +Qed. + +Lemma bw_cast_equiv_iff {n n' m m'} (f : n ⟶ m) (g : n' ⟶ m') Hn Hm : + cast_bwarr Hn Hm f ≅ g <-> f ≅ cast_bwarr (eq_sym Hn) (eq_sym Hm) g. +Proof. + subst. + easy. +Qed. + +Lemma bw_cast_arrid {n n'} H H' : + cast_bwarr H H' (arrid n) = arrid n'. +Proof. + now subst; rewrite bw_cast_id. +Qed. + +Lemma bw_cast_tens_top {n n' m m' p q} (Hn : n = n') (Hm : m = m') + f (g : p ⟶ q) : + cast_bwarr Hn Hm f ⊠ g = + cast_bwarr + (f_equal (fun a => a ⨂ p) Hn) + (f_equal (fun a => a ⨂ q) Hm) + (f ⊠ g). +Proof. + subst; easy. +Qed. + +Lemma bw_cast_comp_l {n n' m p q} (Hn : n = n') (Hm : m = p) + f (g : p ⟶ q) : + cast_bwarr Hn Hm f ○ g = + cast_bwarr Hn eq_refl + (f ○ cast_bwarr (eq_sym Hm) eq_refl g). +Proof. + subst; easy. +Qed. + +Lemma bw_cast_comp_r {n m p q q'} (Hp : p = m) (Hq : q = q') + (f : n ⟶ m) g : + f ○ cast_bwarr Hp Hq g = + cast_bwarr eq_refl Hq + (cast_bwarr eq_refl (eq_sym Hp) f ○ g). +Proof. + subst; easy. +Qed. + +End bwcast. + +End bwarr_theory. + +Hint Resolve bwarrinv_linv bwarrinv_rinv : bwarrdb. +Hint Resolve bwarr_invassoc_natl bwarr_invlunitor_natl + bwarr_invrunitor_natl bwarr_invpentagonl bwarr_invtrianglel' : bwarrdb. +Hint Rewrite @bwarr_invassoc_natl + @bwarr_invlunitor_natl @bwarr_invrunitor_natl : bwarrdb. +(* Hint Rewrite <- @bwarr_invassoc_natl + @bwarr_invlunitor_natl @bwarr_invrunitor_natl : bwarrdb_rev. *) + + +Section bw_cat. + +Import CategoryTypeclass. + +#[export] Instance bwcat : Category bw | 10 := {| + morphism := bwarr; + c_equiv := bwarrequiv; + compose := fun _ _ _ => arrcomp; + c_identity := arrid; +|}. + +Obligation Tactic := Tactics.program_simpl; simpl; eauto 3 with bwarrdb. + +#[export, program] Instance bwcath : CategoryCoherence bwcat. +Next Obligation. +split; apply bwarrequiv_setoid. +Qed. +Solve All Obligations. + +#[export, program] Instance bwassoc_iso (a b c : bw) + : Isomorphism (a ⨂ b ⨂ c) (a ⨂ (b ⨂ c)) := { + forward := arrinvassoc a b c; + reverse := arrassoc a b c; +}. + +#[export, program] Instance bwlunitor_iso (a : bw) + : Isomorphism (e ⨂ a) a := { + forward := arrlunitor a; + reverse := arrinvlunitor a; +}. + +#[export, program] Instance bwrunitor_iso (a : bw) + : Isomorphism (a ⨂ e) a := { + forward := arrrunitor a; + reverse := arrinvrunitor a; +}. + +#[export] Instance bwmcat : MonoidalCategory bwcat | 10 := { + obj_tensor := tens; + mor_tensor := @arrtens X; + associator := bwassoc_iso; + left_unitor := bwlunitor_iso; + right_unitor := bwrunitor_iso; +}. + +#[export, program] Instance bwmcath : MonoidalCategoryCoherence bwmcat := {}. + + +(* #[export, program] Instance bwgroupoid : IsGroupoid bwcat := { + groupoid_inv := @bwarrinv +}. *) + +End bw_cat. + + +Section bw_thin. + +(* Note: We only need EQ_RECT_EQ bw and EQ_RECT_EQ bwnorm, but either + of these is equivalent to EQ_RECT_EQ X and thus UIP X, so we just + ask for the latter for convenience. It *may* be possible to use + only EQ_REC_EQ bw for the case of X a set, but this seems to + require a full rewrite to take advantage of (bw lies in Type, + not Set, at present). Moreover, I am not aware of any types for + which we know EQ_REC_EQ but not EQ_RECT_EQ to motivate doing + this work. (I have not even found examples of types for which we + have UIP but not decidable equality.) *) + +Import CategoryTypeclass MCClasses + UIP_facts MCTheory. +Import CatCategory (FunctorCategory). + +Context {UIPX : UIP X}. + +(* This is intended to make typeclass resolution faster; + I do not know if it does. *) +#[local] Instance Eq_rect_eq_bw : EQ_RECT_EQ bw. +typeclasses eauto. +Qed. + +#[local] Instance Eq_rect_eq_bwnorm : EQ_RECT_EQ bwnorm. +typeclasses eauto. +Qed. + +Definition Eq_rect_bw := + fun x P a h => eq_sym (Eq_rect_eq_bw.(eq_rect_eq) x P a h). + +Definition Eq_rect_bwnorm := + fun x P a h => eq_sym (Eq_rect_eq_bwnorm.(eq_rect_eq) x P a h). + +#[global] Add Parametric Morphism {n n' m m'} Hn Hm : + (@cast_bwarr n n' m m' Hn Hm) with signature + bwarrequiv n m ==> bwarrequiv n' m' + as cast_bwarr_mor'. +Proof. + intros; subst. + easy. +Qed. + + +Local Notation "'𝒩'" := (DiscreteCategory bwnorm). +Local Notation "'𝒩h'" := (DiscreteCategoryCoherence bwnorm). + +Local Notation "a '⟶' b" := (@bwarr X a b) + (at level 60) : type_scope. (* \longrightarrow *) + +Obligation Tactic := idtac. + +#[export, program] Instance norm_bw_bifunc : + Bifunctor 𝒩 bwcat bwcat := { + obj_bimap := fun n a => n ⨂ a; + morphism_bimap := fun n n' a b neq f => + cast_bwarr eq_refl _ (arrtens (arrid n) f) +}. +Next Obligation. + simpl. + intros; subst. + reflexivity. +Defined. +Next Obligation. + simpl. + intros. + apply bwarr_tens_id. +Qed. +Next Obligation. + simpl. + intros. + subst. + simpl. + unfold eq_ind_r. + simpl. + apply arrtens_pushout_bot. +Qed. +Next Obligation. + simpl. + intros. + subst. + rewrite !bw_cast_id. + eauto 2 with bwarrdb. +Qed. + +Fixpoint bwbrac_eq_of_arr {a b} (f : a ⟶ b) {struct f} : forall n, ⟦a⟧ n = ⟦b⟧ n. + induction f; intros n. + all: try reflexivity. + - etransitivity; [apply IHf1 | apply IHf2]. + - simpl. + etransitivity; [| apply IHf2]. + apply f_equal. + apply IHf1. +Defined. + +Definition Nf_eq_of_arr {a b : bw} (f : bwarr a b) : Nf a = Nf b := + bwbrac_eq_of_arr f norm_e. + +Arguments Nf_eq_of_arr {_ _} !f /. + +Definition bwarr_tens_cancel_l {a b c : bw} + (f : a ⨂ b ⟶ a ⨂ c) : b ⟶ c := + bwarr_of_Nf_eq (Nf_cancel_tens_l (Nf_eq_of_arr f)). + +Definition bwarr_tens_cancel_r {a b c : bw} + (f : b ⨂ a ⟶ c ⨂ a) : b ⟶ c := + bwarr_of_Nf_eq (Nf_cancel_tens_r (Nf_eq_of_arr f)). + +Obligation Tactic := Tactics.program_simpl; simpl; eauto 3 with bwarrdb. + +#[export, program] Instance Nf_functor : Functor bwcat 𝒩 := { + obj_map := Nf; + morphism_map := fun a b f => (bwbrac_of_bweq a b (bweq_of_arr f) norm_e) +}. + +#[export, program] Instance bwbrac_functor : + Functor bwcat (@FunctorCategory _ _ 𝒩 𝒩h 𝒩 𝒩h) := { + obj_map := fun a => {|obj_map := bwbrac a|}; + morphism_map := fun a b f => + {| component_map := fun c => _ |}; + (* morphism_map := fun a b f => (bwbrac_of_bweq a b (bweq_of_arr f) norm_e) *) +}. +Next Obligation. + apply bwbrac_eq_of_arr, f. +Defined. +(* FIXME: Polymorphism issue. *) +Admit Obligations. +(* Solve All Obligations with constructor. *) + +Definition bwbrac_mor_bimap_pf (n m : bwnorm) (a b : bw) + (H : 𝒩.(morphism) n m) (f : a ⟶ b) : ⟦a⟧ n = ⟦b⟧ m :> bw. +Proof. + rewrite <- (bwbrac_eq_of_arr f m). + f_equal. + f_equal. + apply H. +Defined. + + +Definition bwbrac_mor_bimap (n m : bwnorm) (a b : bw) + (H : 𝒩.(morphism) n m) (f : a ⟶ b) : ⟦a⟧ n ⟶ ⟦b⟧ m := + cast_bwarr eq_refl (bwbrac_mor_bimap_pf n m a b H f) (arrid _). +Arguments bwbrac_mor_bimap _ _ _ _ _ / _. +Arguments eq_rect_r [_] [_] _ _ [_] / _. + +Obligation Tactic := simpl; intros; eauto 3 with bwarrdb. + + +#[export, program] Instance bwbrac_bifunctor : + Bifunctor 𝒩 bwcat bwcat := { + obj_bimap := fun n a => ⟦a⟧ n; + morphism_bimap := bwbrac_mor_bimap; +}. +Next Obligation. + rewrite bw_compose_cast_r. + simpl. + rewrite bwarr_lunit. + rewrite bw_cast_equiv_iff, !bw_cast_cast. + rewrite bw_cast_arrid. + easy. +Qed. +Next Obligation. + apply cast_bwarr_mor; easy. +Qed. + +Obligation Tactic := simpl; eauto 3 with bwarrdb. + + + +#[export, program] Instance bwbinat_trans : + NaturalBiIsomorphism norm_bw_bifunc bwbrac_bifunctor := { + component_biiso := fun n a => {| + forward := xi_comp_map n a; + reverse := (xi_comp_map n a) ^-; + |}; + component_biiso_natural := fun n m a b hnm f => _ +}. +Next Obligation. + Local Ltac gen_casts := + repeat match goal with + |- context[cast_bwarr ?pf1 ?pf2 _] => + assert_fails (is_evar pf1); generalize pf1 pf2 + end. + intros. + subst. + simpl. + revert m; + induction f; intros m. + all : rewrite ?bw_cast_arrid, ?bwarr_tens_id, + ?bwarr_runit, ?bwarr_lunit; simpl. + - rewrite arrtens_pushout_bot, bwarr_assoc, IHf2, + <- bwarr_assoc, IHf1, bwarr_assoc. + apply bwarr_comp; [easy|]. + rewrite bw_compose_cast_r, bwarr_lunit. + simpl. + rewrite bw_cast_equiv_iff, bw_cast_cast, bw_cast_arrid. + easy. + - simpl. + rewrite <- !bwarr_assoc. + rewrite <- bwarr_assoc_nat, (bwarr_assoc (arrassoc _ _ _)). + rewrite <- bwarr_tens_comp, IHf1, bwarr_runit. + rewrite arrtens_pushin_top, !bwarr_assoc. + repeat apply bwarr_compose_cancel_l_iff. + rewrite bw_cast_tens_top. + rewrite bw_cast_comp_l. + simpl. + specialize (IHf2 (⟦a⟧ m)). + rewrite bwarr_compose_r in IHf2. + rewrite IHf2. + rewrite !bwarr_assoc. + apply bwarr_compose_cancel_l_iff. + symmetry. + rewrite bw_cast_equiv_iff. + rewrite !bw_cast_comp_r, !bw_cast_comp_l, bwarr_lunit. + simpl. + rewrite !bw_cast_cast. + pose proof (bw_of_bwnorm_inj _ _ + (bwbrac_mor_bimap_pf m m a a' eq_refl f1)) as H. + gen_casts. + clear f1 f2 IHf1 IHf2. + revert H. + generalize (⟦a⟧ m). + generalize (⟦a'⟧ m). + intros; subst. + now rewrite !bw_cast_id, bwarrinv_linv, bw_cast_arrid. + - easy. + - rewrite !arrtens_pushin_top. + rewrite <- !bwarr_assoc. + repeat apply bwarr_compose_cancel_r_iff. + rewrite bwarr_compose_r. simpl. + rewrite !bwarr_assoc, bwarr_assoc_nat. + rewrite <- 2!(bwarr_assoc (_ ⊠ _)), bwarr_tens_id. + rewrite bwarrinv_rinv, bwarr_lunit. + now rewrite bwarr_pentagon. + - rewrite !arrtens_pushin_top. + rewrite <- !bwarr_assoc. + repeat apply bwarr_compose_cancel_r_iff. + rewrite bwarr_compose_r. simpl. + rewrite !bwarr_assoc, bwarr_invassoc_natl. + rewrite <- !bwarr_assoc, bwarr_tens_id. + repeat apply bwarr_compose_cancel_r_iff. + rewrite bwarr_compose_r', bwarr_assoc, bwarr_compose_l. + now rewrite bwarr_pentagon, bwarr_assoc. + - now rewrite bwarr_triangle. + - rewrite bwarr_triangle. + rewrite <- bwarr_assoc. + rewrite <- bwarr_tens_comp. + rewrite (bwarrinv_linv (arrlunitor a)). + now rewrite bwarr_lunit, bwarr_tens_id, bwarr_lunit. + - rewrite bwarr_assoc, <- bwarr_runitor_nat, <- bwarr_assoc. + now rewrite <- bwarr_runitor_tri. + - rewrite bwarr_assoc, bwarr_compose_l, + <- bwarr_runitor_nat, <- bwarr_assoc. + now rewrite <- bwarr_runitor_tri. +Qed. + + +#[export, program] Instance Nf_bwcat_functor : Functor bwcat bwcat := { + obj_map := Nf; + morphism_map := fun a b f => + cast_bwarr eq_refl (f_equal _ (bwbrac_eq_of_arr f norm_e)) (arrid _) +}. +Next Obligation. + intros. + rewrite bw_cast_equiv_iff, bw_cast_comp_l, bwarr_lunit, !bw_cast_cast. + now rewrite bw_cast_arrid. +Qed. +Next Obligation. + intros. + apply cast_bwarr_mor; easy. +Qed. + +#[export, program] Instance toNf_natiso : + NaturalIsomorphism (CatCategory.IdentityFunctor bwcat) Nf_bwcat_functor := { + component_iso := fun a => + CatCategory.ComposeIsomorphisms + {| forward := arrinvlunitor a : bwcat.(morphism) _ _; reverse := arrlunitor a |} + (bwbinat_trans norm_e a) +}. +Next Obligation. + intros. + (* rewrite bw_cast_comp_r, bwarr_runit. *) + simpl. + rewrite <- bwarr_assoc. + rewrite bwarr_invlunitor_natl, 2!bwarr_assoc. + rewrite bwarr_compose_cancel_l_iff. + epose proof (bwbinat_trans.(component_biiso_natural) norm_e norm_e + A B eq_refl f) as en. + simpl in en. + rewrite en. + erewrite cast_bwarr_indep. + reflexivity. +Qed. + +Theorem bw_thin {a b : bw} (f g : a ⟶ b) : f ≅ g. +Proof. + pose proof ((toNf_natiso).(component_iso_natural) f) as Hf. + rewrite compose_iso_r in Hf. + pose proof ((toNf_natiso).(component_iso_natural) g) as Hg. + rewrite compose_iso_r in Hg. + simpl in *. + rewrite Hf, Hg. + rewrite bwarr_compose_cancel_r_iff, + bwarr_compose_cancel_l_iff. + erewrite cast_bwarr_indep; reflexivity. +Qed. + +End bw_thin. + +End bw_theory. \ No newline at end of file diff --git a/ViCaR/MonoidalCoherence/MCconsequences.v b/ViCaR/MonoidalCoherence/MCconsequences.v new file mode 100644 index 0000000..909dc64 --- /dev/null +++ b/ViCaR/MonoidalCoherence/MCconsequences.v @@ -0,0 +1,608 @@ +Require Import Setoid. +Require MCDefinitions UIP_facts MCbw. + + + +Section MonoidalCoherenceConsequences. + +Set Universe Polymorphism. + +Import CategoryTypeclass MCDefinitions + MC_notations MCClasses MC_setoids UIP_facts MCbw. + +Context {X : Type} {UIPX : UIP X}. +Context {cC : Category X} {cCh : CategoryCoherence cC} + {mC : MonoidalCategory cC} {mCh : MonoidalCategoryCoherence mC}. + +Local Notation bw := (bw X). +Local Notation bwnorm := (bwnorm X). + +Section bwconsequences. + +Import bwarr_notation. + +Open Scope bw_scope. + +Lemma bwarr_tens_cancel_l_equiv {a b c : bw} (f : a ⨂ b ⟶ a ⨂ c) : + f ≅ arrid a ⊠ bwarr_tens_cancel_l f. +Proof. apply bw_thin. Qed. + +Lemma bwarr_tens_cancel_r_equiv {a b c : bw} (f : b ⨂ a ⟶ c ⨂ a) : + f ≅ bwarr_tens_cancel_r f ⊠ arrid a. +Proof. apply bw_thin. Qed. + +End bwconsequences. + +Section MonoidalCoherence. + +Local Open Scope Cat_scope. + +Local Obligation Tactic := Tactics.program_simpl; simpl; eauto 3 with bwarrdb; try easy. + +Local Definition realize_bw' := (@realize_bw X cC mC). + +(* Local Coercion realize_bw' : bw >-> X. *) + +Local Notation "a '⟶' b" := (@bwarr X a b) + (at level 60) : type_scope. (* \longrightarrow *) + + +#[export, program] Instance RealizationFunctor : + Functor bwcat cC := { + obj_map := realize_bw; + morphism_map := @realize_bwarr X cC mC; +}. +Next Obligation. + induction H; simpl; + rewrite ?iso_inv_l, ?iso_inv_r; + try ((idtac + symmetry); solve [eauto using assoc, compose_compat, + left_unit, right_unit, tensor_id, tensor_compose, tensor_compat, + associator_cohere, left_unitor_cohere, right_unitor_cohere, + (equiv_refl _ _ c_equiv_rel); try easy]). + - rewrite <- compose_iso_l', <- assoc, <- compose_iso_r. + symmetry. + apply associator_cohere. + - rewrite <- left_unit, <- assoc. + rewrite <- 2!compose_iso_r'. + rewrite !assoc. + rewrite <- pentagon. + rewrite <- 2!(assoc (α_ realize_bw a, _, _ ⁻¹ ⊗ id_ realize_bw d)). + rewrite <- tensor_compose, iso_inv_l, right_unit, tensor_id, left_unit. + rewrite <- (assoc _ (α_ _, _ ,_)). + rewrite iso_inv_l, left_unit. + now rewrite <- tensor_compose, iso_inv_l, left_unit, tensor_id. + - rewrite <- triangle. + now rewrite <- assoc, iso_inv_l, left_unit. + - etransitivity; eassumption. +Qed. + + +Theorem monoidal_coherence {a b : bw} (f g : a ⟶ b) : + realize_bwarr f ≃ realize_bwarr g. +Proof. + apply RealizationFunctor.(morphism_compat). + apply bw_thin. +Qed. + +End MonoidalCoherence. + +Section monarr_theory. + +Section monarr_cat. + +Local Open Scope bw_scope. +Import monarr_notation. + +Definition monbwcat : Category bw := {| + morphism := monarr; + c_equiv := monarrequiv; + compose := fun _ _ _ => monarrcomp; + c_identity := arrid; +|}. + +#[local] Existing Instance monbwcat | 9. + +#[export, program] Instance monbwcath : CategoryCoherence monbwcat. +Next Obligation. +split; apply monarrequiv_setoid. +Qed. +Obligation Tactic := + Tactics.program_simpl; eauto 4 with monarrdb bwarrdb; try easy. +Solve All Obligations. + +#[export, program] Instance monassoc_iso (a b c : bw) + : Isomorphism (a ⨂ b ⨂ c) (a ⨂ (b ⨂ c)) := { + forward := monarrstruct (arrinvassoc a b c); + reverse := arrassoc a b c; +}. + +#[export, program] Instance monlunitor_iso (a : bw) + : Isomorphism (e ⨂ a) a := { + forward := monarrstruct (arrlunitor a); + reverse := arrinvlunitor a; +}. + +#[export, program] Instance monrunitor_iso (a : bw) + : Isomorphism (a ⨂ e) a := { + forward := arrrunitor a; + reverse := arrinvrunitor a; +}. + +#[export] Instance monbwmcat : MonoidalCategory monbwcat | 10 := { + obj_tensor := tens; + mor_tensor := @monarrtens X cC mC; + associator := monassoc_iso; + left_unitor := monlunitor_iso; + right_unitor := monrunitor_iso; +}. + +#[export, program] Instance monbwmcath : + MonoidalCategoryCoherence monbwmcat := {}. +Next Obligation. + apply (compose_iso_l (monassoc_iso _ _ _)). + simpl. + rewrite monarr_assoc. + rewrite monarr_assoc_nat. + rewrite <- monarr_assoc, monarr_arrcomp, (monarr_struct _ (arrid _)). + now rewrite monarr_runit. +Qed. +Next Obligation. + rewrite !monarr_arrtens, monarr_arrcomp. + apply monarr_struct. +Qed. +Next Obligation. + rewrite !monarr_arrtens, !monarr_arrcomp. + apply monarr_struct. +Qed. + +#[export, program] Instance GeneralRealizationFunctor : + Functor monbwcat cC := { + obj_map := realize_bw; + morphism_map := @realize_monarr X cC mC; +}. +Next Obligation. + induction H; try reflexivity; simpl. + - apply compose_compat; auto. + - symmetry; apply assoc. + - apply left_unit. + - apply right_unit. + - apply tensor_compat; easy. + - apply tensor_compose. + - apply monoidal_coherence. + - rewrite <- compose_iso_l', <- assoc, associator_cohere. + rewrite assoc, iso_inv_r, right_unit. + easy. + - apply left_unitor_cohere. + - apply right_unitor_cohere. + - symmetry; easy. + - etransitivity; eauto. +Qed. + +#[global] Add Parametric Morphism {a b} : (@realize_monarr X cC mC a b) + with signature + monarrequiv a b ==> cC.(c_equiv) + as realize_monarr_mor. +Proof. + apply GeneralRealizationFunctor.(morphism_compat). +Qed. + +Section monarr_lemmas. + +Section tens_comp_split. + +Context {a b c m n o : bw} (f : a ⟶ b) (g : b ⟶ c) + (h : m ⟶ n) (j : n ⟶ o). + +Lemma monarr_tens_comp_split_diag : + f ⧆ h ≊ (f ⧆ arrid m) ◌ (arrid b ⧆ h). +Proof. + now rewrite <- monarr_tens_comp, monarr_lunit, monarr_runit. +Qed. + +Lemma monarr_tens_comp_split_antidiag : + f ⧆ h ≊ (arrid a ⧆ h) ◌ (f ⧆ arrid n). +Proof. + now rewrite <- monarr_tens_comp, monarr_lunit, monarr_runit. +Qed. + + +Lemma monarr_tens_comp_split_bot_l : + f ⧆ (h ◌ j) ≊ + (arrid a ⧆ h) ◌ (f ⧆ j). +Proof. + now rewrite <- monarr_tens_comp, monarr_lunit. +Qed. + +Lemma monarr_tens_comp_split_bot_r : + f ⧆ (h ◌ j) ≊ + (f ⧆ h) ◌ (arrid b ⧆ j). +Proof. + now rewrite <- monarr_tens_comp, monarr_runit. +Qed. + +Lemma monarr_tens_comp_split_top_l : + (f ◌ g) ⧆ h ≊ + (f ⧆ arrid m) ◌ (g ⧆ h). +Proof. + now rewrite <- monarr_tens_comp, monarr_lunit. +Qed. + +Lemma monarr_tens_comp_split_top_r : + (f ◌ g) ⧆ h ≊ + (f ⧆ h) ◌ (g ⧆ arrid n). +Proof. + now rewrite <- monarr_tens_comp, monarr_runit. +Qed. + +End tens_comp_split. + +Lemma monarr_struct_Nf_eq {a b} (f : bwarr a b) : + f ≊ bwarr_of_Nf_eq (Nf_eq_of_arr f). +Proof. + apply monarr_struct. +Qed. + +(* Section monarr_struct_builders. + +Local Notation "a → b" := (bwarr a b) (at level 70) : type_scope. + +Lemma __monarr_struct_assoc {a b c d} + (f : a ⨂ (b ⨂ c) → d) (g : ) *) + +Lemma monarr_struct_id {a} (f : bwarr a a) : + f ≊ arrid a. +Proof. apply monarr_struct. Qed. + +Lemma monarr_struct_assoc {a b c} (f : bwarr _ _) : + f ≊ arrassoc a b c. +Proof. apply monarr_struct. Qed. + +Lemma monarr_struct_invassoc {a b c} (f : bwarr _ _) : + f ≊ arrinvassoc a b c. +Proof. apply monarr_struct. Qed. + +Lemma monarr_struct_lunitor {a} (f : bwarr _ _) : + f ≊ arrlunitor a. +Proof. apply monarr_struct. Qed. + +Lemma monarr_struct_invlunitor {a} (f : bwarr _ _) : + f ≊ arrinvlunitor a. +Proof. apply monarr_struct. Qed. + +Lemma monarr_struct_runitor {a} (f : bwarr _ _) : + f ≊ arrrunitor a. +Proof. apply monarr_struct. Qed. + +Lemma monarr_struct_invrunitor {a} (f : bwarr _ _) : + f ≊ arrinvrunitor a. +Proof. apply monarr_struct. Qed. + + + +Lemma monarrcomp_struct_r {a b m} (f : a ⟶ b) + (g : bwarr b m) (h : a ⟶ m) : + monarrcomp f g ≊ h <-> f ≊ monarrcomp h (g ^-). +Proof. + split; + intros H; + [rewrite <- H | rewrite H]; rewrite <- monarr_assoc; + rewrite monarr_arrcomp, (monarr_struct _ (arrid _)); + now rewrite ?monarr_runit, ?monarr_lunit. +Qed. + +Lemma monarrcomp_struct_l {a b m} (f : bwarr a b) + (g : b ⟶ m) (h : a ⟶ m) : + monarrcomp f g ≊ h <-> g ≊ monarrcomp (f ^-) h. +Proof. + split; + intros H; + [rewrite <- H | rewrite H]; rewrite monarr_assoc; + rewrite monarr_arrcomp, (monarr_struct _ (arrid _)); + now rewrite ?monarr_runit, ?monarr_lunit. +Qed. + +Lemma monarrcomp_struct_r' {a b m} (f : a ⟶ b) + (g : bwarr b m) (h : a ⟶ m) : + h ≊ monarrcomp f g <-> monarrcomp h (g ^-) ≊ f. +Proof. + split; + intros H; + symmetry; + apply monarrcomp_struct_r; easy. +Qed. + +Lemma monarrcomp_struct_l' {a b m} (f : bwarr a b) + (g : b ⟶ m) (h : a ⟶ m) : + h ≊ monarrcomp f g <-> monarrcomp (f ^-) h ≊ g. +Proof. + split; + intros H; + symmetry; + apply monarrcomp_struct_l; easy. +Qed. + +Lemma monarr_struct_inv {a b} (f : bwarr a b) (g : bwarr b a) : + f ◌ g ≊ arrid a. +Proof. + eauto with monarrdb. +Qed. + +Lemma monarr_id_l {a b} (f : a ⟶ b) (g : bwarr a a) : + g ◌ f ≊ f. +Proof. now rewrite monarr_struct_id, monarr_lunit. Qed. + +Lemma monarr_id_r {a b} (f : a ⟶ b) (g : bwarr b b) : + f ◌ g ≊ f. +Proof. now rewrite monarr_struct_id, monarr_runit. Qed. + +Lemma monarr_assoc_nat_l {a b m n p q} (f : a ⟶ b) (g : m ⟶ n) + (h : p ⟶ q) (assoc' : bwarr (a ⨂ (m ⨂ p)) (a ⨂ m ⨂ p)) : + assoc' ◌ f ⧆ g ⧆ h + ≊ f ⧆ (g ⧆ h) ◌ arrassoc b n q. +Proof. + rewrite <- monarr_assoc_nat. + eauto with monarrdb. +Qed. + +Lemma monarr_assoc_nat_r {a b m n p q} (f : a ⟶ b) (g : m ⟶ n) + (h : p ⟶ q) (assoc' : bwarr (b ⨂ (n ⨂ q)) (b ⨂ n ⨂ q)) : + f ⧆ (g ⧆ h) ◌ assoc' + ≊ arrassoc a m p ◌ f ⧆ g ⧆ h. +Proof. + rewrite monarr_assoc_nat. + eauto with monarrdb. +Qed. + +Lemma monarr_invassoc_nat_l {a b m n p q} (f : a ⟶ b) (g : m ⟶ n) + (h : p ⟶ q) (assoc' : bwarr (a ⨂ m ⨂ p) (a ⨂ (m ⨂ p))) : + assoc' ◌ f ⧆ (g ⧆ h) + ≊ f ⧆ g ⧆ h ◌ arrinvassoc b n q. +Proof. + rewrite monarrcomp_struct_l, monarr_assoc, monarr_assoc_nat_l. + rewrite <- monarr_assoc, monarr_struct_inv. + eauto with monarrdb. +Qed. + +Lemma monarr_invassoc_nat_r {a b m n p q} (f : a ⟶ b) (g : m ⟶ n) + (h : p ⟶ q) (assoc' : bwarr (b ⨂ n ⨂ q) (b ⨂ (n ⨂ q))) : + f ⧆ g ⧆ h ◌ assoc' + ≊ arrinvassoc a m p ◌ f ⧆ (g ⧆ h). +Proof. + rewrite monarr_invassoc_nat_l. + eauto with monarrdb. +Qed. + +Lemma monarr_lunitor_nat_l {a b} (f : a ⟶ b) (unitor' : bwarr (e ⨂ a) a) : + unitor' ◌ f ≊ arrid e ⧆ f ◌ arrlunitor b. +Proof. + rewrite <- monarr_lunitor_nat. + eauto with monarrdb. +Qed. + +Lemma monarr_lunitor_nat_r {a b} (f : a ⟶ b) (unitor' : bwarr (e ⨂ b) b) : + arrid e ⧆ f ◌ unitor' ≊ arrlunitor a ◌ f. +Proof. + rewrite monarr_lunitor_nat. + eauto with monarrdb. +Qed. + +Lemma monarr_invlunitor_nat_l {a b} (f : a ⟶ b) (unitor' : bwarr _ _) : + unitor' ◌ arrid e ⧆ f ≊ f ◌ arrinvlunitor b. +Proof. + rewrite monarrcomp_struct_l, monarr_assoc, monarrcomp_struct_r'. + eauto with monarrdb. +Qed. + +Lemma monarr_invlunitor_nat_r {a b} (f : a ⟶ b) (unitor' : bwarr _ _) : + f ◌ unitor' ≊ arrinvlunitor a ◌ arrid e ⧆ f. +Proof. + rewrite monarrcomp_struct_l', monarr_assoc, monarrcomp_struct_r. + eauto with monarrdb. +Qed. + +Lemma monarr_runitor_nat_l {a b} (f : a ⟶ b) (unitor' : bwarr _ _) : + unitor' ◌ f ≊ f ⧆ arrid e ◌ arrrunitor b. +Proof. + rewrite <- monarr_runitor_nat. + eauto with monarrdb. +Qed. + +Lemma monarr_runitor_nat_r {a b} (f : a ⟶ b) (unitor' : bwarr _ _) : + f ⧆ arrid e ◌ unitor' ≊ arrrunitor a ◌ f. +Proof. + rewrite monarr_runitor_nat. + eauto with monarrdb. +Qed. + +Lemma monarr_invrunitor_nat_l {a b} (f : a ⟶ b) (unitor' : bwarr _ _) : + unitor' ◌ f ⧆ arrid e ≊ f ◌ arrinvrunitor b. +Proof. + rewrite monarrcomp_struct_l, monarr_assoc, monarrcomp_struct_r'. + eauto with monarrdb. +Qed. + +Lemma monarr_invrunitor_nat_r {a b} (f : a ⟶ b) (unitor' : bwarr _ _) : + f ◌ unitor' ≊ arrinvrunitor a ◌ f ⧆ arrid e. +Proof. + rewrite monarrcomp_struct_l', monarr_assoc, monarrcomp_struct_r. + eauto with monarrdb. +Qed. + +Lemma monarr_tens_id_split_bot {a b c d} + (f : a ⟶ b) (g : b ⟶ c) (arrid' : bwarr d d) : + monarrtens arrid' (monarrcomp f g) + ≊ monarrcomp (monarrtens arrid' f) (monarrtens (arrid d) g). +Proof. + now rewrite <- monarr_tens_comp, monarr_id_r. +Qed. + +Lemma monarr_tens_id_split_top {a b c d} + (f : a ⟶ b) (g : b ⟶ c) (arrid' : bwarr d d): + monarrtens (monarrcomp f g) arrid' + ≊ monarrcomp (monarrtens f arrid') (monarrtens g (arrid d)). +Proof. + now rewrite <- monarr_tens_comp, monarr_runit. +Qed. + +Lemma monarr_comp_tens_struct_r {a b c d n m} (f : a ⟶ b) (g : c ⟶ d) + (h : bwarr (b ⨂ d) (n ⨂ m)) (htop : bwarr b n) (hbot : bwarr d m) : + f ⧆ g ◌ h ≊ (f ◌ htop) ⧆ (g ◌ hbot). +Proof. + rewrite monarr_tens_comp, monarr_arrtens. + auto with monarrdb. +Qed. + +Lemma monarr_comp_tens_struct_l {a b c d n m} (f : a ⟶ b) (g : c ⟶ d) + (h : bwarr (n ⨂ m) (a ⨂ c)) (htop : bwarr n a) (hbot : bwarr m c) : + h ◌ f ⧆ g ≊ (htop ◌ f) ⧆ (hbot ◌ g). +Proof. + rewrite monarr_tens_comp, monarr_arrtens. + auto with monarrdb. +Qed. + +Lemma monarr_comp_tens_id_top_struct_r {a b c d n} + (f : a ⟶ b) (g : c ⟶ d) + (h : bwarr (b ⨂ d) (b ⨂ n)) : + f ⧆ g ◌ h ≊ f ⧆ (g ◌ bwarr_tens_cancel_l h). +Proof. + erewrite monarr_struct, <- monarr_arrtens. + rewrite <- monarr_tens_comp. + rewrite monarr_runit. + easy. +Qed. + +Lemma monarr_comp_tens_id_bot_struct_r {a b c d n} + (f : a ⟶ b) (g : c ⟶ d) + (h : bwarr (b ⨂ d) (n ⨂ d)) : + f ⧆ g ◌ h ≊ (f ◌ bwarr_tens_cancel_r h) ⧆ g. +Proof. + erewrite monarr_struct, <- monarr_arrtens. + rewrite <- monarr_tens_comp. + rewrite monarr_runit. + easy. +Qed. + +Lemma monarr_comp_tens_id_top_struct_l {a b c d n} + (f : a ⟶ b) (g : c ⟶ d) + (h : bwarr (a ⨂ n) (a ⨂ c)) : + h ◌ f ⧆ g ≊ f ⧆ (bwarr_tens_cancel_l h ◌ g). +Proof. + erewrite monarr_struct, <- monarr_arrtens. + rewrite <- monarr_tens_comp. + rewrite monarr_lunit. + easy. +Qed. + +Lemma monarr_comp_tens_id_bot_struct_l {a b c d n} + (f : a ⟶ b) (g : c ⟶ d) + (h : bwarr (n ⨂ c) (a ⨂ c)) : + h ◌ f ⧆ g ≊ (bwarr_tens_cancel_r h ◌ f) ⧆ g. +Proof. + erewrite monarr_struct, <- monarr_arrtens. + rewrite <- monarr_tens_comp. + rewrite monarr_lunit. + easy. +Qed. + +End monarr_lemmas. + +Section monarr_transformations. + + +Fixpoint monarr_rassoc_app {a b c} (f : a ⟶ b) : b ⟶ c -> a ⟶ c := + match f with + | monarrcomp h f' => fun g => monarrcomp h (monarr_rassoc_app f' g) + | f' => fun g => monarrcomp f' g + end. + +Fixpoint rassoc_monarr {a b} (f : a ⟶ b) : a ⟶ b := + match f with + | monarrcomp g h => monarr_rassoc_app (rassoc_monarr g) (rassoc_monarr h) + | monarrtens g h => monarrtens (rassoc_monarr g) (rassoc_monarr h) + | monarrstruct g => monarrstruct g + | mongeneric g => mongeneric g + end. + +Lemma monarr_rassoc_app_correct {a b c} (f : a ⟶ b) (g : b ⟶ c) : + monarr_rassoc_app f g ≊ monarrcomp f g. +Proof. + induction f; eauto with monarrdb. +Qed. + +Lemma rassoc_monarr_correct {a b} (f : a ⟶ b) : + rassoc_monarr f ≊ f. +Proof. + induction f; simpl; [rewrite monarr_rassoc_app_correct|..]; + eauto 2 with monarrdb. +Qed. + +Definition structify_acc_structs {a b} (f : a ⟶ b) : forall {c}, b ⟶ c -> a ⟶ c := + match f in a' ⟶ b' return forall {c}, b' ⟶ c -> a' ⟶ c with + | monarrstruct f' => fun c g => (match g in b' ⟶ c' return forall {a'}, bwarr a' b' -> a' ⟶ c' with + | monarrstruct g' => fun _ f' => monarrstruct (f' ○ g') + | monarrcomp g1 g2 => + (match g1 in b' ⟶ c' return forall {d'} (_ : c' ⟶ d') {a'}, bwarr a' b' -> a' ⟶ d' with + | monarrstruct h => fun d' g' a' f' => monarrcomp (f' ○ h) g' + | g1' => fun d' g' a' f' => monarrcomp f' (g1' ◌ g') + end) _ g2 + | monarrtens g1 g2 => fun a' f' => monarrcomp f' (g1 ⧆ g2) + | mongeneric g' => fun a' f' => monarrcomp f' (mongeneric g') + end) _ f' + | monarrtens f1 f2 => fun c g => monarrcomp (monarrtens f1 f2) g + | monarrcomp f1 f2 => fun c g => monarrcomp (monarrcomp f1 f2) g + | mongeneric f' => fun c g => monarrcomp (mongeneric f') g + end. + +Fixpoint structify_rassoc_monarr {a b} (f : a ⟶ b) : a ⟶ b := + match f with + | monarrcomp g1 g2 => + structify_acc_structs (structify_rassoc_monarr g1) (structify_rassoc_monarr g2) + | monarrtens g1 g2 => + match structify_rassoc_monarr g1, structify_rassoc_monarr g2 with + | monarrstruct h1, monarrstruct h2 => monarrstruct (arrtens h1 h2) + | h1, h2 => monarrtens h1 h2 + end + | monarrstruct g => monarrstruct g + | mongeneric g => mongeneric g + end. + +Lemma structify_acc_structs_correct {a b c} (f : a ⟶ b) (g : b ⟶ c) : + structify_acc_structs f g ≊ f ◌ g. +Proof. + induction f; [easy..|]. + destruct g; [|easy..|]. + - destruct g1; [easy..|]. + simpl. + rewrite <- monarr_arrcomp. + symmetry; apply monarr_assoc. + - symmetry; apply monarr_arrcomp. +Qed. + +Lemma structify_rassoc_monarr_correct {a b} (f : a ⟶ b) : + structify_rassoc_monarr f ≊ f. +Proof. + induction f; [| |easy..]. + - eauto using structify_acc_structs_correct with monarrdb. + - simpl. + revert IHf1 IHf2. + destruct (structify_rassoc_monarr f1); [eauto 2 with monarrdb..|]. + destruct (structify_rassoc_monarr f2); eauto with monarrdb. +Qed. + +Definition structify_monarr {a b} (f : a ⟶ b) : a ⟶ b := + structify_rassoc_monarr (rassoc_monarr f). + +Lemma structify_monarr_correct {a b} (f : a ⟶ b) : + structify_monarr f ≊ f. +Proof. + transitivity (rassoc_monarr f). + - apply structify_rassoc_monarr_correct. + - apply rassoc_monarr_correct. +Qed. + + +End monarr_transformations. + +End monarr_cat. +End monarr_theory. +End MonoidalCoherenceConsequences. \ No newline at end of file diff --git a/ViCaR/MonoidalCoherence/MCmonarrlist.v b/ViCaR/MonoidalCoherence/MCmonarrlist.v new file mode 100644 index 0000000..3ff50ec --- /dev/null +++ b/ViCaR/MonoidalCoherence/MCmonarrlist.v @@ -0,0 +1,1793 @@ +Require Import Setoid. +Require MCDefinitions MCbw MCconsequences Lia. + +Section monarrlist_theory. + +Set Universe Polymorphism. + +Import MCDefinitions MC_setoids MCbw + MCconsequences MC_notations MCClasses UIP_facts. +Import CategoryTypeclass. + +Open Scope bw_scope. + +Context {X : Type} {UIPX : UIP X}. +Context {cC : Category X} {cCh : CategoryCoherence cC} + {mC : MonoidalCategory cC} {mCh : MonoidalCategoryCoherence mC}. + +Local Notation bw := (bw X). +Local Notation "a ⟶ b" := (@monarr X cC mC a b) (at level 60). + +Section monarr_norm_equiv_theory. + +Definition monarr_norm_equiv {a a' b b' : bw} + (f : a ⟶ a') (g : b ⟶ b') : Prop := + exists (Hin : Nf b = Nf a) (Hout : Nf a' = Nf b'), + bwarr_of_Nf_eq Hin ◌ f ◌ bwarr_of_Nf_eq Hout ≊ g. + +Lemma monarr_norm_equiv_refl {a b : bw} : + reflexive _ (@monarr_norm_equiv a b a b). +Proof. + intros f. + exists eq_refl, eq_refl. + rewrite monarr_id_l, monarr_id_r. + easy. +Qed. + +Lemma monarr_norm_equiv_symm {a a' b b' : bw} + (f : monarr a a') (g : monarr b b') : + monarr_norm_equiv f g -> monarr_norm_equiv g f. +Proof. + intros (H & H' & Hequiv). + exists (eq_sym H), (eq_sym H'). + rewrite <- Hequiv. + rewrite <- 2!monarr_assoc, monarr_arrcomp, monarr_id_r. + rewrite monarr_assoc, monarr_arrcomp, monarr_id_l. + easy. +Qed. + +Lemma monarr_norm_equiv_symmetric {a a' b b' : bw} + (f : monarr a a') (g : monarr b b') : + monarr_norm_equiv f g <-> monarr_norm_equiv g f. +Proof. + split; apply monarr_norm_equiv_symm. +Qed. + +Lemma monarr_norm_equiv_trans {a a' b b' c c' : bw} + {f : monarr a a'} {g : monarr b b'} {h : monarr c c'} : + monarr_norm_equiv f g -> monarr_norm_equiv g h -> + monarr_norm_equiv f h. +Proof. + rewrite (monarr_norm_equiv_symmetric f g), + (monarr_norm_equiv_symmetric g h). + intros (Hab & Hba' & Hgf) (Hbc & Hcb' & Hhg). + exists (eq_sym (eq_trans Hab Hbc)), (eq_sym (eq_trans Hcb' Hba')). + rewrite <- Hgf, <- Hhg. + rewrite <- 4!monarr_assoc, 2!monarr_arrcomp, monarr_id_r. + rewrite 2!monarr_assoc, 2!monarr_arrcomp, monarr_id_l. + easy. +Qed. + +Lemma monarr_norm_equiv_trans' {a b a' b' c d c' d' : bw} + {f : monarr a b} {f' : monarr a' b'} + {g : monarr c d} {g' : monarr c' d'} + (Hf : monarr_norm_equiv f f') + (Hg : monarr_norm_equiv g g') + (Hfg : monarr_norm_equiv f' g') : + monarr_norm_equiv f g. +Proof. + exact (monarr_norm_equiv_trans + (monarr_norm_equiv_trans Hf Hfg) + (monarr_norm_equiv_symm _ _ Hg)). +Qed. + +#[global] Add Parametric Relation {a b : bw} : (a ⟶ b) (@monarr_norm_equiv a b a b) + reflexivity proved by monarr_norm_equiv_refl + symmetry proved by monarr_norm_equiv_symm + transitivity proved by (fun _ _ _ => monarr_norm_equiv_trans) + as monarr_norm_equiv_rel. + +#[global] Add Parametric Morphism {a a' b b' : bw} : (@monarr_norm_equiv a a' b b') + with signature + monarrequiv a a' ==> monarrequiv b b' ==> iff + as monarr_norm_equiv_mor. +Proof. + intros ? ? H ? ? H'. + unfold monarr_norm_equiv. + setoid_rewrite H. + setoid_rewrite H'. + easy. +Qed. + +#[export] Instance monarr_norm_equiv_subrel (a b : bw) : + subrelation (@monarr_norm_equiv a b a b) (monarrequiv a b). +Proof. + intros f g (? & ? & <-). + now rewrite monarr_id_l, monarr_id_r. +Qed. + +#[export] Instance subrel_monarr_norm_equiv (a b : bw) : + subrelation (monarrequiv a b) (@monarr_norm_equiv a b a b). +Proof. + intros f g ->. + easy. +Qed. + +Lemma monarr_norm_equiv_of_monarrequiv {a b} + {f g : a ⟶ b} (H : f ≊ g) : + monarr_norm_equiv f g. +Proof. + exists eq_refl, eq_refl; now rewrite monarr_id_l, monarr_id_r. +Qed. + +Lemma Nf_eq_in_of_norm_equiv {a a' b b' : bw} + {f : a ⟶ a'} {g : b ⟶ b'} : + monarr_norm_equiv f g -> Nf a = Nf b. +Proof. + intros (? & _ & _); easy. +Qed. + +Lemma Nf_eq_out_of_norm_equiv {a a' b b' : bw} + {f : monarr a a'} {g : monarr b b'} : + monarr_norm_equiv f g -> Nf a' = Nf b'. +Proof. + intros (_ & ? & _); easy. +Qed. + +Lemma monarr_norm_equiv_struct_l {a b c : bw} (f : bwarr a b) (g : b ⟶ c) : + monarr_norm_equiv (monarrcomp f g) g. +Proof. + exists (eq_sym (Nf_eq_of_arr f)), eq_refl. + rewrite monarr_id_r. + rewrite monarr_assoc, monarr_arrcomp, monarr_id_l. + easy. +Qed. + +Lemma monarr_norm_equiv_struct_r {a b c : bw} (f : a ⟶ b) (g : bwarr b c) : + monarr_norm_equiv (monarrcomp f g) f. +Proof. + exists eq_refl, (eq_sym (Nf_eq_of_arr g)). + rewrite monarr_id_l. + rewrite <- monarr_assoc, monarr_arrcomp, monarr_id_r. + easy. +Qed. + +Lemma monarr_norm_equiv_struct_l_iff {a b c m n : bw} + (f : bwarr a b) (g : b ⟶ c) (h : m ⟶ n) : + monarr_norm_equiv (monarrcomp f g) h <-> monarr_norm_equiv g h. +Proof. + split; intros (H & H' & Hrw). + - exists (eq_trans H (Nf_eq_of_arr f)), H'. + rewrite monarr_assoc, monarr_arrcomp in Hrw. + erewrite monarr_struct. + apply Hrw. + - exists (eq_trans H (eq_sym (Nf_eq_of_arr f))), H'. + rewrite monarr_assoc, monarr_arrcomp. + erewrite monarr_struct. + apply Hrw. +Qed. + +Lemma monarr_norm_equiv_struct_r_iff {a b c m n : bw} + (f : bwarr b c) (g : a ⟶ b) (h : m ⟶ n) : + monarr_norm_equiv (monarrcomp g f) h <-> monarr_norm_equiv g h. +Proof. + split; intros (H & H' & Hrw). + - exists H, (eq_trans (Nf_eq_of_arr f) H'). + rewrite <- 2!monarr_assoc, monarr_arrcomp, monarr_assoc in Hrw. + rewrite <- Hrw. + eauto with monarrdb. + - exists H, (eq_trans (eq_sym (Nf_eq_of_arr f)) H'). + rewrite <- 2!monarr_assoc, monarr_arrcomp. + rewrite <- Hrw. + eauto with monarrdb. +Qed. + +Lemma monarr_norm_equiv_struct_l'_iff {a b c m n : bw} + (f : bwarr a b) (g : b ⟶ c) (h : m ⟶ n) : + monarr_norm_equiv h (monarrcomp f g) <-> monarr_norm_equiv h g. +Proof. + rewrite monarr_norm_equiv_symmetric, monarr_norm_equiv_struct_l_iff. + apply monarr_norm_equiv_symmetric. +Qed. + +Lemma monarr_norm_equiv_struct_r'_iff {a b c m n : bw} + (f : bwarr b c) (g : a ⟶ b) (h : m ⟶ n) : + monarr_norm_equiv h (monarrcomp g f) <-> monarr_norm_equiv h g. +Proof. + rewrite monarr_norm_equiv_symmetric, monarr_norm_equiv_struct_r_iff. + apply monarr_norm_equiv_symmetric. +Qed. + +Lemma monarr_norm_equiv_conj_struct_iff {a b c d m n : bw} + (f : bwarr a b) (g : b ⟶ c) (f' : bwarr c d) (h : m ⟶ n) : + monarr_norm_equiv (f ◌ g ◌ f') h <-> monarr_norm_equiv g h. +Proof. + now rewrite monarr_norm_equiv_struct_r_iff, monarr_norm_equiv_struct_l_iff. +Qed. + +Lemma monarr_norm_equiv_conj_struct'_iff {a b c d m n : bw} + (f : bwarr a b) (g : b ⟶ c) (f' : bwarr c d) (h : m ⟶ n) : + monarr_norm_equiv h (f ◌ g ◌ f') <-> monarr_norm_equiv h g. +Proof. + now rewrite monarr_norm_equiv_symmetric, + monarr_norm_equiv_conj_struct_iff, monarr_norm_equiv_symmetric. +Qed. + +Lemma monarr_norm_equiv_comp {a a' a'' b b' b''} + (f : a ⟶ a') (f' : a' ⟶ a'') + (g : b ⟶ b') (g' : b' ⟶ b'') : + monarr_norm_equiv f g -> + monarr_norm_equiv f' g' -> + monarr_norm_equiv (f ◌ f') (g ◌ g'). +Proof. + intros (Hin & Hout & Hfg) (Hin' & Hout' & Hfg'). + rewrite <- Hfg, <- Hfg'. + rewrite monarr_assoc, monarr_norm_equiv_struct_r'_iff. + rewrite <- 2!monarr_assoc, monarr_norm_equiv_struct_l'_iff. + rewrite (monarr_assoc (monarrstruct _)), monarr_arrcomp, monarr_id_l. + apply monarr_norm_equiv_refl. +Qed. + +Lemma monarr_norm_equiv_tens {a a' b b' c c' d d'} + (f : monarr a a') (f' : monarr b b') + (g : monarr c c') (g' : monarr d d') : + monarr_norm_equiv f g -> + monarr_norm_equiv f' g' -> + monarr_norm_equiv (monarrtens f f') (monarrtens g g'). +Proof. + intros (Hin & Hout & Hfg) (Hin' & Hout' & Hfg'). + rewrite <- Hfg, <- Hfg'. + rewrite 2!monarr_tens_comp, 2!monarr_arrtens. + rewrite monarr_norm_equiv_conj_struct'_iff. + reflexivity. +Qed. + +Lemma monarr_norm_equiv_arrid_e {a b} (hd : a ⟶ b) : + monarr_norm_equiv (arrid e) hd -> + forall (g : bwarr a b), hd ≊ g. +Proof. + intros (Ha & Hb & Hhid). + intros g. + rewrite <- Hhid. + rewrite !monarr_arrcomp. + apply monarr_struct. +Qed. + +Lemma monarrequiv_iff_monarr_norm_equiv {a b} (f g : a ⟶ b) : + f ≊ g <-> monarr_norm_equiv f g. +Proof. + split; [intros ->; easy|]. + intros (? & ? & <-). + now rewrite monarr_id_l, monarr_id_r. +Qed. + +Lemma monarr_norm_equiv_of_Nfs_eq_equiv {a b a' b'} {f : a ⟶ b} {g : a' ⟶ b'} + (Ha : Nf a = Nf a') (Hb : Nf b = Nf b') : + f ≊ bwarr_of_Nf_eq Ha ◌ g ◌ bwarr_of_Nf_eq (eq_sym Hb) + -> monarr_norm_equiv f g. +Proof. + intros H; symmetry in H; apply monarr_norm_equiv_symmetric. + eexists; eexists; eauto. +Qed. + +Lemma monarr_norm_equiv_tens_bwarr_e_l {a b c} (f : a ⟶ b) (h : bwarr e c) : + monarr_norm_equiv (h ⧆ f) f. +Proof. + exists eq_refl, (f_equal (⟦ b ⟧) (eq_sym (Nf_eq_of_arr h))). + rewrite (monarr_struct _ (h ^- ⊠ arrid b ○ arrlunitor b)). + rewrite <- monarr_arrcomp, <- monarr_assoc, (monarr_assoc (h ⧆ f)). + rewrite <- monarr_arrtens, <- monarr_tens_comp, monarr_struct_inv. + rewrite monarr_id_r, monarr_lunitor_nat_r, monarr_assoc. + now rewrite monarr_arrcomp, monarr_id_l. +Qed. + +Lemma monarr_norm_equiv_tens_bwarr_e_r {a b c} (f : a ⟶ b) (h : bwarr e c) : + monarr_norm_equiv (f ⧆ h) f. +Proof. + exists eq_refl. + unshelve (eexists). + - rewrite Nf_tens_bwnormapp, <- (Nf_eq_of_arr h). + easy. + - rewrite (monarr_struct _ (arrid b ⊠ h ^- ○ arrrunitor b)). + rewrite <- monarr_arrcomp, <- monarr_assoc, (monarr_assoc (f ⧆ h)). + rewrite <- monarr_arrtens, <- monarr_tens_comp, monarr_struct_inv. + rewrite monarr_id_r, monarr_runitor_nat_r, monarr_assoc. + now rewrite monarr_arrcomp, monarr_id_l. +Qed. + +Lemma monarr_norm_equiv_tens_assoc {a b c d m n} + (f : a ⟶ b) (g : c ⟶ d) (h : m ⟶ n) : + monarr_norm_equiv (f ⧆ g ⧆ h) (f ⧆ (g ⧆ h)). +Proof. + rewrite <- (monarr_norm_equiv_struct_r'_iff (arrassoc _ _ _)). + rewrite monarr_assoc_nat_r. + apply monarr_norm_equiv_struct_l'_iff. + easy. +Qed. + +Lemma monarr_norm_equiv_struct_eq_in + {a b a' b'} (f : bwarr a b) (g : bwarr a' b') + (H : Nf a = Nf a') : + monarr_norm_equiv f g. +Proof. + exists (eq_sym H), + (eq_trans (eq_trans (eq_sym (Nf_eq_of_arr f)) H) (Nf_eq_of_arr g)). + rewrite !monarr_arrcomp; apply monarr_struct. +Qed. + +Lemma monarr_norm_equiv_struct_eq_out + {a b a' b'} (f : bwarr a b) (g : bwarr a' b') + (H : Nf b = Nf b') : + monarr_norm_equiv f g. +Proof. + exists + (eq_sym (eq_trans (eq_trans (Nf_eq_of_arr f) H) + (eq_sym (Nf_eq_of_arr g)))), H. + rewrite !monarr_arrcomp; apply monarr_struct. +Qed. + +Lemma equiv_conj_of_monarr_norm_equiv {a b m n} (f : a ⟶ b) + (g : m ⟶ n) (h : bwarr a m) (h' : bwarr n b) : + monarr_norm_equiv f g -> + f ≊ h ◌ g ◌ h'. +Proof. + rewrite monarr_norm_equiv_symmetric. + intros (Hin & Hout & Hfg). + rewrite <- Hfg. + eauto with monarrdb. +Qed. + +End monarr_norm_equiv_theory. + +Section monarrlist_definitions. + +Import List ListNotations. + +#[universes(polymorphic=yes,cumulative=yes)] +Inductive monarrlistelt := + | monarrlist_id (a : bw) : monarrlistelt + | monarrlist_arr (a b : bw) (f : a ⟶ b) : monarrlistelt. + +Definition source (f : monarrlistelt) : bw := + match f with + | monarrlist_id a => a + | monarrlist_arr a b g => a + end. + +Definition target (f : monarrlistelt) : bw := + match f with + | monarrlist_id a => a + | monarrlist_arr a b g => b + end. + +Definition realize_monarrlistelt (f : monarrlistelt) : + monarr (source f) (target f) := + match f with + | monarrlist_id a => arrid a + | monarrlist_arr a b g => g + end. + +Definition monarrlist := (list monarrlistelt). + +Definition monarrlist_source (f : monarrlist) : bw := + fold_right (fun a n => tens n a) e (map source f). + +Definition monarrlist_target (f : monarrlist) : bw := + fold_right (fun a n => tens n a) e (map target f). + +Definition source_vars (fs : monarrlist) := + flat_map (Basics.compose bw_to_varlist source) fs. + +Definition target_vars (fs : monarrlist) := + flat_map (Basics.compose bw_to_varlist target) fs. + +Fixpoint stack_monarrlist (f : monarrlist) : + monarr (monarrlist_source f) (monarrlist_target f) := + match f with + | nil => arrid norm_e + | mae :: f' => + monarrtens + (stack_monarrlist f') + (realize_monarrlistelt mae) + end. + +Definition composable (fs gs : monarrlist) := + Nf (monarrlist_target fs) = Nf (monarrlist_source gs). + +Definition compose_composable_monarr {a a' b b'} + (f : a ⟶ a') (g : b ⟶ b') + (H : Nf a' = Nf b) := + f ◌ bwarr_of_Nf_eq H ◌ g. + +Definition compose_composable (fs gs : monarrlist) + (H : composable fs gs) := + stack_monarrlist fs ◌ bwarr_of_Nf_eq H + ◌ stack_monarrlist gs. + +Definition monarrlistelt_map (f : forall a b : bw, a ⟶ b -> a ⟶ b) + (g : monarrlistelt) := + match g with + | monarrlist_id a => monarrlist_id a + | monarrlist_arr a b g' => monarrlist_arr a b (f a b g') + end. + +Definition monarrlist_list_source (fss : list monarrlist) : bw := + match fss with + | nil => e + | fs :: fss' => monarrlist_source fs + end. + +Fixpoint monarrlist_list_target (fss : list monarrlist) : bw := + match fss with + | nil => e + | fs :: nil => monarrlist_target fs + | fs :: fss' => monarrlist_list_target fss' + end. + +Fixpoint totally_composable_helper + (fst : monarrlist) (fss : list monarrlist) := + match fss with + | nil => True + | fst' :: fss' => composable fst fst' /\ + totally_composable_helper fst' fss' + end. + +Definition totally_composable (fss : list monarrlist) := + match fss with + | nil => True + | fst :: fss' => totally_composable_helper fst fss' + end. + +Fixpoint compose_totally_composable_helper + fs (fss : list monarrlist) : totally_composable_helper fs fss -> + (monarrlist_list_source (fs::fss)) ⟶ (monarrlist_list_target (fs :: fss)) := + match fss with + | nil => fun _ => stack_monarrlist fs + | fs' :: fss => + fun H => + compose_composable_monarr + (stack_monarrlist fs) + (compose_totally_composable_helper fs' fss (proj2 H)) + (proj1 H) + end. + +Definition compose_totally_composable + (fss : list monarrlist) : totally_composable fss -> + monarr + (monarrlist_list_source fss) + (monarrlist_list_target fss) := + match fss with + | nil => fun _ => arrid e + | fs :: fss' => + compose_totally_composable_helper fs fss' + end. + +Definition zip_with_default_l {A B C} (f : A -> B -> C) + (xs : list A) (ydef : B) : list C := + map (fun x => f x ydef) xs. + +Definition zip_with_default_r {A B C} (f : A -> B -> C) + (xdef : A) (ys : list B) : list C := + map (f xdef) ys. + +Fixpoint zip_defaults {A B C} (f : A -> B -> C) + (xs : list A) (ys : list B) (xdef : A) (ydef : B) : list C := + match xs, ys with + | nil, ys' => zip_with_default_r f xdef ys' + | xs', nil => zip_with_default_l f xs' ydef + | x::xs', y::ys' => (f x y) :: zip_defaults f xs' ys' xdef ydef + end. + +Definition monarrlistelt_equiv (f g : monarrlistelt) := + monarr_norm_equiv + (realize_monarrlistelt f) + (realize_monarrlistelt g). + +Fixpoint ForallF {A} (P : A -> Prop) (l : list A) : Prop := + match l with + | nil => True + | a :: l' => P a /\ ForallF P l' + end. + +Definition all_monarrlist_equiv (fs gs : monarrlist) := + ForallF (uncurry monarrlistelt_equiv) + (zip_defaults pair fs gs (monarrlist_id e) (monarrlist_id e)). + +Definition all_monarrlist_list_equiv (fss gss : list monarrlist) := + ForallF (uncurry all_monarrlist_equiv) + (zip_defaults pair fss gss nil nil). + +End monarrlist_definitions. + +Import List ListNotations. + +(* Notation monarrlist := (list monarrlistelt). *) + +Section composable_theory. + +Lemma composable_iff_bweq (fs gs : monarrlist) : + composable fs gs <-> bweq (monarrlist_target fs) (monarrlist_source gs). +Proof. + apply Nf_eq_iff_bweq. +Qed. + +Lemma composable_iff_varlist_eq (fs gs : monarrlist) : + composable fs gs <-> target_vars fs = source_vars gs. +Proof. + unfold composable. + rewrite Nf_eq_iff_varlist_eq. + unfold monarrlist_target, monarrlist_source. + rewrite 2!bw_to_varlist_fold_right. + rewrite 2!map_map. + rewrite <- 2!flat_map_concat_map. + easy. +Qed. + +Lemma composable_of_app_composable {fs fs' gs gs' : monarrlist} : + composable (fs ++ fs') (gs ++ gs') -> + composable fs gs -> composable fs' gs'. +Proof. + rewrite 3!composable_iff_varlist_eq. + unfold source_vars, target_vars. + rewrite 2!flat_map_app. + intros H H'. + rewrite H' in H. + apply app_inv_head in H. + easy. +Qed. + +Lemma composable_independence {fs gs : monarrlist} (H H' : composable fs gs) : + H = H'. +Proof. apply uip. Qed. + +Lemma compose_composable_indep {fs gs : monarrlist} (H H' : composable fs gs) : + compose_composable fs gs H = compose_composable fs gs H'. +Proof. + f_equal; apply uip. +Qed. + +Lemma compose_composable_monarr_indep {a a' b b'} + (fs : a ⟶ b) (gs : a' ⟶ b') (H G : Nf b = Nf a') : + compose_composable_monarr fs gs H = + compose_composable_monarr fs gs G. +Proof. + f_equal; apply uip. +Qed. + +#[global] Add Parametric Morphism (a a' b b' : bw) : + (@compose_composable_monarr a a' b b') + with signature + (monarrequiv a a') ==> (monarrequiv b b') ==> true_rel ==> + (monarrequiv a b') as compose_composable_monarr_mor. +Proof. + intros f f' Hf g g' Hg H1 H2 _. + unfold compose_composable_monarr. + apply monarr_comp; [apply monarr_comp|]; + assumption + apply monarr_struct. +Qed. + +Lemma forall_composable_map {f} (Hc : forall fs gs, + composable fs gs -> composable (fst (f fs gs)) (snd (f fs gs))) + {fsgs} (Hfsgs : Forall (uncurry composable) fsgs) : + Forall (uncurry composable) (map (uncurry f) fsgs). +Proof. + induction fsgs as [|[g g'] fsgs IHfsgs]; [easy|]. + simpl. + inversion Hfsgs; subst. + constructor; [|apply IHfsgs; assumption]. + - rewrite (surjective_pairing (f g g')). + apply Hc; assumption. +Qed. + + +Lemma totally_composable_of_cons {fst : monarrlist} {fss : list monarrlist} : + totally_composable (fst :: fss) -> totally_composable fss. +Proof. + destruct fss; easy + intros H; apply H. +Qed. + +Lemma totally_composable_indep (fss : list monarrlist) + (H G : totally_composable fss) : H = G. +Proof. + induction fss; [now destruct H, G|]. + destruct fss; [now destruct H, G|]. + simpl in H, G. + destruct H, G. + f_equal. + * apply uip. + * apply IHfss. +Qed. + +Lemma compose_totally_composable_indep + (fss : list monarrlist) (H G : totally_composable fss) : + compose_totally_composable fss H = + compose_totally_composable fss G. +Proof. + rewrite (totally_composable_indep fss H G). + easy. +Qed. + +Lemma totally_composable_app_restrict_r {fsts fss} : + totally_composable (fsts ++ fss) -> + totally_composable fss. +Proof. + destruct fss; [easy|]. + induction fsts; [easy|]. + intros H; apply IHfsts. + simpl in H. + destruct fsts; apply H. +Qed. + +Lemma totally_composable_app_restrict_l {fsts fss} : + totally_composable (fsts ++ fss) -> + totally_composable fsts. +Proof. + (* destruct fss; [rewrite app_nil_r; easy|]. *) + induction fsts; [easy|]. + destruct fsts; [easy|]. + intros H; split. + - apply H. + - apply IHfsts, H. +Qed. + +Lemma middle_arr_composable_of_totally_composable + {a b fss fss'} (H : totally_composable ((a :: fss) ++ (b :: fss'))) : + Nf (monarrlist_list_target (a :: fss)) + = Nf (monarrlist_list_source (b :: fss')). +Proof. + simpl monarrlist_list_source. + revert a H; + induction fss as [|fst fss IHfss]; + intros a H; [apply H|]. + simpl totally_composable in *. + specialize (IHfss _ (proj2 H)). + apply IHfss. +Qed. + +Lemma monarrlist_list_target_app {fss b fss'} : + monarrlist_list_target (fss ++ b :: fss') = + monarrlist_list_target (b :: fss'). +Proof. + induction fss; [easy|]. + destruct fss; + apply IHfss. +Qed. + +Lemma monarrlist_list_source_app fss fss' : + monarrlist_list_source (fss ++ fss') = + match fss with + | nil => monarrlist_list_source fss' + | x::_ => monarrlist_source x + end. +Proof. + destruct fss; easy. +Qed. + +Lemma Nf_monarrlist_source_app hd tl : + Nf (monarrlist_source (hd ++ tl)) = + bwnormapp (Nf (monarrlist_source tl)) + (Nf (monarrlist_source hd)). +Proof. + unfold monarrlist_source. + rewrite map_app. + apply Nf_fold_right_tens. +Qed. + +Lemma Nf_monarrlist_target_app hd tl : + Nf (monarrlist_target (hd ++ tl)) = + bwnormapp (Nf (monarrlist_target tl)) + (Nf (monarrlist_target hd)). +Proof. + unfold monarrlist_target. + rewrite map_app. + apply Nf_fold_right_tens. +Qed. + +Lemma Nf_monarrlist_source_app' hd tl : + Nf (monarrlist_source (hd ++ tl)) = + Nf (monarrlist_source tl ⨂ monarrlist_source hd). +Proof. + rewrite Nf_tens_bwnormapp. + apply Nf_monarrlist_source_app. +Qed. + +Lemma Nf_monarrlist_target_app' hd tl : + Nf (monarrlist_target (hd ++ tl)) = + Nf (monarrlist_target tl ⨂ monarrlist_target hd). +Proof. + rewrite Nf_tens_bwnormapp. + apply Nf_monarrlist_target_app. +Qed. + +Lemma composable_app {hd1 hd2 tl1 tl2} + (Hh : composable hd1 hd2) + (Ht : composable tl1 tl2) : + composable (hd1 ++ tl1) (hd2 ++ tl2). +Proof. + unfold composable in *. + rewrite Nf_monarrlist_source_app, Nf_monarrlist_target_app. + f_equal; easy. +Qed. + +Lemma app_composable_of_composable_mid {fss gss} + (Hf : totally_composable fss) (Hg : totally_composable gss) + (Hfg : Nf (monarrlist_list_target fss) = Nf (monarrlist_list_source gss)) : + totally_composable (fss ++ gss). +Proof. + induction fss. + - apply Hg. + - simpl. + destruct fss. + + simpl. + destruct gss; [easy|]. + split; + [apply Hfg|]. + apply Hg. + + split; + [apply Hf|]. + apply IHfss; [apply Hf|]. + apply Hfg. +Qed. + +Lemma zip_defaults_cons {A B C} (f : A -> B -> C) x y xs ys xdef ydef : + zip_defaults f (x :: xs) (y :: ys) xdef ydef = + (f x y) :: zip_defaults f xs ys xdef ydef. +Proof. + easy. +Qed. + +Lemma monarrlist_list_target_cons_cons f f' fs : + monarrlist_list_target (f :: f' :: fs) = monarrlist_list_target (f' :: fs). +Proof. + easy. +Qed. + +Lemma middle_arr_composable_of_totally_composable_nonempty + {fss fss'} (H : totally_composable (fss ++ fss')) + (Hfss : fss <> []) (Hfss' : fss' <> []) : + Nf (monarrlist_list_target fss) + = Nf (monarrlist_list_source fss'). +Proof. + destruct fss, fss'; [easy..|]. + apply middle_arr_composable_of_totally_composable; assumption. +Qed. + +Lemma monarrlist_list_source_app_not_nil {fss} (H : fss <> []) gss : + monarrlist_list_source (fss ++ gss) = monarrlist_list_source fss. +Proof. + rewrite monarrlist_list_source_app. + destruct fss; easy. +Qed. + +Lemma monarrlist_list_target_app_not_nil {gss} (H : gss <> []) fss : + monarrlist_list_target (fss ++ gss) = monarrlist_list_target gss. +Proof. + destruct gss; [easy|]. + now rewrite monarrlist_list_target_app. +Qed. + +Lemma monarrlist_target_singleton fss : + monarrlist_target [fss] = e ⨂ target fss. +Proof. + easy. +Qed. + +Lemma monarrlist_source_cons f fs : + monarrlist_source (f :: fs) = monarrlist_source fs ⨂ source f. +Proof. + easy. +Qed. + +Lemma monarrlist_target_cons f fs : + monarrlist_target (f :: fs) = monarrlist_target fs ⨂ target f. +Proof. + easy. +Qed. + +Lemma zip_with_default_l_totally_composable {fss} + (Hf : totally_composable fss) a : + totally_composable + (zip_with_default_l (@app monarrlistelt) fss [monarrlist_id a]). +Proof. + induction fss; [easy|destruct fss; [easy|]]. + split. + - apply composable_app; [apply Hf | easy]. + - apply IHfss, Hf. +Qed. + +Lemma zip_with_default_r_totally_composable {fss} + (Hf : totally_composable fss) a : + totally_composable + (zip_with_default_r (@app monarrlistelt) [monarrlist_id a] fss). +Proof. + induction fss; [easy|destruct fss; [easy|]]. + split. + - apply composable_app; [easy | apply Hf]. + - apply IHfss, Hf. +Qed. + +Lemma zip_with_default_l_source_nonempty {fss} (H : fss <> []) a : + Nf (monarrlist_list_source (zip_with_default_l (@app monarrlistelt) fss a)) + = Nf (monarrlist_source a ⨂ monarrlist_list_source fss). +Proof. + destruct fss; [easy|]. + apply Nf_monarrlist_source_app'. +Qed. + +Lemma zip_with_default_l_target_nonempty {fss} (H : fss <> []) a : + Nf (monarrlist_list_target (zip_with_default_l (@app monarrlistelt) fss a)) + = Nf (monarrlist_target a ⨂ monarrlist_list_target fss). +Proof. + induction fss; [easy|]. + destruct fss; [apply Nf_monarrlist_target_app'|]. + apply IHfss; easy. +Qed. + +Lemma zip_with_default_r_source_nonempty {fss} (H : fss <> []) a : + Nf (monarrlist_list_source (zip_with_default_r (@app monarrlistelt) a fss)) + = Nf (monarrlist_list_source fss ⨂ monarrlist_source a). +Proof. + destruct fss; [easy|]. + apply Nf_monarrlist_source_app'. +Qed. + +Lemma zip_with_default_r_target_nonempty {fss} (H : fss <> []) a : + Nf (monarrlist_list_target (zip_with_default_r (@app monarrlistelt) a fss)) + = Nf (monarrlist_list_target fss ⨂ monarrlist_target a). +Proof. + induction fss; [easy|]. + destruct fss; [apply Nf_monarrlist_target_app'|]. + apply IHfss; easy. +Qed. + +Lemma zip_defaults_target_nonempty {fss gss} (Hf : fss <> []) (Hg : gss <> []) : + Nf (monarrlist_list_target (zip_defaults (@app monarrlistelt) + fss gss + [monarrlist_id (monarrlist_list_target fss)] + [monarrlist_id (monarrlist_list_target gss)])) + = Nf (monarrlist_list_target gss ⨂ monarrlist_list_target fss). +Proof. + revert gss Hg; + induction fss; + intros gss Hg; [ + apply (zip_with_default_r_target_nonempty Hg)|]. + destruct fss. + - destruct gss; [easy|]. + destruct gss; [apply Nf_monarrlist_target_app'|]. + rewrite zip_defaults_cons. + change (?a :: ?l) with ((a::nil) ++ l). + rewrite monarrlist_list_target_app_not_nil by easy. + etransitivity. + apply zip_with_default_r_target_nonempty; easy. + rewrite monarrlist_list_target_app. + easy. + - destruct gss; [easy|]. + destruct gss. + + rewrite zip_defaults_cons. + change (?a :: ?l) with ((a::nil) ++ l). + rewrite monarrlist_list_target_app_not_nil by easy. + etransitivity. + apply zip_with_default_l_target_nonempty; easy. + rewrite monarrlist_list_target_app. + easy. + + rewrite zip_defaults_cons. + change (?a :: ?l) with ((a::nil) ++ l). + rewrite monarrlist_list_target_app_not_nil by easy. + rewrite 2!monarrlist_list_target_cons_cons. + etransitivity; + [apply (IHfss ltac:(easy) (l1 :: gss) ltac:(easy))|]. + easy. +Qed. + +Lemma zip_defaults_target_defaults_indep {fss gss df dg} + (Hf : fss <> []) (Hg : gss <> []) + (Hdf : Nf df = Nf (monarrlist_list_target fss)) + (Hdg : Nf dg = Nf (monarrlist_list_target gss)) : + Nf (monarrlist_list_target (zip_defaults (@app monarrlistelt) + fss gss + [monarrlist_id df] + [monarrlist_id dg])) + = Nf (monarrlist_list_target (zip_defaults (@app monarrlistelt) + fss gss + [monarrlist_id (monarrlist_list_target fss)] + [monarrlist_id (monarrlist_list_target gss)])). +Proof. + revert gss Hg Hdg; + induction fss; intros gss Hg Hdg; [induction gss; [easy|]|destruct gss]. + - destruct gss. + + cbn -[app Nf]. + rewrite 2!Nf_monarrlist_target_app; f_equal. + easy. + + apply IHgss; easy. + - destruct fss. + + cbn -[monarrlist_target app Nf]. + rewrite 2!Nf_monarrlist_target_app; f_equal. + easy. + + refine (IHfss _ _ [] _ _); easy. + - destruct fss. + + cbn -[monarrlist_target app Nf]. + destruct gss; [easy|]. + simpl. + etransitivity; + [apply (zip_with_default_r_target_nonempty + (fss:=m0::gss) ltac:(easy))| + etransitivity; + [|symmetry; apply (zip_with_default_r_target_nonempty + (fss:=m0::gss) ltac:(easy))]]. + rewrite 2!Nf_tens_bwnormapp. + f_equal; easy. + + destruct gss. + * simpl. + etransitivity; + [apply (zip_with_default_l_target_nonempty + (fss:=m0::fss) ltac:(easy))| + etransitivity; + [|symmetry; apply (zip_with_default_l_target_nonempty + (fss:=m0::fss) ltac:(easy))]]. + rewrite 2!Nf_tens_bwnormapp. + f_equal; easy. + * refine (IHfss _ _ (m1 :: gss) _ _); easy. +Qed. + +Lemma zip_defaults_source_nonempty {fss gss} (Hf : fss <> []) (Hg : gss <> []) : + Nf (monarrlist_list_source (zip_defaults (@app monarrlistelt) + fss gss + [monarrlist_id (monarrlist_list_target fss)] + [monarrlist_id (monarrlist_list_target gss)])) + = Nf (monarrlist_list_source gss ⨂ monarrlist_list_source fss). +Proof. + destruct fss, gss; try easy. + apply Nf_monarrlist_source_app'. +Qed. + +Lemma totally_composable_helper_swap {a b fss} : + Nf (monarrlist_target a) = Nf (monarrlist_target b) -> + totally_composable_helper a fss -> + totally_composable_helper b fss. +Proof. + intros Heq. + destruct fss; [easy|]. + intros [Hl Hr]. + split; [|apply Hr]. + unfold composable in *. + rewrite <- Heq. + easy. +Qed. + +End composable_theory. + +Section composition_theory. + +Lemma monarr_norm_equiv_comp_composable {a b c d a' b' c' d'} + (f : a ⟶ b) (g : c ⟶ d) (f' : a' ⟶ b') (g' : c' ⟶ d') Hc Hc' : + monarr_norm_equiv f f' -> + monarr_norm_equiv g g' -> + monarr_norm_equiv + (compose_composable_monarr f g Hc) + (compose_composable_monarr f' g' Hc'). +Proof. + intros Hf Hg. + unfold compose_composable_monarr. + apply monarr_norm_equiv_comp; + [apply monarr_norm_equiv_struct_r_iff, + monarr_norm_equiv_struct_r'_iff|]; + assumption. +Qed. + +Lemma ForallFE {A} (P : A -> Prop) (l : list A) : + ForallF P l = fold_right and True (map P l). +Proof. + induction l; [easy|]. + simpl. + f_equal; apply IHl. +Qed. + +Lemma ForallF_iff_Forall {A} (P : A -> Prop) (l : list A) : + ForallF P l <-> Forall P l. +Proof. + induction l; [easy|]. + rewrite Forall_cons_iff, <- IHl. + easy. +Qed. + +Lemma zip_defaults_nil_r {A B C} (f : A -> B -> C) xs xdef ydef : + zip_defaults f xs [] xdef ydef = + zip_with_default_l f xs ydef. +Proof. + destruct xs; easy. +Qed. + +Lemma stack_monarrlist_app' fs fss : + monarr_norm_equiv + (stack_monarrlist (fs ++ fss)) + (stack_monarrlist fss ⧆ stack_monarrlist fs). +Proof. + apply monarr_norm_equiv_symm. + induction fs. + - apply monarr_norm_equiv_tens_bwarr_e_r. + - simpl. + eapply monarr_norm_equiv_trans; + [apply monarr_norm_equiv_symm, monarr_norm_equiv_tens_assoc|]. + apply monarr_norm_equiv_tens; easy. +Qed. + +Lemma stack_monarrlist_app fs fss : + stack_monarrlist (fs ++ fss) + ≊ bwarr_of_Nf_eq (Nf_monarrlist_source_app' fs fss) ◌ + stack_monarrlist fss ⧆ stack_monarrlist fs ◌ + bwarr_of_Nf_eq (eq_sym (Nf_monarrlist_target_app' fs fss)). +Proof. + apply equiv_conj_of_monarr_norm_equiv. + apply stack_monarrlist_app'. +Qed. + +Lemma tens_stack_monarrlist fs fss : + stack_monarrlist fss ⧆ stack_monarrlist fs + ≊ bwarr_of_Nf_eq (eq_sym (Nf_monarrlist_source_app' fs fss)) ◌ + stack_monarrlist (fs ++ fss) ◌ + bwarr_of_Nf_eq (Nf_monarrlist_target_app' fs fss). +Proof. + apply equiv_conj_of_monarr_norm_equiv. + apply monarr_norm_equiv_symm. + apply stack_monarrlist_app'. +Qed. + + + +Lemma compose_composable_monarr_comp_l {a a' a'' b b'} (f : a ⟶ a') + (g : a' ⟶ a'') (h : b ⟶ b') H : + compose_composable_monarr (f ◌ g) h H ≊ + f ◌ compose_composable_monarr g h H. +Proof. + unfold compose_composable_monarr. + now rewrite !monarr_assoc. +Qed. + +Lemma compose_composable_monarr_comp_r {a a' a'' b b'} (f : a ⟶ a') + (g : a' ⟶ a'') (h : b ⟶ b') H : + compose_composable_monarr h (f ◌ g) H ≊ + compose_composable_monarr h f H ◌ g. +Proof. + unfold compose_composable_monarr. + now rewrite !monarr_assoc. +Qed. + +Lemma compose_composable_monarr_struct_l' {a a' a'' b b'} (f : bwarr a a') + (g : a' ⟶ a'') (h : b ⟶ b') H : + compose_composable_monarr h (f ◌ g) H ≊ + compose_composable_monarr h g (eq_trans H (Nf_eq_of_arr f)). +Proof. + unfold compose_composable_monarr. + rewrite !monarr_assoc. + apply monarr_comp; [|easy]. + rewrite <- monarr_assoc, monarr_arrcomp. + auto with monarrdb. +Qed. + +Lemma compose_composable_monarr_struct_r {a a' a'' b b'} (f : a ⟶ a') + (g : bwarr a' a'') (h : b ⟶ b') H : + compose_composable_monarr (f ◌ g) h H ≊ + compose_composable_monarr f h (eq_trans (Nf_eq_of_arr g) H). +Proof. + unfold compose_composable_monarr. + apply monarr_comp; [|easy]. + rewrite <- monarr_assoc, monarr_arrcomp. + auto with monarrdb. +Qed. + +Lemma compose_totally_composable_cons + {a b fss} (H : totally_composable (a :: b :: fss)) : + (compose_totally_composable _ H) + ≊ + compose_composable_monarr (stack_monarrlist a) + (compose_totally_composable (b::fss) (proj2 H)) + (proj1 H). +Proof. + revert a b H. + induction fss as [|fst fss IHfss]; + intros a b H; [easy|]. + simpl in *. + specialize (IHfss _ _ (proj2 H)). + easy. +Qed. + +Lemma compose_totally_composable_app_helper + {a b fss fss'} (H : totally_composable ((a :: fss) ++ (b :: fss'))) : + monarrequiv _ _ + (compose_totally_composable _ H) + (monarrcomp + (compose_composable_monarr + (compose_totally_composable (a :: fss) + (totally_composable_app_restrict_l H)) + (compose_totally_composable (b :: fss') + (totally_composable_app_restrict_r H)) + (middle_arr_composable_of_totally_composable H)) + (cast_bwarr eq_refl (eq_sym monarrlist_list_target_app) (arrid _))). +Proof. + revert a H; + induction fss as [|fst fss IHfss]; + intros a H. + - simpl app. + rewrite compose_totally_composable_cons. + unfold compose_composable_monarr. + rewrite <- !monarr_assoc. + apply monarr_comp; [easy|]. + apply monarr_comp; [apply monarr_struct|]. + rewrite (monarr_struct _ (arrid _)), monarr_runit. + erewrite compose_totally_composable_indep. + reflexivity. + - simpl app. + rewrite compose_totally_composable_cons. + unfold compose_composable_monarr. + rewrite compose_totally_composable_cons. + unfold compose_composable_monarr. + rewrite <- !monarr_assoc. + apply monarr_comp; [easy|]. + apply monarr_comp; [apply monarr_struct|]. + change (fst :: fss ++ b :: fss') with ((fst :: fss) ++ b :: fss'). + rewrite IHfss. + unfold compose_composable_monarr. + rewrite <- !monarr_assoc. + repeat apply monarr_comp; + apply monarr_struct + + (erewrite compose_totally_composable_indep; + reflexivity). +Qed. + +Lemma compose_monarrlist_list_cancel_suffix_helper_nil + hd fss + (H1 : totally_composable fss) + (H2 : totally_composable (hd ++ fss)) + (Hh : totally_composable hd) : + monarr_norm_equiv + (arrid e) + (compose_totally_composable + hd Hh) -> + monarr_norm_equiv + (compose_totally_composable _ H1) + (compose_totally_composable _ H2). +Proof. + destruct hd; [|destruct fss]. + - intros _. + apply monarr_norm_equiv_of_monarrequiv. + erewrite compose_totally_composable_indep. + reflexivity. + - revert H2. + rewrite app_nil_r. + intros H2. + intros H. + erewrite (compose_totally_composable_indep (_::hd)). + apply H. + - intros H. + rewrite compose_totally_composable_app_helper. + rewrite monarr_norm_equiv_struct_r'_iff. + unfold compose_composable_monarr. + rewrite <- monarr_lunit. + apply monarr_norm_equiv_comp; + [|erewrite compose_totally_composable_indep; easy]. + rewrite monarr_norm_equiv_struct_r'_iff. + erewrite compose_totally_composable_indep. + eapply monarr_norm_equiv_trans; [|apply H]. + apply monarr_norm_equiv_struct_eq_in. + rewrite <- (middle_arr_composable_of_totally_composable H2). + symmetry. + apply (Nf_eq_out_of_norm_equiv H). +Qed. + +Lemma compose_monarrlist_list_cancel_suffix_helper + hd1 hd2 fss + (H1 : totally_composable (hd1 ++ fss)) + (H2 : totally_composable (hd2 ++ fss)) + (H1h : totally_composable hd1) + (H2h : totally_composable hd2) : + monarr_norm_equiv + (compose_totally_composable + hd1 H1h) + (compose_totally_composable + hd2 H2h) -> + monarr_norm_equiv + (compose_totally_composable _ H1) + (compose_totally_composable _ H2). +Proof. + revert H1 H2. + induction fss as [|fst fss]. + 1: { + rewrite 2!app_nil_r. + intros H1 H2. + rewrite (compose_totally_composable_indep _ H1 H1h). + rewrite (compose_totally_composable_indep _ H2 H2h). + easy. + } + destruct hd1. + 1: { + clear IHfss. + simpl app. + simpl (compose_totally_composable [] _). + intros H1 H2. + apply compose_monarrlist_list_cancel_suffix_helper_nil. + } + destruct hd2. + 1: { + clear IHfss. + simpl app. + setoid_rewrite monarr_norm_equiv_symmetric. + simpl (compose_totally_composable [] _). + intros H1 H2. + apply compose_monarrlist_list_cancel_suffix_helper_nil. + } + intros H1 H2. + rewrite 2!compose_totally_composable_app_helper. + intros (Hl & Hhd & Hequiv). + clear IHfss. + symmetry in Hequiv. + erewrite compose_totally_composable_indep in Hequiv. + rewrite Hequiv. + rewrite monarr_norm_equiv_struct_r_iff, monarr_norm_equiv_struct_r'_iff. + exists Hl, eq_refl. + rewrite monarr_id_r. + rewrite compose_composable_monarr_struct_r, compose_composable_monarr_comp_l. + apply monarr_comp; [easy|]. + apply compose_composable_monarr_mor; [..|easy]; + erewrite compose_totally_composable_indep; easy. +Qed. + +Lemma compose_monarrlist_list_cancel_suffix + hd1 hd2 fss + (H1 : totally_composable (hd1 ++ fss)) + (H2 : totally_composable (hd2 ++ fss)) : + monarr_norm_equiv + (compose_totally_composable + hd1 (totally_composable_app_restrict_l H1)) + (compose_totally_composable + hd2 (totally_composable_app_restrict_l H2)) -> + monarr_norm_equiv + (compose_totally_composable _ H1) + (compose_totally_composable _ H2). +Proof. + intros H. + eapply compose_monarrlist_list_cancel_suffix_helper. + apply H. +Qed. + +Lemma compose_monarrlist_list_cancel_prefix_helper + hd1 hd2 fss + (H1 : totally_composable (fss ++ hd1)) + (H2 : totally_composable (fss ++ hd2)) + (H1h : totally_composable hd1) + (H2h : totally_composable hd2) : + monarr_norm_equiv + (compose_totally_composable + hd1 H1h) + (compose_totally_composable + hd2 H2h) -> + monarr_norm_equiv + (compose_totally_composable _ H1) + (compose_totally_composable _ H2). +Proof. + intros H. + induction fss. + - rewrite (compose_totally_composable_indep _ _ H1h). + rewrite (compose_totally_composable_indep _ _ H2h). + easy. + - destruct fss. + 1: { + destruct hd1, hd2. + 1: easy. + 1: { + specialize (IHfss (Logic.I) (proj2 H2)). + simpl app. + rewrite compose_totally_composable_cons. + rewrite <- monarr_runit. + unfold compose_composable_monarr. + apply monarr_norm_equiv_comp. + - rewrite monarr_norm_equiv_struct_r'_iff. + easy. + - eapply monarr_norm_equiv_trans; + [|apply IHfss]. + apply monarr_norm_equiv_struct_eq_in. + simpl. + rewrite (proj1 H2). + apply Nf_eq_in_of_norm_equiv in H. + symmetry; apply H. + } + 1: { + specialize (IHfss (proj2 H1) (Logic.I)). + simpl app. + rewrite compose_totally_composable_cons. + apply monarr_norm_equiv_symm. + rewrite <- monarr_runit. + unfold compose_composable_monarr. + apply monarr_norm_equiv_comp. + - rewrite monarr_norm_equiv_struct_r'_iff. + easy. + - eapply monarr_norm_equiv_trans; + [|apply monarr_norm_equiv_symm, IHfss]. + apply monarr_norm_equiv_struct_eq_in. + simpl. + rewrite (proj1 H1). + apply (Nf_eq_in_of_norm_equiv H). + } + simpl app. + rewrite 2!compose_totally_composable_cons. + apply monarr_norm_equiv_comp_composable; [easy|]. + erewrite (compose_totally_composable_indep (m0 :: _)), + compose_totally_composable_indep. + apply H. + } + simpl app. + rewrite 2!compose_totally_composable_cons. + apply monarr_norm_equiv_comp_composable; [easy|]. + apply IHfss. +Qed. + +Lemma compose_monarrlist_list_cancel_prefix + hd1 hd2 fss + (H1 : totally_composable (fss ++ hd1)) + (H2 : totally_composable (fss ++ hd2)) : + monarr_norm_equiv + (compose_totally_composable + hd1 (totally_composable_app_restrict_r H1)) + (compose_totally_composable + hd2 (totally_composable_app_restrict_r H2)) -> + monarr_norm_equiv + (compose_totally_composable _ H1) + (compose_totally_composable _ H2). +Proof. + intros H. + eapply compose_monarrlist_list_cancel_prefix_helper. + apply H. +Qed. + +Lemma compose_totally_composable_app_nonempty + {fss fss'} (H : totally_composable (fss ++ fss')) + (Hfss : fss <> []) (Hfss' : fss' <> []) : + compose_totally_composable _ H + ≊ (cast_bwarr eq_refl + (monarrlist_list_source_app_not_nil Hfss _) (arrid _)) + ◌ (compose_composable_monarr + (compose_totally_composable fss + (totally_composable_app_restrict_l H)) + (compose_totally_composable fss' + (totally_composable_app_restrict_r H)) + (middle_arr_composable_of_totally_composable_nonempty H Hfss Hfss')) + ◌ (cast_bwarr eq_refl + (eq_sym (monarrlist_list_target_app_not_nil Hfss' _)) (arrid _)). +Proof. + destruct fss, fss'; [easy..|]. + rewrite compose_totally_composable_app_helper. + rewrite monarr_struct_id, monarr_lunit. + apply monarr_comp; [|apply monarr_struct]. + erewrite compose_composable_monarr_indep. + easy. +Qed. + +End composition_theory. + + +Section monarrlist_equivalences. + + + +Lemma all_monarrlist_cons_equiv_nil {a fs} : + all_monarrlist_equiv (a :: fs) [] -> + all_monarrlist_equiv fs []. +Proof. + unfold all_monarrlist_equiv. + rewrite 2!zip_defaults_nil_r. + simpl. + easy. +Qed. + +Lemma all_monarrlist_list_cons_equiv_nil {a fss} : + all_monarrlist_list_equiv (a :: fss) [] -> + all_monarrlist_list_equiv fss []. +Proof. + unfold all_monarrlist_list_equiv. + rewrite 2!zip_defaults_nil_r. + simpl. + easy. +Qed. + +Lemma stack_monarrlist_equiv {fs gs} + (H : all_monarrlist_equiv fs gs) : + monarr_norm_equiv + (stack_monarrlist fs) + (stack_monarrlist gs). +Proof. + revert gs H; + induction fs; intros gs H. + - induction gs. + + apply monarr_norm_equiv_refl. + + change [] with (@nil monarrlistelt ++ nil). + rewrite stack_monarrlist_app. + change (a :: gs) with ((a::nil) ++ gs). + rewrite stack_monarrlist_app. + rewrite monarr_norm_equiv_conj_struct_iff, + monarr_norm_equiv_conj_struct'_iff. + apply monarr_norm_equiv_tens; + [apply IHgs, H|]. + destruct H. + simpl. + eapply monarr_norm_equiv_trans; + [apply H|]. + apply monarr_norm_equiv_symm. + apply monarr_norm_equiv_tens_bwarr_e_l. + - destruct gs. + + change [] with (@nil monarrlistelt ++ nil). + change (a :: fs) with ((a::nil) ++ fs). + rewrite 2!stack_monarrlist_app. + rewrite monarr_norm_equiv_conj_struct_iff, + monarr_norm_equiv_conj_struct'_iff. + apply monarr_norm_equiv_tens; + [apply IHfs, (all_monarrlist_cons_equiv_nil H)|]. + eapply monarr_norm_equiv_trans; + [| apply H]. + apply monarr_norm_equiv_tens_bwarr_e_l. + + apply monarr_norm_equiv_tens; + [apply IHfs, H|]. + apply H. +Qed. + +Lemma compose_composable_app_helper {fs gs} + (Hfg : composable fs gs) {fss gss} + (Hfsgs : composable (fs ++ fss) (gs ++ gss)) + (Hfssgss : composable fss gss) : + monarr_norm_equiv + (compose_composable (fs ++ fss) (gs ++ gss) Hfsgs) + (monarrtens + (compose_composable fss gss Hfssgss) + (compose_composable fs gs Hfg)). +Proof. + unfold compose_composable. + rewrite 2!monarr_tens_comp. + rewrite 2!tens_stack_monarrlist. + rewrite <- !monarr_assoc. + rewrite monarr_norm_equiv_struct_l'_iff. + apply monarr_norm_equiv_comp; [easy|]. + rewrite monarr_arrtens. + rewrite !monarr_norm_equiv_struct_l'_iff, + monarr_norm_equiv_struct_l_iff, + monarr_norm_equiv_struct_r'_iff. + easy. +Qed. + +Lemma compose_composable_app {fs gs} + (Hfg : composable fs gs) {fss gss} + (Hfsgs : composable (fs ++ fss) (gs ++ gss)): + monarr_norm_equiv + (compose_composable (fs ++ fss) (gs ++ gss) Hfsgs) + (monarrtens + (compose_composable fss gss + (composable_of_app_composable Hfsgs Hfg)) + (compose_composable fs gs Hfg)). +Proof. + apply compose_composable_app_helper. +Qed. + +Lemma monarrlistelt_equiv_refl : + Reflexive monarrlistelt_equiv. +Proof. + intros a; + apply monarr_norm_equiv_refl. +Qed. + +Lemma monarrlistelt_equiv_symm : + Symmetric monarrlistelt_equiv. +Proof. + intros a b; + apply monarr_norm_equiv_symm. +Qed. + +Lemma monarrlistelt_equiv_trans : + Transitive monarrlistelt_equiv. +Proof. + intros a b c; + apply monarr_norm_equiv_trans. +Qed. + + +Lemma monarrlist_equiv_refl : + Reflexive all_monarrlist_equiv. +Proof. + intros f. + induction f; [easy|]. + split; [apply monarrlistelt_equiv_refl|]. + apply IHf. +Qed. + +Lemma monarrlist_equiv_symm : + Symmetric all_monarrlist_equiv. +Proof. + intros fss. + induction fss; intros gss H; [|destruct gss]. + - induction gss; [easy|]. + split. + + apply monarr_norm_equiv_symm, H. + + specialize (IHgss (proj2 H)). + destruct gss; + apply IHgss. + - split. + + apply monarr_norm_equiv_symm, H. + + apply (IHfss nil). + destruct fss; apply H. + - split. + + apply monarr_norm_equiv_symm. + apply H. + + apply IHfss, H. +Qed. + +Import Lia. + +Lemma monarrlist_equiv_trans : + Transitive all_monarrlist_equiv. +Proof. + intros fs gs hs. + remember (length fs + length gs + length hs) as k eqn:Heqk. + assert (Hle : length fs + length gs + length hs <= k) by (subst; easy). + clear Heqk. + revert fs gs hs Hle. + induction k; + [intros [] [] []; easy|]. + intros [|f fs] [|g gs] [|h hs]; try reflexivity + (intros; assumption). + - simpl. + intros Hle Hnil Hgh. + split. + + eapply monarrlistelt_equiv_trans; [apply Hnil|]. + apply Hgh. + + apply (IHk [] gs hs); [simpl; lia|apply Hnil|apply Hgh]. + - intros Hle Hnil Hgh. + split. + + eapply monarrlistelt_equiv_trans; [apply Hnil|]. + apply Hgh. + + apply (IHk fs [] hs); [simpl in *; lia| |apply Hgh]. + apply (all_monarrlist_cons_equiv_nil Hnil). + - intros Hle Hnil Hgh. + split. + + eapply monarrlistelt_equiv_trans; [apply Hnil|]. + apply Hgh. + + pose proof (IHk fs gs [] ltac:(simpl in *;lia) (proj2 Hnil) + (all_monarrlist_cons_equiv_nil Hgh)) as e. + destruct fs; apply e. + - simpl. + intros Hle Hfg Hgh. + split. + + eapply monarrlistelt_equiv_trans; [apply Hfg|apply Hgh]. + + apply (IHk fs gs hs); [lia|..]. + * apply Hfg. + * apply Hgh. +Qed. + + +Lemma monarrlist_list_equiv_refl : + Reflexive all_monarrlist_list_equiv. +Proof. + intros fss. + induction fss; [easy|]. + split; [| apply IHfss]. + apply monarrlist_equiv_refl. +Qed. + + +Lemma monarrlist_list_equiv_symm : + Symmetric all_monarrlist_list_equiv. +Proof. + intros fss. + induction fss; intros gss H; [|destruct gss]. + - induction gss; [easy|]. + split. + + apply monarrlist_equiv_symm, H. + + specialize (IHgss (proj2 H)). + destruct gss; + apply IHgss. + - split. + + apply monarrlist_equiv_symm, H. + + apply (IHfss nil). + destruct fss; apply H. + - split. + + apply monarrlist_equiv_symm. + apply H. + + apply IHfss, H. +Qed. + +Lemma monarrlist_list_equiv_trans: + Transitive all_monarrlist_list_equiv. +Proof. + intros fss gss hss. + remember (length fss + length gss + length hss) as k eqn:Heqk. + assert (Hle : length fss + length gss + length hss <= k) by lia. + clear Heqk. + revert fss gss hss Hle. + induction k; + [intros [] [] []; easy|]. + intros [|fs fss] [|gs gss] [|hs hss]; try reflexivity + (intros; assumption). + - simpl. + intros Hle Hnil Hgh. + split. + + eapply monarrlist_equiv_trans; [apply Hnil|]. + apply Hgh. + + apply (IHk [] gss hss); [simpl; lia|apply Hnil|apply Hgh]. + - intros Hle Hnil Hgh. + split. + + eapply monarrlist_equiv_trans; [apply Hnil|]. + apply Hgh. + + apply (IHk fss [] hss); [simpl in *; lia| |apply Hgh]. + apply (all_monarrlist_list_cons_equiv_nil Hnil). + - intros Hle Hnil Hgh. + split. + + eapply monarrlist_equiv_trans; [apply Hnil|]. + apply Hgh. + + pose proof (IHk fss gss [] ltac:(simpl in *;lia) (proj2 Hnil) + (all_monarrlist_list_cons_equiv_nil Hgh)) as e. + unfold all_monarrlist_list_equiv in e. + rewrite zip_defaults_nil_r in e. + apply e. + - simpl. + intros Hle Hfg Hgh. + split. + + eapply monarrlist_equiv_trans; [apply Hfg|apply Hgh]. + + apply (IHk fss gss hss); [lia|..]. + * apply Hfg. + * apply Hgh. +Qed. + +Lemma compose_composable_monarrlist_list_equiv_nil {fss} + (H : all_monarrlist_list_equiv fss []) + (Hf : totally_composable fss) : + monarr_norm_equiv + (compose_totally_composable fss Hf) + (arrid e). +Proof. + induction fss; + [ apply monarr_norm_equiv_refl | destruct fss; + [apply (stack_monarrlist_equiv (proj1 H))|]]. + rewrite compose_totally_composable_cons. + unfold compose_composable_monarr. + rewrite <- (monarr_lunit (arrid e)). + apply monarr_norm_equiv_comp. + + eapply monarr_norm_equiv_trans; + [apply monarr_norm_equiv_struct_r|]. + apply (stack_monarrlist_equiv (proj1 H)). + + apply IHfss, H. +Qed. + +#[global] Add Parametric Relation : monarrlistelt monarrlistelt_equiv + reflexivity proved by monarrlistelt_equiv_refl + symmetry proved by monarrlistelt_equiv_symm + transitivity proved by monarrlistelt_equiv_trans + as monarrlistelt_equiv_equivalence. + +#[global] Add Parametric Relation : monarrlist all_monarrlist_equiv + reflexivity proved by monarrlist_equiv_refl + symmetry proved by monarrlist_equiv_symm + transitivity proved by monarrlist_equiv_trans + as monarrlist_equiv_equivalence. + +#[global] Add Parametric Relation : (list monarrlist) all_monarrlist_list_equiv + reflexivity proved by monarrlist_list_equiv_refl + symmetry proved by monarrlist_list_equiv_symm + transitivity proved by monarrlist_list_equiv_trans + as monarrlist_list_equiv_equivalence. + + + +Lemma all_monarrlist_list_equiv_iff_ex {fss gss} : + all_monarrlist_list_equiv fss gss <-> + exists fshd fsnil gshd gsnil, + fss = fshd ++ fsnil /\ gss = gshd ++ gsnil /\ + all_monarrlist_list_equiv fsnil [] /\ + all_monarrlist_list_equiv gsnil [] /\ + length fshd = length gshd /\ + all_monarrlist_list_equiv fshd gshd. +Proof. + split. + 2: { + intros (fshd & fsnil & gshd & gsnil & Hfss & Hgss & Hfnil & Hgnil & Hlen & Hfg). + subst. + revert gshd Hlen Hfg; + induction fshd; intros gshd Hlen Hfg. + - destruct gshd; [|easy]. + etransitivity; eauto using monarrlist_list_equiv_symm. + - destruct gshd; [easy|]. + simpl in Hlen. + injection Hlen; clear Hlen; intros Hlen. + specialize (IHfshd _ Hlen). + specialize (IHfshd (proj2 Hfg)). + split; [apply Hfg | apply IHfshd]. + } + revert gss; + induction fss as [|fs fss IHfss]; + intros gss; [|destruct gss as [|gs gss]]. + - intros Hgss. + exists [], [], [], gss. + repeat split; easy + symmetry; easy. + - intros Hfss. + exists [], (fs :: fss), [], []. + rewrite app_nil_r. + repeat easy + split. + - intros [Hfgs Hfgss]. + destruct (IHfss gss Hfgss) as + (fshd & fsnil & gshd & gsnil + & Hfss & Hgss & Hfnil & Hgnil & Hlen & Hequiv). + exists (fs :: fshd), fsnil, (gs :: gshd), gsnil. + (repeat first [easy | split]); simpl; subst; easy + f_equal; easy. +Qed. + +Lemma compose_composable_monarrlist_list_equiv {fss gss} + (H : all_monarrlist_list_equiv fss gss) + (Hf : totally_composable fss) + (Hg : totally_composable gss) : + monarr_norm_equiv + (compose_totally_composable fss Hf) + (compose_totally_composable gss Hg). +Proof. + remember (length fss + length gss) as k. + assert (Hk : length fss + length gss <= k) by (subst; easy). + clear Heqk. + revert fss gss H Hf Hg Hk. + induction k; intros fss gss; + destruct fss, gss; + try (solve [simpl; intros; easy]); + intros H Hf Hg Hk; + [..|destruct fss, gss]. + - apply monarr_norm_equiv_symm. + apply compose_composable_monarrlist_list_equiv_nil. + apply monarrlist_list_equiv_symm, H. + - apply compose_composable_monarrlist_list_equiv_nil, H. + - apply stack_monarrlist_equiv, H. + - rewrite <- monarr_runit. + rewrite compose_totally_composable_cons. + apply monarr_norm_equiv_comp. + + rewrite monarr_norm_equiv_struct_r'_iff. + apply stack_monarrlist_equiv, H. + + eapply monarr_norm_equiv_trans; + [|apply (IHk [] (m1::gss) (proj2 H) ltac:(easy) (proj2 Hg))]; + [|simpl in *; lia]. + apply monarr_norm_equiv_struct_eq_in. + simpl. + etransitivity; + [apply ( (Nf_eq_out_of_norm_equiv + (stack_monarrlist_equiv (proj1 H))))|]. + specialize (IHk [] (m1::gss) (proj2 H) ltac:(easy) + (proj2 Hg) ltac:(simpl in Hk |- *; lia)). + apply (Nf_eq_in_of_norm_equiv) in IHk. + simpl in IHk. + rewrite IHk. + apply Hg. + - apply monarr_norm_equiv_symm. + rewrite <- monarr_runit. + rewrite compose_totally_composable_cons. + apply monarr_norm_equiv_comp. + + eapply monarr_norm_equiv_symm, monarr_norm_equiv_trans; + [apply monarr_norm_equiv_struct_r|]; + apply stack_monarrlist_equiv, H. + + eapply monarr_norm_equiv_trans; + [|apply (IHk [] (m1::fss) + (monarrlist_list_equiv_symm (m1::fss) [] (proj2 H)) + ltac:(easy) (proj2 Hf) )]; + [|simpl in *; lia]. + apply monarr_norm_equiv_struct_eq_in. + simpl. + etransitivity; + [apply (eq_sym (Nf_eq_out_of_norm_equiv + (stack_monarrlist_equiv (proj1 H))))|]. + specialize (IHk (m1::fss) [] (proj2 H) (proj2 Hf) + ltac:(easy) ltac:(simpl in Hk |- *; lia)). + apply (Nf_eq_in_of_norm_equiv) in IHk. + simpl in IHk. + rewrite <- IHk. + apply Hf. + - rewrite 2!compose_totally_composable_cons. + apply monarr_norm_equiv_comp. + + eapply monarr_norm_equiv_trans; + [|apply monarr_norm_equiv_symm; + eapply monarr_norm_equiv_trans]; + [apply monarr_norm_equiv_struct_r.. |]. + apply monarr_norm_equiv_symm. + apply stack_monarrlist_equiv, H. + + apply (IHk (m1 :: fss) (m2 :: gss)); [|simpl in *; lia]. + apply H. +Qed. + +End monarrlist_equivalences. + +End monarrlist_theory. \ No newline at end of file diff --git a/ViCaR/MonoidalCoherence/MCprocessing.v b/ViCaR/MonoidalCoherence/MCprocessing.v new file mode 100644 index 0000000..57d2c8c --- /dev/null +++ b/ViCaR/MonoidalCoherence/MCprocessing.v @@ -0,0 +1,1896 @@ +Require Import Setoid. +Require MCproperop. + +Section ProperOperationInstances. + +Set Universe Polymorphism. + +Open Scope bw_scope. + +Import MCDefinitions MCClasses MC_setoids + MC_notations UIP_facts MCbw + MCconsequences MCmonarrlist MCproperop. +Import CategoryTypeclass. +Import List ListNotations. + +Context {X : Type} {UIPX : UIP X}. +Context {cC : Category X} {cCh : CategoryCoherence cC} + {mC : MonoidalCategory cC} {mCh : MonoidalCategoryCoherence mC}. + +Local Notation bw := (bw X). +Local Notation "a ⟶ b" := (@monarr X cC mC a b) (at level 60). +Local Notation monarrlist := (@monarrlist X cC mC). +Local Notation monarrlistelt := (@monarrlistelt X cC mC). + +Section proper_map_definitions. + +Fixpoint foliate_monarr {a b} (f : a ⟶ b) : list monarrlist := + match f with + | monarrcomp g h => foliate_monarr g ++ foliate_monarr h + | monarrtens g h => + zip_defaults (@app monarrlistelt) (foliate_monarr h) (foliate_monarr g) + [monarrlist_id (monarrlist_list_target (foliate_monarr h))] + [monarrlist_id (monarrlist_list_target (foliate_monarr g))] + | monarrstruct f => [[monarrlist_arr _ _ (monarrstruct f)]] + | mongeneric f => [[monarrlist_arr _ _ (mongeneric f)]] + end. + +Definition sum_to_bool {A B} (x : A + B) : bool := + match x with + | inl _ => true + | inr _ => false + end. + +Definition proj_to_monarrlist_list (f : (list monarrlist) * bw) : list monarrlist := + match fst f with + | [] => [[monarrlist_id (snd f)]] + | fs :: fss => fs :: fss + end. + +Definition trim_foliate_monarr_helper_compose (f g : (list monarrlist) * bw) : + (list monarrlist) * bw := + (fst f ++ fst g, snd g). + +Global Arguments trim_foliate_monarr_helper_compose !_ !_ /. + +Definition trim_foliate_monarr_helper_stack (f g : (list monarrlist) * bw) : + (list monarrlist) * bw := + (zip_defaults (@app monarrlistelt) (fst g) (fst f) + [monarrlist_id (snd g)] + [monarrlist_id (snd f)], tens (snd f) (snd g)). + +Global Arguments trim_foliate_monarr_helper_stack !_ !_ /. + +Fixpoint trim_foliate_monarr_helper {a b} (f : a ⟶ b) : + (list monarrlist) * bw := + (* inl lst is the usual output, + inr a means a structural arrow starting at a *) + match f with + | monarrcomp g h => + trim_foliate_monarr_helper_compose + (trim_foliate_monarr_helper g) + (trim_foliate_monarr_helper h) + | monarrtens g h => + trim_foliate_monarr_helper_stack + (trim_foliate_monarr_helper g) + (trim_foliate_monarr_helper h) + | @monarrstruct _ _ _ a b f => ([], b) + | @mongeneric _ _ _ a b f => ([[monarrlist_arr _ _ (mongeneric f)]], b) + end. + +Definition trim_foliate_monarr {a b} (f : a ⟶ b) : list monarrlist := + proj_to_monarrlist_list (trim_foliate_monarr_helper f). + +Fixpoint is_empty (a : bw) : bool := + match a with + | e => true + | var _ => false + | tens a' b' => is_empty a' && is_empty b' + end. + +Fixpoint remove_empty_structs (fs : monarrlist) : monarrlist := + match fs with + | nil => nil + | monarrlist_id a :: fs' => + (if is_empty a then [] else [monarrlist_id a]) + ++ remove_empty_structs fs' + | f :: fs' => f :: remove_empty_structs fs' + end. + +Definition trim_foliate_monarr_no_empty {a b} (f : a ⟶ b) : list monarrlist := + map remove_empty_structs (trim_foliate_monarr f). +(* +Definition proj_to_monarrlist_list (f : (list monarrlist) * bw) : list monarrlist := + match fst f with + | inl fs => fs + | inr a => [[monarrlist_id a]] + end. + +Definition trim_foliate_monarr_helper_compose (f g : (list monarrlist) + bw) : + (list monarrlist) + bw := + match f, g with + | inl fs, inl gs => inl (fs ++ gs) + | inl fs, inr b => inl fs + | inr a, inl gs => inl gs + | inr a, inr b => inr a + end. + +Definition trim_foliate_monarr_helper_stack (f g : (list monarrlist) + bw) : + (list monarrlist) + bw := + match f, g with + | inl fs, inl gs => inl ( + zip_defaults (@app monarrlistelt) gs fs + [monarrlist_id (monarrlist_list_target gs)] + [monarrlist_id (monarrlist_list_target fs)]) + | inl fs, inr b => inl ( + zip_with_default_l (@app monarrlistelt) + fs [monarrlist_id b]) + | inr a, inl gs => inl ( + zip_with_default_r (@app monarrlistelt) + [monarrlist_id a] gs) + | inr a, inr b => inr (tens a b) + end. + +Fixpoint trim_foliate_monarr_helper {a b} (f : a ⟶ b) : + (list monarrlist) + bw := + (* inl lst is the usual output, + inr a means a structural arrow starting at a *) + match f with + | monarrcomp g h => + trim_foliate_monarr_helper_compose + (trim_foliate_monarr_helper g) + (trim_foliate_monarr_helper h) + | monarrtens g h => + trim_foliate_monarr_helper_stack + (trim_foliate_monarr_helper g) + (trim_foliate_monarr_helper h) + | @monarrstruct _ _ _ a b f => inr a + | mongeneric f => inl [[monarrlist_arr _ _ (mongeneric f)]] + end. + +Definition trim_foliate_monarr {a b} (f : a ⟶ b) : list monarrlist := + proj_to_monarrlist_list (trim_foliate_monarr_helper f). +*) + +(* +Fixpoint foliate_monarr_alt {a b} (f : a ⟶ b) : list monarrlist := + match f with + | monarrcomp g h => foliate_monarr g ++ foliate_monarr h + | monarrtens g h => + zip_defaults (@app monarrlistelt) (foliate_monarr h) (foliate_monarr g) + (monarrlist_id (monarrlist_list_target (foliate_monarr h)) :: nil) + (monarrlist_id (monarrlist_list_target (foliate_monarr g)) :: nil) + | monarrstruct f => (monarrlist_arr _ _ (monarrstruct f) :: nil) :: nil + | mongeneric f => (monarrlist_arr _ _ (mongeneric f) :: nil) :: nil + end + (* with foliate_monarr_alt_def {a b} (f : a ⟶ b) : list monarrlist *) + . *) + +Definition is_idarr (fs : monarrlistelt) := + match fs with + | monarrlist_id _ => true + | monarrlist_arr _ _ _ => false + end. + +Definition is_idarrlist (fs : monarrlist) := + forallb is_idarr fs. + +Definition idarr_of_vars (xs : list X) : monarrlist := + map (fun x => monarrlist_id (var x)) xs. + +Definition shift_left (fs gs : monarrlist) := + if is_idarrlist fs then + (gs, idarr_of_vars (target_vars gs)) + else (fs, gs). + +Definition shift_right (fs gs : monarrlist) := + if is_idarrlist gs then + (idarr_of_vars (source_vars fs), fs) + else (fs, gs). + +Fixpoint is_struct {a b} (f : a ⟶ b) : bool := + match f with + | mongeneric _ => false + | monarrstruct _ => true + | monarrcomp f' g' => is_struct f' && is_struct g' + | monarrtens f' g' => is_struct f' && is_struct g' + end. + +Fixpoint struct_of_is_struct {a b} (f : a ⟶ b) + (H : is_struct f = true) : bwarr a b. + destruct f. + - refine (struct_of_is_struct _ _ _ _ ○ struct_of_is_struct _ _ _ _). + apply (proj1 (proj1 (Bool.andb_true_iff _ _) H)). + apply (proj2 (proj1 (Bool.andb_true_iff _ _) H)). + - refine (struct_of_is_struct _ _ _ _ ⊠ struct_of_is_struct _ _ _ _). + apply (proj1 (proj1 (Bool.andb_true_iff _ _) H)). + apply (proj2 (proj1 (Bool.andb_true_iff _ _) H)). + - discriminate H. + - apply f. +Defined. + +Fixpoint structify {a b} (f : a ⟶ b) : a ⟶ b := + match f with + | mongeneric f' => mongeneric f' + | monarrstruct f' => monarrstruct f' + | monarrcomp f' g' => + let fs := structify f' in + let gs := structify g' in + match Bool.bool_dec (is_struct fs) true with + | right _ => monarrcomp fs gs + | left Hfstruct => + match Bool.bool_dec (is_struct gs) true with + | right _ => monarrcomp fs gs + | left Hgstruct => + arrcomp + (struct_of_is_struct fs Hfstruct) + (struct_of_is_struct gs Hgstruct) + end + end + | monarrtens f' g' => + let fs := structify f' in + let gs := structify g' in + match Bool.bool_dec (is_struct fs) true with + | right _ => monarrtens fs gs + | left Hfstruct => + match Bool.bool_dec (is_struct gs) true with + | right _ => monarrtens fs gs + | left Hgstruct => + arrtens + (struct_of_is_struct fs Hfstruct) + (struct_of_is_struct gs Hgstruct) + end + end + end. + +Definition remove_struct (g : monarrlistelt) := + match g with + | monarrlist_id a => monarrlist_id a + | monarrlist_arr a b f => + match f with + | monarrstruct _ => monarrlist_id a + | f' => monarrlist_arr a b f' + end + end. + +Definition remove_structs := map remove_struct. + +Fixpoint struct_to_id (fss : monarrlist) := + match fss with + | nil => nil + | fs :: fss' => match fs with + | monarrlist_id a => monarrlist_id a + | monarrlist_arr a b (monarrstruct f) => monarrlist_id a + | monarrlist_arr a b f => monarrlist_arr a b f + end :: struct_to_id fss' + end. + +Fixpoint split_ids (fss : monarrlist) := + match fss with + | nil => nil + | fs :: fss' => match fs with + | monarrlist_id a => map monarrlist_id (map var (bw_to_varlist a)) + | monarrlist_arr a b f => monarrlist_arr a b f :: nil + end ++ split_ids fss' + end. + +Fixpoint combine_ids_helper (a : bw) (fss : monarrlist) := + match fss with + | nil => monarrlist_id a :: nil + | monarrlist_id b :: fss' => combine_ids_helper (tens b a) fss' + | fs :: fss' => monarrlist_id a :: fs :: combine_ids fss' + end +with combine_ids (fss : monarrlist) := + match fss with + | nil => nil + | monarrlist_id a :: fss' => combine_ids_helper a fss' + | fs :: fss' => fs :: combine_ids fss' + end. + +Definition is_idarr_singleton (fs : monarrlist) := + match fs with + | monarrlist_id _ :: nil => true + | _ => false + end. + +Definition filter_idarrs (fss : list monarrlist) := + match fss with + | nil => nil + | fs :: fss' => + fs :: filter (fun x => negb (is_idarrlist x)) fss' + (* if is_idarrlist fs then + (if length (filter_idarrs fss') > 0 then filter_idarrs fss' else fs :: nil) + else fs :: (filter_idarrs fss') *) + end. + +Fixpoint trim_idarrs_helper (fs : monarrlist) (fss : list monarrlist) := + match fss with + | nil => [fs] + | fs' :: fss' => if is_idarrlist fs' + then trim_idarrs_helper fs fss' + else fs :: fss + end. + +Definition trim_idarrs (fss : list monarrlist) := + match fss with + | nil => nil + | fs :: fss' => trim_idarrs_helper fs fss' + end. + +Definition trim_idarr (fss : list monarrlist) := + match fss with + | nil => nil + | fs :: nil => [fs] + | fs :: fs' :: fss' => + if is_idarrlist fs then fs' :: fss' else fs :: fs' :: fss' + end. + +Definition structify_monarrlistelt (f : monarrlistelt) := + monarrlistelt_map (@structify) f. + +Definition structify_monarrlist (fs : monarrlist) := + map structify_monarrlistelt fs. + +Local Notation "g |> f" := (Basics.compose f g) + (at level 40, left associativity). + +Definition full_right_shift (fss : list monarrlist) := + Nat.iter (pred (length fss)) ( + pairwise_apply (pairify_apply shift_right) |> + map split_ids + ) fss. + +Definition full_process_monarrlist_list := + map structify_monarrlist |> + map remove_structs |> + map split_ids |> + full_right_shift |> + trim_idarrs |> + trim_idarr. + +End proper_map_definitions. + +Section foliation. + +Section zip_with_defaults. + +Lemma tens_comp_composable_split_bot_r {a b c d x y} (f : a ⟶ b) (g : c ⟶ d) + (H : Nf b = Nf c) (h : x ⟶ y) : + h ⧆ compose_composable_monarr f g H ≊ + compose_composable_monarr (h ⧆ f) (arrid y ⧆ g) + (Nf_tens_eq eq_refl H). +Proof. + unfold compose_composable_monarr. + rewrite !monarr_tens_comp_split_bot_r. + rewrite monarr_arrtens. + eauto with monarrdb. +Qed. + +Lemma tens_comp_composable_split_bot_l {a b c d x y} (f : a ⟶ b) (g : c ⟶ d) + (H : Nf b = Nf c) (h : x ⟶ y) : + h ⧆ compose_composable_monarr f g H ≊ + compose_composable_monarr (arrid x ⧆ f) (h ⧆ g) + (Nf_tens_eq eq_refl H). +Proof. + unfold compose_composable_monarr. + rewrite !monarr_tens_comp_split_bot_l. + rewrite monarr_arrtens. + eauto with monarrdb. +Qed. + +Lemma tens_comp_composable_split_top_pf {a b} (H : Nf a = Nf b) (x : bw) : + Nf (a ⨂ x) = Nf (b ⨂ x). +Proof. + cbn. + f_equal. + apply H. +Qed. + +Lemma tens_comp_composable_split_top_l {a b c d x y} (f : a ⟶ b) (g : c ⟶ d) + (H : Nf b = Nf c) (h : x ⟶ y): + compose_composable_monarr f g H ⧆ h ≊ + compose_composable_monarr (f ⧆ arrid x) (g ⧆ h) + (Nf_tens_eq H eq_refl). +Proof. + unfold compose_composable_monarr. + rewrite !monarr_tens_comp_split_top_l. + rewrite monarr_arrtens. + eauto with monarrdb. +Qed. + +Lemma tens_comp_composable_split_top_r {a b c d x y} (f : a ⟶ b) (g : c ⟶ d) + (H : Nf b = Nf c) (h : x ⟶ y): + compose_composable_monarr f g H ⧆ h ≊ + compose_composable_monarr (f ⧆ h) (g ⧆ arrid y) + (Nf_tens_eq H eq_refl). +Proof. + unfold compose_composable_monarr. + rewrite !monarr_tens_comp_split_top_r. + rewrite monarr_arrtens. + eauto with monarrdb. +Qed. + +Lemma monarr_tens_comp_composable {a b c d m n o p} (f : a ⟶ b) (g : c ⟶ d) + (h : m ⟶ n) (k : o ⟶ p) (H : Nf b = Nf c) (H' : Nf n = Nf o) : + compose_composable_monarr f g H ⧆ compose_composable_monarr h k H' ≊ + compose_composable_monarr (f ⧆ h) (g ⧆ k) (Nf_tens_eq H H'). +Proof. + unfold compose_composable_monarr. + rewrite !monarr_tens_comp, monarr_arrtens. + eauto with monarrdb. +Qed. + +Lemma zip_with_default_l_norm_equiv_nonempty {fss} (Hne : fss <> []) + (Hf : totally_composable fss) a : + monarr_norm_equiv (compose_totally_composable + _ (zip_with_default_l_totally_composable Hf a)) + (arrid (e⨂a) ⧆ compose_totally_composable fss Hf). +Proof. + induction fss; [|destruct fss]; [easy|..]. + - simpl. + apply (monarr_norm_equiv_trans (stack_monarrlist_app' _ _)). + simpl. + apply monarrequiv_iff_monarr_norm_equiv. + apply monarr_tens; [|easy]. + rewrite monarr_arrtens. + apply monarr_struct. + - simpl zip_with_default_l in *. + rewrite 2!compose_totally_composable_cons. + (* erewrite compose_totally_composable_indep. *) + eapply monarr_norm_equiv_trans. + 1:{ + eapply monarr_norm_equiv_comp_composable; + [apply stack_monarrlist_app'|]. + erewrite compose_totally_composable_indep. + refine (IHfss _ (proj2 Hf)). + easy. + } + rewrite tens_comp_composable_split_bot_l. + unshelve (instantiate (1:=_)); + [rewrite 2!Nf_tens_bwnormapp; f_equal; apply Hf|]. + apply monarr_norm_equiv_comp_composable; [|easy]. + apply monarr_norm_equiv_tens; [|easy]. + cbn. + now rewrite monarr_arrtens, monarr_struct_id. +Qed. + +Lemma zip_with_default_l_equiv_nonempty {fss} (Hne : fss <> []) + (Hf : totally_composable fss) a : + compose_totally_composable + _ (zip_with_default_l_totally_composable Hf a) + ≊ bwarr_of_Nf_eq (zip_with_default_l_source_nonempty Hne [monarrlist_id a]) + ◌ arrid (e⨂a) ⧆ compose_totally_composable fss Hf + ◌ bwarr_of_Nf_eq (eq_sym + (zip_with_default_l_target_nonempty Hne [monarrlist_id a])). +Proof. + apply equiv_conj_of_monarr_norm_equiv. + apply zip_with_default_l_norm_equiv_nonempty, Hne. +Qed. + +Lemma zip_with_default_r_norm_equiv_nonempty {fss} (Hne : fss <> []) + (Hf : totally_composable fss) a : + monarr_norm_equiv (compose_totally_composable + _ (zip_with_default_r_totally_composable Hf a)) + (compose_totally_composable fss Hf ⧆ arrid (e ⨂ a)). +Proof. + induction fss; [easy|destruct fss]. + - simpl. + exists eq_refl. + exists eq_refl. + cbn. + rewrite monarr_comp_tens_id_top_struct_l, monarr_comp_tens_id_top_struct_r. + apply monarr_tens; [easy|]. + rewrite !monarr_arrcomp; apply monarr_struct. + - simpl zip_with_default_r in *. + rewrite 2!compose_totally_composable_cons. + (* erewrite compose_totally_composable_indep. *) + eapply monarr_norm_equiv_trans. + 1:{ + eapply monarr_norm_equiv_comp_composable; + [apply (stack_monarrlist_app' [_] _)|]. + erewrite compose_totally_composable_indep. + refine (IHfss _ (proj2 Hf)). + easy. + } + rewrite tens_comp_composable_split_top_l. + unshelve (instantiate (1:=_)); + [rewrite 2!Nf_tens_bwnormapp; f_equal; apply Hf|]. + apply monarr_norm_equiv_comp_composable; [|easy]. + apply monarr_norm_equiv_tens; [easy|]. + cbn. + now rewrite monarr_arrtens, monarr_struct_id. +Qed. + +Lemma zip_with_default_r_equiv_nonempty {fss} (Hne : fss <> []) + (Hf : totally_composable fss) a : + compose_totally_composable + _ (zip_with_default_r_totally_composable Hf a) + ≊ bwarr_of_Nf_eq (zip_with_default_r_source_nonempty Hne [monarrlist_id a]) + ◌ compose_totally_composable fss Hf ⧆ arrid (e ⨂ a) + ◌ bwarr_of_Nf_eq (eq_sym + (zip_with_default_r_target_nonempty Hne [monarrlist_id a])). +Proof. + apply equiv_conj_of_monarr_norm_equiv, + zip_with_default_r_norm_equiv_nonempty, Hne. +Qed. + +Lemma zip_defaults_totally_composable' {fss gss df dg} + (Hf : totally_composable fss) + (Hg : totally_composable gss) + (Hdf : Nf df = Nf (monarrlist_list_target fss)) + (Hdg : Nf dg = Nf (monarrlist_list_target gss)) : + totally_composable + (zip_defaults (@app monarrlistelt) + fss gss + (monarrlist_id df :: nil) + (monarrlist_id dg :: nil)). +Proof. + revert gss Hg Hdg; + induction fss; intros gss Hg Hdg; + [|revert a Hf Hdf; induction gss as [|l gss IHgss]; + intros a Hf Hdf; [|destruct fss; [|destruct gss]]]; + [..|destruct fss, gss]. + - apply zip_with_default_r_totally_composable, Hg. + - apply zip_with_default_l_totally_composable, Hf. + - destruct gss; [easy|]. + split. + + apply composable_app; [easy | apply Hg]. + + apply (zip_with_default_r_totally_composable (fss:=m :: gss) (proj2 Hg)). + - split. + + apply composable_app; [apply Hf | easy]. + + apply (zip_with_default_l_totally_composable (fss:=m :: fss) (proj2 Hf)). + - split; [|easy]. apply composable_app; [apply Hf | apply Hg]. + - split; [apply composable_app; [apply Hf | apply Hg]|]. + split. + + apply composable_app; [|apply Hg]. + unfold composable. + cbn. + change (⟦ df ⟧ norm_e) with (Nf df). + rewrite Hdf; easy. + + apply (zip_with_default_r_totally_composable + (fss:=m1 :: gss) (proj2 (proj2 Hg))). + - split; [apply composable_app; [apply Hf | apply Hg]|]. + split. + + apply composable_app; [apply Hf|]. + unfold composable. + cbn. + change (⟦ dg ⟧ norm_e) with (Nf dg). + rewrite Hdg; easy. + + apply (zip_with_default_l_totally_composable + (fss:=m1 :: fss) (proj2 (proj2 Hf))). + - split; [apply composable_app; [apply Hf | apply Hg]|]. + apply (IHfss (proj2 Hf) Hdf (m0 :: m2 :: gss) (proj2 Hg) Hdg). +Qed. + +Lemma zip_defaults_totally_composable {fss gss} + (Hf : totally_composable fss) + (Hg : totally_composable gss) : + totally_composable + (zip_defaults (@app monarrlistelt) + fss gss + (monarrlist_id (monarrlist_list_target fss) :: nil) + (monarrlist_id (monarrlist_list_target gss) :: nil)). +Proof. + revert gss Hg; + induction fss; intros gss Hg; [|revert a Hf; induction gss as [|l gss IHgss]; + intros a Hf; [|destruct fss; [|destruct gss]]]. + - apply zip_with_default_r_totally_composable, Hg. + - apply zip_with_default_l_totally_composable, Hf. + - destruct gss; [easy|]. + split. + + apply composable_app; [easy | apply Hg]. + + apply (zip_with_default_r_totally_composable (fss:=m :: gss) (proj2 Hg)). + - split. + + apply composable_app; [apply Hf | easy]. + + apply (zip_with_default_l_totally_composable (fss:=m :: fss) (proj2 Hf)). + - split. + + apply composable_app; [apply Hf | apply Hg]. + + rewrite 2!monarrlist_list_target_cons_cons. + apply (IHfss (proj2 Hf) (m0 :: gss)), Hg. +Qed. + +Lemma tens_zip_defaults_correct_helper'' {fss gss} + (Hnf : fss <> []) (Hng : gss <> []) + (Hf : totally_composable fss) (Hg : totally_composable gss) {df dg} + (Hdf : Nf df = _) (Hdg : Nf dg = _) : + monarr_norm_equiv (compose_totally_composable + _ (zip_defaults_totally_composable' Hf Hg Hdf Hdg)) + (compose_totally_composable gss Hg ⧆ compose_totally_composable fss Hf). +Proof. + revert gss Hg Hng Hdg; + induction fss; + intros gss Hg Hng Hdg; [easy|destruct gss; [easy|]]. + destruct fss, gss(* ; [..|destruct fss, gss] *). + - apply stack_monarrlist_app'. + - simpl zip_defaults. + rewrite 2!compose_totally_composable_cons. + rewrite tens_comp_composable_split_top_r. + apply monarr_norm_equiv_comp_composable; + [apply stack_monarrlist_app'|]. + pose proof (@zip_with_default_r_norm_equiv_nonempty (m0 :: gss) + (fun x => nil_cons (eq_sym x)) (proj2 Hg) df) as H. + simpl zip_with_default_r in H. + erewrite compose_totally_composable_indep. + apply (monarr_norm_equiv_trans H). + apply monarr_norm_equiv_tens; [easy|]. + apply monarr_norm_equiv_struct_eq_in. + easy. + - simpl zip_defaults. + rewrite 2!compose_totally_composable_cons. + rewrite tens_comp_composable_split_bot_r. + apply monarr_norm_equiv_comp_composable; + [apply stack_monarrlist_app'|]. + pose proof (@zip_with_default_l_norm_equiv_nonempty (m0 :: fss) + (fun x => nil_cons (eq_sym x)) (proj2 Hf) dg) as H. + simpl zip_with_default_l in H. + erewrite compose_totally_composable_indep. + apply (monarr_norm_equiv_trans H). + apply monarr_norm_equiv_tens; [|easy]. + apply monarr_norm_equiv_struct_eq_in. + easy. + - simpl zip_defaults. + rewrite 3!compose_totally_composable_cons. + rewrite monarr_tens_comp_composable. + apply monarr_norm_equiv_comp_composable; + [apply stack_monarrlist_app'|]. + (* erewrite (compose_totally_composable_indep (_::_) (proj2 Hg)). + erewrite (compose_totally_composable_indep (_::_) (proj2 Hf)). *) + refine (monarr_norm_equiv_trans _ (IHfss ltac:(easy) + (proj2 Hf) Hdf (m1::gss) (proj2 Hg) ltac:(easy) Hdg)). + erewrite compose_totally_composable_indep. + easy. +Qed. + +Lemma tens_zip_defaults_correct_helper' {fss gss} + (Hnf : fss <> []) (Hng : gss <> []) + (Hf : totally_composable fss) (Hg : totally_composable gss) : + monarr_norm_equiv (compose_totally_composable + _ (zip_defaults_totally_composable Hf Hg)) + (compose_totally_composable gss Hg ⧆ compose_totally_composable fss Hf). +Proof. + erewrite compose_totally_composable_indep. + apply (tens_zip_defaults_correct_helper'' Hnf Hng Hf Hg eq_refl eq_refl). +Qed. + +Lemma tens_zip_defaults_correct_helper {fss gss} (Hnf : fss <> []) (Hng : gss <> []) + (Hf : totally_composable fss) (Hg : totally_composable gss) : + compose_totally_composable + _ (zip_defaults_totally_composable Hf Hg) + ≊ bwarr_of_Nf_eq (zip_defaults_source_nonempty Hnf Hng) + ◌ compose_totally_composable gss Hg ⧆ compose_totally_composable fss Hf + ◌ bwarr_of_Nf_eq (eq_sym (zip_defaults_target_nonempty Hnf Hng)). +Proof. + apply equiv_conj_of_monarr_norm_equiv, + tens_zip_defaults_correct_helper'; easy. +Qed. + +Lemma tens_zip_defaults_correct' {fss gss} (Hnf : fss <> []) (Hng : gss <> []) + (Hf : totally_composable fss) (Hg : totally_composable gss) + Hcomp : + monarr_norm_equiv (compose_totally_composable + (zip_defaults (app (A:=monarrlistelt)) fss gss + [monarrlist_id (monarrlist_list_target fss)] + [monarrlist_id (monarrlist_list_target gss)]) + Hcomp) + (compose_totally_composable gss Hg ⧆ compose_totally_composable fss Hf). +Proof. + erewrite compose_totally_composable_indep. + apply tens_zip_defaults_correct_helper'; easy. +Qed. + +Lemma tens_zip_defaults_correct {fss gss} (Hnf : fss <> []) (Hng : gss <> []) + (Hf : totally_composable fss) (Hg : totally_composable gss) + Hcomp : + compose_totally_composable + (zip_defaults (app (A:=monarrlistelt)) fss gss + [monarrlist_id (monarrlist_list_target fss)] + [monarrlist_id (monarrlist_list_target gss)]) + Hcomp + ≊ bwarr_of_Nf_eq (zip_defaults_source_nonempty Hnf Hng) + ◌ compose_totally_composable gss Hg ⧆ compose_totally_composable fss Hf + ◌ bwarr_of_Nf_eq (eq_sym (zip_defaults_target_nonempty Hnf Hng)). +Proof. + erewrite compose_totally_composable_indep. + apply tens_zip_defaults_correct_helper. +Qed. + +End zip_with_defaults. + +Lemma foliate_monarr_not_nil {a b} (f : a ⟶ b) : + foliate_monarr f <> []. +Proof. + induction f; try easy; simpl; + destruct (foliate_monarr f1); + destruct (foliate_monarr f2); easy. +Qed. + +Lemma monarrlist_list_source_foliate {a b} (f : a ⟶ b) : + Nf (monarrlist_list_source (foliate_monarr f)) = Nf a. +Proof. + induction f; simpl; try easy. + - rewrite monarrlist_list_source_app_not_nil by apply foliate_monarr_not_nil. + easy. + - pose proof (foliate_monarr_not_nil f1). + pose proof (foliate_monarr_not_nil f2). + destruct (foliate_monarr f1); + destruct (foliate_monarr f2); try easy. + simpl in *. + rewrite Nf_monarrlist_source_app, Nf_tens_bwnormapp. + f_equal; easy. +Qed. + +Lemma monarrlist_list_target_foliate_helper fss gss : + Nf (monarrlist_list_target + (zip_defaults (app (A:=monarrlistelt)) gss fss + [monarrlist_id (monarrlist_list_target gss)] + [monarrlist_id (monarrlist_list_target fss)])) = + bwnormapp + (Nf (monarrlist_list_target fss)) (Nf (monarrlist_list_target gss)). +Proof. + revert fss; + induction gss; intro fss. + - simpl. + induction fss; [easy|]. + simpl. + destruct fss; [easy|]. + apply IHfss. + - destruct gss. + + induction fss. + * simpl. + rewrite Nf_monarrlist_target_app. + easy. + * simpl. + destruct fss; + [apply Nf_monarrlist_target_app|]. + simpl (monarrlist_list_target [a]) in IHfss. + rewrite <- IHfss. + simpl. + destruct (zip_with_default_r (app (A:=monarrlistelt)) [monarrlist_id (monarrlist_target a)] fss). + change (?x :: l) with ((x::nil) ++ l). + rewrite 2!Nf_monarrlist_target_app. + easy. + easy. + + destruct fss; + [apply (IHgss nil)|]. + rewrite zip_defaults_cons. + change (monarrlist_list_target (?a :: ?l :: ?gss)) + with (monarrlist_list_target (l :: gss)). + destruct fss. + simpl zip_defaults. + simpl monarrlist_list_target. + destruct gss. + simpl. + rewrite Nf_monarrlist_target_app; easy. + destruct gss. + simpl. + rewrite Nf_monarrlist_target_app; easy. + do 2 change (monarrlist_list_target (?a :: ?l :: ?gss)) + with (monarrlist_list_target (l :: gss)) in *. + (* rewrite 2!monarrlist_list_target_cons_cons in *. *) + specialize (IHgss (l0::nil)). + simpl (monarrlist_list_target (l0 :: nil)) in IHgss. + rewrite <- IHgss. + easy. + change (monarrlist_list_target (?a :: ?l :: ?gss)) + with (monarrlist_list_target (l :: gss)). + rewrite <- IHgss. + easy. +Qed. + +Lemma monarrlist_list_target_foliate {a b} (f : a ⟶ b) : + Nf (monarrlist_list_target (foliate_monarr f)) = Nf b. +Proof. + induction f; simpl; try easy. + - rewrite monarrlist_list_target_app_not_nil by apply foliate_monarr_not_nil. + easy. + - rewrite monarrlist_list_target_foliate_helper. + rewrite Nf_tens_bwnormapp. + f_equal; easy. +Qed. + +Lemma foliate_monarr_totally_composable {a b} (f : a ⟶ b) : + totally_composable (foliate_monarr f). +Proof. + induction f; try easy. + simpl. + apply app_composable_of_composable_mid; try easy. + rewrite monarrlist_list_target_foliate, monarrlist_list_source_foliate. + easy. + simpl. + revert IHf1 IHf2. + generalize (foliate_monarr f2). + generalize (foliate_monarr f1). + clear f1 f2. + intros fss. + induction fss; intros gss; [|destruct fss]; intros Hf Hg. + - induction gss; [easy|]. + simpl. + destruct gss; [easy|]. + split; [|apply IHgss, Hg]. + apply composable_app; + [apply Hg|]. + easy. + - induction gss; [easy|]. + simpl. + destruct gss; [easy|]. + split; + [apply composable_app; [apply Hg| easy]|]. + simpl in IHgss. + destruct gss; try (apply IHgss); try assumption + easy. + split; [|apply IHgss]. + apply composable_app. + apply Hg. + easy. + intros. + apply Hg. + - induction gss; [|destruct gss]. + + clear IHfss. + split; [apply composable_app; [easy|apply Hf]|]. + apply (totally_composable_app_restrict_r (fsts:=a0::nil)) in Hf. + revert m Hf. + induction fss; + intros l Hf; [easy|]. + split; [apply composable_app; [easy|apply Hf]|]. + apply IHfss, Hf. + + clear IHfss IHgss. + split; [apply composable_app; [easy|apply Hf]|]. + apply (totally_composable_app_restrict_r (fsts:=a0::nil)) in Hf. + revert m Hf. + induction fss; + split; [apply composable_app; [easy|apply Hf]|]. + apply IHfss, Hf. + + split; [apply composable_app; [apply Hg|apply Hf]|]. + rewrite 2!monarrlist_list_target_cons_cons. + apply (IHfss (m0 :: gss)); + [apply Hf | apply Hg]. +Qed. + +Lemma foliate_correct' {a b} (f : a ⟶ b) : + monarr_norm_equiv f + (compose_totally_composable (foliate_monarr f) + (foliate_monarr_totally_composable f)). +Proof. + induction f. + - simpl. + rewrite (compose_totally_composable_app_nonempty + (fss:=foliate_monarr f1) (fss':=foliate_monarr f2) + (foliate_monarr_totally_composable (monarrcomp f1 f2)) + (foliate_monarr_not_nil f1) + (foliate_monarr_not_nil f2)). + rewrite monarr_norm_equiv_conj_struct'_iff. + apply monarr_norm_equiv_comp; + [|erewrite (compose_totally_composable_indep (_ _)); + eassumption]. + rewrite monarr_norm_equiv_struct_r'_iff. + erewrite (compose_totally_composable_indep (_ _)); + eassumption. + - simpl. + rewrite (tens_zip_defaults_correct + (foliate_monarr_not_nil f2) (foliate_monarr_not_nil f1) + (foliate_monarr_totally_composable f2) + (foliate_monarr_totally_composable f1)). + rewrite monarr_norm_equiv_conj_struct'_iff. + apply monarr_norm_equiv_tens; easy. + - apply monarr_norm_equiv_symm. + apply monarr_norm_equiv_tens_bwarr_e_l. + - apply monarr_norm_equiv_symm. + apply monarr_norm_equiv_tens_bwarr_e_l. +Qed. + + +Lemma foliate_correct {a b} (f : a ⟶ b) : + f ≊ + bwarr_of_Nf_eq (eq_sym (monarrlist_list_source_foliate f)) ◌ + compose_totally_composable (foliate_monarr f) + (foliate_monarr_totally_composable f) ◌ + bwarr_of_Nf_eq (monarrlist_list_target_foliate f). +Proof. + apply equiv_conj_of_monarr_norm_equiv. + apply foliate_correct'. +Qed. + +Lemma compose_equiv_of_eq {fss gss} (H : fss = gss) + (Hf : totally_composable fss) (Hg : totally_composable gss) : + compose_totally_composable _ Hf ≊ + cast_bwarr eq_refl (f_equal monarrlist_list_source H) (arrid _) ◌ + compose_totally_composable _ Hg ◌ + cast_bwarr eq_refl (f_equal monarrlist_list_target (eq_sym H)) (arrid _). +Proof. + subst. + rewrite monarr_struct_id, monarr_lunit. + rewrite monarr_struct_id, monarr_runit. + erewrite compose_totally_composable_indep. + easy. +Qed. + +Lemma equiv_of_foliate_eq {a b} (f g : a ⟶ b) : + foliate_monarr f = foliate_monarr g -> + f ≊ g. +Proof. + intros Heq. + rewrite foliate_correct, (foliate_correct g). + rewrite monarrcomp_struct_r, monarrcomp_struct_l. + rewrite <- monarr_assoc, monarr_arrcomp. + rewrite !monarr_assoc, monarr_arrcomp. + erewrite (compose_equiv_of_eq Heq _ (foliate_monarr_totally_composable g)). + apply monarr_comp; [apply monarr_comp|]; + easy + apply monarr_struct. +Qed. + +Lemma equiv_of_foliate_equiv {a b} (f g : a ⟶ b) : + monarr_norm_equiv + (compose_totally_composable + (foliate_monarr f) + (foliate_monarr_totally_composable f)) + (compose_totally_composable + (foliate_monarr g) + (foliate_monarr_totally_composable g)) -> + f ≊ g. +Proof. + intros Heq. + rewrite foliate_correct, (foliate_correct g). + apply monarrequiv_iff_monarr_norm_equiv. + apply monarr_norm_equiv_conj_struct'_iff, + monarr_norm_equiv_conj_struct_iff. + apply Heq. +Qed. + +Lemma equiv_of_foliate_all_equiv {a b} (f g : a ⟶ b) : + all_monarrlist_list_equiv (foliate_monarr f) (foliate_monarr g) -> + f ≊ g. +Proof. + intros Heq. + apply equiv_of_foliate_equiv. + apply compose_composable_monarrlist_list_equiv, Heq. +Qed. + +End foliation. + +Section trim_foliation. + +Lemma totally_composable_of_middle_arr {fss gss} + (Hf : totally_composable fss) (Hg : totally_composable gss) + (Hmid : Nf (monarrlist_list_target fss) = Nf (monarrlist_list_source gss)) : + totally_composable (fss ++ gss). +Proof. + revert gss Hg Hmid; induction fss; intros gss Hg Hmid; [easy|]. + destruct fss. + - destruct gss; [easy|]. + split; + [apply Hmid | apply Hg]. + - split; [apply Hf|]. + apply IHfss; apply Hf + easy. +Qed. + +Lemma trim_foliate_spec {a b} (f : a ⟶ b) : + snd (trim_foliate_monarr_helper f) = b /\ + Nf (monarrlist_list_target (trim_foliate_monarr f)) + = Nf (snd (trim_foliate_monarr_helper f)) /\ + exists H : totally_composable (trim_foliate_monarr f), + monarr_norm_equiv f + (compose_totally_composable (trim_foliate_monarr f) H). +Proof. + unfold trim_foliate_monarr. + induction f. + - split; [easy|]. + simpl. + destruct (trim_foliate_monarr_helper f1) as [f1s f1_targ]. + destruct (trim_foliate_monarr_helper f2) as [f2s f2_targ]. + simpl in *. + destruct IHf1 as (Hf1targ & H1targ & H1comp & Heq1). + destruct IHf2 as (Hf2targ & H2targ & H2comp & Heq2). + destruct f1s as [|f1h f1s], f2s as[|f2h f2s]; + (split; [try easy|]); + cbn [proj_to_monarrlist_list fst] in *. + + simpl in *. + exists Logic.I. + rewrite <- (monarr_lunit (_ ⧆ _)). + apply monarr_norm_equiv_comp; [|apply Heq2]. + apply (monarr_norm_equiv_trans Heq1). + rewrite monarr_arrtens. + apply monarr_norm_equiv_struct_eq_in. + cbn. + rewrite Hf1targ. + apply (eq_trans (Nf_eq_in_of_norm_equiv Heq2)). + easy. + + simpl in *. + exists H2comp. + rewrite <- (monarr_lunit (compose_totally_composable_helper _ _ _)). + apply monarr_norm_equiv_comp; [|apply Heq2]. + apply (monarr_norm_equiv_trans Heq1). + rewrite monarr_arrtens. + apply monarr_norm_equiv_struct_eq_in. + cbn. + rewrite Hf1targ. + apply (eq_trans (Nf_eq_in_of_norm_equiv Heq2)). + easy. + + simpl in *. + rewrite !app_nil_r. + etransitivity; + [|apply (Nf_eq_in_of_norm_equiv Heq2)]. + rewrite <- Hf1targ. + apply H1targ. + + simpl in *. + rewrite app_nil_r. + exists H1comp. + rewrite <- (monarr_runit (compose_totally_composable_helper _ _ _)). + apply monarr_norm_equiv_comp; [apply Heq1|]. + apply (monarr_norm_equiv_trans Heq2). + rewrite monarr_arrtens. + apply monarr_norm_equiv_struct_eq_in. + simpl. + rewrite H1targ. + symmetry. + etransitivity; [|apply (Nf_eq_in_of_norm_equiv Heq2)]. + now rewrite Hf1targ. + + simpl in *. + rewrite monarrlist_list_target_app. + destruct f1s; apply H2targ. + + unshelve (eexists). + 1: { + abstract (apply (totally_composable_of_middle_arr + (fss:=_::_) (gss:=_::_) H1comp H2comp); + rewrite H1targ, Hf1targ; + apply (Nf_eq_in_of_norm_equiv Heq2)). + } + cbn [app proj_to_monarrlist_list fst]. + change (?x :: ?l ++ ?l') with ((x :: l) ++ l'). + unshelve (rewrite compose_totally_composable_app_nonempty); + [abstract(easy)..|]. + apply monarr_norm_equiv_conj_struct'_iff. + unfold compose_composable_monarr. + apply monarr_norm_equiv_comp. + * apply monarr_norm_equiv_struct_r'_iff. + erewrite compose_totally_composable_indep. + apply Heq1. + * erewrite compose_totally_composable_indep. + apply Heq2. + - split; [simpl; f_equal; easy|]. + simpl. + destruct (trim_foliate_monarr_helper f1) as [f1s f1_targ]. + destruct (trim_foliate_monarr_helper f2) as [f2s f2_targ]. + simpl in *. + destruct IHf1 as (Hf1targ & H1targ & H1comp & Heq1). + destruct IHf2 as (Hf2targ & H2targ & H2comp & Heq2). + destruct f1s as [|f1h f1s], f2s as[|f2h f2s]; + (split; [try easy|]); + cbn [proj_to_monarrlist_list fst snd zip_defaults] in *. + + simpl in *. + cbn [monarrlist_target fold_right map] in *. + exists Logic.I. + simpl in *. + apply (monarr_norm_equiv_trans + (monarr_norm_equiv_tens _ _ _ _ Heq1 Heq2)). + rewrite 4!monarr_arrtens. + apply monarr_norm_equiv_struct_eq_in. + easy. + + simpl. + clear -H2targ. + induction f2s; simpl. + * etransitivity; + [apply Nf_monarrlist_target_app'|]. + apply Nf_tens_eq; easy. + * destruct f2s; + [|apply IHf2s,H2targ]. + etransitivity; + [apply Nf_monarrlist_target_app'|]. + apply Nf_tens_eq; easy. + + unfold proj_to_monarrlist_list at 2. + exists (zip_with_default_l_totally_composable H2comp _). + change (proj_to_monarrlist_list (?f, ?g)) with f. + apply (monarr_norm_equiv_trans + (monarr_norm_equiv_tens _ _ _ _ Heq1 Heq2)). + (* simpl (compose_totally_composable [[_]]). *) + (* cbn beta. *) + (* apply (monarr_norm_equiv_trans (monarr_norm_equiv_tens_assoc _ _ _)). *) + refine (monarr_norm_equiv_trans _ + (monarr_norm_equiv_symm _ _ + (zip_with_default_l_norm_equiv_nonempty (fss:=_::_) + ltac:(easy) H2comp f1_targ))). + apply monarr_norm_equiv_tens; [|easy]. + simpl. + rewrite monarr_arrtens. + apply monarr_norm_equiv_struct_eq_in; easy. + + change (proj_to_monarrlist_list (?f, ?g)) with f. + clear -H1targ. + induction f1s; simpl. + * etransitivity; + [apply (Nf_monarrlist_target_app' [_])|]. + apply Nf_tens_eq; easy. + * destruct f1s; + [|apply IHf1s,H1targ]. + etransitivity; + [apply (Nf_monarrlist_target_app' [_])|]. + apply Nf_tens_eq; easy. + + change (proj_to_monarrlist_list (?f, ?g)) with f. + exists (zip_with_default_r_totally_composable H1comp _). + apply (monarr_norm_equiv_trans + (monarr_norm_equiv_tens _ _ _ _ Heq1 Heq2)). + simpl (compose_totally_composable [[_]]). + cbn beta. + (* apply (monarr_norm_equiv_trans ( + monarr_norm_equiv_symm _ _ (monarr_norm_equiv_tens_assoc _ _ _))). *) + refine (monarr_norm_equiv_trans _ + (monarr_norm_equiv_symm _ _ + (zip_with_default_r_norm_equiv_nonempty (fss:=_::_) + ltac:(easy) H1comp f2_targ))). + apply monarr_norm_equiv_tens; [easy|]. + rewrite monarr_arrtens. + apply monarr_norm_equiv_struct_eq_in; easy. + + rewrite <- zip_defaults_cons. + pose proof (zip_defaults_target_nonempty + (fss:=f2h::f2s) (gss:=f1h::f1s) ltac:(easy) ltac:(easy)) as e. + rewrite Nf_tens_bwnormapp, H1targ, H2targ, <- Nf_tens_bwnormapp in e. + rewrite <- e. + apply zip_defaults_target_defaults_indep; easy. + + eexists (zip_defaults_totally_composable' H2comp H1comp + (eq_sym H2targ) (eq_sym H1targ)). + apply (monarr_norm_equiv_trans + (monarr_norm_equiv_tens _ _ _ _ Heq1 Heq2)). + apply ( + (monarr_norm_equiv_symm _ _ + (tens_zip_defaults_correct_helper'' (fss:=f2h::f2s) (gss:=f1h::f1s) + ltac:(easy) ltac:(easy) H2comp H1comp + (eq_sym H2targ) (eq_sym H1targ)))). + - repeat split; try easy. + exists Logic.I. + apply monarr_norm_equiv_symm. + apply monarr_norm_equiv_tens_bwarr_e_l. + - repeat split; try easy. + exists Logic.I. + apply monarr_norm_equiv_symm. + simpl. + rewrite monarr_arrtens. + apply monarr_norm_equiv_struct_eq_out; easy. +Qed. + +Import Lia. + +Lemma trim_foliate_totally_composable {a b} (f : a ⟶ b) : + totally_composable (trim_foliate_monarr f). +Proof. + apply trim_foliate_spec. +Qed. + +Lemma trim_foliate_correct' {a b} (f : a ⟶ b) : + monarr_norm_equiv f + (compose_totally_composable (trim_foliate_monarr f) + (trim_foliate_totally_composable f)). +Proof. + destruct (trim_foliate_spec f) as (? & ? & ? & ?). + erewrite compose_totally_composable_indep. + eassumption. +Qed. + +Lemma trim_foliate_correct {a b} (f : a ⟶ b) : + f ≊ + bwarr_of_Nf_eq (Nf_eq_in_of_norm_equiv (trim_foliate_correct' f)) + ◌ compose_totally_composable (trim_foliate_monarr f) + (trim_foliate_totally_composable f) + ◌ bwarr_of_Nf_eq (eq_sym + (Nf_eq_out_of_norm_equiv (trim_foliate_correct' f))). +Proof. + apply equiv_conj_of_monarr_norm_equiv, trim_foliate_correct'. +Qed. + +Lemma is_empty_iff (a : bw) : is_empty a = true <-> Nf a = norm_e. +Proof. + induction a; simpl in *; try easy. + rewrite Bool.andb_true_iff. + rewrite IHa1, IHa2. + split. + - rewrite Nf_tens_bwnormapp. + intros [-> ->]; easy. + - rewrite Nf_tens_bwnormapp. + destruct (Nf a1), (Nf a2); easy. +Qed. + +Lemma remove_empty_structs_proper : + unary_monarrlist_proper remove_empty_structs. +Proof. + intros fs. + induction fs as [|[] fs' IHfs]; [easy|..]. + - simpl. + destruct (is_empty a) eqn:E. + + apply (monarr_norm_equiv_trans + (monarr_norm_equiv_symm _ _ + (monarr_norm_equiv_tens_bwarr_e_r _ (arrid e)))). + apply monarr_norm_equiv_tens; [easy|]. + apply monarr_norm_equiv_struct_eq_in. + simpl. + rewrite is_empty_iff in E; rewrite E. + easy. + + apply monarr_norm_equiv_tens; easy. + - apply monarr_norm_equiv_tens; easy. +Qed. + +Lemma trim_foliate_no_empty_composable {a b} (f : a ⟶ b) : + totally_composable (trim_foliate_monarr_no_empty f). +Proof. + unfold trim_foliate_monarr_no_empty. + apply monarrlist_list_proper_composable; + [|apply trim_foliate_totally_composable]. + apply map_unary_proper_monarrlist_list_proper, + remove_empty_structs_proper. +Qed. + +Lemma trim_foliate_no_empty_correct' {a b} (f : a ⟶ b) : + monarr_norm_equiv f + (compose_totally_composable (trim_foliate_monarr_no_empty f) + (trim_foliate_no_empty_composable f)). +Proof. + apply (monarr_norm_equiv_trans (trim_foliate_correct' f)). + erewrite (compose_totally_composable_indep (trim_foliate_monarr_no_empty f)). + unfold trim_foliate_monarr_no_empty. + refine (proper_preserves _ _). + apply map_unary_proper_monarrlist_list_proper, + remove_empty_structs_proper. +Qed. + +End trim_foliation. + +Section idarrs. + +Lemma source_idarr_of_vars (xs : list X) : + source_vars (idarr_of_vars xs) = xs. +Proof. + induction xs; [easy|]. + simpl; f_equal; easy. +Qed. + +Lemma target_idarr_of_vars (xs : list X) : + target_vars (idarr_of_vars xs) = xs. +Proof. + induction xs; [easy|]. + simpl; f_equal; easy. +Qed. + +Lemma composable_idarr_of_target (fs : monarrlist) : + composable fs (idarr_of_vars (target_vars fs)). +Proof. + rewrite composable_iff_varlist_eq. + now rewrite source_idarr_of_vars. +Qed. + +Lemma composable_idarr_of_source (fs : monarrlist) : + composable (idarr_of_vars (source_vars fs)) fs. +Proof. + rewrite composable_iff_varlist_eq. + now rewrite target_idarr_of_vars. +Qed. + +Lemma target_source_idarr_of_vars (xs : list X) : + monarrlist_target (idarr_of_vars xs) + = monarrlist_source (idarr_of_vars xs). +Proof. + induction xs; [easy|]. + simpl. + unfold monarrlist_target, monarrlist_source in *. + simpl. + f_equal; easy. +Qed. + +Lemma source_target_idarr_of_vars (xs : list X) : + monarrlist_source (idarr_of_vars xs) + = monarrlist_target (idarr_of_vars xs). +Proof. + now rewrite target_source_idarr_of_vars. +Qed. + +Lemma realize_idarr (xs : list X) : + stack_monarrlist (idarr_of_vars xs) ≊ + cast_bwarr eq_refl (source_target_idarr_of_vars xs) (arrid _). +Proof. + induction xs; [apply monarr_struct|]. + simpl. + rewrite IHxs. + rewrite monarr_arrtens. + apply monarr_struct. +Qed. + +Lemma compose_composable_monarr_idarr_r + (gs : monarrlist) (xs : list X) (H : composable gs (idarr_of_vars xs)) : + monarr_norm_equiv + (compose_composable gs (idarr_of_vars xs) H) + (stack_monarrlist gs). +Proof. + exists eq_refl. + unshelve (eexists). + rewrite target_source_idarr_of_vars; apply (eq_sym H). + unfold compose_composable. + rewrite monarr_id_l. + rewrite realize_idarr. + rewrite <- !monarr_assoc, !monarr_arrcomp. + apply monarr_id_r. +Qed. + +Lemma compose_composable_monarr_idarr_l + (gs : monarrlist) (xs : list X) (H : composable (idarr_of_vars xs) gs) : + monarr_norm_equiv + (compose_composable (idarr_of_vars xs) gs H) + (stack_monarrlist gs). +Proof. + unshelve (eexists _, eq_refl). + rewrite source_target_idarr_of_vars; apply (eq_sym H). + unfold compose_composable. + rewrite monarr_id_r. + rewrite realize_idarr. + rewrite !monarr_assoc, !monarr_arrcomp. + apply monarr_id_l. +Qed. + +Lemma source_target_is_idarrlist (fs : monarrlist) + (H : is_idarrlist fs = true) : + monarrlist_source fs + = monarrlist_target fs. +Proof. + induction fs as [|a fs IHfs]; [easy|]. + destruct a; [|easy]. + simpl in H. + unfold monarrlist_source, monarrlist_target in *. + simpl. + f_equal; auto. +Qed. + +Lemma input_output_is_idarrlist {fs} (H : is_idarrlist fs = true) : + monarrlist_source fs = monarrlist_target fs. +Proof. + induction fs as [|[] fs]; [easy|..]; simpl in *; [|easy]. + unfold monarrlist_source, monarrlist_target. + simpl. + f_equal. + apply IHfs, H. +Qed. + +Lemma realize_idarrlist (fs : monarrlist) + (H : is_idarrlist fs = true) : + stack_monarrlist fs ≊ + cast_bwarr eq_refl (source_target_is_idarrlist fs H) (arrid _). +Proof. + induction fs as [|a fs IHfs]; [apply monarr_struct|]. + simpl. + destruct a; [|easy]. + simpl in H. + rewrite (IHfs H). + simpl. + rewrite monarr_arrtens. + apply monarr_struct. +Qed. + +Lemma compose_composable_monarr_is_idarr_r {fs} (Hfs : is_idarrlist fs = true) + {a b} (gs : a ⟶ b) (H : Nf b = Nf (monarrlist_source fs)) : + monarr_norm_equiv + (compose_composable_monarr gs (stack_monarrlist fs) H) + gs. +Proof. + exists eq_refl. + unshelve (eexists). + rewrite <- (source_target_is_idarrlist _ Hfs); apply (eq_sym H). + unfold compose_composable_monarr. + rewrite (realize_idarrlist fs Hfs). + rewrite <- !monarr_assoc, !monarr_arrcomp, monarr_id_l. + apply monarr_id_r. +Qed. + +Lemma compose_composable_monarr_is_idarr_l {fs} (Hfs : is_idarrlist fs = true) + {a b} (gs : a ⟶ b) (H : Nf (monarrlist_target fs) = Nf a) : + monarr_norm_equiv + (compose_composable_monarr (stack_monarrlist fs) gs H) + gs. +Proof. + unshelve (eexists _, eq_refl). + rewrite (source_target_is_idarrlist _ Hfs); apply (eq_sym H). + unfold compose_composable_monarr. + rewrite (realize_idarrlist fs Hfs). + rewrite !monarr_assoc, !monarr_arrcomp, monarr_id_r. + apply monarr_id_l. +Qed. + +Lemma split_id_norm_equiv (a : bw) : + monarr_norm_equiv + (stack_monarrlist (map monarrlist_id (map var (bw_to_varlist a)))) + (arrid a). +Proof. + induction a; + [apply monarr_norm_equiv_refl | apply monarr_norm_equiv_tens_bwarr_e_l | ]. + simpl. + rewrite 2!map_app. + apply (monarr_norm_equiv_trans (stack_monarrlist_app' _ _)). + rewrite (monarr_struct _ (arrid a1 ⊠ arrid a2)). + rewrite <- monarr_arrtens. + apply monarr_norm_equiv_tens; + assumption. +Qed. + +Lemma split_ids_proper_unary : + unary_monarrlist_proper split_ids. +Proof. + intros fss. + induction fss as [|[] fss]; + [apply monarr_norm_equiv_refl|..]; + simpl; + [|apply monarr_norm_equiv_tens; + assumption + apply monarr_norm_equiv_refl]. + apply (monarr_norm_equiv_trans (stack_monarrlist_app' _ _)). + apply monarr_norm_equiv_tens; [apply IHfss|]. + apply split_id_norm_equiv. +Qed. + +End idarrs. + +Section shift_left_right. + +Lemma shift_left_proper : binary_monarrlist_proper shift_left. +Proof. + intros fs gs H. + unfold shift_left. + destruct (is_idarrlist fs) eqn:E; + [|exists H; easy]. + exists (composable_idarr_of_target gs). + simpl. + eapply (monarr_norm_equiv_trans); + [apply (compose_composable_monarr_is_idarr_l E)|]. + apply monarr_norm_equiv_symm. + apply compose_composable_monarr_idarr_r. +Qed. + +Lemma shift_right_proper : binary_monarrlist_proper shift_right. +Proof. + intros fs gs H. + unfold shift_right. + destruct (is_idarrlist gs) eqn:E; + [|exists H; easy]. + exists (composable_idarr_of_source fs). + simpl. + eapply (monarr_norm_equiv_trans); + [apply (compose_composable_monarr_is_idarr_r E)|]. + apply monarr_norm_equiv_symm. + apply compose_composable_monarr_idarr_l. +Qed. + +Lemma full_right_shift_proper : + monarrlist_list_proper full_right_shift. +Proof. + apply monarrlist_list_proper_iter'. + apply monarrlist_list_proper_compose. + + apply pairwise_apply_proper_binary_proper, + pairify_apply_proper, shift_right_proper. + + apply map_unary_proper_monarrlist_list_proper, + split_ids_proper_unary. +Qed. + +End shift_left_right. + +Section structify. + +Import Bool. + +Lemma struct_of_is_struct_indep {a b} (f : a ⟶ b) + (H G : is_struct f = true) : + struct_of_is_struct f H = struct_of_is_struct f G. +Proof. + rewrite (uip H G). + easy. +Qed. + +Lemma is_struct_structify {a b} (f : a ⟶ b) : + is_struct f = is_struct (structify f). +Proof. + induction f; try reflexivity; + simpl; + destruct (bool_dec (is_struct (structify f1)) true); + destruct (bool_dec (is_struct (structify f2)) true); + simpl; try congruence; + apply andb_true_iff; split; congruence. +Qed. + +Lemma structify_of_is_struct {a b} (f : a ⟶ b) + (H : is_struct f = true) : + structify f = struct_of_is_struct f H. +Proof. + induction f; try reflexivity + discriminate; + simpl in H; + pose proof H as G; + rewrite (andb_true_iff _ _) in G; + destruct G as [Hl Hr]; + simpl structify; + rewrite (IHf1 Hl), (IHf2 Hr); + repeat match goal with + |- context[match ?eqab with | left _ => _ | right _ => _ end] => + destruct eqab; [|pose proof @is_struct_structify; congruence] + end; + simpl; + erewrite struct_of_is_struct_indep, (struct_of_is_struct_indep f2); + reflexivity. +Qed. + + +Lemma struct_of_is_struct_equiv {a b} (f : a ⟶ b) + (H : is_struct f = true) : + struct_of_is_struct f H ≊ f. +Proof. + symmetry. + induction f; [..|discriminate|reflexivity]; + pose proof H as G; + simpl in G; + rewrite andb_true_iff in G; destruct G as [Hl Hr]; + rewrite (IHf1 Hl), (IHf2 Hr) at 1; + simpl; + [ rewrite <- monarr_arrcomp + | rewrite <- monarr_arrtens ]; + erewrite struct_of_is_struct_indep, (struct_of_is_struct_indep f2); + reflexivity. +Qed. + +Lemma structify_equiv {a b} (f : a ⟶ b) : + f ≊ structify f. +Proof. + induction f; try reflexivity; + rewrite IHf1, IHf2 at 1; + simpl; + repeat match goal with + |- context[match ?eqab with | left _ => _ | right _ => _ end] => + destruct eqab + end; try reflexivity. + - rewrite <- monarr_arrcomp. + rewrite 2!struct_of_is_struct_equiv by easy. + easy. + - rewrite <- monarr_arrtens. + rewrite 2!struct_of_is_struct_equiv by easy. + easy. +Qed. + +Lemma structify_proper : unary_monarrlist_proper structify_monarrlist. +Proof. + apply map_monarr_proper_unary_proper. + intros; apply structify_equiv. +Qed. + +Lemma remove_struct_proper (g : monarrlistelt) : + monarr_norm_equiv + (realize_monarrlistelt (remove_struct g)) + (realize_monarrlistelt g). +Proof. + destruct g; [apply monarr_norm_equiv_refl|]. + simpl. + destruct f; [apply monarr_norm_equiv_refl..|]. + unfold monarr_norm_equiv. + exists eq_refl, (Nf_eq_of_arr f). + simpl. + rewrite !monarr_arrcomp. + apply monarr_struct. +Qed. + +Lemma remove_structs_proper : + unary_monarrlist_proper remove_structs. +Proof. + apply map_elt_proper_unary_proper. + exact remove_struct_proper. +Qed. + +End structify. + +Section trim_idarrs. + +Lemma source_trim_idarrs_helper fs fss : + monarrlist_list_source (trim_idarrs_helper fs fss) + = monarrlist_list_source (fs :: fss). +Proof. + revert fs; + induction fss; intros fs; [easy|simpl]. + destruct (is_idarrlist a); + easy + apply IHfss. +Qed. + +Lemma target_trim_idarrs_helper fs fss (H : totally_composable (fs :: fss)) : + Nf (monarrlist_list_target (trim_idarrs_helper fs fss)) + = Nf (monarrlist_list_target (fs :: fss)). +Proof. + revert fs H; + induction fss; intros fs H; [easy|simpl]. + destruct (is_idarrlist a) eqn:e; [|easy]. + apply source_target_is_idarrlist in e. + rewrite <- e. + rewrite IHfss. + - destruct fss; easy + apply H. + - apply (totally_composable_helper_swap + (eq_sym (eq_trans (proj1 H) (f_equal Nf e)))). + apply H. +Qed. + +Lemma trim_idarrs_helper_composable {fs fss} + (H : totally_composable (fs :: fss)) : + totally_composable (trim_idarrs_helper fs fss). +Proof. + revert fs H; induction fss; intros fs H; [easy|]. + simpl. + destruct (is_idarrlist a) eqn:e; [|apply H]. + apply source_target_is_idarrlist in e. + apply IHfss. + apply (totally_composable_helper_swap + (eq_sym (eq_trans (proj1 H) (f_equal Nf e)))). + apply H. +Qed. + +Lemma trim_idarrs_helper_equiv {fs fss} + (H : totally_composable (fs :: fss)) : + monarr_norm_equiv + (compose_totally_composable _ H) + (compose_totally_composable _ (trim_idarrs_helper_composable H)). +Proof. + revert fs H; induction fss; intros fs H; [easy|]. + generalize (trim_idarrs_helper_composable H). + destruct fss. + - simpl. + destruct (is_idarrlist a) eqn:e. + + simpl. + intros _. + apply (compose_composable_monarr_is_idarr_r e). + + intros ?. + simpl. + erewrite compose_composable_monarr_indep; easy. + - simpl. + destruct (is_idarrlist a) eqn:e; intros Hc. + + erewrite (totally_composable_indep _ Hc). + eapply monarr_norm_equiv_trans; + [|refine (IHfss fs _)]. + 2: { + abstract (split; [|apply H]; + unfold composable; + rewrite (proj1 H); + rewrite (source_target_is_idarrlist _ e); + apply H). + } + simpl. + apply monarr_norm_equiv_comp_composable; [easy|]. + erewrite (totally_composable_indep (m::fss) (proj2 (proj2 H))). + apply (compose_composable_monarr_is_idarr_l e). + + erewrite (totally_composable_indep _ H Hc). + easy. +Qed. + +Lemma trim_idarrs_composable {fss} (H : totally_composable fss) : + totally_composable (trim_idarrs fss). +Proof. + destruct fss; [easy|apply trim_idarrs_helper_composable, H]. +Qed. + +Lemma trim_idarrs_equiv {fss} (H : totally_composable fss) : + monarr_norm_equiv + (compose_totally_composable fss H) + (compose_totally_composable (trim_idarrs fss) (trim_idarrs_composable H)). +Proof. + destruct fss; [easy|]. + erewrite (compose_totally_composable_indep (trim_idarrs _)). + apply trim_idarrs_helper_equiv. +Qed. + +Lemma trim_idarrs_proper : monarrlist_list_proper trim_idarrs. +Proof. + intros fss H. + exists (trim_idarrs_composable H). + apply trim_idarrs_equiv. +Qed. + +Lemma trim_idarr_proper : monarrlist_list_proper trim_idarr. +Proof. + intros fs Hfs. + unshelve (eexists). + - abstract (destruct fs; [easy|]; destruct fs; [easy|]; + simpl; + destruct (is_idarrlist m); + apply Hfs). + - destruct fs; [easy|]; destruct fs; [easy|]; + simpl. + generalize (trim_idarr_proper_subproof (m::m0::fs) Hfs). + simpl. + destruct (is_idarrlist m) eqn:e; + [|intros; erewrite (compose_totally_composable_indep _ t Hfs); easy]. + intros. + erewrite (compose_totally_composable_indep _ t (proj2 Hfs)). + apply compose_composable_monarr_is_idarr_l, e. +Qed. + +End trim_idarrs. + +Section filter_idarrs. + +Lemma totally_composable_filter_idarrs {fss} : + totally_composable (fss) -> + totally_composable (filter_idarrs fss). +Proof. + destruct fss as [|fs fss]; [easy|]. + revert fs. + (* intros H. *) + induction fss as [|fs' fss]; [easy|]; + intros fs H. + simpl. + destruct (is_idarrlist fs') eqn:e; simpl. + - apply IHfss. + refine (totally_composable_helper_swap _ (proj2 H)). + rewrite <- (input_output_is_idarrlist e). + symmetry. + apply H. + - split; [apply H|]. + apply IHfss, H. +Qed. + +Fixpoint list_iter_map {A} (f : list A -> list A) (l : list A) := + match l with + | nil => nil + | a :: l' => f (a :: list_iter_map f l') + end. + +Lemma trim_idarrs_helper_filter_not_idarrlist fs fss : + trim_idarrs_helper fs (filter (fun x => negb (is_idarrlist x)) fss) = + fs :: filter (fun x => negb (is_idarrlist x)) fss. +Proof. + revert fs; + induction fss; [easy|]; intros fs. + simpl. + destruct (is_idarrlist a) eqn:e; + [apply IHfss|]. + simpl. + rewrite e. + easy. +Qed. + +Lemma filter_idarrs_trim fss : + filter_idarrs fss = list_iter_map trim_idarrs fss. +Proof. + induction fss as [|fs fss IHfss]; [easy|]. + simpl. + rewrite <- IHfss. + clear IHfss. + destruct fss as [|fs' fss]; [easy|]. + simpl. + destruct (is_idarrlist fs') eqn:e; simpl; [|easy]. + now rewrite trim_idarrs_helper_filter_not_idarrlist. +Qed. + +Lemma totally_composable_helper_of_totally_composable {a fs} : + totally_composable fs -> + Nf (monarrlist_target a) = Nf (monarrlist_list_source fs) -> + totally_composable_helper a fs. +Proof. + destruct fs; easy. +Qed. + +Lemma list_iter_map_proper {f} (Hf : monarrlist_list_proper f) : + monarrlist_list_proper (list_iter_map f). +Proof. + intros fss Hfss. + match goal with |- ?G => enough ( + (forall a, totally_composable (a :: fss) -> + totally_composable (a :: list_iter_map f fss)) /\ G) + by easy + end. + induction fss; [intros; split; try exists Logic.I; easy|]. + split; [|unshelve (eexists)]. + - intros b. + intros Hcompba. + destruct (Hf (a :: fss) Hfss) as (Hcomp & Hin & _). + apply totally_composable_helper_of_totally_composable. + + apply Hf. + apply (IHfss (totally_composable_of_cons Hfss)), Hfss. + + simpl. + destruct (Hf _ (proj1 (IHfss (totally_composable_of_cons Hfss)) a Hfss)) + as (Hcomp' & Hin' & Hout' & _). + rewrite Hin'. + apply Hcompba. + - abstract (destruct (IHfss (totally_composable_of_cons Hfss)) as + [_ (Hcomp & Hin & Hout & Heq)]; + simpl; + apply Hf; + destruct fss; [easy|]; + apply (totally_composable_helper_of_totally_composable Hcomp); + rewrite Hin; + apply Hfss). + - specialize (IHfss (totally_composable_of_cons Hfss)) as IHfss'. + destruct (Hf _ (proj1 IHfss' a Hfss)) as + (Hcomp & Hequiv). + erewrite (compose_totally_composable_indep _ (_ _)). + refine (monarr_norm_equiv_trans _ Hequiv). + generalize (proj1 IHfss' a Hfss). + destruct (list_iter_map f fss) eqn:e. + + destruct fss; [easy|]. + intros ?. + rewrite compose_totally_composable_cons. + rewrite <- (monarr_runit (compose_totally_composable [a] _)). + unfold compose_composable_monarr. + apply monarr_norm_equiv_comp. + * apply monarr_norm_equiv_struct_r. + * destruct (proj2 (IHfss (proj2 Hfss))) as [? Heq]. + refine (monarr_norm_equiv_trans Heq _). + apply monarr_norm_equiv_struct_eq_in. + pose proof (proj2 (IHfss (proj2 Hfss))) as H. + rewrite <- e in H. + destruct H as (? & Hin & _). + rewrite e in Hin. + rewrite Hin. + symmetry. + apply Hfss. + + destruct fss; [easy|]. + intros t. + rewrite 2!compose_totally_composable_cons. + apply monarr_norm_equiv_comp_composable; [easy|]. + destruct (proj2 IHfss') as [? Heq]. + erewrite (compose_totally_composable_indep (m::l) (proj2 t)). + erewrite (compose_totally_composable_indep (m0::fss) (proj2 Hfss)). + apply Heq. +Qed. + +Lemma filter_idarrs_proper : + monarrlist_list_proper filter_idarrs. +Proof. + intros fss. + rewrite filter_idarrs_trim. + apply list_iter_map_proper. + apply trim_idarrs_proper. +Qed. + +End filter_idarrs. + +Section full_process. + +Lemma full_process_monarrlist_list_proper : + monarrlist_list_proper full_process_monarrlist_list. +Proof. + unfold full_process_monarrlist_list. + repeat match goal with + | |- monarrlist_list_proper (Basics.compose _ _) => + apply monarrlist_list_proper_compose + | |- monarrlist_list_proper (map _) => + apply map_unary_proper_monarrlist_list_proper + end. + - apply structify_proper. + - apply remove_structs_proper. + - apply split_ids_proper_unary. + - apply full_right_shift_proper. + - apply trim_idarrs_proper. + - apply trim_idarr_proper. +Qed. + +Lemma full_process_in_pf {a b} (f : a ⟶ b) : + Nf a = Nf (monarrlist_list_source + (full_process_monarrlist_list (foliate_monarr f))). +Proof. + rewrite <- (Nf_eq_in_of_norm_equiv + (proper_preserves full_process_monarrlist_list_proper + (foliate_monarr_totally_composable f))). + rewrite monarrlist_list_source_foliate. + easy. +Qed. + +Lemma full_process_out_pf {a b} (f : a ⟶ b) : + Nf (monarrlist_list_target + (full_process_monarrlist_list (foliate_monarr f))) + = Nf b. +Proof. + rewrite <- (Nf_eq_out_of_norm_equiv + (proper_preserves full_process_monarrlist_list_proper + (foliate_monarr_totally_composable f))). + rewrite monarrlist_list_target_foliate. + easy. +Qed. + +Definition full_process_monarr {a b} (f : a ⟶ b) := + compose_totally_composable (full_process_monarrlist_list (foliate_monarr f)) + (monarrlist_list_proper_composable full_process_monarrlist_list_proper + (foliate_monarr_totally_composable f)). + +Lemma full_process_monarr_equiv {a b} (f : a ⟶ b) : + f ≊ + bwarr_of_Nf_eq (full_process_in_pf f) + ◌ full_process_monarr f + ◌ bwarr_of_Nf_eq (full_process_out_pf f). +Proof. + apply equiv_conj_of_monarr_norm_equiv. + unfold full_process_monarr. + apply (monarr_norm_equiv_trans (foliate_correct' f)). + apply proper_preserves. +Qed. + +End full_process. + +End ProperOperationInstances. \ No newline at end of file diff --git a/ViCaR/MonoidalCoherence/MCproperop.v b/ViCaR/MonoidalCoherence/MCproperop.v new file mode 100644 index 0000000..3dc9c31 --- /dev/null +++ b/ViCaR/MonoidalCoherence/MCproperop.v @@ -0,0 +1,761 @@ +Require Import Setoid. +Require MCmonarrlist. + +Section ProperOperation. + +Set Universe Polymorphism. + +Open Scope bw_scope. + +Import MCDefinitions MCClasses MC_setoids + MC_notations UIP_facts MCbw MCconsequences MCmonarrlist. +Import CategoryTypeclass. +Import List ListNotations. + +Context {X : Type} {UIPX : UIP X}. +Context {cC : Category X} {cCh : CategoryCoherence cC} + {mC : MonoidalCategory cC} {mCh : MonoidalCategoryCoherence mC}. + +Local Notation bw := (bw X). +Local Notation "a ⟶ b" := (@monarr X cC mC a b) (at level 60). +Local Notation monarrlist := (@monarrlist X cC mC). +Local Notation monarrlistelt := (@monarrlistelt X cC mC). + +Section properop_definitions. + +Definition monarrlistelt_proper (f : monarrlistelt -> monarrlistelt) : Prop := + forall g : monarrlistelt, + monarr_norm_equiv (realize_monarrlistelt (f g)) (realize_monarrlistelt g). + +Definition unary_monarrlist_proper (f : monarrlist -> monarrlist) : Prop := + forall fs : monarrlist, monarr_norm_equiv + (stack_monarrlist (f fs)) (stack_monarrlist fs). + +Definition binary_monarrlist_proper + (f : monarrlist -> monarrlist -> monarrlist * monarrlist) : Prop := + forall (fs gs : monarrlist) (Hfg : composable fs gs), + exists (H : composable (fst (f fs gs)) (snd (f fs gs))), + monarr_norm_equiv + (compose_composable fs gs Hfg) + (compose_composable _ _ H). + +Definition bin_to_kary_monarrlist_proper + (f : monarrlist -> monarrlist -> monarrlist * list monarrlist) := + forall fs gs (Hfg : composable fs gs), + exists H : totally_composable (fst (f fs gs) :: snd (f fs gs)), + monarr_norm_equiv + (compose_composable_monarr + (stack_monarrlist fs) + (stack_monarrlist gs) Hfg) + (compose_totally_composable + (fst (f fs gs) :: snd (f fs gs)) H). + +Definition monarrlist_list_proper (f : list monarrlist -> list monarrlist) := + forall (fss : list monarrlist) (Hfss : totally_composable fss), + exists (H : totally_composable (f fss)), + monarr_norm_equiv + (compose_totally_composable fss Hfss) + (compose_totally_composable (f fss) H). + +Fixpoint pairwise_apply_helper + (f : monarrlist -> monarrlist -> monarrlist * monarrlist) + (fs : monarrlist) (fss : list monarrlist) := + match fss with + | nil => fs::nil + | fs' :: fss' => + fst (f fs fs') :: (pairwise_apply_helper f (snd (f fs fs')) fss') + end. + +Definition pairwise_apply f (fss : list monarrlist) := + match fss with + | nil => nil + | fs :: fss' => pairwise_apply_helper f fs fss' + end. + +End properop_definitions. + +Section elt_properop. + +Lemma map_elt_proper_unary_proper (f : monarrlistelt -> monarrlistelt) + (Hf : monarrlistelt_proper f) : + unary_monarrlist_proper (map f). +Proof. + intros fs; induction fs; [easy|]. + apply monarr_norm_equiv_tens; easy. +Qed. + +Lemma map_monarr_proper_unary_proper + (f : forall a b : bw, a ⟶ b -> a ⟶ b) + (Hf : forall {a b : bw} (g : a ⟶ b), g ≊ f a b g) : + unary_monarrlist_proper + (map (monarrlistelt_map f)). +Proof. + apply map_elt_proper_unary_proper. + intros []; apply monarr_norm_equiv_of_monarrequiv; + simpl; easy. +Qed. + +End elt_properop. + +Section unary_properop. + +Lemma unary_proper_id : unary_monarrlist_proper id. +Proof. + unfold id. easy. +Qed. + +Lemma unary_proper_compose {f g} + (Hf : unary_monarrlist_proper f) + (Hg : unary_monarrlist_proper g) : + unary_monarrlist_proper (Basics.compose f g). +Proof. + intros fs. + apply (monarr_norm_equiv_trans (g := stack_monarrlist (g fs))); + apply Hf + apply Hg. +Qed. + +Lemma unary_proper_iter {f} (Hf : unary_monarrlist_proper f) n : + unary_monarrlist_proper (Nat.iter n f). +Proof. + induction n; [apply unary_proper_id|]. + apply unary_proper_compose; assumption. +Qed. + +Lemma unary_proper_iter' {f} (Hf : unary_monarrlist_proper f) + (fn : monarrlist -> nat) : + unary_monarrlist_proper (fun fs => Nat.iter (fn fs) f fs). +Proof. + intros fs; apply (unary_proper_iter Hf). +Qed. + +Lemma unary_proper_composable {f} (H : unary_monarrlist_proper f) + {fs gs} : + composable fs gs -> composable (f fs) (f gs). +Proof. + unfold composable. + intros Hc. + rewrite (Nf_eq_out_of_norm_equiv (H fs)). + rewrite (Nf_eq_in_of_norm_equiv (H gs)). + apply Hc. +Qed. + +Lemma unary_proper_totally_composable {f} (H : unary_monarrlist_proper f) + {fss : list monarrlist} (Hf : totally_composable fss) : + totally_composable (map f fss). +Proof. + induction fss; [easy|]. + destruct fss; [easy|]. + specialize (IHfss (proj2 Hf)). + split. + - apply (unary_proper_composable H), Hf. + - apply IHfss. +Qed. + +Lemma unary_proper_preserves {f} (H : unary_monarrlist_proper f) + {fss : list monarrlist} (Hf : totally_composable fss) : + monarr_norm_equiv + (compose_totally_composable fss Hf) + (compose_totally_composable (map f fss) + (unary_proper_totally_composable H Hf)). +Proof. + induction fss; [apply monarr_norm_equiv_refl|]; destruct fss. + - apply monarr_norm_equiv_symm. + apply H. + - simpl map. + rewrite 2!compose_totally_composable_cons. + apply monarr_norm_equiv_comp_composable; + [apply monarr_norm_equiv_symm, (H a)|]. + erewrite (compose_totally_composable_indep (f _ :: _)). + apply IHfss. +Qed. + +Lemma map_unary_proper_monarrlist_list_proper {f} + (Hf : unary_monarrlist_proper f) : + monarrlist_list_proper (map f). +Proof. + intros fss Hfss. + exists (unary_proper_totally_composable Hf Hfss). + apply unary_proper_preserves. +Qed. + +End unary_properop. + + +Section binary_properop. + +Lemma binary_proper_composable f (Hf : binary_monarrlist_proper f) + (fs gs : monarrlist) (H : composable fs gs) : + composable (fst (f fs gs)) (snd (f fs gs)). +Proof. + destruct (Hf fs gs H); easy. +Qed. + +Lemma Nf_in_eq_of_proper_binary_op f (Hf : binary_monarrlist_proper f) + (fs gs : monarrlist) (H : composable fs gs) : + Nf (monarrlist_source (fst (f fs gs))) = Nf (monarrlist_source fs). +Proof. + destruct (Hf fs gs H) as (? & ? & _); easy. +Qed. + +Lemma Nf_out_eq_of_proper_binary_op f (Hf : binary_monarrlist_proper f) + (fs gs : monarrlist) (H : composable fs gs) : + Nf (monarrlist_target (snd (f fs gs))) = Nf (monarrlist_target gs). +Proof. + destruct (Hf fs gs H) as (? & ? & ? & _); easy. +Qed. + +Lemma binary_proper_composable_prefix f (Hf : binary_monarrlist_proper f) + fs fs' fss (H : totally_composable (fs :: fs' :: fss)) : + totally_composable (fst (f fs fs') :: snd (f fs fs') :: fss). +Proof. + destruct (Hf fs fs' (proj1 H)) as (? & ? & Hout & ?). + split; [easy|]. + apply (totally_composable_helper_swap Hout). + apply H. +Qed. + +Lemma binary_proper_of_composable_proper + (f : monarrlist -> monarrlist -> monarrlist * monarrlist) + (Hc : forall fs gs, composable fs gs -> composable (fst (f fs gs)) (snd (f fs gs))) + (Hprop : forall fs gs (Hfg : composable fs gs) + (Hc' : forall fs gs, composable fs gs -> composable (fst (f fs gs)) (snd (f fs gs))), + monarr_norm_equiv + (compose_composable fs gs Hfg) + (compose_composable (fst (f fs gs)) (snd (f fs gs)) + (Hc' fs gs Hfg)) + ): + binary_monarrlist_proper f. +Proof. + intros fs gs H; exists (Hc fs gs H). + apply Hprop. +Qed. + +Lemma binary_proper_iff_composable_proper + (f : monarrlist -> monarrlist -> monarrlist * monarrlist) : + binary_monarrlist_proper f <-> + (forall fs gs, composable fs gs -> composable (fst (f fs gs)) (snd (f fs gs))) + /\ forall fs gs (Hfg : composable fs gs) + (Hc' : forall fs gs, composable fs gs -> composable (fst (f fs gs)) (snd (f fs gs))), + monarr_norm_equiv + (compose_composable fs gs Hfg) + (compose_composable (fst (f fs gs)) (snd (f fs gs)) + (Hc' fs gs Hfg)). +Proof. + split; [|intros []; apply binary_proper_of_composable_proper; easy]. + intros Hf. + split. + - intros; apply Hf; easy. + - intros. + destruct (Hf fs gs Hfg) as (? & H). + erewrite (composable_independence (_ _)). + apply H. +Qed. + +Lemma binary_proper_id : binary_monarrlist_proper pair. +Proof. + intros fs gs H. + exists H; easy. +Qed. + +Lemma binary_proper_compose {f g} (Hf : binary_monarrlist_proper f) + (Hg : binary_monarrlist_proper g) : + binary_monarrlist_proper (fun fs gs => uncurry g (f fs gs)). +Proof. + intros fs gs H. + destruct (Hf fs gs H) as [Hcf Heqf]. + destruct (Hg _ _ Hcf) as [Hcg Heqg]. + unfold uncurry. + rewrite (surjective_pairing (f fs gs)). + exists Hcg. + eapply monarr_norm_equiv_trans; eassumption. +Qed. + +Lemma binary_proper_iter {f} (Hf : binary_monarrlist_proper f) n : + binary_monarrlist_proper (curry (Nat.iter n (uncurry f))). +Proof. + induction n. + - apply binary_proper_id. + - apply binary_proper_compose; assumption. +Qed. + +Section pairwise_proper. + +Lemma pairwise_apply_helper_totally_composable fs fss f + (Hf : binary_monarrlist_proper f) : + totally_composable_helper fs fss -> + totally_composable (pairwise_apply_helper f fs fss). +Proof. + revert fs; induction fss; intros fs; [easy|]. + intros H; pose proof H as G; destruct G as [Hl Hr]. + simpl. + destruct fss as [|fs' fss]. + - split; [|easy]. + apply Hf; easy. + - split. + + simpl in Hr. + destruct (Hf _ _ (proj1 H)) as [H' Heq]. + unfold composable in *. + rewrite H'. + symmetry. + apply (Nf_in_eq_of_proper_binary_op f Hf _ fs'). + unfold composable. + rewrite (Nf_out_eq_of_proper_binary_op f Hf _ _ (proj1 H)). + apply H. + + apply IHfss. + eapply totally_composable_helper_swap; [|apply Hr]. + symmetry. + apply (Nf_out_eq_of_proper_binary_op f Hf _ _ Hl). +Qed. + +Lemma pairwise_apply_totally_composable fss f + (Hf : binary_monarrlist_proper f) : + totally_composable fss -> + totally_composable (pairwise_apply f fss). +Proof. + destruct fss; [easy|]. + apply pairwise_apply_helper_totally_composable, Hf. +Qed. + +Lemma pairwise_apply_proper_binary_proper f + (Hf : binary_monarrlist_proper f) : + monarrlist_list_proper (pairwise_apply f). +Proof. + intros fss Hfss. + exists (pairwise_apply_totally_composable fss f Hf Hfss). + destruct fss as [|fs fss]; [easy|]. + simpl. + revert fs Hfss; + induction fss; intros fs H; [apply monarr_norm_equiv_refl|]. + pose proof (binary_proper_composable_prefix f Hf fs a fss H) as Hc'. + apply (monarr_norm_equiv_trans (g:=compose_totally_composable _ Hc')). + - apply (compose_monarrlist_list_cancel_suffix [fs;a] [fst _;snd _] fss). + simpl. + destruct (Hf fs a (proj1 H)) as (? & H'). + erewrite (compose_composable_monarr_indep (stack_monarrlist (fst _))), + compose_composable_monarr_indep. + apply H'. + - apply (compose_monarrlist_list_cancel_prefix (snd _ ::fss) + (pairwise_apply_helper f _ fss) [_]). + apply monarr_norm_equiv_symm. + erewrite compose_totally_composable_indep. + apply monarr_norm_equiv_symm. + erewrite compose_totally_composable_indep. + apply (IHfss _ (proj2 Hc')). +Qed. + +End pairwise_proper. +End binary_properop. + +Section monarrlist_list_proper. + +Lemma monarrlist_list_proper_composable {f} + (Hf : monarrlist_list_proper f) {fss} + (H : totally_composable fss) : + totally_composable (f fss). +Proof. + apply Hf, H. +Qed. + +Lemma proper_preserves {f} (Hf : monarrlist_list_proper f) + {fss} (H : totally_composable fss) : + monarr_norm_equiv + (compose_totally_composable fss H) + (compose_totally_composable (f fss) + (monarrlist_list_proper_composable Hf H)). +Proof. + destruct (Hf fss H) as (? & H'). + erewrite (compose_totally_composable_indep (f fss)). + apply H'. +Qed. + +(* Making these special cases to make terms look more palatable + when applying proper_preserves'. *) +Lemma monarrlist_list_proper_in_pf {f} (Hf : monarrlist_list_proper f) + {fss} (H : totally_composable fss) : + Nf (monarrlist_list_source fss) = + Nf (monarrlist_list_source (f fss)). +Proof. + exact (Nf_eq_in_of_norm_equiv (proper_preserves Hf H)). +Qed. + +Lemma monarrlist_list_proper_out_pf {f} (Hf : monarrlist_list_proper f) + {fss} (H : totally_composable fss) : + Nf (monarrlist_list_target (f fss)) = + Nf (monarrlist_list_target fss). +Proof. + symmetry. + exact (Nf_eq_out_of_norm_equiv (proper_preserves Hf H)). +Qed. + +Lemma proper_preserves' {f} (Hf : monarrlist_list_proper f) + {fss} (H : totally_composable fss) : + compose_totally_composable fss H ≊ + bwarr_of_Nf_eq (monarrlist_list_proper_in_pf Hf H) ◌ + compose_totally_composable (f fss) (monarrlist_list_proper_composable Hf H) ◌ + bwarr_of_Nf_eq (monarrlist_list_proper_out_pf Hf H). +Proof. + destruct (monarr_norm_equiv_symm _ _ (proper_preserves Hf H)) + as (Hin & Hout & Heq). + rewrite <- Heq. + eauto with monarrdb. +Qed. + +Lemma monarrlist_list_proper_totally_composable {f} + (Hf : monarrlist_list_proper f) {fss} : + totally_composable fss -> totally_composable (f fss). +Proof. + intros Hfss. + destruct (Hf fss Hfss); + easy. +Qed. + +Lemma monarrlist_list_proper_iff {f} : + monarrlist_list_proper f <-> + {Hc : forall fss, totally_composable fss -> + totally_composable (f fss) & + forall fss (Hfss : totally_composable fss) + (Hffss : totally_composable (f fss)), + monarr_norm_equiv + (compose_totally_composable fss Hfss) + (compose_totally_composable (f fss) Hffss)}. +Proof. + split. + - intros Hf. + exists (@monarrlist_list_proper_totally_composable f Hf). + intros fss Hfss Hc. + destruct (Hf fss Hfss). + erewrite (compose_totally_composable_indep (f fss)). + eassumption. + - intros [Hc Hprop]. + intros fss Hfss. + exists (Hc fss Hfss). + apply Hprop. +Qed. + +Lemma monarrlist_list_proper_preserves {f} (Hf : monarrlist_list_proper f) + {fss} (Hfss : totally_composable fss) : + monarr_norm_equiv + (compose_totally_composable fss Hfss) + (compose_totally_composable (f fss) + (monarrlist_list_proper_totally_composable Hf Hfss)). +Proof. + destruct (proj1 monarrlist_list_proper_iff Hf) as [Hc Hprop]. + apply Hprop. +Qed. + + +Lemma monarrlist_list_proper_id : + monarrlist_list_proper id. +Proof. + unfold id. + intros fss Hfss; exists Hfss; easy. +Qed. + +Lemma monarrlist_list_proper_compose' {f g} : + monarrlist_list_proper f -> + monarrlist_list_proper g -> + monarrlist_list_proper (fun fss => g (f fss)). +Proof. + intros Hf Hg fss Hfss. + destruct (Hf fss Hfss) as [Hcf Heqf]. + destruct (Hg (f fss) Hcf) as [Hcg Heqg]. + exists Hcg. + eapply monarr_norm_equiv_trans; eassumption. +Qed. + +Lemma monarrlist_list_proper_compose {f g} : + monarrlist_list_proper f -> + monarrlist_list_proper g -> + monarrlist_list_proper (Basics.compose g f). +Proof. + unfold Basics.compose. + apply monarrlist_list_proper_compose'. +Qed. + +Lemma monarrlist_list_proper_iter {f} (Hf : monarrlist_list_proper f) n : + monarrlist_list_proper (Nat.iter n f). +Proof. + induction n. + - apply monarrlist_list_proper_id. + - apply monarrlist_list_proper_compose; + [apply IHn | apply Hf]. +Qed. + +Lemma monarrlist_list_proper_iter' {f} (Hf : monarrlist_list_proper f) + (fn : list monarrlist -> nat) : + monarrlist_list_proper (fun fs => Nat.iter (fn fs) f fs). +Proof. + intros fs Hfs. + apply monarrlist_list_proper_iter, Hf. +Qed. + +End monarrlist_list_proper. + +Section pairify. + +Import Lia. + +Definition compble_lengths (fs gs : monarrlist) := + length (bw_to_varlist (monarrlist_target fs)) + = length (bw_to_varlist (monarrlist_source gs)). + +Definition compble_lengths_dec (fs gs : monarrlist) : + {compble_lengths fs gs} + {~ compble_lengths fs gs} := + PeanoNat.Nat.eq_dec + (length (bw_to_varlist (monarrlist_target fs))) + (length (bw_to_varlist (monarrlist_source gs))). + +Fixpoint pairify_helper (facc gacc fs gs : monarrlist) : + list (monarrlist * monarrlist) := + match fs, gs with + | nil, gs' => (facc, gacc ++ gs') :: nil + | fs', nil => (facc ++ fs', gacc) :: nil + | f::fs', g::gs' => + let facc' := facc ++ [f] in + let gacc' := gacc ++ [g] in + if (compble_lengths_dec facc' gacc') + then (facc', gacc') :: pairify_helper [] [] fs' gs' + else pairify_helper facc' gacc' fs' gs' + end. + +Definition pairify (fs gs : monarrlist) := + pairify_helper [] [] fs gs. + +Fixpoint depairify (fsgs : list (monarrlist * monarrlist)) + : monarrlist * monarrlist := + match fsgs with + | nil => (nil, nil) + | fg :: fsgs' => + (* let '(fs, gs) := depairify fsgs' in *) + (fst fg ++ fst (depairify fsgs'), snd fg ++ snd (depairify fsgs')) + end. + +Definition pairify_apply + (f : monarrlist -> monarrlist -> monarrlist * monarrlist) + (fs : monarrlist) (gs : monarrlist) : monarrlist * monarrlist := + depairify (map (uncurry f) (pairify fs gs)). + +Lemma app_inj_len {A} {x1 x2 y1 y2 : list A} : + length x1 = length y1 -> x1 ++ x2 = y1 ++ y2 -> + x1 = y1 /\ x2 = y2. +Proof. + intros Hlen Heq. + destruct (app_eq_app _ _ _ _ Heq) as (lnil & H). + destruct lnil; cycle 1. + - destruct H as [H | H]; + destruct H as [Hl Hr]; + apply (f_equal (@length A)) in Hl; + apply (f_equal (@length A)) in Hr; + rewrite !app_length in *; + simpl in *; lia. + - rewrite !app_nil_r in *; simpl in *. + destruct H as [H | H]; + destruct H; easy. +Qed. + +Lemma composable_of_composable_lengths {fs gs fss gss : monarrlist} + (Hc : composable (fs ++ fss) (gs ++ gss)) + (Hlen : compble_lengths fs gs) : + composable fs gs /\ composable fss gss. +Proof. + unfold composable in *. + rewrite Nf_monarrlist_target_app, + Nf_monarrlist_source_app in Hc. + rewrite bwnorm_eq_iff_varlist_eq in Hc. + rewrite 2!bwnorm_to_varlist_app in Hc. + unfold compble_lengths in Hlen. + apply app_inj_len in Hc. + - destruct Hc as [Hl Hr]. + apply bwnorm_to_varlist_inj in Hl, Hr. + easy. + - rewrite 2!bwnorm_to_varlist_Nf. + easy. +Qed. + +Lemma app_cons {A} (l1 l2 : list A) a : + l1 ++ a :: l2 = (l1 ++ [a]) ++ l2. +Proof. + rewrite <- app_assoc; easy. +Qed. + +Lemma pairify_helper_composable (facc gacc fs gs : monarrlist) + (H : composable (facc ++ fs) (gacc ++ gs)) : + Forall (uncurry composable) (pairify_helper facc gacc fs gs). +Proof. + revert gs facc gacc H; + induction fs as [| f fs' Hfs]; + intros gs facc gacc H; + [|destruct gs as [|g gs']]. + - constructor; [|easy]. + rewrite app_nil_r in H. + apply H. + - constructor; [|easy]. + rewrite app_nil_r in H. + apply H. + - simpl. + destruct (compble_lengths_dec (facc ++ [f]) (gacc ++ [g])) as [Hc | Hnc]. + + rewrite (app_cons gacc), app_cons in H. + destruct (composable_of_composable_lengths H Hc). + constructor; [easy|]. + apply Hfs; easy. + + apply Hfs. + rewrite <- 2!app_assoc. + apply H. +Qed. + +Lemma pairify_composable (fs gs : monarrlist) : + composable fs gs -> + Forall (uncurry composable) (pairify fs gs). +Proof. + intros H. + apply pairify_helper_composable. + apply H. +Qed. + +Lemma depairify_composable fsgs + (H : Forall (uncurry composable) fsgs) : + composable (fst (depairify fsgs)) (snd (depairify fsgs)). +Proof. + induction fsgs as [|[f g] fsgs IHfsgs]; [easy|]. + simpl. + destruct (depairify fsgs) as (fs, gs). + simpl. + inversion H; subst. + apply composable_app; [assumption|]. + apply IHfsgs; assumption. +Qed. + +Lemma pairify_apply_composable f (Hc : forall fs gs, + composable fs gs -> composable (fst (f fs gs)) (snd (f fs gs))) : + forall fs gs, composable fs gs -> + composable (fst (pairify_apply f fs gs)) (snd (pairify_apply f fs gs)). +Proof. + intros fs gs Hfg. + unfold pairify_apply. + apply depairify_composable. + apply (forall_composable_map Hc). + apply pairify_composable; assumption. +Qed. + +Lemma depairify_apply_compat f (Hf : binary_monarrlist_proper f) + (Hc : forall fs gs, composable fs gs -> composable (fst (f fs gs)) (snd (f fs gs))) + (Hprop : forall fs gs (Hfg : composable fs gs) + (Hc' : forall fs gs, composable fs gs -> composable (fst (f fs gs)) (snd (f fs gs))), + monarr_norm_equiv + (compose_composable fs gs Hfg) + (compose_composable (fst (f fs gs)) (snd (f fs gs)) + (Hc' fs gs Hfg))) + fsgs (Hfsgs : Forall (uncurry composable) fsgs) : + monarr_norm_equiv + (compose_composable + (fst (depairify fsgs)) (snd (depairify fsgs)) + (depairify_composable fsgs Hfsgs)) + + (compose_composable + _ _ + (depairify_composable _ + (forall_composable_map Hc Hfsgs))). +Proof. + induction fsgs as [|[fs gs] fsgs IHfsgs]. + - apply monarr_norm_equiv_of_monarrequiv. + unfold compose_composable. + simpl. + rewrite !monarr_arrcomp. + apply monarr_struct. + - simpl. + inversion Hfsgs; subst. + simpl in * |-. + eapply monarr_norm_equiv_trans. + apply (@compose_composable_app _ _ _ fs gs ltac:(assumption)). + eapply monarr_norm_equiv_trans. + 2: { + apply monarr_norm_equiv_symm. + apply (compose_composable_app (Hc fs gs ltac:(easy))). + } + apply monarr_norm_equiv_tens. + + erewrite compose_composable_indep. + apply monarr_norm_equiv_symm. + erewrite compose_composable_indep. + apply monarr_norm_equiv_symm. + apply (IHfsgs ltac:(easy)). + + apply Hprop. +Qed. + +Lemma depairify_pairify_helper_id {fs fss gs gss} + (H : composable (fs ++ fss) (gs ++ gss)) : + depairify (pairify_helper fs gs fss gss) = (fs ++ fss, gs ++ gss). +Proof. + revert fs gs gss H. + induction fss; intros fs gs gss H; + [|destruct gss]. + - simpl. + now rewrite !app_nil_r. + - simpl. + now rewrite !app_nil_r. + - simpl. + change (fs ++ a :: fss) with (fs ++ [a] ++ fss) in *. + change (gs ++ m :: gss) with (gs ++ [m] ++ gss) in *. + rewrite 2!app_assoc in *. + destruct (compble_lengths_dec (fs ++ [a]) (gs ++ [m])) as [c | c]. + + simpl. + pose proof (composable_of_composable_lengths H c) as [Hc Hc']. + change fss with (nil ++ fss) in Hc'. + change gss with (nil ++ gss) in Hc'. + rewrite (IHfss _ _ _ Hc'). + easy. + + rewrite (IHfss _ _ _ H). + easy. +Qed. + +Lemma depairify_pairify_id {fss gss} + (H : composable fss gss) : + depairify (pairify fss gss) = (fss, gss). +Proof. + unfold pairify. + rewrite depairify_pairify_helper_id; easy. +Qed. + +Lemma depairify_pairify_monarr_norm_equiv {fss gss} + (H : composable fss gss) : + monarr_norm_equiv + (compose_composable fss gss H) + (compose_composable + (fst (depairify (pairify fss gss))) + (snd (depairify (pairify fss gss))) + (depairify_composable _ (pairify_composable _ _ H))). +Proof. + generalize + (depairify_composable (pairify fss gss) (pairify_composable fss gss H)). + rewrite (depairify_pairify_id H). + simpl. + intros ?. + erewrite compose_composable_indep. + easy. +Qed. + +Lemma pairify_apply_proper f (Hf : binary_monarrlist_proper f) : + binary_monarrlist_proper (pairify_apply f). +Proof. + pose proof Hf as Hf'. + rewrite binary_proper_iff_composable_proper in Hf. + apply binary_proper_of_composable_proper. + - intros fs gs Hfg. + apply pairify_apply_composable; [apply Hf | apply Hfg]. + - intros fs gs Hfg Hc. + eapply monarr_norm_equiv_trans. + apply depairify_pairify_monarr_norm_equiv. + apply monarr_norm_equiv_symm. + erewrite compose_composable_indep. + apply monarr_norm_equiv_symm. + erewrite compose_composable_indep. + apply (depairify_apply_compat _ Hf'). + intros. + apply Hf. + Unshelve. + + apply pairify_composable, Hfg. + + apply Hf. +Qed. + +End pairify. + +End ProperOperation. diff --git a/ViCaR/MonoidalCoherence/UIP_facts.v b/ViCaR/MonoidalCoherence/UIP_facts.v new file mode 100644 index 0000000..48cffcf --- /dev/null +++ b/ViCaR/MonoidalCoherence/UIP_facts.v @@ -0,0 +1,532 @@ +Require MCDefinitions List. +Import MCDefinitions.MCClasses. + +Section Eq_rect_facts. + +Set Universe Polymorphism. + +Lemma eq_rect_comp {A} (x : A) (P : A -> Type) y z p Hxy Hyz : + eq_rect y P (eq_rect x P p y Hxy) z Hyz = + eq_rect x P p z (eq_trans Hxy Hyz). +Proof. + revert z Hyz. + case Hxy. + intros z Hyz. + case Hyz. + easy. +Qed. + +End Eq_rect_facts. + + +Section UIP_list. + +Set Universe Polymorphism. + +Context {X : Type} {UIPX : UIP X}. + +Import List.ListNotations. + +Local Open Scope list_scope. + +Definition list_eq_proj_ty (a b : list X) : Prop := + match a, b with + | nil, nil => True + | x :: a', y :: b' => x = y /\ a' = b' + | _, _ => False + end. + +Definition list_eq_proj {a : list X} : forall {b} (H : a = b), + list_eq_proj_ty a b. + refine ( + match a with + | nil => fun b => + match b with + | nil => fun _ => Logic.I + | y :: b' => fun H => False_ind _ _ + end + | x :: a' => fun b => + match b with + | nil => fun H => False_ind _ _ + | y :: b' => fun H => + conj + (f_equal (fun c => match c with + | nil => x + | z :: c' => z + end) H) + (f_equal (fun c => match c with + | nil => nil + | z :: c' => c' + end) H) + end + end); + congruence. +Defined. + +Definition list_eq_proj_inv {a : list X} : forall {b} + (H : list_eq_proj_ty a b), a = b := + match a with + | nil => fun b => + match b with + | nil => fun _ => eq_refl + | y :: b' => False_ind _ + end + | x :: a' => fun b => + match b with + | nil => False_ind _ + | y :: b' => fun H => + eq_trans + (f_equal (cons x) (proj2' H)) + (f_equal (fun z => cons z b') (proj1' H)) + end + end. + +Lemma list_nu_linv {a b} (H : a = b) : + list_eq_proj_inv (list_eq_proj H) = H. +Proof. + case H, a; easy. +Qed. + +Lemma list_eq_proj_const_of_UIP_on {a b : list X} (y : X) + (HUIP : EqdepFacts.UIP_refl_on_ (list X) b) + (H H' : a = y :: b) : + list_eq_proj H = list_eq_proj H'. +Proof. + destruct a; [easy|]. + simpl. + f_equal. + - apply uip. + - pose proof (eq_sym (proj2 (list_eq_proj H))) as Hs. + induction Hs. + etransitivity; + [|symmetry]; + apply HUIP. +Qed. + +Lemma UIP_refl_list : EqdepFacts.UIP_refl_ (list X). +Proof. + intros a; induction a; intros H; + match goal with + |- ?LHS = ?RHS => + rewrite <- (list_nu_linv LHS), <- (list_nu_linv RHS); f_equal + end. + apply (list_eq_proj_const_of_UIP_on _ IHa). +Qed. + +Lemma UIP_list : UIP (list X). +Proof. + constructor. + intros a b H H'. + pose proof H as Hs; + induction Hs. + etransitivity; + [|symmetry]; + apply UIP_refl_list. +Qed. + +End UIP_list. + +Section UIP_bw. + +Set Universe Polymorphism. + +Import MCDefinitions. + +Context {X : Type} {UIPX : UIP X}. + +Local Notation bw := (@bw X). +Local Notation bwnorm := (@bwnorm X). + +(* Definition bwnorm_norm_e_eq (a : bwnorm) : {a = norm_e} + {~ a = norm_e}. +destruct a; (left; reflexivity) + (right; congruence). +Defined. + +Definition norm_e_nu {b : bwnorm} : forall {a : bwnorm} (H : a = b), a = b := + match b with + | norm_e => + fun a H => + match bwnorm_norm_e_eq a with + | left Heq => Heq + | right Hneq => False_ind _ (Hneq H) + end + | _ => fun a H => H + end. + +Definition norm_e_nu_inv {b : bwnorm} : forall {a : bwnorm} (H : a = b), a = b := + match b with + | norm_e => fun a H => eq_ind _ (fun a => a = norm_e) H _ eq_refl + | _ => fun a H => H + end. + +Lemma norm_e_nu_linv {a b : bwnorm} (H : a = b) : + norm_e_nu_inv (norm_e_nu H) = H. +Proof. + case H. + destruct a; easy. +Qed. + +Lemma norm_e_nu_const {a : bwnorm} (H H' : a = norm_e) : + norm_e_nu H = norm_e_nu H'. +Proof. + unfold norm_e_nu. + destruct (bwnorm_norm_e_eq a); easy. +Qed. + +Lemma norm_e_UIP {a : bwnorm} (H H' : a = norm_e) : + H = H'. +Proof. + rewrite <- (norm_e_nu_linv H), <- (norm_e_nu_linv H'). + f_equal. + apply norm_e_nu_const. +Qed. *) + +Definition bwnorm_rtens_eq_l_r {a b : bwnorm} {x y : X} + (H : norm_rtens a x = norm_rtens b y) : + a = b /\ x = y := + conj (f_equal + (fun e => + match e with + | norm_e => a + | norm_rtens n _ => n + end) H) + (f_equal + (fun e : bwnorm => + match e with + | norm_e => x + | norm_rtens _ x0 => x0 + end) H). + +Definition bwnorm_rtens_eq_l_r_inv {a b : bwnorm} {x y : X} + (H : a = b /\ x = y) : + norm_rtens a x = norm_rtens b y := + eq_trans (f_equal (fun a => norm_rtens a x) (proj1' H)) + (f_equal (norm_rtens b) (proj2' H)). + (* f_equal2 norm_rtens (proj1' H) (proj2' H). *) + +Definition bwnorm_eq_proj_type (b a : bwnorm) : Prop := + match b with + | norm_e => + match a with + | norm_e => True + | norm_rtens a' x => + False + end + | norm_rtens b' y => + match a with + | norm_e => False + | norm_rtens a' x => + a' = b' /\ x = y + end + end. + +Definition bwnorm_proj_eq {b : bwnorm} : forall {a : bwnorm} (H : a = b), + bwnorm_eq_proj_type b a. + refine ( + match b as _b return forall a (H : a = _b), bwnorm_eq_proj_type _b a with + | norm_e => + fun a => match a as _a return forall (H : _a = norm_e), + bwnorm_eq_proj_type norm_e _a with + | norm_e => fun _ => Logic.I + | norm_rtens a' y' => _ + end + | norm_rtens b' x => + fun a => + match a as _a return forall (H : _a = norm_rtens b' x), + bwnorm_eq_proj_type (norm_rtens b' x) _a with + | norm_e => _ + | norm_rtens a' y' => fun H => + conj (f_equal + (fun e => + match e with + | norm_e => a + | norm_rtens n _ => n + end) H) + (f_equal + (fun e : bwnorm => + match e with + | norm_e => x + | norm_rtens _ x0 => x0 + end) H) + end + end); + congruence. +Defined. + +Definition bwnorm_proj_eq_inv {b : bwnorm} : forall {a : bwnorm} + (H: bwnorm_eq_proj_type b a), a = b := + match b with + | norm_e => + fun a => + match a with + | norm_e => fun _ => eq_refl + | norm_rtens _ _ => False_ind _ + end + | norm_rtens b' y => + fun a => + match a with + | norm_e => False_ind _ + | norm_rtens a' y => fun H => + bwnorm_rtens_eq_l_r_inv H + end + end. + +Lemma bwnorm_proj_eq_linv {a b : bwnorm} (H : a = b) : + bwnorm_proj_eq_inv (bwnorm_proj_eq H) = H. +Proof. + case H. + destruct a, b; easy. +Qed. + +Lemma bwnorm_proj_eq_const_norm_e {a : bwnorm} (H H' : a = norm_e) : + bwnorm_proj_eq H = bwnorm_proj_eq H'. +Proof. + generalize (bwnorm_proj_eq H) (bwnorm_proj_eq H'). + destruct a; intros [] []; easy. +Qed. + +Lemma bwnorm_proj_eq_const_of_UIP_on {a b : bwnorm} (y : X) + (HUIP : EqdepFacts.UIP_refl_on_ bwnorm b) + (H H' : a = norm_rtens b y) : + bwnorm_proj_eq H = bwnorm_proj_eq H'. +Proof. + destruct a; [easy|]. + simpl. + unfold bwnorm_rtens_eq_l_r. + f_equal; simpl. + - pose proof (eq_sym (proj1 (bwnorm_rtens_eq_l_r H))) as Hab. + induction Hab. + etransitivity; [|symmetry]; apply HUIP. + - apply UIPX. +Qed. + +Lemma UIP_on_norm_rtens_of_UIP_on {b : bwnorm} (y : X) + (HUIP : EqdepFacts.UIP_refl_on_ bwnorm b) : + EqdepFacts.UIP_refl_on_ bwnorm (norm_rtens b y). +Proof. + intros H. + match goal with + |- ?LHS = ?RHS => + rewrite <- (bwnorm_proj_eq_linv LHS), <- (bwnorm_proj_eq_linv RHS) + end. + f_equal. + apply bwnorm_proj_eq_const_of_UIP_on; easy. +Qed. + +Lemma UIP_refl_bwnorm (b : bwnorm) : EqdepFacts.UIP_refl_on_ bwnorm b. +Proof. + induction b. + - intros H. + match goal with + |- ?LHS = ?RHS => + rewrite <- (bwnorm_proj_eq_linv LHS), <- (bwnorm_proj_eq_linv RHS) + end. + f_equal; apply bwnorm_proj_eq_const_norm_e. + - apply UIP_on_norm_rtens_of_UIP_on, IHb. +Qed. + +#[export, program] Instance UIP_bwnorm : UIP bwnorm := {}. +Next Obligation. + etransitivity; + [|symmetry]; + apply UIP_refl_bwnorm. +Qed. + +Definition bw_eq_proj_type (b a : bw) : Prop := + match b with + | e => + match a with + | e => True + | var x => False + | tens a0 a1 => False + end + | var y => + match a with + | e => False + | var x => x = y + | tens a0 a1 => False + end + | tens b0 b1 => + match a with + | e => False + | var x => False + | tens a0 a1 => a0 = b0 /\ a1 = b1 + end + end. + +Definition bw_eq_proj {b : bw} : forall {a : bw} (H : a = b), + bw_eq_proj_type b a. + refine ( + match b as _b return forall a, a = _b -> bw_eq_proj_type _b a with + | e => fun a => + match a with + | e => fun _ => Logic.I + | var x => fun H => _ + | tens a0 a1 => fun H => _ + end + | var y => fun a => + match a as _a return _a = var y -> bw_eq_proj_type (var y) _a with + | e => fun H => _ + | var x => fun H => + f_equal (fun c : bw => + match c with + | e => x + | var z => z + | tens a' b' => x + end) H + | tens a0 a1 => fun H => _ + end + | tens b0 b1 => fun a => + match a with + | e => fun H => _ + | var x => fun H => _ + | tens a0 a1 => fun H => + conj + (f_equal (fun c : bw => + match c with + | e => a0 + | var z => a0 + | tens a' b' => a' + end) H) + (f_equal (fun c : bw => + match c with + | e => a0 + | var z => a0 + | tens a' b' => b' + end) H) + end + end); + congruence. +Defined. + +Definition bw_eq_proj_inv {b : bw} : forall {a : bw} + (H : bw_eq_proj_type b a), a = b := + match b with + | e => fun a => + match a with + | e => fun _ => eq_refl + | var x => fun H => False_ind _ H + | tens a0 a1 => fun H => False_ind _ H + end + | var y => fun a => + match a with + | e => fun H => False_ind _ H + | var x => fun H => f_equal var H + | tens a0 a1 => fun H => False_ind _ H + end + | tens b0 b1 => fun a => + match a with + | e => fun H => False_ind _ H + | var x => fun H => False_ind _ H + | tens a0 a1 => fun H => + eq_trans (f_equal (fun a => tens a a1) (proj1' H)) + (f_equal (tens b0) (proj2' H)) + end + end. + +Lemma bw_eq_proj_linv {a b : bw} (H : a = b) : + bw_eq_proj_inv (bw_eq_proj H) = H. +Proof. + case H. + destruct a, b; easy. +Qed. + +Lemma UIP_refl_bw : EqdepFacts.UIP_refl_ bw. +Proof. + intros b. + induction b. + - intros H. + rewrite <- (bw_eq_proj_linv H), <- (bw_eq_proj_linv eq_refl). + f_equal. + - intros H. + rewrite <- (bw_eq_proj_linv H), <- (bw_eq_proj_linv eq_refl). + f_equal. + apply UIPX. + - intros H. + rewrite <- (bw_eq_proj_linv H), <- (bw_eq_proj_linv eq_refl). + f_equal. + simpl. + f_equal. + + apply IHb1. + + apply IHb2. +Qed. + +#[export, program] Instance UIP_bw : UIP bw. +Next Obligation. + etransitivity; + [|symmetry]; + apply UIP_refl_bw. +Qed. + +End UIP_bw. + +(* Section UIP_isomorphism. + +Context {X Y : Type} {UIPX : UIP X} + (f : X -> Y) (g : Y -> X) + (Hfg : forall y, f (g y) = y) + (Hgf : forall x, g (f x) = x). *) + +(* Definition normalize_eq {x y : X} (H : f x = f y) : f x = f y := + eq_ind _ (fun a => f a = f y) (H y) _ _. +pose (f_equal f (f_equal g H)) as H'. +rewrite 2!Hgf in H'. +exact H'. +Defined. + +Definition normalize_eq_inv {x y : X} (H : f x = f y) : f x = f y := + eq_ind _ (fun a => a = f y) H _ (normalize_eq eq_refl). + +Lemma eq_ind_eq_rect {A} (x : A) (P : A -> Prop) (p : P x) (y : A) (H : x = y) : + eq_ind x P p y H = eq_rect x P p y H. +Proof. + easy. +Qed. + +Lemma normalize_eq_linv {x y : X} (H : f x = f y) : + normalize_eq_inv (normalize_eq H) = H. +Proof. + pose proof H as H'. + apply (f_equal g) in H'. + rewrite 2!Hgf in H'. + induction H'. + unfold normalize_eq_inv, normalize_eq. + induction H. + rewrite (UIPX.(uip) _ eq_refl). + simpl. + rewrite !eq_ind_eq_rect. + generalize (Hgf x). + generalize (g (f x)). + intros ? e; induction (eq_sym e). + rewrite 2!(eq_rect_eq (X:=X)). + revert H. + generalize (f x). + intros y H. + induction H. + + lazymatch goal with + |- context[ ?T ] => lazymatch type of T with + | @eq X _ _ => generalize T + end + end. + generalize (Hgf x) (Hgf y). + generalize (f x) (f y). + unfold normalize_eq_inv + case H. + simpl. + case (Hfg x). + simpl. + unfold eq_ind. + (* rewrite eq_rect_eq. *) + rewrite (UIPX.(uip) _ eq_refl). + + +Lemma f_equal_inj + +Lemma UIP_of_isomorphism {X Y} {UIPX : UIP X} + (f : X -> Y) (g : Y -> X) + (Hfg : forall y, f (g y) = y) + (Hgf : forall x, g (f x) = x) : + UIP Y. +Proof. + constructor. *) +