├── AXIOMS └── functional_extensionality.v ├── CAT ├── CatFunct.v ├── IO.v ├── Misc.v ├── NT.v ├── SO.v ├── bifunctor.v ├── cat_DISCRETE.v ├── cat_INDEXED_TYPE.v ├── cat_SET.v ├── cat_TYPE.v ├── cat_TYPE_option.v ├── cat_gen_monad.v ├── category.v ├── category_hom_transport.v ├── coproduct.v ├── delta.v ├── enriched_cat.v ├── eq_fibre.v ├── equiv_Monad_MonaD.v ├── free_cats.v ├── functor.v ├── functor_leibniz_eq.v ├── functor_properties.v ├── het_rewrite.v ├── horcomp.v ├── ind_potype.v ├── initial_terminal.v ├── ipo_modules.v ├── limits.v ├── module_postcomp_functor.v ├── mon_cats.v ├── monad_def.v ├── monad_h_module.v ├── monad_h_morphism.v ├── monad_h_morphism_gen.v ├── monad_haskell.v ├── monad_monad_h.v ├── monad_morphism.v ├── monic_epi.v ├── my_lib.v ├── nat_iso.v ├── order_classes.v ├── orders.v ├── orders_bis.v ├── orders_lemmas.v ├── parallel_monads.v ├── preorder_top.v ├── prods.v ├── product.v ├── pts.v ├── retype_functor.v ├── retype_functor_po.v ├── rich_monads.v ├── rmodule_TYPE.v ├── rmodule_pb_rmonad_hom.v ├── rmonad.v ├── rmonad_gen.v ├── rmonad_gen_comp.v ├── rmonad_gen_type_po.v ├── rmonad_hom.v ├── simplicial_cat.v ├── small_cat.v ├── smon_cats.v ├── subcategories.v ├── type_unit.v ├── unit_mod.v ├── unit_type_monad.v └── unit_type_rmonad.v ├── COMP ├── PCF.v ├── PCF_ULC_comp.v ├── PCF_ULC_nounit.v ├── PCF_eq.v ├── PCF_quant.v ├── PCF_rep.v ├── PCF_rep_cat.v ├── PCF_rep_cat_quant.v ├── PCF_rep_comp.v ├── PCF_rep_comp_quant.v ├── PCF_rep_eq_quant.v ├── PCF_rep_hom_quant.v ├── PCF_rep_quant.v ├── TLC.v ├── TLC_rep.v ├── TLC_rep_cat.v ├── TLC_rep_comp.v ├── ULC_PCF_rep.v ├── f_induced_monad.v └── type_arities.v ├── DESCRIPTION ├── LICENSE.md ├── Make ├── ORDER ├── pcf_order_rep.v ├── pcf_order_rep_eq.v ├── tlc.v ├── tlc_init.v ├── tlc_order_rep.v ├── tlc_order_rep_eq.v ├── tlc_properties.v ├── ulc.v ├── ulc_init.v ├── ulc_order_rep.v └── ulc_order_rep_eq.v ├── PCF ├── PCF_Monad.v ├── PCF_RMonad.v ├── PCF_semantics.v ├── PCF_syntax.v ├── PCF_types.v ├── pcf.v ├── pcf_mod.v ├── pcf_rep.v ├── pcf_rep_init.v ├── pcf_rich_monad.v ├── pcf_sub_rep.v ├── pcf_sub_rep_init.v ├── pcf_subpo.v ├── pcfpo.v ├── pcfpo_exp.v ├── pcfpo_mod.v ├── pcfpo_rep.v └── pcfpo_rep_init.v ├── PCF_o_c ├── RPCF_INIT.v ├── RPCF_ULC.v ├── RPCF_ULC_nounit.v ├── RPCF_ULC_sandbox1.v ├── RPCF_ULC_sandbox2.v ├── RPCF_rep.v ├── RPCF_rep_cat.v ├── RPCF_rep_comp.v ├── RPCF_rep_eq.v ├── RPCF_rep_hom.v ├── RPCF_rep_id.v ├── RPCF_syntax.v ├── RPCF_syntax_init.v ├── RPCF_syntax_rep.v ├── RTLC_rep.v ├── ULC_comp.v ├── ULC_syntax1.v ├── ULC_syntax2.v └── ULC_syntax_nounit.v ├── PCF_order_comp ├── RPCF_INIT.v ├── RPCF_ULC.v ├── RPCF_ULC_alt.v ├── RPCF_ULC_alt2.v ├── RPCF_ULC_sandbox.v ├── RPCF_rep.v ├── RPCF_rep_cat.v ├── RPCF_rep_comp.v ├── RPCF_rep_eq.v ├── RPCF_rep_hom.v ├── RPCF_rep_id.v ├── RPCF_syntax.v ├── RPCF_syntax_init.v ├── RPCF_syntax_rep.v ├── RTLC_rep.v ├── ULC_comp.v └── ULC_syntax.v ├── PROP_untyped ├── arities.v ├── ex_ulcbeta.v ├── ex_ulcbeta_variant.v ├── inequations_syntactically.v ├── initial.v ├── prop_arities.v ├── prop_arities_initial.v ├── prop_arities_initial_variant.v └── representations.v ├── README.md ├── RPCF ├── RPCF_INIT.v ├── RPCF_ULC_nounit.v ├── RPCF_rep.v ├── RPCF_rep_cat.v ├── RPCF_rep_comp.v ├── RPCF_rep_eq.v ├── RPCF_rep_hom.v ├── RPCF_rep_id.v ├── RPCF_syntax_init.v ├── RPCF_syntax_rep.v └── ULC_comp.v ├── STS ├── STS_arities.v ├── STS_ex.v ├── STS_initial.v ├── STS_initial_nested.v └── STS_representations.v ├── SYS_F └── F_types.v ├── TLC ├── TLC_Monad.v ├── TLC_RMonad.v ├── TLC_semantics.v ├── TLC_syntax.v └── TLC_types.v └── ULC ├── ULC_Monad.v ├── ULC_RMonad.v ├── ULC_semantics.v ├── ULC_syntax.v ├── ULC_terms.v ├── ULC_terms_print.v └── subst_experiment.v /AXIOMS/functional_extensionality.v: -------------------------------------------------------------------------------- 1 | 2 | Set Implicit Arguments. 3 | Unset Strict Implicit. 4 | 5 | Axiom functional_extensionality : forall (A B : Type)(f g : A -> B), 6 | (forall x, f x = g x) -> f = g. 7 | 8 | 9 | -------------------------------------------------------------------------------- /CAT/CatFunct.v: -------------------------------------------------------------------------------- 1 | 2 | Require Export CatSem.CAT.NT. 3 | 4 | Set Implicit Arguments. 5 | Unset Strict Implicit. 6 | Unset Transparent Obligations. 7 | Unset Automatic Introduction. 8 | 9 | Section CatFunct. 10 | 11 | (** we build a category of functors and natural transformations *) 12 | 13 | Variables obC obD: Type. 14 | Variable morC : obC -> obC -> Type. 15 | Variable morD : obD -> obD -> Type. 16 | Variable C: Cat_struct morC. 17 | Variable D: Cat_struct morD. 18 | 19 | Obligation Tactic := cat; unfold vcompNT1; 20 | try (match goal with | [ |- Proper _ _ ] => 21 | unfold Proper; do 2 red end) ; 22 | cat; unfold vcompNT1 ; try rewrite assoc; 23 | repeat (match goal with [ H : forall a, _ == _ |- _ ] => 24 | rewrite H end); 25 | cat. 26 | 27 | Program Instance FunctCat_struct : 28 | @Cat_struct (Functor C D) (NT (C:=C)(D:=D)) := { 29 | mor_oid F G := EXT_NT_oid F G; 30 | id F := vidNT F; 31 | comp a b c f g := g # f 32 | }. 33 | 34 | Definition FunctCat : Cat := Build_Cat FunctCat_struct. 35 | 36 | End CatFunct. 37 | 38 | Existing Instance FunctCat_struct. 39 | 40 | Canonical Structure FunctCat. 41 | 42 | Notation "[ C , D ]" := (FunctCat C D) (at level 75). 43 | 44 | 45 | (*Definition END (C:Category) := [ C , C ].*) 46 | 47 | 48 | Section FunctCat_lemmata. 49 | 50 | (** we prove an extensionality principle on natural transformations *) 51 | 52 | Variables obC obD: Type. 53 | Variable morC : obC -> obC -> Type. 54 | Variable morD : obD -> obD -> Type. 55 | Variable C: Cat_struct morC. 56 | Variable D: Cat_struct morD. 57 | 58 | Variables F G H K: Functor C D. 59 | Variable alpha: F ---> G. 60 | Variable beta: H ---> K. 61 | 62 | Hypothesis h: F = H. 63 | Hypothesis h': G = K. 64 | 65 | Lemma NT_Extensionality: 66 | (forall x: obC, alpha x === beta x) -> alpha === beta. 67 | Proof. 68 | generalize dependent alpha; clear alpha. 69 | rewrite h. 70 | generalize dependent beta; clear beta. 71 | rewrite h'. 72 | clear h h' F G. 73 | intros beta alpha H'. 74 | apply Build_Equal_hom. 75 | simpl. 76 | intro c. 77 | apply Equal_hom_imp_setoideq2_jmq_eq. 78 | apply (H' c). 79 | Qed. 80 | 81 | End FunctCat_lemmata. 82 | 83 | 84 | 85 | 86 | 87 | 88 | 89 | 90 | 91 | 92 | -------------------------------------------------------------------------------- /CAT/Misc.v: -------------------------------------------------------------------------------- 1 | Require Export Coq.Setoids.Setoid. 2 | Require Export Coq.Classes.SetoidClass. 3 | Require Export Coq.Classes.Morphisms. 4 | 5 | Set Implicit Arguments. 6 | Unset Strict Implicit. 7 | Unset Automatic Introduction. 8 | 9 | Ltac app P := let H:= fresh "YMCA" in 10 | assert (H := P); 11 | simpl in H; 12 | apply H; 13 | clear H; simpl; auto. 14 | 15 | Ltac rew P := let H:=fresh "YMCA" in 16 | assert (H := P); 17 | simpl in H; 18 | rewrite H; 19 | clear H; 20 | simpl; 21 | auto. 22 | 23 | Ltac rerew P := let H:=fresh "YMCA" in 24 | assert (H := P); 25 | simpl in H; 26 | rewrite <- H; 27 | clear H; 28 | simpl; 29 | auto. 30 | 31 | Ltac app_all := match goal with 32 | | [H:_|-_] => apply H end. 33 | 34 | Ltac unf_Proper := match goal with 35 | | [ |- Proper _ _ ] => unfold Proper; repeat red end. 36 | 37 | Ltac rew_hyp_2 := match goal with 38 | | [ H : forall _ _ , _ = _ |- _ ] => rewrite H end. 39 | 40 | Ltac rew_hyp := match goal with 41 | | [ H : forall _ , _ = _ |- _ ] => rewrite H end. 42 | 43 | Ltac rew_setoid := match goal with 44 | | [H : forall _ , _ == _ |- _ ] => rewrite H end. 45 | 46 | Ltac rew_all := match goal with 47 | | [H : _ |- _ ] => rewrite H end. 48 | 49 | Ltac rerew_all := match goal with 50 | | [H : _ |- _ ] => rewrite <- H end. 51 | 52 | Ltac rew_set := match goal with 53 | | [H : _ == _ |- _ ] => rewrite H end. 54 | 55 | Ltac elim_unit := match goal with 56 | | [ H : unit |- _ ] => destruct H end. 57 | 58 | Ltac elim_conj := match goal with 59 | | [ H : _ /\ _ |- _ ] => destruct H; intros end. 60 | 61 | Ltac elim_conjs := repeat elim_conj. 62 | 63 | Ltac mauto := simpl; intros; try unf_Proper; simpl; intros; 64 | elim_conjs; auto. 65 | 66 | 67 | Section setoid_lemmata. 68 | 69 | Variable A:Type. 70 | Variable R: Setoid A. 71 | 72 | Lemma setoid_refl : forall a, a == a. 73 | Proof. 74 | destruct R. 75 | apply setoid_equiv. 76 | Qed. 77 | 78 | Lemma setoid_sym : forall a b, a == b -> b == a. 79 | Proof. 80 | destruct R. 81 | apply setoid_equiv. 82 | Qed. 83 | 84 | Lemma setoid_trans : forall a b c, 85 | a == b -> b == c -> a == c. 86 | Proof. 87 | destruct R. 88 | apply setoid_equiv. 89 | Qed. 90 | 91 | End setoid_lemmata. 92 | 93 | Hint Resolve setoid_refl setoid_sym setoid_trans : setoid. 94 | Hint Rewrite setoid_trans : setoid. 95 | 96 | Ltac setoid := intros; mauto; eauto with setoid. 97 | 98 | Lemma eq_trans_2 (A:Type)(a b c : A) : a = b -> b = c -> c = a. 99 | Proof. 100 | congruence. 101 | Qed. 102 | 103 | 104 | Lemma eq_eq_rect A (z : A) (P : A -> Type) 105 | (f g : P z) (y : A) (e : z = y): 106 | f = g -> eq_rect z P f y e = eq_rect z P g y e. 107 | Proof. 108 | intros; 109 | subst; 110 | auto. 111 | Qed. 112 | 113 | 114 | 115 | -------------------------------------------------------------------------------- /CAT/NT.v: -------------------------------------------------------------------------------- 1 | 2 | Require Export CatSem.CAT.functor. 3 | 4 | Set Implicit Arguments. 5 | Unset Strict Implicit. 6 | 7 | Unset Automatic Introduction. 8 | 9 | Section NT_def. 10 | 11 | Variables obC obD: Type. 12 | Variable morC : obC -> obC -> Type. 13 | Variable morD : obD -> obD -> Type. 14 | Variable C: Cat_struct morC. 15 | Variable D: Cat_struct morD. 16 | 17 | Definition trafo_comm (F G: Functor C D) 18 | (alpha: forall c:obC, morD (F c) (G c)):= 19 | forall c d f, alpha c ;; #G f == #F f ;; alpha d. 20 | 21 | Class NT_struct {F G: Functor C D} 22 | (alpha: forall c: obC, morD (F c) (G c)) := { 23 | trafo_ax: forall c d f, alpha c ;; #G f == #F f ;; alpha d 24 | }. 25 | 26 | Record NT (F G: Functor C D) := { 27 | trafo:> forall c: obC, morD (F c) (G c); 28 | nt_struct:> NT_struct trafo 29 | }. 30 | 31 | 32 | Existing Instance nt_struct. 33 | 34 | Lemma NTcomm (F G: Functor C D)(alpha: forall c, morD (F c) (G c)) 35 | (H:NT_struct alpha) 36 | (c d: obC)(f: morC c d): 37 | alpha c ;; #G f == #F f ;; alpha d. 38 | Proof. 39 | intros F G a c d f. 40 | apply trafo_ax. 41 | Qed. 42 | 43 | 44 | 45 | Variables F G: Functor C D. 46 | 47 | Definition EXT_NT : relation (NT F G) := 48 | fun a b => forall c: obC, a c == b c. 49 | 50 | Lemma EXT_NT_refl: Reflexive (EXT_NT). 51 | Proof. 52 | unfold Reflexive. 53 | intros x c. 54 | cat. 55 | Qed. 56 | 57 | Lemma EXT_NT_sym: Symmetric EXT_NT. 58 | Proof. 59 | unfold Symmetric. 60 | intros x y H c. 61 | apply hom_sym. 62 | apply H. 63 | Qed. 64 | 65 | Lemma EXT_NT_trans: Transitive EXT_NT. 66 | Proof. 67 | unfold Transitive. 68 | intros x y z H H' c. 69 | apply hom_trans with (y c); 70 | auto. 71 | Qed. 72 | 73 | Lemma eq_NT_equiv : @Equivalence (NT F G) 74 | (fun a b => forall c: obC, a c == b c). 75 | Proof. 76 | split. 77 | assert (H:=EXT_NT_refl); 78 | auto. 79 | assert (H:=EXT_NT_sym); 80 | auto. 81 | assert (H:=EXT_NT_trans); 82 | auto. 83 | Qed. 84 | 85 | 86 | Definition EXT_NT_oid := Build_Setoid eq_NT_equiv. 87 | 88 | End NT_def. 89 | 90 | Existing Instance nt_struct. 91 | 92 | (* 93 | Notation "a =NT= b" := (eq_NT1 a b) (at level 80). 94 | *) 95 | 96 | (** the vertical identity NT *) 97 | 98 | Section VIDNT. 99 | 100 | Variables obC obD: Type. 101 | Variable morC : obC -> obC -> Type. 102 | Variable morD : obD -> obD -> Type. 103 | Variable C: Cat_struct morC. 104 | Variable D: Cat_struct morD. 105 | Variable F: Functor C D. 106 | 107 | Program Definition vidNT : NT F F := Build_NT 108 | (trafo := fun c => id (F c)) _ . 109 | Next Obligation. 110 | Proof. 111 | apply Build_NT_struct. 112 | cat. 113 | Qed. 114 | 115 | End VIDNT. 116 | 117 | Section vcompNT. 118 | (* Variables C D: Category. *) 119 | Variables obC obD: Type. 120 | Variable morC : obC -> obC -> Type. 121 | Variable morD : obD -> obD -> Type. 122 | Variable C: Cat_struct morC. 123 | Variable D: Cat_struct morD. 124 | 125 | Variables F G H: Functor C D. 126 | Variables (a:NT F G) (b:NT G H). 127 | 128 | Definition vcompNT1 := fun c => a c ;; b c. 129 | 130 | Lemma vcompNT1_ax: forall c d f, 131 | vcompNT1 c ;; #H f == #F f ;; vcompNT1 d. 132 | Proof. 133 | intros c d f. 134 | unfold vcompNT1. 135 | rewrite assoc. 136 | rewrite (NTcomm b). 137 | rewrite <- assoc. 138 | rewrite (NTcomm a). 139 | rewrite assoc. 140 | apply hom_refl. 141 | Qed. 142 | 143 | Definition vcompNT := Build_NT (Build_NT_struct vcompNT1_ax). 144 | 145 | End vcompNT. 146 | 147 | 148 | Notation "a # b" := (vcompNT b a) (at level 50, left associativity). 149 | 150 | (** lemmata about vertical composition *) 151 | Section vcompNT_lemmata. 152 | Variables obC obD: Type. 153 | Variable morC : obC -> obC -> Type. 154 | Variable morD : obD -> obD -> Type. 155 | Variable C: Cat_struct morC. 156 | Variable D: Cat_struct morD. 157 | 158 | Variable F G: Functor C D. 159 | Variable alpha: NT F G. 160 | 161 | Notation "a =NT= b" := (EXT_NT a b) (at level 80). 162 | 163 | Lemma vidNTl: alpha # vidNT F =NT= alpha. 164 | Proof. 165 | unfold EXT_NT, vidNT. 166 | intro c; simpl. 167 | unfold vcompNT1; 168 | simpl. apply idl. 169 | Qed. 170 | 171 | Lemma vidNTr: vidNT G # alpha =NT= alpha. 172 | Proof. 173 | unfold EXT_NT, vidNT. 174 | intro c; simpl. 175 | unfold vcompNT1; simpl. 176 | apply idr. 177 | Qed. 178 | 179 | Variables H J: Functor C D. 180 | Variable beta: NT G H. 181 | Variable gamma: NT H J. 182 | 183 | Lemma vcompNT_assoc: 184 | gamma # (beta # alpha) =NT= (gamma # beta) # alpha. 185 | Proof. 186 | unfold EXT_NT, vcompNT. 187 | simpl. 188 | unfold vcompNT1; simpl. 189 | intro c; apply assoc. 190 | Qed. 191 | 192 | End vcompNT_lemmata. 193 | -------------------------------------------------------------------------------- /CAT/SO.v: -------------------------------------------------------------------------------- 1 | Require Export CatSem.CAT.cat_TYPE. 2 | Require Export CatSem.CAT.orders. 3 | Require Export CatSem.CAT.subcategories. 4 | Require Export CatSem.CAT.product. 5 | Require Export CatSem.CAT.initial_terminal. 6 | 7 | Set Implicit Arguments. 8 | Unset Strict Implicit. 9 | 10 | Unset Automatic Introduction. 11 | 12 | (** a category of preorders and maps (not necessarily monotone) 13 | between them 14 | *) 15 | 16 | Obligation Tactic := cat; unfold Proper; do 2 red; cat; 17 | repeat rew_all; cat. 18 | 19 | Program Instance wPO_struct : Cat_struct (fun a b : PO_obj => a -> b) := { 20 | mor_oid a b := TYPE_hom_oid a b; 21 | id a := fun x: a => x; 22 | comp a b c := fun (f:a -> b) (g:b -> c) => fun x => g (f x) 23 | }. 24 | 25 | Definition wOrd := Build_Cat wPO_struct. 26 | 27 | Class Monotone (a b : wOrd) (f : a ---> b) := 28 | monotone : Proper (Rel ==> Rel) f. 29 | 30 | Obligation Tactic := mauto ; unfold Monotone; cat; 31 | unfold Proper, respectful in *; mauto. 32 | (* 33 | Obligation Tactic := idtac. 34 | *) 35 | Program Instance wPO_monotone_SubCat : SubCat_compat wOrd 36 | (fun x => True) (fun a b f => Monotone f). 37 | 38 | Definition wPO_monotone := SubCat wPO_monotone_SubCat. 39 | 40 | Program Instance wPO_prod : Cat_Prod wOrd := { 41 | product a b := PO_product a b; 42 | prl a b := fst; 43 | prr a b := snd; 44 | prod_mor a b c f g := fun z => (f z, g z) 45 | }. 46 | 47 | Program Instance wPO_Terminal : Terminal wOrd := { 48 | Term := PO_TERM; 49 | TermMor a := fun x => tt 50 | }. 51 | 52 | Obligation Tactic := simpl; intros; 53 | repeat match goal with [H:TEMPTY |- _ ] => elim H end; 54 | cat. 55 | 56 | Program Instance TYPE_init : Initial wOrd := { 57 | Init := PO_INIT 58 | }. 59 | 60 | 61 | 62 | 63 | 64 | 65 | -------------------------------------------------------------------------------- /CAT/cat_DISCRETE.v: -------------------------------------------------------------------------------- 1 | Require Export CatSem.CAT.category. 2 | 3 | Set Implicit Arguments. 4 | Unset Strict Implicit. 5 | Unset Transparent Obligations. 6 | Unset Automatic Introduction. 7 | 8 | Section DISCRETE_CAT. 9 | 10 | Variable T : Type. 11 | 12 | (* Definition Discrete_hom (a b:T) := a = b. 13 | *) 14 | Lemma Discrete_hom_equiv (a b : T) : @Equivalence (a = b) 15 | (fun p q => True). 16 | Proof. 17 | intros a b; 18 | constructor; simpl; auto. 19 | Qed. 20 | 21 | Instance Discrete_hom_oid a b : Setoid (a = b) := 22 | Build_Setoid (Discrete_hom_equiv a b). 23 | (* 24 | Obligation Tactic := intuition; auto. 25 | *) 26 | Program Instance DISCRETE_struct : Cat_struct (obj:=T) (fun a b => a = b). 27 | (* 28 | Next Obligation. 29 | apply (Discrete_hom_oid). 30 | Defined. 31 | *) 32 | Next Obligation. 33 | Proof. 34 | unfold Proper; repeat red. 35 | auto. 36 | Qed. 37 | 38 | Definition DISCRETE := Build_Cat DISCRETE_struct. 39 | 40 | End DISCRETE_CAT. 41 | 42 | 43 | 44 | 45 | 46 | 47 | 48 | 49 | -------------------------------------------------------------------------------- /CAT/cat_SET.v: -------------------------------------------------------------------------------- 1 | 2 | Require Import Coq.Relations.Relations. 3 | 4 | Require Export Misc. 5 | Require Export CatSem.CAT.category. 6 | Require Import CatSem.CAT.product. 7 | Require Import CatSem.CAT.initial_terminal. 8 | Require Import CatSem.CAT.coproduct. 9 | 10 | Set Implicit Arguments. 11 | Unset Strict Implicit. 12 | Unset Transparent Obligations. 13 | Unset Automatic Introduction. 14 | 15 | 16 | Section SET. 17 | 18 | Section SET_data. 19 | 20 | Variable A B:Set. 21 | 22 | Definition SET_hom_equiv: 23 | relation (A -> B) := fun f g => forall x, f x = g x. 24 | 25 | Lemma SET_hom_oid_prf: @Equivalence (A -> B) 26 | (fun f g => forall x, f x = g x). 27 | Proof. 28 | constructor; 29 | unfold Reflexive; 30 | unfold Symmetric; 31 | unfold Transitive; 32 | intros; 33 | etransitivity; 34 | eauto. 35 | Qed. 36 | 37 | Definition SET_hom_oid := Build_Setoid SET_hom_oid_prf. 38 | 39 | End SET_data. 40 | 41 | Obligation Tactic := cat; try unf_Proper; 42 | intros; repeat rew_hyp; auto. 43 | 44 | Program Instance SET_catstruct : Cat_struct (fun a b : Set => a -> b) := { 45 | mor_oid a b := SET_hom_oid a b; 46 | id a := fun x: a => x; 47 | comp a b c f g := fun x => g (f x) 48 | }. 49 | 50 | Definition SET := Build_Cat SET_catstruct. 51 | 52 | End SET. 53 | 54 | Section SET_INIT_TERM. 55 | 56 | Inductive Empty : Set := . 57 | 58 | Obligation Tactic := simpl; intros; 59 | repeat match goal with [H:Empty |- _ ] => elim H end; 60 | cat. 61 | 62 | Program Instance SET_INIT : Initial SET := { 63 | Init := Empty }. 64 | 65 | Hint Extern 3 (?a = ?b) => elim a. 66 | 67 | Program Instance SET_TERM : Terminal SET := { 68 | Term := unit; 69 | TermMor A := fun x => tt 70 | }. 71 | 72 | End SET_INIT_TERM. 73 | 74 | Section SET_COPROD. 75 | 76 | Set Implicit Arguments. 77 | Set Contextual Implicit. 78 | 79 | 80 | 81 | Inductive SET_COPROD_ob (A B: Set): Set := 82 | | INL : A -> SET_COPROD_ob A B 83 | | INR : B -> SET_COPROD_ob A B. 84 | 85 | (* Global Arguments INL [_ _] _. *) 86 | (* Global Arguments INR [_ _] _. *) 87 | 88 | Obligation Tactic := simpl; intros; 89 | repeat unf_Proper; 90 | intros; 91 | repeat match goal with 92 | [x : SET_COPROD_ob _ _ |- _ ] => destruct x end; 93 | elim_conjs; 94 | auto. 95 | 96 | Program Instance SET_COPROD : Cat_Coprod SET := { 97 | coprod a b := SET_COPROD_ob a b; 98 | inl a b x := INL (A:=a) (B:=b) x; 99 | inr a b x := INR (A:=a) (B:=b) x; 100 | coprod_mor a b d f g := 101 | fun x => match x with INL a => f a | 102 | INR b => g b end 103 | }. 104 | 105 | End SET_COPROD. 106 | 107 | Section SET_PROD. 108 | 109 | Obligation Tactic := simpl; intros; 110 | try unf_Proper; 111 | intros; elim_conjs; 112 | repeat rew_hyp; 113 | cat. 114 | 115 | Hint Extern 1 (_ = _ ) => apply injective_projections; simpl. 116 | 117 | Program Instance SET_PRODUCT: Cat_Prod (SET) := { 118 | product a b := prod a b; 119 | prl a b x := fst x; 120 | prr a b x := snd x; 121 | prod_mor a c d f g := fun x => (f x, g x) 122 | }. 123 | 124 | End SET_PROD. 125 | 126 | 127 | 128 | 129 | 130 | 131 | 132 | 133 | 134 | 135 | 136 | 137 | 138 | 139 | 140 | 141 | 142 | -------------------------------------------------------------------------------- /CAT/cat_TYPE.v: -------------------------------------------------------------------------------- 1 | Require Import Coq.Relations.Relations. 2 | 3 | Require Export CatSem.CAT.Misc. 4 | Require Export CatSem.CAT.product. 5 | Require Export CatSem.CAT.initial_terminal. 6 | 7 | Set Implicit Arguments. 8 | Unset Strict Implicit. 9 | Unset Transparent Obligations. 10 | 11 | Unset Automatic Introduction. 12 | 13 | Section TYPE. 14 | 15 | Section TYPE_data. 16 | 17 | Variable A B:Type. 18 | (*Definition TYPE_hom := A -> B.*) 19 | (*Definition TYPE_hom_equiv: 20 | relation := fun f g => forall x, f x = g x.*) 21 | 22 | Lemma TYPE_hom_oid_prf: 23 | Equivalence (A:= A -> B) (fun f g => forall x, f x = g x). 24 | Proof. 25 | constructor; 26 | unfold Reflexive; 27 | unfold Symmetric; 28 | unfold Transitive; 29 | intros; 30 | etransitivity; 31 | eauto. 32 | Qed. 33 | 34 | Definition TYPE_hom_oid := Build_Setoid TYPE_hom_oid_prf. 35 | 36 | End TYPE_data. 37 | 38 | Obligation Tactic := cat; try unf_Proper; 39 | intros; repeat rew_hyp; auto. 40 | 41 | Program Instance TYPE_struct : Cat_struct (fun a b => a -> b) := { 42 | mor_oid a b := TYPE_hom_oid a b; 43 | id a := fun x: a => x; 44 | comp a b c := fun (f:a -> b) (g:b -> c) => fun x => g (f x) 45 | }. 46 | 47 | Definition TYPE: Cat := Build_Cat 48 | TYPE_struct. 49 | 50 | End TYPE. 51 | 52 | Canonical Structure TYPE. 53 | 54 | Obligation Tactic := simpl; intros; 55 | try unf_Proper; 56 | intros; elim_conjs; 57 | repeat rew_hyp; 58 | cat. 59 | 60 | Hint Extern 1 (_ = _ ) => apply injective_projections; simpl. 61 | 62 | 63 | Program Instance TYPE_prod : Cat_Prod TYPE := { 64 | product := prod; 65 | prl a b := fst; 66 | prr a b := snd; 67 | prod_mor a b c f g := fun z => (f z, g z) 68 | }. 69 | 70 | Inductive TEMPTY : Type := . 71 | 72 | Obligation Tactic := simpl; intros; 73 | repeat match goal with [H:TEMPTY |- _ ] => elim H end; 74 | cat. 75 | 76 | Program Instance TYPE_init : Initial TYPE := { 77 | Init := TEMPTY 78 | }. 79 | 80 | Hint Extern 3 (?a = ?b) => elim a. 81 | 82 | Program Instance TYPE_term : Terminal TYPE := { 83 | Term := unit; 84 | TermMor a := fun x => tt 85 | }. 86 | 87 | Existing Instance TYPE_struct. 88 | 89 | 90 | 91 | 92 | 93 | 94 | -------------------------------------------------------------------------------- /CAT/cat_TYPE_option.v: -------------------------------------------------------------------------------- 1 | Require Export CatSem.CAT.cat_TYPE. 2 | Require Export CatSem.CAT.monad_haskell. 3 | Require Export CatSem.CAT.rmonad. 4 | 5 | Set Implicit Arguments. 6 | Unset Strict Implicit. 7 | Unset Automatic Introduction. 8 | Unset Transparent Obligations. 9 | 10 | Obligation Tactic := simpl; intros; 11 | try unf_Proper; intros; 12 | repeat match goal with [H : option _ |- _ ] => destruct H end; 13 | auto. 14 | 15 | Program Instance option_monad_s : 16 | Monad_struct (C:=TYPE) (option) := { 17 | weta := @Some ; 18 | kleisli a b f := fun t => match t with 19 | | Some y => f y 20 | | None => None 21 | end 22 | }. 23 | 24 | Definition option_monad := Build_Monad option_monad_s. 25 | 26 | Definition option_default (V W : TYPE) (f : V ---> W) 27 | (w: W): option V ---> W := fun v => 28 | match v with 29 | | None => w 30 | | Some v' => (f v') 31 | end. 32 | 33 | 34 | 35 | -------------------------------------------------------------------------------- /CAT/cat_gen_monad.v: -------------------------------------------------------------------------------- 1 | Require Export CatSem.CAT.monad_h_morphism_gen. 2 | 3 | Set Implicit Arguments. 4 | Unset Strict Implicit. 5 | Unset Transparent Obligations. 6 | Unset Automatic Introduction. 7 | 8 | 9 | 10 | Section id. 11 | 12 | Variable C : Cat. 13 | Variable P : Monad C. 14 | 15 | Obligation Tactic := cat. 16 | 17 | Program Instance colax_Monad_Hom_id_s : colax_Monad_Hom_struct (P:=P) 18 | (F0 := Id C) 19 | (fun c => id _ ). 20 | 21 | Definition colax_Monad_Hom_id := Build_colax_Monad_Hom colax_Monad_Hom_id_s. 22 | 23 | End id. 24 | 25 | Section comp. 26 | 27 | Variable C : Cat. 28 | Variable P : Monad C. 29 | Variable D : Cat. 30 | Variable Q : Monad D. 31 | Variable E : Cat. 32 | Variable R : Monad E. 33 | 34 | Variable F : Functor C D. 35 | Variable S : colax_Monad_Hom P Q F. 36 | Variable G : Functor D E. 37 | Variable T : colax_Monad_Hom Q R G. 38 | 39 | Program Instance colax_Monad_Hom_comp_s : colax_Monad_Hom_struct (P:=P) 40 | (F0 := G O F) 41 | (fun c => #G(S c) ;; T (F c)). 42 | Next Obligation. 43 | Proof. 44 | repeat rewrite <-assoc. 45 | rewrite <- FComp. 46 | rewrite <- FComp. 47 | rewrite assoc. 48 | rerew (gen_monad_hom_kl (colax_Monad_Hom_struct := T)). 49 | rewrite <- assoc. 50 | rewrite <- FComp. 51 | rerew (gen_monad_hom_kl (colax_Monad_Hom_struct := S)). 52 | apply hom_refl. 53 | Qed. 54 | Next Obligation. 55 | Proof. 56 | rewrite <- assoc. 57 | rewrite <- FComp. 58 | rewrite (gen_monad_hom_weta (colax_Monad_Hom_struct := S)). 59 | rewrite (gen_monad_hom_weta (colax_Monad_Hom_struct := T)). 60 | apply hom_refl. 61 | Qed. 62 | 63 | Definition colax_Monad_Hom_comp := 64 | Build_colax_Monad_Hom colax_Monad_Hom_comp_s. 65 | 66 | End comp. 67 | 68 | Section eq. 69 | 70 | Variable C : Cat. 71 | Variable P : Monad C. 72 | Variable D : Cat. 73 | Variable Q : Monad D. 74 | 75 | 76 | Record gen_Monad_Hom_type := { 77 | base_f : Functor C D ; 78 | gen_mon_hom :> colax_Monad_Hom P Q base_f 79 | }. 80 | 81 | (* 82 | Section eq_rel. 83 | 84 | Variable a b : gen_Monad_Hom_type. 85 | 86 | Record eq_gen_Monad_Hom_type : Prop := { 87 | b_NT : NatIso (base_f a)(base_f b) ; 88 | b_comm : forall c, b_NT (P c) ;; b c == a c ;; lift (b_NT c) 89 | }. 90 | 91 | End eq_rel. 92 | 93 | Lemma eq_gen_eq : Equivalence eq_gen_Monad_Hom_type. 94 | Proof. 95 | constructor. 96 | unfold Reflexive. 97 | intros. 98 | apply (@Build_eq_gen_Monad_Hom_type _ _ (vidNT _ )). 99 | simpl. 100 | cat. 101 | unfold Symmetric. 102 | intros. destruct H. 103 | simpl in *. 104 | 105 | Definition eq_gen_Monad_Hom_type : relation gen_Monad_Hom_type := 106 | fun a b => 107 | Section eq. 108 | 109 | Variable C : Cat. 110 | Variable P : Monad C. 111 | Variable D : Cat. 112 | Variable Q : Monad D. 113 | 114 | Lemma gen_Monad_Hom_eq : 115 | Equivalence (A:=gen_Monad_Hom_type P Q) 116 | (fun S T => forall c, S c == T c). 117 | 118 | 119 | Variable F : Functor C D. 120 | Variable F' : Functor C D. 121 | 122 | Inductive eq_gen_Monad_Hom 123 | 124 | Lemma gen_Monad_Hom_eq : 125 | Equivalence (A:=gen_Monad_Hom P Q F) 126 | (fun S T => forall c, S c == T c). 127 | Proof. 128 | constructor. 129 | unfold Reflexive. 130 | cat. 131 | unfold Symmetric. 132 | intros. 133 | rew_all; cat. 134 | unfold Transitive. 135 | intros. 136 | rew_all. 137 | rew_all. 138 | cat. 139 | Qed. 140 | 141 | Definition gen_Monad_Hom_oid := Build_Setoid gen_Monad_Hom_eq. 142 | 143 | End eq. 144 | 145 | Program Instance gen_Monad_s : Cat 146 | 147 | rewrite <- assoc. 148 | cat. 149 | 150 | apply Foid. 151 | rerew (FComp G). 152 | simpl. 153 | cat. 154 | 155 | 156 | (Tau : forall c, F (P c) ---> Q (F c)) := { 157 | gen_monad_hom_kl : forall c d (f : c ---> P d), 158 | #F (kleisli f) ;; Tau _ == 159 | Tau _ ;; (kleisli (#F f ;; Tau _ )) ; 160 | gen_monad_hom_weta : forall c : C, 161 | #F (weta c) ;; Tau _ == weta _ 162 | }. 163 | 164 | 165 | Variable D : Cat. 166 | Variable Q : Monad D. 167 | Variable F : Functor C D 168 | 169 | *) 170 | 171 | End eq. -------------------------------------------------------------------------------- /CAT/category.v: -------------------------------------------------------------------------------- 1 | 2 | Require Export Coq.Setoids.Setoid. 3 | Require Export Coq.Classes.SetoidClass. 4 | 5 | Require Export Misc. 6 | 7 | Set Implicit Arguments. 8 | Unset Strict Implicit. 9 | Unset Transparent Obligations. 10 | (* for coq 8.3 rc1 *) 11 | 12 | Unset Automatic Introduction. 13 | 14 | 15 | 16 | (** Categories *) 17 | 18 | (** a categorical structure Cat on "obj" and "mor" consists of 19 | - a setoid on each mor-set 20 | - an id mor 21 | - a composition which is compatible with setoid struct 22 | - id properties 23 | - associativity 24 | *) 25 | Class Cat_struct (obj:Type)(mor: obj -> obj -> Type) := { 26 | mor_oid:> forall a b, Setoid (mor a b); 27 | id: forall a, mor a a; 28 | comp: forall {a b c}, 29 | mor a b -> mor b c -> mor a c; 30 | comp_oid:> forall a b c, 31 | Proper (equiv ==> equiv ==> equiv) (@comp a b c); 32 | id_r: forall a b (f: mor a b), 33 | comp f (id b) == f; 34 | id_l: forall a b (f: mor a b), 35 | comp (id a) f == f; 36 | Assoc: forall a b c d (f: mor a b) 37 | (g:mor b c) (h: mor c d), 38 | comp (comp f g) h == comp f (comp g h) 39 | }. 40 | 41 | 42 | 43 | Notation "x ';;' y" := (comp x y) (at level 50, left associativity). 44 | 45 | (** Category *) 46 | 47 | Record Cat := { 48 | obj :> Type; 49 | mor : obj -> obj -> Type; 50 | cat_struct :> Cat_struct mor}. 51 | 52 | Notation "x ---> y" := (mor x y) (at level 50). 53 | 54 | Existing Instance cat_struct. 55 | 56 | 57 | (** Trivial lemmata for easier use in proofs *) 58 | 59 | Section cat_lemmata. 60 | 61 | Variable ob: Type. 62 | Variable hom: ob -> ob -> Type. 63 | Variable C: Cat_struct hom. 64 | Variables a b: ob. 65 | Variable f: hom a b. 66 | 67 | Lemma idr: f ;; id b == f. 68 | Proof. 69 | apply (id_r). 70 | Qed. 71 | 72 | Lemma idl: id a ;; f == f. 73 | Proof. 74 | apply (id_l). 75 | Qed. 76 | 77 | Variables (c d: ob) (g: hom b c) (h: hom c d). 78 | 79 | Lemma assoc: (f ;; g) ;; h == f ;; (g ;; h). 80 | Proof. 81 | apply (Assoc). 82 | Qed. 83 | 84 | (** helpful lemmata, equivalence properties for the morphisms *) 85 | 86 | Lemma hom_refl : f == f. 87 | Proof. 88 | apply (Equivalence_Reflexive). 89 | Qed. 90 | 91 | Variables f' f'': hom a b. 92 | 93 | Lemma hom_sym : f == f' -> f' == f. 94 | Proof. 95 | apply (Equivalence_Symmetric). 96 | Qed. 97 | 98 | 99 | Lemma hom_trans : 100 | f == f' -> f' == f'' -> f == f''. 101 | Proof. 102 | apply (Equivalence_Transitive). 103 | Qed. 104 | 105 | 106 | Lemma praecomp (g' : hom b c): 107 | g == g' -> f ;; g == f ;; g'. 108 | Proof. 109 | intros; repeat rew_set; 110 | apply (Equivalence_Reflexive). 111 | Qed. 112 | 113 | Lemma postcomp (g' : hom b c): 114 | g == g' -> g ;; h == g' ;; h. 115 | Proof. 116 | intros; repeat rew_set; 117 | apply (Equivalence_Reflexive). 118 | Qed. 119 | 120 | End cat_lemmata. 121 | 122 | 123 | Hint Resolve hom_refl idl idr 124 | praecomp postcomp : cat. 125 | 126 | Hint Rewrite idl idr : cat. 127 | 128 | Ltac cat := simpl in *; 129 | intros; autorewrite with cat; 130 | auto with cat; 131 | try reflexivity. 132 | 133 | 134 | (** equality predicate for ALL morphisms of a category *) 135 | 136 | Section Equal_hom. 137 | 138 | Variable obC:Type. 139 | Variable morC : obC -> obC -> Type. 140 | Variable C: Cat_struct morC. 141 | 142 | Inductive Equal_hom (a b: obC) (f : morC a b) : 143 | forall c d : obC, (morC c d) -> Prop := 144 | Build_Equal_hom : forall g : morC a b, f == g -> Equal_hom f g. 145 | 146 | 147 | Notation "f === g" := (Equal_hom f g) (at level 55). 148 | 149 | (** Equal_hom has Equivalence properties *) 150 | 151 | Hint Extern 2 (_ === _) => constructor. 152 | 153 | Lemma Equal_hom_refl (a b: obC) (f: morC a b) : 154 | f === f. 155 | Proof. 156 | cat. 157 | Qed. 158 | 159 | Lemma Equal_hom_sym (a b c d: obC) (f: morC a b) (g: morC c d): 160 | f === g -> g === f. 161 | Proof. 162 | intros; 163 | match goal with [H : _ === _ |- _ ] => elim H end; 164 | cat; 165 | auto using hom_sym. 166 | Qed. 167 | 168 | Lemma Equal_hom_trans (a b c d e e': obC) 169 | (f: morC a b) 170 | (g: morC c d) 171 | (h: morC e e'): 172 | f === g -> g === h -> f === h. 173 | Proof. 174 | destruct 1; 175 | destruct 1; 176 | constructor; 177 | etransitivity; eauto. 178 | Qed. 179 | 180 | End Equal_hom. 181 | 182 | Implicit Arguments Equal_hom [obC morC C a b c d]. 183 | 184 | Notation "f === g" := (Equal_hom f g) (at level 55). 185 | 186 | 187 | 188 | 189 | -------------------------------------------------------------------------------- /CAT/category_hom_transport.v: -------------------------------------------------------------------------------- 1 | 2 | Require Import Coq.Program.Equality. 3 | Require Import Coq.Logic.Eqdep. 4 | 5 | Require Export CatSem.CAT.category. 6 | 7 | Set Implicit Arguments. 8 | Unset Strict Implicit. 9 | Unset Transparent Obligations. 10 | Unset Automatic Introduction. 11 | 12 | 13 | (** a = a' /\ b = b' gives mor a b -> mor a' b' *) 14 | 15 | Section transport. 16 | 17 | Variable ob: Type. 18 | Variable hom: ob -> ob -> Type. 19 | Variable C: Cat_struct hom. 20 | 21 | 22 | Definition hom_transport (a b: ob)(f: hom a b) (a' b': ob) 23 | (H: a' = a) (H': b' = b) : hom a' b'. 24 | intros a b f a' b' H H'. 25 | rewrite H; 26 | rewrite H'. 27 | exact f. 28 | Defined. 29 | 30 | Lemma hom_transport_pi (a b: ob) (f: hom a b) (a' b': ob) 31 | (H1 H2 : a' = a) (Hb Hb2: b' = b): 32 | hom_transport f H1 Hb = hom_transport f H2 Hb2. 33 | Proof. 34 | intros a b f a' b' H1 H2 H3 H4. 35 | destruct H1. 36 | destruct H3. 37 | rewrite (UIP_refl _ _ H2). 38 | rewrite (UIP_refl _ _ H4). 39 | auto. 40 | Qed. 41 | 42 | Lemma hom_transport_id (a b: ob) (f: hom a b) 43 | (Ha: a = a) (Hb: b = b): 44 | hom_transport f Ha Hb = f. 45 | Proof. 46 | intros a b f Ha Hb. 47 | rewrite (UIP_refl _ _ Ha). 48 | rewrite (UIP_refl _ _ Hb). 49 | auto. 50 | Qed. 51 | 52 | 53 | Lemma Equal_hom_imp_setoideq (a b: ob)(f: hom a b) 54 | (a' b': ob) (g: hom a' b')(H: f === g)(Ha: a' = a) (Hb: b' = b) : 55 | hom_transport f Ha Hb == g . 56 | Proof. 57 | destruct 1. 58 | intros Ha Hb. 59 | rewrite hom_transport_id. 60 | auto. 61 | Qed. 62 | 63 | 64 | 65 | Lemma Equal_hom_imp_setoideq2_jmq_eq (a b: ob) (f g: hom a b): 66 | f === g -> f == g. 67 | Proof. 68 | intros a b f g H. 69 | dependent destruction H. 70 | (* dependent induction H.*) 71 | auto. 72 | Qed. 73 | 74 | 75 | Lemma Equal_hom_imp_setoideq2 (a b: ob) (f g: hom a b): 76 | f === g -> f == g. 77 | Proof. 78 | intros a b f g H. 79 | rewrite <- (@hom_transport_id a b f 80 | (refl_equal) (refl_equal)). 81 | apply Equal_hom_imp_setoideq. 82 | auto. 83 | Qed. 84 | 85 | (** heterogeneous equality of morphisms entails 86 | equality of source and target 87 | *) 88 | 89 | Section Equal_hom_lemmata. 90 | 91 | Variables a b c d: ob. 92 | Variable f: hom a b. 93 | Variable g: hom c d. 94 | 95 | Hypothesis H: f === g. 96 | 97 | Lemma Equal_hom_src: a = c. 98 | Proof. 99 | elim H. 100 | auto. 101 | Qed. 102 | 103 | Lemma Equal_hom_tgt: b = d. 104 | Proof. 105 | elim H. 106 | auto. 107 | Qed. 108 | 109 | End Equal_hom_lemmata. 110 | 111 | 112 | End transport. 113 | -------------------------------------------------------------------------------- /CAT/enriched_cat.v: -------------------------------------------------------------------------------- 1 | Require Export CatSem.CAT.mon_cats. 2 | 3 | Set Implicit Arguments. 4 | Unset Strict Implicit. 5 | Unset Transparent Obligations. 6 | Unset Automatic Introduction. 7 | 8 | Section enriched_cat_def. 9 | 10 | Notation "a 'X' b" := (build_prod_mor _ _ a b) (at level 45). 11 | 12 | Variable obM:Type. 13 | Variable morM: obM -> obM -> Type. 14 | Variable M: Cat_struct morM. 15 | Variable MonM: mon_cat (C:=M). 16 | 17 | Class Enriched_Cat (ob: Type) (hom: ob -> ob -> obM) := { 18 | idM: forall a : ob, morM (*@I obM morM M MonM*) I (hom a a); 19 | compM: forall a b c: ob, morM (tens (hom a b, hom b c)) (hom a c); 20 | tens_assoc: forall a b c d, 21 | Alpha ((hom a b,hom b c),hom c d) ;; 22 | #tens (compM _ _ _ X id (hom _ _ )) ;; compM _ _ _ == 23 | #tens (id (hom _ _ ) X compM _ _ _ ) ;; 24 | compM _ _ _ ; 25 | idMl: forall a b, 26 | #tens (idM a X id (hom a b) ) ;; compM _ _ _ == 27 | Lambda _ ; 28 | idMr: forall a b, 29 | #tens (id (hom a b) X idM b ) ;; compM _ _ _ == 30 | Rho _ 31 | }. 32 | 33 | 34 | 35 | End enriched_cat_def. 36 | 37 | 38 | (** a category enriched over SET is an ordinary category *) 39 | 40 | Section enriched_over_SET_MONOIDAL. 41 | Variable obC : Type. 42 | Variable morC : obC -> obC -> SET. 43 | Variable C : Enriched_Cat (SET_MONOIDAL) morC. 44 | 45 | Lemma default_setoid_eq (A:SET) : @Equivalence A (fun a b => a = b). 46 | Proof. 47 | intro A; 48 | constructor; 49 | auto. 50 | intros a b c H; 51 | rewrite H; 52 | auto. 53 | Qed. 54 | 55 | Definition default_setoid A := Build_Setoid (default_setoid_eq A). 56 | 57 | Program Instance Enriched_cat_Cat : Cat_struct morC := { 58 | mor_oid A B := default_setoid (morC A B); 59 | id a := idM (Enriched_Cat:=C) a tt ; 60 | comp a b c H H0 := compM (Enriched_Cat:=C) a b c (H,H0) 61 | }. 62 | (* 63 | Next Obligation. 64 | Proof. 65 | unfold Proper. 66 | intros x y H x' y' H'. 67 | rewrite H,H'. 68 | auto. 69 | Qed. 70 | *) 71 | Next Obligation. 72 | Proof. 73 | set (H:= idMr (Enriched_Cat := C)). 74 | simpl in *. 75 | set (H':= H a b (f, tt)). 76 | simpl in *. 77 | auto. 78 | Qed. 79 | Next Obligation. 80 | Proof. 81 | set (H:= idMl (Enriched_Cat := C)). 82 | simpl in *. 83 | set (H':= H a b (tt, f)). 84 | simpl in *. 85 | auto. 86 | Qed. 87 | Next Obligation. 88 | Proof. 89 | set (H:= tens_assoc (Enriched_Cat := C)). 90 | simpl in H. 91 | set (H' := H a b c d (f, (g,h))). 92 | auto. 93 | Qed. 94 | 95 | End enriched_over_SET_MONOIDAL. 96 | 97 | -------------------------------------------------------------------------------- /CAT/eq_fibre.v: -------------------------------------------------------------------------------- 1 | Require Export CatSem.CAT.rmonad_gen_type_po. 2 | 3 | Set Implicit Arguments. 4 | Unset Strict Implicit. 5 | Unset Automatic Introduction. 6 | 7 | (* contains a "smaller" version of the eq fibre transport 8 | morphism of modules 9 | *) 10 | 11 | Section gen_pb_fib_and_eq. 12 | 13 | Variables obC obD obC' obD': Type. 14 | Variable morC : obC -> obC -> Type. 15 | Variable morD : obD -> obD -> Type. 16 | Variable morC' : obC' -> obC' -> Type. 17 | Variable morD' : obD' -> obD' -> Type. 18 | (*Variable C : Cat morC. 19 | Variable D : Cat morD.*) 20 | Variable C' : Cat_struct morC'. 21 | Variable D' : Cat_struct morD'. 22 | 23 | (*Variable F : Functor C D.*) 24 | Variable F' : Functor C' D'. 25 | 26 | (*Variable P : RMonad F.*) 27 | Variable Q : RMonad F'. 28 | 29 | (** for a monad hom we need two functors *) 30 | 31 | (*Variable G1 : Functor C C'. 32 | Variable G2 : Functor D D'. 33 | *) 34 | (*Variable N : NT (CompF G1 F') (CompF F G2).*) 35 | 36 | (*Variable h : gen_RMonad_Hom P Q N.*) 37 | 38 | Variable T : Type. 39 | 40 | Variables u t : T. 41 | 42 | Variable M : RMOD Q (IPO T). 43 | 44 | Hypothesis H : u = t. 45 | 46 | Definition bla c: 47 | (FIB_RMOD Q u) M c -> ((FIB_RMOD Q t) M) c. 48 | simpl. 49 | intros c s. 50 | rewrite <- H. 51 | exact s. 52 | Defined. 53 | (* 54 | Obligation Tactic := idtac. 55 | *) 56 | Program Instance eq_rect_small_po_s c: 57 | PO_mor_struct 58 | (a:=(FIB_RMOD Q u) M c) (b:=(FIB_RMOD Q t) M c) 59 | (fun (s:M c u) => 60 | eq_rect u (fun t : T => (M c) t) s t H). 61 | Next Obligation. 62 | Proof. 63 | subst. 64 | simpl. 65 | auto. 66 | Qed. 67 | 68 | Definition eq_rect_small_po c:=Build_PO_mor (eq_rect_small_po_s c). 69 | 70 | Program Instance FIB_RMOD_small_eq_s : RModule_Hom_struct 71 | (M:=FIB_RMOD _ u M) (N:=FIB_RMOD _ t M) 72 | eq_rect_small_po. 73 | Next Obligation. 74 | Proof. 75 | subst. 76 | simpl. 77 | auto. 78 | Qed. 79 | 80 | 81 | Definition Fib_eq_RMod : 82 | FIB_RMOD _ u M ---> FIB_RMOD _ t M := 83 | Build_RModule_Hom FIB_RMOD_small_eq_s. 84 | 85 | End gen_pb_fib_and_eq. 86 | -------------------------------------------------------------------------------- /CAT/functor_properties.v: -------------------------------------------------------------------------------- 1 | Require Export CatSem.CAT.functor. 2 | 3 | Set Implicit Arguments. 4 | Unset Strict Implicit. 5 | Unset Transparent Obligations. 6 | Unset Automatic Introduction. 7 | 8 | Section full_faithful. 9 | 10 | Variables obC obD : Type. 11 | Variable morC : obC -> obC -> Type. 12 | Variable morD : obD -> obD -> Type. 13 | Variables (C : Cat_struct morC) (D : Cat_struct morD). 14 | 15 | Class Faithful (F : Functor C D) := { 16 | faithful : forall c d (f f' : morC c d), 17 | #F f == #F f' -> f == f' 18 | }. 19 | 20 | Class Full (F : Functor C D) := { 21 | preimg : forall c d (f : morD (F c) (F d)), morC c d; 22 | preimg_cond : forall c d (f : morD (F c) (F d)), 23 | #F (preimg f) == f 24 | }. 25 | 26 | End full_faithful. 27 | 28 | Section comp_of_functors. 29 | 30 | Variables obC obD obE : Type. 31 | Variable morC : obC -> obC -> Type. 32 | Variable morD : obD -> obD -> Type. 33 | Variable morE : obE -> obE -> Type. 34 | Variables (C : Cat_struct morC) (D : Cat_struct morD) (E : Cat_struct morE). 35 | Variables (F : Functor C D) (G : Functor D E). 36 | 37 | Instance comp_faithful (HF : Faithful F) (HG : Faithful G) : 38 | Faithful (CompF F G). 39 | Proof. 40 | intros; 41 | constructor; 42 | intros; 43 | apply (faithful (Faithful:=HF)); 44 | apply (faithful (Faithful:=HG)); 45 | auto. 46 | Qed. 47 | 48 | Program Instance comp_full (HF : Full F) (HG : Full G) : 49 | Full (CompF F G) := { 50 | preimg c d f := preimg (preimg (Full := HG) f) 51 | }. 52 | Next Obligation. 53 | Proof. 54 | intros; simpl; 55 | repeat rewrite preimg_cond; 56 | cat. 57 | Qed. 58 | 59 | 60 | End comp_of_functors. 61 | 62 | Existing Instance comp_full. 63 | 64 | -------------------------------------------------------------------------------- /CAT/het_rewrite.v: -------------------------------------------------------------------------------- 1 | Require Export Coq.Logic.Eqdep. 2 | 3 | Ltac rewr H := rewrite (inj_pair2 _ _ _ _ _ H). 4 | Ltac rerewr H := rewrite <- (inj_pair2 _ _ _ _ _ H). -------------------------------------------------------------------------------- /CAT/initial_terminal.v: -------------------------------------------------------------------------------- 1 | Require Export CatSem.CAT.category. 2 | 3 | Set Implicit Arguments. 4 | Unset Strict Implicit. 5 | Unset Transparent Obligations. 6 | Unset Automatic Introduction. 7 | 8 | Section Initial_Terminal. 9 | 10 | Variable obC: Type. 11 | Variable morC: obC -> obC -> Type. 12 | 13 | Class Initial (C:Cat_struct morC) := { 14 | Init : obC; 15 | InitMor: forall a:obC, morC Init a; 16 | InitMorUnique: forall b (f:morC Init b), f == InitMor b 17 | }. 18 | 19 | Class Terminal (C:Cat_struct morC) := { 20 | Term: obC; 21 | TermMor: forall a: obC, morC a Term; 22 | TermMorUnique: forall b f, f == TermMor b 23 | }. 24 | 25 | 26 | End Initial_Terminal. 27 | 28 | Implicit Arguments Init [obC morC C Initial]. 29 | Implicit Arguments Term [obC morC C Terminal]. 30 | 31 | (*Notation "'0'" := Init. 32 | Notation "1" := Term.*) 33 | 34 | 35 | 36 | (* 37 | Definition initial (C : Cat) (O : C) := 38 | forall a : C, sig (fun f:mor O a => forall g:mor O a, f == g). 39 | 40 | (** witness extraction *) 41 | 42 | Definition zerone (C:Category) (O:C): 43 | initial O -> forall a, mor O a. 44 | intros C O X a. 45 | unfold initial in X. 46 | destruct (X a) as [x p]. 47 | exact x. 48 | Defined. 49 | 50 | 51 | 52 | Definition terminal (C:Category) (l:C):= 53 | forall a : C, sig (fun f:mor a l => forall g:mor a l, f == g). 54 | 55 | Definition termone (C:Category)(l:C): 56 | terminal l -> forall a, mor a l. 57 | intros C l X a. 58 | destruct (X a) as [x p]. 59 | exact x. 60 | Defined. 61 | *) -------------------------------------------------------------------------------- /CAT/limits.v: -------------------------------------------------------------------------------- 1 | Require Export CatSem.CAT.NT. 2 | Require Export CatSem.CAT.CatFunct. 3 | Require Import CatSem.CAT.initial_terminal. 4 | 5 | Set Implicit Arguments. 6 | Unset Strict Implicit. 7 | 8 | Unset Automatic Introduction. 9 | 10 | (** content : 11 | - categories J, C 12 | - ConstFunc : constant functor (but won't be used in the following) 13 | - DiagFunctor : Functor C [J, C] 14 | - LIMIT A := Terminal (Cone A) 15 | - some functions for direct access to interesting parts of a limit 16 | *) 17 | 18 | 19 | Section defs. 20 | 21 | Variable obC : Type. 22 | Variable morC : obC -> obC -> Type. 23 | Variable C : Cat_struct morC. 24 | 25 | Variable obJ : Type. 26 | Variable morJ : obJ -> obJ -> Type. 27 | Variable J : Cat_struct morJ. 28 | 29 | (** constant functor *) 30 | 31 | Program Definition ConstFunc (c : obC) := Build_Functor 32 | (Fobj := fun j: obJ => c)(Fmor := fun a b f => id c) _ . 33 | Next Obligation. 34 | Proof. 35 | constructor; 36 | [unfold Proper; red | 37 | idtac | idtac ]; 38 | cat. 39 | Qed. 40 | 41 | (** NT between constant functors, induced by f : a --> b *) 42 | 43 | Section on_morphisms. 44 | Variables a b : obC. 45 | Variable f : morC a b. 46 | 47 | Program Instance ConstFuncNatTrans : 48 | NT_struct (F:=ConstFunc a) (G:=ConstFunc b) (fun _ => f). 49 | Next Obligation. 50 | Proof. 51 | cat. 52 | Qed. 53 | 54 | Definition ConstFuncNT := Build_NT ConstFuncNatTrans. 55 | 56 | End on_morphisms. 57 | 58 | (** Diagonal functor [C ---> [J, C]] *) 59 | 60 | Program Instance DiagFunc : Functor_struct (C:=C) (D:= FunctCat J C) 61 | (Fobj := ConstFunc) 62 | (ConstFuncNT). 63 | Next Obligation. 64 | Proof. 65 | unfold Proper; red. 66 | unfold ConstFuncNT. 67 | simpl. 68 | auto. 69 | Qed. 70 | Next Obligation. 71 | Proof. 72 | cat. 73 | Qed. 74 | Next Obligation. 75 | Proof. 76 | unfold vcompNT1. 77 | cat. 78 | Qed. 79 | 80 | Definition DiagFunctor := Build_Functor DiagFunc. 81 | 82 | (** category of cones *) 83 | 84 | Section Cone. 85 | 86 | Variable F : Functor J C. 87 | 88 | Class ConeClass (a : obC) := { 89 | cone_mor : forall j, morC a (F j) ; 90 | cone_prop : forall j j' (f : morJ j j'), 91 | cone_mor j ;; #F f == cone_mor j' 92 | }. 93 | 94 | Record Cone := { 95 | ConeTop :> obC ; 96 | cone_struct :> ConeClass ConeTop 97 | }. 98 | 99 | Existing Instance cone_struct. 100 | 101 | Section Cone_Homs. 102 | 103 | Variables M N : Cone. 104 | 105 | Class Cone_Hom_struct (f : morC M N) := { 106 | cone_comm : forall j:obJ, f ;; cone_mor j == cone_mor j 107 | }. 108 | 109 | Record Cone_Hom := { 110 | cone_hom_carrier :> morC M N; 111 | cone_hom_struct :> Cone_Hom_struct cone_hom_carrier 112 | }. 113 | 114 | Lemma Cone_Hom_equiv : @Equivalence Cone_Hom 115 | (fun A B => cone_hom_carrier A == cone_hom_carrier B). 116 | Proof. 117 | split. 118 | unfold Reflexive; 119 | cat. 120 | unfold Symmetric; 121 | intros x y; 122 | apply hom_sym. 123 | unfold Transitive; 124 | intros x y z; 125 | apply hom_trans. 126 | Qed. 127 | 128 | Definition Cone_Hom_oid := Build_Setoid Cone_Hom_equiv. 129 | 130 | End Cone_Homs. 131 | 132 | Existing Instance cone_hom_struct. 133 | 134 | Section Cone_id_comp. 135 | 136 | Variable A : Cone. 137 | 138 | Program Definition Cone_id : Cone_Hom A A := 139 | Build_Cone_Hom (cone_hom_carrier := id _ ) _ . 140 | Next Obligation. 141 | Proof. 142 | constructor. 143 | cat. 144 | Qed. 145 | 146 | Variables B D : Cone. 147 | Variable f : Cone_Hom A B. 148 | Variable g : Cone_Hom B D. 149 | 150 | Program Definition Cone_comp : Cone_Hom A D := 151 | Build_Cone_Hom (cone_hom_carrier := cone_hom_carrier f ;; 152 | cone_hom_carrier g) _ . 153 | Next Obligation. 154 | Proof. 155 | constructor. 156 | intro j. 157 | rewrite assoc. 158 | rewrite (cone_comm j). 159 | rewrite (cone_comm j). 160 | cat. 161 | Qed. 162 | 163 | End Cone_id_comp. 164 | 165 | Obligation Tactic := cat ; try apply assoc. 166 | 167 | Program Instance CONE_struct : Cat_struct Cone_Hom := { 168 | mor_oid := Cone_Hom_oid; 169 | id := Cone_id; 170 | comp := Cone_comp 171 | }. 172 | Next Obligation. 173 | Proof. 174 | unfold Proper; 175 | do 2 red. 176 | simpl. 177 | intros x y H x' y' H'. 178 | rewrite H, H'. 179 | cat. 180 | Qed. 181 | 182 | Definition CONE := Build_Cat CONE_struct. 183 | 184 | End Cone. 185 | 186 | (** a limit is a terminal object in (CONE A) *) 187 | 188 | 189 | 190 | Definition LIMIT A := Terminal (CONE A). 191 | 192 | (** easy access to interesting parts of a limit *) 193 | 194 | Section Limit_defs. 195 | 196 | Variable A : Functor J C. 197 | 198 | Hypothesis H : LIMIT A. 199 | 200 | Definition Limit : Cone A := Term (Terminal := H). 201 | 202 | Definition LimitVertex : obC := ConeTop Limit. 203 | 204 | Definition LimitMor (j : obJ) := cone_mor (ConeClass := Limit) j. 205 | 206 | End Limit_defs. 207 | 208 | End defs. 209 | 210 | Existing Instance cone_struct. 211 | Existing Instance cone_hom_struct. 212 | Existing Instance CONE_struct. 213 | 214 | 215 | 216 | 217 | 218 | 219 | 220 | 221 | 222 | 223 | -------------------------------------------------------------------------------- /CAT/module_postcomp_functor.v: -------------------------------------------------------------------------------- 1 | Require Export CatSem.CAT.monad_h_module. 2 | 3 | Set Implicit Arguments. 4 | Unset Strict Implicit. 5 | Unset Transparent Obligations. 6 | Unset Automatic Introduction. 7 | 8 | Section postcomp_functor. 9 | (* 10 | Variable obC : Type. 11 | Variable morC : obC -> obC -> Type. 12 | Variable C : Cat morC. 13 | *) 14 | Variable C : Cat. 15 | 16 | Variable P : Monad C. 17 | (* 18 | Variable obD : Type. 19 | Variable morD : obD -> obD -> Type. 20 | Variable D : Cat morD. 21 | *) 22 | Variables D E : Cat. 23 | (* 24 | Variable obE : Type. 25 | Variable morE : obE -> obE -> Type. 26 | Variable E : Cat morE. 27 | *) 28 | Variable F : Functor D E. 29 | 30 | Section on_modules. 31 | 32 | Variable M : Module P D. 33 | 34 | 35 | Obligation Tactic := idtac. 36 | 37 | Program Instance Mod_postcomp_F_struct : Module_struct P 38 | (fun x => F (M x)) := { 39 | mkleisli a b f := #F (mkleisli f) 40 | }. 41 | Next Obligation. 42 | Proof. 43 | unfold Proper; red; 44 | intros; 45 | rew_all; 46 | cat. 47 | Qed. 48 | Next Obligation. 49 | Proof. 50 | intros. 51 | rewrite <- FComp. 52 | rew mklmkl. 53 | cat. 54 | Qed. 55 | Next Obligation. 56 | Proof. 57 | intros. 58 | simpl. 59 | rew mklweta. 60 | cat. 61 | Qed. 62 | 63 | Definition Mod_postcomp_F := Build_Module Mod_postcomp_F_struct. 64 | 65 | End on_modules. 66 | 67 | Section on_module_morphisms. 68 | 69 | Variables M N : Module P D. 70 | 71 | Variable S : M ---> N. 72 | 73 | Obligation Tactic := idtac. 74 | 75 | Program Instance Mod_postcomp_F_mor_struct : Module_Hom_struct 76 | (S := Mod_postcomp_F M) 77 | (T := Mod_postcomp_F N) (fun x => #F (S x)). 78 | Next Obligation. 79 | Proof. 80 | simpl; intros. 81 | repeat rewrite <- FComp. 82 | rew (mod_hom_mkl). 83 | cat. 84 | app_all. 85 | Qed. 86 | 87 | Definition Mod_postcomp_F_mor : Mod_postcomp_F M ---> 88 | Mod_postcomp_F N := 89 | Build_Module_Hom Mod_postcomp_F_mor_struct. 90 | 91 | End on_module_morphisms. 92 | 93 | Obligation Tactic := simpl; intros; try unf_Proper; 94 | cat; rew_all; cat. 95 | 96 | Program Instance MOD_pc_F : 97 | Functor_struct (Fobj := Mod_postcomp_F) 98 | Mod_postcomp_F_mor. 99 | 100 | End postcomp_functor. 101 | 102 | 103 | 104 | 105 | 106 | 107 | 108 | 109 | 110 | 111 | 112 | 113 | 114 | 115 | 116 | 117 | 118 | 119 | 120 | 121 | 122 | 123 | -------------------------------------------------------------------------------- /CAT/monad_h_morphism.v: -------------------------------------------------------------------------------- 1 | Require Export CatSem.CAT.monad_haskell. 2 | Require Export CatSem.CAT.NT. 3 | 4 | Set Implicit Arguments. 5 | Unset Strict Implicit. 6 | Unset Transparent Obligations. 7 | Unset Automatic Introduction. 8 | 9 | (** definition of a morphism of monads in haskell style *) 10 | 11 | Section Monad_Hom. 12 | (* 13 | Variable obC: Type. 14 | Variable morC: obC -> obC -> Type. 15 | Variable C: Cat_struct morC. 16 | *) 17 | Variable C : Cat. 18 | 19 | Section Monad_Hom_def. 20 | 21 | Variables P Q: Monad C. 22 | 23 | Class Monad_Hom_struct (Tau: forall c, P c ---> Q c) := { 24 | monad_hom_kl: forall X Y (f: X ---> P Y), 25 | kleisli f ;; Tau Y == Tau X ;; kleisli (f ;; Tau Y ) ; 26 | monad_hom_weta: forall c:C, 27 | weta c ;; Tau c == weta c 28 | }. 29 | 30 | Section lemmata. 31 | Variable Tau : forall c, P c ---> Q c. 32 | Variable H : Monad_Hom_struct Tau. 33 | 34 | Lemma mh_weta c : weta c ;; Tau c == weta c. 35 | Proof. 36 | apply H. 37 | Qed. 38 | 39 | Lemma mh_kleisli X Y f : 40 | kleisli f ;; Tau Y == Tau X ;; kleisli (f ;; Tau Y ). 41 | Proof. 42 | apply H. 43 | Qed. 44 | 45 | End lemmata. 46 | 47 | Record Monad_Hom := { 48 | monad_hom:> forall c, P c ---> Q c; 49 | monad_hom_struct :> Monad_Hom_struct monad_hom 50 | }. 51 | 52 | Section Monad_Hom_NT. 53 | 54 | Variable M : Monad_Hom. 55 | 56 | Obligation Tactic := idtac. 57 | 58 | Program Instance Monad_Hom_NatTrans : NT_struct (fun c => M c). 59 | Next Obligation. 60 | Proof. 61 | intros c d f. 62 | simpl. 63 | unfold lift. 64 | rewrite (monad_hom_kl (Monad_Hom_struct := M) (f;; weta d)). 65 | apply praecomp. 66 | apply kleisli_oid. 67 | rewrite assoc. 68 | rewrite (monad_hom_weta (Monad_Hom_struct := M)). 69 | cat. 70 | Qed. 71 | 72 | Canonical Structure Monad_Hom_NT := Build_NT Monad_Hom_NatTrans. 73 | 74 | End Monad_Hom_NT. 75 | 76 | End Monad_Hom_def. 77 | (* 78 | Definition eq_Monad_Hom (P Q: Monad C) : relation (Monad_Hom P Q) := 79 | fun M N => forall x, M x == N x. 80 | *) 81 | Lemma Monad_Hom_equiv (P Q: Monad C) : 82 | Equivalence (A:=Monad_Hom P Q) (fun M N => forall x, M x == N x). 83 | Proof. 84 | intros P Q; 85 | constructor. 86 | unfold Reflexive. 87 | cat. 88 | unfold Symmetric; 89 | intros; apply hom_sym; 90 | auto. 91 | unfold Transitive. 92 | intros. 93 | etransitivity; auto. 94 | Qed. 95 | 96 | Definition Monad_Hom_oid (P Q: Monad C) := 97 | Build_Setoid (Monad_Hom_equiv P Q). 98 | 99 | 100 | Hint Rewrite monad_hom_kl monad_hom_weta : monad. 101 | Existing Instance monad_hom_struct. 102 | 103 | Section Monad_Hom_comp. 104 | 105 | Variables P Q R : Monad C. 106 | Variable TT: Monad_Hom P Q. 107 | Variable SS: Monad_Hom Q R. 108 | 109 | Obligation Tactic := idtac. 110 | 111 | Program Instance Monad_Hom_comp_struct : 112 | Monad_Hom_struct (fun c => TT c ;; SS c). 113 | Next Obligation. 114 | Proof. 115 | intros X Y f; simpl. 116 | rewrite <- assoc. 117 | rewrite (monad_hom_kl); try apply TT. 118 | repeat rewrite assoc. 119 | apply praecomp. 120 | rewrite <- assoc. 121 | rewrite <- monad_hom_kl; 122 | try apply SS. 123 | monad; try apply TT; try apply SS. 124 | Qed. 125 | Next Obligation. 126 | Proof. 127 | intro c. 128 | rewrite <- assoc. 129 | rewrite monad_hom_weta; 130 | monad; 131 | try apply TT; 132 | try apply SS. 133 | Qed. 134 | 135 | Canonical Structure Monad_Hom_comp := 136 | Build_Monad_Hom Monad_Hom_comp_struct. 137 | 138 | End Monad_Hom_comp. 139 | 140 | Section Monad_Hom_id. 141 | 142 | Variable P: Monad C. 143 | 144 | Obligation Tactic := monad; (try constructor); monad. 145 | 146 | Program Instance Monad_Hom_id_struct : 147 | Monad_Hom_struct (fun c => id (P c)). 148 | 149 | 150 | Canonical Structure Monad_Hom_id := 151 | Build_Monad_Hom Monad_Hom_id_struct. 152 | 153 | End Monad_Hom_id. 154 | 155 | Existing Instance monad_hom_struct. 156 | 157 | Obligation Tactic := simpl; intros; try apply assoc; cat. 158 | 159 | Program Instance MONAD_struct : Cat_struct (fun P Q => Monad_Hom P Q) := { 160 | mor_oid P Q := Monad_Hom_oid P Q; 161 | id P := Monad_Hom_id P; 162 | comp a b c f g := Monad_Hom_comp f g 163 | }. 164 | Next Obligation. 165 | Proof. 166 | unfold Proper; do 2 red. 167 | simpl. 168 | intros x y H x' y' H' z; 169 | rewrite H; 170 | rewrite H'; 171 | cat. 172 | Qed. 173 | 174 | Definition MONAD := Build_Cat MONAD_struct. 175 | 176 | Section Facts. 177 | 178 | Variables P Q:MONAD. 179 | 180 | Variable h : P ---> Q. 181 | 182 | Existing Instance monad_hom_struct. 183 | 184 | Hint Rewrite monad_hom_kl monad_hom_weta : monad. 185 | 186 | Lemma monad_hom_lift c d (f: c ---> d) : 187 | lift f ;; h _ == h _ ;; lift f. 188 | Proof. 189 | unfold lift. 190 | monad; apply h. 191 | Qed. 192 | 193 | Lemma monad_hom_join (c : C): 194 | join _ ;; h _ == h _ ;; lift (h _ ) ;; join c . 195 | Proof. 196 | unfold join; 197 | monad; 198 | try apply h. 199 | Qed. 200 | 201 | End Facts. 202 | 203 | End Monad_Hom. 204 | 205 | Existing Instance monad_hom_struct. 206 | Existing Instance MONAD_struct. 207 | Existing Instance Monad_Hom_oid. 208 | Existing Instance Monad_Hom_NatTrans. 209 | 210 | Hint Rewrite monad_hom_lift monad_hom_join monad_hom_kl 211 | monad_hom_weta: monad. 212 | 213 | Coercion Monad_Hom_NatTrans : Monad_Hom >-> NT_struct. 214 | 215 | 216 | 217 | -------------------------------------------------------------------------------- /CAT/monad_h_morphism_gen.v: -------------------------------------------------------------------------------- 1 | Require Export CatSem.CAT.monad_h_module. 2 | Require Export CatSem.CAT.NT. 3 | Require Export CatSem.CAT.product. 4 | 5 | Set Implicit Arguments. 6 | Unset Strict Implicit. 7 | Unset Transparent Obligations. 8 | Unset Automatic Introduction. 9 | 10 | Section gen_monad_morphism. 11 | (* 12 | Variables obC obD : Type. 13 | Variable morC : obC -> obC -> Type. 14 | Variable morD : obD -> obD -> Type. 15 | Variable C : Cat_struct morC. 16 | Variable D : Cat_struct morD. 17 | *) 18 | Variables C D : Cat. 19 | Variable P : Monad C. 20 | Variable Q : Monad D. 21 | 22 | Variable F : Functor C D. 23 | 24 | (** is P;F a module C -> D over P ?*) 25 | 26 | Program Instance F_mod_s : Module_struct P (D:=D) 27 | (fun x => F (P x)) := { 28 | mkleisli a b f := #F (mkleisli f) 29 | }. 30 | Next Obligation. 31 | Proof. 32 | intros; simpl. 33 | unfold Proper; 34 | red. 35 | intros f g H. 36 | apply (functor_map_morphism (kleisli_oid H)). 37 | Qed. 38 | Next Obligation. 39 | Proof. 40 | intros; simpl. 41 | rewrite <- FComp. 42 | apply (functor_map_morphism (dist _ _ )). 43 | Qed. 44 | Next Obligation. 45 | Proof. 46 | intros; simpl. 47 | monad; 48 | cat. 49 | Qed. 50 | 51 | Definition F_mod : MOD P D := Build_Module F_mod_s. 52 | 53 | 54 | (** generalizing the previos one *) 55 | 56 | 57 | Class colax_Monad_Hom_struct (Tau : forall c, F (P c) ---> Q (F c)) := { 58 | gen_monad_hom_kl : forall c d (f : c ---> P d), 59 | #F (kleisli f) ;; Tau _ == 60 | Tau _ ;; (kleisli (#F f ;; Tau _ )) ; 61 | gen_monad_hom_weta : forall c : C, 62 | #F (weta c) ;; Tau _ == weta _ 63 | }. 64 | 65 | Record colax_Monad_Hom := { 66 | gen_monad_hom:> forall c, F (P c) ---> Q (F c); 67 | gen_monad_hom_struct :> colax_Monad_Hom_struct gen_monad_hom 68 | }. 69 | 70 | 71 | Existing Instance gen_monad_hom_struct. 72 | 73 | Section Monad_Hom_NT_PB. 74 | 75 | Variable M : colax_Monad_Hom. 76 | 77 | Program Instance colax_Monad_Hom_NatTrans : 78 | NT_struct (F:=CompF (MFunc P) F) 79 | (G:=CompF F (MFunc Q)) (fun c : C => M c). 80 | Next Obligation. 81 | Proof. 82 | simpl; intros. 83 | unfold lift. 84 | simpl. 85 | assert (H:=gen_monad_hom_kl (colax_Monad_Hom_struct := M) (f;; weta d)). 86 | simpl in *. 87 | rewrite H. 88 | apply praecomp. 89 | apply kleisli_oid. 90 | rewrite FComp. 91 | rewrite assoc. 92 | rewrite (gen_monad_hom_weta (colax_Monad_Hom_struct := M)). 93 | cat. 94 | Qed. 95 | 96 | (*End Monad_Hom_NT.*) 97 | 98 | (*Section pullback.*) 99 | (* 100 | Variable obE : Type. 101 | Variable morE : obE -> obE -> Type. 102 | Variable E : Cat morE. 103 | *) 104 | Variable E : Cat. 105 | Section pullback_ob. 106 | 107 | Variable N : MOD Q E. 108 | 109 | Program Instance PModule_struct : Module_struct P (D:=E) 110 | (fun c => N (F c)) := { 111 | mkleisli a b f := mkleisli (Module_struct := N)(#F f ;; M _ ) 112 | }. 113 | Next Obligation. 114 | Proof. 115 | intros; simpl. 116 | unfold Proper; 117 | red. 118 | intros x y H. 119 | apply mkleisli_oid. 120 | apply postcomp. 121 | apply (functor_map_morphism H). 122 | Qed. 123 | Next Obligation. 124 | Proof. 125 | intros; simpl. 126 | rewrite mkl_mkl. 127 | apply mkleisli_oid. 128 | rewrite FComp. 129 | repeat rewrite assoc. 130 | apply praecomp. 131 | assert (H:=gen_monad_hom_kl (colax_Monad_Hom_struct := M)). 132 | simpl in *. 133 | mod. 134 | Qed. 135 | Next Obligation. 136 | Proof. 137 | intros; simpl. 138 | rewrite gen_monad_hom_weta. 139 | mod. 140 | Qed. 141 | 142 | Definition PModule : MOD P E := Build_Module PModule_struct. 143 | 144 | End pullback_ob. 145 | 146 | (** pullback along a gen monad hom is functorial *) 147 | 148 | Section pullback_mor. 149 | 150 | Variables N N' : MOD Q E. 151 | Variable r : N ---> N'. 152 | 153 | Program Instance PMod_Hom_s : Module_Hom_struct 154 | (S := PModule N) (T:= PModule N') (fun c => r (F c) ). 155 | Next Obligation. 156 | simpl; intros. 157 | mod; 158 | try apply r. 159 | Qed. 160 | 161 | Definition PMod_Hom := Build_Module_Hom PMod_Hom_s. 162 | 163 | End pullback_mor. 164 | 165 | Variables N N' : MOD Q E. 166 | 167 | Variable eprod : Cat_Prod E. 168 | 169 | Obligation Tactic := cat. 170 | 171 | Definition prod_mod : (forall x : C, 172 | (product (PModule N) (PModule N')) x ---> (PModule (product N N')) x) := (fun x => id (Cat_struct := E) (product (N (F x)) (N' (F x)))). 173 | 174 | 175 | Program Instance PROD_PM_s : Module_Hom_struct 176 | (S := product (PModule N) (PModule N')) 177 | (T := PModule (product N N')) 178 | prod_mod. 179 | 180 | Definition PROD_PM := Build_Module_Hom PROD_PM_s. 181 | 182 | End Monad_Hom_NT_PB. 183 | 184 | 185 | 186 | (*End gen_monad_morphism.*) 187 | 188 | 189 | Variable M : colax_Monad_Hom. 190 | 191 | Program Instance PMod_ind_Hom_s : Module_Hom_struct 192 | (S:= F_mod) (T:=PModule M Q) M. 193 | Next Obligation. 194 | Proof. 195 | simpl; intros; 196 | rewrite gen_monad_hom_kl. 197 | cat. 198 | Qed. 199 | 200 | Definition PMod_ind_Hom := Build_Module_Hom PMod_ind_Hom_s. 201 | 202 | End gen_monad_morphism. 203 | 204 | Existing Instance gen_monad_hom_struct. 205 | 206 | 207 | 208 | 209 | 210 | 211 | 212 | -------------------------------------------------------------------------------- /CAT/monic_epi.v: -------------------------------------------------------------------------------- 1 | Require Export CatSem.CAT.category. 2 | 3 | Set Implicit Arguments. 4 | Unset Strict Implicit. 5 | 6 | Unset Automatic Introduction. 7 | 8 | Section defs. 9 | 10 | Variable obC : Type. 11 | Variable morC : obC -> obC -> Type. 12 | Variable C : Cat_struct morC. 13 | 14 | (** definition of 15 | - monic 16 | - epi 17 | - invertible 18 | - ismorphic objects 19 | *) 20 | 21 | 22 | Section morphisms. 23 | 24 | Variables a b : obC. 25 | 26 | Class Monic (f: morC a b) := { 27 | monic_cond : forall d (g1 g2 : morC d a), 28 | g1 ;; f == g2 ;; f -> g1 == g2 29 | }. 30 | 31 | Class Epi (f: morC a b) := { 32 | epi_cond : forall d (g1 g2 : morC b d), 33 | f ;; g1 == f ;; g2 -> g1 == g2 34 | }. 35 | 36 | Class Invertible (f: morC a b) := { 37 | inverse : morC b a ; 38 | inv_prae : inverse ;; f == id _ ; 39 | inv_post : f ;; inverse == id _ 40 | }. 41 | 42 | Class isomorphic := { 43 | inv_morphism : morC a b; 44 | invertible :> Invertible inv_morphism 45 | }. 46 | 47 | End morphisms. 48 | 49 | Implicit Arguments inverse [a b Invertible]. 50 | 51 | Notation "-- f" := (inverse f) (at level 30). 52 | 53 | (** lemmata, such as "composition of monics..." etc bla bla *) 54 | 55 | Section lemmata. 56 | 57 | Variables a b c : obC. 58 | Variable f : morC a b. 59 | Variable g : morC b c. 60 | 61 | Global Instance comp_of_monics 62 | (Hf : Monic f) (Hg : Monic g) : Monic (f ;; g). 63 | Proof. 64 | intros Hf Hg. 65 | constructor. 66 | intros d g1 g2. 67 | repeat rewrite <- assoc. 68 | intro H. 69 | assert (H': g1 ;; f == g2 ;; f). 70 | apply Hg; auto. 71 | apply Hf; auto. 72 | Qed. 73 | 74 | Global Instance comp_of_epis (Hf : Epi f) (Hg : Epi g) : Epi (f ;; g). 75 | Proof. 76 | intros Hf Hg. 77 | constructor. 78 | intros d g1 g2. 79 | repeat rewrite assoc. 80 | intro H. 81 | assert (H': g ;; g1 == g ;; g2). 82 | apply Hf; auto. 83 | apply Hg; auto. 84 | Qed. 85 | 86 | Global Instance comp_monic_fst_monic (H : Monic (f ;; g)) : Monic f. 87 | Proof. 88 | intro H. 89 | set (H':= monic_cond (Monic:=H)). 90 | constructor. 91 | intros d g1 g2 H2. 92 | apply (H' _ g1 g2). 93 | repeat rewrite <- assoc. 94 | rewrite H2. 95 | cat. 96 | Qed. 97 | 98 | Global Instance comp_epi_snd_epi (H : Epi (f ;; g)) : Epi g. 99 | Proof. 100 | intro H. 101 | set (H':= epi_cond (Epi:=H)). 102 | constructor. 103 | intros d g1 g2 H2. 104 | apply (H' _ g1 g2). 105 | repeat rewrite assoc. 106 | rewrite H2. 107 | cat. 108 | Qed. 109 | 110 | (** uniqueness of the inverse *) 111 | 112 | Lemma inverse_unique (H : Invertible f) h 113 | (H1 : f ;; h == id _ ) (H2 : h ;; f == id _ ) : 114 | h == -- f. 115 | Proof. 116 | intros H h H1 H2. 117 | transitivity (h ;; (f ;; --f)). 118 | rewrite inv_post; cat. 119 | rewrite <- assoc. 120 | rewrite H2. 121 | cat. 122 | Qed. 123 | 124 | Program Instance inverse_invertible (H : Invertible f) : 125 | Invertible (--f):= { 126 | inverse := f; 127 | inv_prae := inv_post (Invertible := H); 128 | inv_post := inv_prae (Invertible := H) 129 | }. 130 | 131 | End lemmata. 132 | 133 | Existing Instance inverse_invertible. 134 | 135 | Section more_lemmata. 136 | 137 | Variables a b c : obC. 138 | Variable f : morC a b. 139 | Variable g : morC b c. 140 | 141 | 142 | Lemma inverse_inverse (H : Invertible f) : -- (-- f) == f. 143 | Proof. 144 | intro H. 145 | apply hom_sym. 146 | apply inverse_unique. 147 | apply inv_prae. 148 | apply inv_post. 149 | Qed. 150 | 151 | 152 | Lemma put_inv_to_right (H : Invertible f) h : 153 | f ;; g == h <-> g == -- f ;; h. 154 | Proof. 155 | intros H h. 156 | split; intro H'. 157 | transitivity (--f ;; f ;; g). 158 | rewrite inv_prae; cat. 159 | rewrite <- H'; apply assoc. 160 | 161 | transitivity (f ;; --f ;; h). 162 | repeat rewrite assoc. 163 | apply praecomp; auto. 164 | rewrite inv_post; cat. 165 | Qed. 166 | 167 | End more_lemmata. 168 | 169 | 170 | Section still_more_lemmata. 171 | 172 | Variables a b c : obC. 173 | Variable f : morC a b. 174 | Variable g : morC b c. 175 | 176 | (** inverse of a composition *) 177 | 178 | Program Instance inv_of_comp (Hf : Invertible f) (Hg : Invertible g) : 179 | Invertible (f ;; g) := { 180 | inverse := --g ;; --f 181 | }. 182 | Next Obligation. 183 | Proof. 184 | apply hom_sym. 185 | rewrite assoc. 186 | rewrite <- put_inv_to_right. 187 | rewrite <- assoc. 188 | rewrite inv_prae. 189 | cat. 190 | Qed. 191 | 192 | Next Obligation. 193 | apply hom_sym. 194 | 195 | 196 | rewrite <- assoc. 197 | set (z := f ;; g ;; -- g). 198 | 199 | assert (eqz :z == f). 200 | unfold z. 201 | rewrite assoc . 202 | rewrite inv_post. 203 | cat. 204 | rewrite eqz. 205 | rewrite inv_post. 206 | 207 | cat. 208 | Qed. 209 | 210 | 211 | End still_more_lemmata. 212 | 213 | End defs. 214 | 215 | Implicit Arguments Invertible [obC morC C a b]. 216 | Implicit Arguments inverse [obC morC C a b Invertible]. 217 | 218 | 219 | 220 | -------------------------------------------------------------------------------- /CAT/my_lib.v: -------------------------------------------------------------------------------- 1 | Require Import Coq.Logic.ProofIrrelevance. 2 | (* 3 | Require Import Coq.Logic.FunctionalExtensionality. 4 | *) 5 | Set Implicit Arguments. 6 | Unset Strict Implicit. 7 | 8 | Unset Automatic Introduction. 9 | 10 | 11 | Definition PI := proof_irrelevance. 12 | 13 | 14 | Lemma break_Sig_Prop : forall A : Type, forall P : A -> Prop, 15 | forall a a' : A, forall (pa : P a)(pa': P a'), 16 | a = a' -> 17 | exist P a pa = exist P a' pa'. 18 | Proof. intros. assert (P a = P a'). 19 | rewrite H; auto. 20 | generalize pa. 21 | rewrite H0 in pa. 22 | rewrite H. 23 | intros. 24 | rewrite (@PI (P a') pa0 pa'). 25 | reflexivity. 26 | Qed. 27 | 28 | Lemma break_Sig_Prop2: forall A : Type, forall P : A -> Prop, 29 | forall a a': sig P, 30 | proj1_sig a = proj1_sig a' -> a = a'. 31 | Proof. intros A P a a' H. 32 | destruct a. destruct a'. 33 | apply break_Sig_Prop. 34 | auto. 35 | Qed. 36 | 37 | Lemma reduce_to_sigma: forall (A : Type)(B : A -> Type)(C : Type), 38 | forall f: forall a:A, B a -> C, 39 | forall a a' b b', 40 | existT _ a b = existT _ a' b' -> 41 | f a b = f a' b'. 42 | Proof. intros. dependent rewrite H; auto. Qed. 43 | 44 | Lemma simpl_eq_sig: forall (A:Type) (B: A -> Type), 45 | forall (a a':A)(b:B a)(b':B a'), 46 | existT B a b = existT B a' b' -> 47 | a = a'. 48 | Proof. intros A B a a' b b' H. 49 | dependent rewrite H. auto. 50 | Qed. 51 | 52 | 53 | (* 54 | Lemma simpl_eq_sig2: forall A:Type, forall (B:A -> Type), 55 | forall (a:A) b b', existT B a b = existT B a b' -> 56 | projT2 (existT B a b) = projT2 (existT B a b'). 57 | Proof. intros. simpl. 58 | inversion H. 59 | injection H. auto. 60 | dependent rewrite <- H. simpl. 61 | inversion H. 62 | injection H. 63 | intros. 64 | assumption. 65 | *) -------------------------------------------------------------------------------- /CAT/nat_iso.v: -------------------------------------------------------------------------------- 1 | Require Export CatSem.CAT.NT. 2 | Require Export CatSem.CAT.monic_epi. 3 | 4 | Set Implicit Arguments. 5 | Unset Strict Implicit. 6 | Unset Automatic Introduction. 7 | 8 | Section Nat_Iso_def. 9 | 10 | Variables obC obD: Type. 11 | Variable morC : obC -> obC -> Type. 12 | Variable morD : obD -> obD -> Type. 13 | Variable C: Cat_struct morC. 14 | Variable D: Cat_struct morD. 15 | Variables F G: Functor C D. 16 | 17 | Section struct_def. 18 | 19 | Variable alpha: NT F G. 20 | (* 21 | Class NISO_struct := { 22 | inv: forall c, morD (G c) (F c); 23 | inv_right: forall c, alpha c ;; inv c == id _; 24 | inv_left: forall c, inv c ;; alpha c == id _ 25 | }. 26 | *) 27 | Class NISO_struct := { 28 | NT_inv :> forall c, Invertible (alpha c) 29 | }. 30 | 31 | 32 | End struct_def. 33 | 34 | Record NISO := { 35 | tra :> NT F G; 36 | niso_struct :> NISO_struct tra 37 | }. 38 | 39 | End Nat_Iso_def. 40 | 41 | Existing Instance niso_struct. 42 | 43 | 44 | Section Nat_Iso_gives_NT. 45 | 46 | Notation "-- f" := (inverse f) (at level 30). 47 | 48 | Variables obC obD: Type. 49 | Variable morC : obC -> obC -> Type. 50 | Variable morD : obD -> obD -> Type. 51 | Variable C: Cat_struct morC. 52 | Variable D: Cat_struct morD. 53 | Variables F G: Functor C D. 54 | 55 | Variable alpha: NISO F G. 56 | 57 | Program Definition NTinv: NT G F := Build_NT 58 | (trafo := fun c => -- (alpha c)) _ . 59 | Next Obligation. 60 | Proof. 61 | constructor. 62 | intros c d f. 63 | destruct alpha as [al NIS]. simpl. 64 | destruct NIS as [INV]. 65 | simpl in *|-*. 66 | set (H := trafo_ax (NT_struct := al)). 67 | assert (H': --al c ;; (al c ;; #G f) == 68 | --al c ;; (#F f ;; al d)). 69 | apply praecomp. auto. 70 | assert (H'': (--al c ;; (al c ;; #G f)) ;; --al d == 71 | (--al c ;; (#F f ;; al d)) ;; --al d). 72 | apply postcomp. auto. 73 | repeat rewrite assoc in H''. 74 | rewrite inv_post in H''. 75 | repeat rewrite <- assoc in H''. 76 | rewrite inv_prae in H''. 77 | rewrite idl in H''. 78 | rewrite idr in H''. 79 | apply hom_sym. auto. 80 | Qed. 81 | 82 | 83 | (* 84 | Program Definition NTinv: NT G F := Build_NT 85 | (trafo := @inv _ _ _ _ _ _ F G alpha (niso_struct alpha) ) _ . 86 | Next Obligation. 87 | Proof. 88 | apply Build_NatTrans. 89 | unfold trafo_comm, inv. 90 | intros c d f. 91 | destruct alpha as [al da]. simpl. 92 | destruct da as [invv r l]. simpl in *|-*. 93 | destruct al as [al H]. simpl in *|-*. 94 | destruct H as [H]. 95 | unfold trafo_comm in H. simpl in H. 96 | assert (H': invv c ;; (al c ;; !G f) == 97 | invv c ;; (!F f ;; al d)). 98 | apply praecomp. auto. 99 | assert (H'': (invv c ;; (al c ;; !G f)) ;; invv d == 100 | (invv c ;; (!F f ;; al d)) ;; invv d). 101 | apply postcomp. auto. 102 | repeat rewrite assoc in H''. 103 | rewrite r in H''. 104 | repeat rewrite <- assoc in H''. 105 | rewrite l in H''. 106 | rewrite idl in H''. 107 | rewrite idr in H''. 108 | apply hom_sym. auto. 109 | Qed. 110 | *) 111 | End Nat_Iso_gives_NT. 112 | 113 | 114 | (* 115 | Section ISO. 116 | 117 | Variables obC obD: Type. 118 | Variable morC : obC -> obC -> Type. 119 | Variable morD : obD -> obD -> Type. 120 | Variable C: Cat morC. 121 | Variable D: Cat morD. 122 | Variables F G: Functor C D. 123 | 124 | Class 125 | 126 | *) 127 | 128 | 129 | 130 | 131 | 132 | 133 | 134 | 135 | 136 | 137 | 138 | 139 | 140 | 141 | 142 | 143 | 144 | 145 | 146 | 147 | 148 | 149 | 150 | 151 | 152 | 153 | 154 | -------------------------------------------------------------------------------- /CAT/order_classes.v: -------------------------------------------------------------------------------- 1 | Require Export Coq.Classes.RelationClasses. 2 | Require Export Coq.Classes.Morphisms. 3 | Require Import Coq.Relations.Relation_Definitions. 4 | 5 | Require Export CatSem.CAT.cat_INDEXED_TYPE. 6 | 7 | Set Implicit Arguments. 8 | Unset Strict Implicit. 9 | 10 | Unset Automatic Introduction. 11 | 12 | Section Classes. 13 | 14 | Class Rel (X : Type) := rel : relation X. 15 | 16 | Variable T : Type. 17 | 18 | Class IRel (X : T -> Type) := irel :> forall t, Rel (X t). 19 | 20 | Class IPreOrder {A} (R : IRel A) := 21 | ipreorder :> forall t : T, PreOrder (R t) . 22 | 23 | 24 | Class IMor {A}{B}(Ra : IRel A) (Rb : IRel B) (f : A ---> B) : Prop := 25 | imor : forall t, Proper (irel (t:=t) ++> irel (t:=t)) (f t). 26 | 27 | Inductive optrelT (u:T) (V: T -> Type) (R : IRel V) : IRel (opt u V) := 28 | | optrelTP_none : optrelT _ (none u V) (none u V) 29 | | optrelTP_some : forall t (y z:V t), R _ y z -> 30 | optrelT _ (some _ y) (some _ z). 31 | 32 | 33 | End Classes. 34 | 35 | Existing Instance optrelT. 36 | -------------------------------------------------------------------------------- /CAT/orders_bis.v: -------------------------------------------------------------------------------- 1 | 2 | Require Export CatSem.CAT.orders. 3 | 4 | 5 | Set Implicit Arguments. 6 | Unset Strict Implicit. 7 | 8 | Unset Automatic Introduction. 9 | 10 | 11 | Inductive smallest_rel (V : TYPE) : relation V := 12 | | t_eq : forall (x : V), smallest_rel x x. 13 | 14 | Program Instance sm_po_struct (V : TYPE) : PO_obj_struct V := { 15 | Rel := smallest_rel (V:=V) 16 | }. 17 | Next Obligation. 18 | Proof. 19 | constructor. 20 | unfold Reflexive. 21 | constructor. 22 | unfold Transitive. 23 | intros x y z H H'. 24 | induction H. 25 | auto. 26 | Qed. 27 | 28 | Canonical Structure sm_po (V : TYPE) : Ord := 29 | Build_PO_obj (sm_po_struct V). 30 | 31 | Section sm_po_mor. 32 | 33 | Variables V W : TYPE. 34 | Variable f : V ---> W. 35 | 36 | Program Instance sm_po_mor_s : 37 | PO_mor_struct (a:=sm_po V) (b:=sm_po W) f. 38 | Next Obligation. 39 | Proof. 40 | unfold Proper. 41 | red. 42 | intros x y H. 43 | induction H. 44 | constructor. 45 | Qed. 46 | 47 | Canonical Structure sm_po_mor : sm_po V ---> sm_po W := 48 | Build_PO_mor sm_po_mor_s. 49 | 50 | End sm_po_mor. 51 | 52 | Program Instance SM_po_s : Functor_struct sm_po_mor. 53 | Next Obligation. 54 | Proof. 55 | unfold Proper. 56 | red. 57 | intros. 58 | simpl. 59 | auto. 60 | Qed. 61 | 62 | Definition SM_po := Build_Functor SM_po_s. 63 | Definition Delta := SM_po. 64 | 65 | Section Sm_ind. 66 | 67 | Variable V : TYPE. 68 | Variable W : Ord. 69 | 70 | Variable f : V -> W. 71 | 72 | Program Instance Sm_ind_s : PO_mor_struct 73 | (a:=SM_po V) (b:=W) f. 74 | Next Obligation. 75 | Proof. 76 | unfold Proper; 77 | red. 78 | intros x y H. 79 | induction H. 80 | reflexivity. 81 | Qed. 82 | 83 | Canonical Structure Sm_ind := Build_PO_mor Sm_ind_s. 84 | 85 | End Sm_ind. 86 | -------------------------------------------------------------------------------- /CAT/orders_lemmas.v: -------------------------------------------------------------------------------- 1 | Require Export CatSem.CAT.orders. 2 | 3 | Set Implicit Arguments. 4 | Unset Strict Implicit. 5 | Unset Transparent Obligations. 6 | Unset Automatic Introduction. 7 | 8 | Instance Rel_Trans (A : PO_obj) : Transitive (Rel (T:=A)). 9 | Proof. 10 | intros; 11 | apply POprf. 12 | Qed. 13 | 14 | Lemma Rel_refl (V : PO_obj) (x : V) : x << x. 15 | Proof. 16 | reflexivity. 17 | Qed. 18 | 19 | Lemma Rel_eq (V : PO_obj) (x y : V) : x = y -> x << y. 20 | Proof. 21 | intros; 22 | subst; 23 | reflexivity. 24 | Qed. 25 | 26 | Lemma Rel_eq_r (V : PO_obj) (x y z : V) : 27 | x << y -> y = z -> x << z. 28 | Proof. 29 | intros; 30 | subst; 31 | auto. 32 | Qed. 33 | 34 | Lemma IRel_eq_l (V : PO_obj) (x y z : V) : 35 | x << z -> x = y -> y << z. 36 | Proof. 37 | intros; subst; auto. 38 | Qed. 39 | 40 | Lemma Rel_trans (V : PO_obj) (x y z : V) : 41 | x << y -> y << z -> x << z. 42 | Proof. 43 | intros; 44 | etransitivity; 45 | eauto. 46 | Qed. 47 | 48 | Lemma Rel_trans_2 (V : PO_obj) (x y z : V) : 49 | x << z -> y << x -> y << z. 50 | Proof. 51 | intros; 52 | etransitivity; 53 | eauto. 54 | Qed. 55 | -------------------------------------------------------------------------------- /CAT/rich_monads.v: -------------------------------------------------------------------------------- 1 | (*Require Coq.Program.*) 2 | 3 | Require Export CatSem.CAT.monad_h_module. 4 | Require Export CatSem.CAT.cat_INDEXED_TYPE. 5 | Require Export CatSem.CAT.orders. 6 | Require Export CatSem.CAT.ind_potype. 7 | 8 | Generalizable All Variables. 9 | 10 | Set Implicit Arguments. 11 | Unset Strict Implicit. 12 | 13 | Unset Automatic Introduction. 14 | 15 | Section defs. 16 | 17 | Variable T : Type. 18 | 19 | Notation "'IT'" := (ITYPE T). 20 | Notation "'IP'" := (IPO T). 21 | 22 | Section order_stuff. 23 | 24 | Class Rel (X : Type) := rel : relation X. 25 | 26 | Class ITrel (X : IT) := itrel :> forall t, Rel (X t). 27 | 28 | Class ITpo (X : IT)(H:ITrel X) : Prop := 29 | itpo :> forall t, PreOrder (@itrel _ _ t). 30 | 31 | Variables A B : IT. 32 | Variable f : A ---> B. 33 | 34 | Class ITproper (Xa : ITrel A) (Xb : ITrel B) := 35 | itproper :> forall t, Proper (itrel (t:=t) ==> @itrel _ _ t) (@f t). 36 | 37 | End order_stuff. 38 | (* 39 | Section rich_monad. 40 | 41 | Variable P : Monad IT. 42 | Variable Ind_Rel : forall X : IP, ITrel (P X). 43 | 44 | Class Enriched_Monad := { 45 | ipo_preserve : forall X : IP, (ITpo (@Ind_Rel X)); 46 | kl_preserve : forall (X Y : IP) (f : forall t, X t -> P Y t), 47 | ITproper f (@IRel _ _ X) (@Ind_Rel Y) -> 48 | ITproper (kleisli f) (@Ind_Rel _)(@Ind_Rel _ ); 49 | we_preserve : forall X : IP, 50 | ITproper (weta (Monad_struct := P)X) (@IRel T _ X) (@Ind_Rel X ) 51 | }. 52 | 53 | 54 | Variable H : Enriched_Monad. 55 | 56 | Program Instance ind_ipo_struct (X : IPO T) : ipo_obj_struct (P X) := { 57 | IRel := @Ind_Rel X 58 | }. 59 | Next Obligation. 60 | Proof. 61 | apply ipo_preserve. 62 | Qed. 63 | 64 | Definition ind_ipo X := Build_ipo_obj (ind_ipo_struct X). 65 | 66 | Program Instance we_ipo (X : IPO T) : 67 | ipo_mor_struct (a:=X)(b:= ind_ipo (X))(T:=T) 68 | (weta (Monad_struct := P) X). 69 | Next Obligation. 70 | Proof. 71 | apply we_preserve. 72 | Qed. 73 | 74 | Program Instance kl_ipo (X Y: IPO T) (f : X ---> ind_ipo Y) : 75 | ipo_mor_struct (a:=ind_ipo X) (b:=ind_ipo Y) (T:=T) 76 | (kleisli (Monad_struct := P) f). 77 | Next Obligation. 78 | Proof. 79 | apply kl_preserve. 80 | unfold ITproper. 81 | apply f. 82 | Qed. 83 | 84 | Program Instance IPO_Monad : Monad_struct IP 85 | (fun X => Build_ipo_obj (ind_ipo X)) := { 86 | weta X := Build_ipo_mor (we_ipo X) ; 87 | kleisli X Y f := Build_ipo_mor (kl_ipo X Y f) 88 | }. 89 | Next Obligation. 90 | Proof. 91 | unfold Proper; red. 92 | simpl. 93 | intros. 94 | assert (H':=kleisli_oid (Monad_struct:=P)). 95 | unfold Proper in H'; 96 | red in H'. 97 | apply H'. 98 | simpl. 99 | apply H0. 100 | Qed. 101 | Next Obligation. 102 | Proof. 103 | assert (H':=eta_kl (Monad_struct:=P)). 104 | simpl in H'. 105 | apply H'. 106 | Qed. 107 | Next Obligation. 108 | Proof. 109 | assert (H':=kl_eta (Monad_struct:=P)). 110 | simpl in H'. 111 | apply H'. 112 | Qed. 113 | Next Obligation. 114 | Proof. 115 | assert (H':=dist (Monad_struct:=P)). 116 | simpl in H'. 117 | apply H'. 118 | Qed. 119 | 120 | End rich_monad. 121 | *) 122 | 123 | End defs. -------------------------------------------------------------------------------- /CAT/rmodule_pb_rmonad_hom.v: -------------------------------------------------------------------------------- 1 | Require Export CatSem.CAT.rmonad. 2 | Require Export CatSem.CAT.rmonad_hom. 3 | 4 | Set Implicit Arguments. 5 | Unset Strict Implicit. 6 | Unset Transparent Obligations. 7 | Unset Automatic Introduction. 8 | 9 | (** pullback *) 10 | 11 | 12 | Section pullback. 13 | Variables obC obD : Type. 14 | Variable morC : obC -> obC -> Type. 15 | Variable morD : obD -> obD -> Type. 16 | Variable C : Cat_struct morC. 17 | Variable D : Cat_struct morD. 18 | Variable F : Functor C D. 19 | 20 | Variables P Q : RMonad F. 21 | Variable h : RMonad_Hom P Q. 22 | 23 | Section pb_def. 24 | 25 | Variable obE : Type. 26 | Variable morE : obE -> obE -> Type. 27 | Variable E : Cat_struct morE. 28 | 29 | Section PbRMod. 30 | 31 | Variable N : RMOD Q E. 32 | 33 | Obligation Tactic := idtac. 34 | 35 | Program Instance PbRMod_struct : RModule_struct P E N := { 36 | rmkleisli c d f := rmkleisli (f ;; h d ) 37 | }. 38 | Next Obligation. 39 | Proof. 40 | intros c d. 41 | unfold Proper; red. 42 | intros x y H. 43 | rewrite H. 44 | rmonad. 45 | Qed. 46 | Next Obligation. 47 | Proof. 48 | rmonad. 49 | apply h. 50 | Qed. 51 | Next Obligation. 52 | Proof. 53 | rmonad. 54 | apply h. 55 | Qed. 56 | 57 | Definition PbRMod : RMOD P E := Build_RModule PbRMod_struct. 58 | 59 | End PbRMod. 60 | 61 | Section PbRMod_Hom. 62 | 63 | Variables M N : RMOD Q E. 64 | Variable S : M ---> N. 65 | 66 | Obligation Tactic := rmonad. 67 | 68 | Program Instance PbRMod_Hom_struct : RModule_Hom_struct 69 | (M := PbRMod M) (N:=PbRMod N) S. 70 | (*Next Obligation. 71 | Proof. 72 | simpl. 73 | rewrite (rmod_hom_rmkl); 74 | try cat. 75 | apply S. 76 | Qed. 77 | *) 78 | Definition PbRMod_Hom : PbRMod M ---> PbRMod N := 79 | Build_RModule_Hom PbRMod_Hom_struct. 80 | 81 | End PbRMod_Hom. 82 | 83 | Obligation Tactic := unfold Proper; red; simpl; rmonad. 84 | 85 | Program Instance PbRMOD_struct : Functor_struct PbRMod_Hom. 86 | 87 | Canonical Structure PbRMOD := Build_Functor PbRMOD_struct. 88 | 89 | Section pb_prod. 90 | 91 | Variable H : Cat_Prod E. 92 | 93 | Existing Instance RMOD_PROD. 94 | 95 | Variables M N : RModule Q E. 96 | 97 | Notation "a 'x' b" := (product a b)(at level 60). 98 | 99 | Obligation Tactic := cat. 100 | 101 | Program Instance PROD_RPB_s : RModule_Hom_struct 102 | (M:= product (C:=RMOD P E) (PbRMod M) (PbRMod N)) 103 | (N:=PbRMod (product (C:=RMOD Q E) M N)) 104 | (fun c => id (M c x N c) ). 105 | 106 | Definition PROD_RPB : 107 | (PbRMod M x PbRMod N) ---> PbRMod (M x N) := 108 | Build_RModule_Hom PROD_RPB_s. 109 | 110 | End pb_prod. 111 | 112 | End pb_def. 113 | 114 | Program Instance PbRMod_ind_Hom_struct : 115 | RModule_Hom_struct (M:= P) (N:=PbRMOD _ Q) (fun x => h x). 116 | Next Obligation. 117 | Proof. 118 | rmonad; 119 | try apply h. 120 | Qed. 121 | 122 | Definition PbRMod_ind_Hom : Taut_RMod P ---> PbRMod Q := 123 | Build_RModule_Hom PbRMod_ind_Hom_struct. 124 | 125 | 126 | End pullback. 127 | 128 | Coercion PbRMod_ind_Hom : RMonad_Hom >-> mor. 129 | -------------------------------------------------------------------------------- /CAT/rmonad_gen_comp.v: -------------------------------------------------------------------------------- 1 | Require Export CatSem.CAT.retype_functor_po. 2 | Require Export CatSem.CAT.rmonad_gen_type_po. 3 | 4 | 5 | Set Implicit Arguments. 6 | Unset Strict Implicit. 7 | Unset Transparent Obligations. 8 | Unset Automatic Introduction. 9 | 10 | 11 | Section transp_lift_id. 12 | 13 | Variables U U' : Type. 14 | Variable f : U -> U'. 15 | (*Variable P : Monad (ITYPE U).*) 16 | Variable Q : RMonad (IDelta U'). 17 | 18 | Hypothesis H : forall t, f t = f t. 19 | Variable V : IPO U. 20 | Variable t' : U'. 21 | 22 | 23 | Lemma rlift_transp_id : forall y : Q (retype (fun t => f t) V) t', 24 | rlift Q (transp (f:= f)(g:= f) H (V:=V)) 25 | _ y = y. 26 | Proof. 27 | intro y. 28 | assert (H'' := rlift_eq_id Q). 29 | apply H''. 30 | apply transp_id. 31 | Qed. 32 | 33 | End transp_lift_id. 34 | 35 | Section comp. 36 | 37 | Variables U U' U'' : Type. 38 | Variable f : U -> U'. 39 | Variable f' : U' -> U''. 40 | 41 | Variable P : RMonad (IDelta U). 42 | Variable Q : RMonad (IDelta U'). 43 | Variable R : RMonad (IDelta U''). 44 | 45 | Variable M : colax_RMonad_Hom P Q 46 | (G1:=RETYPE (fun t => f t)) 47 | (G2:=RETYPE_PO (fun t => f t)) 48 | (RT_NT (fun t => f t)). 49 | Variable N : colax_RMonad_Hom Q R 50 | (G1:=RETYPE (fun t => f' t)) 51 | (G2:=RETYPE_PO (fun t => f' t)) 52 | (RT_NT (fun t => f' t)). 53 | 54 | Obligation Tactic := idtac. 55 | 56 | Definition RMon_comp_data c: 57 | (forall t, 58 | (retype_ipo (fun t0 => f' (f t0)) (P c)) t -> 59 | (R (retype (fun t0 => f' (f t0)) c)) t) := 60 | fun t 61 | (y : retype (fun t => f' (f t)) (P c) t) => 62 | match y with ctype _ z => 63 | (rlift R (double_retype_1 (f:=f) 64 | (f':=f') (V:=c))) _ 65 | (N _ _ (ctype (fun t => f' t) 66 | (M _ _ (ctype (fun t => f t) z )))) 67 | end. 68 | 69 | Program Instance RMon_cc c : 70 | ipo_mor_struct 71 | (a:=retype_ipo (fun t => f' (f t)) (P c)) 72 | (b:=R (retype (fun t => f' (f t)) c)) 73 | (RMon_comp_data (c:=c)). 74 | Next Obligation. 75 | Proof. 76 | intros c t. 77 | unfold Proper; 78 | red. 79 | simpl. 80 | intros y z H. 81 | induction H. 82 | simpl. 83 | apply (rlift R _ ). 84 | apply (N (retype ( fun t0 => f t0) c)). 85 | constructor. 86 | apply (M c). 87 | constructor; 88 | auto. 89 | Qed. 90 | 91 | 92 | Definition RMon_car: 93 | (forall c : ITYPE U, 94 | (RETYPE_PO (fun t => f' (f t))) (P c) ---> 95 | R ((RETYPE (fun t => f' (f t))) c)) := 96 | fun c => Build_ipo_mor (RMon_cc c). 97 | 98 | Program Instance RMon_comp_s : 99 | colax_RMonad_Hom_struct 100 | (P:=P) (Q:=R) (G1:=RETYPE (fun t => f' (f t))) 101 | (G2:=RETYPE_PO (fun t => f' (f t))) 102 | (RT_NT (fun t => f' (f t))) 103 | RMon_car. 104 | Next Obligation. 105 | Proof. 106 | simpl. 107 | intros c t z. 108 | destruct z as [t z]; 109 | simpl. 110 | rew (gen_rmonad_hom_rweta M _ _ (ctype _ z)). 111 | simpl. 112 | assert (H:=gen_rmonad_hom_rweta N ). 113 | simpl in H. 114 | set (y := ctype (fun t0 => f t0) z). 115 | simpl in *. 116 | rew (H _ _ (ctype (*fun t => tcomp N t*) _ y)). 117 | rew (rlift_rweta R). 118 | Qed. 119 | Next Obligation. 120 | Proof. 121 | simpl. 122 | intros V W g t z. 123 | destruct z as [t y]. 124 | simpl. 125 | rew (gen_rh_rkl M). 126 | rew (gen_rh_rkl N). 127 | rew (rlift_rkleisli (M:=R)). 128 | rew (rkleisli_rlift (M:=R)). 129 | apply (rkl_eq R). 130 | 131 | intros t0 z. 132 | destruct z as [t0 z]. 133 | simpl. 134 | destruct z as [t0 z]. 135 | simpl. 136 | auto. 137 | Qed. 138 | 139 | 140 | Definition RMon_comp := Build_colax_RMonad_Hom RMon_comp_s. 141 | 142 | End comp. 143 | -------------------------------------------------------------------------------- /CAT/rmonad_hom.v: -------------------------------------------------------------------------------- 1 | 2 | Require Export CatSem.CAT.rmonad. 3 | 4 | Set Implicit Arguments. 5 | Unset Strict Implicit. 6 | Unset Transparent Obligations. 7 | Unset Automatic Introduction. 8 | 9 | Section rmonad_hom. 10 | 11 | Variables obC obD : Type. 12 | Variable morC : obC -> obC -> Type. 13 | Variable morD : obD -> obD -> Type. 14 | Variable C : Cat_struct morC. 15 | Variable D : Cat_struct morD. 16 | Variable F : Functor C D. 17 | 18 | Section rmonad_hom. 19 | 20 | Variables P Q : RMonad F. 21 | 22 | 23 | Class RMonad_Hom_struct (tau : forall c : obC, morD (P c) (Q c)) := { 24 | rmonad_hom_rweta : forall c:obC, 25 | rweta c ;; tau c == rweta c ; 26 | rmonad_hom_rkl : forall c d (f : morD (F c) (P d)), 27 | rkleisli f ;; tau d == tau c ;; rkleisli (f ;; tau d) 28 | }. 29 | 30 | Record RMonad_Hom := { 31 | rmonad_hom :> forall c, morD (P c) (Q c); 32 | rmonad_hom_struct :> RMonad_Hom_struct rmonad_hom 33 | }. 34 | 35 | Section RMonad_Hom_NT. 36 | 37 | Variable tau : forall c : obC, morD (P c) (Q c). 38 | 39 | Variable tua : RMonad_Hom_struct tau. 40 | 41 | (** trivial lemmata *) 42 | 43 | Lemma rmon_hom_rweta (c : obC) : rweta c ;; tau c == rweta c. 44 | Proof. 45 | apply rmonad_hom_rweta. 46 | Qed. 47 | 48 | Lemma rmon_hom_rkl (c d : obC) (f : morD (F c) (P d)) : 49 | rkleisli f ;; tau d == tau c ;; rkleisli (f ;; tau d). 50 | Proof. 51 | apply rmonad_hom_rkl. 52 | Qed. 53 | 54 | Hint Rewrite rmon_hom_rweta : rmonad. 55 | Hint Resolve rmon_hom_rkl : rmonad. 56 | Hint Rewrite rmon_hom_rkl. 57 | 58 | Obligation Tactic := idtac. 59 | 60 | Program Instance RMonad_Hom_NT : 61 | NT_struct (F:=RMFunc P) (G:=RMFunc Q) tau. 62 | Next Obligation. 63 | Proof. 64 | simpl. 65 | intros c d f. 66 | unfold rlift. 67 | rewrite (rmonad_hom_rkl). 68 | rewrite assoc. 69 | rmonad. 70 | Qed. 71 | 72 | End RMonad_Hom_NT. 73 | 74 | End rmonad_hom. 75 | 76 | 77 | Section RMonad_id_comp. 78 | 79 | Variable P : RMonad F. 80 | 81 | Obligation Tactic := rmonad. 82 | 83 | Program Instance RMonad_id_s : 84 | RMonad_Hom_struct (P:=P) (fun _ => id _ ). 85 | 86 | Definition RMonad_id := Build_RMonad_Hom RMonad_id_s. 87 | 88 | Variables Q R : RMonad F. 89 | Variable S : RMonad_Hom P Q. 90 | Variable T : RMonad_Hom Q R. 91 | 92 | Program Instance RMonad_comp_s : 93 | RMonad_Hom_struct (fun c => S c ;; T c). 94 | Next Obligation. 95 | Proof. 96 | rewrite <- assoc. 97 | rewrite (rmon_hom_rweta S). 98 | rewrite (rmon_hom_rweta T). 99 | cat. 100 | Qed. 101 | Next Obligation. 102 | Proof. 103 | rewrite <- assoc. 104 | rewrite (rmon_hom_rkl S). 105 | rewrite assoc. 106 | rewrite (rmon_hom_rkl T). 107 | rmonad. 108 | Qed. 109 | 110 | Definition RMonad_comp := Build_RMonad_Hom RMonad_comp_s. 111 | 112 | End RMonad_id_comp. 113 | 114 | Lemma RMonad_Hom_equiv (P Q: RMonad F) : 115 | Equivalence (A:=RMonad_Hom P Q) (fun M N => forall x, M x == N x). 116 | Proof. 117 | intros P Q; 118 | constructor. 119 | unfold Reflexive. 120 | cat. 121 | unfold Symmetric; 122 | intros; apply hom_sym; 123 | auto. 124 | unfold Transitive. 125 | intros. 126 | etransitivity; auto. 127 | Qed. 128 | 129 | Definition RMonad_Hom_oid (P Q: RMonad F) := 130 | Build_Setoid (RMonad_Hom_equiv P Q). 131 | 132 | Obligation Tactic := simpl; intros; try apply assoc; cat. 133 | 134 | Program Instance RMONAD_struct : Cat_struct RMonad_Hom := { 135 | mor_oid := RMonad_Hom_oid ; 136 | comp := RMonad_comp ; 137 | id := RMonad_id 138 | }. 139 | Next Obligation. 140 | Proof. 141 | unfold Proper; do 2 red. 142 | simpl. 143 | intros x y H x' y' H' z; 144 | rewrite H; 145 | rewrite H'; 146 | cat. 147 | Qed. 148 | 149 | Definition RMONAD := Build_Cat RMONAD_struct. 150 | 151 | End rmonad_hom. 152 | 153 | Existing Instance rmonad_hom_struct. 154 | 155 | Hint Rewrite rmon_hom_rweta : rmonad. 156 | Hint Resolve rmon_hom_rkl : rmonad. 157 | Hint Rewrite rmon_hom_rkl : rmonad. -------------------------------------------------------------------------------- /CAT/small_cat.v: -------------------------------------------------------------------------------- 1 | Require Export CatSem.CAT.CatFunct. 2 | Require Export CatSem.CAT.functor_leibniz_eq. 3 | 4 | Set Implicit Arguments. 5 | Unset Strict Implicit. 6 | 7 | Unset Automatic Introduction. 8 | 9 | Class SmallCat := { 10 | sobj: Set; 11 | smor: sobj -> sobj -> Set; 12 | sstruct:> Cat_struct smor 13 | }. 14 | 15 | Definition Cat_from_SmallCat (C: SmallCat) : Cat := Build_Cat 16 | (obj := sobj) (mor := smor) sstruct. 17 | 18 | Coercion Cat_from_SmallCat: SmallCat >-> Cat. 19 | 20 | Obligation Tactic := simpl; intros; 21 | try apply Equal_hom_refl; 22 | try apply CompF_oid. 23 | 24 | Program Instance SMALLCAT_struct : 25 | @Cat_struct SmallCat (fun a b => FunctCat a b) := { 26 | mor_oid a b := EXT_Functor_oid a b; 27 | id a := Id a; 28 | comp a b c f g := CompF f g 29 | }. 30 | 31 | Definition SMALLCAT := Build_Cat SMALLCAT_struct. 32 | 33 | 34 | -------------------------------------------------------------------------------- /CAT/type_unit.v: -------------------------------------------------------------------------------- 1 | Require Export CatSem.CAT.cat_INDEXED_TYPE. 2 | Require Export CatSem.CAT.ind_potype. 3 | Require Export CatSem.CAT.cat_TYPE_option. 4 | 5 | Set Implicit Arguments. 6 | Unset Strict Implicit. 7 | Unset Transparent Obligations. 8 | Unset Automatic Introduction. 9 | 10 | 11 | 12 | Definition unit_passing (c : unit -> Type) t u 13 | (y : (opt t c u)): 14 | (option (c tt)). 15 | intros. 16 | destruct y. 17 | destruct t0. 18 | apply (Some c0). 19 | apply (None). 20 | Defined. 21 | 22 | 23 | Definition unit_passing2 c t u (y : opt t c u) : 24 | (option (c tt)). 25 | intros. 26 | destruct u. 27 | destruct y. 28 | apply (Some c0). 29 | apply None. 30 | Defined. 31 | 32 | 33 | Section bla. 34 | 35 | Notation "'TT'":= (ITYPE unit). 36 | 37 | Obligation Tactic := mauto. 38 | 39 | (** a functor (unit -> Type) -> Type *) 40 | 41 | Program Instance sunit_s : Functor_struct (C:=TT) (D:=TYPE) 42 | (Fobj:= fun V => V tt) (fun a b f => f tt). 43 | 44 | Definition sunit : Functor TT TYPE := IT_proj tt. 45 | 46 | (** a functor Type -> (unit -> Type) *) 47 | 48 | Program Instance wunit_s : Functor_struct (C:=TYPE) (D:=TT) 49 | (Fobj := fun V => fun _ => V)(fun a b f => fun _ => f). 50 | 51 | Definition wunit := Build_Functor wunit_s. 52 | 53 | (** a functor (unit -> PO) -> PO = (IPO unit) -> PO*) 54 | 55 | Definition sunit_po : Functor (IPO unit) Ord := IP_proj tt. 56 | 57 | (** a functor PO -> (IPO unit) *) 58 | 59 | Program Instance wunit_ob_s (V : Ord): 60 | ipo_obj_struct (T:=unit)(fun _ => V) := { 61 | IRel t := Rel ; 62 | IPOprf t := POprf 63 | }. 64 | 65 | Definition wunit_ob V := Build_ipo_obj (wunit_ob_s V). 66 | 67 | Obligation Tactic := cat; try unf_Proper; cat; 68 | repeat apply PO_mor_monotone; cat. 69 | 70 | Program Instance wunit_mor_s (V W : Ord)(f : V ---> W) : 71 | ipo_mor_struct (a:=wunit_ob V) (b:=wunit_ob W) 72 | (fun t => f). 73 | 74 | Definition wunit_mor V W (f:V--->W) := 75 | Build_ipo_mor (wunit_mor_s f). 76 | 77 | Obligation Tactic := mauto. 78 | 79 | Program Instance wunit_po_s : Functor_struct (C:=Ord)(D:=IPO unit) 80 | (Fobj:=wunit_ob) 81 | (fun a b f => wunit_mor f). 82 | 83 | Definition wunit_po : Functor Ord (IPO unit) := Build_Functor wunit_po_s. 84 | 85 | End bla. 86 | 87 | 88 | 89 | 90 | 91 | 92 | 93 | -------------------------------------------------------------------------------- /CAT/unit_mod.v: -------------------------------------------------------------------------------- 1 | Require Export CatSem.CAT.monad_h_morphism_gen. 2 | Require Export CatSem.CAT.cat_TYPE. 3 | 4 | 5 | Set Implicit Arguments. 6 | Unset Strict Implicit. 7 | Unset Transparent Obligations. 8 | Unset Automatic Introduction. 9 | 10 | 11 | Section unit_mod. 12 | (* 13 | Variable obC : Type. 14 | Variable morC : obC -> obC -> Type. 15 | Variable C : Cat morC. 16 | 17 | Variable obD : Type. 18 | Variable morD : obD -> obD -> Type. 19 | Variable D : Cat morD. 20 | *) 21 | Variables C D : Cat. 22 | Variable P : Monad C. 23 | Variable R : Monad D. 24 | Variable F : Functor C D. 25 | 26 | Variable M : colax_Monad_Hom P R F. 27 | 28 | 29 | Program Instance unit_mod_s : Module_Hom_struct 30 | (S:=Term (C:=MOD P TYPE)) 31 | (T:=PModule M (Term (C:=MOD R TYPE))) 32 | (fun V y => y). 33 | 34 | Definition unit_mod := Build_Module_Hom unit_mod_s. 35 | 36 | End unit_mod. 37 | -------------------------------------------------------------------------------- /CAT/unit_type_monad.v: -------------------------------------------------------------------------------- 1 | 2 | Require Export CatSem.CAT.type_unit. 3 | Require Export CatSem.CAT.monad_haskell. 4 | 5 | Section bla. 6 | 7 | Variable R : Monad TYPE. 8 | 9 | Definition unit_weta_car : 10 | forall (c : ITYPE unit) (t : unit), 11 | (c) t -> (wunit (R (c tt))) t := 12 | fun (c : unit -> Type) (t : unit) (X : c t) => 13 | match t as u return (c u -> R (c tt)) with 14 | | tt => fun X0 : c tt => (weta (Monad_struct:=R)(c tt)) X0 15 | end X. 16 | 17 | 18 | 19 | 20 | Definition unit_kleisli : 21 | forall a b : ITYPE unit, 22 | a ---> wunit (R (sunit b)) -> 23 | wunit (R (sunit a)) ---> wunit (R (sunit b)) := 24 | fun a b f => 25 | fun t => kleisli (Monad_struct := R) (#sunit f). 26 | 27 | Obligation Tactic := cat; try unf_Proper; 28 | unfold unit_weta_car, unit_kleisli; 29 | cat; repeat elim_unit; 30 | try apply (kl_eq R); try rew (etakl R); 31 | try rerew (kleta R); 32 | repeat rew (klkl R); 33 | repeat apply (kl_eq R); cat; 34 | rew (kleta R); cat. 35 | 36 | Program Instance unit_Monad_struct : 37 | Monad_struct (C:=ITYPE unit) (fun V => wunit (R (sunit V))) := { 38 | weta := unit_weta_car ; 39 | kleisli := unit_kleisli 40 | }. 41 | 42 | Definition unit_Monad := Build_Monad unit_Monad_struct. 43 | 44 | End bla. 45 | 46 | -------------------------------------------------------------------------------- /CAT/unit_type_rmonad.v: -------------------------------------------------------------------------------- 1 | Require Export CatSem.CAT.ind_potype. 2 | Require Export CatSem.CAT.orders_bis. 3 | Require Export CatSem.CAT.type_unit. 4 | Require Export CatSem.CAT.rmonad. 5 | 6 | Section bla. 7 | 8 | (** starting from a rmonad Set -> PO, 9 | we define a rmonad (unit -> Set) -> (unit -> PO) 10 | *) 11 | 12 | Variable R : RMonad (SM_po). 13 | 14 | Definition unit_rweta_car : 15 | forall (c : ITYPE unit) (t : unit), 16 | (sm_ipo (T:=unit) c) t -> (wunit_ob (R (c tt))) t := 17 | fun (c : unit -> Type) (t : unit) (X : c t) => 18 | match t as u return (c u -> R (c tt)) with 19 | | tt => fun X0 : c tt => (rweta (RMonad_struct:=R)(c tt)) X0 20 | end X. 21 | 22 | Obligation Tactic := cat; unf_Proper; cat; 23 | repeat match goal with [H:smallest_irel _ _ |- _ ] => destruct H end; 24 | reflexivity. 25 | 26 | Program Instance unit_rweta_po c: 27 | ipo_mor_struct 28 | (a:=sm_ipo (T:=unit) c) (b:=wunit_ob (R (c tt))) 29 | (unit_rweta_car c). 30 | 31 | Definition unit_rweta c: 32 | (IDelta unit) c ---> wunit_po (R (sunit c)) := 33 | Build_ipo_mor (unit_rweta_po c). 34 | 35 | Obligation Tactic := cat; try unf_Proper; cat; 36 | repeat apply PO_mor_monotone; cat. 37 | 38 | Program Instance unit_rkleisli_po : 39 | forall (a b : unit -> Type) 40 | (f:ipo_mor (sm_ipo (T:=unit) a) (wunit_ob (R (b tt)))), 41 | ipo_mor_struct (a:=wunit_ob (R (a tt))) (b:=wunit_ob (R (b tt))) 42 | (fun t => rkleisli (RMonad_struct := R) (Sm_ind (#sunit f))). 43 | 44 | Definition unit_rkleisli : 45 | forall (a b : ITYPE unit) 46 | (f : (IDelta unit) a ---> wunit_po (R (sunit b))), 47 | wunit_po (R (sunit a)) ---> wunit_po (R (sunit b)) := 48 | fun a b f => Build_ipo_mor (unit_rkleisli_po a b f). 49 | 50 | Obligation Tactic := cat; try unf_Proper; 51 | unfold unit_rweta_car; 52 | cat; repeat elim_unit; 53 | try apply (rkl_eq R); try rew (retakl R); 54 | try rerew (rkleta R); 55 | repeat rew (rklkl R); 56 | repeat apply (rkl_eq R); cat; 57 | rew (rkleta R); cat. 58 | 59 | Program Instance unit_RMonad_struct : 60 | RMonad_struct (IDelta unit) (fun V => wunit_po (R (sunit V))) := { 61 | rweta := unit_rweta ; 62 | rkleisli := unit_rkleisli 63 | }. 64 | 65 | Definition unit_RMonad := Build_RMonad unit_RMonad_struct. 66 | 67 | End bla. 68 | 69 | -------------------------------------------------------------------------------- /COMP/PCF_ULC_comp.v: -------------------------------------------------------------------------------- 1 | Require Export CatSem.COMP.PCF_quant. 2 | Require Export CatSem.COMP.PCF_ULC_nounit. 3 | 4 | Set Implicit Arguments. 5 | Unset Strict Implicit. 6 | Unset Printing Implicit Defensive. 7 | 8 | Definition PCF_ULC_compilation := 9 | InitMor PCF_ULC. 10 | 11 | Definition PCF_ULC_c := 12 | rep_Hom_monad PCF_ULC_compilation. 13 | 14 | (** some examples *) 15 | (* 16 | Eval compute in 17 | (PCF_ULC_c (fun t => False) tt 18 | (ctype _ (Bottom (fun _ => False) Nat))). 19 | Eval compute in 20 | (PCF_ULC_c (fun t => False) tt 21 | (ctype _ (Const (fun _ => False) (succ )))). 22 | Eval compute in 23 | (PCF_ULC_c (fun t => False) tt 24 | (ctype _ (Const (fun _ => False) (preds )))). 25 | 26 | Eval compute in 27 | (PCF_ULC_c (fun t => False) tt 28 | (ctype _ (Const (fun _ => False) (condB )))). 29 | 30 | 31 | Eval compute in 32 | (PCF_ULC_c ( (fun t => False)) tt 33 | (ctype _ ( 34 | Lam (PCF_syntax.Var (none Bool (fun t => False)))))). 35 | *) 36 | -------------------------------------------------------------------------------- /COMP/PCF_rep_quant.v: -------------------------------------------------------------------------------- 1 | Require Import Coq.Logic.FunctionalExtensionality. 2 | (*Require Import Coq.Program.Equality.*) 3 | 4 | Require Export CatSem.CAT.cat_INDEXED_TYPE. 5 | Require Export CatSem.CAT.retype_functor. 6 | Require Export CatSem.CAT.monad_h_morphism_gen. 7 | Require Export CatSem.CAT.monad_h_module. 8 | Require Export CatSem.PCF.PCF_types. 9 | 10 | Set Implicit Arguments. 11 | Unset Strict Implicit. 12 | Unset Transparent Obligations. 13 | Unset Automatic Introduction. 14 | 15 | Notation "'TY'" := PCF.Sorts. 16 | Notation "'Bool'":= PCF.Bool. 17 | Notation "'Nat'":= PCF.Nat. 18 | 19 | Section rep. 20 | 21 | Notation "'IT'" := (ITYPE TY). 22 | Notation "a '~>' b" := (PCF.Arrow a b) 23 | (at level 69, right associativity). 24 | Notation "a 'x' b" := (product (*C:= MOD _ _*) a b) (at level 30). 25 | Notation "P [ z ]" := (ITFIB_MOD _ z P) (at level 35). 26 | Notation "'d' P // s" := (ITDER_MOD _ _ s P) (at level 25). 27 | Notation "'*'" := (Term (C:=MOD _ TYPE)). 28 | Notation "f 'X' g" := (product_mor _ f g)(at level 30). 29 | 30 | Section PCF_rep. 31 | 32 | Variable U : Type. 33 | Variable P : Monad (ITYPE U). 34 | Variable f : TY -> U. 35 | 36 | Variable Arrow : U -> U -> U. 37 | Notation "a ~~> b" := (Arrow a b) (at level 60, right associativity). 38 | 39 | (* don't put it here, but we need it in the record, 40 | for the initial morphism has to have this 41 | property 42 | Hypothesis type_arrow : forall s t, f (s ~> t) = f s ~~> f t. 43 | *) 44 | 45 | (** a monad is a representation if it is accompagnied by 46 | - a lot of morphisms of modules 47 | - beta-reduction 48 | *) 49 | 50 | Class PCF_rep_struct := { 51 | app : forall u v, (P[u ~~> v]) x (P[u]) ---> P[v]; 52 | abs : forall u v, (d P //u)[v] ---> P[u ~~> v]; 53 | rec : forall t, P[t ~~> t] ---> P[t]; 54 | tttt : * ---> P[f Bool]; 55 | ffff : * ---> P[f Bool]; 56 | nats : forall m:nat, * ---> P[f Nat]; 57 | Succ : * ---> P[f Nat ~~> f Nat]; 58 | Pred : * ---> P[f Nat ~~> f Nat]; 59 | Zero : * ---> P[f Nat ~~> f Bool]; 60 | CondN: * ---> P[f Bool ~~> f Nat ~~> f Nat ~~> f Nat]; 61 | CondB: * ---> P[f Bool ~~> f Bool ~~> f Bool ~~> f Bool]; 62 | bottom: forall t, * ---> P[t] 63 | }. 64 | 65 | End PCF_rep. 66 | 67 | Record PCF_rep := { 68 | type_type : Type; 69 | type_arrow : type_type -> type_type -> type_type; 70 | pcf_rep_monad :> Monad (ITYPE type_type); 71 | type_mor : TY -> type_type; 72 | type_arrow_dist : forall s t, type_mor (s ~> t) = 73 | type_arrow (type_mor s) (type_mor t); 74 | pcf_rep_struct :> PCF_rep_struct pcf_rep_monad type_mor type_arrow 75 | 76 | }. 77 | 78 | End rep. 79 | 80 | Existing Instance pcf_rep_struct. 81 | Notation "a ~~> b" := (type_arrow a b) 82 | (at level 60, right associativity). 83 | 84 | 85 | 86 | 87 | 88 | 89 | 90 | 91 | 92 | 93 | 94 | 95 | 96 | 97 | 98 | 99 | 100 | 101 | 102 | 103 | 104 | 105 | 106 | 107 | 108 | -------------------------------------------------------------------------------- /COMP/f_induced_monad.v: -------------------------------------------------------------------------------- 1 | Require Import CatSem.CAT.monad_h_module. 2 | Require Import CatSem.CAT.cat_INDEXED_TYPE. 3 | Require Import CatSem.CAT.retype_functor. 4 | 5 | Set Implicit Arguments. 6 | Unset Strict Implicit. 7 | 8 | Unset Automatic Introduction. 9 | 10 | Section f. 11 | 12 | Variables T U : Type. 13 | 14 | Variable f : T -> U. 15 | 16 | Variable P : Monad (ITYPE U). 17 | 18 | Inductive f_P (V : ITYPE T)(t : T) : Type := 19 | | copy : P (retype f V) (f t) -> f_P V t. 20 | 21 | Definition inv : forall V t, f_P V t -> P (retype f V) (f t). 22 | intros. 23 | destruct X. 24 | apply f0. 25 | Defined. 26 | 27 | Hint Resolve copy ctype. 28 | 29 | Notation "' V" := (retype _ V) (at level 20). 30 | 31 | 32 | 33 | Definition bla V W (x : forall t : T, V t -> f_P W t): 34 | forall t : U, (retype f V) t -> P (retype f W) t. 35 | intros. 36 | destruct X. 37 | apply inv. 38 | auto. 39 | Defined. 40 | 41 | Obligation Tactic := idtac. 42 | 43 | Program Instance f_P_Mon_s : Monad_struct (ITYPE T)(f_P). 44 | Next Obligation. 45 | simpl. 46 | intros. 47 | apply copy. 48 | apply (weta (Monad_struct := P)). 49 | apply ctype. 50 | apply X. 51 | Defined. 52 | Next Obligation. 53 | simpl. 54 | intros V W x. 55 | set (Z := kleisli (Monad_struct := P)). 56 | set (z := Z (retype f V) (retype f W)). 57 | simpl in *. 58 | set (z':= z (bla x)). 59 | intros. 60 | 61 | destruct X. 62 | apply copy. 63 | apply z'. 64 | apply f0. 65 | Defined. 66 | Next Obligation. 67 | Proof. 68 | unfold Proper,f_P_Mon_s_obligation_2. 69 | red. 70 | intros V W g g' H. 71 | intros t x. 72 | induction x. 73 | apply f_equal. 74 | apply (kleisli_oid (Monad_struct := P)). 75 | simpl. 76 | intros t0 x. 77 | induction x. 78 | simpl. 79 | rewrite H. 80 | auto. 81 | Qed. 82 | Next Obligation. 83 | Proof. 84 | simpl. 85 | intros. 86 | assert (H:=etakl P). 87 | simpl in *. 88 | rewrite H. 89 | simpl. 90 | generalize (f0 t x). 91 | intros. 92 | induction f1. 93 | auto. 94 | Qed. 95 | Next Obligation. 96 | Proof. 97 | simpl. 98 | intros V t x. 99 | induction x. 100 | simpl. 101 | apply f_equal. 102 | assert (H:=kleta P). 103 | simpl in H. 104 | unfold f_P_Mon_s_obligation_1. 105 | simpl in *. 106 | rewrite <- H. 107 | apply (kl_eq P). 108 | simpl. 109 | intros. 110 | induction x. 111 | simpl. 112 | auto. 113 | Qed. 114 | Next Obligation. 115 | Proof. 116 | simpl. 117 | intros. 118 | induction x. 119 | simpl. 120 | apply f_equal. 121 | assert (H:=dist (Monad_struct := P)). 122 | simpl in H. 123 | rewrite H. 124 | apply (kl_eq P). 125 | simpl. 126 | intros. 127 | induction x. 128 | simpl. 129 | generalize (f0 t0 v). 130 | intros f2. 131 | induction f2. 132 | simpl. 133 | auto. 134 | Qed. 135 | 136 | Definition f_P_Mon := Build_Monad f_P_Mon_s. 137 | 138 | 139 | Definition bla2: 140 | (forall c : ITYPE T, 141 | (RETYPE (U:=T) (U':=U) f) (f_P_Mon c) ---> 142 | P ((RETYPE (U:=T) (U':=U) f) c)). 143 | intros V u x. 144 | simpl in *. 145 | induction x. 146 | induction v. 147 | apply f0. 148 | Defined. 149 | 150 | Program Instance pb_mon_s : 151 | gen_Monad_Hom_struct (P:=f_P_Mon) (Q:=P) (F0:=RETYPE f) bla2. 152 | Next Obligation. 153 | Proof. 154 | simpl. 155 | intros V W g u x. 156 | induction x. 157 | simpl. 158 | induction v. 159 | simpl. 160 | apply (kl_eq P). 161 | simpl. 162 | intros t0 x. 163 | induction x. 164 | simpl. 165 | generalize (g t0 v). 166 | intros. 167 | induction f1. 168 | simpl. 169 | auto. 170 | Qed. 171 | Next Obligation. 172 | Proof. 173 | simpl. 174 | intros V t x. 175 | induction x. 176 | simpl. 177 | auto. 178 | Qed. 179 | 180 | Definition pb_mon := Build_gen_Monad_Hom pb_mon_s. 181 | 182 | Variable W : Monad (ITYPE T). 183 | 184 | Variable r : gen_Monad_Hom W P (RETYPE f). 185 | 186 | Definition car : (forall c : ITYPE T, W c ---> f_P_Mon c). 187 | simpl. 188 | intros. 189 | set (r' := r c). 190 | simpl in *. 191 | apply copy. 192 | apply r. simpl. 193 | apply (ctype _ X). 194 | Defined. 195 | (* 196 | Program Instance fac : Monad_Hom_struct 197 | (P := W) (Q := f_P_Mon) car. 198 | Next Obligation. 199 | Proof. 200 | simpl. 201 | intros. 202 | unfold car. 203 | apply f_equal. 204 | assert (H:=gen_monad_hom_kl (gen_Monad_Hom_struct := r)). 205 | simpl in H. 206 | unfold retype_map in H. 207 | simpl in H. 208 | rewrite <- H. 209 | simpl. 210 | *) 211 | End f. 212 | -------------------------------------------------------------------------------- /LICENSE.md: -------------------------------------------------------------------------------- 1 | This is free and unencumbered software released into the public domain. 2 | 3 | Anyone is free to copy, modify, publish, use, compile, sell, or 4 | distribute this software, either in source code form or as a compiled 5 | binary, for any purpose, commercial or non-commercial, and by any 6 | means. 7 | 8 | In jurisdictions that recognize copyright laws, the author or authors 9 | of this software dedicate any and all copyright interest in the 10 | software to the public domain. We make this dedication for the benefit 11 | of the public at large and to the detriment of our heirs and 12 | successors. We intend this dedication to be an overt act of 13 | relinquishment in perpetuity of all present and future rights to this 14 | software under copyright law. 15 | 16 | THE SOFTWARE IS PROVIDED "AS IS", WITHOUT WARRANTY OF ANY KIND, 17 | EXPRESS OR IMPLIED, INCLUDING BUT NOT LIMITED TO THE WARRANTIES OF 18 | MERCHANTABILITY, FITNESS FOR A PARTICULAR PURPOSE AND NONINFRINGEMENT. 19 | IN NO EVENT SHALL THE AUTHORS BE LIABLE FOR ANY CLAIM, DAMAGES OR 20 | OTHER LIABILITY, WHETHER IN AN ACTION OF CONTRACT, TORT OR OTHERWISE, 21 | ARISING FROM, OUT OF OR IN CONNECTION WITH THE SOFTWARE OR THE USE OR 22 | OTHER DEALINGS IN THE SOFTWARE. 23 | 24 | For more information, please refer to 25 | -------------------------------------------------------------------------------- /ORDER/pcf_order_rep_eq.v: -------------------------------------------------------------------------------- 1 | Require Export CatSem.ORDER.pcf_order_rep. 2 | 3 | Set Implicit Arguments. 4 | Unset Strict Implicit. 5 | Unset Transparent Obligations. 6 | Unset Automatic Introduction. 7 | 8 | 9 | Lemma eq_Rep_equiv (P R: PCFPO_rep) : 10 | Equivalence (A:=PCFPO_rep_Hom P R) 11 | (fun a c => forall r, a r == c r). 12 | Proof. 13 | intros P R. 14 | constructor; 15 | simpl in *. 16 | unfold Reflexive; intros. auto. 17 | unfold Symmetric; simpl; intros r s. 18 | intros. 19 | auto. 20 | 21 | unfold Transitive; intros. 22 | transitivity (y r t z0); auto. 23 | Qed. 24 | 25 | Definition eq_Rep_oid (P R : PCFPO_rep) := Build_Setoid (eq_Rep_equiv P R). 26 | 27 | 28 | (** Identity Representation *) 29 | 30 | Lemma Rep_id_struct (P : PCFPO_rep) : 31 | PCFPO_rep_Hom_struct (RMonad_Hom_id P). 32 | Proof. 33 | intro P; 34 | unfold Monad_Hom_id. 35 | simpl. 36 | constructor; 37 | simpl; 38 | intros; 39 | try rewrite <- surjective_pairing; 40 | auto. 41 | Qed. 42 | 43 | Definition Rep_id (P: PCFPO_rep) := Build_PCFPO_rep_Hom (Rep_id_struct P). 44 | 45 | (** Composition of Representations *) 46 | Section Rep_comp. 47 | Variables P Q R: PCFPO_rep. 48 | Variable M: PCFPO_rep_Hom P Q. 49 | Variable N: PCFPO_rep_Hom Q R. 50 | 51 | Program Instance Rep_comp_struct : 52 | PCFPO_rep_Hom_struct (Monad_Hom_comp M N). 53 | Next Obligation. 54 | Proof. 55 | set (HM:=App_Hom (Sig := M)). 56 | set (HN:=App_Hom (Sig := N)). 57 | simpl in *. 58 | rewrite HM. 59 | rewrite HN. 60 | auto. 61 | Qed. 62 | Next Obligation. 63 | set (HM:=Abs_Hom (Sig := M)). 64 | set (HN:=Abs_Hom (Sig := N)). 65 | simpl in *. 66 | 67 | rewrite HM. 68 | rewrite HN. 69 | auto. 70 | Qed. 71 | Next Obligation. 72 | set (HM:=Rec_Hom (Sig := M)). 73 | set (HN:=Rec_Hom (Sig := N)). 74 | simpl in *. 75 | 76 | rewrite HM. 77 | rewrite HN. 78 | auto. 79 | Qed. 80 | Next Obligation. 81 | set (HM:=ttt_Hom (Sig := M)). 82 | set (HN:=ttt_Hom (Sig := N)). 83 | simpl in *. 84 | 85 | rewrite HM. 86 | rewrite HN. 87 | auto. 88 | Qed. 89 | Next Obligation. 90 | set (HM:=fff_Hom (Sig := M)). 91 | set (HN:=fff_Hom (Sig := N)). 92 | simpl in *. 93 | 94 | rewrite HM. 95 | rewrite HN. 96 | auto. 97 | Qed. 98 | Next Obligation. 99 | set (HM:=nats_Hom (Sig := M)). 100 | set (HN:=nats_Hom (Sig := N)). 101 | simpl in *. 102 | 103 | rewrite HM. 104 | rewrite HN. 105 | auto. 106 | Qed. 107 | Next Obligation. 108 | set (HM:=succ_Hom (Sig := M)). 109 | set (HN:=succ_Hom (Sig := N)). 110 | simpl in *. 111 | 112 | rewrite HM. 113 | rewrite HN. 114 | auto. 115 | Qed. 116 | Next Obligation. 117 | set (HM:=zero_Hom (Sig := M)). 118 | set (HN:=zero_Hom (Sig := N)). 119 | simpl in *. 120 | 121 | rewrite HM. 122 | rewrite HN. 123 | auto. 124 | Qed. 125 | Next Obligation. 126 | set (HM:=condN_Hom (Sig := M)). 127 | set (HN:=condN_Hom (Sig := N)). 128 | simpl in *. 129 | 130 | rewrite HM. 131 | rewrite HN. 132 | auto. 133 | Qed. 134 | Next Obligation. 135 | set (HM:=condB_Hom (Sig := M)). 136 | set (HN:=condB_Hom (Sig := N)). 137 | simpl in *. 138 | 139 | rewrite HM. 140 | rewrite HN. 141 | auto. 142 | Qed. 143 | Next Obligation. 144 | set (HM:=Bottom_Hom (Sig := M)). 145 | set (HN:=Bottom_Hom (Sig := N)). 146 | simpl in *. 147 | 148 | rewrite HM. 149 | rewrite HN. 150 | auto. 151 | Qed. 152 | 153 | Definition Rep_comp := Build_PCFPO_rep_Hom Rep_comp_struct. 154 | 155 | End Rep_comp. 156 | 157 | (** category of representations *) 158 | 159 | Program Instance PCFPO_REP_struct : Cat (fun a c => PCFPO_rep_Hom a c) := { 160 | mor_oid a c := eq_Rep_oid a c; 161 | id a := Rep_id a; 162 | comp P Q R f g := Rep_comp f g 163 | }. 164 | Next Obligation. 165 | Proof. 166 | unfold Proper in *; do 2 red. 167 | simpl. 168 | intros y z h y' z' h' e r t. 169 | rewrite h. 170 | rewrite h'. 171 | auto. 172 | Qed. 173 | 174 | Definition PCFPO_REP := Build_Category PCFPO_REP_struct. 175 | 176 | End PCFPO_representation. 177 | 178 | Existing Instance pcfpo_rep_struct. 179 | 180 | -------------------------------------------------------------------------------- /ORDER/tlc.v: -------------------------------------------------------------------------------- 1 | Require Import Coq.Relations.Relations. 2 | 3 | Require Export CatSem.ORDER.tlc_order_rep_eq. 4 | Require Export CatSem.TLC.TLC_RMonad. 5 | 6 | Set Implicit Arguments. 7 | Unset Strict Implicit. 8 | Unset Transparent Obligations. 9 | Unset Automatic Introduction. 10 | 11 | Obligation Tactic := simpl; intros. 12 | 13 | Program Instance TLC_rep : TLCPO_rep_struct TLCB := { 14 | App := tlc_app ; 15 | Abs := tlc_abs 16 | }. 17 | Next Obligation. 18 | Proof. 19 | simpl; intros. 20 | apply clos_refl_trans_1n_contains. 21 | apply relorig. 22 | constructor. 23 | Qed. 24 | 25 | Canonical Structure TLC_Rep : TLCPO_REP := Build_TLCPO_rep TLC_rep. 26 | 27 | 28 | 29 | 30 | 31 | 32 | 33 | 34 | 35 | 36 | -------------------------------------------------------------------------------- /ORDER/tlc_init.v: -------------------------------------------------------------------------------- 1 | Require Export CatSem.ORDER.tlc. 2 | Require Export CatSem.ORDER.tlc_order_rep_eq. 3 | 4 | Set Implicit Arguments. 5 | Unset Strict Implicit. 6 | Unset Transparent Obligations. 7 | Unset Automatic Introduction. 8 | 9 | 10 | 11 | Section initiality. 12 | 13 | Variable R : TLCPO_REP. 14 | 15 | Fixpoint init_d V t (y : TLC_Rep V t) : R V t := 16 | match y in TLC _ t return R V t with 17 | | var _ v => rweta (RMonad_struct := R ) _ _ v 18 | | abs _ _ v => Abs _ _ _ (init_d v) 19 | | app _ _ v1 v2 => App _ _ _ (init_d v1, init_d v2) 20 | end. 21 | 22 | Lemma init_rename V t (y : TLC_Rep V t) W (f : V ---> W) : 23 | init_d (y //- f) = 24 | rlift R f _ (init_d y). 25 | Proof. 26 | intros V t y. 27 | induction y; 28 | simpl; intros. 29 | 30 | assert (H:=rlift_rweta R). 31 | simpl in H; auto. 32 | 33 | unfold rlift. simpl. 34 | rerew (rmhom_rmkl (Abs (TLCPO_rep_struct := R) r s)). 35 | apply f_equal. 36 | rewrite IHy. 37 | unfold rlift. 38 | simpl. 39 | apply (rkl_eq R). 40 | simpl; intros. 41 | assert (H3:=sshift_rweta R). 42 | simpl in *. 43 | destruct z as [t z | ]; 44 | simpl; auto. 45 | unfold rlift; simpl; 46 | rew (retakl R). 47 | 48 | unfold rlift. simpl. 49 | rerew (rmhom_rmkl (App (TLCPO_rep_struct := R) r s)). 50 | apply f_equal. 51 | rewrite IHy1. 52 | rewrite IHy2. 53 | auto. 54 | Qed. 55 | 56 | Hint Rewrite init_rename : fin. 57 | 58 | Lemma init_subst V t (y : TLC_Rep V t) W (f : V ---> TLC W): 59 | init_d (y >>= f) = 60 | rkleisli (RMonad_struct := R) 61 | (SM_ind (fun t v => init_d (f t v))) _ (init_d y). 62 | Proof. 63 | intros V t y. 64 | induction y; 65 | simpl; intros. 66 | rew (retakl R). 67 | 68 | rerew (rmhom_rmkl (Abs (TLCPO_rep_struct := R) r s)). 69 | apply f_equal. 70 | rewrite IHy. 71 | simpl. 72 | apply (rkl_eq R). 73 | simpl; intros t z. 74 | destruct z; fin. 75 | unfold _inj. fin. 76 | rerew (rmhom_rmkl (App (TLCPO_rep_struct := R) r s)). 77 | fin. 78 | Qed. 79 | 80 | Hint Rewrite init_subst : fin. 81 | Hint Resolve init_subst : fin. 82 | 83 | Lemma beta1_init_d V t (v v': TLC V t) 84 | (H : beta1 v v') : 85 | init_d v <<< init_d v'. 86 | Proof. 87 | intros V t v v' H. 88 | induction H. 89 | simpl. 90 | unfold _substar. 91 | rewrite init_subst. 92 | assert (H3:=beta_red (TLCPO_rep_struct:=R) 93 | (r:=s) (s:=t) (V:=V) (init_d M) (init_d N)). 94 | simpl in *. 95 | apply (IRel_eq_r H3). 96 | unfold Rsubstar. 97 | apply (rkl_eq R). 98 | simpl. 99 | intros t0 z. 100 | unfold Rsubst_star_map. 101 | destruct z as [t0 z | ]; 102 | simpl; auto. 103 | Qed. 104 | 105 | Lemma beta_star_init_d V t (v v': TLC V t) 106 | (H : beta_star v v') : 107 | init_d v <<< init_d v'. 108 | Proof. 109 | intros V t v v' H. 110 | induction H. 111 | 112 | apply beta1_init_d. 113 | auto. 114 | 115 | assert (H':=PO_mor_monotone (App (TLCPO_rep_struct:=R) s t V)). 116 | simpl in H'. 117 | apply H'. 118 | constructor. 119 | auto. 120 | reflexivity. 121 | 122 | assert (H':=PO_mor_monotone (App (TLCPO_rep_struct:=R) s t V)). 123 | simpl in H'. 124 | apply H'. 125 | constructor. 126 | reflexivity. 127 | auto. 128 | 129 | assert (H':=PO_mor_monotone (Abs (TLCPO_rep_struct:=R) s t V)). 130 | simpl in H'. 131 | apply H'. 132 | auto. 133 | Qed. 134 | 135 | Obligation Tactic := idtac. 136 | 137 | Program Instance init_s V : ipo_mor_struct (init_d (V:=V)). 138 | Next Obligation. 139 | Proof. 140 | unfold Proper; 141 | red. 142 | intros V t x y H. 143 | induction H; 144 | simpl. 145 | apply IRel_refl. 146 | 147 | transitivity (init_d y); auto. 148 | 149 | apply beta_star_init_d. 150 | auto. 151 | Qed. 152 | 153 | Definition init V := Build_ipo_mor (init_s V). 154 | 155 | Obligation Tactic := fin; try apply (rkl_eq R); fin. 156 | 157 | Program Instance init_mon_s : 158 | RMonad_Hom_struct (P:=TLCB) (Q:=R) init. 159 | 160 | Definition initM := Build_RMonad_Hom init_mon_s. 161 | 162 | Program Instance init_rep_s : TLCPO_rep_Hom_struct initM. 163 | 164 | Definition initR := Build_TLCPO_rep_Hom init_rep_s. 165 | 166 | Section unicity. 167 | 168 | Variable f : TLC_Rep ---> R. 169 | 170 | Lemma unique : f == initR. 171 | Proof. 172 | intros V t z; 173 | induction z; 174 | simpl; 175 | try rew (rmon_hom_rweta f); 176 | try rew (Abs_Hom (TLCPO_rep_Hom_struct := f)); 177 | try match goal with [H:TLC _ (?a ~> _), H1:TLC _ ?a |- _ ] => 178 | rew (App_Hom (TLCPO_rep_Hom_struct := f) (H,H1)) end; 179 | fin. 180 | Qed. 181 | 182 | End unicity. 183 | 184 | End initiality. 185 | 186 | Program Instance TLCPO_init : Initial TLCPO_REP := { 187 | Init := TLC_Rep ; 188 | InitMor := initR ; 189 | InitMorUnique := unique 190 | }. 191 | 192 | Print Assumptions TLCPO_init. 193 | 194 | 195 | 196 | 197 | 198 | -------------------------------------------------------------------------------- /ORDER/tlc_order_rep.v: -------------------------------------------------------------------------------- 1 | Require Export CatSem.CAT.rmodule_TYPE. 2 | Require Export CatSem.TLC.TLC_types. 3 | 4 | Set Implicit Arguments. 5 | Unset Strict Implicit. 6 | Unset Transparent Obligations. 7 | Unset Automatic Introduction. 8 | 9 | Section TLC_rep. 10 | 11 | (** a lot of notation to make things readable *) 12 | 13 | (*Notation "'TY'" := PCF.types.*) 14 | 15 | Notation "'TY'" := TLCtypes. 16 | 17 | Notation "'IP'" := (IPO TY). 18 | Notation "a '~>' b" := (TLCarrow a b) 19 | (at level 69, right associativity). 20 | Notation "a 'x' b" := (product a b) (at level 30). 21 | Notation "M [ z ]" := (FIB_RMOD _ z M) (at level 35). 22 | Notation "'d' M // s" := (DER_RMOD _ _ s M) (at level 25). 23 | Notation "'*'" := (Term (C:=RMOD _ Ord)). 24 | 25 | Notation "f 'X' g" := (product_mor _ f g)(at level 30). 26 | 27 | Notation "'FM'" := (#(FIB_RMOD _ _ )). 28 | Notation "'FPB'":= (FIB_RPB _ _ _ ). 29 | Notation "'PRPB'":= (PROD_RPB _ _ _ _ ). 30 | Notation "'PBM'":= (#(PbRMOD _ _ )). 31 | Notation "'DM'":= (#(DER_RMOD _ _ _ )). 32 | Notation "'DPB'":= (DER_RPB _ _ _ ). 33 | 34 | (* this might be changed for the module-morphism-version, 35 | if we decide to use it *) 36 | Notation "y [* := z ]":= (Rsubstar z _ y)(at level 55). 37 | 38 | (** a monad is a representation if it is accompagnied by 39 | - a lot of morphisms of modules 40 | - beta-reduction 41 | *) 42 | 43 | Class TLCPO_rep_struct (P : RMonad (IDelta TY)) := { 44 | App : forall r s, (P[r~>s]) x (P[r]) ---> P[s]; 45 | Abs : forall r s, (d P // r)[s] ---> P[r ~> s]; 46 | 47 | beta_red : forall r s V y z, 48 | App r s V (Abs r s V y, z) << y[*:= z] 49 | (*; 50 | these are already verified since App and Abs 51 | are module morphisms in (RMOD P PO) 52 | propag_App1: forall r s V y y' z, 53 | y << y' -> App r s V (y,z) << App _ _ _ (y',z) ; 54 | 55 | propag_App2: forall r s V y z z', 56 | z << z' -> App r s V (y,z) << App _ _ _ (y,z') ; 57 | 58 | propag_Abs: forall r s V y y', 59 | y << y' -> Abs r s V y << Abs _ _ _ y' 60 | *) 61 | }. 62 | 63 | 64 | Record TLCPO_rep := { 65 | tlcpo_rep_monad :> RMonad (IDelta TY) ; 66 | tlcpo_rep_struct :> TLCPO_rep_struct tlcpo_rep_monad 67 | }. 68 | 69 | Existing Instance tlcpo_rep_struct. 70 | 71 | (** morphisms of representations *) 72 | 73 | Section TLCPO_rep_Hom. 74 | 75 | Variables P R : TLCPO_rep. 76 | 77 | Section Rep_Hom_Class. 78 | 79 | Variable S : RMonad_Hom P R. 80 | 81 | Notation "'SS'":= (PbRMod_ind_Hom S). 82 | 83 | 84 | Class TLCPO_rep_Hom_struct := { 85 | 86 | App_Hom: forall r s, 87 | App r s ;; FM SS ;; FPB == 88 | (FM SS ;; FPB) X (FM SS ;; FPB);; 89 | PRPB ;; PBM (App r s) 90 | ; 91 | 92 | Abs_Hom: forall r s, 93 | Abs r s ;; FM SS ;; FPB == 94 | FM (DM SS ;; DPB) ;; FPB ;; PBM (Abs r s) 95 | 96 | }. 97 | 98 | End Rep_Hom_Class. 99 | 100 | (** the type of morphismes of representations P -> R *) 101 | 102 | Record TLCPO_rep_Hom := { 103 | rep_Hom_monad :> RMonad_Hom P R ; 104 | rep_Hom_monad_struct :> TLCPO_rep_Hom_struct rep_Hom_monad 105 | }. 106 | 107 | End TLCPO_rep_Hom. 108 | 109 | End TLC_rep. 110 | Existing Instance rep_Hom_monad_struct. 111 | Existing Instance tlcpo_rep_struct. 112 | 113 | -------------------------------------------------------------------------------- /ORDER/tlc_order_rep_eq.v: -------------------------------------------------------------------------------- 1 | Require Export CatSem.ORDER.tlc_order_rep. 2 | 3 | Set Implicit Arguments. 4 | Unset Strict Implicit. 5 | Unset Transparent Obligations. 6 | Unset Automatic Introduction. 7 | 8 | (** we define two morphisms of representations to be equal 9 | if they agree on each set of variables *) 10 | 11 | Lemma eq_Rep_equiv (P R: TLCPO_rep) : 12 | Equivalence (A:=TLCPO_rep_Hom P R) 13 | (fun a c => forall r, a r == c r). 14 | Proof. 15 | intros; 16 | constructor; 17 | unfold Reflexive, 18 | Symmetric, 19 | Transitive; 20 | simpl; intros; auto; 21 | etransitivity; eauto. 22 | Qed. 23 | 24 | Definition eq_Rep_oid (P R : TLCPO_rep) := Build_Setoid (eq_Rep_equiv P R). 25 | 26 | (** Identity Representation *) 27 | 28 | Program Instance Rep_id_struct (P : TLCPO_rep) : 29 | TLCPO_rep_Hom_struct (RMonad_id P). 30 | Next Obligation. 31 | match goal with [H:prod _ _ |- _] => 32 | destruct H end; 33 | auto. 34 | Qed. 35 | 36 | Definition Rep_id (P: TLCPO_rep) := Build_TLCPO_rep_Hom (Rep_id_struct P). 37 | 38 | (** Composition of Representations *) 39 | Section Rep_comp. 40 | 41 | Variables P Q R: TLCPO_rep. 42 | Variable M: TLCPO_rep_Hom P Q. 43 | Variable N: TLCPO_rep_Hom Q R. 44 | 45 | Obligation Tactic := cat; 46 | repeat rew (App_Hom (S:=M)); 47 | repeat rew (App_Hom (S:=N)); 48 | repeat rew (Abs_Hom (S:=M)); 49 | repeat rew (Abs_Hom (S:=N)); 50 | cat. 51 | 52 | Program Instance Rep_comp_struct : 53 | TLCPO_rep_Hom_struct (RMonad_comp M N). 54 | 55 | Definition Rep_comp := Build_TLCPO_rep_Hom Rep_comp_struct. 56 | 57 | End Rep_comp. 58 | 59 | (** category of representations *) 60 | 61 | Ltac t := match goal with [H : forall _ _ _ , _ = _ |- _ ] => 62 | rewrite H end. 63 | 64 | Obligation Tactic := cat; try unf_Proper; 65 | cat; repeat t; cat. 66 | 67 | Program Instance TLCPO_REP_struct : Cat_struct (TLCPO_rep_Hom) := { 68 | mor_oid a c := eq_Rep_oid a c; 69 | id a := Rep_id a; 70 | comp P Q R f g := Rep_comp f g 71 | }. 72 | 73 | Definition TLCPO_REP := Build_Cat TLCPO_REP_struct. 74 | 75 | Existing Instance tlcpo_rep_struct. 76 | 77 | -------------------------------------------------------------------------------- /ORDER/tlc_properties.v: -------------------------------------------------------------------------------- 1 | Require Export CatSem.ORDER.tlc. 2 | 3 | 4 | Set Implicit Arguments. 5 | Unset Strict Implicit. 6 | Unset Transparent Obligations. 7 | Unset Automatic Introduction. 8 | 9 | 10 | Lemma subst_monotone_on_functions (V : IT) 11 | (t : TY)(y : TLCB V t) W (f g : V ---> TLCB W) 12 | (H : forall t v, f t v <<< g t v) : 13 | beta (y >>= f) (y >>= g). 14 | Proof. 15 | intros V t y. 16 | induction y; 17 | simpl; 18 | intros. 19 | apply H. 20 | apply cp_Lam. 21 | apply IHy. 22 | induction v; 23 | simpl. 24 | unfold _inj. 25 | apply beta_rename. 26 | auto. 27 | reflexivity. 28 | transitivity ( 29 | app (y1 >>= f) (y2 >>= g)). 30 | apply cp_App2. 31 | apply IHy2. 32 | auto. 33 | apply cp_App1. 34 | apply IHy1. 35 | auto. 36 | Qed. 37 | 38 | 39 | 40 | 41 | 42 | 43 | 44 | 45 | -------------------------------------------------------------------------------- /ORDER/ulc.v: -------------------------------------------------------------------------------- 1 | Require Import Coq.Relations.Relations. 2 | 3 | Require Export CatSem.ORDER.ulc_order_rep_eq. 4 | Require Export CatSem.ULC.ULC_RMonad. 5 | 6 | Set Implicit Arguments. 7 | Unset Strict Implicit. 8 | Unset Transparent Obligations. 9 | Unset Automatic Introduction. 10 | 11 | Obligation Tactic := simpl; intros. 12 | 13 | Program Instance ULC_rep : ULCPO_rep_struct ULCBETA := { 14 | app := ulc_app ; 15 | abs := ulc_abs 16 | }. 17 | Next Obligation. 18 | Proof. 19 | simpl; intros. 20 | apply clos_refl_trans_1n_contains. 21 | apply relorig. 22 | constructor. 23 | Qed. 24 | 25 | Canonical Structure ULC_Rep : ULCPO_REP := Build_ULCPO_rep ULC_rep. 26 | 27 | 28 | 29 | 30 | 31 | 32 | 33 | 34 | 35 | 36 | -------------------------------------------------------------------------------- /ORDER/ulc_init.v: -------------------------------------------------------------------------------- 1 | Require Export CatSem.ORDER.ulc. 2 | Require Export CatSem.ORDER.ulc_order_rep_eq. 3 | Require Export CatSem.CAT.orders_lemmas. 4 | 5 | Set Implicit Arguments. 6 | Unset Strict Implicit. 7 | Unset Transparent Obligations. 8 | Unset Automatic Introduction. 9 | 10 | 11 | 12 | Section initiality. 13 | 14 | Variable R : ULCPO_REP. 15 | 16 | Fixpoint init_d V (y : ULC_Rep V) : R V := 17 | match y in ULC _ return R V with 18 | | Var v => rweta (RMonad_struct := R ) _ v 19 | | Abs v => abs _ (init_d v) 20 | | App v1 v2 => app _ (init_d v1, init_d v2) 21 | end. 22 | 23 | Lemma init_rename V (y : ULC_Rep V) W (f : V ---> W) : 24 | init_d (y //- f) = 25 | rlift R f (init_d y). 26 | Proof. 27 | intros V y. 28 | induction y; 29 | simpl; intros. 30 | 31 | assert (H:=rlift_rweta R). 32 | simpl in H; auto. 33 | 34 | unfold rlift. simpl. 35 | rerew (rmhom_rmkl (abs (ULCPO_rep_struct := R))). 36 | apply f_equal. 37 | rewrite IHy. 38 | unfold rlift. 39 | simpl. 40 | apply (rkl_eq R). 41 | simpl; intros. 42 | assert (H3:=shift_not_rweta R). 43 | simpl in *. 44 | destruct x as [t z | ]; 45 | simpl; auto. 46 | unfold rlift; simpl; 47 | rew (retakl R). 48 | 49 | unfold rlift. simpl. 50 | rerew (rmhom_rmkl (app (ULCPO_rep_struct := R))). 51 | apply f_equal. 52 | rewrite IHy1. 53 | rewrite IHy2. 54 | auto. 55 | Qed. 56 | 57 | Hint Rewrite init_rename : fin. 58 | 59 | Lemma init_subst V (y : ULC_Rep V) W (f : V ---> ULC W): 60 | init_d (y >>= f) = 61 | rkleisli (RMonad_struct := R) 62 | (Sm_ind (fun v => init_d (f v))) (init_d y). 63 | Proof. 64 | induction y; 65 | simpl; intros. 66 | rew (retakl R). 67 | 68 | rerew (rmhom_rmkl (abs (ULCPO_rep_struct := R))). 69 | apply f_equal. 70 | rewrite IHy. 71 | simpl. 72 | apply (rkl_eq R). 73 | simpl; intros z. 74 | destruct z; fin. 75 | unfold inj. fin. 76 | rerew (rmhom_rmkl (app (ULCPO_rep_struct := R))). 77 | fin. 78 | Qed. 79 | 80 | Hint Rewrite init_subst : fin. 81 | Hint Resolve init_subst : fin. 82 | 83 | Lemma beta_init_d V (v v': ULC V) 84 | (H : beta v v') : 85 | init_d v << init_d v'. 86 | Proof. 87 | intros V v v' H. 88 | induction H. 89 | simpl. 90 | unfold substar. 91 | rewrite init_subst. 92 | assert (H3:=beta_red (ULCPO_rep_struct:=R) 93 | (V:=V) (init_d M) (init_d N)). 94 | simpl in *. 95 | apply (Rel_eq_r H3). 96 | unfold Rsubstar. 97 | apply (rkl_eq R). 98 | simpl. 99 | intros z. 100 | unfold Rsubst_star_map. 101 | destruct z as [t0 z | ]; 102 | simpl; auto. 103 | Qed. 104 | 105 | Lemma beta_star_init_d V (v v': ULC V) 106 | (H : beta_star v v') : 107 | init_d v << init_d v'. 108 | Proof. 109 | intros V v v' H. 110 | induction H. 111 | 112 | apply beta_init_d. 113 | auto. 114 | 115 | assert (H':=PO_mor_monotone (app (ULCPO_rep_struct:=R) V)). 116 | simpl in H'. 117 | apply H'. 118 | constructor. 119 | auto. 120 | reflexivity. 121 | 122 | assert (H':=PO_mor_monotone (app (ULCPO_rep_struct:=R) V)). 123 | simpl in H'. 124 | apply H'. 125 | constructor. 126 | reflexivity. 127 | auto. 128 | 129 | assert (H':=PO_mor_monotone (abs (ULCPO_rep_struct:=R) V)). 130 | simpl in H'. 131 | apply H'. 132 | auto. 133 | Qed. 134 | 135 | Obligation Tactic := idtac. 136 | 137 | Program Instance init_s V : PO_mor_struct (init_d (V:=V)). 138 | Next Obligation. 139 | Proof. 140 | unfold Proper; 141 | red. 142 | intros V x y H. 143 | induction H; 144 | simpl. 145 | apply Rel_refl. 146 | 147 | transitivity (init_d y); auto. 148 | 149 | apply beta_star_init_d. 150 | auto. 151 | Qed. 152 | 153 | Definition init V := Build_PO_mor (init_s V). 154 | 155 | Obligation Tactic := fin; try apply (rkl_eq R); fin. 156 | 157 | Program Instance init_mon_s : 158 | RMonad_Hom_struct (P:=ULCBETA) (Q:=R) init. 159 | 160 | Definition initM := Build_RMonad_Hom init_mon_s. 161 | 162 | Program Instance init_rep_s : ULCPO_rep_Hom_struct initM. 163 | 164 | Definition initR := Build_ULCPO_rep_Hom init_rep_s. 165 | 166 | Section unicity. 167 | 168 | Variable f : ULC_Rep ---> R. 169 | 170 | Lemma unique : f == initR. 171 | Proof. 172 | intros V z; 173 | induction z; 174 | simpl; 175 | try rew (rmon_hom_rweta f); 176 | try rew (abs_Hom (ULCPO_rep_Hom_struct := f)); 177 | try match goal with [H:ULC _ , H1:ULC _ |- _ ] => 178 | rew (app_Hom (ULCPO_rep_Hom_struct := f) (H,H1)) end; 179 | fin. 180 | Qed. 181 | 182 | End unicity. 183 | 184 | End initiality. 185 | 186 | Program Instance ULCPO_init : Initial ULCPO_REP := { 187 | Init := ULC_Rep ; 188 | InitMor := initR ; 189 | InitMorUnique := unique 190 | }. 191 | 192 | Print Assumptions ULCPO_init. 193 | 194 | 195 | 196 | 197 | 198 | -------------------------------------------------------------------------------- /ORDER/ulc_order_rep.v: -------------------------------------------------------------------------------- 1 | Require Export CatSem.CAT.rmodule_TYPE. 2 | 3 | Set Implicit Arguments. 4 | Unset Strict Implicit. 5 | Unset Transparent Obligations. 6 | Unset Automatic Introduction. 7 | 8 | Section ULC_rep. 9 | 10 | (** a lot of notation to make things readable *) 11 | 12 | Notation "a 'x' b" := (product (C:=RMOD _ _) a b) (at level 30). 13 | Notation "'d' M" := (DER_RMOD_not _ _ M) (at level 25). 14 | Notation "'*'" := (Term (C:=RMOD _ Ord)). 15 | 16 | Notation "f 'X' g" := (product_mor _ f g)(at level 30). 17 | 18 | Notation "'PRPB'":= (PROD_RPB _ _ _ _ ). 19 | Notation "'PBM'":= (#(PbRMOD _ _ )). 20 | Notation "'DM'":= (#(DER_RMOD_not _ _ )). 21 | Notation "'DPB'":= (DER_RPB_not _ _ ). 22 | 23 | (* this might be changed for the module-morphism-version, 24 | if we decide to use it *) 25 | Notation "y [* := z ]":= (Rsubstar_not z y)(at level 55). 26 | 27 | (** a monad is a representation if it is accompagnied by 28 | - a lot of morphisms of modules 29 | - beta-reduction 30 | *) 31 | 32 | Class ULCPO_rep_struct (P : RMonad SM_po) := { 33 | app : P x P ---> P; 34 | abs : d P ---> P; 35 | beta_red : forall V y z, 36 | app V (abs V y, z) << y[*:= z] 37 | }. 38 | 39 | 40 | Record ULCPO_rep := { 41 | ulcpo_rep_monad :> RMonad SM_po ; 42 | ulcpo_rep_struct :> ULCPO_rep_struct ulcpo_rep_monad 43 | }. 44 | 45 | Existing Instance ulcpo_rep_struct. 46 | 47 | (** morphisms of representations *) 48 | 49 | Section ULCPO_rep_Hom. 50 | 51 | Variables P R : ULCPO_rep. 52 | 53 | Section Rep_Hom_Class. 54 | 55 | Variable S : RMonad_Hom P R. 56 | 57 | Notation "'ST'":= (PbRMod_ind_Hom S). 58 | 59 | 60 | Class ULCPO_rep_Hom_struct := { 61 | 62 | app_Hom: 63 | app ;; ST == 64 | ST X ST ;; PRPB ;; PBM (app) 65 | ; 66 | 67 | abs_Hom: 68 | abs ;; ST == 69 | DM ST ;; DPB ;; PBM (abs) 70 | 71 | }. 72 | 73 | End Rep_Hom_Class. 74 | 75 | (** the type of morphismes of representations P -> R *) 76 | 77 | Record ULCPO_rep_Hom := { 78 | rep_Hom_monad :> RMonad_Hom P R ; 79 | rep_Hom_monad_struct :> ULCPO_rep_Hom_struct rep_Hom_monad 80 | }. 81 | 82 | End ULCPO_rep_Hom. 83 | 84 | End ULC_rep. 85 | Existing Instance rep_Hom_monad_struct. 86 | Existing Instance ulcpo_rep_struct. 87 | 88 | -------------------------------------------------------------------------------- /ORDER/ulc_order_rep_eq.v: -------------------------------------------------------------------------------- 1 | Require Export CatSem.ORDER.ulc_order_rep. 2 | 3 | Set Implicit Arguments. 4 | Unset Strict Implicit. 5 | Unset Transparent Obligations. 6 | Unset Automatic Introduction. 7 | 8 | (** we define two morphisms of representations to be equal 9 | if they agree on each set of variables *) 10 | 11 | Lemma eq_Rep_equiv (P R: ULCPO_rep) : 12 | Equivalence (A:=ULCPO_rep_Hom P R) 13 | (fun a c => forall r, a r == c r). 14 | Proof. 15 | intros; 16 | constructor; 17 | unfold Reflexive, 18 | Symmetric, 19 | Transitive; 20 | simpl; intros; auto; 21 | etransitivity; eauto. 22 | Qed. 23 | 24 | Definition eq_Rep_oid (P R : ULCPO_rep) := Build_Setoid (eq_Rep_equiv P R). 25 | 26 | (** Identity Representation *) 27 | 28 | Program Instance Rep_id_struct (P : ULCPO_rep) : 29 | ULCPO_rep_Hom_struct (RMonad_id P). 30 | Next Obligation. 31 | match goal with [H:prod _ _ |- _] => 32 | destruct H end; 33 | auto. 34 | Qed. 35 | 36 | Definition Rep_id (P: ULCPO_rep) := Build_ULCPO_rep_Hom (Rep_id_struct P). 37 | 38 | (** Composition of Representations *) 39 | Section Rep_comp. 40 | 41 | Variables P Q R: ULCPO_rep. 42 | Variable M: ULCPO_rep_Hom P Q. 43 | Variable N: ULCPO_rep_Hom Q R. 44 | 45 | Obligation Tactic := cat; 46 | repeat rew (app_Hom (S:=M)); 47 | repeat rew (app_Hom (S:=N)); 48 | repeat rew (abs_Hom (S:=M)); 49 | repeat rew (abs_Hom (S:=N)); 50 | cat. 51 | 52 | Program Instance Rep_comp_struct : 53 | ULCPO_rep_Hom_struct (RMonad_comp M N). 54 | 55 | Definition Rep_comp := Build_ULCPO_rep_Hom Rep_comp_struct. 56 | 57 | End Rep_comp. 58 | 59 | (** category of representations *) 60 | 61 | Ltac t := match goal with [H : _ |- _ ] => 62 | rewrite H end. 63 | 64 | Obligation Tactic := cat; try unf_Proper; 65 | cat; repeat t; cat. 66 | 67 | Program Instance ULCPO_REP_struct : Cat_struct (ULCPO_rep_Hom) := { 68 | mor_oid a c := eq_Rep_oid a c; 69 | id a := Rep_id a; 70 | comp P Q R f g := Rep_comp f g 71 | }. 72 | 73 | Definition ULCPO_REP := Build_Cat ULCPO_REP_struct. 74 | 75 | Existing Instance ulcpo_rep_struct. 76 | 77 | -------------------------------------------------------------------------------- /PCF/PCF_types.v: -------------------------------------------------------------------------------- 1 | 2 | Set Implicit Arguments. 3 | Unset Strict Implicit. 4 | 5 | Unset Automatic Introduction. 6 | 7 | Module PCF. 8 | 9 | Inductive Sorts := 10 | | Nat : Sorts 11 | | Bool : Sorts 12 | | Arrow : Sorts -> Sorts -> Sorts. 13 | 14 | Notation "a '~>' b" := (Arrow a b) (at level 60, right associativity). 15 | 16 | End PCF. -------------------------------------------------------------------------------- /PCF/pcf_rich_monad.v: -------------------------------------------------------------------------------- 1 | Require Export rich_monads pcfpo_mod. 2 | 3 | Set Implicit Arguments. 4 | Unset Strict Implicit. 5 | 6 | Unset Automatic Introduction. 7 | 8 | Section pcf_is_rich. 9 | 10 | Program Instance PCF_enriched : Enriched_Monad (T:=PCF.TY) (P:=PCF_monad) 11 | (fun V => @PCF.betaS V). 12 | Next Obligation. 13 | Proof. 14 | assert (H:= PCF.BETA_struct X). 15 | assert (H':=Clos_RT_1n_prf). 16 | destruct H. 17 | simpl. 18 | constructor; unfold itrel. 19 | unfold PCF.betaS. 20 | apply H'. 21 | apply H'. 22 | Qed. 23 | Next Obligation. 24 | Proof. 25 | assert (H':=PCF.subst_h (f:=f)). 26 | unfold ITproper in *. 27 | unfold itrel. 28 | apply H'. 29 | constructor. 30 | auto. 31 | Qed. 32 | Next Obligation. 33 | Proof. 34 | unfold ITproper. 35 | unfold Proper; repeat red. 36 | intros t x y H. 37 | constructor 2 with (PCF.Var y); 38 | try constructor. 39 | constructor. auto. 40 | Qed. 41 | End pcf_is_rich. 42 | 43 | 44 | 45 | 46 | 47 | 48 | 49 | 50 | 51 | 52 | 53 | 54 | 55 | 56 | 57 | 58 | 59 | 60 | 61 | 62 | 63 | 64 | -------------------------------------------------------------------------------- /PCF/pcfpo_exp.v: -------------------------------------------------------------------------------- 1 | Require Export pcfpo_mod monad_h_module. 2 | 3 | Set Implicit Arguments. 4 | Unset Strict Implicit. 5 | 6 | Unset Automatic Introduction. 7 | 8 | (** a PCF exponential monad is a monad together with, 9 | for each u, an isomorphism of modules 10 | [M -> M*u] 11 | *) 12 | 13 | (* TODO: replace M in PCF_abs and PCF_app1 by (M_u), the u-Module of M *) 14 | 15 | Class PCFMonad_struct (M : Monad (IPO PCF.TY)) := { 16 | PCF_abs : forall (u:PCF.TY), 17 | Module_Hom (IPO_Der_Mod M u) M; 18 | PCF_app1 : forall (u:PCF.TY), 19 | Module_Hom M (IPO_Der_Mod M u); 20 | PCF_eta : forall u c, 21 | PCF_abs u c ;; PCF_app1 u c == id _ ; 22 | PCF_beta: forall u c, 23 | PCF_app1 u c ;; PCF_abs u c == id _ 24 | }. 25 | 26 | Record PCFMonad := { 27 | pcfmonad :> Monad (IPO PCF.TY); 28 | pcfmonad_struct :> PCFMonad_struct pcfmonad 29 | }. 30 | 31 | Existing Instance pcfmonad_struct. 32 | 33 | (** a morphism of PCF exp monads is a monad morphism which is compat with 34 | - app1 35 | - abs 36 | *) 37 | 38 | Section PCFMonad_Hom. 39 | 40 | Variables P Q: PCFMonad. 41 | 42 | Class PCFMonad_Hom_struct (h: Monad_Hom P Q):= { 43 | expmonad_hom_abs: forall u c, 44 | PCF_abs u _ ;; h _ == h _ ;; PCF_abs u c; 45 | expmonad_hom_app: forall u c, PCF_app1 u _ ;; h _ == 46 | h c ;; PCF_app1 u _ 47 | }. 48 | 49 | Record PCFMonad_Hom := { 50 | pcfmonad_hom :> Monad_Hom P Q; 51 | pcfmonad_hom_struct :> PCFMonad_Hom_struct pcfmonad_hom 52 | }. 53 | 54 | End PCFMonad_Hom. 55 | 56 | Existing Instance pcfmonad_hom_struct. 57 | 58 | 59 | 60 | -------------------------------------------------------------------------------- /PCF_o_c/RPCF_INIT.v: -------------------------------------------------------------------------------- 1 | Require Import Coq.Logic.FunctionalExtensionality. 2 | Require Import Coq.Logic.Eqdep. 3 | 4 | Require Export CatSem.PCF_o_c.RPCF_rep_cat. 5 | Require Export CatSem.PCF_o_c.RPCF_syntax_init. 6 | 7 | Set Implicit Arguments. 8 | Unset Strict Implicit. 9 | Unset Transparent Obligations. 10 | Unset Automatic Introduction. 11 | 12 | 13 | Section unique. 14 | 15 | Variable R : REP. 16 | 17 | Variable g : PCFE_rep ---> R. 18 | 19 | Lemma initR_unique : g == initR R. 20 | Proof. 21 | simpl. 22 | assert (H : forall t, tcomp (initR R) t = tcomp g t). 23 | assert (H':= ttriag g). 24 | simpl in *. 25 | auto. 26 | apply (eq_rep (H:=H)). 27 | simpl. 28 | intros V t y. 29 | destruct y as [t y]. 30 | simpl. 31 | destruct g as [f Hfarrow Hfcomm T Ts]; 32 | simpl in *. 33 | clear g. 34 | 35 | assert (H' : type_mor R = f) by 36 | (apply functional_extensionality; 37 | intros; auto). 38 | generalize (H t). 39 | generalize H. 40 | generalize dependent T. 41 | clear H. 42 | generalize dependent Hfarrow. 43 | generalize dependent Hfcomm. 44 | rewrite <- H'. 45 | intros. 46 | rewrite (UIP_refl _ _ e). 47 | simpl. 48 | 49 | assert (Hl := rlift_transp_id (Q:=R) H). 50 | simpl in *. 51 | rewrite Hl. 52 | clear Hl. 53 | clear e H. 54 | 55 | induction y. 56 | simpl. 57 | 58 | assert (Hb:=bottom_hom (PCFPO_rep_Hom_struct := Ts) 59 | t (c:=V) tt). 60 | simpl in *. 61 | auto. 62 | 63 | destruct c. 64 | 65 | assert (Hb:=nats_hom n (PCFPO_rep_Hom_struct := 66 | Ts) (c:=V) tt). 67 | simpl in *. 68 | rewrite (UIP_refl _ _ (Hfcomm _)) in Hb. 69 | simpl in * ; auto. 70 | 71 | assert (Hb:=ttt_hom (PCFPO_rep_Hom_struct := 72 | Ts) (c:=V) tt). 73 | simpl in *. 74 | rewrite (UIP_refl _ _ (Hfcomm _)) in Hb. 75 | simpl in * ; auto. 76 | 77 | assert (Hb:=fff_hom (PCFPO_rep_Hom_struct := 78 | Ts) (c:=V) tt). 79 | simpl in *. 80 | rewrite (UIP_refl _ _ (Hfcomm _)) in Hb. 81 | simpl in * ; auto. 82 | 83 | assert (Hb:=Succ_hom (PCFPO_rep_Hom_struct := 84 | Ts) (c:=V) tt). 85 | simpl in *. 86 | rewrite <- Hb. 87 | clear Hb. 88 | symmetry. 89 | apply double_eq_rect. 90 | 91 | assert (Hb:=Pred_hom (PCFPO_rep_Hom_struct := 92 | Ts) (c:=V) tt). 93 | simpl in *. 94 | rewrite <- Hb. 95 | symmetry. 96 | apply double_eq_rect. 97 | 98 | assert (Hb:=Zero_hom (PCFPO_rep_Hom_struct := 99 | Ts) (c:=V) tt). 100 | simpl in *. 101 | rewrite <- Hb. 102 | clear Hb. 103 | symmetry. 104 | apply double_eq_rect. 105 | 106 | assert (Hb:=CondN_hom (PCFPO_rep_Hom_struct := 107 | Ts) (c:=V) tt). 108 | simpl in *. 109 | rewrite <- Hb. 110 | clear Hb. 111 | symmetry. 112 | apply double_eq_rect. 113 | 114 | assert (Hb:=CondB_hom (PCFPO_rep_Hom_struct := 115 | Ts) (c:=V) tt). 116 | simpl in *. 117 | rewrite <- Hb. 118 | clear Hb. 119 | symmetry. 120 | apply double_eq_rect. 121 | 122 | (*1*) 123 | assert (Hw:=gen_rmonad_hom_rweta 124 | ((*gen_RMonad_Hom_struct := *)T)). 125 | simpl in Hw. 126 | assert (Hw' := Hw _ _ (ctype _ v)). 127 | simpl in Hw'. 128 | auto. 129 | (*2*) 130 | simpl. 131 | rewrite <- IHy1. 132 | rewrite <- IHy2. 133 | assert (Happ:=app_hom (PCFPO_rep_Hom_struct := 134 | Ts) (u:=s) (v:=t) (y1,y2)). 135 | simpl in *. 136 | rewrite <- Happ. 137 | clear Happ. 138 | apply f_equal. 139 | simpl. 140 | apply injective_projections; 141 | simpl; 142 | auto. 143 | rewrite (UIP _ _ _ 144 | (type_arrow_dist R s t)(Hfarrow s t)). 145 | auto. 146 | 147 | (*3*) 148 | simpl. 149 | rewrite <- IHy. 150 | clear IHy. 151 | assert (Habs:=abs_hom2 (PCFPO_rep_Hom_struct := 152 | Ts) (u:=t) (v:=s) y ). 153 | simpl in *. 154 | rewrite Habs. 155 | rewrite (UIP _ _ _ (eq_sym (Hfarrow t s)) 156 | (eq_sym (type_arrow_dist R t s))). 157 | reflexivity. 158 | 159 | simpl. 160 | rewrite <- IHy. 161 | assert (Happ:=rec_hom (PCFPO_rep_Hom_struct := 162 | Ts) (t:=t) y). 163 | simpl in *. 164 | rewrite Happ. 165 | rewrite (UIP _ _ _ (Hfarrow t t) 166 | (type_arrow_dist R t t)). 167 | reflexivity. 168 | Qed. 169 | 170 | End unique. 171 | 172 | Hint Resolve initR_unique : fin. 173 | 174 | Obligation Tactic := fin. 175 | 176 | Program Instance PCF_initial : Initial REP := { 177 | Init := PCFE_rep ; 178 | InitMor R := initR R ; 179 | InitMorUnique R := @initR_unique R 180 | }. 181 | 182 | Print Assumptions PCF_initial. 183 | -------------------------------------------------------------------------------- /PCF_o_c/RPCF_rep.v: -------------------------------------------------------------------------------- 1 | Require Export CatSem.CAT.retype_functor_po. 2 | Require Export CatSem.CAT.rmonad_gen_type_po. 3 | Require Export CatSem.PCF.PCF_types. 4 | 5 | Set Implicit Arguments. 6 | Unset Strict Implicit. 7 | Unset Transparent Obligations. 8 | Unset Automatic Introduction. 9 | 10 | Notation "'TY'" := PCF.types. 11 | Notation "'Bool'" := PCF.Bool. 12 | Notation "'Nat'" := PCF.Nat. 13 | 14 | Notation "'IP'" := (IPO TY). 15 | Notation "a '~>' b" := (PCF.arrow a b) 16 | (at level 69, right associativity). 17 | Notation "a 'x' b" := (product a b) (at level 30). 18 | Notation "M [ z ]" := (FIB_RMOD _ z M) (at level 35). 19 | Notation "'d' M // s" := (DER_RMOD _ _ s M) (at level 25). 20 | Notation "'*'" := (Term (C:=RMOD _ PO)). 21 | 22 | Notation "f 'X' g" := (product_mor _ f g)(at level 30). 23 | 24 | Notation "'FM'" := (#(FIB_RMOD _ _ )). 25 | Notation "'FPB'":= (FIB_RPB _ _ _ ). 26 | Notation "'PRPB'":= (PROD_RPB _ _ _ _ ). 27 | (*Notation "'PBF'":= (PB_FIB _ _ _ ).*) 28 | Notation "'PBM'":= (#(PbRMOD _ _ )). 29 | Notation "'DM'":= (#(DER_RMOD _ _ _ )). 30 | Notation "'DPB'":= (DER_RPB _ _ _ ). 31 | 32 | Notation "y [* := z ]":= (Rsubstar z _ y)(at level 55). 33 | 34 | 35 | 36 | Section PCF_rep. 37 | 38 | Variable U : Type. 39 | Variable P : RMonad (SM_ipo U). 40 | Variable f : TY -> U. 41 | 42 | Variable Arrow : U -> U -> U. 43 | Notation "a ~~> b" := (Arrow a b) (at level 60, right associativity). 44 | 45 | (* don't put it here, but we need it in the record, 46 | for the initial morphism has to have this 47 | property 48 | Hypothesis type_arrow : forall s t, f (s ~> t) = f s ~~> f t. 49 | *) 50 | 51 | (** a monad is a representation if it is accompagnied by 52 | - a lot of morphisms of modules 53 | - beta-reduction 54 | *) 55 | 56 | Class PCFPO_rep_struct := { 57 | app : forall u v, (P[u ~~> v]) x (P[u]) ---> P[v]; 58 | abs : forall u v, (d P //u)[v] ---> P[u ~~> v]; 59 | rec : forall t, P[t ~~> t] ---> P[t]; 60 | tttt : * ---> P[f Bool]; 61 | ffff : * ---> P[f Bool]; 62 | nats : forall m:nat, * ---> P[f Nat]; 63 | Succ : * ---> P[f Nat ~~> f Nat]; 64 | Pred : * ---> P[f Nat ~~> f Nat]; 65 | Zero : * ---> P[f Nat ~~> f Bool]; 66 | CondN: * ---> P[f Bool ~~> f Nat ~~> f Nat ~~> f Nat]; 67 | CondB: * ---> P[f Bool ~~> f Bool ~~> f Bool ~~> f Bool]; 68 | bottom: forall t, * ---> P[t]; 69 | 70 | beta_red : forall r s V y z, 71 | app r s V (abs r s V y, z) << y[*:= z] ; 72 | (* 73 | propag_app1: forall r s V y y' z, 74 | y << y' -> app r s V (y,z) << app _ _ _ (y',z) ; 75 | 76 | propag_app2: forall r s V y z z', 77 | z << z' -> app r s V (y,z) << app _ _ _ (y,z') ; 78 | 79 | propag_abs: forall r s V y y', 80 | y << y' -> abs r s V y << abs _ _ _ y' ; 81 | 82 | propag_rec: forall t V y y', 83 | y << y' -> rec t V y << rec _ _ y' ; 84 | *) 85 | CondN_t: forall V n m, 86 | app _ _ _ (app _ _ _ 87 | (app _ _ _ (CondN V tt, tttt _ tt), 88 | n), m) << n ; 89 | 90 | CondN_f: forall V n m, 91 | app _ _ _ (app _ _ _ 92 | (app _ _ _ (CondN V tt, ffff _ tt), 93 | n), m) << m ; 94 | 95 | CondB_t: forall V u v, 96 | app _ _ _ (app _ _ _ 97 | (app _ _ _ (CondB V tt, tttt _ tt), 98 | u), v) << u ; 99 | 100 | CondB_f: forall V u v, 101 | app _ _ _ (app _ _ _ 102 | (app _ _ _ (CondB V tt, ffff _ tt), 103 | u), v) << v ; 104 | 105 | Succ_red: forall V n, 106 | app _ _ _ (Succ V tt, nats n _ tt) << nats (S n) _ tt ; 107 | 108 | Zero_t: forall V, 109 | app _ _ _ (Zero V tt, nats 0 _ tt) << tttt _ tt ; 110 | 111 | Zero_f: forall V n, 112 | app _ _ _ (Zero V tt, nats (S n) _ tt) << ffff _ tt ; 113 | 114 | Pred_Succ: forall V n, 115 | app _ _ _ (Pred V tt, app _ _ _ (Succ V tt, nats n _ tt)) << 116 | nats n _ tt; 117 | 118 | Pred_Z: forall V, 119 | app _ _ _ (Pred V tt, nats 0 _ tt) << nats 0 _ tt ; 120 | 121 | Rec_A: forall V t g, 122 | rec _ _ g << app t t V (g, rec _ _ g) 123 | 124 | }. 125 | 126 | 127 | End PCF_rep. 128 | 129 | Record PCFPO_rep := { 130 | type_type : Type; 131 | type_arrow : type_type -> type_type -> type_type; 132 | pcf_rep_monad :> RMonad (SM_ipo type_type); 133 | type_mor : TY -> type_type; 134 | type_arrow_dist : forall s t, type_mor (s ~> t) = 135 | type_arrow (type_mor s) (type_mor t); 136 | pcf_rep_struct :> PCFPO_rep_struct pcf_rep_monad type_mor type_arrow 137 | 138 | }. 139 | 140 | 141 | Existing Instance pcf_rep_struct. 142 | Notation "a ~~> b" := (type_arrow a b) 143 | (at level 60, right associativity). 144 | 145 | 146 | 147 | 148 | 149 | 150 | -------------------------------------------------------------------------------- /PCF_o_c/RPCF_rep_eq.v: -------------------------------------------------------------------------------- 1 | Require Import Coq.Logic.FunctionalExtensionality. 2 | Require Import Coq.Logic.Eqdep. 3 | 4 | Require Export CatSem.PCF_o_c.RPCF_rep_hom. 5 | Require Import CatSem.CAT.rmonad_gen_comp. 6 | 7 | Set Implicit Arguments. 8 | Unset Strict Implicit. 9 | Unset Transparent Obligations. 10 | Unset Automatic Introduction. 11 | 12 | 13 | Inductive eq_Rep (P R : PCFPO_rep) : relation (PCFPO_rep_Hom P R) := 14 | | eq_rep : forall (a c : PCFPO_rep_Hom P R), 15 | forall H : (forall t, tcomp c t = tcomp a t), 16 | (forall V, rep_Hom_monad a V ;; rlift R (Transp H V) 17 | == 18 | Transp_PO H (P V) ;; rep_Hom_monad c V ) -> 19 | 20 | eq_Rep a c. 21 | 22 | (** the defined relation is an equivalence and 23 | may hence serve as equality 24 | *) 25 | 26 | Lemma eq_Rep_equiv (P R: PCFPO_rep) : 27 | @Equivalence (PCFPO_rep_Hom P R) 28 | (@eq_Rep P R). 29 | Proof. 30 | intros P R. 31 | split. 32 | 33 | unfold Reflexive. 34 | intro M. 35 | assert (H': forall t, tcomp M t = tcomp M t) by 36 | (intros; reflexivity). 37 | 38 | apply (eq_rep (H:=H')). 39 | 40 | simpl. 41 | intros V t y. 42 | destruct y as [t y]. 43 | 44 | simpl. 45 | rewrite (UIP_refl _ _ (H' t)). 46 | simpl. 47 | assert (H:= rlift_transp_id). 48 | simpl in *. 49 | rewrite H. 50 | auto. 51 | 52 | (* now symmetric *) 53 | 54 | unfold Symmetric. 55 | intros M N H. 56 | destruct H. 57 | assert (H': forall t, tcomp a t = tcomp c t) by auto. 58 | apply (eq_rep (H:=H')). 59 | simpl; intros V t y. 60 | destruct a. 61 | destruct c. 62 | simpl in *. 63 | assert (H3 : tcomp = tcomp0). 64 | apply (functional_extensionality). 65 | auto. 66 | 67 | generalize dependent y. 68 | generalize dependent H'. 69 | generalize dependent rep_Hom_monad0. 70 | generalize dependent rep_Hom_monad. 71 | generalize dependent ttriag. 72 | generalize dependent ttriag0. 73 | generalize dependent H. 74 | generalize dependent tcomp_arrow. 75 | rewrite H3. 76 | intros; simpl in *. 77 | rewrite transp_id. 78 | assert (H2:= rlift_transp_id). 79 | simpl in H2. 80 | rewrite H2. 81 | assert (H4 := H0 V t y). 82 | rewrite transp_id in H4. 83 | rewrite H2 in H4. 84 | simpl in *; auto. 85 | 86 | (* finally transitive *) 87 | 88 | unfold Transitive. 89 | intros a b c H H'. 90 | destruct H; 91 | destruct H'. 92 | assert (Ht : forall t, tcomp c t = tcomp a t). 93 | intro t. transitivity (tcomp a0 t); auto. 94 | 95 | apply (eq_rep (H:=Ht)). 96 | simpl; intros V t y. 97 | destruct a; 98 | destruct a0; 99 | destruct c. 100 | simpl in *. 101 | assert (H5 : tcomp0 = tcomp1) by 102 | (apply functional_extensionality; intro; auto). 103 | 104 | assert (H6 : tcomp1 = tcomp) by 105 | (apply functional_extensionality; intro; auto). 106 | 107 | generalize dependent H2. 108 | generalize dependent H1. 109 | generalize dependent rep_Hom_monad. 110 | generalize dependent rep_Hom_monad1. 111 | generalize dependent rep_Hom_monad0. 112 | generalize dependent ttriag. 113 | generalize dependent ttriag1. 114 | generalize dependent ttriag0. 115 | generalize dependent H. 116 | generalize dependent Ht. 117 | generalize dependent tcomp_arrow. 118 | generalize dependent tcomp_arrow0. 119 | generalize dependent tcomp_arrow1. 120 | rewrite H5. 121 | rewrite H6. 122 | 123 | clear H6 H5. 124 | clear tcomp0. 125 | clear tcomp1. 126 | 127 | intros. 128 | assert (H7:=H0 V t y). 129 | assert (H9:=H2 V t y). 130 | rewrite transp_id in *. 131 | simpl in *. 132 | assert (H8 := rlift_transp_id). 133 | simpl in H8. 134 | rewrite H8 in *. 135 | transitivity (rep_Hom_monad0 V t y); 136 | auto. 137 | Qed. 138 | 139 | Definition eq_Rep_oid (P R : PCFPO_rep) := Build_Setoid (eq_Rep_equiv P R). 140 | 141 | -------------------------------------------------------------------------------- /PCF_o_c/RPCF_syntax_rep.v: -------------------------------------------------------------------------------- 1 | 2 | Require Export CatSem.PCF.PCF_RMonad. 3 | Require Export CatSem.PCF_o_c.RPCF_rep. 4 | 5 | Set Implicit Arguments. 6 | Unset Strict Implicit. 7 | Unset Transparent Obligations. 8 | Unset Automatic Introduction. 9 | 10 | 11 | Program Instance PCFE_rep_struct : 12 | PCFPO_rep_struct PCFEM (fun t => t) (PCF.arrow):= { 13 | app r s := PCFApp r s; 14 | abs r s := PCFAbs r s; 15 | rec t := PCFRec t ; 16 | tttt := PCFconsts ttt ; 17 | ffff := PCFconsts fff; 18 | Succ := PCFconsts succ; 19 | Pred := PCFconsts preds; 20 | CondN := PCFconsts condN; 21 | CondB := PCFconsts condB; 22 | Zero := PCFconsts zero ; 23 | nats m := PCFconsts (Nats m); 24 | bottom t := PCFbottom t 25 | }. 26 | Next Obligation. 27 | Proof. 28 | unfold Rsubst_star_map. 29 | simpl. 30 | apply clos_refl_trans_1n_contains. 31 | apply relorig. 32 | apply app_abs. 33 | Qed. 34 | (* 35 | Next Obligation. 36 | Proof. 37 | apply cp_App1. 38 | auto. 39 | Qed. 40 | Next Obligation. 41 | Proof. 42 | apply cp_App2. 43 | auto. 44 | Qed. 45 | Next Obligation. 46 | Proof. 47 | apply cp_Lam. 48 | auto. 49 | Qed. 50 | Next Obligation. 51 | Proof. 52 | apply cp_Rec. 53 | auto. 54 | Qed. 55 | *) 56 | Next Obligation. 57 | Proof. 58 | apply clos_refl_trans_1n_contains. 59 | apply relorig. 60 | constructor. 61 | Qed. 62 | Next Obligation. 63 | Proof. 64 | apply clos_refl_trans_1n_contains. 65 | apply relorig. 66 | constructor. 67 | Qed. 68 | Next Obligation. 69 | Proof. 70 | apply clos_refl_trans_1n_contains. 71 | apply relorig. 72 | constructor. 73 | Qed. 74 | Next Obligation. 75 | Proof. 76 | apply clos_refl_trans_1n_contains. 77 | apply relorig. 78 | constructor. 79 | Qed. 80 | Next Obligation. 81 | Proof. 82 | apply clos_refl_trans_1n_contains. 83 | apply relorig. 84 | constructor. 85 | Qed. 86 | Next Obligation. 87 | Proof. 88 | apply clos_refl_trans_1n_contains. 89 | apply relorig. 90 | constructor. 91 | Qed. 92 | Next Obligation. 93 | Proof. 94 | apply clos_refl_trans_1n_contains. 95 | apply relorig. 96 | constructor. 97 | Qed. 98 | Next Obligation. 99 | Proof. 100 | apply clos_refl_trans_1n_contains. 101 | apply relorig. 102 | constructor. 103 | Qed. 104 | Next Obligation. 105 | Proof. 106 | apply clos_refl_trans_1n_contains. 107 | apply relorig. 108 | constructor. 109 | Qed. 110 | Next Obligation. 111 | Proof. 112 | apply clos_refl_trans_1n_contains. 113 | apply relorig. 114 | constructor. 115 | Qed. 116 | 117 | 118 | Definition PCFE_rep : PCFPO_rep := 119 | Build_PCFPO_rep (type_type:=TY) (type_arrow:=PCF.arrow) 120 | (pcf_rep_monad:=PCFEM)(type_mor:=fun t => t) 121 | (fun s t => eq_refl) 122 | PCFE_rep_struct. 123 | 124 | 125 | 126 | 127 | 128 | 129 | 130 | 131 | 132 | 133 | 134 | -------------------------------------------------------------------------------- /PCF_o_c/ULC_comp.v: -------------------------------------------------------------------------------- 1 | Require Export CatSem.PCF_o_c.RPCF_INIT. 2 | Require Export CatSem.PCF_o_c.RPCF_ULC_nounit. 3 | 4 | Set Implicit Arguments. 5 | Unset Strict Implicit. 6 | (*Unset Printing Implicit Defensive.*) 7 | Unset Automatic Introduction. 8 | 9 | 10 | Definition PCF_ULC_compilation := 11 | InitMor PCF_ULC. 12 | 13 | Definition PCF_ULC_c := 14 | rep_Hom_monad PCF_ULC_compilation. 15 | 16 | (** some examples *) 17 | 18 | Eval compute in 19 | (PCF_ULC_c (fun t => False) tt 20 | (ctype _ (Bottom (fun _ => False) Nat))). 21 | Eval compute in 22 | (PCF_ULC_c (fun t => False) tt 23 | (ctype _ (Const (fun _ => False) (succ )))). 24 | Eval compute in 25 | (PCF_ULC_c (fun t => False) tt 26 | (ctype _ (Const (fun _ => False) (preds )))). 27 | 28 | Eval compute in 29 | (PCF_ULC_c (fun t => False) tt 30 | (ctype _ (Const (fun _ => False) (condB )))). 31 | 32 | 33 | Eval compute in 34 | (PCF_ULC_c ( (fun t => False)) tt 35 | (ctype _ ( 36 | Lam (PCF_syntax.Var (none Bool (fun t => False)))))). 37 | 38 | Check ctype. 39 | Check (Const (fun _ => False) condB). 40 | 41 | Check (fun V t (a : PCF V t) => PCF_ULC_c (V) tt 42 | (ctype (fun t => tt) a)). 43 | 44 | Require Import Coq.Program.Equality. 45 | 46 | Lemma aaaa V t (a : PCF V t) b : 47 | (PCF_ULC_c _ _ (ctype (fun t => tt) a)) >>> b -> 48 | (exists b', a >> b' /\ PCF_ULC_c _ _ (ctype _ b') = b). 49 | intros V t a. 50 | intro b. 51 | dependent induction b; 52 | simpl; intros. 53 | inversion H. 54 | exists (PCF_syntax.Var (ct ). 55 | exists (Bottom V t). 56 | split. reflexivity. 57 | simpl. 58 | inversion H. auto. 59 | subst. clear H. 60 | inversion H0. 61 | subst. 62 | exists (Bottom V t). 63 | split. 64 | reflexivity. 65 | simpl; auto. 66 | subst. 67 | exists (Bottom V t). 68 | split. reflexivity. 69 | ssiauto. 70 | 71 | destruct H. 72 | intros. 73 | elim H. 74 | clear H. 75 | intros. 76 | simpl in *. 77 | inversion H. 78 | subst. 79 | exists a. 80 | reflexivity. 81 | subst. 82 | 83 | 84 | 85 | dependent induction H. 86 | 87 | 88 | 89 | -------------------------------------------------------------------------------- /PCF_order_comp/RPCF_INIT.v: -------------------------------------------------------------------------------- 1 | Require Import Coq.Logic.FunctionalExtensionality. 2 | Require Import Coq.Logic.Eqdep. 3 | 4 | Require Export CatSem.PCF_order_comp.RPCF_rep_cat. 5 | Require Export CatSem.PCF_order_comp.RPCF_syntax_init. 6 | 7 | Set Implicit Arguments. 8 | Unset Strict Implicit. 9 | Unset Transparent Obligations. 10 | Unset Automatic Introduction. 11 | 12 | 13 | Section unique. 14 | 15 | Variable R : REP. 16 | 17 | Variable g : PCFE_rep ---> R. 18 | 19 | Lemma initR_unique : g == initR R. 20 | Proof. 21 | simpl. 22 | assert (H : forall t, tcomp (initR R) t = tcomp g t). 23 | assert (H':= ttriag g). 24 | simpl in *. 25 | auto. 26 | apply (eq_rep (H:=H)). 27 | simpl. 28 | intros V t y. 29 | destruct y as [t y]. 30 | simpl. 31 | 32 | destruct g. 33 | simpl in *. 34 | 35 | assert (H' : type_mor R = tcomp) by 36 | (apply functional_extensionality; 37 | intros; auto). 38 | 39 | generalize dependent rep_Hom_monad. 40 | generalize dependent ttriag. 41 | generalize dependent H. 42 | 43 | rewrite <- H'. 44 | 45 | simpl. 46 | intros. 47 | 48 | assert (Hl := rlift_transp_id (Q:=R) H). 49 | simpl in *. 50 | rewrite Hl. 51 | rewrite (UIP_refl _ _ (H t)). 52 | simpl. 53 | 54 | induction y; 55 | simpl. 56 | 57 | assert (Hb:=bottom_hom (PCFPO_rep_Hom_struct := 58 | rep_gen_Hom_monad_struct) t (c:=V) tt). 59 | simpl in *. 60 | 61 | rewrite (UIP_refl _ _ (ttriag _)) in Hb. 62 | simpl in *. 63 | auto. 64 | 65 | destruct c. 66 | 67 | assert (Hb:=nats_hom n (PCFPO_rep_Hom_struct := 68 | rep_gen_Hom_monad_struct) (c:=V) tt). 69 | simpl in *. 70 | 71 | rewrite (UIP_refl _ _ (ttriag _)) in Hb. 72 | simpl in * ; auto. 73 | 74 | assert (Hb:=ttt_hom (PCFPO_rep_Hom_struct := 75 | rep_gen_Hom_monad_struct) (c:=V) tt). 76 | simpl in *. 77 | 78 | rewrite (UIP_refl _ _ (ttriag _)) in Hb. 79 | simpl in * ; auto. 80 | 81 | assert (Hb:=fff_hom (PCFPO_rep_Hom_struct := 82 | rep_gen_Hom_monad_struct) (c:=V) tt). 83 | simpl in *. 84 | 85 | rewrite (UIP_refl _ _ (ttriag _)) in Hb. 86 | simpl in * ; auto. 87 | 88 | assert (Hb:=Succ_hom (PCFPO_rep_Hom_struct := 89 | rep_gen_Hom_monad_struct) (c:=V) tt). 90 | simpl in *. 91 | 92 | rewrite (UIP_refl _ _ (ttriag _)) in Hb. 93 | simpl in * ; auto. 94 | 95 | assert (Hb:=Pred_hom (PCFPO_rep_Hom_struct := 96 | rep_gen_Hom_monad_struct) (c:=V) tt). 97 | simpl in *. 98 | 99 | rewrite (UIP_refl _ _ (ttriag _)) in Hb. 100 | simpl in * ; auto. 101 | 102 | assert (Hb:=Zero_hom (PCFPO_rep_Hom_struct := 103 | rep_gen_Hom_monad_struct) (c:=V) tt). 104 | simpl in *. 105 | 106 | rewrite (UIP_refl _ _ (ttriag _)) in Hb. 107 | simpl in * ; auto. 108 | 109 | assert (Hb:=CondN_hom (PCFPO_rep_Hom_struct := 110 | rep_gen_Hom_monad_struct) (c:=V) tt). 111 | simpl in *. 112 | 113 | rewrite (UIP_refl _ _ (ttriag _)) in Hb. 114 | simpl in * ; auto. 115 | 116 | assert (Hb:=CondB_hom (PCFPO_rep_Hom_struct := 117 | rep_gen_Hom_monad_struct) (c:=V) tt). 118 | simpl in *. 119 | 120 | rewrite (UIP_refl _ _ (ttriag _)) in Hb. 121 | simpl in * ; auto. 122 | 123 | 124 | (*1*) 125 | assert (Hw:=gen_rmonad_hom_rweta 126 | ((*gen_RMonad_Hom_struct := *)rep_Hom_monad)). 127 | simpl in Hw. 128 | assert (Hw' := Hw _ _ (ctype _ v)). 129 | simpl in Hw'. 130 | auto. 131 | 132 | rewrite <- IHy1. 133 | rewrite <- IHy2. 134 | assert (Happ:=app_hom (PCFPO_rep_Hom_struct := 135 | rep_gen_Hom_monad_struct) (r:=s) (s:=t) (y1,y2)). 136 | simpl in *. 137 | 138 | rewrite (UIP_refl _ _ (ttriag s)) in Happ. 139 | rewrite (UIP_refl _ _ (ttriag t)) in Happ. 140 | rewrite (UIP_refl _ _ (ttriag (s ~> t))) in Happ. 141 | simpl in *. 142 | simpl in *. 143 | auto. 144 | (*2*) 145 | rewrite <- IHy. 146 | assert (Habs:=abs_hom (PCFPO_rep_Hom_struct := 147 | rep_gen_Hom_monad_struct) (r:=t) (s:=s) y ). 148 | simpl in *. 149 | 150 | rewrite (UIP_refl _ _ (ttriag s)) in Habs. 151 | rewrite (UIP_refl _ _ (ttriag t)) in Habs. 152 | rewrite (UIP_refl _ _ (ttriag _)) in Habs. 153 | simpl in *. 154 | auto. 155 | 156 | rewrite <- IHy. 157 | assert (Happ:=rec_hom (PCFPO_rep_Hom_struct := 158 | rep_gen_Hom_monad_struct) (t:=t) y). 159 | simpl in *. 160 | 161 | rewrite (UIP_refl _ _ (ttriag t)) in Happ. 162 | rewrite (UIP_refl _ _ (ttriag (t ~> t))) in Happ. 163 | simpl in *. 164 | auto. 165 | Qed. 166 | 167 | End unique. 168 | 169 | Hint Resolve initR_unique : fin. 170 | 171 | Obligation Tactic := fin. 172 | 173 | Program Instance PCF_initial : Initial REP := { 174 | Init := PCFE_rep ; 175 | InitMor R := initR R ; 176 | InitMorUnique R := @initR_unique R 177 | }. 178 | 179 | Print Assumptions PCF_initial. 180 | -------------------------------------------------------------------------------- /PCF_order_comp/RPCF_rep.v: -------------------------------------------------------------------------------- 1 | Require Export CatSem.CAT.retype_functor_po. 2 | Require Export CatSem.CAT.rmonad_gen_type_po. 3 | Require Export CatSem.PCF.PCF_types. 4 | 5 | Set Implicit Arguments. 6 | Unset Strict Implicit. 7 | Unset Transparent Obligations. 8 | Unset Automatic Introduction. 9 | 10 | Notation "'TY'" := PCF.types. 11 | Notation "'Bool'" := PCF.Bool. 12 | Notation "'Nat'" := PCF.Nat. 13 | 14 | Notation "'IP'" := (IPO TY). 15 | Notation "a '~>' b" := (PCF.arrow a b) 16 | (at level 69, right associativity). 17 | Notation "a 'x' b" := (product a b) (at level 30). 18 | Notation "M [ z ]" := (FIB_RMOD _ z M) (at level 35). 19 | Notation "'d' M // s" := (DER_RMOD _ _ s M) (at level 25). 20 | Notation "'*'" := (Term (C:=RMOD _ PO)). 21 | 22 | Notation "f 'X' g" := (product_mor _ f g)(at level 30). 23 | 24 | Notation "'FM'" := (#(FIB_RMOD _ _ )). 25 | Notation "'FPB'":= (FIB_RPB _ _ _ ). 26 | Notation "'PRPB'":= (PROD_RPB _ _ _ _ ). 27 | (*Notation "'PBF'":= (PB_FIB _ _ _ ).*) 28 | Notation "'PBM'":= (#(PbRMOD _ _ )). 29 | Notation "'DM'":= (#(DER_RMOD _ _ _ )). 30 | Notation "'DPB'":= (DER_RPB _ _ _ ). 31 | 32 | Notation "y [* := z ]":= (Rsubstar z _ y)(at level 55). 33 | 34 | 35 | 36 | Section PCF_rep. 37 | 38 | Variable U : Type. 39 | Variable P : RMonad (SM_ipo U). 40 | Variable f : TY -> U. 41 | 42 | (** a monad is a representation if it is accompagnied by 43 | - a lot of morphisms of modules 44 | - beta-reduction 45 | *) 46 | 47 | Class PCFPO_rep_struct := { 48 | app : forall r s, (P[f (r~>s)]) x (P[f r]) ---> P[f s] 49 | where "A @ B" := (app _ _ _ (A,B)); 50 | abs : forall r s, (d P //(f r))[f s] ---> P[f (r ~> s)]; 51 | rec : forall t, P[f (t ~> t)] ---> P[f t]; 52 | tttt : * ---> P[f Bool]; 53 | ffff : * ---> P[f Bool]; 54 | nats : forall m:nat, * ---> P[f Nat]; 55 | Succ : * ---> P[f (Nat ~> Nat)]; 56 | Pred : * ---> P[f (Nat ~> Nat)]; 57 | Zero : * ---> P[f (Nat ~> Bool)]; 58 | CondN: * ---> P[f (Bool ~> Nat ~> Nat ~> Nat)]; 59 | CondB: * ---> P[f (Bool ~> Bool ~> Bool ~> Bool)]; 60 | bottom: forall t, * ---> P[f t]; 61 | 62 | beta_red : forall r s V y z, 63 | app r s V (abs r s V y, z) << y[*:= z] ; 64 | (* 65 | propag_app1: forall r s V y y' z, 66 | y << y' -> app r s V (y,z) << app _ _ _ (y',z) ; 67 | 68 | propag_app2: forall r s V y z z', 69 | z << z' -> app r s V (y,z) << app _ _ _ (y,z') ; 70 | 71 | propag_abs: forall r s V y y', 72 | y << y' -> abs r s V y << abs _ _ _ y' ; 73 | 74 | propag_rec: forall t V y y', 75 | y << y' -> rec t V y << rec _ _ y' ; 76 | *) 77 | CondN_t: forall V n m, 78 | app _ _ _ (app _ _ _ 79 | (app _ _ _ (CondN V tt, tttt _ tt), 80 | n), m) << n ; 81 | 82 | CondN_f: forall V n m, 83 | app _ _ _ (app _ _ _ 84 | (app _ _ _ (CondN V tt, ffff _ tt), 85 | n), m) << m ; 86 | 87 | CondB_t: forall V u v, 88 | app _ _ _ (app _ _ _ 89 | (app _ _ _ (CondB V tt, tttt _ tt), 90 | u), v) << u ; 91 | 92 | CondB_f: forall V u v, 93 | app _ _ _ (app _ _ _ 94 | (app _ _ _ (CondB V tt, ffff _ tt), 95 | u), v) << v ; 96 | 97 | Succ_red: forall V n, 98 | app _ _ _ (Succ V tt, nats n _ tt) << nats (S n) _ tt ; 99 | 100 | Zero_t: forall V, 101 | app _ _ _ (Zero V tt, nats 0 _ tt) << tttt _ tt ; 102 | 103 | Zero_f: forall V n, 104 | app _ _ _ (Zero V tt, nats (S n) _ tt) << ffff _ tt ; 105 | 106 | (*Pred_Sn: forall V n, 107 | app _ _ _ (Pred V tt, nats (S n) _ tt) << nats n _ tt ;*) 108 | 109 | Pred_Succ: forall V n, 110 | app _ _ _ (Pred V tt, app _ _ _ (Succ V tt, nats n _ tt)) << 111 | nats n _ tt; 112 | 113 | Pred_Z: forall V, 114 | app _ _ _ (Pred V tt, nats 0 _ tt) << nats 0 _ tt ; 115 | 116 | Rec_A: forall V t g, 117 | rec _ _ g << app t t V (g, rec _ _ g) 118 | 119 | (* 120 | Rec_B: forall V t g z, 121 | rec _ _ g << z -> app t t V (g , rec _ _ g) << z; 122 | 123 | Rec_C: forall V t g, exists z, 124 | rec _ _ g << z /\ app t t V (g, rec _ _ g) << z ; 125 | 126 | Rec_D: forall V t g M, (forall s a b, M <> app s t _ (a, b)) -> 127 | app t t V (g, rec _ _ g) << M -> rec _ _ g << M 128 | *) 129 | }. 130 | 131 | 132 | End PCF_rep. 133 | 134 | Record PCFPO_rep := { 135 | type_type : Type; 136 | pcf_rep_monad :> RMonad (SM_ipo type_type); 137 | type_mor : TY -> type_type; 138 | pcf_rep_struct :> PCFPO_rep_struct pcf_rep_monad type_mor 139 | }. 140 | 141 | 142 | Existing Instance pcf_rep_struct. 143 | 144 | 145 | 146 | 147 | 148 | 149 | 150 | -------------------------------------------------------------------------------- /PCF_order_comp/RPCF_rep_eq.v: -------------------------------------------------------------------------------- 1 | Require Import Coq.Logic.FunctionalExtensionality. 2 | Require Import Coq.Logic.Eqdep. 3 | 4 | Require Export CatSem.PCF_order_comp.RPCF_rep_hom. 5 | Require Import CatSem.CAT.rmonad_gen_comp. 6 | 7 | Set Implicit Arguments. 8 | Unset Strict Implicit. 9 | Unset Transparent Obligations. 10 | Unset Automatic Introduction. 11 | 12 | 13 | Inductive eq_Rep (P R : PCFPO_rep) : relation (PCFPO_rep_Hom P R) := 14 | | eq_rep : forall (a c : PCFPO_rep_Hom P R), 15 | forall H : (forall t, tcomp c t = tcomp a t), 16 | (forall V, rep_Hom_monad a V ;; rlift R (Transp H V) 17 | == 18 | Transp_PO H (P V) ;; rep_Hom_monad c V ) -> 19 | 20 | eq_Rep a c. 21 | 22 | (** the defined relation is an equivalence and 23 | may hence serve as equality 24 | *) 25 | 26 | Lemma eq_Rep_equiv (P R: PCFPO_rep) : 27 | @Equivalence (PCFPO_rep_Hom P R) 28 | (@eq_Rep P R). 29 | Proof. 30 | intros P R. 31 | split. 32 | 33 | unfold Reflexive. 34 | intro M. 35 | assert (H': forall t, tcomp M t = tcomp M t) by 36 | (intros; reflexivity). 37 | 38 | apply (eq_rep (H:=H')). 39 | 40 | simpl. 41 | intros V t y. 42 | destruct y as [t y]. 43 | 44 | simpl. 45 | rewrite (UIP_refl _ _ (H' t)). 46 | simpl. 47 | assert (H:= rlift_transp_id). 48 | simpl in *. 49 | rewrite H. 50 | auto. 51 | 52 | (* now symmetric *) 53 | 54 | unfold Symmetric. 55 | intros M N H. 56 | destruct H. 57 | assert (H': forall t, tcomp a t = tcomp c t) by auto. 58 | apply (eq_rep (H:=H')). 59 | simpl; intros V t y. 60 | destruct a. 61 | destruct c. 62 | simpl in *. 63 | assert (H3 : tcomp = tcomp0). 64 | apply (functional_extensionality). 65 | auto. 66 | 67 | generalize dependent y. 68 | generalize dependent H'. 69 | generalize dependent rep_Hom_monad0. 70 | generalize dependent rep_Hom_monad. 71 | generalize dependent ttriag. 72 | generalize dependent ttriag0. 73 | generalize dependent H. 74 | rewrite H3. 75 | intros; simpl in *. 76 | rewrite transp_id. 77 | assert (H2:= rlift_transp_id). 78 | simpl in H2. 79 | rewrite H2. 80 | assert (H4 := H0 V t y). 81 | rewrite transp_id in H4. 82 | rewrite H2 in H4. 83 | simpl in *; auto. 84 | 85 | (* finally transitive *) 86 | 87 | unfold Transitive. 88 | intros a b c H H'. 89 | destruct H; 90 | destruct H'. 91 | assert (Ht : forall t, tcomp c t = tcomp a t). 92 | intro t. transitivity (tcomp a0 t); auto. 93 | 94 | apply (eq_rep (H:=Ht)). 95 | simpl; intros V t y. 96 | destruct a; 97 | destruct a0; 98 | destruct c. 99 | simpl in *. 100 | assert (H5 : tcomp0 = tcomp1) by 101 | (apply functional_extensionality; intro; auto). 102 | 103 | assert (H6 : tcomp1 = tcomp) by 104 | (apply functional_extensionality; intro; auto). 105 | 106 | generalize dependent H2. 107 | generalize dependent H1. 108 | generalize dependent rep_Hom_monad. 109 | generalize dependent rep_Hom_monad1. 110 | generalize dependent rep_Hom_monad0. 111 | generalize dependent ttriag. 112 | generalize dependent ttriag1. 113 | generalize dependent ttriag0. 114 | generalize dependent H. 115 | generalize dependent Ht. 116 | rewrite H5. 117 | rewrite H6. 118 | 119 | clear H6 H5. 120 | clear tcomp0. 121 | clear tcomp1. 122 | 123 | intros. 124 | assert (H7:=H0 V t y). 125 | assert (H9:=H2 V t y). 126 | rewrite transp_id in *. 127 | simpl in *. 128 | assert (H8 := rlift_transp_id). 129 | simpl in H8. 130 | rewrite H8 in *. 131 | transitivity (rep_Hom_monad0 V t y); 132 | auto. 133 | Qed. 134 | 135 | Definition eq_Rep_oid (P R : PCFPO_rep) := Build_Setoid (eq_Rep_equiv P R). 136 | 137 | -------------------------------------------------------------------------------- /PCF_order_comp/RPCF_rep_id.v: -------------------------------------------------------------------------------- 1 | Require Export CatSem.PCF_order_comp.RPCF_rep_hom. 2 | 3 | Set Implicit Arguments. 4 | Unset Strict Implicit. 5 | Unset Transparent Obligations. 6 | Unset Automatic Introduction. 7 | 8 | Section id_rep. 9 | 10 | Variable P : PCFPO_rep. 11 | 12 | Definition id_rep_car c: 13 | (forall t : type_type P, 14 | (retype_ipo (fun u : type_type P => u) (P c)) t -> 15 | (P (retype (fun u : type_type P => u) c)) t) := 16 | fun t y => 17 | match y with ctype _ z => 18 | rlift P (@id_retype _ c) _ z end. 19 | 20 | Obligation Tactic := 21 | intros; do 2 red; 22 | simpl; intros; 23 | try match goal with [H: retype_ord _ _ |- _ ] => induction H end; 24 | try apply IPO_mor_monotone; auto. 25 | 26 | Program Instance id_rep_pos c: 27 | ipo_mor_struct (a:=retype_ipo (fun u : type_type P => u) (P c)) 28 | (b:=P (retype (fun u : type_type P => u) c)) (@id_rep_car c). 29 | 30 | Definition id_rep_po c := Build_ipo_mor (id_rep_pos c). 31 | 32 | Definition id_rep_c : 33 | (forall c : ITYPE (type_type P), 34 | (RETYPE_PO (fun u : type_type P => u)) (P c) ---> 35 | P ((RETYPE (fun u : type_type P => u)) c)) := 36 | id_rep_po. 37 | 38 | Ltac elim_retype := match goal with 39 | [y : retype _ _ _ |- _ ] => destruct y end. 40 | 41 | Obligation Tactic := cat; repeat elim_retype; 42 | cat; 43 | try rew (rlift_rweta P); 44 | try rew (rlift_rkleisli (M:=P)); 45 | try rew (rkleisli_rlift (M:=P)); 46 | try apply (rkl_eq P); cat. 47 | 48 | Program Instance RMon_id_s : 49 | gen_RMonad_Hom_struct (P:=P) (Q:=P) 50 | (G1:=RETYPE (fun u : type_type P => u)) 51 | (G2:=RETYPE_PO (fun u : type_type P => u)) 52 | (NNNT1 (fun u : type_type P => u)) id_rep_c. 53 | 54 | Definition RMon_id := Build_gen_RMonad_Hom RMon_id_s. 55 | 56 | Obligation Tactic := simpl; intros; unfold rlift; 57 | match goal with [|- _ = (_ _ _ ((rmodule_hom ?H) _ )) _ ] => 58 | rerew (rmod_hom_rmkl H) end; 59 | try apply f_equal; try rew (rklkl P); try apply (rkl_eq P); 60 | simpl; intros; try rew (retakl P); elim_opt; 61 | try rew (rlift_rweta P); auto. 62 | 63 | Program Instance PCFPO_id_struct : PCFPO_rep_Hom_struct 64 | (P:=P) (R:=P) 65 | (f:=fun u => u) (fun u => eq_refl) RMon_id. 66 | 67 | Definition PCFPO_id := Build_PCFPO_rep_Hom PCFPO_id_struct. 68 | 69 | End id_rep. 70 | 71 | 72 | 73 | -------------------------------------------------------------------------------- /PCF_order_comp/RPCF_syntax_rep.v: -------------------------------------------------------------------------------- 1 | (*Require Export CatSem.PCF_order_comp.RPCF_syntax.*) 2 | Require Export CatSem.PCF.PCF_RMonad. 3 | Require Export CatSem.PCF_order_comp.RPCF_rep. 4 | 5 | Set Implicit Arguments. 6 | Unset Strict Implicit. 7 | Unset Transparent Obligations. 8 | Unset Automatic Introduction. 9 | 10 | 11 | Program Instance PCFE_rep_struct : 12 | PCFPO_rep_struct PCFEM (fun t => t) := { 13 | app r s := PCFApp r s; 14 | abs r s := PCFAbs r s; 15 | rec t := PCFRec t ; 16 | tttt := PCFconsts ttt ; 17 | ffff := PCFconsts fff; 18 | Succ := PCFconsts succ; 19 | Pred := PCFconsts preds; 20 | CondN := PCFconsts condN; 21 | CondB := PCFconsts condB; 22 | Zero := PCFconsts zero ; 23 | nats m := PCFconsts (Nats m); 24 | bottom t := PCFbottom t 25 | }. 26 | Next Obligation. 27 | Proof. 28 | unfold Rsubst_star_map. 29 | simpl. 30 | apply clos_refl_trans_1n_contains. 31 | apply relorig. 32 | apply app_abs. 33 | Qed. 34 | Next Obligation. 35 | Proof. 36 | apply clos_refl_trans_1n_contains. 37 | apply relorig. 38 | constructor. 39 | Qed. 40 | Next Obligation. 41 | Proof. 42 | apply clos_refl_trans_1n_contains. 43 | apply relorig. 44 | constructor. 45 | Qed. 46 | Next Obligation. 47 | Proof. 48 | apply clos_refl_trans_1n_contains. 49 | apply relorig. 50 | constructor. 51 | Qed. 52 | Next Obligation. 53 | Proof. 54 | apply clos_refl_trans_1n_contains. 55 | apply relorig. 56 | constructor. 57 | Qed. 58 | Next Obligation. 59 | Proof. 60 | apply clos_refl_trans_1n_contains. 61 | apply relorig. 62 | constructor. 63 | Qed. 64 | Next Obligation. 65 | Proof. 66 | apply clos_refl_trans_1n_contains. 67 | apply relorig. 68 | constructor. 69 | Qed. 70 | Next Obligation. 71 | Proof. 72 | apply clos_refl_trans_1n_contains. 73 | apply relorig. 74 | constructor. 75 | Qed. 76 | Next Obligation. 77 | Proof. 78 | apply clos_refl_trans_1n_contains. 79 | apply relorig. 80 | constructor. 81 | Qed. 82 | Next Obligation. 83 | Proof. 84 | apply clos_refl_trans_1n_contains. 85 | apply relorig. 86 | constructor. 87 | Qed. 88 | Next Obligation. 89 | Proof. 90 | apply clos_refl_trans_1n_contains. 91 | apply relorig. 92 | constructor. 93 | Qed. 94 | 95 | Definition PCFE_rep : PCFPO_rep := Build_PCFPO_rep PCFE_rep_struct. 96 | 97 | 98 | 99 | 100 | 101 | 102 | 103 | 104 | 105 | 106 | 107 | -------------------------------------------------------------------------------- /PCF_order_comp/ULC_comp.v: -------------------------------------------------------------------------------- 1 | Require Export CatSem.PCF_order_comp.RPCF_ULC_sandbox. 2 | Require Export CatSem.PCF_order_comp.RPCF_INIT. 3 | 4 | Unset Printing Implicit Defensive. 5 | 6 | Definition PCF_ULC_compilation := 7 | InitMor PCF_ULC. 8 | 9 | Definition PCF_ULC_c := 10 | rep_Hom_monad PCF_ULC_compilation. 11 | (** example 12 | Eval compute in 13 | (PCF_ULC_c (fun t => False) tt 14 | (ctype _ (Bottom (fun _ => False) Nat))). 15 | *) -------------------------------------------------------------------------------- /PROP_untyped/ex_ulcbeta.v: -------------------------------------------------------------------------------- 1 | Require Export CatSem.PROP_untyped.prop_arities_initial. 2 | 3 | Require Import Coq.Program.Equality. 4 | 5 | Set Implicit Arguments. 6 | Unset Strict Implicit. 7 | Unset Automatic Introduction. 8 | Unset Transparent Obligations. 9 | 10 | (* Check INITIAL_INEQ_REP. *) 11 | 12 | Inductive Lambda_index := ABS | APP. 13 | 14 | Definition Lambda : Signature := {| 15 | sig_index := Lambda_index ; 16 | sig := fun x => match x with 17 | | ABS => [[ 1 ]] 18 | | APP => [[ 0 ; 0]] 19 | end 20 | |}. 21 | 22 | (*Definition Lambda_subst := subst_half_eq Lambda. 23 | Check Lambda_subst. 24 | *) 25 | 26 | Section App_Abs_half_eq. 27 | 28 | 29 | Section for_a_representation. 30 | 31 | Variable R : REP Lambda. 32 | 33 | Definition beta_carrier : 34 | (forall c : TYPE, (S_Mod_classic_ob [[1; 0]] R) c ---> (S_Mod_classic_ob [[0]] R) c) . 35 | simpl. 36 | intros. 37 | simpl in *. 38 | inversion X. 39 | simpl in *. 40 | inversion X1. 41 | simpl in X2. 42 | constructor. 43 | simpl. 44 | destruct R as [Rr Repr]. 45 | simpl in *. 46 | apply (Repr APP). 47 | simpl. 48 | simpl in *. 49 | constructor. 50 | simpl. 51 | apply (Repr ABS). 52 | simpl. 53 | constructor. 54 | simpl. 55 | apply X0. 56 | constructor. 57 | constructor. 58 | simpl. 59 | apply X2. 60 | constructor. 61 | constructor. 62 | Defined. 63 | 64 | Program Instance beta_struct : RModule_Hom_struct 65 | (M:=S_Mod_classic_ob [[1;0]] R) (N:=S_Mod_classic_ob [[0]] R) beta_carrier. 66 | Next Obligation. 67 | Proof. 68 | dependent destruction x. 69 | simpl. 70 | dependent destruction x. 71 | simpl in *. 72 | apply CONSTR_eq; auto. 73 | destruct R. 74 | simpl in *. 75 | rerew (rmhom_rmkl (repr APP)). 76 | apply f_equal. 77 | simpl in *. 78 | apply CONSTR_eq; auto. 79 | clear x. 80 | clear p0. 81 | rerew (rmhom_rmkl (repr ABS)). 82 | Qed. 83 | 84 | Definition beta_module_mor := Build_RModule_Hom beta_struct. 85 | 86 | End for_a_representation. 87 | 88 | Program Instance beta_half_s : half_equation_struct 89 | (U:=S_Mod_classic Lambda [[1 ; 0]]) 90 | (V:=S_Mod_classic Lambda [[0]]) 91 | beta_module_mor. 92 | Next Obligation. 93 | Proof. 94 | 95 | dependent destruction x. 96 | dependent destruction x. 97 | dependent destruction x. 98 | 99 | simpl. 100 | apply CONSTR_eq; auto. 101 | destruct T; simpl. 102 | destruct R; simpl. 103 | assert (H:=@repr_hom_s _ _ _ _ f). 104 | simpl in *. 105 | assert (Habs := H ABS). 106 | simpl in *. 107 | unfold commute in Habs; simpl in *. 108 | assert (Happ := H APP). 109 | simpl in *. 110 | unfold commute in Happ; simpl in *. 111 | rewrite <- Happ. 112 | apply f_equal. 113 | simpl in *. 114 | apply CONSTR_eq; auto. 115 | rewrite <- Habs. 116 | auto. 117 | Qed. 118 | 119 | 120 | Definition beta_half_eq := Build_half_equation beta_half_s. 121 | 122 | 123 | End App_Abs_half_eq. 124 | 125 | 126 | Definition beta_rule : ineq_classic Lambda := {| 127 | half_eq_l := beta_half_eq ; 128 | half_eq_r := subst_half_eq Lambda |}. 129 | 130 | 131 | 132 | Definition Lambda_beta_Cat := INEQ_REP 133 | (S:=Lambda)(A:=unit)(fun x : unit => beta_rule). 134 | 135 | 136 | Definition Lambda_beta_SynSem := (INITIAL_INEQ_REP (fun x : unit => beta_rule)). 137 | 138 | Definition Lambda_beta := @Init _ _ _ (INITIAL_INEQ_REP (fun x : unit => beta_rule)). 139 | 140 | 141 | 142 | 143 | 144 | 145 | 146 | 147 | 148 | 149 | 150 | 151 | 152 | 153 | 154 | 155 | -------------------------------------------------------------------------------- /PROP_untyped/ex_ulcbeta_variant.v: -------------------------------------------------------------------------------- 1 | Require Export CatSem.PROP_untyped.prop_arities_initial_variant. 2 | 3 | Require Import Coq.Program.Equality. 4 | 5 | Set Implicit Arguments. 6 | Unset Strict Implicit. 7 | Unset Automatic Introduction. 8 | Unset Transparent Obligations. 9 | 10 | (* Check INITIAL_INEQ_REP. *) 11 | 12 | Inductive Lambda_index := ABS | APP. 13 | 14 | Definition Lambda : Signature := {| 15 | sig_index := Lambda_index ; 16 | sig := fun x => match x with 17 | | ABS => [[ 1 ]] 18 | | APP => [[ 0 ; 0]] 19 | end 20 | |}. 21 | 22 | (*Definition Lambda_subst := subst_half_eq Lambda. 23 | Check Lambda_subst. 24 | *) 25 | 26 | Section App_Abs_half_eq. 27 | 28 | 29 | Section for_a_representation. 30 | 31 | Variable R : REP Lambda. 32 | 33 | Definition beta_carrier : 34 | (forall c : TYPE, (S_Mod_classic_ob [[1; 0]] R) c ---> (S_Mod_classic_ob [[0]] R) c) . 35 | simpl. 36 | intros. 37 | simpl in *. 38 | inversion X. 39 | simpl in *. 40 | inversion X1. 41 | simpl in X2. 42 | constructor. 43 | simpl. 44 | destruct R as [Rr Repr]. 45 | simpl in *. 46 | apply (Repr APP). 47 | simpl. 48 | simpl in *. 49 | constructor. 50 | simpl. 51 | apply (Repr ABS). 52 | simpl. 53 | constructor. 54 | simpl. 55 | apply X0. 56 | constructor. 57 | constructor. 58 | simpl. 59 | apply X2. 60 | constructor. 61 | constructor. 62 | Defined. 63 | 64 | Program Instance beta_struct : RModule_Hom_struct 65 | (M:=S_Mod_classic_ob [[1;0]] R) (N:=S_Mod_classic_ob [[0]] R) beta_carrier. 66 | Next Obligation. 67 | Proof. 68 | dependent destruction x. 69 | simpl. 70 | dependent destruction x. 71 | simpl in *. 72 | apply CONSTR_eq; auto. 73 | destruct R. 74 | simpl in *. 75 | rerew (rmhom_rmkl (repr APP)). 76 | apply f_equal. 77 | simpl in *. 78 | apply CONSTR_eq; auto. 79 | clear x. 80 | clear p0. 81 | rerew (rmhom_rmkl (repr ABS)). 82 | Qed. 83 | 84 | Definition beta_module_mor := Build_RModule_Hom beta_struct. 85 | 86 | End for_a_representation. 87 | Check S_Mod_classic. 88 | Program Instance beta_half_s : half_equation_struct 89 | (U:=S_Mod_classic [[1 ; 0]]) 90 | (V:=S_Mod_classic [[0]]) 91 | beta_module_mor. 92 | Next Obligation. 93 | Proof. 94 | 95 | dependent destruction x. 96 | dependent destruction x. 97 | dependent destruction x. 98 | 99 | simpl. 100 | apply CONSTR_eq; auto. 101 | destruct T; simpl. 102 | destruct R; simpl. 103 | assert (H:=@repr_hom_s _ _ _ _ f). 104 | simpl in *. 105 | assert (Habs := H ABS). 106 | simpl in *. 107 | unfold commute in Habs; simpl in *. 108 | assert (Happ := H APP). 109 | simpl in *. 110 | unfold commute in Happ; simpl in *. 111 | rewrite <- Happ. 112 | apply f_equal. 113 | simpl in *. 114 | apply CONSTR_eq; auto. 115 | rewrite <- Habs. 116 | auto. 117 | Qed. 118 | 119 | 120 | Definition beta_half_eq := Build_half_equation beta_half_s. 121 | 122 | 123 | End App_Abs_half_eq. 124 | 125 | 126 | Definition beta_rule : ineq_classic Lambda := {| 127 | half_eq_l := beta_half_eq ; 128 | half_eq_r := subst_half_eq Lambda |}. 129 | 130 | 131 | 132 | Definition Lambda_beta_Cat := INEQ_REP 133 | (S:=Lambda)(A:=unit)(fun x : unit => beta_rule). 134 | 135 | 136 | Definition Lambda_beta_SynSem := (INITIAL_INEQ_REP (fun x : unit => beta_rule)). 137 | 138 | Definition Lambda_beta := @Init _ _ _ (INITIAL_INEQ_REP (fun x : unit => beta_rule)). 139 | 140 | 141 | 142 | 143 | 144 | 145 | 146 | 147 | 148 | 149 | 150 | 151 | 152 | 153 | 154 | 155 | -------------------------------------------------------------------------------- /PROP_untyped/inequations_syntactically.v: -------------------------------------------------------------------------------- 1 | Require Export CatSem.PROP_untyped.prop_arities_initial_variant. 2 | 3 | Set Implicit Arguments. 4 | Unset Strict Implicit. 5 | Unset Automatic Introduction. 6 | Unset Transparent Obligations. 7 | 8 | Section type_of_inequations. 9 | 10 | Variable S : Signature. 11 | Print Signature. 12 | 13 | 14 | 15 | Inductive inequation : [nat] -> nat -> Type := 16 | | subst : inequation [[1;0]] 0. 17 | | ident : 18 | 19 | with 20 | inequation_product : 21 | 22 | Print inequation. -------------------------------------------------------------------------------- /PROP_untyped/representations.v: -------------------------------------------------------------------------------- 1 | Require Export CatSem.PROP_untyped.arities. 2 | 3 | Set Implicit Arguments. 4 | Unset Strict Implicit. 5 | Unset Transparent Obligations. 6 | Unset Automatic Introduction. 7 | 8 | (** ** Category of representations *) 9 | (** we define a category of representations of a signature Sig *) 10 | (** main work has been done in ./arities.v , remains to do: 11 | - identity representation morphism 12 | - composition of rep morphs 13 | *) 14 | 15 | (** Morphisms of representations are just monad morphisms which verify 16 | some property. It is hence sufficient to check those properties. 17 | No new data is defined here. *) 18 | 19 | Section cat_of_reps. 20 | 21 | Notation "[ T ]" := (list T) (at level 5). 22 | 23 | Variable S : Signature. 24 | 25 | (** ** Identity 26 | identity representation morphism *) 27 | 28 | Section id. 29 | 30 | Variable P : Representation S. 31 | 32 | Hint Extern 1 (CONSTR _ _ = CONSTR _ _) => apply CONSTR_eq. 33 | 34 | Lemma Prod_mor_c_id V i (x : prod_mod_c _ V (sig (s:= S) i)): 35 | Prod_mor_c (RMonad_id P) (l:=sig (s:=S) i) (V:=V) x = x. 36 | Proof. 37 | induction x; simpl; intros; auto. 38 | Qed. 39 | 40 | Hint Rewrite Prod_mor_c_id : bla. 41 | 42 | Obligation Tactic := unfold commute, commute_left, commute_right; 43 | simpl; intros; autorewrite with bla; auto. 44 | 45 | Hint Extern 1 (_ = _) => apply f_equal. 46 | 47 | Program Instance Rep_Id_struct : Representation_Hom_struct (RMonad_id P). 48 | 49 | Definition Rep_Id := Build_Representation_Hom Rep_Id_struct. 50 | 51 | End id. 52 | 53 | Hint Extern 1 (CONSTR _ _ = CONSTR _ _) => apply CONSTR_eq. 54 | 55 | (** ** Composition 56 | composition of rep homs, preparation *) 57 | 58 | Section comp_prepar. 59 | 60 | Variables P Q R : RMonad SM_po. 61 | Variable f : RMonad_Hom P Q. 62 | Variable g : RMonad_Hom Q R. 63 | 64 | Lemma prod_ind_mod_mor_comp l (V : TYPE) (t : prod_mod_c _ V l) : 65 | Prod_mor_c (RMonad_comp f g) t = Prod_mor_c g (Prod_mor_c f t). 66 | Proof. 67 | induction t; simpl; auto. 68 | Qed. 69 | 70 | Hint Rewrite prod_ind_mod_mor_comp : bla. 71 | 72 | Lemma comp_hophop l 73 | (MR : modhom_from_arity R (l)) 74 | (MP : modhom_from_arity P (l)) 75 | (MQ : modhom_from_arity Q (l)) 76 | (HMf : commute f MP MQ) 77 | (HMg : commute g MQ MR): 78 | commute (RMonad_comp f g) MP MR. 79 | Proof. 80 | intros; 81 | unfold commute, commute_left in *; 82 | simpl in *; intros; 83 | rerew_all; 84 | autorewrite with bla; auto. 85 | Qed. 86 | 87 | End comp_prepar. 88 | 89 | (** composition of rep homs *) 90 | 91 | Section comp. 92 | 93 | Variables P Q R : Representation S. 94 | Variable f : Representation_Hom P Q. 95 | Variable g : Representation_Hom Q R. 96 | 97 | Obligation Tactic := simpl; intros; 98 | apply comp_hophop with (repr Q _ ); 99 | match goal with 100 | [H:Representation_Hom _ _ |- _ ] => apply H end. 101 | 102 | Program Instance Rep_comp_struct : 103 | Representation_Hom_struct (RMonad_comp f g). 104 | 105 | Definition Rep_Comp := Build_Representation_Hom Rep_comp_struct. 106 | 107 | End comp. 108 | 109 | (** ** Equality of Representation Morphisms 110 | rep homs are equal if their resp carriers (monad homs) are *) 111 | 112 | Section Req_equiv. 113 | 114 | Variables P R : Representation S. 115 | 116 | Ltac equiv := match goal with 117 | | [|- Reflexive _ ] => unfold Reflexive; intro; 118 | apply Equivalence_Reflexive 119 | | [|- Symmetric _] => unfold Symmetric; do 2 intro; 120 | apply Equivalence_Symmetric 121 | | [|- Transitive _ ] => unfold Transitive; do 3 intro; 122 | apply Equivalence_Transitive end. 123 | 124 | Existing Instance RMonad_Hom_oid. 125 | 126 | Lemma eq_Rep_equiv : 127 | @Equivalence (Representation_Hom P R) 128 | (fun a c => repr_hom_c a == repr_hom_c c). 129 | Proof. 130 | constructor; equiv. 131 | Qed. 132 | 133 | Definition eq_Rep_oid := Build_Setoid (eq_Rep_equiv). 134 | 135 | End Req_equiv. 136 | 137 | Existing Instance RMONAD_struct. 138 | 139 | Obligation Tactic := simpl; intros; try unf_Proper; 140 | simpl; intros; 141 | repeat match goal with [H:_ |-_]=>rewrite H end; 142 | auto. 143 | 144 | (** ** Category of Representations *) 145 | 146 | Program Instance REP_struct : 147 | Cat_struct (@Representation_Hom S) := { 148 | mor_oid a c := eq_Rep_oid a c; 149 | id a := Rep_Id a; 150 | comp P Q R f g := Rep_Comp f g }. 151 | 152 | Definition REP := Build_Cat REP_struct. 153 | 154 | End cat_of_reps. 155 | 156 | 157 | 158 | 159 | 160 | 161 | 162 | 163 | 164 | -------------------------------------------------------------------------------- /README.md: -------------------------------------------------------------------------------- 1 | 2 | ## COMPILATION 3 | 4 | This Coq theory compiles under Coq 8.3pl5, available from https://coq.inria.fr/distrib/V8.3pl5/files/ . 5 | Earlier patch levels should also work; I have tested with 8.3pl2. 6 | 7 | Create a Makefile by calling 8 | ```bash 9 | $ coq_makefile -f Make > Makefile 10 | ``` 11 | and compile by calling 12 | ```bash 13 | $ make 14 | ``` 15 | 16 | WARNING : The compilation of some of the files consumes up to 2GB of memory. Make sure you dispose of sufficient reserves of ram before compiling the code. 17 | 18 | 19 | ## WORK WITH THE CODE 20 | 21 | Call coqide as follows from the root of the library: 22 | ```bash 23 | $ coqide -R . CatSem 24 | ``` 25 | 26 | ## CONTENT 27 | 28 | Read the file "./DESCRIPTION" for a description of the content of each file. 29 | 30 | ## BRANCHES 31 | 32 | Each branch below, printed in **boldface**, corresponds to an article, printed in _italic_. 33 | 34 | * [**JFR**](https://github.com/benediktahrens/monads/tree/JFR): [_Initial Semantics for higher-order typed syntax in Coq_](http://jfr.cib.unibo.it/article/view/2066) 35 | * [**MSCS**](https://github.com/benediktahrens/monads/tree/MSCS): [_Modules over relative monads for syntax and semantics_](http://dx.doi.org/10.1017/S0960129514000103) 36 | * [**REDUCTION_RULES**](https://github.com/benediktahrens/monads/tree/REDUCTION_RULES): [_Initial Semantics for Reduction Rules_](http://arxiv.org/abs/1212.5668) 37 | 38 | All the articles are also available from my [webpage](http://benedikt-ahrens.de/publications). 39 | 40 | 41 | -------------------------------------------------------------------------------- /RPCF/RPCF_INIT.v: -------------------------------------------------------------------------------- 1 | (*Require Import Coq.Logic.FunctionalExtensionality.*) 2 | Require Import Coq.Logic.Eqdep. 3 | 4 | Require Import CatSem.AXIOMS.functional_extensionality. 5 | 6 | Require Export CatSem.RPCF.RPCF_rep_cat. 7 | Require Export CatSem.RPCF.RPCF_syntax_init. 8 | 9 | Set Implicit Arguments. 10 | Unset Strict Implicit. 11 | Unset Transparent Obligations. 12 | Unset Automatic Introduction. 13 | 14 | 15 | Section unique. 16 | 17 | Variable R : REP. 18 | 19 | Variable g : PCFE_rep ---> R. 20 | 21 | Lemma initR_unique : g == initR R. 22 | Proof. 23 | simpl. 24 | assert (H : forall t, Sorts_map (initR R) t = Sorts_map g t). 25 | simpl. 26 | induction t; simpl; 27 | destruct g; simpl in *; 28 | try rerew_all; auto. 29 | repeat rew_all; auto. 30 | apply (eq_rep (H:=H)). 31 | 32 | intros V t y. 33 | destruct y as [t y]; 34 | simpl in *. 35 | (* generalize (H t). *) 36 | assert (H' : Init_Sorts_map R = Sorts_map g). 37 | apply functional_extensionality; intros. 38 | destruct g. 39 | rew_all. auto. 40 | destruct g as [f garrow gnat gbool gM gMs]; 41 | simpl in *. 42 | generalize (H t). 43 | generalize H. 44 | generalize dependent gM. 45 | clear H. 46 | generalize dependent gnat. 47 | generalize dependent gbool. 48 | generalize dependent garrow. 49 | rewrite <- H'. 50 | intros; simpl. 51 | rewrite (UIP_refl _ _ _ ). 52 | simpl. 53 | clear g. 54 | assert (Hl := rlift_transp_id (Q:=R) H). 55 | simpl in *. 56 | rewrite Hl. 57 | clear Hl. 58 | clear e H. 59 | induction y; 60 | simpl. 61 | 62 | assert (Hb:=bottom_hom (PCFPO_rep_Hom_struct := gMs) 63 | t (c:=V) tt). 64 | simpl in *. 65 | auto. 66 | 67 | destruct c. 68 | 69 | assert (Hb:=nats_hom n (PCFPO_rep_Hom_struct := 70 | gMs) (c:=V) tt). 71 | simpl in *. 72 | rewrite (UIP_refl _ _ _) in Hb. 73 | simpl in * ; auto. 74 | 75 | assert (Hb:=ttt_hom (PCFPO_rep_Hom_struct := 76 | gMs) (c:=V) tt). 77 | simpl in *. 78 | rewrite (UIP_refl _ _ _) in Hb. 79 | simpl in * ; auto. 80 | 81 | assert (Hb:=fff_hom (PCFPO_rep_Hom_struct := 82 | gMs) (c:=V) tt). 83 | simpl in *. 84 | rewrite (UIP_refl _ _ _) in Hb. 85 | simpl in * ; auto. 86 | 87 | assert (Hb:=Succ_hom (PCFPO_rep_Hom_struct := 88 | gMs) (c:=V) tt). 89 | simpl in *. 90 | rewrite (UIP_refl _ _ _ ) in Hb. 91 | simpl in Hb. 92 | auto. 93 | 94 | assert (Hb:=Pred_hom (PCFPO_rep_Hom_struct := 95 | gMs) (c:=V) tt). 96 | simpl in *. 97 | rewrite (UIP_refl _ _ _ ) in Hb. 98 | simpl in Hb. 99 | auto. 100 | 101 | assert (Hb:=Zero_hom (PCFPO_rep_Hom_struct := 102 | gMs) (c:=V) tt). 103 | simpl in *. 104 | rewrite (UIP_refl _ _ _ ) in Hb. 105 | simpl in Hb. 106 | auto. 107 | 108 | assert (Hb:=CondN_hom (PCFPO_rep_Hom_struct := 109 | gMs) (c:=V) tt). 110 | simpl in *. 111 | rewrite (UIP_refl _ _ _ ) in Hb. 112 | simpl in Hb. 113 | auto. 114 | 115 | assert (Hb:=CondB_hom (PCFPO_rep_Hom_struct := 116 | gMs) (c:=V) tt). 117 | simpl in *. 118 | rewrite (UIP_refl _ _ _ ) in Hb. 119 | simpl in Hb. 120 | auto. 121 | 122 | (*1*) 123 | assert (Hw:=gen_rmonad_hom_rweta 124 | ((*gen_RMonad_Hom_struct := *)gM)). 125 | simpl in Hw. 126 | assert (Hw' := Hw _ _ (ctype _ v)). 127 | simpl in Hw'. 128 | auto. 129 | (*2*) 130 | simpl. 131 | rewrite <- IHy1. 132 | rewrite <- IHy2. 133 | assert (Happ:=app_hom (PCFPO_rep_Hom_struct := 134 | gMs) (u:=s) (v:=t) (y1,y2)). 135 | simpl in *. 136 | rewrite <- Happ. 137 | clear Happ. 138 | apply f_equal. 139 | simpl. 140 | apply injective_projections; 141 | simpl; 142 | auto. 143 | rewrite (UIP_refl _ _ (garrow s t)). 144 | auto. 145 | 146 | (*3*) 147 | simpl. 148 | rewrite <- IHy. 149 | clear IHy. 150 | assert (Habs:=abs_hom (PCFPO_rep_Hom_struct := 151 | gMs) (u:=t) (v:=s) y ). 152 | simpl in *. 153 | rewrite Habs. 154 | rewrite (UIP_refl _ _ _ ). 155 | reflexivity. 156 | 157 | simpl. 158 | rewrite <- IHy. 159 | assert (Happ:=rec_hom (PCFPO_rep_Hom_struct := 160 | gMs) (t:=t) y). 161 | simpl in *. 162 | rewrite Happ. 163 | rewrite (UIP_refl _ _ _ ). 164 | reflexivity. 165 | Qed. 166 | 167 | End unique. 168 | 169 | Hint Resolve initR_unique : fin. 170 | 171 | Program Instance PCF_initial : Initial REP := { 172 | Init := PCFE_rep ; 173 | InitMor R := initR R ; 174 | InitMorUnique R := @initR_unique R 175 | }. 176 | 177 | (* Print Assumptions PCF_initial. *) 178 | -------------------------------------------------------------------------------- /RPCF/RPCF_rep_eq.v: -------------------------------------------------------------------------------- 1 | 2 | (*Require Import Coq.Logic.FunctionalExtensionality.*) 3 | Require Import Coq.Logic.Eqdep. 4 | 5 | Require Import CatSem.AXIOMS.functional_extensionality. 6 | 7 | Require Export CatSem.RPCF.RPCF_rep_hom. 8 | Require Import CatSem.CAT.rmonad_gen_comp. 9 | 10 | Set Implicit Arguments. 11 | Unset Strict Implicit. 12 | Unset Transparent Obligations. 13 | Unset Automatic Introduction. 14 | 15 | 16 | Inductive eq_Rep (P R : PCFPO_rep) : relation (PCFPO_rep_Hom P R) := 17 | | eq_rep : forall (a c : PCFPO_rep_Hom P R), 18 | forall H : (forall t, Sorts_map c t = Sorts_map a t), 19 | (forall V, (*rep_Hom_monad*) a V ;; rlift R (Transp H V) 20 | == 21 | Transp_ord H (P V) ;; (*rep_Hom_monad*) c V ) -> 22 | 23 | eq_Rep a c. 24 | 25 | (** the defined relation is an equivalence and 26 | may hence serve as equality 27 | *) 28 | 29 | Lemma eq_Rep_equiv (P R: PCFPO_rep) : 30 | @Equivalence (PCFPO_rep_Hom P R) 31 | (@eq_Rep P R). 32 | Proof. 33 | intros P R. 34 | split. 35 | 36 | unfold Reflexive. 37 | intro M. 38 | assert (H': forall t, Sorts_map M t = Sorts_map M t) by 39 | (intros; reflexivity). 40 | 41 | apply (eq_rep (H:=H')). 42 | 43 | simpl. 44 | intros V t y. 45 | destruct y as [t y]. 46 | 47 | simpl. 48 | rewrite (UIP_refl _ _ (H' t)). 49 | simpl. 50 | assert (H:= rlift_transp_id). 51 | simpl in *. 52 | rewrite H. 53 | auto. 54 | 55 | (* now symmetric *) 56 | 57 | unfold Symmetric. 58 | intros M N H. 59 | destruct H. 60 | assert (H': forall t, Sorts_map a t = Sorts_map c t) by auto. 61 | apply (eq_rep (H:=H')). 62 | simpl; intros V t y. 63 | destruct a. 64 | destruct c. 65 | simpl in *. 66 | assert (H3 : Sorts_map = Sorts_map0). 67 | apply (functional_extensionality). 68 | auto. 69 | 70 | generalize dependent y. 71 | generalize dependent H'. 72 | generalize dependent rep_Hom_monad0. 73 | generalize dependent rep_Hom_monad. 74 | generalize dependent HNat. 75 | generalize dependent HNat0. 76 | generalize dependent HBool. 77 | generalize dependent HBool0. 78 | generalize dependent H. 79 | generalize dependent HArrow. 80 | rewrite H3. 81 | intros; simpl in *. 82 | rewrite transp_id. 83 | assert (H2:= rlift_transp_id). 84 | simpl in H2. 85 | rewrite H2. 86 | assert (H4 := H0 V t y). 87 | rewrite transp_id in H4. 88 | rewrite H2 in H4. 89 | simpl in *; auto. 90 | 91 | (* finally transitive *) 92 | 93 | unfold Transitive. 94 | intros a b c H H'. 95 | destruct H; 96 | destruct H'. 97 | assert (Ht : forall t, Sorts_map c t = Sorts_map a t). 98 | intro t. transitivity (Sorts_map a0 t); auto. 99 | 100 | apply (eq_rep (H:=Ht)). 101 | simpl; intros V t y. 102 | destruct a; 103 | destruct a0; 104 | destruct c. 105 | simpl in *. 106 | assert (H5 : Sorts_map0 = Sorts_map1) by 107 | (apply functional_extensionality; intro; auto). 108 | 109 | assert (H6 : Sorts_map1 = Sorts_map) by 110 | (apply functional_extensionality; intro; auto). 111 | 112 | generalize dependent H2. 113 | generalize dependent H1. 114 | generalize dependent rep_Hom_monad. 115 | generalize dependent rep_Hom_monad1. 116 | generalize dependent rep_Hom_monad0. 117 | generalize dependent HBool. 118 | generalize dependent HBool0. 119 | generalize dependent HBool1. 120 | generalize dependent HNat. 121 | generalize dependent HNat0. 122 | generalize dependent HNat1. 123 | 124 | generalize dependent H. 125 | generalize dependent Ht. 126 | generalize dependent HArrow. 127 | generalize dependent HArrow0. 128 | generalize dependent HArrow1. 129 | rewrite H5. 130 | rewrite H6. 131 | 132 | clear H6 H5. 133 | clear Sorts_map0. 134 | clear Sorts_map1. 135 | 136 | intros. 137 | assert (H7:=H0 V t y). 138 | assert (H9:=H2 V t y). 139 | rewrite transp_id in *. 140 | simpl in *. 141 | assert (H8 := rlift_transp_id). 142 | simpl in H8. 143 | rewrite H8 in *. 144 | transitivity (rep_Hom_monad0 V t y); 145 | auto. 146 | Qed. 147 | 148 | Definition eq_Rep_oid (P R : PCFPO_rep) := Build_Setoid (eq_Rep_equiv P R). 149 | 150 | -------------------------------------------------------------------------------- /RPCF/RPCF_syntax_rep.v: -------------------------------------------------------------------------------- 1 | 2 | Require Export CatSem.PCF.PCF_RMonad. 3 | Require Export CatSem.RPCF.RPCF_rep. 4 | 5 | Set Implicit Arguments. 6 | Unset Strict Implicit. 7 | Unset Transparent Obligations. 8 | Unset Automatic Introduction. 9 | 10 | Obligation Tactic := simpl; intros; 11 | unfold Rsubst_star_map; simpl; 12 | apply clos_refl_trans_1n_contains; 13 | apply relorig; 14 | try apply app_abs; 15 | try constructor. 16 | 17 | Program Instance PCFE_rep_struct : 18 | PCFPO_rep_struct PCFEM PCF.Arrow PCF.Bool PCF.Nat := { 19 | app r s := PCFApp r s; 20 | abs r s := PCFAbs r s; 21 | rec t := PCFRec t ; 22 | tttt := PCFconsts ttt ; 23 | ffff := PCFconsts fff; 24 | Succ := PCFconsts succ; 25 | Pred := PCFconsts preds; 26 | CondN := PCFconsts condN; 27 | CondB := PCFconsts condB; 28 | Zero := PCFconsts zero ; 29 | nats m := PCFconsts (Nats m); 30 | bottom t := PCFbottom t 31 | }. 32 | (* 33 | Next Obligation. 34 | Proof. 35 | unfold Rsubst_star_map. 36 | simpl. 37 | apply clos_refl_trans_1n_contains. 38 | apply relorig. 39 | apply app_abs. 40 | Qed. 41 | (* 42 | Next Obligation. 43 | Proof. 44 | apply cp_App1. 45 | auto. 46 | Qed. 47 | Next Obligation. 48 | Proof. 49 | apply cp_App2. 50 | auto. 51 | Qed. 52 | Next Obligation. 53 | Proof. 54 | apply cp_Lam. 55 | auto. 56 | Qed. 57 | Next Obligation. 58 | Proof. 59 | apply cp_Rec. 60 | auto. 61 | Qed. 62 | *) 63 | Next Obligation. 64 | Proof. 65 | apply clos_refl_trans_1n_contains. 66 | apply relorig. 67 | constructor. 68 | Qed. 69 | Next Obligation. 70 | Proof. 71 | apply clos_refl_trans_1n_contains. 72 | apply relorig. 73 | constructor. 74 | Qed. 75 | Next Obligation. 76 | Proof. 77 | apply clos_refl_trans_1n_contains. 78 | apply relorig. 79 | constructor. 80 | Qed. 81 | Next Obligation. 82 | Proof. 83 | apply clos_refl_trans_1n_contains. 84 | apply relorig. 85 | constructor. 86 | Qed. 87 | Next Obligation. 88 | Proof. 89 | apply clos_refl_trans_1n_contains. 90 | apply relorig. 91 | constructor. 92 | Qed. 93 | Next Obligation. 94 | Proof. 95 | apply clos_refl_trans_1n_contains. 96 | apply relorig. 97 | constructor. 98 | Qed. 99 | Next Obligation. 100 | Proof. 101 | apply clos_refl_trans_1n_contains. 102 | apply relorig. 103 | constructor. 104 | Qed. 105 | Next Obligation. 106 | Proof. 107 | apply clos_refl_trans_1n_contains. 108 | apply relorig. 109 | constructor. 110 | Qed. 111 | Next Obligation. 112 | Proof. 113 | apply clos_refl_trans_1n_contains. 114 | apply relorig. 115 | constructor. 116 | Qed. 117 | Next Obligation. 118 | Proof. 119 | apply clos_refl_trans_1n_contains. 120 | apply relorig. 121 | constructor. 122 | Qed. 123 | *) 124 | 125 | Definition PCFE_rep : PCFPO_rep := 126 | Build_PCFPO_rep (Sorts:=PCF.Sorts) (Arrow:=PCF.Arrow) 127 | (pcf_rep_monad:=PCFEM) 128 | PCFE_rep_struct. 129 | 130 | 131 | 132 | 133 | 134 | 135 | 136 | 137 | 138 | 139 | 140 | -------------------------------------------------------------------------------- /RPCF/ULC_comp.v: -------------------------------------------------------------------------------- 1 | Require Export CatSem.RPCF.RPCF_INIT. 2 | Require Export CatSem.RPCF.RPCF_ULC_nounit. 3 | 4 | Set Implicit Arguments. 5 | Unset Strict Implicit. 6 | (*Unset Printing Implicit Defensive.*) 7 | Unset Automatic Introduction. 8 | 9 | 10 | 11 | Definition PCF_ULC_compilation := 12 | InitMor PCF_ULC. 13 | 14 | Definition PCF_ULC_c := 15 | rep_Hom_monad PCF_ULC_compilation. 16 | Locate "@". 17 | (** some examples *) 18 | 19 | Notation "a @ b" := (App a b) (at level 20, left associativity). 20 | Notation "a @@ b" := (PCF_syntax.App a b) (at level 20, left associativity). 21 | Notation "'1'" := (Var None) (at level 33). 22 | Notation "'2'" := (Var (Some None)) (at level 24). 23 | Notation "'3'" := (Var (Some (Some None))) (at level 23). 24 | Notation "'4'" := (Var (Some (Some (Some None)))) (at level 22). 25 | 26 | Notation "'y'" := (PCF_syntax.Var None) (at level 33). 27 | 28 | Eval compute in 29 | (PCF_ULC_c (fun t => False) tt 30 | (ctype _ (Bottom (fun _ => False) PCF.Nat))). 31 | Eval compute in 32 | (PCF_ULC_c (fun t => False) tt 33 | (ctype _ (succ ' ))). 34 | 35 | Eval compute in 36 | (PCF_ULC_c (fun t => False) tt 37 | (ctype _ (Const (fun _ => False) (succ )))). 38 | Eval compute in 39 | (PCF_ULC_c (fun t => False) tt 40 | (ctype _ (preds '))). 41 | 42 | Eval compute in 43 | (PCF_ULC_c (fun t => False) tt 44 | (ctype _ (condB '))). 45 | 46 | Eval compute in 47 | (PCF_ULC_c (fun t => False) tt (ctype _ 48 | (condB ' @@ ttt ' @@ fff ' @@ ttt '))). 49 | 50 | Eval compute in 51 | (PCF_ULC_c (opt PCF.Bool (fun t => False)) tt (ctype _ 52 | (condB ' @@ (PCF_syntax.Var (none PCF.Bool (fun t => False))) @@ fff ' @@ ttt '))). 53 | 54 | Notation "'x_bool'" := (PCF_syntax.Var (none PCF.Bool (fun t => False))). 55 | 56 | Eval compute in 57 | (PCF_ULC_c ((fun t => False)) tt (ctype _ 58 | (Lam (condB ' @@ (PCF_syntax.Var (none PCF.Bool (fun t => False))) @@ fff ' @@ ttt ')))). 59 | 60 | Eval compute in 61 | (PCF_ULC_c ((fun t => False)) tt (ctype _ 62 | (Lam (condB ' @@ x_bool @@ fff ' @@ ttt ')))). 63 | 64 | Print Assumptions PCF_initial. 65 | 66 | 67 | Check Lam. 68 | (* 69 | Eval compute in 70 | (PCF_ULC_c (fun t => False) tt (ctype _ 71 | (Lam (condB ' @@ PCF_syntax.Var None @@ fff ' @@ ttt ')))). 72 | *) 73 | Eval compute in 74 | (PCF_ULC_c ( (fun t => False)) tt 75 | (ctype _ ( 76 | Lam (PCF_syntax.Var (none PCF.Bool (fun t => False)))))). 77 | 78 | Check ctype. 79 | Check (Const (fun _ => False) condB). 80 | 81 | Check (fun V t (a : PCF V t) => PCF_ULC_c (V) tt 82 | (ctype (fun t => tt) a)). 83 | 84 | Require Import Coq.Program.Equality. 85 | 86 | Lemma aaaa V t (a : PCF V t) b : 87 | (PCF_ULC_c _ _ (ctype (fun t => tt) a)) >>> b -> 88 | (exists b', a >> b' /\ PCF_ULC_c _ _ (ctype _ b') = b). 89 | intros V t a. 90 | intro b. 91 | dependent induction b; 92 | simpl; intros. 93 | inversion H. 94 | exists (PCF_syntax.Var (ct ). 95 | exists (Bottom V t). 96 | split. reflexivity. 97 | simpl. 98 | inversion H. auto. 99 | subst. clear H. 100 | inversion H0. 101 | subst. 102 | exists (Bottom V t). 103 | split. 104 | reflexivity. 105 | simpl; auto. 106 | subst. 107 | exists (Bottom V t). 108 | split. reflexivity. 109 | ssiauto. 110 | 111 | destruct H. 112 | intros. 113 | elim H. 114 | clear H. 115 | intros. 116 | simpl in *. 117 | inversion H. 118 | subst. 119 | exists a. 120 | reflexivity. 121 | subst. 122 | 123 | 124 | 125 | dependent induction H. 126 | 127 | 128 | 129 | -------------------------------------------------------------------------------- /STS/STS_ex.v: -------------------------------------------------------------------------------- 1 | Require Export CatSem.STS.STS_initial. 2 | Require Export CatSem.TLC.TLC_types. 3 | Require Export CatSem.PCF.PCF_types. 4 | 5 | Set Implicit Arguments. 6 | Unset Strict Implicit. 7 | Unset Transparent Obligations. 8 | Unset Automatic Introduction. 9 | 10 | Section example_signatures. 11 | 12 | (** two examples of signatures : 13 | - simply typed lambda calculus 14 | - a fragment of PCF (just to make it shorter) 15 | *) 16 | Notation "'T'" := TLCtypes. 17 | Notation "s ~> t" := (TLCarrow s t)(at level 40). 18 | Notation "[ S ]" := (list S) (at level 5). 19 | 20 | Inductive TLC_index : T -> Type := 21 | | TLC_abs : forall s t : T, TLC_index (s ~> t) 22 | | TLC_app : forall s t : T, TLC_index t. 23 | 24 | Definition TLC_arguments : forall t, TLC_index t -> [[T] * T] := 25 | fun t r => match r with 26 | | TLC_abs u v => (u::nil,v)::nil 27 | | TLC_app u v => (nil,u ~> v)::(nil,u)::nil 28 | end. 29 | 30 | Definition TLC_sig t := Build_Signature_t t 31 | (@TLC_arguments t). 32 | 33 | (* 34 | Definition TLC_sig t := Build_Signature_t t 35 | (fun r : TLC_index t => match r with 36 | | TLC_abs s t => (s::nil,t)::nil 37 | | TLC_app s t => (nil,s ~> t)::(nil,s)::nil 38 | end). 39 | *) 40 | 41 | Notation "'Bool'" := PCF.Bool. 42 | Notation "'Nat'" := PCF.Nat. 43 | Notation "'T'" := PCF.Sorts. 44 | Notation "a ~> b" := (PCF.Arrow a b)(at level 40). 45 | 46 | Inductive PCF_index_t : T -> Type := 47 | | PCF_abs : forall s t : T, PCF_index_t (s ~> t) 48 | | PCF_app : forall s t : T, PCF_index_t t 49 | | PCF_rec : forall t : T, PCF_index_t t 50 | | PCF_true : PCF_index_t Bool 51 | | PCF_false : PCF_index_t Bool 52 | | PCF_nat : forall n : nat, PCF_index_t Nat 53 | | PCF_succ : PCF_index_t (Nat ~> Nat) 54 | | PCF_zero : PCF_index_t (Nat ~> Bool) 55 | | PCF_condN : PCF_index_t (Bool ~> Nat ~> Nat ~> Nat) 56 | | PCF_condB : PCF_index_t (Bool ~> Bool ~> Bool ~> Bool) 57 | | PCF_bottom : forall t : T, PCF_index_t t. 58 | 59 | Definition PCF_sig t := Build_Signature_t t 60 | (fun r : PCF_index_t t => match r in PCF_index_t t with 61 | | PCF_abs s t => (s::nil,t)::nil 62 | | PCF_app s t => (nil, s ~> t)::(nil,s)::nil 63 | | PCF_rec t => (nil, t ~> t)::nil 64 | | _ => nil 65 | end). 66 | 67 | End example_signatures. 68 | 69 | 70 | (** we are interested in TLC *) 71 | 72 | (** short name for its types *) 73 | Definition T := TLCtypes. 74 | 75 | Notation "a ~> b" := (TLCarrow a b) (at level 60). 76 | 77 | 78 | (** the category of reps of TLC_sig *) 79 | Definition TLCreps : Cat := REPRESENTATION TLC_sig. 80 | 81 | (** initiality of TLCreps *) 82 | Definition TLCInitial : Representation TLC_sig := 83 | Init (Initial := STS_initial TLC_sig). 84 | 85 | (** we isolate the monad *) 86 | Definition TLC : Monad (ITYPE T) := rep_monad TLCInitial. 87 | 88 | (** we want to have a short name for its simultaneous substitution *) 89 | Definition bind V W (f : V ---> TLC W) t x 90 | := kleisli f t x. 91 | 92 | (** and the substitution of a single variable, for beta *) 93 | Definition TLCsubstar r s V (M : TLC (opt r V) s) 94 | (N : TLC V r) : TLC V s := 95 | substar (S:=TLC_sig) N M. 96 | 97 | (** and a nice notation *) 98 | Notation "x >>= f" := (bind f x) (at level 60). 99 | Notation "M [*:= N ]" := (TLCsubstar M N) (at level 40). 100 | 101 | Definition app r s V (M : TLC V (r ~> s)) (N : TLC V r) : TLC V s := 102 | Build (S:=TLC_sig) (i:=TLC_app r s) 103 | (constr (b:=(nil,r~>s))(S:=TLC_sig) M 104 | (constr (b:=(nil,r)) N (TT (T:=T) TLC_sig V))). 105 | 106 | Definition abs r s V (M : TLC (opt r V) s) : TLC V (r ~> s) := 107 | Build (S:=TLC_sig) (t:=r ~> s) (i:=TLC_abs r s) 108 | (constr (b:=(r :: nil, s)) M (TT (T:=T) TLC_sig V)). 109 | 110 | Notation "M @ N" := (app M N) (at level 30). 111 | Notation "'\*' M" := (abs M) (at level 32). 112 | 113 | Reserved Notation "x >> y" (at level 50). 114 | 115 | Inductive beta : forall V t, relation (TLC V t) := 116 | | app_abs : forall V r s (M : TLC (opt r V) s) N, 117 | (\* M) @ N >> M [*:= N] 118 | | beta_app1 : forall V r s (M M' : TLC V (r ~> s)) N, 119 | M >> M' -> M @ N >> M' @ N 120 | | beta_app2 : forall V r s (M : TLC V (r ~> s)) N N', 121 | N >> N' -> M @ N >> M @ N' 122 | | beta_abs : forall V r s (M M': TLC (opt r V) s), 123 | M >> M' -> \* M >> \* M' 124 | 125 | where "x >> y" := (beta x y). 126 | 127 | 128 | (** we could now start to define reduction relations etc. *) 129 | 130 | -------------------------------------------------------------------------------- /STS/STS_representations.v: -------------------------------------------------------------------------------- 1 | Require Export CatSem.STS.STS_arities. 2 | 3 | Set Implicit Arguments. 4 | Unset Strict Implicit. 5 | Unset Transparent Obligations. 6 | Unset Automatic Introduction. 7 | 8 | 9 | (** we define a category of representations of a signature Sig over T *) 10 | (** main work has been done in STS_arities, remains to do: 11 | - identity representation morphism 12 | - composition of rep morphs 13 | *) 14 | 15 | Section cat_of_reps. 16 | 17 | Notation "[ T ]" := (list T) (at level 5). 18 | 19 | Variable T : Type. 20 | 21 | Variable S : Signature T. 22 | 23 | (** identity representation morphism *) 24 | 25 | Section id. 26 | 27 | Variable P : Representation S. 28 | 29 | Hint Extern 1 (CONSTR _ _ = CONSTR _ _) => apply CONSTR_eq. 30 | 31 | Lemma Prod_mor_c_id V t i (x : prod_mod_c P V (sig (s:=S t) i)): 32 | Prod_mor_c (Monad_Hom_id P) (l:=sig (s:=S t) i) (V:=V) x = x. 33 | Proof. 34 | induction x; simpl; intros; auto. 35 | Qed. 36 | 37 | Hint Rewrite Prod_mor_c_id : bla. 38 | 39 | Obligation Tactic := unfold commute, commute_left, commute_right; 40 | simpl; intros; autorewrite with bla; auto. 41 | 42 | Hint Extern 1 (_ = _) => apply f_equal. 43 | 44 | Program Instance Rep_Id_struct : 45 | Representation_Hom_struct (Monad_Hom_id P). 46 | 47 | Definition Rep_Id := Build_Representation_Hom Rep_Id_struct. 48 | 49 | End id. 50 | 51 | Hint Extern 1 (CONSTR _ _ = CONSTR _ _) => apply CONSTR_eq. 52 | 53 | (** composition of rep homs, preparation *) 54 | 55 | Section comp_prepar. 56 | 57 | Variables P Q R : Monad (ITYPE T). 58 | Variable f : Monad_Hom P Q. 59 | Variable g : Monad_Hom Q R. 60 | 61 | Lemma prod_ind_mod_mor_comp l 62 | (c : T -> Type) 63 | (t : prod_mod_c P c l) : 64 | Prod_mor_c (Monad_Hom_comp f g) t = 65 | Prod_mor_c g (Prod_mor_c f t). 66 | Proof. 67 | induction t; simpl; auto. 68 | Qed. 69 | 70 | Hint Rewrite prod_ind_mod_mor_comp : bla. 71 | 72 | Lemma comp_hophop l a 73 | (MR : modhom_from_arity R (a,l)) 74 | (MP : modhom_from_arity P (a,l)) 75 | (MQ : modhom_from_arity Q (a,l)) 76 | (HMf : commute f MP MQ) 77 | (HMg : commute g MQ MR): 78 | commute (Monad_Hom_comp f g) MP MR. 79 | Proof. 80 | intros; 81 | unfold commute, commute_left in *; 82 | simpl in *; intros. 83 | rerew_all. 84 | autorewrite with bla; auto. 85 | Qed. 86 | 87 | End comp_prepar. 88 | 89 | (** composition of rep homs *) 90 | 91 | Section comp. 92 | 93 | Variables P Q R : Representation S. 94 | Variable f : Representation_Hom P Q. 95 | Variable g : Representation_Hom Q R. 96 | 97 | Obligation Tactic := simpl; intros; 98 | apply comp_hophop with (repr Q _ ); 99 | match goal with 100 | [H:Representation_Hom _ _ |- _ ] => apply H end. 101 | 102 | Program Instance Rep_comp_struct : 103 | Representation_Hom_struct (Monad_Hom_comp f g). 104 | 105 | Definition Rep_Comp := Build_Representation_Hom Rep_comp_struct. 106 | 107 | End comp. 108 | 109 | (** rep homs are equal if their resp carriers (monad homs) are *) 110 | 111 | Section Req_equiv. 112 | 113 | Variables P R : Representation S. 114 | 115 | Ltac equiv := match goal with 116 | | [|- Reflexive _ ] => unfold Reflexive; intro; 117 | apply Equivalence_Reflexive 118 | | [|- Symmetric _] => unfold Symmetric; do 2 intro; 119 | apply Equivalence_Symmetric 120 | | [|- Transitive _ ] => unfold Transitive; do 3 intro; 121 | apply Equivalence_Transitive end. 122 | 123 | Lemma eq_Rep_equiv : 124 | @Equivalence (Representation_Hom P R) 125 | (fun a c => repr_hom_c a == repr_hom_c c). 126 | Proof. 127 | constructor; equiv. 128 | Qed. 129 | 130 | Definition eq_Rep_oid := Build_Setoid (eq_Rep_equiv). 131 | 132 | End Req_equiv. 133 | 134 | Existing Instance MONAD_struct. 135 | 136 | Obligation Tactic := simpl; intros; try unf_Proper; 137 | simpl; intros; 138 | repeat match goal with [H:_ |-_]=>rewrite H end; 139 | auto. 140 | 141 | Program Instance REPRESENTATION_struct : 142 | Cat_struct (@Representation_Hom _ S) := { 143 | mor_oid a c := eq_Rep_oid a c; 144 | id a := Rep_Id a; 145 | comp P Q R f g := Rep_Comp f g 146 | }. 147 | 148 | Definition REPRESENTATION := Build_Cat REPRESENTATION_struct. 149 | 150 | End cat_of_reps. 151 | 152 | 153 | 154 | 155 | 156 | 157 | 158 | 159 | 160 | -------------------------------------------------------------------------------- /SYS_F/F_types.v: -------------------------------------------------------------------------------- 1 | Require Export CatSem.CAT.monad_haskell. 2 | Require Export CatSem.CAT.orders. 3 | Require Import Coq.Relations.Relations. 4 | 5 | Set Implicit Arguments. 6 | Unset Strict Implicit. 7 | 8 | Inductive F (V : Ord) : Type := 9 | | Var : V -> F V 10 | | Ar : F V -> F V -> F V. 11 | 12 | 13 | Inductive FPO V : relation (F V) := 14 | | arp1 : forall s s' t, FPO s s' -> FPO (Ar s' t) (Ar s t) 15 | | arp2 : forall s t t', FPO t t' -> FPO (Ar s t) (Ar s t') 16 | (* | arp : forall s t s' t', FPO s s' -> FPO t t' -> FPO (Ar s' t) (Ar s t') *) 17 | | varp: forall v v', v << v' -> FPO (Var v) (Var v'). 18 | 19 | Definition FP (V : Ord) := Clos_RT_1n_prf ((@FPO V)). 20 | 21 | 22 | Instance FPP (V : Ord) : PO_obj_struct (F V) := { 23 | POprf := FP V 24 | }. 25 | 26 | Definition FF (V : Ord) : Ord := Build_PO_obj (FPP V). 27 | 28 | Lemma Ar2 V (s t t' : FF V) : t << t' -> (Ar s t) << (Ar s t'). 29 | Proof. 30 | intro. 31 | generalize dependent s. 32 | induction H. 33 | reflexivity. 34 | intros. 35 | constructor 2 with (Ar s y). 36 | Focus 2. apply IHclos_refl_trans_1n. 37 | constructor. 38 | auto. 39 | Qed. 40 | 41 | Lemma nnn V (x z : F V): 42 | clos_refl_trans_1n (F V) (@FPO V) z x -> 43 | clos_refl_trans_n1 (F V) (@FPO V) z x. 44 | Proof. 45 | intro H. 46 | apply clos_rt_rtn1. 47 | apply clos_rt1n_rt. 48 | auto. 49 | Qed. 50 | 51 | Lemma kkk V (x z : F V): 52 | clos_refl_trans_n1 (F V) (@FPO V) z x -> 53 | clos_refl_trans_1n (F V) (@FPO V) z x. 54 | Proof. 55 | intro H. 56 | apply clos_rt_rt1n. 57 | apply clos_rtn1_rt. 58 | auto. 59 | Qed. 60 | 61 | 62 | Lemma Ar1 V (s s' t : FF V) : s' << s -> (Ar s t) << (Ar s' t). 63 | Proof. 64 | intro. 65 | generalize dependent t. 66 | induction H. 67 | reflexivity. 68 | intros. 69 | simpl. 70 | apply kkk. 71 | constructor 2 with (Ar y t). 72 | constructor. 73 | auto. 74 | apply nnn. 75 | apply IHclos_refl_trans_1n. 76 | Qed. 77 | 78 | 79 | Lemma bla : forall V (s t s' t' : FF V) , s << s' -> t << t' -> 80 | Ar s' t << Ar s t'. 81 | Proof. 82 | intros. 83 | transitivity (Ar s' t'). 84 | apply Ar2. 85 | auto. 86 | apply Ar1. 87 | auto. 88 | Qed. 89 | 90 | 91 | Fixpoint subst (V W : Ord) (f : V ---> FF W) (x : FF V) : FF W := 92 | match x with 93 | | Var v => f v 94 | | Ar s t => Ar (subst f s) (subst f t) 95 | end . 96 | 97 | Program Instance _subst V W (f : V ---> FF W) : PO_mor_struct (subst f). 98 | Next Obligation. 99 | Proof. 100 | unfold Proper; red. 101 | intros. 102 | induction H. 103 | reflexivity. 104 | transitivity (subst f y); 105 | auto. 106 | clear dependent z. 107 | induction H; simpl. 108 | apply Ar1. 109 | auto. 110 | apply Ar2. 111 | auto. 112 | apply f. 113 | auto. 114 | Qed. 115 | 116 | Definition subst_po V W (f : V ---> FF W) := Build_PO_mor (_subst V W f). 117 | 118 | Program Instance var V : PO_mor_struct (b:=FF V) (@Var V). 119 | Next Obligation. 120 | Proof. 121 | unfold Proper; red. 122 | intros. 123 | apply clos_refl_trans_1n_contains. 124 | constructor. 125 | auto. 126 | Qed. 127 | 128 | Definition VAR V := Build_PO_mor (var V). 129 | 130 | Program Instance FFMonad : Monad_struct FF := { 131 | weta := VAR; 132 | kleisli := subst_po 133 | }. 134 | Next Obligation. 135 | Proof. 136 | unfold Proper; red. 137 | simpl. 138 | intros f g H x. 139 | induction x; simpl. 140 | auto. 141 | rewrite IHx1. 142 | rewrite IHx2. 143 | auto. 144 | Qed. 145 | Next Obligation. 146 | Proof. 147 | induction x; 148 | simpl. 149 | auto. 150 | rewrite IHx2. 151 | rewrite IHx1. 152 | auto. 153 | Qed. 154 | Next Obligation. 155 | Proof. 156 | induction x; 157 | simpl. 158 | auto. 159 | rewrite IHx1. 160 | rewrite IHx2. 161 | auto. 162 | Qed. 163 | 164 | (* 165 | Definition subst V W f : FF V ---> FF W 166 | 167 | Focus 2. 168 | apply f. 169 | auto. 170 | 171 | simpl. 172 | 173 | transitivit 174 | 175 | apply clos_refl_trans_1n_contains. 176 | simpl. 177 | apply 178 | apply clos_refl_trans_1n_contains. 179 | induction H. 180 | simpl in *. 181 | constructor; 182 | auto. 183 | simpl. 184 | apply PO_mor_monotone. 185 | apply f. 186 | apply 187 | constructor. 188 | apply clos_refl_trans_1n_contains. 189 | constructor. 190 | app 191 | simpl. 192 | apply reflexivity. 193 | apply PO_refl. 194 | simpl. 195 | 196 | *) 197 | 198 | 199 | 200 | 201 | 202 | 203 | 204 | 205 | 206 | 207 | 208 | 209 | 210 | 211 | 212 | 213 | 214 | 215 | 216 | 217 | 218 | 219 | 220 | -------------------------------------------------------------------------------- /TLC/TLC_Monad.v: -------------------------------------------------------------------------------- 1 | Require Export CatSem.TLC.TLC_syntax. 2 | Require Export CatSem.CAT.monad_haskell. 3 | 4 | Set Implicit Arguments. 5 | Unset Strict Implicit. 6 | Unset Transparent Obligations. 7 | Unset Automatic Introduction. 8 | 9 | Obligation Tactic := fin. 10 | 11 | Program Instance TLCm : 12 | Monad_struct (C:=IT) TLC := { 13 | weta := var; 14 | kleisli := _subst 15 | }. 16 | 17 | Canonical Structure TLCM := Build_Monad TLCm. 18 | 19 | 20 | (** some equalities : 21 | - lift = rename 22 | - shift = _shift 23 | *) 24 | 25 | Lemma rename_lift V t (v : TLC V t) W (f : V ---> W) : 26 | v //- f = lift (M:=TLCM) f _ v. 27 | Proof. 28 | unfold lift; 29 | fin. 30 | Qed. 31 | 32 | Lemma shift_shift a V W (f : V ---> TLC W) t (y : opt a V t) : 33 | y >- f = shift f y. 34 | Proof. 35 | intros a V W f t y. 36 | destruct y; 37 | simpl; 38 | auto. 39 | unfold lift; 40 | fin. 41 | Qed. 42 | 43 | Hint Resolve shift_shift rename_lift : fin. 44 | Hint Rewrite shift_shift rename_lift : fin. 45 | -------------------------------------------------------------------------------- /TLC/TLC_RMonad.v: -------------------------------------------------------------------------------- 1 | 2 | Require Export CatSem.CAT.rmonad. 3 | Require Export CatSem.CAT.rmodule_TYPE. 4 | Require Export CatSem.TLC.TLC_semantics. 5 | 6 | Set Implicit Arguments. 7 | Unset Strict Implicit. 8 | Unset Transparent Obligations. 9 | Unset Automatic Introduction. 10 | 11 | 12 | Program Instance TLCBs : RMonad_struct (IDelta TY) BETA := { 13 | rweta c := SM_ind (var (V:=c)) ; 14 | rkleisli := subst 15 | }. 16 | 17 | Canonical Structure TLCB := Build_RMonad TLCBs. 18 | 19 | Lemma shift_shift r V t (y : opt r V t) W 20 | (f : ipo_mor (sm_ipo (T:=TY) V) (BETA W)): 21 | sshift_ (P:=TLCB) (W:=W) f y = y >- f. 22 | Proof. 23 | intros r v t y. 24 | destruct y as [t y | ]; 25 | simpl; 26 | intros; 27 | fin. 28 | Qed. 29 | 30 | 31 | Section app_po. 32 | 33 | Variables r s : TLCtypes. 34 | 35 | Obligation Tactic := idtac. 36 | 37 | Program Instance app_po_mors c : PO_mor_struct 38 | (a := PO_product (ipo_proj (BETA c) (r ~> s)) (ipo_proj (BETA c) r)) 39 | (b := ipo_proj (BETA c) s) 40 | (fun z => app (fst z) (snd z)). 41 | Next Obligation. 42 | Proof. 43 | intros c. 44 | unfold Proper; 45 | red. 46 | intros x y H. 47 | destruct H. 48 | simpl in *. 49 | transitivity (app v' w). 50 | apply cp_App1; auto. 51 | apply cp_App2; auto. 52 | Qed. 53 | 54 | Definition app_po_mor c := Build_PO_mor (app_po_mors c). 55 | 56 | Obligation Tactic := fin. 57 | 58 | Program Instance tlc_app_s : RModule_Hom_struct 59 | (M:=product ((FIB_RMOD TLCB (r ~> s)) TLCB) 60 | ((FIB_RMOD TLCB r) TLCB)) 61 | (N:=FIB_RMOD TLCB s TLCB) (app_po_mor). 62 | 63 | Definition tlc_app := Build_RModule_Hom tlc_app_s. 64 | 65 | End app_po. 66 | 67 | Section abs_po. 68 | 69 | Variables r s : TLCtypes. 70 | 71 | Obligation Tactic := unfold Proper; red; 72 | simpl; intros; apply cp_Lam; auto. 73 | 74 | Program Instance abs_po_mors c : PO_mor_struct 75 | (a := (ipo_proj (BETA (c * r)) s)) 76 | (b := ipo_proj (BETA c) (r ~> s)) (@abs _ _ _ ). 77 | 78 | Definition abs_po_mor c := Build_PO_mor (abs_po_mors c). 79 | 80 | Obligation Tactic := idtac. 81 | 82 | Program Instance tlc_abs_s : RModule_Hom_struct 83 | (M:=(FIB_RMOD TLCB s) ((DER_RMOD _ TLCB r) TLCB)) 84 | (N:=(FIB_RMOD TLCB (r ~> s)) TLCB) abs_po_mor. 85 | Next Obligation. 86 | Proof. 87 | fin. 88 | apply f_equal. 89 | apply subst_eq. 90 | intros. 91 | apply (shift_shift y f ). 92 | Qed. 93 | 94 | Definition tlc_abs := Build_RModule_Hom tlc_abs_s. 95 | 96 | End abs_po. 97 | 98 | 99 | 100 | -------------------------------------------------------------------------------- /TLC/TLC_types.v: -------------------------------------------------------------------------------- 1 | Set Implicit Arguments. 2 | Unset Strict Implicit. 3 | Unset Transparent Obligations. 4 | Unset Automatic Introduction. 5 | 6 | Inductive TLCtypes : Type := 7 | | TLCbase : TLCtypes 8 | | TLCarrow : TLCtypes -> TLCtypes -> TLCtypes. 9 | 10 | 11 | -------------------------------------------------------------------------------- /ULC/ULC_Monad.v: -------------------------------------------------------------------------------- 1 | Require Export CatSem.CAT.monad_haskell. 2 | Require Export CatSem.ULC.ULC_syntax. 3 | 4 | Set Implicit Arguments. 5 | Unset Strict Implicit. 6 | Unset Automatic Introduction. 7 | 8 | Obligation Tactic := fin. 9 | 10 | Program Instance ULCm : 11 | Monad_struct (C:=TT) ULC := { 12 | weta := Var; 13 | kleisli := _subst 14 | }. 15 | 16 | Canonical Structure ULCM := Build_Monad ULCm. 17 | 18 | Lemma rename_lift V (v : ULC V) W (f : V ---> W) : 19 | v //- f = lift (M:=ULCM) f v. 20 | Proof. 21 | unfold lift; 22 | fin. 23 | Qed. 24 | 25 | Lemma shift_shift V W (f : V ---> ULC W) (y : V*) : 26 | y >- f = _shift f y. 27 | Proof. 28 | intros V W f y. 29 | destruct y; 30 | simpl; 31 | auto. 32 | Qed. -------------------------------------------------------------------------------- /ULC/ULC_RMonad.v: -------------------------------------------------------------------------------- 1 | 2 | Require Export CatSem.CAT.cat_TYPE_option. 3 | Require Export CatSem.CAT.rmonad. 4 | Require Export CatSem.CAT.rmodule_TYPE. 5 | Require Export CatSem.ULC.ULC_semantics. 6 | 7 | Set Implicit Arguments. 8 | Unset Strict Implicit. 9 | Unset Automatic Introduction. 10 | 11 | Obligation Tactic := fin. 12 | 13 | Program Instance ULCBETA_s : RMonad_struct (SM_po) ULCBETA := { 14 | rweta := VAR; 15 | rkleisli a b f := SUBST f 16 | }. 17 | Next Obligation. 18 | Proof. 19 | unfold Proper; 20 | red; 21 | fin. 22 | Qed. 23 | 24 | Definition ULCBETA : RMonad (SM_po) := Build_RMonad ULCBETA_s. 25 | 26 | Lemma rename_lift V (v : ULC V) W (f : V ---> W) : 27 | v //- f = rlift ULCBETA f v. 28 | Proof. 29 | unfold lift; 30 | fin. 31 | Qed. 32 | 33 | 34 | Section app_po. 35 | 36 | Obligation Tactic := idtac. 37 | 38 | Program Instance app_po_mors c : PO_mor_struct 39 | (a := PO_product (ULCBETA c ) (ULCBETA c)) 40 | (b := ULCBETA c) 41 | (fun z => App (fst z) (snd z)). 42 | Next Obligation. 43 | Proof. 44 | intros c. 45 | unfold Proper; 46 | red. 47 | intros x y H. 48 | destruct H. 49 | simpl in *. 50 | transitivity (App v' w). 51 | apply cp_App1; auto. 52 | apply cp_App2; auto. 53 | Qed. 54 | 55 | Definition app_po_mor c := Build_PO_mor (app_po_mors c). 56 | 57 | Obligation Tactic := fin. 58 | 59 | Program Instance ulc_app_s : RModule_Hom_struct 60 | (M:=product (C:=RMOD _ Ord) ULCBETA ULCBETA) 61 | (N:=ULCBETA) (app_po_mor). 62 | 63 | Definition ulc_app := Build_RModule_Hom ulc_app_s. 64 | 65 | End app_po. 66 | 67 | Section abs_po. 68 | 69 | 70 | Obligation Tactic := unfold Proper; red; 71 | simpl; intros; apply cp_Abs; auto. 72 | 73 | Program Instance abs_po_mors c : PO_mor_struct 74 | (a := (ULCBETA (option c))) 75 | (b := (ULCBETA c)) (@Abs _ ). 76 | 77 | Definition abs_po_mor c := Build_PO_mor (abs_po_mors c). 78 | 79 | Obligation Tactic := idtac. 80 | 81 | Program Instance ulc_abs_s : RModule_Hom_struct 82 | (M:= DER_RMOD_not _ _ ULCBETA) 83 | (N:=ULCBETA) abs_po_mor. 84 | Next Obligation. 85 | Proof. 86 | fin. 87 | apply f_equal. 88 | apply subst_eq. 89 | intros. 90 | match goal with [H:option _ |- _ ]=> destruct H end; 91 | simpl; fin. 92 | Qed. 93 | 94 | Definition ulc_abs := Build_RModule_Hom ulc_abs_s. 95 | 96 | End abs_po. -------------------------------------------------------------------------------- /ULC/ULC_terms_print.v: -------------------------------------------------------------------------------- 1 | Require Export CatSem.ULC.ULC_terms. 2 | 3 | Eval compute in ULC_theta. 4 | 5 | Notation "a @ b" := (App a b) (at level 42, left associativity). 6 | Notation "'1'" := (Var None) (at level 33). 7 | Notation "'2'" := (Var (Some None)) (at level 24). 8 | Notation "'3'" := (Var (Some (Some None))) (at level 23). 9 | Notation "'4'" := (Var (Some (Some (Some None)))) (at level 22). 10 | Eval compute in ULC_theta. 11 | Eval compute in ULC_True. 12 | Eval compute in ULC_False. 13 | Eval compute in ULC_Nat 0. 14 | Eval compute in ULC_Nat 2. 15 | Eval compute in ULCsucc. 16 | Eval compute in ULC_cond. 17 | Eval compute in ULC_omega. 18 | Eval compute in ULC_zero. 19 | Eval compute in ULC_pred. 20 | Eval compute in ULC_Y. -------------------------------------------------------------------------------- /ULC/subst_experiment.v: -------------------------------------------------------------------------------- 1 | 2 | Require Export CatSem.CAT.monad_haskell. 3 | Require Export CatSem.CAT.cat_TYPE. 4 | 5 | Notation "X *" := (option X) (at level 5). 6 | 7 | Section freshness. 8 | 9 | (** b fresh for z => z{a:=x} == z{a:=b}{b:=x} *) 10 | 11 | Variable P : Monad TYPE. 12 | Variable X : Type. 13 | Variable z : P (X*). 14 | Variable a : P (X). 15 | Implicit Arguments kleisli [C F a b]. 16 | Implicit Arguments weta [C F c]. 17 | 18 | Definition fresh_subst := kleisli P (a:=X*) (b:=X * *) 19 | (fun x => match x in option _ return P (X * *) with 20 | | Some x => weta P (Some (Some x)) 21 | | None => weta P None 22 | end). 23 | 24 | Definition left_subst := kleisli P 25 | (fun x => match x in option _ return P X with 26 | | Some x => weta P x 27 | | None => a 28 | end). 29 | 30 | Definition right_subst := kleisli P 31 | (fun x => match x in option _ return P X with 32 | | Some x' => match x' with 33 | | Some x'' => weta P x'' 34 | | None => a 35 | end 36 | | None => a 37 | end). 38 | 39 | Lemma don: left_subst z = right_subst (fresh_subst z). 40 | Proof. 41 | unfold left_subst. 42 | unfold right_subst. 43 | unfold fresh_subst. 44 | rew (klkl P) . 45 | assert (H:= kl_eq P). 46 | simpl in H. 47 | apply H. 48 | intros. 49 | destruct x; simpl; auto. 50 | rew (etakl P). 51 | rew (etakl P). 52 | Qed. 53 | 54 | End freshness. 55 | 56 | Section freshness2. 57 | 58 | (** b fresh for z => z{b:=x} == z *) 59 | 60 | Variable P : Monad TYPE. 61 | 62 | Variable X : TYPE. 63 | 64 | Variable z : P X. 65 | 66 | Definition z' : P(X*) := lift (M:=P) (@Some X) z. 67 | 68 | Variable a : P X. 69 | 70 | Implicit Arguments kleisli [C F a b]. 71 | Implicit Arguments weta [C F c]. 72 | 73 | Definition freshsubst := kleisli P (a:=X*)(b:=X) 74 | (fun x => match x with 75 | | Some x => weta P x 76 | | None => a 77 | end). 78 | 79 | 80 | Lemma don2 : freshsubst z' = z. 81 | Proof. 82 | unfold freshsubst. 83 | unfold z'. 84 | unfold lift. 85 | simpl. 86 | rew (klkl P). 87 | app (kleta_id (FM:=P)). 88 | intros. 89 | rew (etakl P). 90 | Qed. 91 | 92 | End freshness2. 93 | 94 | (* 95 | simpl. 96 | simpl. 97 | rew (etakl P). 98 | 99 | Check inj. 100 | 101 | 102 | 103 | simpl. 104 | generalize a; clear a. 105 | 106 | 107 | Require Export CatSem.ULC.ULC_syntax. 108 | 109 | Section freshness. 110 | 111 | Variable X : Type. 112 | 113 | Variable z : ULC (X* ). 114 | Variable a : ULC X. 115 | 116 | Definition fresh_subst := _subst 117 | (fun x => match x in option _ return ULC (X * * ) with 118 | | Some x => Var (Some (Some x)) 119 | | None => Var None 120 | end). 121 | 122 | Definition left_subst := _subst 123 | (fun x => match x in option _ return ULC X with 124 | | Some x => Var x 125 | | None => a 126 | end). 127 | 128 | Definition right_subst := _subst 129 | (fun x => match x in option _ return ULC X with 130 | | Some x' => match x' with 131 | | Some x'' => Var x'' 132 | | None => a 133 | end 134 | | None => a 135 | end). 136 | Check _subst. 137 | Check left_subst. 138 | 139 | Require Import Coq.Program.Equality. 140 | 141 | Lemma don: left_subst z = right_subst (fresh_subst z). 142 | unfold left_subst. 143 | unfold right_subst. 144 | unfold fresh_subst. 145 | generalize a. 146 | clear a. 147 | generalize z. 148 | clear z. 149 | intro z. 150 | dependent inversion z; 151 | simpl; intros. 152 | destruct o; 153 | simpl; auto. 154 | apply f_equal. 155 | *) 156 | 157 | --------------------------------------------------------------------------------