├── .gitignore ├── README.md └── src ├── Categories.v ├── CategoriesAux.v ├── Fin.v ├── Finmap.v ├── FinmapMetric.v ├── KSOp.v ├── KSSem.v ├── KSTm.v ├── KSTy.v ├── KSTyping.v ├── KnasterTarski.v ├── Makefile ├── Makefile.win ├── MetricCore.v ├── MetricRec.v ├── NSetoid.v ├── Predom.v ├── PredomAll.v ├── PredomAux.v ├── PredomCore.v ├── PredomFix.v ├── PredomKleisli.v ├── PredomLift.v ├── PredomProd.v ├── PredomRec.v ├── PredomSum.v ├── Vars.v ├── VarsRec.v ├── buildcoqdoc.bat ├── builddistrib.bat ├── coqdoc.sty ├── linecount.bat ├── mmsem.v ├── moperational.v ├── mpremet.v ├── msem.v ├── msyntax.v ├── numbers ├── numbers.bat ├── typedadequacy.v ├── typeddensem.v ├── typedlambda.v ├── typedlambdapat.v ├── typedopsem.v ├── typedopsempat.v ├── typedsoundness.v ├── typedsubst.v ├── typedsubstalt.v ├── unii.v ├── uniiade.v ├── uniiop.v ├── uniirec.v ├── uniisem.v ├── uniisound.v ├── untypeddensem.v ├── untypedlambda.v ├── untypedopsem.v ├── untypedsubst.v └── utility.v /.gitignore: -------------------------------------------------------------------------------- 1 | *.vo 2 | *.glob 3 | *.v.d 4 | -------------------------------------------------------------------------------- /README.md: -------------------------------------------------------------------------------- 1 | coqdomains 2 | ========== 3 | 4 | Domain theory and denotational semantics in Coq 5 | -------------------------------------------------------------------------------- /src/Fin.v: -------------------------------------------------------------------------------- 1 | (********************************************************************************** 2 | * Fin.v * 3 | * Formalizing Domains, Ultrametric Spaces and Semantics of Programming Languages * 4 | * Nick Benton, Lars Birkedal, Andrew Kennedy and Carsten Varming * 5 | * Jan 2012 * 6 | * Build with Coq 8.3pl2 plus SSREFLECT * 7 | **********************************************************************************) 8 | 9 | (* Finite nats and maps *) 10 | 11 | Require Export ssreflect ssrnat. 12 | Require Export NaryFunctions. 13 | Require Import Program. 14 | Set Implicit Arguments. 15 | Unset Strict Implicit. 16 | Import Prenex Implicits. 17 | 18 | Inductive Fin : nat -> Type := 19 | | FinZ : forall n, Fin n.+1 20 | | FinS : forall n, Fin n -> Fin n.+1. 21 | 22 | Definition FMap n D := Fin n -> D. 23 | 24 | (* Head, tail and cons *) 25 | Definition tl n D (m:FMap n.+1 D) : FMap n D := fun v => m (FinS v). 26 | Definition hd n D (m:FMap n.+1 D) : D := m (FinZ _). 27 | 28 | Definition cons n D (hd:D) (tl : FMap n D) : FMap n.+1 D := 29 | (fun var => 30 | match var in Fin n return FMap n.-1 D -> D with 31 | | FinZ _ => fun _ => hd 32 | | FinS _ var' => fun tl' => tl' var' 33 | end tl). 34 | 35 | Fixpoint nth n D (v : Fin n) : FMap n D -> D := 36 | match v with 37 | | FinZ _ => fun m => hd m 38 | | FinS _ i => fun m => nth i (tl m) 39 | end. 40 | 41 | Fixpoint asTuple n D : FMap n D -> D ^ n := 42 | match n with 43 | | 0 => fun m => tt 44 | | n.+1 => fun m => (hd m, asTuple (tl m)) 45 | end. 46 | 47 | Lemma hdCons : forall n D (v : D) (m : FMap n D), hd (cons v m) = v. 48 | Proof. by []. Qed. 49 | 50 | Axiom Extensionality : forall n D (m1 m2 : FMap n D), (forall i, m1 i = m2 i) -> m1 = m2. 51 | 52 | Lemma tlCons : forall n D (v : D) (s : FMap n D), tl (cons v s) = s. 53 | Proof. move => n D v s. by apply Extensionality. Qed. 54 | 55 | Ltac extFMap i := (apply Extensionality; intros i; dependent destruction i). 56 | 57 | Lemma consEta : forall n D (m:FMap n.+1 D), m = cons (hd m) (tl m). 58 | Proof. move => n D m. extFMap i; by []. Qed. 59 | 60 | 61 | 62 | -------------------------------------------------------------------------------- /src/FinmapMetric.v: -------------------------------------------------------------------------------- 1 | (********************************************************************************** 2 | * Finmap.v * 3 | * Formalizing Domains, Ultrametric Spaces and Semantics of Programming Languages * 4 | * Nick Benton, Lars Birkedal, Andrew Kennedy and Carsten Varming * 5 | * Jan 2012 * 6 | * Build with Coq 8.3pl2 plus SSREFLECT * 7 | **********************************************************************************) 8 | 9 | Require Export NSetoid MetricCore Finmap. 10 | 11 | Section SET. 12 | Variable T:compType. 13 | Variable S:setoidType. 14 | 15 | Lemma finmap_setoid_axiom : Setoid.axiom (fun f f' : FinDom T S => dom f =i dom f' /\ forall t, t \in dom f -> f t =-= f' t). 16 | split ; last split. 17 | - by move => f ; simpl ; split ; last by []. 18 | - move => x y z ; simpl => X Y. case: X => D X. case: Y => D' Y. 19 | split. move => a. rewrite (D a). by rewrite (D' a). 20 | move => a A. rewrite -> (X a A). rewrite D in A. by apply (Y _ A). 21 | - move => x y ; simpl => X. case: X => D X. 22 | split => a. by rewrite (D a). move => A. rewrite <- D in A. by rewrite -> (X a A). 23 | Qed. 24 | 25 | Definition findom_setoidMixin := SetoidMixin finmap_setoid_axiom. 26 | Canonical Structure findom_setoidType := Eval hnf in SetoidType findom_setoidMixin. 27 | 28 | Lemma indom_app_respect (f f': findom_setoidType) x (P:x \in dom f) (P':x \in dom f') : f =-= f' -> indom_app P =-= indom_app P'. 29 | move => e. have X:Some (indom_app P) =-= Some (indom_app P') ; last by apply X. 30 | do 2 rewrite indom_app_eq. case: f e P => s A. case: f' P' => s' A'. move => P e P'. 31 | unfold tset_eq in e. simpl in e. by apply (proj2 e). 32 | Qed. 33 | 34 | Lemma respect_domeq (f f':findom_setoidType) : f =-= f' -> dom f =i dom f'. 35 | case:f => s P. case:f' => s' P'. 36 | unfold dom. simpl. unfold tset_eq. simpl. unfold dom. simpl. by move => [A _]. 37 | Qed. 38 | 39 | Lemma findom_sapp_respect (x:T) : setoid_respect (fun f:findom_setoidType => f x). 40 | move => f f' e. case_eq (x \in dom f) => D. by apply (proj2 e x D). 41 | rewrite findom_undef ; last by rewrite D. rewrite (proj1 e) in D. 42 | rewrite findom_undef ; last by rewrite D. by []. 43 | Qed. 44 | 45 | Definition findom_sapp (x:T) : findom_setoidType =-> option_setoidType S := 46 | Eval hnf in mk_fset (findom_sapp_respect x). 47 | 48 | 49 | End SET. 50 | 51 | Section MET. 52 | Variable T:compType. 53 | Variable M:metricType. 54 | 55 | Lemma findom_metric_axiom : Metric.axiom (fun n => (fun f f' : FinDom T M => match n with O => True | S _ => 56 | dom f =i dom f' /\ forall i, i \in dom f -> i \in dom f' -> f i = n = f' i end)). 57 | move => n x y z. split ; last split ; last split ; last split ; clear. 58 | - split => X. 59 | + split. specialize (X 1). simpl in X. by apply (proj1 X). 60 | move => t I. apply: (proj1 (Mrefl _ _)) => n. case: n ; first by []. move => n. specialize (X n.+1). 61 | simpl in X. apply (proj2 X). by apply I. rewrite (proj1 X) in I. by apply I. 62 | + case ; first by []. move => n. split ; first by apply (proj1 X). move => i I I'. 63 | apply (proj2 (Mrefl _ _)). by apply (proj2 X _ I). 64 | - move => A. simpl. simpl in A. case: n A ; first by []. move => n [A B]. split ; first by move => a ; rewrite -> A. 65 | move => i I I'. specialize (B i I' I). by apply (Msym B). 66 | - move => A B. simpl in A,B. simpl. case: n A B ; first by []. 67 | move => n [A D] [A' D']. split ; first by move => a ; rewrite A; rewrite A'. 68 | move => i I I'. specialize (D i I). rewrite A in I. specialize (D I). specialize (D' i I I'). 69 | by apply (Mrel_trans D D'). 70 | - move => A. simpl. simpl in A. case: n A ; first by []. move => n [D A]. split ; first by []. 71 | move => i I I'. specialize (A i I I'). by apply Mmono. 72 | - by []. 73 | Qed. 74 | 75 | Canonical Structure findom_metricMixin := MetricMixin findom_metric_axiom. 76 | Canonical Structure findom_metricType := Eval hnf in MetricType findom_metricMixin. 77 | 78 | Lemma findom_f_nonexp (f f':findom_metricType) n x : f = n = f' -> f x = n = f' x. 79 | case: n x ; first by []. 80 | move => n x [E P]. case_eq (x \in dom f) => D. 81 | - have D':=D. rewrite E in D'. specialize (P x D D'). by apply P. 82 | rewrite findom_undef ; last by rewrite D. 83 | rewrite E in D. rewrite findom_undef ; last by rewrite D. by []. 84 | Qed. 85 | 86 | Lemma findom_sapp_ne (x:T) : nonexpansive (fun f:findom_metricType => f x). 87 | move => n f f' e. by apply: findom_f_nonexp. 88 | Qed. 89 | 90 | Definition findom_napp (x:T) : findom_metricType =-> option_metricType M := 91 | Eval hnf in mk_fmet (findom_sapp_ne x). 92 | 93 | End MET. 94 | 95 | Section CMET. 96 | Variable T:compType. 97 | Variable M:cmetricType. 98 | 99 | Lemma findom_chain_dom x n (c:cchain (findom_metricType T M)) : x \in dom (c 1) -> x \in dom (c n.+1). 100 | elim: n c ; first by []. 101 | move => n IH c X. specialize (IH (cutn 1 c)). simpl in IH. do 2 rewrite -> addSn, -> add0n in IH. 102 | apply IH. clear IH. have C:=cchain_cauchy c. specialize (C 1 1 2 (ltnSn _) (ltnW (ltnSn _))). 103 | by rewrite (proj1 C) in X. 104 | Qed. 105 | 106 | Definition findom_chainx (c:cchain (findom_metricType T M)) x (P:x \in dom (c 1)) : cchain M. 107 | exists (fun n => @indom_app _ _ (c n.+1) _ (findom_chain_dom _ n _ P)). 108 | case ; first by []. 109 | move => n i j l l'. 110 | have X: Some (indom_app (findom_chain_dom x i c P)) = n.+1 = Some (indom_app (findom_chain_dom x j c P)) ; last by apply X. 111 | do 2 rewrite indom_app_eq. have a:= (@cchain_cauchy _ c n.+1 i.+1 j.+1 (ltnW l) (ltnW l')). 112 | by apply findom_f_nonexp. 113 | Defined. 114 | 115 | Definition findom_lub (c:cchain (findom_metricType T M)) : findom_metricType T M := 116 | @findom_map _ _ _ (c 1) (fun x X => umet_complete (@findom_chainx c x X)). 117 | 118 | Lemma ff (P:T -> nat -> Prop) (A:forall x n m, n <= m -> P x n -> P x m) (s:seq T) : 119 | (forall x, x \in s -> exists m, P x m) -> exists m, forall x, x \in s -> P x m. 120 | elim: s ; first by move => X; exists 0. 121 | move => t s IH X. specialize (IH (fun x A => X x (mem_cons t A))). case: IH => [m IH]. 122 | case: (X t (mem_head t _)) => [m' Pm']. exists (maxn m m'). move => x I. 123 | rewrite in_cons in I. case_eq (x == t) => E. 124 | - rewrite (eqP E). apply: (A _ _ _ _ Pm'). by rewrite leq_maxr. 125 | - rewrite E in I. simpl in I. specialize (IH _ I). apply: (A _ _ _ _ IH). by rewrite leq_maxl. 126 | Qed. 127 | 128 | Lemma findom_lubP (c:cchain (findom_metricType T M)) : mconverge c (findom_lub c). 129 | unfold findom_lub. case ; first by exists 0. move => n. 130 | have A:exists m, forall x, x \in dom (c 1) -> forall (X: x \in dom (c 1)) i, m < i -> c i x = n.+1 = Some (umet_complete (findom_chainx _ _ X)). 131 | apply (@ff (fun x m => forall (X:x \in dom (c 1)) i, m < i -> c i x = n.+1 = Some (umet_complete (findom_chainx _ _ X)))). 132 | clear. move => x j i A Y X k L. apply Y. by apply (@ssrnat.leq_ltn_trans _ _ _ A L). 133 | move => x I. destruct (cumet_conv (findom_chainx _ _ I) n.+1) as [m P]. 134 | exists m. move => X. case ; first by []. move => i L. specialize (P i L). simpl in P. 135 | rewrite <- (indom_app_eq (findom_chain_dom _ i _ I)). 136 | apply: (Mrel_trans P). apply umet_complete_extn. move=> j. simpl. 137 | have A:Some (indom_app (findom_chain_dom _ j _ I)) = n.+1 = Some (indom_app (findom_chain_dom _ j _ X)) ; last by apply A. 138 | rewrite (indom_app_eq (findom_chain_dom _ j _ I)). by rewrite (indom_app_eq (findom_chain_dom _ j _ X)). 139 | destruct A as [m A]. exists m.+1. case ; first by []. move => i L. split. rewrite <- dom_findom_map. 140 | have B:=cchain_cauchy c. specialize (B 1 i.+1 1 (ltn0Sn i) (ltn0Sn _)). apply (proj1 B). 141 | move => x I I'. specialize (A x). 142 | have J:x \in dom (c 1). have J:=cchain_cauchy c. specialize (J 1 i.+1 1 (ltn0Sn _) (ltn0Sn _)). have X:= (proj1 J). 143 | rewrite X in I. by []. 144 | specialize (A J J i.+1 L). apply (Mrel_trans A). 145 | by rewrite -> (findom_map_app J). 146 | Qed. 147 | 148 | Canonical Structure findom_cmetricMixin := CMetricMixin findom_lubP. 149 | Canonical Structure findom_cmetricType := Eval hnf in CMetricType findom_cmetricMixin. 150 | 151 | Lemma dom_findom_lub (c:cchain findom_cmetricType) : dom (umet_complete c) = dom (c 1). 152 | unfold umet_complete. simpl. unfold findom_lub. 153 | by rewrite <- dom_findom_map. 154 | Qed. 155 | 156 | Lemma findom_lub_eq (c:cchain findom_cmetricType) x : umet_complete c x =-= umet_complete (liftc (findom_napp _ _ x) c). 157 | apply: (proj1 (Mrefl _ _)) => n. 158 | case: (cumet_conv c n). move => m X. 159 | case: (cumet_conv (liftc (findom_napp T M x) c) n). move => m' Y. 160 | specialize (X ((m+m')%N.+1)%N). specialize (X (leqW (leq_addr _ _))). specialize (Y ((m+m')%N.+1) (leqW (leq_addl _ _))). 161 | case: n X Y ; first by []. 162 | move => n X Y. case: X => D X. specialize (X x). 163 | - case_eq (x \in dom (umet_complete c)) => e. 164 | have e':=e. rewrite dom_findom_lub in e'. specialize (X (findom_chain_dom _ _ _ e') e). 165 | rewrite <- X. by rewrite <- Y. 166 | - rewrite findom_undef ; last by rewrite e. clear X. rewrite <- Y. clear Y. simpl. 167 | rewrite findom_undef ; first by []. rewrite D. by rewrite e. 168 | Qed. 169 | 170 | End CMET. 171 | 172 | 173 | 174 | -------------------------------------------------------------------------------- /src/KSOp.v: -------------------------------------------------------------------------------- 1 | (********************************************************************************** 2 | * KSOp.v * 3 | * Formalizing Domains, Ultrametric Spaces and Semantics of Programming Languages * 4 | * Nick Benton, Lars Birkedal, Andrew Kennedy and Carsten Varming * 5 | * Jan 2012 * 6 | * Build with Coq 8.3pl2 plus SSREFLECT * 7 | **********************************************************************************) 8 | 9 | (* Operational semantics of the kitchen sink language *) 10 | 11 | Require Import Finmap KSTm Fin. 12 | 13 | Set Implicit Arguments. 14 | Unset Strict Implicit. 15 | Import Prenex Implicits. 16 | 17 | Import Tm. 18 | 19 | Definition Heap := FinDom [compType of nat] (Value O). 20 | 21 | Definition subSingle E (v:Value E) e := subExp (cons v (@idSub _)) e. 22 | 23 | (*=EV *) 24 | Inductive EV : nat -> Exp 0 -> Heap -> Value O -> Heap -> Type := 25 | | EvVAL h v : EV O (VAL v) h v h 26 | | EvFST h v0 v1 : EV O (FST (PAIR v0 v1)) h v0 h 27 | | EvSND h v0 v1 : EV O (SND (PAIR v0 v1)) h v1 h 28 | | EvOP h op n0 n1 : EV O (OP op (PAIR (INT n0) (INT n1))) h (INT (op n0 n1)) h 29 | | EvUNFOLD h v : EV 1 (UNFOLD (FOLD v)) h v h 30 | | EvREF (h:Heap) v (l:nat) : l \notin dom h -> EV 1 (REF v) h (LOC l) (updMap l v h) 31 | | EvBANG (h:Heap) v (l:nat) : h l = Some v -> EV 1 (BANG (LOC l)) h v h 32 | | EvASSIGN (h:Heap) v l : h l -> EV 1 (ASSIGN (LOC l) v) h UNIT (updMap l v h) 33 | | EvLET h n0 e0 v0 h0 n1 e1 v h1 : EV n0 e0 h v0 h0 -> 34 | EV n1 (subSingle v0 e1) h0 v h1 -> EV (n0 + n1) (LET e0 e1) h v h1 35 | | EvAPP h n e v0 v h0 : EV n (subSingle v e) h v0 h0 -> 36 | EV n (APP (LAM e) v) h v0 h0 37 | | EvTAPP h n e v0 h0 : EV n e h v0 h0 -> EV n (TAPP (TLAM e)) h v0 h0 38 | | EvCASEL h n v e0 e1 v0 h0 : EV n (subSingle v e0) h v0 h0 -> 39 | EV n (CASE (INL v) e0 e1) h v0 h0 40 | | EvCASER h n v e0 e1 v0 h0 : EV n (subSingle v e1) h v0 h0 -> 41 | EV n (CASE (INR v) e0 e1) h v0 h0. 42 | (*=End *) 43 | 44 | -------------------------------------------------------------------------------- /src/KSTm.v: -------------------------------------------------------------------------------- 1 | (********************************************************************************** 2 | * KSTy.v * 3 | * Formalizing Domains, Ultrametric Spaces and Semantics of Programming Languages * 4 | * Nick Benton, Lars Birkedal, Andrew Kennedy and Carsten Varming * 5 | * Jan 2012 * 6 | * Build with Coq 8.3pl2 plus SSREFLECT * 7 | **********************************************************************************) 8 | 9 | (* Kitchen sink language, well-scoped by construction *) 10 | 11 | Require Export ssreflect ssrnat. 12 | Require Import Program. 13 | Set Implicit Arguments. 14 | Unset Strict Implicit. 15 | Import Prenex Implicits. 16 | Require Import Fin. 17 | 18 | Ltac Rewrites E := 19 | (intros; simpl; try rewrite E; 20 | repeat (match goal with | [H : context[eq _ _] |- _] => rewrite H end); 21 | auto). 22 | 23 | Module Tm. 24 | 25 | Definition Env := nat. 26 | Definition Var := Fin. 27 | 28 | (*=ValueExp *) 29 | Inductive Value E := 30 | | VAR: Var E -> Value E 31 | | LOC: nat -> Value E 32 | | INT: nat -> Value E 33 | | TLAM: Exp E -> Value E 34 | | LAM: Exp E.+1 -> Value E 35 | | UNIT : Value E 36 | | PAIR: Value E -> Value E -> Value E 37 | | INL : Value E -> Value E 38 | | INR : Value E -> Value E 39 | | FOLD : Value E -> Value E 40 | with Exp E := 41 | | VAL: Value E -> Exp E 42 | | FST: Value E -> Exp E 43 | | SND: Value E -> Exp E 44 | | OP: (nat -> nat -> nat) -> Value E -> Exp E 45 | | UNFOLD: Value E -> Exp E 46 | | REF: Value E -> Exp E 47 | | BANG: Value E -> Exp E 48 | | ASSIGN: Value E -> Value E -> Exp E 49 | | LET: Exp E -> Exp E.+1 -> Exp E 50 | | APP: Value E -> Value E -> Exp E 51 | | TAPP: Value E -> Exp E 52 | | CASE: Value E -> Exp E.+1 -> Exp E.+1 -> Exp E. 53 | (*=End *) 54 | 55 | 56 | Implicit Arguments INT [E]. 57 | Implicit Arguments LOC [E]. 58 | Implicit Arguments UNIT [E]. 59 | 60 | Scheme Value_ind2 := Induction for Value Sort Prop 61 | with Exp_ind2 := Induction for Exp Sort Prop. 62 | Combined Scheme ExpValue_ind from Value_ind2, Exp_ind2. 63 | 64 | (*========================================================================== 65 | Variable-domain maps. 66 | By instantiating P with Var we get renamings. 67 | By instantiating P with Value we get substitutions. 68 | ==========================================================================*) 69 | Module Map. 70 | 71 | Section MAP. 72 | 73 | Variable P : Env -> Type. 74 | Definition Map E E' := FMap E (P E'). 75 | 76 | (*========================================================================== 77 | Package of operations used with a Map 78 | vr maps a Var into Var or Value (so is either the identity or TVAR) 79 | vl maps a Var or Value to a Value (so is either TVAR or the identity) 80 | wk weakens a Var or Value (so is either SVAR or renaming through SVAR on a value) 81 | ==========================================================================*) 82 | Record Ops := 83 | { 84 | vr : forall E, Var E -> P E; 85 | vl : forall E, P E -> Value E; 86 | wk : forall E, P E -> P E.+1; 87 | wkvr : forall E (var : Var E), wk (vr var) = vr (FinS var); 88 | vlvr : forall E (var : Var E), vl (vr var) = VAR var 89 | }. 90 | Variable ops : Ops. 91 | 92 | Definition lift E E' (m : Map E E') : Map E.+1 E'.+1 := 93 | (fun var => match var in Fin E return Map E.-1 E' -> P E'.+1 with 94 | | FinZ _ => fun _ => vr ops (FinZ _) 95 | | FinS _ x => fun m => wk ops (m x) 96 | end m). 97 | 98 | Definition shift E E' (m : Map E E') : Map E E'.+1 := fun var => wk ops (m var). 99 | Definition id E : Map E E := fun (var : Var E) => vr ops var. 100 | 101 | Lemma shiftCons : forall E E' (m : Map E E') (x : P E'), shift (cons x m) = cons (wk ops x) (shift m). 102 | Proof. move => E E' m x. apply Extensionality. intros var. by dependent destruction var. Qed. 103 | 104 | Lemma liftAsCons : forall E E' (m : Map E' E), lift m = cons (vr ops (FinZ _)) (shift m). 105 | Proof. move => E E' m. apply Extensionality. intros var. by dependent destruction var. Qed. 106 | 107 | Fixpoint mapVal E E' (m : Map E E') (v : Value E) : Value E' := 108 | match v with 109 | | VAR v => vl ops (m v) 110 | | INT i => INT i 111 | | LAM e => LAM (mapExp (lift m) e) 112 | | TLAM e => TLAM (mapExp m e) 113 | | LOC l => LOC l 114 | | UNIT => UNIT 115 | | PAIR v1 v2 => PAIR (mapVal m v1) (mapVal m v2) 116 | | INL v => INL (mapVal m v) 117 | | INR v => INR (mapVal m v) 118 | | FOLD v => FOLD (mapVal m v) 119 | end 120 | with mapExp E E' (m : Map E E') (e : Exp E) : Exp E' := 121 | match e with 122 | | VAL v => VAL (mapVal m v) 123 | | APP v0 v1 => APP (mapVal m v0) (mapVal m v1) 124 | | LET e0 e1 => LET (mapExp m e0) (mapExp (lift m) e1) 125 | | FST v => FST (mapVal m v) 126 | | SND v => SND (mapVal m v) 127 | | UNFOLD v => UNFOLD (mapVal m v) 128 | | OP op v => OP op (mapVal m v) 129 | | REF v => REF (mapVal m v) 130 | | ASSIGN v1 v2 => ASSIGN (mapVal m v1) (mapVal m v2) 131 | | BANG v => BANG (mapVal m v) 132 | | TAPP v => TAPP (mapVal m v) 133 | | CASE v e1 e2 => CASE (mapVal m v) (mapExp (lift m) e1) (mapExp (lift m) e2) 134 | end. 135 | 136 | Variable E E' : Env. 137 | Variable m : Map E E'. 138 | Lemma mapVAR : forall (var : Var _), mapVal m (VAR var) = vl ops (m var). by []. Qed. 139 | Lemma mapINT : forall n, mapVal m (INT n) = INT n. by []. Qed. 140 | Lemma mapLOC : forall n, mapVal m (LOC n) = LOC n. by []. Qed. 141 | Lemma mapLAM : forall (e : Exp _), mapVal m (LAM e) = LAM (mapExp (lift m) e). by []. Qed. 142 | Lemma mapOP : forall op v, mapExp m (OP op v) = OP op (mapVal m v). by []. Qed. 143 | Lemma mapVAL : forall (v : Value _), mapExp m (VAL v) = VAL (mapVal m v). by []. Qed. 144 | Lemma mapLET : forall (e1 : Exp _) (e2 : Exp _), mapExp m (LET e1 e2) = LET (mapExp m e1) (mapExp (lift m) e2). by []. Qed. 145 | Lemma mapAPP : forall (v1 : Value _) v2, mapExp m (APP v1 v2) = APP (mapVal m v1) (mapVal m v2). by []. Qed. 146 | Lemma mapPAIR : forall v1 v2, mapVal m (PAIR v1 v2) = PAIR (mapVal m v1) (mapVal m v2). by []. Qed. 147 | Lemma mapINL : forall v, mapVal m (INL v) = INL (mapVal m v). by []. Qed. 148 | Lemma mapINR : forall v, mapVal m (INR v) = INR (mapVal m v). by []. Qed. 149 | Lemma mapFOLD : forall v, mapVal m (FOLD v) = FOLD (mapVal m v). by []. Qed. 150 | Lemma mapTLAM : forall e, mapVal m (TLAM e) = TLAM (mapExp m e). by []. Qed. 151 | Lemma mapUNIT : mapVal m UNIT = UNIT. by []. Qed. 152 | Lemma mapFST : forall v, mapExp m (FST v) = FST (mapVal m v). by []. Qed. 153 | Lemma mapSND : forall v, mapExp m (SND v) = SND (mapVal m v). by []. Qed. 154 | Lemma mapBANG : forall v, mapExp m (BANG v) = BANG (mapVal m v). by []. Qed. 155 | Lemma mapREF : forall v, mapExp m (REF v) = REF (mapVal m v). by []. Qed. 156 | Lemma mapASSIGN : forall v1 v2, mapExp m (ASSIGN v1 v2) = ASSIGN (mapVal m v1) (mapVal m v2). by []. Qed. 157 | Lemma mapTAPP : forall v, mapExp m (TAPP v) = TAPP (mapVal m v). by []. Qed. 158 | Lemma mapUNFOLD : forall v, mapExp m (UNFOLD v) = UNFOLD (mapVal m v). by []. Qed. 159 | Lemma mapCASE : forall v e1 e2, mapExp m (CASE v e1 e2) = CASE (mapVal m v) (mapExp (lift m) e1) (mapExp (lift m) e2). by []. Qed. 160 | 161 | Lemma liftId : lift (@id E) = @id E.+1. 162 | Proof. apply Extensionality. intros var. dependent destruction var; [by [] | by apply wkvr]. 163 | Qed. 164 | 165 | Lemma idAsCons : @id E.+1 = cons (vr ops (FinZ _)) (shift (@id E)). 166 | Proof. apply Extensionality. intros var. dependent destruction var; first by []. unfold id, shift. simpl. by rewrite wkvr. Qed. 167 | 168 | End MAP. 169 | 170 | Hint Rewrite mapVAR mapINT mapLAM mapOP mapVAL mapLET mapAPP mapPAIR mapINL mapINR mapFOLD mapTLAM mapUNIT mapFST mapSND mapBANG mapREF mapASSIGN mapTAPP mapUNFOLD mapCASE : mapHints. 171 | Implicit Arguments id [P]. 172 | 173 | Lemma applyId P (ops:Ops P) E : 174 | (forall (v : Value E), mapVal ops (id ops E) v = v) 175 | /\ (forall (e : Exp E), mapExp ops (id ops E) e = e). 176 | Proof. move: E ; apply ExpValue_ind; intros; autorewrite with mapHints; Rewrites liftId. by apply vlvr. 177 | Qed. 178 | 179 | End Map. 180 | 181 | (*========================================================================== 182 | Variable renamings: Map Var 183 | ==========================================================================*) 184 | 185 | Definition Ren := Map.Map Var. 186 | Definition RenOps : Map.Ops Var. refine (@Map.Build_Ops _ (fun _ v => v) VAR FinS _ _). by []. by []. Defined. 187 | 188 | Definition renVal := Map.mapVal RenOps. 189 | Definition renExp := Map.mapExp RenOps. 190 | Definition liftRen := Map.lift RenOps. 191 | Definition shiftRen := Map.shift RenOps. 192 | Definition idRen := Map.id RenOps. 193 | 194 | (*========================================================================== 195 | Composition of renaming 196 | ==========================================================================*) 197 | 198 | Definition composeRen P E E' E'' (m : Map.Map P E' E'') (r : Ren E E') : Map.Map P E E'' := fun var => m (r var). 199 | 200 | Lemma liftComposeRen : forall E E' E'' P ops (m:Map.Map P E' E'') (r:Ren E E'), Map.lift ops (composeRen m r) = composeRen (Map.lift ops m) (liftRen r). 201 | Proof. move => E E' E'' P ops m r. apply Extensionality. intros var. by dependent destruction var. Qed. 202 | 203 | Lemma applyComposeRen E : 204 | (forall (v : Value E) E' E'' P ops (m:Map.Map P E' E'') (s : Ren E E'), 205 | Map.mapVal ops (composeRen m s) v = Map.mapVal ops m (renVal s v)) 206 | /\ (forall (e : Exp E) E' E'' P ops (m:Map.Map P E' E'') (s : Ren E E'), 207 | Map.mapExp ops (composeRen m s) e = Map.mapExp ops m (renExp s e)). 208 | Proof. 209 | move: E ; apply ExpValue_ind; intros; autorewrite with mapHints; Rewrites liftComposeRen. Qed. 210 | 211 | (*========================================================================== 212 | Substitution 213 | ==========================================================================*) 214 | 215 | Definition Sub := Map.Map Value. 216 | Definition SubOps : Map.Ops Value. refine (@Map.Build_Ops _ VAR (fun _ v => v) (fun E => renVal (fun v => FinS v)) _ _). by []. by []. Defined. 217 | 218 | Definition subVal := Map.mapVal SubOps. 219 | Definition subExp := Map.mapExp SubOps. 220 | Definition shiftSub := Map.shift SubOps. 221 | Definition liftSub := Map.lift SubOps. 222 | Definition idSub := Map.id SubOps. 223 | 224 | Ltac UnfoldRenSub := (unfold subVal; unfold subExp; unfold renVal; unfold renExp; unfold liftSub; unfold liftRen). 225 | Ltac FoldRenSub := (fold subVal; fold subExp; fold renVal; fold renExp; fold liftSub; fold liftRen). 226 | Ltac SimplMap := (UnfoldRenSub; autorewrite with mapHints; FoldRenSub). 227 | 228 | (*========================================================================== 229 | Composition of substitution followed by renaming 230 | ==========================================================================*) 231 | Definition composeRenSub E E' E'' (r : Ren E' E'') (s : Sub E E') : Sub E E'' := fun var => renVal r (s var). 232 | 233 | Lemma liftComposeRenSub : forall E E' E'' (r:Ren E' E'') (s:Sub E E'), liftSub (composeRenSub r s) = composeRenSub (liftRen r) (liftSub s). 234 | intros. apply Extensionality. intros var. dependent destruction var; first by []. 235 | unfold composeRenSub, liftSub. simpl. unfold renVal at 1 3. by do 2 rewrite <- (proj1 (applyComposeRen _)). 236 | Qed. 237 | 238 | Lemma applyComposeRenSub E : 239 | (forall (v:Value E) E' E'' (r:Ren E' E'') (s : Sub E E'), 240 | subVal (composeRenSub r s) v = renVal r (subVal s v)) 241 | /\ (forall (e:Exp E) E' E'' (r:Ren E' E'') (s : Sub E E'), 242 | subExp (composeRenSub r s) e = renExp r (subExp s e)). 243 | Proof. move: E ; apply ExpValue_ind; intros; SimplMap; Rewrites liftComposeRenSub. Qed. 244 | 245 | (*========================================================================== 246 | Composition of substitutions 247 | ==========================================================================*) 248 | 249 | Definition composeSub E E' E'' (s' : Sub E' E'') (s : Sub E E') : Sub E E'' := fun var => subVal s' (s var). 250 | 251 | Lemma liftComposeSub : forall E E' E'' (s' : Sub E' E'') (s : Sub E E'), liftSub (composeSub s' s) = composeSub (liftSub s') (liftSub s). 252 | intros. apply Extensionality. intros var. dependent destruction var; first by []. 253 | unfold composeSub. simpl. rewrite <- (proj1 (applyComposeRenSub _)). unfold composeRenSub, subVal. by rewrite <- (proj1 (applyComposeRen _)). 254 | Qed. 255 | 256 | Lemma applyComposeSub E : 257 | (forall (v : Value E) E' E'' (s' : Sub E' E'') (s : Sub E E'), 258 | subVal (composeSub s' s) v = subVal s' (subVal s v)) 259 | /\ (forall (e : Exp E) E' E'' (s' : Sub E' E'') (s : Sub E E'), 260 | subExp (composeSub s' s) e = subExp s' (subExp s e)). 261 | Proof. move: E ; apply ExpValue_ind; intros; SimplMap; Rewrites liftComposeSub. Qed. 262 | 263 | Lemma composeCons : forall E E' E'' (s':Sub E' E'') (s:Sub E E') (v:Value _), 264 | composeSub (cons v s') (liftSub s) = cons v (composeSub s' s). 265 | intros. apply Extensionality. intros var. dependent destruction var; first by []. 266 | unfold composeSub. unfold subVal. simpl. rewrite <- (proj1 (applyComposeRen _)). unfold composeRen. 267 | simpl. replace ((fun (var0:Fin E') => s' var0)) with s' by (by apply Extensionality). by []. 268 | Qed. 269 | 270 | Lemma composeSubIdLeft : forall E E' (s : Sub E E'), composeSub (@idSub _) s = s. 271 | Proof. intros. apply Extensionality. intros var. by apply (proj1 (Map.applyId _ _)). Qed. 272 | 273 | Lemma composeSubIdRight : forall E E' (s:Sub E E'), composeSub s (@idSub _) = s. 274 | Proof. intros. by apply Extensionality. Qed. 275 | 276 | Notation "[ x , .. , y ]" := (cons x .. (cons y (@Map.id _ SubOps _)) ..) : Sub_scope. 277 | Arguments Scope composeSub [_ _ _ Sub_scope Sub_scope]. 278 | Arguments Scope subExp [_ _ Sub_scope]. 279 | Arguments Scope subVal [_ _ Sub_scope]. 280 | Delimit Scope Sub_scope with sub. 281 | 282 | Lemma composeSingleSub : forall E E' (s:Sub E E') (v:Value _), composeSub [v] (liftSub s) = cons v s. 283 | Proof. intros. rewrite composeCons. by rewrite composeSubIdLeft. Qed. 284 | 285 | End Tm. 286 | -------------------------------------------------------------------------------- /src/KSTy.v: -------------------------------------------------------------------------------- 1 | (********************************************************************************** 2 | * KSTy.v * 3 | * Formalizing Domains, Ultrametric Spaces and Semantics of Programming Languages * 4 | * Nick Benton, Lars Birkedal, Andrew Kennedy and Carsten Varming * 5 | * Jan 2012 * 6 | * Build with Coq 8.3pl2 plus SSREFLECT * 7 | **********************************************************************************) 8 | 9 | (* Kitchen sink types *) 10 | 11 | Require Export ssreflect ssrnat. 12 | Require Import Program. 13 | Set Implicit Arguments. 14 | Unset Strict Implicit. 15 | Import Prenex Implicits. 16 | Require Import Fin. 17 | 18 | Ltac Rewrites E := 19 | (intros; simpl; try rewrite E; 20 | repeat (match goal with | [H : context[eq _ _] |- _] => rewrite H end); 21 | auto). 22 | 23 | Module Ty. 24 | 25 | Definition Env := nat. 26 | 27 | (*=Ty *) 28 | Inductive Ty E := 29 | | TVar: Fin E -> Ty E 30 | | Int: Ty E 31 | | Unit: Ty E 32 | | Product: Ty E -> Ty E -> Ty E 33 | | Sum: Ty E -> Ty E -> Ty E 34 | | Mu: Ty E.+1 -> Ty E 35 | | All: Ty E.+1 -> Ty E 36 | | Arrow: Ty E -> Ty E -> Ty E 37 | | Ref: Ty E -> Ty E. 38 | (*=End *) 39 | Implicit Arguments Unit [E]. 40 | Implicit Arguments Int [E]. 41 | 42 | Scheme Ty_ind2 := Induction for Ty Sort Prop. 43 | 44 | (*========================================================================== 45 | Variable-domain maps. 46 | By instantiating P with Var we get renamings. 47 | By instantiating P with Value we get substitutions. 48 | ==========================================================================*) 49 | Module Map. 50 | Section MAP. 51 | 52 | Variable P : Env -> Type. 53 | Definition Map E E' := FMap E (P E'). 54 | 55 | (*========================================================================== 56 | Package of operations used with a Map 57 | vr maps a Var into Var or Value (so is either the identity or TVAR) 58 | vl maps a Var or Value to a Value (so is either TVAR or the identity) 59 | wk weakens a Var or Value (so is either SVAR or renaming through SVAR on a value) 60 | ==========================================================================*) 61 | Record Ops := 62 | { 63 | vr : forall E, Fin E -> P E; 64 | vl : forall E, P E -> Ty E; 65 | wk : forall E, P E -> P E.+1; 66 | wkvr : forall E (var : Fin E), wk (vr var) = vr (FinS var); 67 | vlvr : forall E (var : Fin E), vl (vr var) = TVar var 68 | }. 69 | Variable ops : Ops. 70 | 71 | 72 | Definition lift E E' (m : Map E E') : Map E.+1 E'.+1 := 73 | (fun var => match var in Fin E return Map E.-1 E' -> P E'.+1 with 74 | | FinZ _ => fun _ => vr ops (FinZ _) 75 | | FinS _ x => fun m => wk ops (m x) 76 | end m). 77 | 78 | Definition shift E E' (m : Map E E') : Map E E'.+1 := fun var => wk ops (m var). 79 | Definition id E : Map E E := fun (var : Fin E) => vr ops var. 80 | 81 | Lemma shiftCons : forall E E' (m : Map E E') (x : P E'), shift (cons x m) = cons (wk ops x) (shift m). 82 | Proof. move => E E' m x. apply Extensionality. by dependent destruction i. Qed. 83 | 84 | Lemma liftAsCons : forall E E' (m : Map E' E), lift m = cons (vr ops (FinZ _)) (shift m). 85 | Proof. move => E E' m. apply Extensionality. by dependent destruction i. Qed. 86 | 87 | Fixpoint mapTy E E' (m : Map E E') (t : Ty E) : Ty E' := 88 | match t with 89 | | TVar v => vl ops (m v) 90 | | Int => Int 91 | | Unit => Unit 92 | | Product t1 t2 => Product (mapTy m t1) (mapTy m t2) 93 | | Sum t1 t2 => Sum (mapTy m t1) (mapTy m t2) 94 | | Mu t => Mu (mapTy (lift m) t) 95 | | All t => All (mapTy (lift m) t) 96 | | Arrow t1 t2 => Arrow (mapTy m t1) (mapTy m t2) 97 | | Ref t => Ref (mapTy m t) 98 | end. 99 | 100 | Variable E E' : Env. 101 | Variable m : Map E E'. 102 | Lemma mapTVar : forall (var : Fin _), mapTy m (TVar var) = vl ops (m var). by []. Qed. 103 | Lemma mapInt : mapTy m Int = Int. by []. Qed. 104 | Lemma mapUnit : mapTy m Unit = Unit. by []. Qed. 105 | Lemma mapMu : forall t, mapTy m (Mu t) = Mu (mapTy (lift m) t). by []. Qed. 106 | Lemma mapAll : forall t, mapTy m (All t) = All (mapTy (lift m) t). by []. Qed. 107 | Lemma mapProduct : forall t1 t2, mapTy m (Product t1 t2) = Product (mapTy m t1) (mapTy m t2). by []. Qed. 108 | Lemma mapSum : forall t1 t2, mapTy m (Sum t1 t2) = Sum (mapTy m t1) (mapTy m t2). by []. Qed. 109 | Lemma mapArrow : forall t1 t2, mapTy m (Arrow t1 t2) = Arrow (mapTy m t1) (mapTy m t2). by []. Qed. 110 | Lemma mapRef : forall t, mapTy m (Ref t) = Ref (mapTy m t). by []. Qed. 111 | 112 | Lemma liftId : lift (@id E) = @id E.+1. 113 | Proof. apply Extensionality. dependent destruction i; [by [] | by apply wkvr]. 114 | Qed. 115 | 116 | Lemma idAsCons : @id E.+1 = cons (vr ops (FinZ _)) (shift (@id E)). 117 | Proof. apply Extensionality. dependent destruction i; first by []. unfold id, shift. simpl. by rewrite wkvr. Qed. 118 | 119 | End MAP. 120 | 121 | Hint Rewrite mapTVar mapInt mapUnit mapMu mapAll mapProduct mapSum mapArrow mapRef : mapHints. 122 | Implicit Arguments id [P]. 123 | 124 | Lemma applyId P (ops:Ops P) E : (forall (t : Ty E), mapTy ops (id ops E) t = t). 125 | Proof. induction t; intros; autorewrite with mapHints; Rewrites liftId. by apply vlvr. Qed. 126 | 127 | End Map. 128 | 129 | (*========================================================================== 130 | Variable renamings: Map Var 131 | ==========================================================================*) 132 | 133 | Definition Ren := Map.Map Fin. 134 | Definition RenOps : Map.Ops Fin. refine (@Map.Build_Ops _ (fun _ v => v) TVar FinS _ _). by []. by []. Defined. 135 | 136 | Definition renTy := Map.mapTy RenOps. 137 | Definition liftRen := Map.lift RenOps. 138 | Definition shiftRen := Map.shift RenOps. 139 | Definition idRen := Map.id RenOps. 140 | 141 | (*========================================================================== 142 | Composition of renaming 143 | ==========================================================================*) 144 | 145 | Definition composeRen P E E' E'' (m : Map.Map P E' E'') (r : Ren E E') : Map.Map P E E'' := fun var => m (r var). 146 | 147 | Lemma liftComposeRen : forall E E' E'' P ops (m:Map.Map P E' E'') (r:Ren E E'), Map.lift ops (composeRen m r) = composeRen (Map.lift ops m) (liftRen r). 148 | Proof. move => E E' E'' P ops m r. apply Extensionality. by dependent destruction i. Qed. 149 | 150 | Lemma applyComposeRen E : 151 | (forall (t : Ty E) E' E'' P ops (m:Map.Map P E' E'') (s : Ren E E'), 152 | Map.mapTy ops (composeRen m s) t = Map.mapTy ops m (renTy s t)). 153 | Proof. induction t; intros; autorewrite with mapHints; Rewrites liftComposeRen. Qed. 154 | 155 | (*========================================================================== 156 | Substitution 157 | ==========================================================================*) 158 | 159 | Definition Sub := Map.Map Ty. 160 | Definition SubOps : Map.Ops Ty. refine (@Map.Build_Ops _ TVar (fun _ v => v) (fun E => renTy (fun v => FinS v)) _ _). by []. by []. Defined. 161 | 162 | Definition subTy := Map.mapTy SubOps. 163 | Definition shiftSub := Map.shift SubOps. 164 | Definition liftSub := Map.lift SubOps. 165 | Definition idSub := Map.id SubOps. 166 | 167 | Ltac UnfoldRenSub := (unfold subTy; unfold renTy; unfold liftSub; unfold liftRen). 168 | Ltac FoldRenSub := (fold subTy; fold renTy; fold liftSub; fold liftRen). 169 | Ltac SimplMap := (UnfoldRenSub; autorewrite with mapHints; FoldRenSub). 170 | 171 | (*========================================================================== 172 | Composition of substitution followed by renaming 173 | ==========================================================================*) 174 | Definition composeRenSub E E' E'' (r : Ren E' E'') (s : Sub E E') : Sub E E'' := fun var => renTy r (s var). 175 | 176 | Lemma liftComposeRenSub : forall E E' E'' (r:Ren E' E'') (s:Sub E E'), liftSub (composeRenSub r s) = composeRenSub (liftRen r) (liftSub s). 177 | intros. apply Extensionality. dependent destruction i; first by []. 178 | unfold composeRenSub, liftSub. simpl. unfold renTy at 1 3. by do 2 rewrite <- applyComposeRen. 179 | Qed. 180 | 181 | Lemma applyComposeRenSub E : 182 | (forall (t:Ty E) E' E'' (r:Ren E' E'') (s : Sub E E'), 183 | subTy (composeRenSub r s) t = renTy r (subTy s t)). 184 | Proof. induction t; intros; SimplMap; Rewrites liftComposeRenSub. Qed. 185 | 186 | (*========================================================================== 187 | Composition of substitutions 188 | ==========================================================================*) 189 | 190 | Definition composeSub E E' E'' (s' : Sub E' E'') (s : Sub E E') : Sub E E'' := fun var => subTy s' (s var). 191 | 192 | Lemma liftComposeSub : forall E E' E'' (s' : Sub E' E'') (s : Sub E E'), liftSub (composeSub s' s) = composeSub (liftSub s') (liftSub s). 193 | intros. apply Extensionality. dependent destruction i; first by []. 194 | unfold composeSub. simpl. rewrite <- applyComposeRenSub. unfold composeRenSub, subTy. by rewrite <- applyComposeRen. 195 | Qed. 196 | 197 | Lemma applyComposeSub E : 198 | (forall (t : Ty E) E' E'' (s' : Sub E' E'') (s : Sub E E'), 199 | subTy (composeSub s' s) t = subTy s' (subTy s t)). 200 | Proof. induction t; intros; SimplMap; Rewrites liftComposeSub. Qed. 201 | 202 | Lemma composeCons : forall E E' E'' (s':Sub E' E'') (s:Sub E E') (v:Ty _), 203 | composeSub (cons v s') (liftSub s) = cons v (composeSub s' s). 204 | intros. apply Extensionality. dependent destruction i; first by []. 205 | unfold composeSub. unfold subTy. simpl. rewrite <- applyComposeRen. unfold composeRen. 206 | simpl. replace ((fun (var0:Fin E') => s' var0)) with s' by (by apply Extensionality). by []. 207 | Qed. 208 | 209 | Lemma composeSubIdLeft : forall E E' (s : Sub E E'), composeSub (@idSub _) s = s. 210 | Proof. intros. apply Extensionality. intros var. by apply Map.applyId. Qed. 211 | 212 | Lemma composeSubIdRight : forall E E' (s:Sub E E'), composeSub s (@idSub _) = s. 213 | Proof. intros. by apply Extensionality. Qed. 214 | 215 | Notation "[ x , .. , y ]" := (cons x .. (cons y (@Map.id _ SubOps _)) ..) : Sub_scope. 216 | Arguments Scope composeSub [_ _ _ Sub_scope Sub_scope]. 217 | Arguments Scope subTy [_ _ Sub_scope]. 218 | Delimit Scope Sub_scope with sub. 219 | 220 | Lemma composeSingleSub : forall E E' (s:Sub E E') (t:Ty _), composeSub [t] (liftSub s) = cons t s. 221 | Proof. intros. rewrite composeCons. by rewrite composeSubIdLeft. Qed. 222 | 223 | End Ty. 224 | 225 | Notation "a --> b" := (Ty.Arrow a b) (at level 55, right associativity). 226 | Notation "a ** b" := (Ty.Product a b) (at level 55). 227 | 228 | 229 | 230 | -------------------------------------------------------------------------------- /src/KSTyping.v: -------------------------------------------------------------------------------- 1 | (********************************************************************************** 2 | * KSTyping.v * 3 | * Formalizing Domains, Ultrametric Spaces and Semantics of Programming Languages * 4 | * Nick Benton, Lars Birkedal, Andrew Kennedy and Carsten Varming * 5 | * Jan 2012 * 6 | * Build with Coq 8.3pl2 plus SSREFLECT * 7 | **********************************************************************************) 8 | 9 | (* Typing relation for kitchen sink language *) 10 | 11 | Require Export ssreflect ssrnat. Require Import KSTy. Require Import KSTm. 12 | Set Implicit Arguments. 13 | Unset Strict Implicit. 14 | Import Prenex Implicits. 15 | Require Import Fin. 16 | 17 | Definition StoreType E := nat -> Ty.Ty E. 18 | 19 | Definition TEnv E n := FMap n (Ty.Ty E). 20 | Definition shiftTy E : Ty.Ty E -> Ty.Ty E.+1 := Ty.renTy (fun v => FinS v). 21 | Definition subOneTy E (t' : Ty.Ty E) t := Ty.subTy (cons t' (@Ty.idSub _)) t. 22 | Definition unfoldTy E (t : Ty.Ty E.+1) := subOneTy (Ty.Mu t) t. 23 | 24 | Require Import Program. 25 | 26 | Definition emptyTEnv : TEnv O O. 27 | move => var. dependent destruction var. 28 | Defined. 29 | 30 | Import Tm. 31 | Inductive VTy E (se:StoreType E) : forall n, TEnv E n -> Tm.Value n -> Ty.Ty E -> Type := 32 | | TvVAR: forall n (env : TEnv E n) v, VTy se env (VAR v) (env v) 33 | | TvLOC: forall n (env : TEnv E n) l, VTy se env (LOC l) (Ty.Ref (se l)) 34 | | TvINT: forall n (env : TEnv E n) i , VTy se env (INT i) Ty.Int 35 | | TvUNIT: forall n (env : TEnv E n), VTy se env UNIT Ty.Unit 36 | | TvLAM: forall n (env : TEnv E n) t1 t2 e, ETy se (cons t1 env) e t2 -> VTy se env (LAM e) (t1 --> t2) 37 | | TvTLAM: forall n (env : TEnv E n) t e, ETy (E:=E.+1) (fun l => shiftTy (se l)) (fun v => shiftTy (env v)) e t -> VTy se env (TLAM e) (Ty.All t) 38 | | TvPAIR: forall n (env : TEnv E n) t1 t2 e1 e2, VTy se env e1 t1 -> VTy se env e2 t2 -> VTy se env (PAIR e1 e2) (t1 ** t2) 39 | | TvINL: forall n (env:TEnv E n) v t1 t2, VTy se env v t1 -> VTy se env (INL v) (Ty.Sum t1 t2) 40 | | TvINR: forall n (env:TEnv E n) v t1 t2, VTy se env v t2 -> VTy se env (INR v) (Ty.Sum t1 t2) 41 | | TvFOLD: forall n (env:TEnv E n) v t, VTy se env v (unfoldTy t) -> VTy se env (FOLD v) (Ty.Mu t) 42 | 43 | with ETy E (se:StoreType E) : forall n, TEnv E n -> Tm.Exp n -> Ty.Ty E -> Type := 44 | | TeVAL: forall n (env:TEnv E n) v t, VTy se env v t -> ETy se env (VAL v) t 45 | | TeLET: forall n (env:TEnv E n) e1 e2 t1 t2, ETy se env e1 t1 -> ETy se (cons t1 env) e2 t2 -> ETy se env (LET e1 e2) t2 46 | | TvFST: forall n (env:TEnv E n) v t1 t2, VTy se env v (t1**t2) -> ETy se env (FST v) t1 47 | | TvSND: forall n (env:TEnv E n) v t1 t2, VTy se env v (t1**t2) -> ETy se env (SND v) t2 48 | | TvOP: forall n (env:TEnv E n) op v, VTy se env v (Ty.Int**Ty.Int) -> ETy se env (OP op v) Ty.Int 49 | | TvUNFOLD: forall n (env:TEnv E n) t v, VTy se env v (Ty.Mu t) -> ETy se env (UNFOLD v) (unfoldTy t) 50 | | TvREF: forall n (env:TEnv E n) v t, VTy se env v t -> ETy se env (REF v) (Ty.Ref t) 51 | | TvBANG: forall n (env:TEnv E n) v t, VTy se env v (Ty.Ref t) -> ETy se env (BANG v) t 52 | | TvASSIGN: forall n (env:TEnv E n) v1 v2 t, VTy se env v1 (Ty.Ref t) -> VTy se env v2 t -> ETy se env (ASSIGN v1 v2) Ty.Unit 53 | | TeAPP: forall n (env:TEnv E n) t1 t2 v1 v2, VTy se env v1 (t1 --> t2) -> VTy se env v2 t1 -> ETy se env (APP v1 v2) t2 54 | | TeTAPP: forall n (env:TEnv E n) t v t', VTy se env v (Ty.All t) -> ETy se env (TAPP v) (subOneTy t' t) 55 | | TeCASE: forall n (env:TEnv E n) t1 t2 v e1 e2 t, VTy se env v (Ty.Sum t1 t2) -> ETy se (cons t1 env) e1 t -> ETy se (cons t2 env) e2 t -> ETy se env (CASE v e1 e2) t 56 | . 57 | 58 | Hint Resolve TvINT TeAPP TeCASE TeLET TeVAL TvVAR TvTLAM TLAM TvOP TvINL TvINR TvPAIR TvSND TvFST TvASSIGN TvBANG TvREF. 59 | 60 | Scheme VTy_rec2 := Induction for VTy Sort Set 61 | with ETy_rec2 := Induction for ETy Sort Set. 62 | 63 | Scheme VTy_rect2 := Induction for VTy Sort Type 64 | with ETy_rect2 := Induction for ETy Sort Type. 65 | 66 | Scheme VTy_ind2 := Induction for VTy Sort Prop 67 | with ETy_ind2 := Induction for ETy Sort Prop. 68 | 69 | Combined Scheme Typing_ind from VTy_ind2, ETy_ind2. 70 | 71 | -------------------------------------------------------------------------------- /src/Makefile: -------------------------------------------------------------------------------- 1 | ############################################################################# 2 | ## v # The Coq Proof Assistant ## 3 | ## "$@" || ( RV=$$?; rm -f "$@"; exit $${RV} ) 248 | 249 | %.v.beautified: 250 | $(COQC) $(COQDEBUG) $(COQFLAGS) -beautify $* 251 | 252 | # WARNING 253 | # 254 | # This Makefile has been automagically generated 255 | # Edit at your own risks ! 256 | # 257 | # END OF WARNING 258 | 259 | -------------------------------------------------------------------------------- /src/Makefile.win: -------------------------------------------------------------------------------- 1 | COQFLAGS=-coqlib "$(COQLIB)" 2 | COQDOCFLAGS=-coqlib_path "$(COQLIB)" --latex -g 3 | COQC="$(COQBIN)"/coqc $(COQFLAGS) 4 | COQDOC="$(COQBIN)"/coqdoc $(COQDOCFLAGS) 5 | 6 | 7 | .SUFFIXES: .v .vo .tex 8 | 9 | .v.vo: 10 | $(COQC) $* 11 | .v.tex: 12 | $(COQDOC) $*.v 13 | 14 | 15 | Categories.vo: Categories.v 16 | PredomCore.vo: Categories.vo PredomCore.v 17 | NSetoid.vo: Categories.vo NSetoid.v 18 | PredomSum.vo: PredomSum.v PredomCore.vo 19 | PredomLift.vo: PredomLift.v PredomCore.vo 20 | PredomKleisli.vo: PredomKleisli.v PredomLift.vo PredomCore.vo 21 | PredomFix.vo: PredomFix.v PredomCore.vo PredomLift.vo PredomKleisli.vo 22 | KnasterTarski.vo: KnasterTarski.v PredomCore.vo 23 | PredomAll.vo: PredomAll.v PredomCore.vo PredomSum.vo PredomLift.vo PredomKleisli.vo PredomFix.vo 24 | PredomRec.vo: PredomRec.v Categories.vo PredomCore.vo NSetoid.vo 25 | 26 | typedlambda.vo: typedlambda.v 27 | typedopsem.vo: typedopsem.v typedlambda.vo 28 | typeddensem.vo: typeddensem.v PredomAll.vo typedlambda.vo 29 | typedsubst.vo: typedsubst.v typeddensem.vo 30 | typedsoundness.vo: typedsoundness.v typedsubst.vo typedopsem.vo 31 | typedadequacy.vo: typedadequacy.v typeddensem.vo typedopsem.vo 32 | 33 | unii.vo: unii.v Fin.vo 34 | uniirec.vo: PredomAll.vo PredomRec.vo 35 | uniisem.vo: uniisem.v PredomAll.vo unii.vo uniirec.vo Fin.vo 36 | uniiop.vo: uniiop.v unii.vo 37 | uniisound.vo: uniisound.v uniisem.vo uniiop.vo Fin.vo 38 | uniiade.vo: uniiade.v uniisem.vo uniiop.vo KnasterTarski.vo Fin.vo 39 | 40 | 41 | MetricCore.vo: MetricCore.v NSetoid.vo Categories.vo 42 | Finmap.vo: Finmap.v NSetoid.vo MetricCore.vo 43 | msyntax.vo: msyntax.v Finmap.vo 44 | moperational.vo: moperational.v msyntax.vo Finmap.vo 45 | msem.vo: msem.v mpremet.vo msyntax.vo moperational.vo 46 | mpremet.vo: mpremet.v PredomCore.vo MetricCore.vo Finmap.vo MetricRec.vo FinmapMetric.vo 47 | FinmapMetric.vo: FinmapMetric.v NSetoid.vo MetricCore.vo FinMap.vo 48 | MetricRec.vo: MetricRec.v Categories.vo MetricCore.vo NSetoid.vo 49 | 50 | Fin.vo: Fin.v 51 | KSTy.vo: KSTy.v Fin.vo 52 | KSTm.vo: KSTm.v Fin.vo 53 | KSTyping.vo: KSTyping.v KSTy.vo KSTm.vo Fin.vo 54 | KSOp.vo: KSOp.v KSTm.vo Fin.vo FinMap.vo 55 | KSSem.vo: KSSem.v Fin.vo KSTm.vo KSTy.vo KSTyping.vo KSOp.vo MetricRec.vo mpremet.vo 56 | 57 | all: typedadequacy.vo typedsoundness.vo uniisound.vo uniiade.vo KSSem.vo 58 | 59 | 60 | 61 | 62 | -------------------------------------------------------------------------------- /src/NSetoid.v: -------------------------------------------------------------------------------- 1 | (********************************************************************************** 2 | * NSetoid.v * 3 | * Formalizing Domains, Ultrametric Spaces and Semantics of Programming Languages * 4 | * Nick Benton, Lars Birkedal, Andrew Kennedy and Carsten Varming * 5 | * Jan 2012 * 6 | * Build with Coq 8.3pl2 plus SSREFLECT * 7 | **********************************************************************************) 8 | 9 | Require Export Categories. 10 | 11 | Set Implicit Arguments. 12 | Unset Strict Implicit. 13 | Import Prenex Implicits. 14 | 15 | (** printing >-> %\ensuremath{>}\ensuremath{-}\ensuremath{>}% *) 16 | (** printing :> %\ensuremath{:>}% *) 17 | (** printing Canonical %\coqdockw{Canonical}% *) 18 | 19 | Open Scope S_scope. 20 | 21 | Definition setoid_respect (T T':setoidType) (f: T -> T') := 22 | forall x y, x =-= y -> (f x) =-= (f y). 23 | 24 | Module FSet. 25 | 26 | Section fset. 27 | Variable O1 O2 : setoidType. 28 | 29 | Record mixin_of (f:O1 -> O2) := Mixin { ext :> setoid_respect f}. 30 | 31 | Notation class_of := mixin_of (only parsing). 32 | 33 | Section ClassDef. 34 | Structure type : Type := Pack {sort : O1 -> O2; _ : class_of sort; _ : O1 -> O2}. 35 | Local Coercion sort : type >-> Funclass. 36 | Definition class cT := let: Pack _ c _ := cT return class_of cT in c. 37 | Definition unpack K (k : forall T (c : class_of T), K T c) cT := 38 | let: Pack T c _ := cT return K _ (class cT) in k _ c. 39 | Definition repack cT : _ -> Type -> type := let k T c p := p c in unpack k cT. 40 | Definition pack f (c:class_of f) := @Pack f c f. 41 | End ClassDef. 42 | End fset. 43 | Module Import Exports. 44 | Coercion sort : type >-> Funclass. 45 | Notation fset := type. 46 | Notation fsetMixin := Mixin. 47 | Notation fsetType := pack. 48 | End Exports. 49 | End FSet. 50 | Export FSet.Exports. 51 | 52 | Lemma frespect (S S':setoidType) (f:fset S S') : setoid_respect f. 53 | case:f => f. case. move => a s. by apply a. 54 | Qed. 55 | 56 | Hint Resolve frespect. 57 | 58 | Lemma fset_setoidAxiom (T T' : setoidType) : Setoid.axiom (fun (f g:fset T T') => forall x:T, (f x) =-= (g x)). 59 | split ; last split. 60 | - by move => f. 61 | - move => f g h. simpl. move => X Y e. by apply (tset_trans (X e) (Y e)). 62 | - move => f g. simpl. move => X e. by apply (tset_sym (X e)). 63 | Qed. 64 | 65 | Canonical Structure fset_setoidMixin (T T':setoidType) := Setoid.Mixin (fset_setoidAxiom T T'). 66 | Canonical Structure fset_setoidType T T' := Eval hnf in SetoidType (fset_setoidMixin T T'). 67 | 68 | Definition setmorph_equal (A B: setoidType) : relation (fset A B) := 69 | (@tset_eq (@fset_setoidType A B)). 70 | 71 | Definition setmorph_equal_equivalence (A B: setoidType): Equivalence 72 | (@setmorph_equal A B). 73 | split. 74 | - move => x. by apply: tset_refl. 75 | - move => x y. by apply: tset_sym. 76 | - move => x y z. by apply: tset_trans. 77 | Defined. 78 | 79 | Existing Instance setmorph_equal_equivalence. 80 | 81 | Add Parametric Morphism (A B : setoidType) : 82 | (@FSet.sort A B) with signature (@setmorph_equal A B ==> @tset_eq A ==> 83 | @tset_eq B) 84 | as setmorph_eq_compat. 85 | intros f g fg x y xy. apply tset_trans with (g x) ; first by auto. 86 | by apply frespect. 87 | Qed. 88 | 89 | Definition mk_fsetM (O1 O2:setoidType) (f:O1 -> O2) (m:setoid_respect f) := fsetMixin m. 90 | Definition mk_fset (O1 O2:setoidType) (f:O1 -> O2) (m:setoid_respect f) : fset O1 O2 := 91 | Eval hnf in fsetType (mk_fsetM m). 92 | 93 | Lemma ScompP S S' S'' (f:fset S' S'') (g:fset S S') : setoid_respect (T:=S) (T':=S'') (fun x : S => f (g x)). 94 | move => x y e. have a:=frespect g e. by apply (frespect f a). 95 | Qed. 96 | 97 | Definition Scomp S S' S'' (f:fset S' S'') (g:fset S S') : fset S S'' := 98 | Eval hnf in mk_fset (ScompP f g). 99 | 100 | Lemma SidP (S:setoidType) : setoid_respect (id : S -> S). 101 | move => x y X. by apply X. 102 | Qed. 103 | 104 | Definition Sid S : (fset S S) := Eval hnf in mk_fset (@SidP S). 105 | 106 | Lemma setoidCatAxiom : Category.axiom Scomp Sid. 107 | split ; last split ; last split. 108 | - by move => S S' x. 109 | - by move => S S' x. 110 | - by move => S0 S1 S2 S3 f g h x. 111 | - move => S0 S1 S2 f f' g g' e e' x. apply tset_trans with (y:= f (g' x)). 112 | + apply: (frespect f). by apply e'. 113 | + by apply e. 114 | Qed. 115 | 116 | Canonical Structure setoidCatMixin := CatMixin setoidCatAxiom. 117 | Canonical Structure setoidCat := Eval hnf in CatType setoidCatMixin. 118 | 119 | Open Scope C_scope. 120 | 121 | Lemma PropP : Setoid.axiom (fun f g : Prop => iff f g). 122 | split ; last split. 123 | - by move => x. 124 | - move => x y z. simpl. move => [X Y] [Z W]. by split ; auto. 125 | - move => x y. simpl. move => [X Y]. by split. 126 | Qed. 127 | 128 | Canonical Structure prop_setoidMixin := Setoid.Mixin PropP. 129 | Canonical Structure prop_setoidType := Eval hnf in SetoidType prop_setoidMixin. 130 | 131 | Notation Props := prop_setoidType. 132 | 133 | Section SetoidProducts. 134 | Variable A B:setoidType. 135 | 136 | Lemma sum_setoidAxiom : Setoid.axiom (fun p0 p1 : A + B => match p0,p1 with inl a, inl b => a =-= b | inr a, inr b => a =-= b | _,_ => False end). 137 | split ; last split. 138 | - by case. 139 | - case => a ; case => b ; case => c ; try done ; by apply tset_trans. 140 | - case => a ; case => b ; try done ; by apply tset_sym. 141 | Qed. 142 | 143 | Canonical Structure sum_setoidMixin := Setoid.Mixin sum_setoidAxiom. 144 | Canonical Structure sum_setoidType := Eval hnf in SetoidType sum_setoidMixin. 145 | 146 | Lemma prod_setoidAxiom : Setoid.axiom (fun p0 p1 : A * B => tset_eq (fst p0) (fst p1) /\ tset_eq (snd p0) (snd p1)). 147 | split ; last split. 148 | - case => a b ; by split. 149 | - case => a b ; case => c d ; case => e f ; simpl ; move => [X Y] [Z W]. 150 | by split ; [apply (tset_trans X Z) | apply (tset_trans Y W)]. 151 | - by case => a b ; case => c d ; simpl ; move => [X Y] ; split ; auto. 152 | Qed. 153 | 154 | Canonical Structure prod_setoidMixin := Setoid.Mixin prod_setoidAxiom. 155 | Canonical Structure prod_setoidType := Eval hnf in SetoidType prod_setoidMixin. 156 | 157 | Lemma Sfst_respect : setoid_respect (T:=prod_setoidType) (T':=A) (@fst _ _). 158 | case => a b ; case => c d ; simpl ; unlock tset_eq ; move => [X Y] ; by unlock tset_eq in X ; apply X. 159 | Qed. 160 | 161 | Lemma Ssnd_respect : setoid_respect (T:=prod_setoidType) (T':=B) (@snd _ _). 162 | case => a b ; case => c d ; simpl ; unlock tset_eq ; move => [X Y] ; by unlock tset_eq in Y ; apply Y. 163 | Qed. 164 | 165 | Definition Sfst : prod_setoidType =-> A := Eval hnf in mk_fset Sfst_respect. 166 | 167 | Definition Ssnd : prod_setoidType =-> B := Eval hnf in mk_fset Ssnd_respect. 168 | 169 | Lemma prod_fun_respect C (f:C =-> A) (g:C =-> B) : setoid_respect (T:=C) (T':=prod_setoidType) (fun c : C => (f c, g c)). 170 | move => c c' e. simpl. unlock tset_eq. by split ; [apply (frespect f e) | apply (frespect g e)]. 171 | Qed. 172 | 173 | Definition Sprod_fun C (f:C =-> A) (g:C =-> B) : C =-> prod_setoidType := Eval hnf in mk_fset (prod_fun_respect f g). 174 | 175 | Lemma Sin1_respect : @setoid_respect A (sum_setoidType) (@inl _ _). 176 | move => a a' e. by apply e. 177 | Qed. 178 | 179 | Lemma Sin2_respect : @setoid_respect B (sum_setoidType) (@inr _ _). 180 | move => a a' e. by apply e. 181 | Qed. 182 | 183 | Definition Sin1 : A =-> sum_setoidType := Eval hnf in mk_fset Sin1_respect. 184 | Definition Sin2 : B =-> sum_setoidType := Eval hnf in mk_fset Sin2_respect. 185 | 186 | Lemma Ssum_fun_respect C (f:A =-> C) (g:B =-> C) : @setoid_respect sum_setoidType C (fun x => match x with inl x => f x | inr x => g x end). 187 | case => x ; case => y => e ; try done ; by apply frespect ; apply e. 188 | Qed. 189 | 190 | Definition Ssum_fun C (f:A =-> C) (g:B =-> C) : sum_setoidType =-> C := Eval hnf in mk_fset (Ssum_fun_respect f g). 191 | 192 | 193 | (* 194 | 195 | Lemma Sprod_fun_unique C (f:C -s> A) (g:C -s> B) (h:C -s> prod_setoidType) : 196 | Scomp Sfst h =-= f -> Scomp Ssnd h =-= g -> h =-= Sprod_fun f g. 197 | move => C f g h. unlock tset_eq. move => X Y x. specialize (X x). specialize (Y x). unlock tset_eq. by split ; auto. 198 | Qed. 199 | 200 | Lemma Sprod_fun_fst C f g : Sfst << (@Sprod_fun C f g) =-= f. 201 | move => C f g. unlock tset_eq. by simpl. 202 | Qed. 203 | 204 | Lemma Sprod_fun_snd C f g : Ssnd << (@Sprod_fun C f g) =-= g. 205 | move => C f g. unlock tset_eq. by simpl. 206 | Qed.*) 207 | 208 | End SetoidProducts. 209 | 210 | Lemma setoidProdCatAxiom : @CatProduct.axiom _ prod_setoidType Sfst Ssnd Sprod_fun. 211 | move => S0 S1 S2 f g. split ; [by split | move => m]. 212 | move => [X Y] x. simpl. specialize (X x). by specialize (Y x). 213 | Qed. 214 | 215 | Canonical Structure setoidProdCatMixin := prodCatMixin setoidProdCatAxiom. 216 | Canonical Structure setoidProdCat := Eval hnf in prodCatType (CatProduct.Class setoidProdCatMixin). 217 | 218 | Lemma setoidSumCatAxiom : @CatSum.axiom _ sum_setoidType Sin1 Sin2 Ssum_fun. 219 | move => S0 S1 S2 f g. split ; first by split. 220 | move => h [X Y]. by case => x ; simpl ; [apply: X | apply: Y]. 221 | Qed. 222 | 223 | Canonical Structure setoidSumCatMixin := sumCatMixin setoidSumCatAxiom. 224 | Canonical Structure setoidSumCat := Eval hnf in sumCatType setoidSumCatMixin. 225 | 226 | Section SetoidIProducts. 227 | 228 | Variable I : Type. 229 | Variable P : I -> setoidType. 230 | 231 | Lemma prodI_setoidP : Setoid.axiom (fun (p0 p1:forall i, P i) => forall i:I, @tset_eq _ (p0 i) (p1 i)). 232 | split ; last split. 233 | - move => X ; simpl ; by []. 234 | - move => X Y Z ; simpl ; move => R S i. by apply (tset_trans (R i) (S i)). 235 | - move => X Y ; simpl ; move => R i. by apply (tset_sym (R i)). 236 | Qed. 237 | 238 | Canonical Structure prodI_setoidMixin := Setoid.Mixin prodI_setoidP. 239 | Canonical Structure prodI_setoidType := Eval hnf in SetoidType prodI_setoidMixin. 240 | 241 | Lemma Sproj_respect i : setoid_respect (T:=prodI_setoidType) (T':=P i) 242 | (fun X : prodI_setoidType => X i). 243 | move => a b e. simpl. unlock tset_eq in e. by apply (e i). 244 | Qed. 245 | 246 | Definition Sproj i : prodI_setoidType =-> P i := Eval hnf in mk_fset (Sproj_respect i). 247 | 248 | Lemma SprodI_fun_respect C (f:forall i, C =-> P i) : setoid_respect (T:=C) (T':=prodI_setoidType) (fun (c : C) (i : I) => f i c). 249 | move => c c' e. simpl. unlock tset_eq. by move => i ; apply (frespect (f i) e). 250 | Qed. 251 | 252 | Definition SprodI_fun C (f:forall i, C =-> P i) : C =-> prodI_setoidType := Eval hnf in mk_fset (SprodI_fun_respect f). 253 | 254 | Lemma SprodI_fun_proj C (f:forall i, C =-> P i) i : Scomp (Sproj i) (SprodI_fun f) =-= f i. 255 | by unlock tset_eq ; simpl. 256 | Qed. 257 | 258 | Lemma SprodI_fun_unique C (f:forall i, C =-> P i) (h:C =-> prodI_setoidType) : (forall i, Scomp (Sproj i) h =-= f i) -> h =-= SprodI_fun f. 259 | move => X. unlock tset_eq. move => x. unlock tset_eq. move => i. unlock tset_eq in X. apply (X i x). 260 | Qed. 261 | 262 | End SetoidIProducts. 263 | 264 | Add Parametric Morphism A B : (@pair A B) 265 | with signature (@tset_eq A) ==> (@tset_eq B) ==> (@tset_eq (A * B)) 266 | as pair_eq_compat. 267 | by move => x y e x' y' e'. 268 | Qed. 269 | 270 | Lemma spair_respect (A B : setoidType) (x:A) : setoid_respect (fun y:B => (x,y)). 271 | move => y y' e. by split. 272 | Qed. 273 | 274 | Definition spair (A B : setoidType) (x:A) : B =-> A * B := Eval hnf in mk_fset (spair_respect x). 275 | 276 | Lemma scurry_respect (A B C : setoidType) (f:A * B =-> C) : setoid_respect (fun x => f << spair _ x). 277 | move => x x' e y. simpl. by rewrite -> e. 278 | Qed. 279 | 280 | Definition scurry (A B C : setoidType) (f:A * B =-> C) : A =-> (fset_setoidType B C) := Eval hnf in mk_fset (scurry_respect f). 281 | 282 | Definition setoid_respect2 (S S' S'' : setoidType) (f:S -> S' -> S'') := 283 | (forall s', setoid_respect (fun x => f x s')) /\ forall s, setoid_respect (f s). 284 | 285 | Lemma setoid2_morphP (S S' S'' : setoidType) (f:S -> S' -> S'') (C:setoid_respect2 f) : 286 | setoid_respect (T:=S) (T':=S' =-> S'') 287 | (fun x : S => FSet.Pack (FSet.Mixin (proj2 C x)) (f x)). 288 | move => x y E. unfold tset_eq. simpl. move => s'. apply (proj1 C s' _ _ E). 289 | Qed. 290 | 291 | Definition setoid2_morph (S S' S'' : setoidType) (f:S -> S' -> S'') (C:setoid_respect2 f) : 292 | S =-> (fset_setoidType S' S'') := Eval hnf in mk_fset (setoid2_morphP C). 293 | 294 | Definition Sprod_map (A B C D:setoidType) (f:A =-> C) (g:B =-> D) := <| f << pi1, g << pi2 |>. 295 | 296 | Lemma ev_respect B A : setoid_respect (fun (p:(fset_setoidType B A) * B) => fst p (snd p)). 297 | case => f b. case => f' b'. simpl. 298 | case ; simpl. move => e e'. rewrite -> e. by rewrite -> e'. 299 | Qed. 300 | 301 | Definition Sev B A : (fset_setoidType B A) * B =-> A := Eval hnf in mk_fset (@ev_respect B A). 302 | 303 | Implicit Arguments Sev [A B]. 304 | 305 | Lemma setoidExpAxiom : @CatExp.axiom _ fset_setoidType (@Sev) (@scurry). 306 | move => X Y Z h. split ; first by case. 307 | move => h' A x y. specialize (A (x,y)). by apply A. 308 | Qed. 309 | 310 | Canonical Structure setoidExpMixin := expCatMixin setoidExpAxiom. 311 | Canonical Structure setoidExpCat := Eval hnf in expCatType setoidExpMixin. 312 | 313 | Lemma const_respect (M M':setoidType) (x:M') : setoid_respect (fun _ : M => x). 314 | by []. 315 | Qed. 316 | 317 | Definition sconst (M M':setoidType) (x:M') : M =-> M' := Eval hnf in mk_fset (const_respect x). 318 | 319 | Lemma unit_setoidP : Setoid.axiom (fun _ _ : unit => True). 320 | by []. 321 | Qed. 322 | 323 | Canonical Structure unit_setoidMixin := Setoid.Mixin unit_setoidP. 324 | Canonical Structure unit_setoidType := Eval hnf in SetoidType unit_setoidMixin. 325 | 326 | Lemma setoidTerminalAxiom : CatTerminal.axiom unit_setoidType. 327 | move => C h h' x. case (h x). by case (h' x). 328 | Qed. 329 | 330 | Canonical Structure setoidTerminalMixin := terminalCatMixin (fun T => sconst T tt) setoidTerminalAxiom. 331 | Canonical Structure setoidTerminalCat := Eval hnf in terminalCatType (CatTerminal.Class setoidTerminalMixin). 332 | 333 | Lemma emp_setoid_axiom : Setoid.axiom (fun x y : Empty => True). 334 | split ; last split ; by move => x. 335 | Qed. 336 | 337 | Canonical Structure emp_setoidMixin := SetoidMixin emp_setoid_axiom. 338 | Canonical Structure emp_setoidType := Eval hnf in SetoidType emp_setoidMixin. 339 | 340 | Lemma setoidInitialAxiom : CatInitial.axiom emp_setoidType. 341 | move => C h h'. by case. 342 | Qed. 343 | 344 | Lemma SZero_initial_unique D (f g : emp_setoidType =-> D) : f =-= g. 345 | by move => x. 346 | Qed. 347 | 348 | Lemma SZero_respect (D:setoidType) : setoid_respect (fun (x:Empty) => match x with end : D). 349 | by []. 350 | Qed. 351 | 352 | Definition SZero_initial D : emp_setoidType =-> D := Eval hnf in mk_fset (SZero_respect D). 353 | 354 | Canonical Structure setoidInitialMixin := initialCatMixin SZero_initial setoidInitialAxiom. 355 | Canonical Structure setoidInitialCat := Eval hnf in initialCatType (CatInitial.Class setoidInitialMixin). 356 | 357 | Section Subsetoid. 358 | Variable (S:setoidType) (P:S -> Prop). 359 | 360 | Lemma sub_setoidP : Setoid.axiom (fun (x y:{e : S | P e}) => match x,y with exist x' _, exist y' _ => x' =-= y' end). 361 | split ; last split. 362 | - by case => x Px. 363 | - case => x Px. case => y Py. case => z Pz. by apply tset_trans. 364 | - case => x Px. case => y Py. apply tset_sym. 365 | Qed. 366 | 367 | Canonical Structure sub_setoidMixin := Setoid.Mixin sub_setoidP. 368 | Canonical Structure sub_setoidType := Eval hnf in SetoidType sub_setoidMixin. 369 | 370 | Lemma forget_respect : setoid_respect (fun (x:sub_setoidType) => match x with exist x' _ => x' end). 371 | by case => x Px ; case => y Py. 372 | Qed. 373 | 374 | Definition Sforget : sub_setoidType =-> S := Eval hnf in mk_fset forget_respect. 375 | 376 | Lemma InheretFun_respect (N : setoidType) (f:N =-> S) (c:(forall n, P (f n))) : setoid_respect (fun n => exist (fun x => P x) (f n) (c n)). 377 | move => n n' e. simpl. by apply (frespect f e). 378 | Qed. 379 | 380 | Definition sInheritFun (N : setoidType) (f:N =-> S) (c:forall n, P (f n)) : 381 | N =-> sub_setoidType := Eval hnf in mk_fset (InheretFun_respect c). 382 | 383 | Lemma Sforget_mono M (f:M =-> sub_setoidType) g : Sforget << f =-= Sforget << g -> f =-= g. 384 | move => C x. specialize (C x). case_eq (f x). move => fx Pf. case_eq (g x). move => gx Pg. 385 | move => e e'. have ee:(Sforget (f x) =-= Sforget (g x)) by []. rewrite e in ee. rewrite e' in ee. simpl in ee. 386 | apply ee. 387 | Qed. 388 | 389 | End Subsetoid. 390 | 391 | Section Option. 392 | 393 | Variable S:setoidType. 394 | 395 | Lemma option_setoid_axiom : 396 | Setoid.axiom (fun x y : option S => match x,y with None,None => True | Some x, Some y => x =-= y | _,_ => False end). 397 | split ; last split. 398 | - by case. 399 | - case ; first (move => x ; case ; [move => y ; case ; by [move => z ; apply tset_trans | by []] | by []]). 400 | case ; first by []. by case. 401 | - case ; [move => x ; case ; [move => y ; apply tset_sym | by []] | by case]. 402 | Qed. 403 | 404 | Canonical Structure option_setoidMixin := SetoidMixin option_setoid_axiom. 405 | Canonical Structure option_setoidType := Eval hnf in SetoidType option_setoidMixin. 406 | 407 | End Option. 408 | 409 | Arguments Scope fset_setoidType [S_scope S_scope]. 410 | 411 | Notation "'Option'" := option_setoidType : S_scope. 412 | 413 | Lemma some_respect (S:setoidType) : setoid_respect (@Some S). 414 | by []. 415 | Qed. 416 | 417 | Definition Ssome (S:setoidType) : S =-> Option S := Eval hnf in mk_fset (@some_respect S). 418 | 419 | Lemma sbind_respect (S S':setoidType) (f:S =-> Option S') : setoid_respect (fun x => match x with None => None | Some x => f x end). 420 | case ; last by case. 421 | move => x. case ; last by []. move => y e. apply (frespect f). by apply e. 422 | Qed. 423 | 424 | Definition Soptionbind (S S':setoidType) (f:S =-> Option S') : Option S =-> Option S' := 425 | Eval hnf in mk_fset (sbind_respect f). 426 | 427 | Lemma discrete_setoidAxiom (T:Type) : Setoid.axiom (fun x y : T => x = y). 428 | split ; first by []. 429 | split ; first by apply: trans_equal. 430 | by apply:sym_equal. 431 | Qed. 432 | 433 | Definition discrete_setoidMixin T := SetoidMixin (discrete_setoidAxiom T). 434 | Definition discrete_setoidType T := Eval hnf in SetoidType (discrete_setoidMixin T). 435 | 436 | Lemma setAxiom T : @Setoid.axiom (T -> Prop) (fun X Y => forall x, X x <-> Y x). 437 | split ; last split. 438 | - move => a. simpl. by []. 439 | - move => a. simpl. move => y z A B x. rewrite <- B. by rewrite <- A. 440 | - move => a. simpl. move => b A x. by rewrite A. 441 | Qed. 442 | 443 | Canonical Structure setMixin T := SetoidMixin (setAxiom T). 444 | Canonical Structure setType T := Eval hnf in SetoidType (setMixin T). 445 | 446 | Close Scope S_scope. 447 | Close Scope C_scope. 448 | 449 | -------------------------------------------------------------------------------- /src/PredomAll.v: -------------------------------------------------------------------------------- 1 | (********************************************************************************** 2 | * PredomAll.v * 3 | * Formalizing Domains, Ultrametric Spaces and Semantics of Programming Languages * 4 | * Nick Benton, Lars Birkedal, Andrew Kennedy and Carsten Varming * 5 | * Jan 2012 * 6 | * Build with Coq 8.3pl2 plus SSREFLECT * 7 | **********************************************************************************) 8 | 9 | Require Export PredomCore. 10 | Require Export PredomSum. 11 | Require Export PredomLift. 12 | Require Export PredomKleisli. 13 | Require Export PredomFix. 14 | Implicit Arguments eta [D]. 15 | -------------------------------------------------------------------------------- /src/PredomAux.v: -------------------------------------------------------------------------------- 1 | Require Import utility. 2 | Require Import PredomCore. 3 | Require Import PredomProd. 4 | Require Import PredomLift. 5 | Require Import PredomKleisli. 6 | 7 | Set Implicit Arguments. 8 | Unset Strict Implicit. 9 | 10 | Open Scope O_scope. 11 | 12 | (*========================================================================== 13 | Morphisms, missing from Predom.v 14 | ==========================================================================*) 15 | 16 | (*========================================================================== 17 | Some basic lemmas 18 | ==========================================================================*) 19 | 20 | Lemma eta_injective : 21 | forall D E (f g:D -c> E), eta << f == eta << g -> f == g. 22 | intros D E f g eq. 23 | refine (fcont_eq_intro _). 24 | intro x. refine (vinj _). 25 | assert (H := fcont_eq_elim eq x). auto. 26 | Qed. 27 | 28 | -------------------------------------------------------------------------------- /src/PredomFix.v: -------------------------------------------------------------------------------- 1 | (********************************************************************************** 2 | * PredomFix.v * 3 | * Formalizing Domains, Ultrametric Spaces and Semantics of Programming Languages * 4 | * Nick Benton, Lars Birkedal, Andrew Kennedy and Carsten Varming * 5 | * Jan 2012 * 6 | * Build with Coq 8.3pl2 plus SSREFLECT * 7 | **********************************************************************************) 8 | 9 | (*========================================================================== 10 | Definition of fixpoints and associated lemmas 11 | ==========================================================================*) 12 | Require Import PredomCore. 13 | Require Import PredomLift. 14 | 15 | Set Implicit Arguments. 16 | Unset Strict Implicit. 17 | Import Prenex Implicits. 18 | 19 | (** ** Fixpoints *) 20 | 21 | Section Fixpoints. 22 | Variable D : cppoType. 23 | 24 | Variable f : cpoCatType D D. 25 | 26 | Fixpoint iter_ n : D := match n with O => PBot | S m => f (iter_ m) end. 27 | 28 | Lemma iter_incr : forall n, iter_ n <= f (iter_ n). 29 | elim ; first by apply: leastP. 30 | move => n L. simpl. apply: fmonotonic. by apply L. 31 | Save. 32 | 33 | Hint Resolve iter_incr. 34 | 35 | Lemma iter_m : monotonic iter_. 36 | move => n n'. elim: n' n. 37 | - move => n. move => L. unfold Ole in L. simpl in L. 38 | rewrite -> (leqn0 n) in L. by rewrite (eqP L). 39 | - move => n IH n' L. unfold Ole in L. simpl in L. rewrite leq_eqVlt in L. 40 | case_eq (n' == n.+1) => E ; rewrite E in L. by rewrite (eqP E). 41 | specialize (IH _ L). by rewrite -> IH ; simpl. 42 | Qed. 43 | 44 | Definition iter : natO =-> D := mk_fmono (iter_m). 45 | 46 | Definition fixp : D := lub iter. 47 | 48 | Lemma fixp_le : fixp <= f fixp. 49 | unfold fixp. 50 | apply Ole_trans with (lub (ocomp f iter)). 51 | - apply: lub_le_compat. move => x. simpl. by apply iter_incr. 52 | - by rewrite -> lub_comp_le. 53 | Save. 54 | Hint Resolve fixp_le. 55 | 56 | Lemma fixp_eq : fixp =-= f fixp. 57 | apply: Ole_antisym; first by []. 58 | unfold fixp. rewrite {2} (lub_lift_left iter (S O)). 59 | rewrite -> (fcontinuous f iter). by apply: lub_le_compat => i. 60 | Save. 61 | 62 | Lemma fixp_inv : forall g, f g <= g -> fixp <= g. 63 | unfold fixp; intros g l. 64 | apply: lub_le. elim. simpl. by apply: leastP. 65 | move => n L. apply: (Ole_trans _ l). apply: (fmonotonic f). by apply L. 66 | Save. 67 | 68 | End Fixpoints. 69 | Hint Resolve fixp_le fixp_eq fixp_inv. 70 | 71 | Definition fixp_cte (D:cppoType) : forall (d:D), fixp (const D d) =-= d. 72 | intros; apply fixp_eq with (f:=const D d); red; intros; auto. 73 | Save. 74 | Hint Resolve fixp_cte. 75 | 76 | Add Parametric Morphism (D:cppoType) : (@fixp D) 77 | with signature (@Ole _ : ((D:cpoType) =-> D) -> (D =-> D) -> Prop) ++> (@Ole _) 78 | as fixp_le_compat. 79 | move => x y l. unfold fixp. 80 | apply: lub_le_compat. elim ; first by []. 81 | move => n IH. simpl. rewrite -> IH. by apply l. 82 | Qed. 83 | Hint Resolve fixp_le_compat. 84 | 85 | Add Parametric Morphism (D:cppoType) : (@fixp D) 86 | with signature (@tset_eq _) ==> (@tset_eq D) 87 | as fixp_eq_compat. 88 | by intros x y H; apply Ole_antisym ; apply: (fixp_le_compat _) ; case: H. 89 | Save. 90 | Hint Resolve fixp_eq_compat. 91 | 92 | Lemma fixp_mon (D:cppoType) : monotonic (@fixp D). 93 | move => x y e. by rewrite -> e. 94 | Qed. 95 | 96 | Definition Fixp (D:cppoType) : ordCatType (D -=> D) D := Eval hnf in mk_fmono (@fixp_mon D). 97 | 98 | Lemma Fixp_simpl (D:cppoType) : forall (f:D =-> D), Fixp D f = fixp f. 99 | trivial. 100 | Save. 101 | 102 | Lemma iter_mon (D:cppoType) : monotonic (@iter D). 103 | move => x y l. elim ; first by []. 104 | move => n IH. simpl. rewrite -> IH. by apply l. 105 | Qed. 106 | 107 | Definition Iter (D:cppoType) : ordCatType (D -=> D) (fmon_cpoType natO D) := 108 | Eval hnf in mk_fmono (@iter_mon D). 109 | 110 | Lemma IterS_simpl (D:cppoType) : forall f n, Iter D f (S n) = f (Iter _ f n). 111 | trivial. 112 | Save. 113 | 114 | Lemma iterS_simpl (D:cppoType) : forall (f:cpoCatType D D) n, iter f (S n) = f (iter f n). 115 | trivial. 116 | Save. 117 | 118 | Lemma iter_continuous (D:cppoType) : 119 | forall h : natO =-> D -=> D, 120 | iter (lub h) <= lub (Iter D << h). 121 | move => h. elim ; first by apply: leastP. 122 | move => n IH. simpl. rewrite fcont_app_eq. rewrite -> IH. rewrite <- fcont_app_eq. 123 | rewrite fcont_app_continuous. rewrite lub_diag. by apply lub_le_compat => i. 124 | Qed. 125 | 126 | Hint Resolve iter_continuous. 127 | 128 | Lemma iter_continuous_eq (D:cppoType) : 129 | forall h : natO =-> (D -=> D), 130 | iter (lub h) =-= lub (Iter _ << h). 131 | intros; apply: Ole_antisym; auto. 132 | exact (lub_comp_le (Iter _ (*D*)) h). 133 | Save. 134 | 135 | 136 | Lemma fixp_continuous (D:cppoType) : forall (h : natO =-> (D -=> D)), 137 | fixp (lub h) <= lub (Fixp D << h). 138 | move => h. unfold fixp. rewrite -> iter_continuous_eq. 139 | apply: lub_le => n. simpl. 140 | apply: lub_le => m. simpl. rewrite <- (le_lub _ m). simpl. unfold fixp. by rewrite <- (le_lub _ n). 141 | Save. 142 | Hint Resolve fixp_continuous. 143 | 144 | Lemma fixp_continuous_eq (D:cppoType) : forall (h : natO =-> (D -=> D)), 145 | fixp (lub h) =-= lub (Fixp D << h). 146 | intros; apply: Ole_antisym; auto. 147 | by apply (lub_comp_le (Fixp D) h). 148 | Save. 149 | 150 | Lemma Fixp_cont (D:cppoType) : continuous (@Fixp D). 151 | move => c. 152 | rewrite Fixp_simpl. by rewrite -> fixp_continuous_eq. 153 | Qed. 154 | 155 | Definition FIXP (D:cppoType) : (D -=> D) =-> D := Eval hnf in mk_fcont (@Fixp_cont D). 156 | Implicit Arguments FIXP [D]. 157 | 158 | Lemma FIXP_simpl (D:cppoType) : forall (f:D=->D), FIXP f = fixp f. 159 | trivial. 160 | Save. 161 | 162 | Lemma FIXP_le_compat (D:cppoType) : forall (f g : D =-> D), 163 | f <= g -> FIXP f <= FIXP g. 164 | move => f g l. by rewrite -> l. 165 | Save. 166 | Hint Resolve FIXP_le_compat. 167 | 168 | Lemma FIXP_eq (D:cppoType) : forall (f:D=->D), FIXP f =-= f (FIXP f). 169 | intros; rewrite FIXP_simpl. 170 | by apply: (fixp_eq). 171 | Save. 172 | Hint Resolve FIXP_eq. 173 | 174 | Lemma FIXP_com: forall E D (f:E =-> (D -=> D)) , FIXP << f =-= ev << <| f, FIXP << f |>. 175 | intros E D f. apply: fmon_eq_intro. intros e. simpl. by rewrite {1} fixp_eq. 176 | Qed. 177 | 178 | Lemma FIXP_inv (D:cppoType) : forall (f:D=->D)(g : D), f g <= g -> FIXP f <= g. 179 | intros; rewrite FIXP_simpl; apply: fixp_inv; auto. 180 | Save. 181 | 182 | (** *** Iteration of functional *) 183 | Lemma FIXP_comp_com (D:cppoType) : forall (f g:D=->D), 184 | g << f <= f << g-> FIXP g <= f (FIXP g). 185 | intros; apply FIXP_inv. 186 | apply Ole_trans with (f (g (FIXP g))). 187 | assert (X:=H (FIXP g)). simpl. by apply X. 188 | apply: fmonotonic. 189 | case (FIXP_eq g); trivial. 190 | Save. 191 | 192 | Lemma FIXP_comp (D:cppoType) : forall (f g:D=->D), 193 | g << f <= f << g -> f (FIXP g) <= FIXP g -> FIXP (f << g) =-= FIXP g. 194 | intros; apply: Ole_antisym. 195 | - apply FIXP_inv. simpl. 196 | apply Ole_trans with (f (FIXP g)) ; last by []. by rewrite <- (FIXP_eq g). 197 | - apply: FIXP_inv. 198 | assert (g (f (FIXP (f << g))) <= f (g (FIXP (f << g)))). 199 | specialize (H (FIXP (f << g))). apply H. 200 | case (FIXP_eq (f<D) (n:nat) {struct n} : D =->D := 209 | match n with O => f | S p => fcont_compn f p << f end. 210 | 211 | Add Parametric Morphism (D1 D2 D3 : cpoType) : (@ccomp D1 D2 D3) 212 | with signature (@Ole (D2 -=> D3) : (D2 =-> D3) -> (D2 =-> D3) -> Prop ) ++> (@Ole (D1 -=> D2)) ++> (@Ole (D1 -=> D3)) 213 | as fcont_comp_le_compat. 214 | move => f g l h k l' x. simpl. rewrite -> l. by rewrite -> l'. 215 | Save. 216 | 217 | Lemma fcont_compn_com (D:cppoType) : forall (f:D =->D) (n:nat), 218 | f << (fcont_compn f n) <= fcont_compn f n << f. 219 | induction n; first by []. 220 | simpl fcont_compn. rewrite -> comp_assoc. by apply: (fcont_comp_le_compat _ (Ole_refl f)). 221 | Save. 222 | 223 | Lemma FIXP_compn (D:cppoType) : 224 | forall (f:D =->D) (n:nat), FIXP (fcont_compn f n) =-= FIXP f. 225 | move => f. case ; first by []. simpl. 226 | move => n. apply: FIXP_comp ; first by apply fcont_compn_com. 227 | elim: n. simpl. by rewrite <- (FIXP_eq f). 228 | move => n IH. simpl. rewrite <- (FIXP_eq f). by apply IH. 229 | Save. 230 | 231 | Lemma fixp_double (D:cppoType) : forall (f:D=->D), FIXP (f << f) =-= FIXP f. 232 | intros; exact (FIXP_compn f (S O)). 233 | Save. 234 | 235 | 236 | (** *** Induction principle *) 237 | (*=Adm *) 238 | Definition admissible (D:cpoType) (P:D -> Prop) := 239 | forall f : natO =-> D, (forall n, P (f n)) -> P (lub f). 240 | Lemma fixp_ind (D:cppoType) : forall (F: D =-> D)(P:D -> Prop), 241 | admissible P -> P PBot -> (forall x, P x -> P (F x)) -> P (fixp F). (*CLEAR*) 242 | move => F P Adm B I. 243 | apply Adm. by elim ; simpl ; auto. 244 | Qed. 245 | (*CLEARED*) 246 | (*=End *) 247 | 248 | Definition admissibleT (D:cpoType) (P:D -> Type) := 249 | forall f : natO =-> D, (forall n, P (f n)) -> P (lub f). 250 | 251 | Section SubCPO. 252 | Variable D E:cpoType. 253 | Variable P : D -> Prop. 254 | Variable I:admissible P. 255 | 256 | Definition Subchainlub (c:natO =-> (sub_ordType P)) : sub_ordType P. 257 | exists (@lub D (Forgetm P << c)). 258 | unfold admissible in I. specialize (@I (Forgetm P << c)). 259 | apply I. intros i. simpl. case (c i). auto. 260 | Defined. 261 | 262 | Lemma subCpoAxiom : CPO.axiom Subchainlub. 263 | intros c e n. unfold Subchainlub. split. case_eq (c n). intros x Px cn. 264 | refine (Ole_trans _ (@le_lub _ (Forgetm P << c) n)). 265 | simpl. rewrite cn. auto. 266 | intros C. simpl. 267 | case_eq e. intros dd pd de. 268 | refine (lub_le _). intros i. simpl. specialize (C i). 269 | case_eq (c i). intros d1 pd1 cn. rewrite cn in C. 270 | simpl in C. rewrite de in C. auto. 271 | Qed. 272 | 273 | Canonical Structure sub_cpoMixin := CpoMixin subCpoAxiom. 274 | Canonical Structure sub_cpoType := Eval hnf in CpoType sub_cpoMixin. 275 | 276 | Lemma InheritFun_cont (f:E =-> D) (p:forall d, P (f d)) : continuous (InheritFunm _ p). 277 | move => c. simpl. unfold Ole. simpl. rewrite (fcontinuous f). by apply lub_le_compat => i. 278 | Qed. 279 | 280 | Definition InheritFun (f:E =-> D) (p:forall d, P (f d)) : E =-> sub_cpoType := 281 | Eval hnf in mk_fcont (InheritFun_cont p). 282 | 283 | Lemma InheritFun_simpl (f:E =-> D) (p:forall d, P (f d)) d : InheritFun p d = InheritFunm _ p d. 284 | by []. 285 | Qed. 286 | 287 | Lemma Forgetm_cont : continuous (Forgetm P). 288 | by move => c. 289 | Qed. 290 | 291 | Definition Forget : sub_cpoType =-> D := Eval hnf in mk_fcont Forgetm_cont. 292 | 293 | Lemma forgetlub (c:natO =-> (sub_ordType P)) : 294 | Forget (Subchainlub c) = (@lub D (Forgetm P << c)). 295 | auto. 296 | Qed. 297 | 298 | Lemma ForgetP d : P (Forget d). 299 | intros. case d. auto. 300 | Qed. 301 | 302 | End SubCPO. 303 | 304 | Lemma Forget_leinj: forall (D:cpoType) (P:D -> Prop) (I:admissible P) (d:sub_cpoType I) e, Forget I d <= Forget I e -> d <= e. 305 | intros D P I d e. case e. clear e. intros e Pe. case d. clear d. intros d Pd. 306 | auto. 307 | Qed. 308 | 309 | Hint Resolve Forget_leinj. 310 | 311 | Lemma Forget_inj : forall (D:cpoType) (P:D -> Prop) (I:admissible P) (d:sub_cpoType I) e, Forget I d =-= Forget I e -> d =-= e. 312 | intros. by split ; case: H ; case: d ; case: e. 313 | Qed. 314 | 315 | Hint Resolve Forget_inj. 316 | 317 | Lemma Forget_leinjp: forall (D E:cpoType) (P:D -> Prop) (I:admissible P) (d:E =-> sub_cpoType I) e, 318 | Forget I << d <= Forget I << e -> d <= e. 319 | intros D E P I d e C x. specialize (C x). simpl in C. unfold FCont.fmono. simpl. 320 | case: (e x) C => e' Pe. case: (d x) => d' Pd l. by apply l. 321 | Qed. 322 | 323 | Lemma Forget_injp: forall (D E:cpoType) (P:D -> Prop) (I:admissible P) (d:E =-> sub_cpoType I) e, 324 | Forget I << d =-= Forget I << e -> d =-= e. 325 | intros D E P I d e C. 326 | apply: fmon_eq_intro. intros x. assert (CC:=fmon_eq_elim C x). simpl in CC. 327 | unfold FCont.fmono. simpl. clear C. 328 | case: (e x) CC. clear e. intros e Pe. case (d x). clear d. intros d Pd. 329 | auto. 330 | Qed. 331 | 332 | Lemma InheritFun_eq_compat D E Q Qcc f g X XX : f =-= g -> 333 | (@InheritFun D E Q Qcc f X) =-= (@InheritFun D E Q Qcc g XX). 334 | intros. refine (fmon_eq_intro _). intros d. have A:=fmon_eq_elim H d. clear H. by split ; case: A. 335 | Qed. 336 | 337 | Lemma ForgetInherit (D E:cpoType) (P:D -> Prop) PS f B : Forget PS << @InheritFun D E P PS f B =-= f. 338 | by refine (fmon_eq_intro _). 339 | Qed. 340 | 341 | Lemma InheritFun_comp: forall D E F P I (f:E =-> F) X (g:D =-> E) XX, 342 | @InheritFun F _ P I f X << g =-= @InheritFun _ _ P I (f << g) XX. 343 | intros. refine (fmon_eq_intro _). intros d. simpl. 344 | split ; simpl ; by []. 345 | Qed. 346 | 347 | -------------------------------------------------------------------------------- /src/PredomProd.v: -------------------------------------------------------------------------------- 1 | (*========================================================================== 2 | Definition of products and associated lemmas 3 | ==========================================================================*) 4 | Require Import PredomCore. 5 | Set Implicit Arguments. 6 | Unset Strict Implicit. 7 | Import Prenex Implicits. 8 | 9 | 10 | (*========================================================================== 11 | Domain-theoretic definition 12 | ==========================================================================*) 13 | 14 | 15 | Lemma curry_uncurry (D0 D1 D2:cpoType) : @CURRY D0 D1 D2 << uncurry =o= Id. 16 | intros D0 D1 D2. 17 | refine (fcont_eq_intro _). intros f. refine (fcont_eq_intro _). intros x. refine (fcont_eq_intro _). intros y. by auto. 18 | 19 | (* 20 | assert (forall D0 D1 D2 D3 D4 (f:D0 -c> D1) (g:D1 -C-> D2) (h:D3 -c> D4), (g << f) *f* h == (g *f* ID) << (f *f* h)). 21 | intros. refine (Dprod_unique _ _) ; unfold PROD_map. rewrite FST_PROD_fun. 22 | rewrite <- fcont_comp_assoc. rewrite FST_PROD_fun. repeat (rewrite fcont_comp_assoc). rewrite FST_PROD_fun. auto. 23 | rewrite SND_PROD_fun. 24 | rewrite <- fcont_comp_assoc. rewrite SND_PROD_fun. repeat (rewrite fcont_comp_assoc). rewrite SND_PROD_fun. 25 | rewrite fcont_comp_unitL. auto. 26 | 27 | apply Dext. rewrite H. 28 | rewrite <- fcont_comp_assoc. 29 | unfold CURRY. rewrite <- curry_com. 30 | apply Dext. 31 | rewrite H. 32 | rewrite <- fcont_comp_assoc. 33 | rewrite <- curry_com. 34 | unfold uncurry. 35 | rewrite fcont_comp_assoc. 36 | assert ((<| FST << FST, <| SND << FST, SND |> |> << 37 | (curry (D1:=D0 -C-> D1 -C-> D2) (D2:=D0 *c* D1) (D3:=D2) 38 | (EV << <| EV << <| FST, FST << SND |>, SND << SND |>) *f* 39 | @ID D0) *f* @ID D1) == 40 | (curry (EV << <| EV << <| FST, FST << SND |>, SND << SND |>) *f* (ID *f* ID)) << 41 | <| FST << FST, <| SND << FST, SND |> |>). 42 | refine (Dprod_unique _ _) ; unfold PROD_map. 43 | rewrite <- fcont_comp_assoc. rewrite FST_PROD_fun. rewrite fcont_comp_assoc. 44 | rewrite FST_PROD_fun. rewrite <- fcont_comp_assoc. rewrite FST_PROD_fun. 45 | rewrite <- fcont_comp_assoc. rewrite FST_PROD_fun. repeat (rewrite fcont_comp_assoc). rewrite FST_PROD_fun. auto. 46 | rewrite <- fcont_comp_assoc. rewrite SND_PROD_fun. 47 | rewrite <- fcont_comp_assoc. rewrite SND_PROD_fun. 48 | refine (Dprod_unique _ _). rewrite <- fcont_comp_assoc. rewrite FST_PROD_fun. 49 | rewrite fcont_comp_assoc. rewrite FST_PROD_fun. rewrite <- fcont_comp_assoc. rewrite SND_PROD_fun. 50 | repeat (rewrite <- fcont_comp_assoc). rewrite FST_PROD_fun. 51 | repeat (rewrite fcont_comp_unitL). 52 | repeat (rewrite fcont_comp_assoc). rewrite SND_PROD_fun. rewrite FST_PROD_fun. auto. 53 | rewrite <- fcont_comp_assoc. repeat (rewrite SND_PROD_fun). 54 | repeat (rewrite <- fcont_comp_assoc). rewrite SND_PROD_fun. rewrite fcont_comp_assoc. 55 | rewrite SND_PROD_fun. rewrite fcont_comp_assoc. 56 | rewrite SND_PROD_fun. auto. 57 | rewrite H0. clear H0. rewrite <- fcont_comp_assoc. 58 | assert (@ID D0 *f* @ID D1 == ID) as Ide. refine (Dprod_unique _ _) ; unfold PROD_map. 59 | rewrite FST_PROD_fun. rewrite fcont_comp_unitR. rewrite fcont_comp_unitL. auto. 60 | rewrite SND_PROD_fun. rewrite fcont_comp_unitR. rewrite fcont_comp_unitL. auto. 61 | rewrite (PROD_map_eq_compat (Oeq_refl _) Ide). rewrite <- curry_com. 62 | repeat (rewrite fcont_comp_assoc). 63 | refine (fcont_comp_eq_compat (Oeq_refl _) _). 64 | refine (Dprod_unique _ _) ; unfold PROD_map. 65 | rewrite <- fcont_comp_assoc. repeat (rewrite FST_PROD_fun). 66 | repeat (rewrite fcont_comp_assoc). 67 | refine (fcont_comp_eq_compat (Oeq_refl _) _). 68 | refine (Dprod_unique _ _). repeat (rewrite <- fcont_comp_assoc). repeat (rewrite FST_PROD_fun). 69 | rewrite fcont_comp_unitL. auto. repeat (rewrite <- fcont_comp_assoc). repeat (rewrite SND_PROD_fun). 70 | rewrite fcont_comp_unitL. rewrite fcont_comp_assoc. rewrite SND_PROD_fun. rewrite FST_PROD_fun. auto. 71 | repeat (rewrite <- fcont_comp_assoc). repeat (rewrite SND_PROD_fun). 72 | rewrite fcont_comp_assoc. repeat (rewrite SND_PROD_fun). rewrite fcont_comp_unitL. auto.*) 73 | Qed. 74 | 75 | (* 76 | Lemma uncurry_curry (D0 D1 D2:cpo) : uncurry << @CURRY D0 D1 D2 == ID. 77 | intros. 78 | apply Dext. unfold uncurry. 79 | assert (forall D0 D1 D2 D3 D4 (f:D0 -c> D1) (g:D1 -C-> D2) (h:D3 -c> D4), (g << f) *f* h == (g *f* ID) << (f *f* h)). 80 | intros. refine (Dprod_unique _ _) ; unfold PROD_map. rewrite FST_PROD_fun. 81 | rewrite <- fcont_comp_assoc. rewrite FST_PROD_fun. repeat (rewrite fcont_comp_assoc). rewrite FST_PROD_fun. auto. 82 | rewrite SND_PROD_fun. 83 | rewrite <- fcont_comp_assoc. rewrite SND_PROD_fun. repeat (rewrite fcont_comp_assoc). rewrite SND_PROD_fun. 84 | rewrite fcont_comp_unitL. auto. 85 | 86 | rewrite H. rewrite <- fcont_comp_assoc. rewrite <- curry_com. unfold CURRY. 87 | 88 | *) 89 | 90 | 91 | (** ** Indexed product of cpo's *) 92 | 93 | (* Stuff 94 | (** *** Particular cases with one or two elements *) 95 | 96 | Section Product2. 97 | 98 | Definition I2 := bool. 99 | Variable DI2 : bool -> cpo. 100 | 101 | Definition DP1 := DI2 true. 102 | Definition DP2 := DI2 false. 103 | 104 | Definition PI1 : Dprodi DI2 -c> DP1 := PROJ DI2 true. 105 | Definition pi1 (d:Dprodi DI2) := PI1 d. 106 | 107 | Definition PI2 : Dprodi DI2 -c> DP2 := PROJ DI2 false. 108 | Definition pi2 (d:Dprodi DI2) := PI2 d. 109 | 110 | Definition pair2 (d1:DP1) (d2:DP2) : Dprodi DI2 := bool_rect DI2 d1 d2. 111 | 112 | Lemma pair2_le_compat : forall (d1 d'1:DP1) (d2 d'2:DP2), d1 <= d'1 -> d2 <= d'2 113 | -> pair2 d1 d2 <= pair2 d'1 d'2. 114 | intros; intro b; case b; simpl; auto. 115 | Save. 116 | 117 | Definition Pair2 : DP1 -m> DP2 -m> Dprodi DI2 := le_compat2_mon pair2_le_compat. 118 | 119 | Definition PAIR2 : DP1 -c> DP2 -C-> Dprodi DI2. 120 | apply continuous2_cont with (f:=Pair2). 121 | red; intros; intro b. 122 | case b; simpl; apply lub_le_compat; auto. 123 | Defined. 124 | 125 | Lemma PAIR2_simpl : forall (d1:DP1) (d2:DP2), PAIR2 d1 d2 = Pair2 d1 d2. 126 | trivial. 127 | Save. 128 | 129 | Lemma Pair2_simpl : forall (d1:DP1) (d2:DP2), Pair2 d1 d2 = pair2 d1 d2. 130 | trivial. 131 | Save. 132 | 133 | Lemma pi1_simpl : forall (d1: DP1) (d2:DP2), pi1 (pair2 d1 d2) = d1. 134 | trivial. 135 | Save. 136 | 137 | Lemma pi2_simpl : forall (d1: DP1) (d2:DP2), pi2 (pair2 d1 d2) = d2. 138 | trivial. 139 | Save. 140 | 141 | Definition DI2_map (f1 : DP1 -c> DP1) (f2:DP2 -c> DP2) 142 | : Dprodi DI2 -c> Dprodi DI2 := 143 | DMAPi (bool_rect (fun b:bool => DI2 b -c>DI2 b) f1 f2). 144 | 145 | Lemma Dl2_map_eq : forall (f1 : DP1 -c> DP1) (f2:DP2 -c> DP2) (d:Dprodi DI2), 146 | DI2_map f1 f2 d == pair2 (f1 (pi1 d)) (f2 (pi2 d)). 147 | intros; simpl; apply Oprodi_eq_intro; intro b; case b; trivial. 148 | Save. 149 | End Product2. 150 | Hint Resolve Dl2_map_eq. 151 | 152 | Section Product1. 153 | Definition I1 := unit. 154 | Variable D : cpo. 155 | 156 | Definition DI1 (_:unit) := D. 157 | Definition PI : Dprodi DI1 -c> D := PROJ DI1 tt. 158 | Definition pi (d:Dprodi DI1) := PI d. 159 | 160 | Definition pair1 (d:D) : Dprodi DI1 := unit_rect DI1 d. 161 | 162 | Definition pair1_simpl : forall (d:D) (x:unit), pair1 d x = d. 163 | destruct x; trivial. 164 | Defined. 165 | 166 | Definition Pair1 : D -m> Dprodi DI1. 167 | exists pair1; red; intros; intro d. 168 | repeat (rewrite pair1_simpl);trivial. 169 | Defined. 170 | 171 | 172 | Lemma Pair1_simpl : forall (d:D), Pair1 d = pair1 d. 173 | trivial. 174 | Save. 175 | 176 | Definition PAIR1 : D -c> Dprodi DI1. 177 | exists Pair1; red; intros; repeat (rewrite Pair1_simpl). 178 | intro d; rewrite pair1_simpl. 179 | rewrite (Dprodi_lub_simpl (Di:=DI1)). 180 | apply lub_le_compat; intros. 181 | intro x; simpl; rewrite pair1_simpl; auto. 182 | Defined. 183 | 184 | Lemma pi_simpl : forall (d:D), pi (pair1 d) = d. 185 | trivial. 186 | Save. 187 | 188 | Definition DI1_map (f : D -c> D) 189 | : Dprodi DI1 -c> Dprodi DI1 := 190 | DMAPi (fun t:unit => f). 191 | 192 | Lemma DI1_map_eq : forall (f : D -c> D) (d:Dprodi DI1), 193 | DI1_map f d == pair1 (f (pi d)). 194 | intros; simpl; apply Oprodi_eq_intro; intro b; case b; trivial. 195 | Save. 196 | End Product1. 197 | 198 | Hint Resolve DI1_map_eq. 199 | End of Stuff *) 200 | -------------------------------------------------------------------------------- /src/PredomSum.v: -------------------------------------------------------------------------------- 1 | (********************************************************************************** 2 | * PredomSum.v * 3 | * Formalizing Domains, Ultrametric Spaces and Semantics of Programming Languages * 4 | * Nick Benton, Lars Birkedal, Andrew Kennedy and Carsten Varming * 5 | * Jan 2012 * 6 | * Build with Coq 8.3pl2 plus SSREFLECT * 7 | **********************************************************************************) 8 | 9 | 10 | (*========================================================================== 11 | Definition of co-products and associated lemmas 12 | ==========================================================================*) 13 | Require Import PredomCore. 14 | Set Implicit Arguments. 15 | Unset Strict Implicit. 16 | Import Prenex Implicits. 17 | 18 | 19 | (*========================================================================== 20 | Order-theoretic definitions 21 | ==========================================================================*) 22 | 23 | Section OrdSum. 24 | 25 | Variable O1 O2 : ordType. 26 | 27 | Lemma sum_ordAxiom : PreOrd.axiom (fun (x y:O1+O2) => match x, y with | inl x, inl y => x <= y | inr x, inr y => x <= y | _,_ => False end). 28 | case => x0. split ; first by []. by case => x1 ; case => x2 ; try apply Ole_trans. 29 | split ; first by []. by case => x1 ; case => x2 ; try apply Ole_trans. 30 | Qed. 31 | 32 | Canonical Structure sum_ordMixin := OrdMixin sum_ordAxiom. 33 | Canonical Structure sum_ordType := Eval hnf in OrdType sum_ordMixin. 34 | 35 | Lemma inl_mono : monotonic (@inl O1 O2). 36 | by move => x y l. 37 | Qed. 38 | 39 | Definition Inl : O1 =-> sum_ordType := Eval hnf in mk_fmono inl_mono. 40 | 41 | Lemma inr_mono : monotonic (@inr O1 O2). 42 | by move => x y l. 43 | Qed. 44 | 45 | Definition Inr : O2 =-> sum_ordType := Eval hnf in mk_fmono inr_mono. 46 | 47 | Lemma sum_fun_mono (O3:ordType) (f:O1 =-> O3) (g:O2 =-> O3) : monotonic (fun x => match x with inl x => f x | inr y => g y end). 48 | by case => x ; case => y ; try done ; move => X ; apply fmonotonic. 49 | Qed. 50 | 51 | Definition Osum_fun (O3:ordType) (f:O1 =-> O3) (g:O2 =-> O3) := Eval hnf in mk_fmono (sum_fun_mono f g). 52 | 53 | End OrdSum. 54 | 55 | Lemma sumOrdAxiom : @CatSum.axiom _ (sum_ordType) (@Inl) (@Inr) (@Osum_fun). 56 | move => O0 O1 O2 f g. split ; first by split ; apply: fmon_eq_intro. 57 | move => m [A B]. apply: fmon_eq_intro. case => x. 58 | - by apply (fmon_eq_elim A x). 59 | - by apply (fmon_eq_elim B x). 60 | Qed. 61 | 62 | Canonical Structure sum_ordCatMixin := sumCatMixin sumOrdAxiom. 63 | Canonical Structure sum_ordCatType := Eval hnf in sumCatType sum_ordCatMixin. 64 | 65 | Add Parametric Morphism (D E F:ordType) : (@Osum_fun D E F) 66 | with signature (@Ole _ : (D =-> F) -> (D =-> F) -> Prop) ++> (@Ole _ : (E =-> F) -> (E =-> F) -> Prop) ++> 67 | (@Ole (sum_ordType _ _ -=> _) : ((D + E) =-> F) -> (_ =-> _) -> Prop) 68 | as Osum_fun_le_compat. 69 | intros f f' lf g g' lg. 70 | simpl. intros x ; case x ; clear x ; auto. 71 | Qed. 72 | 73 | Add Parametric Morphism (D E F:ordType) : (@Osum_fun D E F) 74 | with signature (@tset_eq _ : (D =-> F) -> (D =-> F) -> Prop ) ==> 75 | (@tset_eq _ : (E =-> F) -> (E =-> F) -> Prop) ==> (@tset_eq ((sum_ordType D E) -=> F)) 76 | as Osum_fun_eq_compat. 77 | intros f f' lf g g' lg. 78 | simpl. apply fmon_eq_intro. 79 | destruct lg. destruct lf. 80 | intros x ; split ; case x ; clear x ; auto. 81 | Qed. 82 | 83 | Section CPOSum. 84 | 85 | Variable D1 D2 : cpoType. 86 | 87 | Lemma sum_left_mono (c:natO =-> sum_ordType D1 D2) (ex:{d| c O = inl D2 d}) : 88 | monotonic (fun x => match c x with | inl y => y | inr _ => projT1 ex end). 89 | case:ex => x X. simpl. 90 | move => z z' l. simpl. 91 | have H:= fmonotonic c (leq0n z). have HH:=fmonotonic c (leq0n z'). 92 | rewrite -> X in H,HH. 93 | have A:=fmonotonic c l. case: (c z) H A ; last by []. 94 | move => x0 _. by case: (c z') HH. 95 | Qed. 96 | 97 | Definition sum_left (c:natO =-> sum_ordType D1 D2) (ex:{d| c O = inl D2 d}) : natO =-> D1 := 98 | Eval hnf in mk_fmono (sum_left_mono ex). 99 | 100 | Lemma sum_right_mono (c:natO =-> sum_ordType D1 D2) (ex:{d| c O = inr D1 d}) : 101 | monotonic (fun x => match c x with | inr y => y | inl _ => projT1 ex end). 102 | case:ex => x X. simpl. 103 | move => z z' l. simpl. 104 | have H:= fmonotonic c (leq0n z). have HH:=fmonotonic c (leq0n z'). 105 | rewrite -> X in H,HH. 106 | have A:=fmonotonic c l. case: (c z) H A ; first by []. 107 | move => x0 _. by case: (c z') HH. 108 | Qed. 109 | 110 | Definition sum_right (c:natO =-> sum_ordType D1 D2) (ex:{d| c O = inr D1 d}) : natO =-> D2 := 111 | Eval hnf in mk_fmono (sum_right_mono ex). 112 | 113 | Definition SuminjProof (c:natO =-> sum_ordType D1 D2) (n:nat) : 114 | {d | c n = inl D2 d} + {d | c n = inr D1 d}. 115 | case: (c n) => cn. 116 | - left. by exists cn. 117 | - right. by exists cn. 118 | Defined. 119 | 120 | Definition sum_lub (c: natO =-> sum_ordType D1 D2) : sum_ordType D1 D2 := 121 | match (SuminjProof c 0) with 122 | | inl X => (inl D2 (lub (sum_left X))) 123 | | inr X => (inr D1 (lub (sum_right X))) 124 | end. 125 | 126 | (*========================================================================== 127 | Domain-theoretic definitions 128 | ==========================================================================*) 129 | 130 | Lemma Dsum_ub (c : natO =-> sum_ordType D1 D2) (n : nat) : c n <= sum_lub c. 131 | have L:=fmonotonic c (leq0n n). 132 | unfold sum_lub. case: (SuminjProof c 0) ; case => c0 e ; rewrite e in L. 133 | - move: L. case_eq (c n) ; last by []. move => cn e' L. apply: (Ole_trans _ (le_lub _ n)). 134 | simpl. by rewrite e'. 135 | - move: L. case_eq (c n) ; first by []. move => cn e' L. apply: (Ole_trans _ (le_lub _ n)). 136 | simpl. by rewrite e'. 137 | Qed. 138 | 139 | Lemma Dsum_lub (c : natO =-> sum_ordType D1 D2) (x : sum_ordType D1 D2) : 140 | (forall n : natO, c n <= x) -> sum_lub c <= x. 141 | case x ; clear x ; intros x cn. unfold sum_lub. 142 | case (SuminjProof c 0). 143 | - intros s. assert (lub (sum_left s) <= x). 144 | apply (lub_le). intros n. specialize (cn n). destruct s as [d c0]. 145 | simpl. case_eq (c n). 146 | + intros dn dnl. rewrite dnl in cn. by auto. 147 | + intros dn cnr. rewrite cnr in cn. by inversion cn. 148 | by auto. 149 | - intros s. destruct s. simpl. specialize (cn 0). rewrite e in cn. by inversion cn. 150 | - unfold sum_lub. case (SuminjProof c 0) ; first case. 151 | + move => y e. simpl. specialize (cn 0). rewrite e in cn. by inversion cn. 152 | + intros s. assert (lub (sum_right s) <= x). 153 | * apply lub_le. intros n. specialize (cn n). destruct s as [d c0]. 154 | simpl. case_eq (c n). 155 | - intros dl cnl. rewrite cnl in cn. by inversion cn. 156 | - intros dl dll. rewrite dll in cn. by auto. 157 | by auto. 158 | Qed. 159 | 160 | Lemma sum_cpoAxiom : CPO.axiom sum_lub. 161 | move => c s n. 162 | split ; first by apply Dsum_ub. 163 | by apply Dsum_lub. 164 | Qed. 165 | 166 | Canonical Structure sum_cpoMixin := CpoMixin sum_cpoAxiom. 167 | Canonical Structure sum_cpoType := Eval hnf in CpoType sum_cpoMixin. 168 | 169 | (* 170 | (** Printing +c+ %\ensuremath{+}% *) 171 | Infix "+c+" := Dsum (at level 28, left associativity).*) 172 | 173 | Variable D3:cpoType. 174 | 175 | Lemma SUM_fun_cont (f:D1 =-> D3) (g:D2 =-> D3) : @continuous sum_cpoType _ (Osum_fun f g). 176 | intros c. simpl. 177 | case_eq (lub c). 178 | - move => ll sl. unfold lub in sl. simpl in sl. unfold sum_lub in sl. destruct (SuminjProof c 0) ; last by []. 179 | case: sl => e. rewrite <- e. clear ll e. rewrite -> (fcontinuous f (sum_left s)). 180 | apply lub_le_compat => n. simpl. case: s => d s. simpl. move: (fmonotonic c (leq0n n)). 181 | case: (c 0) s ; last by []. move => s e. rewrite -> e ; clear s e. by case: (c n). 182 | - move => ll sl. unfold lub in sl. simpl in sl. unfold sum_lub in sl. destruct (SuminjProof c 0) ; first by []. 183 | case: sl => e. rewrite <- e. clear ll e. rewrite -> (fcontinuous g (sum_right s)). 184 | apply lub_le_compat => n. simpl. case: s => d s. simpl. move: (fmonotonic c (leq0n n)). 185 | case: (c 0) s ; first by []. move => s e. rewrite -> e ; clear s e. by case: (c n). 186 | Qed. 187 | 188 | Definition SUM_funX (f:D1 =-> D3) (g:D2 =-> D3) : sum_cpoType =-> D3 := Eval hnf in mk_fcont (SUM_fun_cont f g). 189 | Definition SUM_fun f g := locked (SUM_funX f g). 190 | 191 | Lemma SUM_fun_simpl f g x : SUM_fun f g x =-= match x with inl x => f x | inr x => g x end. 192 | unlock SUM_fun. simpl. by case: x. 193 | Qed. 194 | 195 | (** printing [| %\ensuremath{[} *) 196 | (** printing |] %\ensuremath{]} *) 197 | (* Notation "'[|' f , g '|]'" := (SUM_fun f g) (at level 30).*) 198 | 199 | Lemma Inl_cont : continuous (Inl D1 D2). 200 | move => c. simpl. by apply: lub_le_compat => n. 201 | Qed. 202 | 203 | Definition INL : D1 =-> sum_cpoType := Eval hnf in mk_fcont Inl_cont. 204 | 205 | Lemma Inr_cont : continuous (Inr D1 D2). 206 | move => c. simpl. by apply: lub_le_compat => n. 207 | Qed. 208 | 209 | Definition INR : D2 =-> sum_cpoType := Eval hnf in mk_fcont Inr_cont. 210 | 211 | End CPOSum. 212 | 213 | Lemma sumCpoAxiom : @CatSum.axiom _ sum_cpoType (@INL) (@INR) (@SUM_fun). 214 | move => D0 D1 D2 f g. split. unlock SUM_fun. apply: (proj1 (@sumOrdAxiom D0 D1 D2 f g)). 215 | move => m. unlock SUM_fun. apply: (proj2 (@sumOrdAxiom D0 D1 D2 f g)). 216 | Qed. 217 | 218 | Canonical Structure cpoSumCatMixin := sumCatMixin sumCpoAxiom. 219 | Canonical Structure cpoSumCatType := Eval hnf in sumCatType cpoSumCatMixin. 220 | 221 | Lemma SUM_fun_simplx (X Y Z : cpoType) (f:X =-> Z) (g : Y =-> Z) x : [|f,g|] x =-= match x with inl x => f x | inr x => g x end. 222 | case:x => x ; unfold sum_fun ; simpl ; by rewrite SUM_fun_simpl. 223 | Qed. 224 | 225 | Definition getmorph (D E : cpoType) (f:D -=> E) : D =-> E := ev << <| const _ f, Id |>. 226 | 227 | Lemma SUM_Fun_mon (D E F:cpoType) : monotonic (fun (P:(D -=> F) * (E -=> F)) => [|getmorph (fst P), getmorph (snd P)|]). 228 | case => f g. case => f' g'. case ; simpl => l l' x. 229 | unfold sum_fun. simpl. do 2 rewrite SUM_fun_simpl. 230 | by case x ; [apply l | apply l']. 231 | Qed. 232 | 233 | Definition SUM_Funm (D E F:cpoType) : fmono_ordType ((D -=> F) * (E -=> F)) (D + E -=> F) := 234 | Eval hnf in mk_fmono (@SUM_Fun_mon D E F). 235 | 236 | Lemma SUM_Fon_cont (D E F:cpoType) : continuous (@SUM_Funm D E F). 237 | move => c. simpl. unfold getmorph. unfold sum_fun. simpl. 238 | by case => x ; simpl ; rewrite SUM_fun_simpl ; apply: lub_le_compat => n ; simpl ; rewrite SUM_fun_simpl. 239 | Qed. 240 | 241 | Definition SUM_Fun (D E F:cpoType) : ((D -=> F) * (E -=> F)) =-> (D + E -=> F) := Eval hnf in mk_fcont (@SUM_Fon_cont D E F). 242 | Implicit Arguments SUM_Fun [D E F]. 243 | 244 | Lemma SUM_Fun_simpl (D E F:cpoType) (f:D =-> F) (g:E =-> F) : SUM_Fun (f,g) =-= SUM_fun f g. 245 | apply: fmon_eq_intro => x. simpl. unfold sum_fun. simpl. by do 2 rewrite SUM_fun_simpl. 246 | Qed. 247 | 248 | 249 | -------------------------------------------------------------------------------- /src/Vars.v: -------------------------------------------------------------------------------- 1 | (*========================================================================== 2 | Variables 3 | ==========================================================================*) 4 | 5 | Require Import List. 6 | 7 | Module Type VAR. 8 | Parameter Ty : Type. 9 | End VAR. 10 | 11 | Module Var (V : VAR). 12 | 13 | Require Import Program.Equality. 14 | Set Implicit Arguments. 15 | Unset Strict Implicit. 16 | Set Printing Implicit Defensive. 17 | 18 | Definition Ty := V.Ty. 19 | Definition Env := list Ty. 20 | 21 | Inductive Var : Env -> Ty -> Type := 22 | | ZVAR : forall env ty, Var (ty :: env) ty 23 | | SVAR : forall env ty ty', Var env ty -> Var (ty' :: env) ty. 24 | 25 | (*========================================================================== 26 | Variable-domain maps. 27 | By instantiating P with Var we get renamings. 28 | By instantiating P with terms we get substitutions. 29 | ==========================================================================*) 30 | 31 | Definition Map (P : Env -> Ty -> Type) env env' := forall t, Var env t -> P env' t. 32 | 33 | (* Head, tail and cons *) 34 | Definition tlMap P env env' ty (m : Map P (ty :: env) env') : Map P env env' := fun ty' v => m ty' (SVAR ty v). 35 | Definition hdMap P env env' ty (m : Map P (ty :: env) env') : P env' ty := m ty (ZVAR _ _). 36 | Implicit Arguments tlMap [P env env' ty]. 37 | Implicit Arguments hdMap [P env env' ty]. 38 | 39 | Program Definition consMap P env env' ty (v : P env' ty) (m:Map P env env') : Map P (ty :: env) env' := 40 | fun ty' (var:Var (ty :: env) ty') => 41 | match var with 42 | | ZVAR _ _ => v 43 | | SVAR _ _ _ var => m _ var 44 | end. 45 | Implicit Arguments consMap [P env env' ty]. 46 | 47 | Axiom MapExtensional : forall P env env' (r1 r2 : Map P env env'), (forall ty var, r1 ty var = r2 ty var) -> r1 = r2. 48 | 49 | Lemma hdConsMap : forall P env env' ty (v : P env' ty) (s : Map P env env'), hdMap (consMap v s) = v. Proof. auto. Qed. 50 | Lemma tlConsMap : forall P env env' ty (v : P env' ty) (s : Map P env env'), tlMap (consMap v s) = s. Proof. intros. apply MapExtensional. auto. Qed. 51 | 52 | Lemma consMapInv : forall P env env' ty (m:Map P (ty :: env) env'), exists m', exists v, m = consMap v m'. 53 | intros. exists (tlMap m). exists (hdMap m). 54 | apply MapExtensional. dependent destruction var; auto. 55 | Qed. 56 | 57 | Definition Renaming := Map Var. 58 | 59 | Definition idRenaming env : Renaming env env := fun ty (var : Var env ty) => var. 60 | Implicit Arguments idRenaming []. 61 | 62 | Definition composeRenaming P env env' env'' (m : Map P env' env'') (r : Renaming env env') : Map P env env'' := fun t var => m _ (r _ var). 63 | 64 | 65 | Lemma composeRenamingAssoc : 66 | forall P env env' env'' env''' (m : Map P env'' env''') r' (r : Renaming env env'), 67 | composeRenaming m (composeRenaming r' r) = composeRenaming (composeRenaming m r') r. 68 | Proof. auto. Qed. 69 | 70 | Lemma composeRenamingIdLeft : forall env env' (r : Renaming env env'), composeRenaming (idRenaming _) r = r. 71 | Proof. intros. apply MapExtensional. auto. Qed. 72 | 73 | Lemma composeRenamingIdRight : forall P env env' (m:Map P env env'), composeRenaming m (idRenaming _) = m. 74 | Proof. intros. apply MapExtensional. auto. Qed. 75 | 76 | End Var. -------------------------------------------------------------------------------- /src/VarsRec.v: -------------------------------------------------------------------------------- 1 | (*========================================================================== 2 | Variables 3 | ==========================================================================*) 4 | 5 | Require Import List. 6 | Require Import Program.Equality. 7 | Set Implicit Arguments. 8 | Unset Strict Implicit. 9 | Set Printing Implicit Defensive. 10 | 11 | Section Vars. 12 | 13 | Variable Ty : Type. 14 | 15 | Definition Env := list Ty. 16 | 17 | Inductive Var : Env -> Ty -> Type := 18 | | ZVAR : forall env ty, Var (ty :: env) ty 19 | | SVAR : forall env ty ty', Var env ty -> Var (ty' :: env) ty. 20 | 21 | (*========================================================================== 22 | Variable-domain maps. 23 | By instantiating P with Var we get renamings. 24 | By instantiating P with terms we get substitutions. 25 | ==========================================================================*) 26 | 27 | Definition Map (P : Env -> Ty -> Type) env env' := forall t, Var env t -> P env' t. 28 | 29 | (* Head, tail and cons *) 30 | Definition tlMap P env env' ty (m : Map P (ty :: env) env') : Map P env env' := fun ty' v => m ty' (SVAR ty v). 31 | Definition hdMap P env env' ty (m : Map P (ty :: env) env') : P env' ty := m ty (ZVAR _ _). 32 | Implicit Arguments tlMap [P env env' ty]. 33 | Implicit Arguments hdMap [P env env' ty]. 34 | 35 | Program Definition consMap P env env' ty (v : P env' ty) (m:Map P env env') : Map P (ty :: env) env' := 36 | fun ty' (var:Var (ty :: env) ty') => 37 | match var with 38 | | ZVAR _ _ => v 39 | | SVAR _ _ _ var => m _ var 40 | end. 41 | Implicit Arguments consMap [P env env' ty]. 42 | 43 | Axiom MapExtensional : forall P env env' (r1 r2 : Map P env env'), (forall ty var, r1 ty var = r2 ty var) -> r1 = r2. 44 | 45 | Lemma hdConsMap : forall P env env' ty (v : P env' ty) (s : Map P env env'), hdMap (consMap v s) = v. Proof. auto. Qed. 46 | Lemma tlConsMap : forall P env env' ty (v : P env' ty) (s : Map P env env'), tlMap (consMap v s) = s. Proof. intros. apply MapExtensional. auto. Qed. 47 | 48 | Lemma consMapInv : forall P env env' ty (m:Map P (ty :: env) env'), exists m', exists v, m = consMap v m'. 49 | intros. exists (tlMap m). exists (hdMap m). 50 | apply MapExtensional. dependent destruction var; auto. 51 | Qed. 52 | 53 | Definition Renaming := Map Var. 54 | 55 | Definition idRenaming env : Renaming env env := fun ty (var : Var env ty) => var. 56 | Implicit Arguments idRenaming []. 57 | 58 | Definition composeRenaming P env env' env'' (m : Map P env' env'') (r : Renaming env env') : Map P env env'' := fun t var => m _ (r _ var). 59 | 60 | 61 | Lemma composeRenamingAssoc : 62 | forall P env env' env'' env''' (m : Map P env'' env''') r' (r : Renaming env env'), 63 | composeRenaming m (composeRenaming r' r) = composeRenaming (composeRenaming m r') r. 64 | Proof. auto. Qed. 65 | 66 | Lemma composeRenamingIdLeft : forall env env' (r : Renaming env env'), composeRenaming (idRenaming _) r = r. 67 | Proof. intros. apply MapExtensional. auto. Qed. 68 | 69 | Lemma composeRenamingIdRight : forall P env env' (m:Map P env env'), composeRenaming m (idRenaming _) = m. 70 | Proof. intros. apply MapExtensional. auto. Qed. 71 | 72 | End Vars. 73 | 74 | 75 | -------------------------------------------------------------------------------- /src/buildcoqdoc.bat: -------------------------------------------------------------------------------- 1 | coqdoc --coqlib_path "%COQLIB%" --latex -g -p "\usepackage{stmaryrd} \input{../conf/macros}" Predom.v PredomAux.v lc.v typedlambda.v typedopsem.v typeddensem.v typedsubst.v typedsoundness.v typedadequacy.v Sets.v KnasterTarski.v uni.v unisem.v unisound.v uniade.v 2 | -------------------------------------------------------------------------------- /src/builddistrib.bat: -------------------------------------------------------------------------------- 1 | tar cf tphols09.tar Makefile utility.v PredomCore.v PredomProd.v PredomSum.v PredomLift.v PredomKleisli.v PredomFix.v PredomAll.v lc.v KnasterTarski.v Sets.v uni.v unisem.v uniade.v unisound.v typeddensem.v typedsubst.v typedlambda.v typedopsem.v typedsoundness.v typedadequacy.v -------------------------------------------------------------------------------- /src/coqdoc.sty: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/nbenton/coqdomains/1ae7ec4af95e4fa44d35d7a5b2452ad123b3a75d/src/coqdoc.sty -------------------------------------------------------------------------------- /src/linecount.bat: -------------------------------------------------------------------------------- 1 | wc utility.v Predom.v PredomAux.v Sets.v KnasterTarski.v lc.v 2 | wc typedlambda.v typedopsem.v typeddensem.v typedsubst.v typedsoundness.v typedadequacy.v 3 | wc uni.v unisem.v unisound.v uniade.v 4 | 5 | -------------------------------------------------------------------------------- /src/numbers: -------------------------------------------------------------------------------- 1 | echo "Domain library:" 2 | coqwc Categories.v NSetoid.v PredomCore.v PredomFix.v PredomKleisli.v PredomLift.v PredomRec.v PredomSum.v PredomAll.v 3 | 4 | echo "Typed PCF:" 5 | coqwc typedadequacy.v typeddensem.v typedlambda.v typedopsem.v typedsoundness.v typedsubst.v 6 | 7 | echo "Uni-Typed CVB Lambda Calculus:" 8 | coqwc uniiop.v unii.v uniirec.v uniisem.v uniisound.v uniiade.v KnasterTarski.v 9 | 10 | echo "Metric library:" 11 | coqwc MetricCore.v MetricRec.v 12 | 13 | echo "Kitchen Sink:" 14 | coqwc mpremet.v KSSem.v Finmap.v FinmapMetric.v KSOp.v KSTyping.v KSTm.v KSTy.v Fin.v 15 | 16 | 17 | 18 | 19 | 20 | 21 | 22 | 23 | -------------------------------------------------------------------------------- /src/numbers.bat: -------------------------------------------------------------------------------- 1 | echo "Domain library:" 2 | coqwc Categories.v NSetoid.v PredomCore.v PredomFix.v PredomKleisli.v PredomLift.v PredomRec.v PredomSum.v PredomAll.v 3 | 4 | echo "Typed PCF:" 5 | coqwc typedadequacy.v typeddensem.v typedlambda.v typedopsem.v typedsoundness.v typedsubst.v 6 | 7 | echo "Uni-Typed CVB Lambda Calculus:" 8 | coqwc uniiop.v unii.v uniirec.v uniisem.v uniisound.v uniiade.v KnasterTarski.v 9 | 10 | echo "Metric library:" 11 | coqwc MetricCore.v MetricRec.v 12 | 13 | echo "Kitchen Sink:" 14 | coqwc mpremet.v KSSem.v Finmap.v FinmapMetric.v KSOp.v KSTyping.v KSTm.v KSTy.v Fin.v 15 | 16 | 17 | 18 | 19 | 20 | 21 | 22 | 23 | -------------------------------------------------------------------------------- /src/typedadequacy.v: -------------------------------------------------------------------------------- 1 | (********************************************************************************** 2 | * typedadequacy.v * 3 | * Formalizing Domains, Ultrametric Spaces and Semantics of Programming Languages * 4 | * Nick Benton, Lars Birkedal, Andrew Kennedy and Carsten Varming * 5 | * Jan 2012 * 6 | * Build with Coq 8.3pl2 plus SSREFLECT * 7 | **********************************************************************************) 8 | 9 | 10 | (* Require Import utility.*) 11 | Require Import PredomAll. 12 | Require Import typedlambda. 13 | Require Import typedopsem. 14 | Require Import typeddensem. 15 | 16 | Set Implicit Arguments. 17 | Unset Strict Implicit. 18 | Import Prenex Implicits. 19 | 20 | 21 | (*========================================================================== 22 | Logical relation between semantics and terms, used to prove adequacy 23 | See Winskel Chapter 11, Section 11.4. 24 | ==========================================================================*) 25 | 26 | (** printing senv %\ensuremath{\rho}% *) 27 | (** printing d1 %\ensuremath{d_1}% *) 28 | 29 | (* Lift a value-relation to a computation-relation *) 30 | (*=liftRel *) 31 | Definition liftRel t (R : SemTy t -> CValue t -> Prop) := 32 | fun d e => forall d', d =-= Val d' -> exists v, e =>> v /\ R d' v. (*CLEAR*) 33 | (* = End *) 34 | 35 | (* Logical relation between semantics and closed value terms *) 36 | Unset Implicit Arguments. 37 | 38 | (* = relValEnvExp *) 39 | (*CLEARED*) 40 | Fixpoint relVal t : SemTy t -> CValue t -> Prop := 41 | match t with 42 | | Int => fun d v => v = TINT d 43 | | Bool => fun d v => v = TBOOL d 44 | | t1 --> t2 => fun d v => exists e, v = TFIX e /\ 45 | forall d1 v1, relVal t1 d1 v1 -> liftRel (relVal t2) (d d1) (subExp [ v1, v ] e) 46 | | t1 ** t2 => fun d v => exists v1, exists v2, 47 | v = TPAIR v1 v2 /\ relVal t1 (fst d) v1 /\ relVal t2 (snd d) v2 48 | end. (*CLEAR*) 49 | 50 | Fixpoint relEnv E : SemEnv E -> Sub E nil -> Prop := 51 | match E with 52 | | nil => fun _ _ => True 53 | | t :: E => fun d s => relVal t (snd d) (hdMap s) /\ relEnv E (fst d) (tlMap s) 54 | end. 55 | (*CLEARED*) 56 | Definition relExp ty := liftRel (relVal ty). 57 | (*=End *) 58 | (* Computation relation, already inlined into relVal *) 59 | 60 | Set Implicit Arguments. 61 | Unset Strict Implicit. 62 | 63 | Lemma relVal_lower_lift : 64 | forall ty, 65 | (forall d d' (v:CValue ty), d <= d' -> relVal ty d' v -> relVal ty d v) -> 66 | (forall d d' (e:CExp ty), d <= d' -> relExp ty d' e -> relExp ty d e). 67 | rewrite /relExp /liftRel. move => t IH d d' e l R d0 eq. 68 | rewrite -> eq in l. clear d eq. case: (DLle_Val_exists_eq l) => d'v [e0 l1]. 69 | specialize (R _ e0). case: R => v [ev rv]. exists v. split ; first by []. 70 | by apply (IH _ _ _ l1). 71 | Qed. 72 | 73 | (* Lemma 11.12 (ii) from Winskel *) 74 | (*=relValLower *) 75 | Lemma relVal_lower: forall t d d' v, d <= d' -> relVal t d' v -> relVal t d v. 76 | (*CLEAR*) 77 | 78 | elim. 79 | 80 | (* Int *) 81 | simpl. move => d d' v. by case => e ; rewrite e. 82 | 83 | (* Bool *) 84 | simpl. move => d d' v. by case: d ; case: d'. 85 | 86 | (* Arrow *) 87 | move => t IH t' IH' d d' v l /=. case => b. case => e ; rewrite e. clear v e. 88 | move => C. exists b. split ; first by []. move => d1 v1 r1. specialize (C d1 v1 r1). 89 | apply (relVal_lower_lift IH' (l d1)). by apply C. 90 | 91 | (* Pair *) 92 | move => t IH t' IH' p1 p2 v L. case => v0 ; case => v1 [e l]. rewrite e. clear v e. 93 | simpl. exists v0. exists v1. split ; first by []. 94 | specialize (IH _ _ v0 (proj1 L)). specialize (IH' _ _ v1 (proj2 L)). split ; [apply IH | apply IH'] ; apply l. 95 | Qed. 96 | 97 | Add Parametric Morphism ty (v:CValue ty) : (fun d => relVal ty d v) 98 | with signature (@tset_eq (SemTy ty)) ==> iff as relVal_eq_compat. 99 | intros x y xyeq. 100 | destruct xyeq as [xy1 xy2]. 101 | split; apply relVal_lower; trivial. 102 | Qed. 103 | 104 | Add Parametric Morphism t (v:CExp t) : (fun d => relExp t d v) 105 | with signature (@tset_eq ((SemTy t)_BOT)) ==> iff as relExp_eq_compat. 106 | intros x y xyeq. 107 | destruct xyeq as [xy1 xy2]. 108 | split; apply (relVal_lower_lift (relVal_lower (t:=t))); trivial. 109 | Qed. 110 | 111 | (* ========================================================================== 112 | Admissibility 113 | ========================================================================== *) 114 | 115 | Lemma rel_admissible_lift : forall ty, 116 | (forall v, admissible (fun d => relVal ty d v)) -> 117 | (forall e, admissible (fun d => relExp ty d e)). 118 | rewrite /admissible /relExp /liftRel. move => t IH e c C d eq. 119 | case: (lubval eq) => k [dv [ck dvd]]. 120 | case: (C _ _ ck) => v [ev rv]. 121 | exists v. split ; first by []. 122 | destruct (DLvalgetchain ck) as [c' P]. 123 | specialize (IH v c'). 124 | have El: ((@lub (SemTy t) c') =-= d). apply: vinj. rewrite <- eq. 125 | apply tset_trans with (y:=eta (lub c')) ; first by []. 126 | rewrite -> (lub_comp_eq). rewrite (lub_lift_left c k). apply: lub_eq_compat. 127 | apply fmon_eq_intro => n. simpl. by rewrite <- P. 128 | 129 | refine (relVal_lower (proj2 El) _). 130 | apply IH => n. 131 | specialize (P n). specialize (C _ _ P). case: C => v' [ev' rv']. 132 | rewrite (Determinacy ev ev'). by apply rv'. 133 | Qed. 134 | 135 | (* Lemma 11.12 (iii) from Winskel *) 136 | 137 | (*CLEARED*)Lemma rel_admissible: forall t v, admissible (fun d => relVal t d v). 138 | (*=End *) 139 | elim. 140 | 141 | (* Int *) 142 | intros v c C. simpl. apply (C 0). 143 | 144 | (* Bool *) 145 | intros v c C. simpl. specialize (C 0). simpl in C. have e:lub c = c 0. 146 | unfold lub. simpl. unfold sum_lub. simpl. case: (SuminjProof c 0). 147 | + simpl. case. case. move => e. by rewrite {2} e. 148 | + simpl. case. case. move => e. by rewrite {2} e. 149 | by rewrite e. 150 | 151 | (* Arrow *) 152 | rewrite /admissible => t IH t' IH' v c C. 153 | simpl. 154 | assert (rv := C 0). simpl in rv. destruct rv as [b [tbeq C0]]. 155 | exists b. 156 | split; first by []. 157 | intros d1 v1 R1. 158 | refine (rel_admissible_lift _ _). 159 | intros vv cc CC. apply IH'. apply CC. 160 | 161 | intros n. unfold relExp. unfold liftRel. 162 | intros d' CC. 163 | assert (Cn := C n). simpl in Cn. 164 | destruct Cn as [bb [tteq Cn]]. rewrite tteq in tbeq C C0 Cn. rewrite tteq. clear v tteq. 165 | assert (b = bb) as bbeq by apply (TFIX_injective (sym_equal tbeq)). 166 | rewrite <- bbeq in C0, C, Cn. rewrite <- bbeq. clear tbeq bb bbeq. 167 | specialize (Cn _ _ R1). 168 | unfold liftRel in Cn. specialize (Cn _ CC). by apply Cn. 169 | 170 | (* Pair *) 171 | unfold admissible. 172 | move => t IH t' IH' v c C. simpl. unfold prod_lub. 173 | generalize C. 174 | destruct (ClosedPair v) as [v1 [v2 veq]]. subst. 175 | intros _. 176 | exists v1. exists v2. 177 | split;first by []. 178 | specialize (IH v1). specialize (IH' v2). 179 | split. 180 | - apply: IH. move => n. specialize (C n). case: C => v3 [v4 [veq [C1 C2]]]. 181 | destruct (TPAIR_injective veq) as [E1 E2]. subst. by apply C1. 182 | - apply: IH'. move => n. specialize (C n). case: C => v3 [v4 [veq [C1 C2]]]. 183 | destruct (TPAIR_injective veq) as [E1 E2]. subst. by apply C2. 184 | Qed. 185 | 186 | Lemma SemEnvCons : forall t E (env : SemEnv (t :: E)), exists d, exists ds, env =(ds, d). 187 | intros. dependent inversion env. exists s0. by exists s. 188 | Qed. 189 | 190 | Lemma Discrete_injective : forall t (v1 v2:discrete_cpoType t), v1 =-= v2 -> v1 = v2. 191 | intros t v1 v2 eq. 192 | destruct eq. destruct H. destruct H0. trivial. 193 | Qed. 194 | 195 | Lemma relEnv_ext env senv s s' : (forall x y, s x y = s' x y) -> relEnv env senv s -> relEnv env senv s'. 196 | elim:env senv s s' ; first by []. 197 | move => a E IH env s s' C r. 198 | simpl. simpl in r. destruct r as [rl rr]. 199 | split. unfold hdMap. rewrite <- (C a (ZVAR E a)). by apply rl. 200 | unfold tlMap. 201 | refine (IH _ _ _ _ rr). intros x y. by specialize (C x (SVAR a y)). 202 | Qed. 203 | 204 | (*=FT *) 205 | Theorem FundamentalTheorem E: 206 | (forall t v senv s, relEnv E senv s -> relVal t (SemVal v senv) (subVal s v)) /\ 207 | (forall t e senv s, relEnv E senv s -> liftRel (relVal t) (SemExp e senv) (subExp s e)). 208 | (*=End *) 209 | Proof. 210 | move: E ; apply ExpVal_ind. 211 | 212 | (* TINT *) 213 | by simpl; auto. 214 | 215 | (* TBOOL *) 216 | simpl. move=> E. by case. 217 | 218 | (* TVAR *) 219 | intros E t v env s Renv. 220 | induction v. 221 | destruct (SemEnvCons env) as [d [ds eq]]. subst. simpl. fold SemEnv in *. 222 | simpl in Renv. destruct Renv as [R1 _]. by trivial. 223 | destruct (SemEnvCons env) as [d [ds eq]]. subst. simpl. fold SemEnv in *. 224 | simpl in Renv. destruct Renv as [_ R2]. 225 | specialize (IHv ds (tlMap s) R2). by auto. 226 | 227 | (* TFIX *) 228 | intros E t1 t2 e IH env s R. 229 | simpl SemVal. 230 | rewrite substTFIX. 231 | apply (@fixp_ind _ ((exp_fun (exp_fun (SemExp e)) env)) _ (@rel_admissible (t1 --> t2) _)). 232 | simpl. exists (subExp (liftSub t1 (liftSub (t1 --> t2) s)) e) ; split ; auto. 233 | intros d1 v1 rv1. unfold liftRel. intros dd. 234 | intros CC. by destruct (PBot_incon_eq (Oeq_sym CC)). 235 | 236 | intros f rf. 237 | exists (subExp (liftSub t1 (liftSub (t1 --> t2) s)) e). split ; first by []. 238 | intros d1 v1 rv1. unfold liftRel. intros dd fd. 239 | assert (SemExp e (env,f,d1) =-= Val dd) as se by auto. clear fd. 240 | rewrite <- (proj2 (substComposeSub _)). 241 | rewrite <- substTFIX. 242 | assert (relEnv ((t1 :: t1 --> t2 :: E)) (env,f,d1) (consMap v1 (consMap (subVal (env':=nil) s (TFIX e)) s))). 243 | simpl relEnv. split. 244 | unfold hdMap. simpl. by apply rv1. split. 245 | unfold tlMap. unfold hdMap. simpl. rewrite substTFIX. exists (subExp 246 | (liftSub t1 (liftSub (t1 --> t2) s)) e). split ; first by []. 247 | intros d2 v2 rv2. 248 | simpl in rf. destruct rf as [bb [bbeq Pb]]. 249 | assert (bb = (subExp (liftSub t1 (liftSub (t1 --> t2) s)) e)) as beq. 250 | apply (TFIX_injective (sym_equal bbeq)). 251 | rewrite beq in bbeq, Pb. clear bb beq bbeq. specialize (Pb d2 v2 rv2). by apply Pb. 252 | 253 | simpl. unfold tlMap. simpl. 254 | refine (relEnv_ext _ R). by intros x y ; auto. 255 | 256 | specialize (IH _ _ H). unfold liftRel in IH. specialize (IH _ se). 257 | rewrite composeCons. rewrite composeCons. 258 | rewrite composeSubIdLeft. by apply IH. 259 | 260 | (* TP *) 261 | intros E t1 t2 v1 IH1 v2 IH2 env s R. 262 | simpl. specialize (IH1 env s R). specialize (IH2 env s R). 263 | exists (subVal s v1). exists (subVal s v2). by split; auto. 264 | 265 | (* TFST *) 266 | intros E t1 t2 v IH env s R. 267 | simpl. specialize (IH env s R). 268 | unfold liftRel. intros d1 eq. 269 | rewrite substTFST. 270 | destruct (ClosedPair (subVal s v)) as [v1 [v2 veq]]. 271 | rewrite veq in IH. rewrite veq. clear veq. 272 | exists v1. split. apply: e_Fst. simpl relVal in IH. 273 | destruct IH as [v3 [v4 [veq [gl1 gl2]]]]. destruct (TPAIR_injective veq) as [v1eq v2eq]. subst. clear veq. 274 | assert (H := vinj eq). 275 | by apply (relVal_lower (proj2 H) gl1). 276 | 277 | (* TSND *) 278 | intros E t1 t2 v IH env s R. 279 | simpl. fold SemTy. specialize (IH env s R). 280 | unfold liftRel. intros d2 eq. 281 | rewrite substTSND. 282 | destruct (ClosedPair (subVal s v)) as [v1 [v2 veq]]. 283 | rewrite veq in IH. rewrite veq. clear veq. 284 | exists v2. split. apply e_Snd. simpl relVal in IH. 285 | destruct IH as [v3 [v4 [veq [gl1 gl2]]]]. destruct (TPAIR_injective veq) as [v1eq v2eq]. subst. clear veq. 286 | assert (H := vinj eq). 287 | by apply (relVal_lower (proj2 H) gl2). 288 | 289 | (* TOP *) 290 | intros E op v1 IH1 v2 IH2 env s R. 291 | simpl. specialize (IH1 env s R). specialize (IH2 env s R). 292 | unfold liftRel. intros d eq. 293 | rewrite substTOP. 294 | destruct (ClosedInt (subVal s v1)) as [i1 eq1]. 295 | destruct (ClosedInt (subVal s v2)) as [i2 eq2]. 296 | rewrite -> eq1 in IH1. rewrite eq1. rewrite eq2 in IH2. rewrite eq2. clear eq1 eq2. 297 | exists (TINT (op i1 i2)). 298 | split. apply e_Op. simpl in IH1, IH2. case: IH1 => IH1. rewrite -> IH1. 299 | case: IH2 => e ; rewrite -> e. case: (vinj eq). simpl => ee. by rewrite <- ee. 300 | 301 | (* TGT *) 302 | intros E v1 IH1 v2 IH2 env s R. 303 | simpl. specialize (IH1 env s R). specialize (IH2 env s R). 304 | unfold liftRel. intros d eq. have H:=vinj eq. 305 | rewrite substTGT. 306 | destruct (ClosedInt (subVal s v1)) as [i1 eq1]. 307 | destruct (ClosedInt (subVal s v2)) as [i2 eq2]. 308 | rewrite -> eq1 in IH1. rewrite eq1. rewrite eq2 in IH2. rewrite eq2. clear eq1 eq2. 309 | exists (TBOOL (i2 <= i1)%N). 310 | split. apply e_Gt. simpl in IH1,IH2. case: IH1 => e ; rewrite e. clear e. case: IH2 => e ; rewrite e. clear e. 311 | f_equal. 312 | have HH:(if (SemVal v2 env <= SemVal v1 env)%N then INL _ _ tt else INR _ _ tt) =-= d by apply H. clear H. 313 | case: (SemVal v2 env <= SemVal v1 env)%N HH ; simpl ; clear eq ; case: d ; simpl ; case ; by case. 314 | 315 | (* VAL *) 316 | intros E t v IH env s R. 317 | simpl. specialize (IH env s R). 318 | unfold liftRel. intros d eq. have H:=vinj eq. clear eq. 319 | exists (subVal s v). 320 | split. apply e_Val. 321 | apply (relVal_lower (proj2 H) IH). 322 | 323 | (* TLET *) 324 | intros E t1 t2 e1 IH1 e2 IH2 env s R. 325 | unfold liftRel. intros d2 d2eq. 326 | specialize (IH1 env s R). unfold liftRel in *. 327 | simpl in d2eq. 328 | 329 | rewrite substTLET. 330 | destruct (KLEISLIR_ValVal d2eq) as [d1 [d1eq d2veq]]. clear d2eq. 331 | specialize (IH1 _ d1eq). 332 | destruct IH1 as [v1 [ev1 R1]]. 333 | specialize (IH2 (env,d1) (consMap v1 s)). 334 | assert (RR : relEnv (t1::E) (env,d1) (consMap v1 s)). 335 | simpl relEnv. split ; first by []. 336 | by rewrite tlConsMap. 337 | specialize (IH2 RR d2 d2veq). 338 | destruct IH2 as [v2 [ev2 R2]]. 339 | exists v2. split. 340 | refine (e_Let ev1 _). 341 | rewrite <- (proj2 (substComposeSub _)). rewrite composeCons. rewrite composeSubIdLeft. apply ev2. by assumption. 342 | 343 | (* TAPP *) 344 | intros E t1 t2 v1 IH1 v2 IH2 env s R. 345 | specialize (IH1 env s R). specialize (IH2 env s R). 346 | unfold liftRel. intros d eq. rewrite substTAPP. 347 | destruct (ClosedFunction (subVal s v1)) as [e1 eq1]. 348 | unfold relVal in IH1. fold relVal in IH1. destruct IH1 as [e1' [eq' R']]. 349 | rewrite eq' in eq1. 350 | assert (H := TFIX_injective eq1). 351 | subst. clear eq1. 352 | specialize (R' (SemVal v2 env) (subVal s v2) IH2). 353 | clear IH2. 354 | unfold liftRel in R'. specialize (R' d eq). 355 | destruct R' as [v [ev Rv]]. 356 | exists v. rewrite eq' in ev. 357 | assert (H := e_App ev). 358 | split. simpl. rewrite eq'. apply H. 359 | by apply Rv. 360 | 361 | (* TIF *) 362 | intros E t v IH e1 IH1 e2 IH2 env s R. 363 | specialize (IH env s R). specialize (IH1 env s R). specialize (IH2 env s R). 364 | unfold liftRel in *. intros d eq. 365 | rewrite substTIF. 366 | destruct (ClosedBool (subVal s v)) as [b beq]. 367 | rewrite beq in IH. 368 | unfold relVal in IH. 369 | inversion IH. clear IH. 370 | simpl subExp. rewrite beq. 371 | rewrite H0 in beq. rewrite H0. clear b H0. simpl in eq. 372 | case: (SemVal v env) eq beq ; case => eq beq. 373 | - specialize (IH1 _ eq). case: IH1 => v0 [S0 R0]. exists v0. split ; last by []. 374 | by apply (e_IfTrue _ S0). 375 | - specialize (IH2 _ eq). case: IH2 => v0 [S0 R0]. exists v0. split ; last by []. 376 | by apply (e_IfFalse _ S0). 377 | Qed. 378 | 379 | (*=Adequacy *) 380 | Corollary Adequacy: forall t (e : CExp t) d, SemExp e tt =-= Val d -> exists v, e =>> v. 381 | (*=End *) 382 | Proof. 383 | intros t e d val. 384 | destruct (FundamentalTheorem nil) as [_ FT]. 385 | specialize (FT t e tt (idSub nil)). 386 | have X: (relEnv nil tt (idSub nil)) by []. 387 | specialize (FT X). unfold liftRel in FT. specialize (FT d val). 388 | destruct FT as [v [ev _]]. exists v. by rewrite (proj2 (applyIdSub _)) in ev. 389 | Qed. 390 | 391 | -------------------------------------------------------------------------------- /src/typeddensem.v: -------------------------------------------------------------------------------- 1 | (********************************************************************************** 2 | * typeddensem.v * 3 | * Formalizing Domains, Ultrametric Spaces and Semantics of Programming Languages * 4 | * Nick Benton, Lars Birkedal, Andrew Kennedy and Carsten Varming * 5 | * Jan 2012 * 6 | * Build with Coq 8.3pl2 plus SSREFLECT * 7 | **********************************************************************************) 8 | 9 | Require Export typedlambda. 10 | Require Import PredomAll. 11 | 12 | Set Implicit Arguments. 13 | Unset Strict Implicit. 14 | Import Prenex Implicits. 15 | 16 | 17 | Open Scope C_scope. 18 | Open Scope D_scope. 19 | Open Scope S_scope. 20 | 21 | (*========================================================================== 22 | Meaning of types and contexts 23 | ==========================================================================*) 24 | 25 | Delimit Scope Ty_scope with Ty. 26 | 27 | Open Scope Ty_scope. 28 | 29 | (*=SemTy *) 30 | Fixpoint SemTy t : cpoType := 31 | match t with 32 | | Int => nat_cpoType 33 | | Bool => (One:cpoType) + One 34 | | t1 --> t2 => SemTy t1 -=> (SemTy t2)_BOT 35 | | t1 ** t2 => SemTy t1 * SemTy t2 36 | end. 37 | Fixpoint SemEnv E : cpoType := 38 | match E with 39 | | nil => One 40 | | t :: E => SemEnv E * SemTy t 41 | end. 42 | (*=End *) 43 | 44 | (* Lookup in an environment *) 45 | (*=SemVarExpVal *) 46 | Fixpoint SemVar E t (var : Var E t) : SemEnv E =-> SemTy t := 47 | match var with 48 | | ZVAR _ _ => pi2 49 | | SVAR _ _ _ v => SemVar v << pi1 50 | end. (*CLEAR*) 51 | 52 | Lemma SimpleB_mon (A B :Type) (C:ordType) (op : A -> B -> C:Type) : monotonic (fun p:discrete_ordType A * discrete_ordType B => op (fst p) (snd p)). 53 | case => x0 y0. case => x1 y1. case ; simpl. move => L L'. 54 | have e:x0 = x1 by []. have e':y0 = y1 by []. rewrite e. by rewrite e'. 55 | Qed. 56 | 57 | Definition SimpleBOpm (A B :Type) (C:ordType) (op : A -> B -> C:Type) : discrete_ordType A * discrete_ordType B =-> C := 58 | mk_fmono (SimpleB_mon op). 59 | 60 | Lemma SimpleB_cont (A B:Type) (C:cpoType) (op : A -> B -> C:Type) : 61 | @continuous (discrete_cpoType A * discrete_cpoType B) C (SimpleBOpm op). 62 | move => c. simpl. apply: (Ole_trans _ (le_lub _ 0)). simpl. 63 | by []. 64 | Qed. 65 | 66 | Definition SimpleBOp (A B:Type) (C:cpoType) (op : A -> B -> C:Type) : discrete_cpoType A * discrete_cpoType B =-> C := 67 | Eval hnf in mk_fcont (SimpleB_cont op). 68 | 69 | Lemma choose_cont (Y Z: cpoType) (a:Z =-> Y) (b : Z =-> Y) (c:Z =-> (One:cpoType) + One) : 70 | forall x, 71 | (fun y => match c y with inl y' => a x | inr y' => b x end) x =-= 72 | (ev << <| SUM_Fun << <| exp_fun ((uncurry (const One a)) << SWAP), 73 | exp_fun ((uncurry (const One b)) << SWAP) |>, c|>) x. 74 | move => x. simpl. rewrite SUM_fun_simplx. by case_eq (c x) ; case => e. 75 | Qed. 76 | 77 | Definition choose (Y Z: cpoType) (a:Z =-> Y) (b : Z =-> Y) (c:Z =-> (One:cpoType) + One) := 78 | Eval hnf in gen_cont (choose_cont a b c). 79 | 80 | Lemma choose_comp (Y Z W : cpoType) (a:Z =-> Y) (b : Z =-> Y) (c:Z =-> (One:cpoType) + One) 81 | (h : W =-> Z) : choose a b c << h =-= choose (a << h) (b << h) (c << h). 82 | by apply: fmon_eq_intro => x. 83 | Qed. 84 | 85 | Add Parametric Morphism (A B:cpoType) : (@choose A B) 86 | with signature (@tset_eq _) ==> (@tset_eq _) ==> (@tset_eq _) ==> (@tset_eq _) 87 | as choose_eq_compat. 88 | move => f f' e g g' e' x x' ex. 89 | apply: fmon_eq_intro => y. simpl. move: (fmon_eq_elim ex y). clear ex. unfold FCont.fmono. simpl. 90 | case: (x y) ; case: (x' y) ; case ; case ; (try by case) ; move => _. by rewrite e. by rewrite e'. 91 | Qed. 92 | 93 | (* SemExpVal *) 94 | (*CLEARED*) 95 | Fixpoint SemExp E t (e : Exp E t) : SemEnv E =-> (SemTy t) _BOT := 96 | match e with 97 | | TOP op v1 v2 => eta << SimpleBOp op << <| SemVal v1 , SemVal v2 |> 98 | | TGT v1 v2 => eta << SimpleBOp (fun x y => if leq x y then inl _ tt else inr _ tt) 99 | << <| SemVal v2 , SemVal v1 |> 100 | | TAPP _ _ v1 v2 => ev << <| SemVal v1 , SemVal v2 |> 101 | | TVAL _ v => eta << SemVal v 102 | | TLET _ _ e1 e2 => KLEISLIR (SemExp e2) << <| Id , SemExp e1 |> 103 | | TIF _ v e1 e2 => choose (SemExp e1) (SemExp e2) (SemVal v) 104 | | TFST _ _ v => eta << pi1 << SemVal v 105 | | TSND _ _ v => eta << pi2 << SemVal v 106 | end with SemVal E t (v : Value E t) : SemEnv E =-> SemTy t := 107 | match v with 108 | | TINT n => const _ n 109 | | TBOOL b => const _ (if b then inl _ tt else inr _ tt) 110 | | TVAR _ i => SemVar i 111 | | TFIX t1 t2 e => (FIXP : cpoCatType _ _) << exp_fun (exp_fun (SemExp e)) 112 | | TPAIR _ _ v1 v2 => <| SemVal v1 , SemVal v2 |> 113 | end. 114 | (*=End *) 115 | -------------------------------------------------------------------------------- /src/typedopsem.v: -------------------------------------------------------------------------------- 1 | (********************************************************************************** 2 | * typedopsem.v * 3 | * Formalizing Domains, Ultrametric Spaces and Semantics of Programming Languages * 4 | * Nick Benton, Lars Birkedal, Andrew Kennedy and Carsten Varming * 5 | * Jan 2012 * 6 | * Build with Coq 8.3pl2 plus SSREFLECT * 7 | **********************************************************************************) 8 | 9 | Require Import Coq.Program.Equality. 10 | Require Import typedlambda. 11 | 12 | Set Implicit Arguments. 13 | Unset Strict Implicit. 14 | Import Prenex Implicits. 15 | 16 | 17 | (*========================================================================== 18 | Evaluation relation 19 | ==========================================================================*) 20 | 21 | (** printing =>> %\ensuremath{\Downarrow}% *) 22 | (** printing n1 %\ensuremath{n_1}% *) 23 | (** printing n2 %\ensuremath{n_2}% *) 24 | (** printing v1 %\ensuremath{v_1}% *) 25 | (** printing v2 %\ensuremath{v_2}% *) 26 | (** printing e1 %\ensuremath{e_1}% *) 27 | (** printing e2 %\ensuremath{e_2}% *) 28 | 29 | Open Scope Sub_scope. 30 | Notation "x '> v" (at level 70, no associativity). 32 | (*=Ev *) 33 | Inductive Ev: forall t, CExp t -> CValue t -> Prop := 34 | | e_Val: forall t (v : CValue t), TVAL v =>> v 35 | | e_Op: forall op n1 n2, TOP op (TINT n1) (TINT n2) =>> TINT (op n1 n2) 36 | | e_Gt : forall n1 n2, TGT (TINT n1) (TINT n2) =>> TBOOL (n2 > v1 38 | | e_Snd : forall t1 t2 (v1 : CValue t1) (v2 : CValue t2), TSND (TPAIR v1 v2) =>> v2 39 | | e_App : forall t1 t2 e (v1 : CValue t1) (v2 : CValue t2), 40 | subExp [ v1, TFIX e ] e =>> v2 -> TAPP (TFIX e) v1 =>> v2 41 | | e_Let : forall t1 t2 e1 e2 (v1 : CValue t1) (v2 : CValue t2), 42 | e1 =>> v1 -> subExp [ v1 ] e2 =>> v2 -> TLET e1 e2 =>> v2 43 | | e_IfTrue : forall t (e1 e2 : CExp t) v, e1 =>> v -> TIF (TBOOL true) e1 e2 =>> v 44 | | e_IfFalse : forall t (e1 e2 : CExp t) v, e2 =>> v -> TIF (TBOOL false) e1 e2 =>> v 45 | where "e '=>>' v" := (Ev e v). 46 | (*=End *) 47 | (*========================================================================== 48 | Determinacy of evaluation 49 | ==========================================================================*) 50 | 51 | Lemma Determinacy: forall ty, forall (e : CExp ty) v1, e =>> v1 -> forall v2, e =>> v2 -> v1 = v2. 52 | Proof. 53 | intros ty e v1 Ev1. 54 | induction Ev1. 55 | 56 | (* e_Val *) 57 | intros v2 Ev2. dependent destruction Ev2; trivial. 58 | (* e_Op *) 59 | intros v2 Ev2. dependent destruction Ev2; trivial. 60 | (* e_Gt *) 61 | intros v2 Ev2. dependent destruction Ev2; trivial. 62 | (* e_Fst *) 63 | intros v3 Ev3. dependent destruction Ev3; trivial. 64 | (* e_Snd *) 65 | intros v3 Ev3. dependent destruction Ev3; trivial. 66 | (* e_App *) 67 | intros v3 Ev3. dependent destruction Ev3; auto. 68 | (* e_Let *) 69 | intros v3 Ev3. dependent destruction Ev3. rewrite (IHEv1_1 _ Ev3_1) in IHEv1_2. auto. 70 | (* e_IfTrue *) 71 | intros v3 Ev3. dependent destruction Ev3; auto. 72 | (* e_IfFalse *) 73 | intros v3 Ev3. dependent destruction Ev3; auto. 74 | Qed. 75 | 76 | Unset Implicit Arguments. 77 | -------------------------------------------------------------------------------- /src/typedopsempat.v: -------------------------------------------------------------------------------- 1 | (* typedopsem.v 2 | Standard operational semantics 3 | *) 4 | Require Import utility. 5 | Require Import Coq.Program.Equality. 6 | Require Import EqdepFacts. 7 | Require Import typedlambdapat. 8 | 9 | Set Implicit Arguments. 10 | Unset Strict Implicit. 11 | 12 | (*========================================================================== 13 | Evaluation relation 14 | ==========================================================================*) 15 | 16 | (** printing =>> %\ensuremath{\Downarrow}% *) 17 | (** printing n1 %\ensuremath{n_1}% *) 18 | (** printing n2 %\ensuremath{n_2}% *) 19 | (** printing v1 %\ensuremath{v_1}% *) 20 | (** printing v2 %\ensuremath{v_2}% *) 21 | (** printing e1 %\ensuremath{e_1}% *) 22 | (** printing e2 %\ensuremath{e_2}% *) 23 | Reserved Notation "e =>> v" (at level 70, no associativity). 24 | 25 | Open Scope Subst_scope. 26 | 27 | Definition deconstructPair : forall ty1 ty2 (v : CValue (ty1 ** ty2)), CValue ty1 * CValue ty2. 28 | unfold CValue. 29 | intros. 30 | dependent destruction v. 31 | dependent destruction v. 32 | exact (v1, v2). 33 | Defined. 34 | 35 | Definition deconstructSum : forall ty1 ty2 (v : CValue (Sum ty1 ty2)), CValue ty1 + CValue ty2. 36 | unfold CValue. 37 | intros. 38 | dependent destruction v. 39 | dependent destruction v. 40 | exact (inl _ v). 41 | exact (inr _ v). 42 | Defined. 43 | 44 | Definition castSubst env env' env'' (s : option (Subst (env'' ++ (env' ++ env)) nil)) : option (Subst ((env'' ++ env') ++ env) nil). 45 | intros. 46 | rewrite app_ass. exact s. 47 | Defined. 48 | 49 | Implicit Arguments castSubst [env env' env'']. 50 | 51 | Lemma castSubstEq : forall env env' env'' s1 s2, @castSubst env env' env'' s1 = castSubst s2 -> s1 = s2. 52 | Proof. 53 | intros. destruct s1. destruct s2. assert (s = s0). apply MapExtensional. 54 | intros t var. inversion var. subst. rewrite <- H1 in var. redestruct env''. unfold castSubst in H. dependent destruction var. auto. 55 | case s1. 56 | case_eq s2. intros s2' s2'eq. subst. case_eq s1. intros s1' s1'eq. subst. assert (s1' = s2'). unfold castSubst in H. unfold eq_rect_r in H. 57 | unfold eq_rect in H. case_eq (sym_eq (app_ass env'' env' env)). intros. destruct H. auto. case H. auto. dependent destruction H. apply H. intros s1' s2'. unfold castSubst in H. 58 | unfold eq_rect_r in H. 59 | unfold eq_rect in H. 60 | dcase (sym_eq (app_ass env'' env' env)). inversion H. simpl in H. 61 | SearchAbout eq_rect. 62 | assert (EQ := eq_rect_eq). 63 | 64 | 65 | simpl_eq H. elim_eq_rect H. unfold eq_rect_r in H. dependent destruction H. auto. destruct H. case H. 66 | 67 | Fixpoint bindPat env' env (s : Subst env nil) ty (p : Pat env' ty) : CValue ty -> option (Subst (env' ++ env) nil) := 68 | match p with 69 | | PVAR _ => fun v => Some (consMap v s) 70 | | PWILD _ => fun v => Some s 71 | | PFAIL _ => fun v => None 72 | | PPAIR env1 env2 ty1 ty2 p1 p2 => fun v => 73 | let (v1, v2) := deconstructPair v 74 | in 75 | match @bindPat env2 _ s ty2 p2 v2 with 76 | | None => None 77 | | Some s' => castSubst (@bindPat env1 _ s' ty1 p1 v1) 78 | end 79 | 80 | | PAS _ _ p => fun v => 81 | match bindPat s p v with 82 | | None => None 83 | | Some s' => Some (consMap v s') 84 | end 85 | 86 | | PINL _ _ _ p' => fun v => 87 | match deconstructSum v with 88 | | inl v' => bindPat s p' v' 89 | | inr _ => None 90 | end 91 | 92 | | PINR _ _ _ p' => fun v => 93 | match deconstructSum v with 94 | | inl _ => None 95 | | inr v' => bindPat s p' v' 96 | end 97 | end. 98 | 99 | Implicit Arguments bindPat [env' env ty]. 100 | 101 | Program Fixpoint appendSubst env1 env2 (s1 : Subst env1 nil) (s2 : Subst env2 nil) : Subst (env1 ++ env2) nil := 102 | match env1 (*as env1 return Subst (env1 ++ env2) nil *) with 103 | | nil => s2 104 | | ty1::env1' => consMap (hdMap (t:=ty1) (E:=env1') s1) (appendSubst (tlMap (t:=ty1) s1) s2) 105 | end. 106 | 107 | Implicit Arguments appendSubst [env1 env2]. 108 | 109 | Notation "s1 ++ s2" := (appendSubst s1 s2) : Subst_scope. 110 | Arguments Scope appendSubst [_ _ Subst_scope Subst_scope]. 111 | 112 | Reserved Notation "v 'matches' p 'by' s" (at level 70, no associativity). 113 | 114 | Inductive EvPat : forall ty (env:Env), CValue ty -> Pat env ty -> Subst env nil -> Prop := 115 | | p_Var : forall ty (v:CValue ty), 116 | v matches PVAR ty by [v] 117 | | p_Pair : forall ty1 ty2 env1 env2 v1 v2 (p1:Pat env1 ty1) (p2:Pat env2 ty2) s1 s2, 118 | v1 matches p1 by s1 -> v2 matches p2 by s2 -> PAIR v1 v2 matches PPAIR p1 p2 by (s1 ++ s2) 119 | | p_Wild : forall ty v, 120 | v matches PWILD ty by idSubst _ 121 | | p_As : forall ty env v (p:Pat env ty) s, 122 | v matches p by s -> v matches (PAS p) by consMap v s 123 | | p_Inl : forall ty1 ty2 env v (p:Pat env ty1) s, 124 | v matches p by s -> INL ty2 v matches PINL _ p by s 125 | | p_Inr : forall ty1 ty2 env v (p:Pat env ty2) s, 126 | v matches p by s -> INR ty1 v matches PINR _ p by s 127 | where "v 'matches' p 'by' s" := (EvPat v p s). 128 | 129 | Definition castExp : forall env ty, Exp env ty -> Exp (env++nil)%list ty. 130 | intros. 131 | rewrite <- app_nil_end. exact H. 132 | Defined. 133 | 134 | Inductive Ev: forall ty, CExp ty -> CValue ty -> Prop := 135 | | e_Val : forall ty (v : CValue ty), VAL v =>> v 136 | | e_Op : forall op n1 n2, OP op (INT n1) (INT n2) =>> INT (op n1 n2) 137 | | e_Gt : forall n1 n2, GT (INT n1) (INT n2) =>> BOOL (ble_nat n2 n1) 138 | | e_Fst : forall ty1 ty2 (v1 : CValue ty1) (v2 : CValue ty2), FST (PAIR v1 v2) =>> v1 139 | | e_Snd : forall ty1 ty2 (v1 : CValue ty1) (v2 : CValue ty2), SND (PAIR v1 v2) =>> v2 140 | | e_App : forall ty1 ty2 e (v1 : CValue ty1) (v2 : CValue ty2), substExp [ v1, FIX e ] e =>> v2 -> APP (FIX e) v1 =>> v2 141 | | e_Let : forall ty1 ty2 e1 e2 (v1 : CValue ty1) (v2 : CValue ty2), e1 =>> v1 -> substExp [ v1 ] e2 =>> v2 -> LET e1 e2 =>> v2 142 | | e_LetPat : forall ty1 ty2 env e1 (p : Pat env ty1) e2 v1 (v2 : CValue ty2), 143 | e1 =>> v1 -> forall s, v1 matches p by s -> substExp s e2 =>> v2 -> LETPAT e1 p (castExp e2) =>> v2 144 | | e_IfTrue : forall ty (e1 e2 : CExp ty) v, e1 =>> v -> COND (BOOL true) e1 e2 =>> v 145 | | e_IfFalse : forall ty (e1 e2 : CExp ty) v, e2 =>> v -> COND (BOOL false) e1 e2 =>> v 146 | where "e '=>>' v" := (Ev e v). 147 | 148 | (*========================================================================== 149 | Determinacy of evaluation 150 | ==========================================================================*) 151 | 152 | Lemma PatDeterminacy: forall ty env (p : Pat env ty) v s1, v matches p by s1 -> forall s2, v matches p by s2 -> s1 = s2. 153 | admit. 154 | Qed. 155 | 156 | Lemma Determinacy: forall ty, forall (e : CExp ty) v1, e =>> v1 -> forall v2, e =>> v2 -> v1 = v2. 157 | Proof. 158 | intros ty e v1 Ev1. 159 | induction Ev1. 160 | 161 | (* e_Val *) 162 | intros v2 Ev2. dependent destruction Ev2; trivial. 163 | (* e_Op *) 164 | intros v2 Ev2. dependent destruction Ev2; trivial. 165 | (* e_Gt *) 166 | intros v2 Ev2. dependent destruction Ev2; trivial. 167 | (* e_Fst *) 168 | intros v3 Ev3. dependent destruction Ev3; trivial. 169 | (* e_Snd *) 170 | intros v3 Ev3. dependent destruction Ev3; trivial. 171 | (* e_App *) 172 | intros v3 Ev3. dependent destruction Ev3; auto. 173 | (* e_Let *) 174 | intros v3 Ev3. dependent destruction Ev3. rewrite (IHEv1_1 _ Ev3_1) in *. auto. 175 | (* e_LetPat *) 176 | intros v3 Ev3. dependent destruction Ev3. rewrite (IHEv1_1 _ Ev3_1) in *. assert (s = s0). apply (PatDeterminacy H H0). subst. clear H H0. 177 | assert (e2 = e3). unfold castExp in H1. dependent destruction H1. rewrite <- app_nil_end in H. destruct H1. simpl in H1. auto. dependent destruction H1. apply IHEv1_2. ; auto. rewrite H in H0. inversion H0. subst. auto. 178 | (* e_IfTrue *) 179 | intros v3 Ev3. dependent destruction Ev3; auto. 180 | (* e_IfFalse *) 181 | intros v3 Ev3. dependent destruction Ev3; auto. 182 | Qed. 183 | 184 | Unset Implicit Arguments. 185 | -------------------------------------------------------------------------------- /src/typedsoundness.v: -------------------------------------------------------------------------------- 1 | (********************************************************************************** 2 | * typedsoundness.v * 3 | * Formalizing Domains, Ultrametric Spaces and Semantics of Programming Languages * 4 | * Nick Benton, Lars Birkedal, Andrew Kennedy and Carsten Varming * 5 | * Jan 2012 * 6 | * Build with Coq 8.3pl2 plus SSREFLECT * 7 | **********************************************************************************) 8 | 9 | Require Import PredomAll. 10 | Require Import typeddensem. 11 | Require Import typedopsem. 12 | Require Import typedsubst. 13 | 14 | Set Implicit Arguments. 15 | Unset Strict Implicit. 16 | Import Prenex Implicits. 17 | 18 | 19 | (*=Soundness *) 20 | Theorem Soundness: forall t (e : CExp t) v, e =>> v -> SemExp e =-= eta << SemVal v. 21 | (*=End *) 22 | Proof. 23 | intros t e v Ev. 24 | induction Ev. 25 | 26 | (* e_Val *) 27 | simpl. trivial. 28 | 29 | (* e_Op *) 30 | simpl. rewrite <- comp_assoc. by apply: fmon_eq_intro. 31 | 32 | (* e_Gt *) 33 | simpl. by apply: fmon_eq_intro. 34 | 35 | (* e_Fst *) 36 | simpl. rewrite <- comp_assoc. by rewrite prod_fun_fst. 37 | 38 | (* e_Snd *) 39 | simpl. rewrite <- comp_assoc. by rewrite prod_fun_snd. 40 | 41 | (* e_App *) 42 | rewrite <- IHEv. clear IHEv. clear Ev. 43 | 44 | rewrite <- (proj2 (SemCommutesWithSub _) _ e nil (consMap v1 (consMap (TFIX e) (idSub _)))). 45 | simpl. rewrite hdConsMap. rewrite -> (terminal_unique _ Id). simpl. 46 | have a:= (@prod_fun_eq_compat cpoProdCatType _ _ _ _ _ (FIXP_com (exp_fun (exp_fun (SemExp e)))) _). 47 | specialize (a _ (SemVal v1) (SemVal v1) (tset_refl _)). rewrite a. clear a. 48 | rewrite - [FIXP << exp_fun _] comp_idL. 49 | rewrite - {1} [exp_fun _] comp_idR. 50 | rewrite - prod_map_prod_fun. rewrite comp_assoc. rewrite exp_com. 51 | rewrite - {1} [SemVal _] comp_idL. rewrite - prod_map_prod_fun. rewrite comp_assoc. by rewrite exp_com. 52 | 53 | (* e_Let *) 54 | simpl. rewrite <- IHEv2. clear IHEv2. 55 | rewrite <- (proj2 (SemCommutesWithSub _) _ _ _). 56 | simpl. rewrite <- (KLEISLIR_unit (SemExp e2) (SemVal v1) _). 57 | rewrite <- IHEv1. clear IHEv1. 58 | by rewrite <- (terminal_unique Id _). 59 | 60 | (* e_Iftrue *) 61 | simpl. rewrite IHEv. by apply: fmon_eq_intro. 62 | 63 | (* e_Iffalse *) 64 | simpl. rewrite IHEv. by apply: fmon_eq_intro. 65 | Qed. 66 | -------------------------------------------------------------------------------- /src/typedsubst.v: -------------------------------------------------------------------------------- 1 | (********************************************************************************** 2 | * typedsubst.v * 3 | * Formalizing Domains, Ultrametric Spaces and Semantics of Programming Languages * 4 | * Nick Benton, Lars Birkedal, Andrew Kennedy and Carsten Varming * 5 | * Jan 2012 * 6 | * Build with Coq 8.3pl2 plus SSREFLECT * 7 | **********************************************************************************) 8 | 9 | Require Import PredomAll. 10 | Require Import typedlambda. 11 | Require Import typeddensem. 12 | 13 | Set Implicit Arguments. 14 | Unset Strict Implicit. 15 | Import Prenex Implicits. 16 | 17 | 18 | (* Need this for nice pretty-printing *) 19 | Unset Printing Implicit Defensive. 20 | 21 | (*========================================================================== 22 | Semantics of substitution and renaming 23 | ==========================================================================*) 24 | 25 | Section SEMMAP. 26 | 27 | Variable P : Env -> Ty -> Type. 28 | Variable ops : MapOps P. 29 | Variable Sem : forall E t, P E t -> SemEnv E =-> SemTy t. 30 | Variable SemVl : forall E t (v : P E t), Sem v =-= SemVal (vl ops v). 31 | Variable SemVr : forall E t, Sem (vr ops (ZVAR E t)) =-= pi2. 32 | Variable SemWk : forall E t (v:P E t) t', Sem (wk ops t' v) =-= Sem v << pi1. 33 | 34 | (* Apply semantic function pointwise to a renaming or substitution *) 35 | (*=SemSub *) 36 | Fixpoint SemMap E E' : Map P E' E -> SemEnv E =-> SemEnv E' := 37 | match E' with 38 | | nil => fun m => terminal_morph (SemEnv E) 39 | | _ :: _ => fun m => <| SemMap (tlMap m) , Sem (hdMap m) |> 40 | end. 41 | (*=End *) 42 | 43 | Lemma SemShift : forall E E' (m : Map P E E') t, SemMap (shiftMap ops t m) =-= SemMap m << pi1. 44 | Proof. elim. 45 | - move => E' m t. by apply: terminal_unique. 46 | - move => t s IH E' m ty. 47 | simpl. destruct (consMapInv m) as [r' [var eq]]. 48 | subst. simpl. rewrite hdConsMap. rewrite tlConsMap. 49 | rewrite -> (IH E' r' ty). rewrite prod_fun_compl. unfold shiftMap, hdMap. by rewrite SemWk. 50 | Qed. 51 | 52 | Lemma SemLift : forall E E' (m : Map P E E') t, SemMap (lift ops t m) =-= SemMap m >< Id. 53 | Proof. move => E E' m t. simpl. unfold tlMap, hdMap. simpl lift. rewrite SemVr. rewrite SemShift. 54 | apply: prod_unique. rewrite prod_fun_fst. by rewrite prod_fun_fst. 55 | rewrite prod_fun_snd. rewrite prod_fun_snd. by rewrite comp_idL. 56 | Qed. 57 | 58 | (* 59 | Lemma SemId : forall E E' (m : Map P E E'), SemMap (fun t (v:Var E t) => vr ops v) =-= Id. 60 | Proof. elim. 61 | - intros. simpl. by apply: terminal_unique. 62 | - intros. simpl. apply:prod_unique. simpl. rewrite <- H. simpl in H. rewrite H. simpl. rewrite H. rewrite tlConsMap hdConsMap. rewrite SemShift. rewrite IH. rewrite comp_idL. rewrite SemVr. 63 | apply: prod_unique. rewrite prod_fun_fst. by rewrite comp_idR. rewrite prod_fun_snd. by rewrite comp_idR. 64 | Qed. 65 | *) 66 | Lemma SemCommutesWithMap E: 67 | (forall t (v : Value E t) E' (r : Map _ E E'), SemVal v << SemMap r =-= SemVal (mapVal ops r v)) 68 | /\ (forall t (exp : Exp E t) E' (r : Map _ E E'), SemExp exp << SemMap r =-= SemExp (mapExp ops r exp)). 69 | Proof. move: E ; apply ExpVal_ind. 70 | 71 | (* TINT *) by []. 72 | (* TBOOL *) intros. simpl. by rewrite <- const_com. 73 | 74 | (* TVAR *) 75 | intros E ty var E' r. 76 | simpl. 77 | induction var. 78 | simpl. rewrite prod_fun_snd. 79 | destruct (consMapInv r) as [r' [v eq]]. subst. 80 | simpl. rewrite hdConsMap. unfold mapVal. unfold travVal. by rewrite SemVl. 81 | destruct (consMapInv r) as [r' [v eq]]. subst. 82 | simpl. 83 | specialize (IHvar r'). 84 | rewrite <- IHvar. rewrite <- comp_assoc. 85 | rewrite tlConsMap. by rewrite prod_fun_fst. 86 | 87 | (* TFIX *) 88 | intros E t t1 body IH E' s. 89 | specialize (IH _ (lift ops _ (lift ops _ s))). 90 | rewrite mapTFIX. 91 | simpl SemVal. rewrite <- comp_assoc. 92 | do 2 rewrite exp_comp. rewrite <- IH. rewrite SemLift. by rewrite SemLift. 93 | 94 | (* TPAIR *) 95 | intros E t1 t2 v1 IH1 v2 IH2 E' s. rewrite mapTPAIR. simpl. rewrite <- IH1. rewrite <- IH2. by rewrite <- prod_fun_compl. 96 | 97 | (* TFST *) 98 | intros E t1 t2 v IH E' s. rewrite mapTFST. simpl. rewrite <- IH. by repeat (rewrite comp_assoc). 99 | 100 | (* TSND *) 101 | intros E t1 t2 v IH E' s. rewrite mapTSND. simpl. rewrite <- IH. by repeat (rewrite comp_assoc). 102 | 103 | (* TOP *) 104 | intros E n v1 IH1 v2 IH2 E' s. rewrite mapTOP. simpl. 105 | repeat (rewrite <- comp_assoc). rewrite <- IH1. rewrite <- IH2. by rewrite prod_fun_compl. 106 | 107 | (* TGT *) 108 | intros E v1 IH1 v2 IH2 E' s. rewrite mapTGT. simpl. 109 | repeat (rewrite <- comp_assoc). rewrite <- IH1. rewrite <- IH2. by rewrite prod_fun_compl. 110 | 111 | (* TVAL *) 112 | intros E t v IH E' s. rewrite mapTVAL. simpl. 113 | rewrite <- IH. by rewrite <- comp_assoc. 114 | 115 | (* TLET *) 116 | intros E t1 t2 e2 IH2 e1 IH1 E' s. rewrite mapTLET. simpl. 117 | rewrite <- comp_assoc. 118 | specialize (IH2 _ s). 119 | specialize (IH1 _ (lift ops _ s)). 120 | rewrite prod_fun_compl. rewrite KLEISLIR_comp. 121 | rewrite <- IH2. rewrite <- IH1. rewrite SemLift. 122 | by do 2 rewrite comp_idL. 123 | 124 | (* TAPP *) 125 | intros E t1 t2 v1 IH1 v2 IH2 E' s. rewrite mapTAPP. simpl. 126 | rewrite <- comp_assoc. rewrite <- IH1. rewrite <- IH2. by rewrite prod_fun_compl. 127 | 128 | (* TIF *) 129 | intros E t ec IHc te1 IH1 te2 IH2 E' s. rewrite mapTIF. simpl. 130 | rewrite choose_comp. rewrite -> IH1. rewrite -> IH2. by rewrite -> IHc. 131 | Qed. 132 | 133 | End SEMMAP. 134 | 135 | Definition SemRen := SemMap SemVar. 136 | Definition SemSub := SemMap SemVal. 137 | 138 | Lemma SemCommutesWithRen E: 139 | (forall t (v : Value E t) E' (r : Ren E E'), SemVal v << SemRen r =-= SemVal (renameVal r v)) 140 | /\ (forall t (e : Exp E t) E' (r : Ren E E'), SemExp e << SemRen r =-= SemExp (renameExp r e)). 141 | Proof. by apply SemCommutesWithMap. Qed. 142 | 143 | Lemma SemShiftRen : forall E E' (r : Ren E E') t, SemRen (shiftRen t r) =-= SemRen r << pi1. 144 | Proof. by apply SemShift. Qed. 145 | 146 | Lemma SemIdRen : forall E, SemRen (idRen E) =-= Id. 147 | elim. 148 | - simpl. by apply: terminal_unique. 149 | - move => t E IH. simpl. rewrite idRenDef. rewrite tlConsMap. 150 | rewrite SemShiftRen. rewrite IH. rewrite comp_idL. 151 | by apply: prod_unique ; [rewrite prod_fun_fst | rewrite prod_fun_snd] ; rewrite comp_idR. 152 | Qed. 153 | 154 | (*=Substitution *) 155 | Lemma SemCommutesWithSub E: 156 | (forall t (v : Value E t) E' (s : Sub E E'), SemVal v << SemSub s =-= SemVal (subVal s v)) 157 | /\ (forall t (e : Exp E t) E' (s : Sub E E'), SemExp e << SemSub s =-= SemExp (subExp s e)). 158 | (*=End *) 159 | move: E. apply SemCommutesWithMap. 160 | + by []. + by []. + intros. simpl. 161 | rewrite <- (proj1 (SemCommutesWithRen E)). assert (SSR := SemShiftRen). unfold shiftRen in SSR. unfold shiftMap in SSR. simpl in SSR. 162 | specialize (SSR E E (fun t v => v) t'). simpl in SSR. 163 | assert ((fun ty' (var : Var E ty') => SVAR t' var) = (fun t0 => SVAR t')). apply MapExtensional. auto. rewrite H in SSR. rewrite SSR. 164 | rewrite SemIdRen. by rewrite comp_idL. 165 | Qed. 166 | -------------------------------------------------------------------------------- /src/typedsubstalt.v: -------------------------------------------------------------------------------- 1 | Require Import utility. 2 | Require Import PredomAux. 3 | Require Import typedlambda. 4 | Require Import typeddensem. 5 | Require Import Program.Equality. 6 | 7 | Set Implicit Arguments. 8 | Unset Strict Implicit. 9 | 10 | (* Need this for nice pretty-printing *) 11 | Unset Printing Implicit Defensive. 12 | 13 | (*========================================================================== 14 | Semantics of substitution and renaming 15 | ==========================================================================*) 16 | 17 | Fixpoint SemMap (P:Env -> Ty -> Type) (sem : forall env ty, P env ty -> SemEnv env -c> SemTy ty) env env' : Map P env' env -> SemEnv env -c> SemEnv env' := 18 | match env' return Map P env' env -> SemEnv env -c> SemEnv env' with 19 | | nil => fun s => K _ (tt:DOne) 20 | | _ => fun s => PROD_fun (SemMap sem (tlMap s)) (sem _ _ (hdMap s)) 21 | end. 22 | 23 | Definition semWk P (k:Kit P) (sem:forall E t, P E t -> SemEnv E -c> SemTy t) := forall E t0 t (v:P E t), sem _ _ (wk k t0 v) == sem E t v @_ FST. 24 | Definition semVr P (k:Kit P) (sem:forall E t, P E t -> SemEnv E -c> SemTy t) := forall E t0, sem _ _ (vr k (ZVAR E t0)) == SND. 25 | Definition semVl P (k:Kit P) (sem:forall E t, P E t -> SemEnv E -c> SemTy t) := forall E t (v:P E t), sem _ _ v = SemVal (vl k v). 26 | 27 | Lemma SemShiftMap : 28 | forall P (k:Kit P) sem, semWk k sem -> forall E E' (s:Map P E E') t, SemMap sem (shiftMap (K:=k) t s) == SemMap sem s @_ FST. 29 | Proof. 30 | intros P k sem semwk. 31 | induction E. intros. simpl. 32 | apply (DOne_final (K (Dprod (SemEnv E') (SemTy t)) (tt:DOne)) _). 33 | 34 | intros. 35 | simpl. destruct (consMapInv s) as [s' [var eq]]. 36 | subst. simpl. rewrite hdConsMap. rewrite tlConsMap. 37 | specialize (IHE E' s' t). 38 | rewrite PROD_fun_comp. 39 | refine (PROD_fun_eq_compat _ _). 40 | rewrite IHE. auto. 41 | 42 | rewrite shiftConsMap. rewrite hdConsMap. 43 | unfold semWk in semwk. 44 | rewrite semwk. auto. 45 | Qed. 46 | 47 | Lemma SemLiftMap : 48 | forall P (k:Kit P) sem, semWk k sem -> forall E E' (s:forall t, Var E t -> P E' t) t0, SemMap sem (tlMap (lift (K:=k) t0 s)) == SemMap sem s @_ FST. 49 | Proof. 50 | induction E. intros. simpl. 51 | apply (DOne_final (K (Dprod (SemEnv E') (SemTy t0)) (tt:DOne)) _). 52 | 53 | intros. 54 | simpl. destruct (consMapInv s) as [s' [var eq]]. 55 | subst. rewrite hdConsMap. rewrite tlConsMap. 56 | specialize (IHE E' s' t0). 57 | rewrite PROD_fun_comp. 58 | refine (PROD_fun_eq_compat _ _). 59 | rewrite IHE. auto.unfold tlMap. simpl. unfold hdMap. simpl. 60 | unfold semWk in H. rewrite H. auto. 61 | Qed. 62 | 63 | Lemma SemLift2Map : 64 | forall P (k:Kit P) sem, semWk k sem -> forall E E' (s:Map P E E') t t', SemMap sem (tlMap (tlMap (lift (K:=k) t' (lift (K:=k) t s)))) == SemMap sem s @_ FST @_ FST. 65 | Proof. 66 | induction E. intros. simpl. 67 | apply (DOne_final (K (Dprod (Dprod (SemEnv E') (SemTy t)) (SemTy t')) (tt:DOne)) _). 68 | 69 | intros. 70 | simpl. destruct (consMapInv s) as [s' [var eq]]. 71 | subst. simpl. rewrite hdConsMap. rewrite tlConsMap. 72 | specialize (IHE E' s' t). 73 | rewrite PROD_fun_comp. rewrite PROD_fun_comp. 74 | refine (PROD_fun_eq_compat _ _). 75 | rewrite IHE. auto. unfold semWk in H. unfold hdMap. unfold tlMap. simpl. rewrite H. rewrite H. auto. 76 | Qed. 77 | 78 | 79 | Lemma SemCommutesWithMap : forall P (k:Kit P) sem, semWk k sem -> semVr k sem -> semVl k sem -> 80 | (forall E t (v : Value E t) E' (r : Map P E E'), 81 | SemVal v @_ SemMap sem r == SemVal (travVal v k r)) 82 | /\ (forall E t (exp : Exp E t) E' (r : Map P E E'), 83 | SemExp exp @_ SemMap sem r == SemExp (travExp exp k r)). 84 | Proof. 85 | intros P k sem semwk semvr semvl. unfold semWk in semwk. unfold semVr in semvr. unfold semVl in semvl. 86 | apply ExpVal_ind. 87 | 88 | (* TINT *) 89 | intros. simpl. rewrite <- K_com; trivial. 90 | 91 | (* TBOOL *) 92 | intros. simpl. rewrite <- K_com; trivial. 93 | 94 | (* TVAR *) 95 | intros E t var E' s. 96 | simpl. 97 | dependent induction var. simpl. rewrite SND_PROD_fun. 98 | destruct (consMapInv s) as [s' [v eq]]. subst. 99 | simpl. rewrite hdConsMap. rewrite semvl. trivial. 100 | destruct (consMapInv s) as [s' [v eq]]. subst. 101 | simpl. 102 | specialize (IHvar s' var (JMeq_refl _)). 103 | rewrite <- IHvar. rewrite fcont_comp_assoc. 104 | rewrite tlConsMap. rewrite FST_PROD_fun. trivial. 105 | 106 | (* TFIX *) 107 | intros E t t1 body IH E' s. 108 | specialize (IH (t :: t --> t1 :: E') (lift (K:=k) t (lift (K:=k) (t --> t1) s))). 109 | simpl SemVal. 110 | rewrite fcont_comp_assoc. rewrite Curry_comp. rewrite Curry_comp. rewrite <- IH. 111 | clear IH. simpl. rewrite PROD_map_fun. rewrite PROD_map_fun. rewrite ID_comp_left. rewrite ID_comp_left. 112 | rewrite PROD_fun_comp. rewrite SemLift2Map. unfold tlMap. unfold hdMap. simpl. 113 | rewrite semvr. rewrite semwk. rewrite semvr. auto. auto. 114 | 115 | (* TPAIR *) 116 | intros E t1 t2 v1 IH1 v2 IH2 E' s. 117 | simpl SemVal. rewrite <- IH1. rewrite <- IH2. rewrite <- PROD_fun_comp. trivial. 118 | 119 | (* TFST *) 120 | intros E t1 t2 v IH E' s. simpl SemExp. 121 | rewrite <- IH. repeat (rewrite fcont_comp_assoc). trivial. 122 | 123 | (* TSND *) 124 | intros E t1 t2 v IH E' s. simpl SemExp. 125 | rewrite <- IH. repeat (rewrite fcont_comp_assoc). trivial. 126 | 127 | (* TOP *) 128 | intros E n v1 IH1 v2 IH2 E' s. simpl SemExp. 129 | repeat (rewrite fcont_comp_assoc). rewrite <- IH1. rewrite <- IH2. 130 | rewrite PROD_fun_comp. trivial. 131 | 132 | (* TGT *) 133 | intros E v1 IH1 v2 IH2 E' s. simpl SemExp. 134 | repeat (rewrite fcont_comp_assoc). rewrite <- IH1. rewrite <- IH2. 135 | rewrite PROD_fun_comp. trivial. 136 | 137 | (* TVAL *) 138 | intros E t v IH E' s. simpl SemExp. 139 | rewrite <- IH. rewrite fcont_comp_assoc. trivial. 140 | 141 | (* TLET *) 142 | intros E t1 t2 e2 IH2 e1 IH1 E' s. simpl SemExp. 143 | rewrite fcont_comp_assoc. 144 | specialize (IH2 _ s). 145 | specialize (IH1 _ (lift (K:=k) _ s)). 146 | rewrite PROD_fun_comp. 147 | rewrite KLEISLIR_comp. 148 | rewrite <- IH2. clear IH2. rewrite <- IH1. clear IH1. simpl. 149 | rewrite PROD_map_fun. rewrite ID_comp_left. rewrite ID_comp_left. 150 | rewrite SemLiftMap. unfold hdMap. simpl. rewrite semvr. trivial. trivial. 151 | 152 | (* TAPP *) 153 | intros E t1 t2 v1 IH1 v2 IH2 E' s. simpl SemExp. 154 | rewrite fcont_comp_assoc. rewrite <- IH1. rewrite <- IH2. 155 | rewrite PROD_fun_comp. trivial. 156 | 157 | (* TIF *) 158 | intros E t ec IHc te1 IH1 te2 IH2 E' s. simpl SemExp. 159 | rewrite fcont_comp3_comp. rewrite <- IHc. rewrite <- IH1. rewrite <- IH2. trivial. 160 | Qed. 161 | 162 | Definition SemRenaming := SemMap SemVar. 163 | Definition SemSubst := SemMap SemVal. 164 | 165 | Corollary SemCommutesWithRenaming: 166 | (forall E t (v : Value E t) E' (r : Renaming E E'), 167 | SemVal v @_ SemRenaming r == SemVal (renameVal r v)) 168 | /\ (forall E t (exp : Exp E t) E' (r : Renaming E E'), 169 | SemExp exp @_ SemRenaming r == SemExp (renameExp r exp)). 170 | Proof. 171 | assert (H1 : semWk RenamingKit SemVar). unfold semWk; intros;auto. 172 | assert (H2 : semVr RenamingKit SemVar). unfold semVr; auto. 173 | assert (H3 : semVl RenamingKit SemVar). unfold semVl; auto. 174 | assert (H := SemCommutesWithMap H1 H2 H3). auto. 175 | Qed. 176 | 177 | (*========================================================================== 178 | Semantic function commutes with substitution 179 | ==========================================================================*) 180 | 181 | Lemma SemIdRenaming : forall E, SemRenaming (idRenaming E) == ID. 182 | Proof. 183 | induction E. 184 | simpl. 185 | apply (DOne_final (K DOne _) _). 186 | 187 | rewrite idRenamingDef. unfold SemRenaming. unfold SemMap. rewrite tlConsMap. fold SemMap. simpl. unfold shiftRenaming. rewrite SemShiftMap. 188 | rewrite IHE. rewrite ID_comp_left. 189 | apply fcont_eq_intro. intros. destruct x. auto. unfold semWk. intros. auto. 190 | Qed. 191 | 192 | 193 | 194 | Lemma SemRenamingShift : forall E t0, SemRenaming (fun s : Ty => SVAR (env:=E) t0) == FST. 195 | induction E. 196 | intros t0. apply fcont_eq_intro. intros x. 197 | case x. intros. simpl in *. dependent destruction x. unfold SemRenaming. unfold SemMap. rewrite K_simpl. case t. auto. 198 | 199 | (* Seems a bit round-about *) 200 | intros t0. 201 | assert ((fun s => SVAR t0) = tlMap (lift (K:=RenamingKit) t0 (idRenaming (a::E)))). 202 | rewrite LiftRenamingDef. rewrite tlConsMap. apply MapExtensional. auto. 203 | rewrite H. unfold SemRenaming. 204 | rewrite SemLiftMap. rewrite SemIdRenaming. rewrite ID_comp_left. auto. 205 | unfold semWk. intros. auto. 206 | Qed. 207 | 208 | Corollary SemCommutesWithSubst: 209 | (forall E t (v : Value E t) E' (s : Subst E E'), 210 | SemVal v @_ SemSubst s == SemVal (substVal s v)) 211 | /\ (forall E t (exp : Exp E t) E' (s : Subst E E'), 212 | SemExp exp @_ SemSubst s == SemExp (substExp s exp)). 213 | Proof. 214 | assert (H1 : semWk SubstKit SemVal). unfold semWk; intros. 215 | simpl. rewrite <- (proj1 SemCommutesWithRenaming). 216 | refine (fcont_comp_eq_compat _ _). trivial. rewrite SemRenamingShift. auto. 217 | assert (H2 : semVr SubstKit SemVal). unfold semVr; auto. 218 | assert (H3 : semVl SubstKit SemVal). unfold semVl; auto. 219 | assert (H := SemCommutesWithMap H1 H2 H3). unfold renameVal. unfold renameExp. 220 | split. intros. apply (proj1 H). 221 | intros. apply (proj2 H). 222 | Qed. 223 | 224 | -------------------------------------------------------------------------------- /src/unii.v: -------------------------------------------------------------------------------- 1 | (********************************************************************************** 2 | * unii.v * 3 | * Formalizing Domains, Ultrametric Spaces and Semantics of Programming Languages * 4 | * Nick Benton, Lars Birkedal, Andrew Kennedy and Carsten Varming * 5 | * Jan 2012 * 6 | * Build with Coq 8.3pl2 plus SSREFLECT * 7 | **********************************************************************************) 8 | 9 | (* Unityped lambda calculus, well-scoped by construction *) 10 | 11 | Require Export ssreflect ssrnat. 12 | Require Import Program. 13 | Require Import Fin. 14 | Set Implicit Arguments. 15 | Unset Strict Implicit. 16 | Import Prenex Implicits. 17 | 18 | Ltac Rewrites E := 19 | (intros; simpl; try rewrite E; 20 | repeat (match goal with | [H : context[eq _ _] |- _] => rewrite H end); 21 | auto). 22 | 23 | Definition Env := nat. 24 | 25 | Inductive Value E := 26 | | VAR: Fin E -> Value E 27 | | INT: nat -> Value E 28 | | LAMBDA: Exp E.+1 -> Value E 29 | with Exp E := 30 | | VAL: Value E -> Exp E 31 | | APP: Value E -> Value E -> Exp E 32 | | LET: Exp E -> Exp E.+1 -> Exp E 33 | | IFZ: Value E -> Exp E -> Exp E -> Exp E 34 | | OP: (nat -> nat -> nat) -> Value E -> Value E -> Exp E. 35 | 36 | Implicit Arguments INT [E]. 37 | 38 | Scheme Value_ind2 := Induction for Value Sort Prop 39 | with Exp_ind2 := Induction for Exp Sort Prop. 40 | Combined Scheme ExpValue_ind from Value_ind2, Exp_ind2. 41 | 42 | (*========================================================================== 43 | Variable-domain maps. 44 | By instantiating P with Var we get renamings. 45 | By instantiating P with Value we get substitutions. 46 | ==========================================================================*) 47 | Module Map. 48 | Section MAP. 49 | 50 | Variable P : Env -> Type. 51 | Definition Map E E' := FMap E (P E'). 52 | 53 | (*========================================================================== 54 | Package of operations used with a Map 55 | vr maps a Var into Var or Value (so is either the identity or TVAR) 56 | vl maps a Var or Value to a Value (so is either TVAR or the identity) 57 | wk weakens a Var or Value (so is either SVAR or renaming through SVAR on a value) 58 | ==========================================================================*) 59 | Record Ops := 60 | { 61 | vr : forall E, Fin E -> P E; 62 | vl : forall E, P E -> Value E; 63 | wk : forall E, P E -> P E.+1; 64 | wkvr : forall E (var : Fin E), wk (vr var) = vr (FinS var); 65 | vlvr : forall E (var : Fin E), vl (vr var) = VAR var 66 | }. 67 | Variable ops : Ops. 68 | 69 | Definition shift E E' (m : Map E E') : Map E E'.+1 := fun var => wk ops (m var). 70 | Definition id E : Map E E := fun (var : Fin E) => vr ops var. 71 | Definition lift E E' (m : Map E E') : Map E.+1 E'.+1 := cons (vr ops (FinZ _)) (shift m). 72 | 73 | Lemma shiftCons : forall E E' (m : Map E E') (x : P E'), shift (cons x m) = cons (wk ops x) (shift m). 74 | Proof. move => E E' m x. extFMap var; by []. Qed. 75 | 76 | Fixpoint mapVal E E' (m : Map E E') (v : Value E) : Value E' := 77 | match v with 78 | | VAR v => vl ops (m v) 79 | | INT i => INT i 80 | | LAMBDA e => LAMBDA (mapExp (lift m) e) 81 | end 82 | with mapExp E E' (m : Map E E') (e : Exp E) : Exp E' := 83 | match e with 84 | | VAL v => VAL (mapVal m v) 85 | | APP v0 v1 => APP (mapVal m v0) (mapVal m v1) 86 | | LET e0 e1 => LET (mapExp m e0) (mapExp (lift m) e1) 87 | | OP op v0 v1 => OP op (mapVal m v0) (mapVal m v1) 88 | | IFZ v e0 e1 => IFZ (mapVal m v) (mapExp m e0) (mapExp m e1) 89 | end. 90 | 91 | Variable E E' : Env. 92 | Variable m : Map E E'. 93 | Lemma mapVAR : forall (var : Fin _), mapVal m (VAR var) = vl ops (m var). by []. Qed. 94 | Lemma mapINT : forall n, mapVal m (INT n) = INT n. by []. Qed. 95 | Lemma mapLAMBDA : forall (e : Exp _), mapVal m (LAMBDA e) = LAMBDA (mapExp (lift m) e). by []. Qed. 96 | Lemma mapOP : forall op v1 v2, mapExp m (OP op v1 v2) = OP op (mapVal m v1) (mapVal m v2). by []. Qed. 97 | Lemma mapVAL : forall (v : Value _), mapExp m (VAL v) = VAL (mapVal m v). by []. Qed. 98 | Lemma mapLET : forall (e1 : Exp _) (e2 : Exp _), mapExp m (LET e1 e2) = LET (mapExp m e1) (mapExp (lift m) e2). by []. Qed. 99 | Lemma mapIFZ : forall v (e1 e2 : Exp _), mapExp m (IFZ v e1 e2) = IFZ (mapVal m v) (mapExp m e1) (mapExp m e2). by []. Qed. 100 | Lemma mapAPP : forall (v1 : Value _) v2, mapExp m (APP v1 v2) = APP (mapVal m v1) (mapVal m v2). by []. Qed. 101 | 102 | Lemma liftId : lift (@id E) = @id E.+1. 103 | Proof. extFMap var; [by [] | by apply wkvr]. 104 | Qed. 105 | 106 | Lemma idAsCons : @id E.+1 = cons (vr ops (FinZ _)) (shift (@id E)). 107 | Proof. extFMap var; first by []. unfold id, shift. simpl. by rewrite wkvr. Qed. 108 | 109 | End MAP. 110 | 111 | Hint Rewrite mapVAR mapINT mapLAMBDA mapOP mapVAL mapLET mapIFZ mapAPP : mapHints. 112 | Implicit Arguments id [P]. 113 | 114 | Lemma applyId P (ops:Ops P) E : 115 | (forall (v : Value E), mapVal ops (id ops E) v = v) 116 | /\ (forall (e : Exp E), mapExp ops (id ops E) e = e). 117 | Proof. move: E ; apply ExpValue_ind; intros; autorewrite with mapHints; Rewrites liftId. by apply vlvr. Qed. 118 | 119 | End Map. 120 | 121 | (*========================================================================== 122 | Variable renamings: Map Var 123 | ==========================================================================*) 124 | 125 | Definition Ren := Map.Map Fin. 126 | Definition RenOps : Map.Ops Fin. refine (@Map.Build_Ops _ (fun _ v => v) VAR FinS _ _). by []. by []. Defined. 127 | 128 | Definition renVal := Map.mapVal RenOps. 129 | Definition renExp := Map.mapExp RenOps. 130 | Definition liftRen := Map.lift RenOps. 131 | Definition shiftRen := Map.shift RenOps. 132 | Definition idRen := Map.id RenOps. 133 | 134 | (*========================================================================== 135 | Composition of renaming 136 | ==========================================================================*) 137 | 138 | Definition composeRen P E E' E'' (m : Map.Map P E' E'') (r : Ren E E') : Map.Map P E E'' := fun var => m (r var). 139 | 140 | Lemma liftComposeRen : forall E E' E'' P ops (m:Map.Map P E' E'') (r:Ren E E'), Map.lift ops (composeRen m r) = composeRen (Map.lift ops m) (liftRen r). 141 | Proof. move => E E' E'' P ops m r. extFMap var; by []. Qed. 142 | 143 | Lemma applyComposeRen E : 144 | (forall (v : Value E) E' E'' P ops (m:Map.Map P E' E'') (s : Ren E E'), 145 | Map.mapVal ops (composeRen m s) v = Map.mapVal ops m (renVal s v)) 146 | /\ (forall (e : Exp E) E' E'' P ops (m:Map.Map P E' E'') (s : Ren E E'), 147 | Map.mapExp ops (composeRen m s) e = Map.mapExp ops m (renExp s e)). 148 | Proof. move: E ; apply ExpValue_ind; intros; autorewrite with mapHints; Rewrites liftComposeRen. Qed. 149 | 150 | (*========================================================================== 151 | Substitution 152 | ==========================================================================*) 153 | 154 | Definition Sub := Map.Map Value. 155 | Definition SubOps : Map.Ops Value. refine (@Map.Build_Ops _ VAR (fun _ v => v) (fun E => renVal (fun v => FinS v)) _ _). by []. by []. Defined. 156 | 157 | Definition subVal := Map.mapVal SubOps. 158 | Definition subExp := Map.mapExp SubOps. 159 | Definition shiftSub := Map.shift SubOps. 160 | Definition liftSub := Map.lift SubOps. 161 | Definition idSub := Map.id SubOps. 162 | 163 | Ltac UnfoldRenSub := (unfold subVal; unfold subExp; unfold renVal; unfold renExp; unfold liftSub; unfold liftRen). 164 | Ltac FoldRenSub := (fold subVal; fold subExp; fold renVal; fold renExp; fold liftSub; fold liftRen). 165 | Ltac SimplMap := (UnfoldRenSub; autorewrite with mapHints; FoldRenSub). 166 | 167 | (*========================================================================== 168 | Composition of substitution followed by renaming 169 | ==========================================================================*) 170 | Definition composeRenSub E E' E'' (r : Ren E' E'') (s : Sub E E') : Sub E E'' := fun var => renVal r (s var). 171 | 172 | Lemma liftComposeRenSub : forall E E' E'' (r:Ren E' E'') (s:Sub E E'), liftSub (composeRenSub r s) = composeRenSub (liftRen r) (liftSub s). 173 | intros. extFMap var; first by []. 174 | unfold composeRenSub, liftSub. simpl. unfold renVal at 1 3. by do 2 rewrite <- (proj1 (applyComposeRen _)). 175 | Qed. 176 | 177 | Lemma applyComposeRenSub E : 178 | (forall (v:Value E) E' E'' (r:Ren E' E'') (s : Sub E E'), 179 | subVal (composeRenSub r s) v = renVal r (subVal s v)) 180 | /\ (forall (e:Exp E) E' E'' (r:Ren E' E'') (s : Sub E E'), 181 | subExp (composeRenSub r s) e = renExp r (subExp s e)). 182 | Proof. move: E ; apply ExpValue_ind; intros; SimplMap; Rewrites liftComposeRenSub. Qed. 183 | 184 | (*========================================================================== 185 | Composition of substitutions 186 | ==========================================================================*) 187 | 188 | Definition composeSub E E' E'' (s' : Sub E' E'') (s : Sub E E') : Sub E E'' := fun var => subVal s' (s var). 189 | 190 | Lemma liftComposeSub : forall E E' E'' (s' : Sub E' E'') (s : Sub E E'), liftSub (composeSub s' s) = composeSub (liftSub s') (liftSub s). 191 | intros. extFMap var; first by []. 192 | unfold composeSub. simpl. rewrite <- (proj1 (applyComposeRenSub _)). unfold composeRenSub, subVal. by rewrite <- (proj1 (applyComposeRen _)). 193 | Qed. 194 | 195 | Lemma applyComposeSub E : 196 | (forall (v : Value E) E' E'' (s' : Sub E' E'') (s : Sub E E'), 197 | subVal (composeSub s' s) v = subVal s' (subVal s v)) 198 | /\ (forall (e : Exp E) E' E'' (s' : Sub E' E'') (s : Sub E E'), 199 | subExp (composeSub s' s) e = subExp s' (subExp s e)). 200 | Proof. move: E ; apply ExpValue_ind; intros; SimplMap; Rewrites liftComposeSub. Qed. 201 | 202 | Lemma composeCons : forall E E' E'' (s':Sub E' E'') (s:Sub E E') (v:Value _), 203 | composeSub (cons v s') (liftSub s) = cons v (composeSub s' s). 204 | Proof. intros. extFMap var; first by []. 205 | unfold composeSub. unfold subVal. simpl. rewrite <- (proj1 (applyComposeRen _)). unfold composeRen. 206 | simpl. replace ((fun (var0:Fin E') => s' var0)) with s' by (by apply Extensionality). by []. 207 | Qed. 208 | 209 | Lemma composeSubIdLeft : forall E E' (s : Sub E E'), composeSub (@idSub _) s = s. 210 | Proof. intros. extFMap var; by apply (proj1 (Map.applyId _ _)). Qed. 211 | 212 | Lemma composeSubIdRight : forall E E' (s:Sub E E'), composeSub s (@idSub _) = s. 213 | Proof. intros. extFMap var; by []. Qed. 214 | 215 | Notation "[ x , .. , y ]" := (cons x .. (cons y (@Map.id _ SubOps _)) ..) : Sub_scope. 216 | Arguments Scope composeSub [_ _ _ Sub_scope Sub_scope]. 217 | Arguments Scope subExp [_ _ Sub_scope]. 218 | Arguments Scope subVal [_ _ Sub_scope]. 219 | Delimit Scope Sub_scope with sub. 220 | 221 | Lemma composeSingleSub : forall E E' (s:Sub E E') (v:Value _), composeSub [v] (liftSub s) = cons v s. 222 | Proof. intros. rewrite composeCons. by rewrite composeSubIdLeft. Qed. 223 | -------------------------------------------------------------------------------- /src/uniiop.v: -------------------------------------------------------------------------------- 1 | (********************************************************************************** 2 | * uniiop.v * 3 | * Formalizing Domains, Ultrametric Spaces and Semantics of Programming Languages * 4 | * Nick Benton, Lars Birkedal, Andrew Kennedy and Carsten Varming * 5 | * Jan 2012 * 6 | * Build with Coq 8.3pl2 plus SSREFLECT * 7 | **********************************************************************************) 8 | 9 | Require Export unii. 10 | 11 | Set Implicit Arguments. 12 | Unset Strict Implicit. 13 | Import Prenex Implicits. 14 | 15 | 16 | Reserved Notation "e '=>>' v" (at level 75). 17 | 18 | (*=Evaluation *) 19 | Inductive Evaluation : Exp O -> Value O -> Prop := 20 | | e_Val : forall v, VAL v =>> v 21 | | e_App : forall e1 v2 v, subExp [v2] e1 =>> v -> APP (LAMBDA e1) v2 =>> v 22 | | e_Let : forall e1 v1 e2 v2, e1 =>> v1 -> subExp [v1] e2 =>> v2 -> LET e1 e2 =>> v2 23 | | e_Ifz1 : forall e1 e2 v1, e1 =>> v1 -> IFZ (INT 0) e1 e2 =>> v1 24 | | e_Ifz2 : forall e1 e2 v2 n, e2 =>> v2 -> IFZ (INT (S n)) e1 e2 =>> v2 25 | | e_Op : forall op n1 n2, OP op (INT n1) (INT n2) =>> INT (op n1 n2) 26 | where "e =>> v" := (Evaluation e v). 27 | 28 | (*=End *) 29 | 30 | -------------------------------------------------------------------------------- /src/uniisem.v: -------------------------------------------------------------------------------- 1 | (********************************************************************************** 2 | * uniisem.v * 3 | * Formalizing Domains, Ultrametric Spaces and Semantics of Programming Languages * 4 | * Nick Benton, Lars Birkedal, Andrew Kennedy and Carsten Varming * 5 | * Jan 2012 * 6 | * Build with Coq 8.3pl2 plus SSREFLECT * 7 | **********************************************************************************) 8 | 9 | (* semantics of unityped lambda calculus in recursive domain *) 10 | 11 | Require Import PredomAll. 12 | Require Import Fin unii uniirec. 13 | 14 | Set Implicit Arguments. 15 | Unset Strict Implicit. 16 | Import Prenex Implicits. 17 | 18 | Export RD. 19 | 20 | (*=SemEnv *) 21 | Fixpoint SemEnv E : cpoType := match E with O => One | S E => SemEnv E * VInf end. 22 | Fixpoint SemVar E (v : Fin E) : SemEnv E =-> VInf := 23 | match v with 24 | | FinZ _ => pi2 25 | | FinS _ v => SemVar v << pi1 26 | end. 27 | (*=End *) 28 | 29 | Canonical Structure nat_cpoType := Eval hnf in discrete_cpoType nat. 30 | Canonical Structure bool_cpoType := Eval hnf in discrete_cpoType bool. 31 | 32 | Lemma zeroCase_mon : monotonic (fun (n:nat_cpoType) => match n with | O => @in1 _ (One:cpoType) _ tt | S m => @in2 _ _ (discrete_cpoType nat) m end). 33 | move => x y. case. move => e ; rewrite e. clear x e. by case: y. 34 | Qed. 35 | 36 | Definition zeroCasem : ordCatType (discrete_cpoType nat) ((One:cpoType) + nat_cpoType) := 37 | Eval hnf in mk_fmono zeroCase_mon. 38 | 39 | Lemma zeroCase_cont : continuous zeroCasem. 40 | move => c. simpl. have e:lub c = c 0. by []. rewrite e. simpl. 41 | by apply: (Ole_trans _ (le_lub _ O)). 42 | Qed. 43 | 44 | (*=zeroCase *) 45 | Definition zeroCase : nat_cpoType =-> (One:cpoType) + nat_cpoType := 46 | Eval hnf in mk_fcont zeroCase_cont. 47 | (*=End *) 48 | 49 | Lemma zeroCase_simplS: forall n, zeroCase (S n) = @in2 _ _ (discrete_cpoType nat) n. 50 | intros n ; auto. 51 | Qed. 52 | 53 | Lemma zeroCase_simplO: zeroCase O = @in1 _ (One:cpoType) _ tt. 54 | auto. 55 | Qed. 56 | 57 | Lemma SimpleB_mon (A B :Type) (C:ordType) (op : A -> B -> C:Type) : monotonic (fun p:discrete_ordType A * discrete_ordType B => op (fst p) (snd p)). 58 | case => x0 y0. case => x1 y1. case ; simpl. move => L L'. 59 | have e:x0 = x1 by []. have e':y0 = y1 by []. rewrite e. by rewrite e'. 60 | Qed. 61 | 62 | Definition SimpleBOpm (A B :Type) (C:ordType) (op : A -> B -> C:Type) : discrete_ordType A * discrete_ordType B =-> C := 63 | Eval hnf in mk_fmono (SimpleB_mon op). 64 | 65 | Lemma SimpleB_cont (A B:Type) (C:cpoType) (op : A -> B -> C:Type) : 66 | @continuous (discrete_cpoType A * discrete_cpoType B) C (SimpleBOpm op). 67 | move => c. simpl. apply: (Ole_trans _ (le_lub _ 0)). simpl. 68 | by []. 69 | Qed. 70 | 71 | Definition SimpleBOp (A B:Type) (C:cpoType) (op : A -> B -> C:Type) : discrete_cpoType A * discrete_cpoType B =-> C := 72 | Eval hnf in mk_fcont (SimpleB_cont op). 73 | 74 | (*=SemValExp *) 75 | Fixpoint SemVal E (v:Value E) : SemEnv E =-> VInf := 76 | match v return SemEnv E =-> VInf with 77 | | INT i => in1 << const _ i 78 | | VAR m => SemVar m 79 | | LAMBDA e => in2 << exp_fun (kleisli (eta << RD.Roll) << SemExp e << Id >< RD.Unroll) 80 | end with SemExp E (e:Exp E) : SemEnv E =-> RD.VInf _BOT := 81 | match e with 82 | | VAL v => eta << SemVal v 83 | | APP v1 v2 => kleisli (eta << RD.Unroll) << ev << 84 | <| [| @const _ (exp_cppoType _ _) PBot , Id|] << SemVal v1, RD.Roll << SemVal v2 |> 85 | | LET e1 e2 => ev << <| exp_fun (KLEISLIR (SemExp e2)), SemExp e1 |> 86 | | OP op v0 v1 => kleisli (eta << in1 << SimpleBOp op) << uncurry (Smash _ _) << 87 | <| [| eta, const _ PBot|] << SemVal v0, [| eta, const _ PBot|] << SemVal v1|> 88 | | IFZ v e1 e2 => ev << 89 | [| [| exp_fun (SemExp e1 << pi2), exp_fun (SemExp e2 << pi2)|] << zeroCase , 90 | @const _ (exp_cppoType _ _) PBot |] >< Id << <|SemVal v, Id|> 91 | end. 92 | (*=End *) 93 | 94 | Lemma Operator2_strictL A B C (f:A * B =-> C _BOT) d : Operator2 f PBot d =-= PBot. 95 | apply: Ole_antisym ; last by apply: leastP. 96 | unlock Operator2. simpl. by do 2 rewrite kleisli_bot. 97 | Qed. 98 | 99 | Lemma Operator2_strictR A B C (f:A * B =-> C _BOT) d : Operator2 f d PBot =-= PBot. 100 | apply: Ole_antisym ; last by apply: leastP. 101 | unlock Operator2. simpl. 102 | apply: (Ole_trans (proj2 (fmon_eq_elim (kleisli_comp2 _ _) d))). 103 | apply: DLless_cond. move => c X. 104 | case: (kleisliValVal X) => a [e P]. rewrite -> e. clear d e X. 105 | simpl in P. rewrite -> kleisli_bot in P. 106 | by case: (PBot_incon (proj2 P)). 107 | Qed. 108 | 109 | Add Parametric Morphism E (e:Exp E) : (SemExp e) 110 | with signature (@tset_eq _ : SemEnv E -> SemEnv E -> Prop) ==> (@tset_eq _ : VInf _BOT -> VInf _BOT -> Prop) 111 | as SemExp_eq_compat. 112 | intros e0 e1 eeq. by apply (fmon_stable (SemExp e)). 113 | Qed. 114 | 115 | Add Parametric Morphism E (v:Value E) : (SemVal v) 116 | with signature (@tset_eq _ : SemEnv E -> SemEnv E -> Prop) ==> (@tset_eq _ : VInf -> VInf -> Prop) 117 | as SemVal_eq_compat. 118 | intros e0 e1 eeq. by apply (fmon_stable (SemVal v)). 119 | Qed. 120 | 121 | 122 | -------------------------------------------------------------------------------- /src/uniisound.v: -------------------------------------------------------------------------------- 1 | (********************************************************************************** 2 | * uniisound.v * 3 | * Formalizing Domains, Ultrametric Spaces and Semantics of Programming Languages * 4 | * Nick Benton, Lars Birkedal, Andrew Kennedy and Carsten Varming * 5 | * Jan 2012 * 6 | * Build with Coq 8.3pl2 plus SSREFLECT * 7 | **********************************************************************************) 8 | 9 | (* soundness of semantics of unityped lambda calculus *) 10 | 11 | Require Import PredomAll. 12 | Require Import ssrnat. 13 | Require Import uniisem uniiop Fin. 14 | 15 | Set Implicit Arguments. 16 | Unset Strict Implicit. 17 | Import Prenex Implicits. 18 | 19 | Section SEMMAP. 20 | 21 | Variable P : Env -> Type. 22 | Variable ops : Map.Ops P. 23 | Variable Sem : forall E, P E -> SemEnv E =-> VInf. 24 | Variable SemVl : forall E (v : P E), Sem v =-= SemVal (Map.vl ops v). 25 | Variable SemVr : forall E, Sem (Map.vr ops (FinZ E)) =-= pi2. 26 | Variable SemWk : forall E (v : P E), Sem (Map.wk ops v) =-= Sem v << pi1. 27 | 28 | Fixpoint SemMap E E' : Map.Map P E' E -> SemEnv E =-> SemEnv E' := 29 | match E' with 30 | | O => fun m => terminal_morph (SemEnv E) 31 | | S _ => fun m => <| SemMap (tl m), Sem (hd m) |> 32 | end. 33 | 34 | Lemma SemShift : forall E E' (m : Map.Map P E E'), SemMap (Map.shift ops m) =-= SemMap m << pi1. 35 | Proof. elim. 36 | - move => E' m. by apply: terminal_unique. 37 | - move => E IH E' m. 38 | rewrite -> (consEta m) at 1. rewrite Map.shiftCons. simpl. rewrite -> (IH E' (tl m)). rewrite prod_fun_compl. unfold Map.shift, hd. by rewrite SemWk. 39 | Qed. 40 | 41 | Lemma SemLift : forall E E' (m : Map.Map P E E'), SemMap (Map.lift ops m) =-= SemMap m >< Id. 42 | Proof. move => E E' m. simpl. unfold tl, hd. simpl Map.lift. rewrite SemVr. rewrite SemShift. 43 | apply: prod_unique. rewrite prod_fun_fst. by rewrite prod_fun_fst. 44 | rewrite prod_fun_snd. rewrite prod_fun_snd. by rewrite comp_idL. 45 | Qed. 46 | 47 | Lemma SemId : forall E, SemMap (@Map.id P ops E) =-= Id. 48 | Proof. elim. 49 | - simpl. by apply: terminal_unique. 50 | - move => env IH. rewrite Map.idAsCons. simpl. rewrite tlCons hdCons. rewrite SemShift. rewrite IH. rewrite comp_idL. rewrite SemVr. 51 | apply: prod_unique. rewrite prod_fun_fst. by rewrite comp_idR. rewrite prod_fun_snd. by rewrite comp_idR. 52 | Qed. 53 | 54 | Lemma SemCommutesWithMap E : 55 | (forall (v : Value E) E' (r : Map.Map _ E E'), SemVal v << SemMap r =-= SemVal (Map.mapVal ops r v)) 56 | /\ (forall (e : Exp E) E' (r : Map.Map _ E E'), SemExp e << SemMap r =-= SemExp (Map.mapExp ops r e)). 57 | Proof. move: E ; apply ExpValue_ind. 58 | (* VAR *) 59 | - move => n v n' s. SimplMap. simpl. 60 | induction v. 61 | + rewrite (consEta s). simpl. rewrite prod_fun_snd hdCons. by rewrite SemVl. 62 | + rewrite (consEta s). simpl. rewrite <- IHv. rewrite <- comp_assoc. by rewrite prod_fun_fst tlCons. 63 | (* INT *) 64 | - move => n body n' s. SimplMap. simpl. rewrite <- comp_assoc. by rewrite <- const_com. 65 | (* LAMBDA *) 66 | - move => n body IH n' s. SimplMap. simpl SemVal. 67 | rewrite <- (IH _ (Map.lift _ s)). rewrite SemLift. do 4 rewrite <- comp_assoc. apply: comp_eq_compat ; first by []. 68 | apply: exp_unique. rewrite <- (comp_assoc pi1 (SemMap s)). rewrite <- prod_map_prod_fun. 69 | rewrite (comp_assoc _ _ ev). rewrite exp_com. do 2 rewrite <- comp_assoc. rewrite prod_map_prod_fun. rewrite comp_idL. 70 | rewrite prod_fun_compl. rewrite <- comp_assoc. rewrite prod_fun_fst. rewrite prod_fun_snd. by do 2 rewrite comp_idL. 71 | (* VAL *) 72 | - move => n v IH n' s. SimplMap. simpl. rewrite <- comp_assoc. by rewrite IH. 73 | (* APP *) 74 | - move => n v IH v' IH' n' s. SimplMap. simpl. rewrite <- IH. rewrite <- IH'. 75 | rewrite <- comp_assoc. rewrite prod_fun_compl. by repeat (rewrite <- comp_assoc). 76 | (* LET *) 77 | - move => n e IH e' IH' n' s. SimplMap. simpl. 78 | rewrite <- comp_assoc. rewrite prod_fun_compl. rewrite IH. rewrite exp_comp. rewrite <- IH'. rewrite SemLift. 79 | rewrite <- (comp_idL (SemExp (Map.mapExp ops s e))). rewrite <- (comp_idR (exp_fun _)). 80 | rewrite <- (prod_map_prod_fun (exp_fun (KLEISLIR (SemExp e') << SemMap s >< Id))). 81 | rewrite comp_assoc. rewrite exp_com. rewrite <- comp_assoc. rewrite prod_map_prod_fun. 82 | rewrite comp_idR. rewrite comp_idL. 83 | rewrite <- (comp_idR (exp_fun _)). rewrite <- (comp_idL (SemExp (Map.mapExp ops s e))). rewrite <- prod_map_prod_fun. 84 | rewrite comp_assoc. rewrite exp_com. rewrite KLEISLIR_comp. by do 2 rewrite comp_idL. 85 | (* IFZ *) 86 | - move => n v IH e0 IH0 e1 IH1 n' s. 87 | SimplMap. simpl. repeat rewrite <- comp_assoc. repeat rewrite prod_fun_compl. 88 | repeat rewrite comp_idL. do 2 rewrite prod_fun_snd. rewrite <- comp_assoc. rewrite prod_fun_fst. 89 | rewrite IH. clear IH. rewrite <- comp_assoc. rewrite prod_fun_fst. 90 | apply: fmon_eq_intro => x. simpl. case: ((SemVal (Map.mapVal _ s v)) x) ; simpl. 91 | + case. 92 | * simpl. do 2 rewrite SUM_fun_simplx. simpl. do 2 rewrite SUM_fun_simplx. apply: (fmon_eq_elim (IH0 _ s) x). 93 | * simpl => nn. do 2 rewrite SUM_fun_simplx. simpl. do 2 rewrite SUM_fun_simplx. apply (fmon_eq_elim (IH1 _ s) x). 94 | + simpl. move => m. do 2 rewrite SUM_fun_simplx. by split ; apply: leastP. 95 | (* OP *) 96 | - move => n op v0 IH0 v1 IH1 v s. 97 | SimplMap. simpl. 98 | by (repeat rewrite <- comp_assoc; repeat rewrite prod_fun_compl; repeat rewrite <- comp_assoc; rewrite IH1; rewrite IH0). 99 | Qed. 100 | 101 | End SEMMAP. 102 | 103 | Definition SemRen := SemMap SemVar. 104 | Definition SemSub := SemMap SemVal. 105 | 106 | Lemma SemCommutesWithRen E: 107 | (forall (v : Value E) E' (r : Ren E E'), SemVal v << SemRen r =-= SemVal (renVal r v)) 108 | /\ (forall (e : Exp E) E' (r : Ren E E'), SemExp e << SemRen r =-= SemExp (renExp r e)). 109 | Proof. by apply SemCommutesWithMap. Qed. 110 | 111 | Lemma SemShiftRen : forall E E' (r : Ren E E'), SemRen (shiftRen r) =-= SemRen r << pi1. 112 | Proof. by apply SemShift. Qed. 113 | 114 | Lemma SemIdRen : forall E, SemRen (@idRen E) =-= Id. 115 | Proof. by apply SemId. Qed. 116 | 117 | 118 | (*=Substitution *) 119 | Lemma SemCommutesWithSub E: 120 | (forall (v : Value E) E' (s : Sub E E'), SemVal v << SemSub s =-= SemVal (subVal s v)) 121 | /\ (forall (e : Exp E) E' (s : Sub E E'), SemExp e << SemSub s =-= SemExp (subExp s e)). 122 | (*=End *) 123 | Proof. move: E ; apply SemCommutesWithMap. 124 | + by []. + by []. + intros. simpl. rewrite <- (proj1 (SemCommutesWithRen E)). rewrite SemShiftRen. rewrite SemIdRen. by rewrite comp_idL. 125 | Qed. 126 | 127 | (*=Soundness *) 128 | Lemma Soundness e v : (e =>> v) -> SemExp e =-= eta << SemVal v. 129 | (*=End *) 130 | Proof. move => D. elim: e v / D. 131 | - by []. 132 | - move => e v v' D IH. simpl. 133 | repeat rewrite <- comp_assoc. rewrite <- (comp_idR ([|_,_|] << _)). rewrite <- (comp_idL (Roll << _)). 134 | rewrite <- prod_map_prod_fun. 135 | rewrite (comp_assoc _ _ ev). rewrite (comp_assoc _ in2). rewrite sum_fun_snd. 136 | rewrite (comp_idL (exp_fun _)). rewrite exp_com. repeat rewrite <- comp_assoc. 137 | rewrite comp_assoc. rewrite <- kleisli_comp2. rewrite <- comp_assoc. 138 | rewrite UR_id. rewrite comp_idR. rewrite kleisli_unit. rewrite comp_idL. 139 | rewrite prod_map_prod_fun. rewrite comp_idR. rewrite comp_assoc. rewrite UR_id. 140 | rewrite comp_idL. 141 | rewrite <- (proj2 (SemCommutesWithSub _) e _ ([v] %sub)) in IH. 142 | rewrite <- IH. simpl. rewrite (terminal_unique Id _). by []. 143 | - move => e0 v0 e1 v1 D0 IH0 D1 IH1. 144 | simpl. rewrite IH0. clear IH0 D0. 145 | rewrite <- (comp_idL (eta << _)). rewrite <- (comp_idR (exp_fun _)). rewrite <- prod_map_prod_fun. 146 | rewrite comp_assoc. rewrite exp_com. rewrite KLEISLIR_unit. 147 | rewrite <- (proj2 (SemCommutesWithSub 1%N) e1 _ ([ v0]%sub)) in IH1. simpl in IH1. 148 | rewrite <- IH1. rewrite (terminal_unique Id (terminal_morph One)). by []. 149 | - move => e0 e1 v0 _ IH. simpl. repeat rewrite <- comp_assoc. rewrite prod_map_prod_fun. 150 | rewrite comp_assoc. rewrite sum_fun_fst. rewrite <- IH. clear IH. 151 | rewrite <- comp_assoc. 152 | have X:(zeroCase << @const One (discrete_cpoType nat) 0%N) =-= in1 by apply: fmon_eq_intro ; case. 153 | rewrite X. clear X. rewrite sum_fun_fst. rewrite <- (comp_idR (exp_fun _)). rewrite <- prod_map_prod_fun. 154 | rewrite comp_assoc. rewrite exp_com. rewrite <- comp_assoc. rewrite prod_fun_snd. by rewrite comp_idR. 155 | - move => e0 e1 v0 n _ IH. simpl. repeat rewrite <- comp_assoc. rewrite prod_map_prod_fun. 156 | rewrite comp_assoc. rewrite sum_fun_fst. rewrite <- IH. clear IH. 157 | rewrite <- comp_assoc. 158 | have X:(zeroCase << @const One (discrete_cpoType nat) n.+1%N) =-= in2 << @const One (discrete_cpoType nat) n by apply: fmon_eq_intro ; case. 159 | rewrite X. clear X. rewrite comp_assoc. rewrite sum_fun_snd. 160 | rewrite <- prod_map_prod_fun. rewrite comp_assoc. rewrite exp_com. rewrite <- comp_assoc. rewrite prod_fun_snd. 161 | by rewrite comp_idR. 162 | - move => op n n'. simpl. repeat rewrite comp_assoc. rewrite sum_fun_fst. simpl. 163 | rewrite <- (prod_map_prod_fun eta). do 2 rewrite <- comp_assoc. 164 | rewrite (comp_assoc _ _ ev). rewrite - {1} (comp_idL pi2). 165 | rewrite (comp_assoc _ (eta >< eta)). 166 | have e:((ev << Smash (discrete_cpoType nat) (discrete_cpoType nat) >< Id) << 167 | eta >< eta) =-= eta by apply: fmon_eq_intro => a; simpl ; unfold Smash ; simpl ; rewrite Operator2_simpl ; case: a. 168 | rewrite -> e. clear e. rewrite comp_assoc. rewrite kleisli_eta_com. 169 | rewrite <- comp_assoc. by apply:fmon_eq_intro. 170 | Qed. 171 | 172 | -------------------------------------------------------------------------------- /src/untypeddensem.v: -------------------------------------------------------------------------------- 1 | Require Import Predom. 2 | Require Import PredomAux. 3 | Require Export lc. 4 | Require Import PredomSum. 5 | Require Export untypedlambda. 6 | Require Import PredomCore. 7 | Require Import PredomLift. 8 | Require Import PredomKleisli. 9 | Require Import PredomProd. 10 | 11 | Set Implicit Arguments. 12 | Unset Strict Implicit. 13 | 14 | 15 | Definition SimpleOp2mm (A B C : Type) (op : A -> B -> C) : Discrete A -> Discrete B -m> Discrete C. 16 | intros. 17 | exists (op X). 18 | unfold monotonic. 19 | intros. 20 | simpl in *. 21 | rewrite H. auto. 22 | Defined. 23 | 24 | Definition SimpleOp2c A B C (op:A -> B -> C) : 25 | Discrete A -> Discrete B -c> Discrete C. 26 | intros. 27 | exists (SimpleOp2mm op X). 28 | simpl in *. 29 | unfold continuous. 30 | intros. 31 | simpl. auto. 32 | Defined. 33 | 34 | Definition SimpleOp2m (A B C:Type) (op : A -> B -> C) : 35 | Discrete A -m> Discrete B -C-> Discrete C. 36 | intros. 37 | exists (SimpleOp2c op). 38 | unfold monotonic. 39 | intros. 40 | simpl in *. 41 | rewrite H. auto. 42 | Defined. 43 | 44 | Definition SimpleOp2 A B C (op:A -> B -> C) : 45 | Discrete A -c> Discrete B -C-> Discrete C. 46 | intros. 47 | exists (SimpleOp2m op). 48 | unfold continuous ; intros ; simpl ; auto. 49 | Defined. 50 | 51 | Definition FS := BiLift_strict (BiSum (BiConst (Discrete nat)) BiArrow). 52 | Definition DInf := DInf FS. 53 | Definition VInf := Dsum (Discrete nat) (DInf -C-> DInf). 54 | 55 | Definition Roll : DL VInf -C-> DInf := ROLL FS. 56 | Definition Unroll : DInf -C-> DL VInf := UNROLL FS. 57 | 58 | Lemma UR_eq : forall x, Unroll (Roll x) == x. 59 | intros y. rewrite <- fcont_comp_simpl. 60 | apply Oeq_trans with (y:= ID y). 61 | refine (app_eq_compat _ (Oeq_refl _)). unfold Unroll. unfold Roll. apply (DIso_ur FS). auto. 62 | Qed. 63 | 64 | Lemma RU_eq : forall x, Roll (Unroll x) == x. 65 | intros x. 66 | assert (xx := fcont_eq_elim (DIso_ru FS) x). rewrite fcont_comp_simpl in xx. rewrite ID_simpl in xx. 67 | simpl in xx. unfold Unroll. unfold Roll. apply xx. 68 | Qed. 69 | 70 | Definition monic D E (f:D -c> E) := forall x y, f x <= f y -> x <= y. 71 | 72 | Lemma Roll_monic: monic Roll. 73 | unfold monic. 74 | assert (zz:=@ROLL_monic _ FS). fold Roll in zz. fold DInf in zz. unfold VInf. 75 | apply zz. 76 | Qed. 77 | 78 | Lemma Unroll_monic: monic Unroll. 79 | assert (yy:=@UNROLL_monic _ FS). fold Unroll in yy. fold DInf in yy. 80 | unfold monic. apply yy. 81 | Qed. 82 | 83 | Lemma inj_monic: forall D E (f:D -c> E), monic f -> forall x y, f x == f y -> x == y. 84 | intros D E f mm. unfold monic in mm. intros x y. 85 | assert (xx:= mm x y). specialize (mm y x). 86 | intros [xy1 xy2]. split ; auto. 87 | Qed. 88 | 89 | Fixpoint SemEnv (env : Env) := 90 | match env with 91 | | O => DOne 92 | | S env' => SemEnv env' *c* VInf 93 | end. 94 | 95 | (* Lookup in an environment *) 96 | Fixpoint SemVar env (var : Var env) : SemEnv env -c> VInf := 97 | match var with 98 | | ZVAR _ => SND 99 | | SVAR _ v => SemVar v << FST 100 | end. 101 | 102 | Lemma K_com2: forall E F (f:E -C-> F), 103 | EV << PROD_fun (@K E _ f) (ID) == f. 104 | intros. apply fcont_eq_intro. auto. 105 | Qed. 106 | 107 | Lemma uncurry_PROD_fun: forall D D' E F G f g1 (g2:G -C-> D') h, 108 | @uncurry D E F f << PROD_fun (g1 << g2) h == @uncurry D' E F (f << g1) << PROD_fun g2 h. 109 | intros. apply fcont_eq_intro. auto. 110 | Qed. 111 | 112 | Add Parametric Morphism D E F : (@curry D E F) 113 | with signature (@Oeq (Dprod D E -C-> F)) ==> (@Oeq (D -C-> E -C-> F)) 114 | as Curry_eq_compat. 115 | intros f0 f1 feq. refine (fcont_eq_intro _). intros d. refine (fcont_eq_intro _). intros e. 116 | repeat (rewrite curry_simpl). auto. 117 | Qed. 118 | 119 | Definition Dlift : (VInf -C-> DInf) -C-> DInf -C-> DInf := curry (Roll << EV << 120 | PROD_map (kleisli_cc _ _ << (fcont_COMP VInf DInf _ Unroll)) Unroll). 121 | 122 | Definition Dapp := EV << PROD_map (SUM_fun (PBot : Discrete nat -C-> DInf -C-> DL VInf) 123 | (fcont_COMP _ _ _ Unroll)) 124 | (Roll << eta) : Dprod VInf VInf -C-> DL VInf. 125 | 126 | Definition zeroCasem : Discrete nat -m> Dsum DOne (Discrete nat). 127 | exists (fun (n:Discrete nat) => match n with | O => INL DOne _ tt | S m => INR _ (Discrete nat) m end). 128 | unfold monotonic. intros x y xy. assert (x = y) as E by auto. 129 | rewrite E in *. auto. 130 | Defined. 131 | 132 | Definition zeroCase : Discrete nat -c> Dsum DOne (Discrete nat). 133 | exists zeroCasem. unfold continuous. intros c. 134 | refine (Ole_trans _ (le_lub _ 0)). auto. 135 | Defined. 136 | 137 | Lemma zeroCase_simplS: forall n, zeroCase (S n) = INR _ (Discrete nat) n. 138 | intros n ; auto. 139 | Qed. 140 | 141 | Lemma zeroCase_simplO: zeroCase O = INL DOne _ tt. 142 | auto. 143 | Qed. 144 | 145 | Fixpoint SemVal env (v : Value env) : SemEnv env -C-> VInf := 146 | match v with 147 | | NUM _ n => INL _ _ << (@K _ (Discrete nat) n) 148 | | VAR _ var => SemVar var 149 | | LAMBDA _ e => INR _ _ << Dlift << curry (Roll << SemExp e) 150 | end 151 | with SemExp env (e : Exp env) : SemEnv env -C-> VInf _BOT := 152 | match e with 153 | | APP _ v1 v2 => Dapp << <| SemVal v1 , SemVal v2 |> 154 | | VAL _ v => eta << (SemVal v) 155 | | LET _ e1 e2 => EV << PROD_fun (curry (KLEISLIR (SemExp e2))) (SemExp e1) 156 | | OP _ op v1 v2 => 157 | uncurry _ _ _ (Operator2 (eta << INL _ _ << uncurry _ _ _ (SimpleOp2 op))) << 158 | PROD_fun (SUM_fun eta (PBot : (DInf -C-> DInf) -C-> DL (Discrete nat)) << SemVal v1) 159 | (SUM_fun eta (PBot : (DInf -C-> DInf) -C-> DL (Discrete nat)) << SemVal v2) 160 | | IFZ _ v e1 e2 => EV << 161 | PROD_fun (SUM_fun (SUM_fun (K _ (SemExp e1)) (K _ (SemExp e2)) << zeroCase) 162 | (PBot : (DInf -C-> DInf) -C-> SemEnv _ -C-> DL VInf ) << (SemVal v)) ID 163 | end. 164 | 165 | Lemma EV_simpl: forall D E (f : D -C-> E) d, EV (f,d) = f d. 166 | intros ; auto. 167 | Qed. 168 | 169 | 170 | Add Parametric Morphism env (e : Exp env) : (SemExp e) 171 | with signature (@Oeq (SemEnv env)) ==> (@Oeq (DL VInf)) 172 | as SemExp_eq_compat. 173 | intros e0 e1 eeq. 174 | destruct eeq. split ; auto. 175 | Qed. 176 | 177 | Add Parametric Morphism env (v : Value env) : (SemVal v) 178 | with signature (@Oeq (SemEnv env)) ==> (@Oeq VInf) 179 | as SemVal_eq_compat. 180 | intros e0 e1 eeq. 181 | destruct eeq. split ; auto. 182 | Qed. 183 | 184 | -------------------------------------------------------------------------------- /src/untypedopsem.v: -------------------------------------------------------------------------------- 1 | (* untypedopsem.v 2 | Standard operational semantics 3 | *) 4 | Require Import utility. 5 | Require Import Coq.Program.Equality. 6 | Require Import EqdepFacts. 7 | Require Import untypedlambda. 8 | 9 | Set Implicit Arguments. 10 | Unset Strict Implicit. 11 | 12 | (*========================================================================== 13 | Evaluation relation 14 | ==========================================================================*) 15 | 16 | (** printing =>> %\ensuremath{\Downarrow}% *) 17 | (** printing n1 %\ensuremath{n_1}% *) 18 | (** printing n2 %\ensuremath{n_2}% *) 19 | (** printing v1 %\ensuremath{v_1}% *) 20 | (** printing v2 %\ensuremath{v_2}% *) 21 | (** printing e1 %\ensuremath{e_1}% *) 22 | (** printing e2 %\ensuremath{e_2}% *) 23 | Reserved Notation "e =>> v" (at level 70, no associativity). 24 | 25 | Open Scope Subst_scope. 26 | Inductive Ev : CExp -> CValue -> Prop := 27 | | e_Val : forall v, VAL v =>> v 28 | | e_App : forall e1 v2 v, substExp (consMap v2 (idSubst _)) e1 =>> v -> APP (LAMBDA e1) v2 =>> v 29 | | e_Let : forall e1 v1 e2 v2, e1 =>> v1 -> substExp (consMap v1 (idSubst _)) e2 =>> v2 -> (LET e1 e2) =>> v2 30 | | e_IfzTrue : forall e1 e2 v1, e1 =>> v1 -> IFZ (NUM 0) e1 e2 =>> v1 31 | | e_IfzFalse : forall e1 e2 v2 n, e2 =>> v2 -> IFZ (NUM (S n)) e1 e2 =>> v2 32 | | e_Op : forall op n1 n2, OP op (NUM n1) (NUM n2) =>> NUM (op n1 n2) 33 | where "e =>> v" := (Ev e v). 34 | -------------------------------------------------------------------------------- /src/untypedsubst.v: -------------------------------------------------------------------------------- 1 | Require Import utility. 2 | Require Import PredomAux. 3 | Require Import PredomSum. 4 | Require Import untypedlambda. 5 | Require Import untypeddensem. 6 | 7 | Set Implicit Arguments. 8 | Unset Strict Implicit. 9 | 10 | (* Need this for nice pretty-printing *) 11 | Unset Printing Implicit Defensive. 12 | 13 | (*========================================================================== 14 | Semantics of substitution and renaming 15 | ==========================================================================*) 16 | 17 | (* Apply semantic function pointwise to a renaming or substitution *) 18 | Fixpoint SemSubst env env' : Subst env' env -> SemEnv env -c> SemEnv env' := 19 | match env' with 20 | | O => fun s => K _ (tt : DOne) 21 | | S _ => fun s => <| SemSubst (tlMap s) , SemVal (hdMap s) |> 22 | end. 23 | 24 | Fixpoint SemRenaming env env' : Renaming env' env -> SemEnv env -c> SemEnv env' := 25 | match env' with 26 | | O => fun s => K _ (tt : DOne) 27 | | S _ => fun s => <| SemRenaming (tlMap s) , SemVar (hdMap s) |> 28 | end. 29 | 30 | Lemma SemLiftRenaming : 31 | forall env env' (r : Renaming env env'), SemRenaming (tlMap (liftRenaming r)) == SemRenaming r << FST. 32 | Proof. 33 | induction env. intros. simpl. 34 | apply (DOne_final (K (Dprod (SemEnv env') VInf) (tt:DOne)) _). 35 | 36 | intros. 37 | simpl. destruct (consMapInv r) as [r' [var eq]]. 38 | subst. simpl. rewrite hdConsMap. rewrite tlConsMap. 39 | specialize (IHenv env' r'). 40 | rewrite PROD_fun_comp. rewrite IHenv. auto. 41 | Qed. 42 | 43 | (*========================================================================== 44 | Semantic function commutes with renaming 45 | ==========================================================================*) 46 | 47 | Implicit Arguments uncurry. 48 | Add Parametric Morphism (D E F:cpo) : (@SUM_fun D E F) 49 | with signature (@Oeq (D -c> F)) ==> (@Oeq (E -c> F)) ==> (@Oeq (Dsum D E -c> F)) 50 | as SUM_fun_eq_compat. 51 | intros f0 f1 feq g0 g1 geq. 52 | refine (fcont_eq_intro _). 53 | intro d. case d. intro dl. rewrite SUM_fun_simpl. rewrite SUM_fun_simpl. simpl. auto. 54 | intro dr. rewrite SUM_fun_simpl. rewrite SUM_fun_simpl. simpl. auto. 55 | Qed. 56 | 57 | 58 | Lemma SemCommutesWithRenaming: 59 | (forall env (v : Value env) env' (r : Renaming env env'), 60 | SemVal v << SemRenaming r == SemVal (renameVal r v)) 61 | /\ (forall env (exp : Exp env) env' (r : Renaming env env'), 62 | SemExp exp << SemRenaming r == SemExp (renameExp r exp)). 63 | Proof. 64 | apply ExpValue_ind. 65 | 66 | (* VAR *) 67 | intros env var env' r. 68 | simpl. 69 | induction var. 70 | simpl. rewrite SND_PROD_fun. 71 | destruct (consMapInv r) as [r' [v eq]]. subst. 72 | simpl. rewrite hdConsMap. trivial. 73 | destruct (consMapInv r) as [r' [v eq]]. subst. 74 | simpl. 75 | specialize (IHvar r'). 76 | rewrite <- IHvar. rewrite fcont_comp_assoc. 77 | rewrite tlConsMap. rewrite FST_PROD_fun. trivial. 78 | 79 | (* TNUM *) 80 | intros. simpl. rewrite fcont_comp_assoc. rewrite <- K_com; trivial. 81 | 82 | 83 | (* LAMBDA *) 84 | intros E body IH E' s. 85 | specialize (IH _ (liftRenaming s)). 86 | rewrite renameLAMBDA. 87 | simpl SemVal. 88 | rewrite fcont_comp_assoc. rewrite fcont_comp_assoc. rewrite fcont_comp_assoc. rewrite fcont_comp_assoc. 89 | rewrite Curry_comp. rewrite fcont_comp_assoc. 90 | rewrite <- IH. rewrite PROD_map_fun. rewrite ID_comp_left. rewrite <- SemLiftRenaming. auto. 91 | 92 | (* OP *) 93 | intros env n v1 IH1 v2 IH2 env' s. rewrite renameOP. simpl. 94 | repeat (rewrite fcont_comp_assoc). rewrite <- IH1. rewrite <- IH2. 95 | Implicit Arguments INR [D1 D2]. Implicit Arguments INL [D1 D2]. Implicit Arguments uncurry [D E F]. 96 | rewrite PROD_fun_compr. rewrite PROD_fun_compr. 97 | refine (fcont_comp_eq_compat _ _). 98 | admit. 99 | 100 | (* IFZ *) 101 | intros env v IH e1 IH1 e2 IH2 env' s. rewrite renameIFZ. simpl. 102 | repeat (rewrite fcont_comp_assoc). rewrite PROD_fun_compl. 103 | rewrite (PROD_fun_eq_compat (fcont_comp_assoc _ _ _) (Oeq_refl _)). 104 | rewrite (PROD_fun_eq_compat (fcont_comp_eq_compat (Oeq_refl _) (IH _ _) ) (Oeq_refl _)). 105 | rewrite (PROD_fun_eq_compat (Oeq_refl _) (fcont_comp_unitL _)). 106 | rewrite <- (PROD_fun_eq_compat (Oeq_refl _) (fcont_comp_unitR _)). 107 | 108 | Lemma xx D E F G (f:E -C-> F -C-> G) (h:D -C-> F) (g: E -C-> D) : 109 | EV << PROD_fun f (h << g) == 110 | EV << PROD_fun (curry (uncurry _ _ _ f << SWAP << PROD_map h ID << SWAP)) g. 111 | intros ; apply fcont_eq_intro ; auto. 112 | Qed. 113 | 114 | rewrite xx. refine (fcont_comp_eq_compat (Oeq_refl _) _). 115 | refine (PROD_fun_eq_compat _ (Oeq_refl _)). apply Oeq_sym. 116 | refine (curry_unique _). rewrite fcont_comp_assoc. 117 | apply fcont_eq_intro. 118 | intros d. repeat (rewrite fcont_comp_simpl). case d ; clear d. intros ds dn. 119 | repeat (rewrite PROD_map_simpl). rewrite ID_simpl. simpl. repeat (rewrite ID_simpl). simpl. 120 | repeat (rewrite EV_simpl). repeat (rewrite fcont_comp_simpl). 121 | rewrite SUM_fun_simpl. 122 | unfold VInf. 123 | admit. 124 | 125 | (* VAL *) 126 | intros env v IH env' s. rewrite renameVAL. simpl. rewrite fcont_comp_assoc. rewrite IH. auto. 127 | 128 | (* LET *) 129 | intros env e2 IH2 e1 IH1 env' s. rewrite renameLET. simpl. 130 | repeat (rewrite fcont_comp_assoc). 131 | rewrite <- IH2. clear IH2. 132 | rewrite <- IH1. clear IH1. 133 | refine (fcont_comp_eq_compat _ _). trivial. simpl. 134 | rewrite SemLiftRenaming. 135 | rewrite PROD_fun_comp. 136 | refine (PROD_fun_eq_compat _ _). 137 | (*simpl. 138 | rewrite Curry_comp. simpl. *) 139 | rewrite PROD_map_fun. 140 | rewrite KLEISLIR_comp. 141 | rewrite ID_comp_left. 142 | admit. trivial. 143 | 144 | 145 | (* APP *) 146 | intros env v1 IH1 v2 IH2 env' s. rewrite renameAPP. simpl. 147 | rewrite fcont_comp_assoc. rewrite <- IH1. rewrite <- IH2. 148 | rewrite PROD_fun_comp. trivial. 149 | Qed. 150 | 151 | Set Printing Implicit Defensive. 152 | 153 | Lemma SemShiftRenaming : 154 | forall env env' (r : Renaming env env'), SemRenaming (shiftRenaming r) == SemRenaming r << FST. 155 | Proof. 156 | induction env. intros. simpl. 157 | apply (DOne_final (K (Dprod (SemEnv env') VInf) (tt:DOne)) _). 158 | 159 | intros. 160 | simpl. destruct (consMapInv r) as [r' [var eq]]. 161 | subst. simpl. rewrite hdConsMap. rewrite tlConsMap. 162 | specialize (IHenv env' r'). 163 | rewrite PROD_fun_comp. rewrite IHenv. auto. 164 | Qed. 165 | 166 | Lemma SemIdRenaming : forall env, SemRenaming (idRenaming env) == ID. 167 | Proof. 168 | induction env. 169 | simpl. 170 | apply (DOne_final (K DOne _) _). 171 | 172 | rewrite idRenamingDef. simpl SemRenaming. rewrite tlConsMap. rewrite SemShiftRenaming. 173 | rewrite IHenv. rewrite ID_comp_left. 174 | apply fcont_eq_intro. intros. destruct x. auto. 175 | Qed. 176 | 177 | 178 | Lemma SemShiftSubst : 179 | forall env env' (s : Subst env env'), SemSubst (shiftSubst s) == SemSubst s << FST. 180 | Proof. 181 | induction env. intros. simpl. 182 | apply (DOne_final (K (Dprod (SemEnv env') VInf) (tt:DOne)) _). 183 | 184 | intros. 185 | simpl. destruct (consMapInv s) as [s' [var eq]]. 186 | subst. simpl. rewrite hdConsMap. rewrite tlConsMap. 187 | specialize (IHenv env' s'). 188 | rewrite PROD_fun_comp. 189 | refine (PROD_fun_eq_compat _ _). 190 | rewrite IHenv. auto. unfold shiftSubst. 191 | rewrite shiftConsMap. rewrite hdConsMap. 192 | simpl. rewrite <- (proj1 SemCommutesWithRenaming). 193 | refine (fcont_comp_eq_compat _ _). trivial. 194 | (* Seems a bit round-about *) 195 | assert ((fun t0 => SVAR t0) = tlMap (liftRenaming (idRenaming env'))). 196 | rewrite LiftRenamingDef. rewrite tlConsMap. apply MapExtensional. auto. 197 | rewrite H. 198 | rewrite SemLiftRenaming. rewrite SemIdRenaming. rewrite ID_comp_left. auto. 199 | Qed. 200 | 201 | 202 | Lemma SemLiftSubst : 203 | forall env env' (s : Subst env env'), SemSubst (tlMap (liftSubst s)) == SemSubst s << FST. 204 | Proof. 205 | intros. 206 | rewrite LiftSubstDef. rewrite tlConsMap. apply SemShiftSubst. 207 | Qed. 208 | 209 | (*========================================================================== 210 | Semantic function commutes with substitution 211 | ==========================================================================*) 212 | 213 | 214 | Lemma SemCommutesWithSubst: 215 | (forall env (v : Value env) env' (s : Subst env env'), 216 | SemVal v << SemSubst s == SemVal (substVal s v)) 217 | /\ (forall env (e : Exp env) env' (s : Subst env env'), 218 | SemExp e << SemSubst s == SemExp (substExp s e)). 219 | Proof. 220 | apply ExpVal_ind. 221 | 222 | (* TINT *) 223 | intros. simpl. rewrite <- K_com; trivial. 224 | 225 | (* TBOOL *) 226 | intros. simpl. rewrite <- K_com; trivial. 227 | 228 | (* TVAR *) 229 | intros env ty var env' s. 230 | simpl. 231 | unfold substVal. simpl travVal. 232 | induction var. simpl. rewrite SND_PROD_fun. 233 | destruct (consMapInv s) as [s' [v eq]]. subst. 234 | simpl. rewrite hdConsMap. trivial. 235 | destruct (consMapInv s) as [s' [v eq]]. subst. 236 | simpl. 237 | specialize (IHvar s'). 238 | rewrite <- IHvar. rewrite fcont_comp_assoc. 239 | rewrite tlConsMap. rewrite FST_PROD_fun. trivial. 240 | 241 | (* TFIX *) 242 | intros E t t1 body IH E' s. rewrite substTFIX. simpl. 243 | rewrite fcont_comp_assoc. 244 | rewrite Curry_comp. rewrite Curry_comp. rewrite <- IH. 245 | clear IH. simpl. rewrite PROD_map_fun. rewrite PROD_map_fun. rewrite ID_comp_left. rewrite ID_comp_left. 246 | rewrite SemLift2Subst. rewrite fcont_comp_assoc. rewrite PROD_fun_comp. rewrite fcont_comp_assoc. trivial. 247 | 248 | (* TPAIR *) 249 | intros E t1 t2 v1 IH1 v2 IH2 E' s. rewrite substTPAIR. simpl. 250 | rewrite <- IH1. rewrite <- IH2. rewrite PROD_fun_comp. trivial. 251 | 252 | (* TFST *) 253 | intros E t1 t2 v IH E' s. rewrite substTFST. simpl. 254 | rewrite <- IH. repeat (rewrite fcont_comp_assoc). trivial. 255 | 256 | (* TSND *) 257 | intros E t1 t2 v IH E' s. rewrite substTSND. simpl. 258 | rewrite <- IH. repeat (rewrite fcont_comp_assoc). trivial. 259 | 260 | (* TOP *) 261 | intros E n v1 IH1 v2 IH2 E' s. rewrite substTOP. simpl. 262 | repeat (rewrite fcont_comp_assoc). rewrite <- IH1. rewrite <- IH2. 263 | rewrite PROD_fun_comp. trivial. 264 | 265 | (* TGT *) 266 | intros E v1 IH1 v2 IH2 E' s. rewrite substTGT. simpl. 267 | repeat (rewrite fcont_comp_assoc). rewrite <- IH1. rewrite <- IH2. 268 | rewrite PROD_fun_comp. trivial. 269 | 270 | (* TVAL *) 271 | intros E t v IH E' s. rewrite substTVAL. simpl. 272 | rewrite <- IH. rewrite fcont_comp_assoc. trivial. 273 | 274 | (* TLET *) 275 | intros E t1 t2 e2 IH2 e1 IH1 E' s. rewrite substTLET. simpl. 276 | rewrite fcont_comp_assoc. 277 | specialize (IH2 _ s). 278 | specialize (IH1 _ (liftSubst _ s)). 279 | rewrite PROD_fun_comp. 280 | rewrite KLEISLIR_comp. 281 | rewrite <- IH2. clear IH2. rewrite <- IH1. clear IH1. simpl. 282 | rewrite PROD_map_fun. rewrite ID_comp_left. rewrite ID_comp_left. rewrite SemLiftSubst. trivial. 283 | 284 | (* TAPP *) 285 | intros E t1 t2 v1 IH1 v2 IH2 E' s. rewrite substTAPP. simpl. 286 | rewrite fcont_comp_assoc. rewrite <- IH1. rewrite <- IH2. 287 | rewrite PROD_fun_comp. trivial. 288 | 289 | (* TIF *) 290 | intros E t ec IHc te1 IH1 te2 IH2 E' s. rewrite substTIF. simpl. 291 | rewrite fcont_comp3_comp. 292 | rewrite <- IHc. rewrite <- IH1. rewrite <- IH2. trivial. 293 | Qed. -------------------------------------------------------------------------------- /src/utility.v: -------------------------------------------------------------------------------- 1 | (* utility.v 2 | 3 | Generally useful stuff that doesn't belong anywhere else for 4 | 5 | Abstracting Allocation : The New New Thing 6 | Nick Benton 7 | 8 | (c) 2006 Microsoft Research 9 | 10 | This also exports a bunch of pervasive stuff from the standard library 11 | *) 12 | 13 | Require Export Arith. 14 | Require Export EqNat. 15 | Require Export Bool. 16 | Require Export BoolEq. 17 | Require Export Setoid. 18 | Require Export Omega. 19 | Require Export Bool_nat. 20 | Require Export List. 21 | Require Export Max. 22 | Require Export Min. 23 | 24 | Lemma generalise_recursion : forall (P:nat->Prop), (forall k , P k) <-> (forall k k0, k0<=k -> P k0). 25 | Proof. 26 | intuition; apply (H k k); trivial. 27 | Qed. 28 | 29 | Ltac dcase x := generalize (refl_equal x); pattern x at -1; case x. 30 | 31 | Notation "[ x ; .. ; y ]" := (cons x .. (cons y nil) ..). 32 | 33 | Tactic Notation "hreplace" constr(h) "with" constr(t) := let id:=fresh in (assert (id := t); clear h; rename id into h). 34 | 35 | Tactic Notation "omega_replace" constr(m) "with" constr(n) := replace m with n ; try omega. 36 | 37 | 38 | (*Tactic Notation "deconstruct" hyp(h) := _deconstruct h.*) 39 | Tactic Notation "deconstruct" hyp(h) "as" simple_intropattern(ip) := generalize h; clear h; intros ip. 40 | 41 | Tactic Notation "witness" ident(p) "in" hyp(h) := elim h; clear h; intros p h. 42 | 43 | Axiom Extensionality : forall (A B : Type) (f g : A->B), (forall a:A, f a = g a) <-> f =g. 44 | 45 | Lemma strictPosSucc: forall n, n>0 -> exists m , S m =n. 46 | Proof. 47 | intros. 48 | induction n. 49 | destruct H. 50 | exists n; trivial. 51 | exists m; trivial. 52 | exists n; trivial. 53 | Qed. 54 | 55 | Ltac unfold_plus:= 56 | try repeat rewrite <- plus_n_Sm; 57 | try rewrite plus_0_r; 58 | try rewrite plus_0_l. 59 | 60 | Tactic Notation (at level 0) "unfold_plus" "in" constr(H) := try repeat rewrite <- plus_n_Sm in H; 61 | try rewrite plus_0_r in H; 62 | try rewrite plus_0_l in H. 63 | 64 | Lemma beq_nat_neq :forall m n, m<>n -> false = (beq_nat m n). 65 | double induction m n; simpl in |-*. 66 | intros. 67 | tauto. 68 | firstorder. 69 | firstorder. 70 | intros. 71 | set(Hn := (H0 n)). 72 | assert (n1 <> n). 73 | omega. 74 | apply (Hn H2). 75 | Qed. 76 | 77 | (* bleq_nat, ble_nat ... *) 78 | 79 | 80 | Fixpoint bleq_nat (n m: nat) {struct n} : bool := 81 | match n with 82 | | 0 => true 83 | | S n1 => match m with 84 | | 0 => false 85 | | S m1 => bleq_nat n1 m1 86 | end 87 | end. 88 | 89 | 90 | Fixpoint ble_nat (n m : nat) {struct n} : bool := 91 | match n with 92 | | 0 => match m with 93 | | 0 => false 94 | | S _ => true 95 | end 96 | | S n1 => match m with 97 | | 0 => false 98 | | S m1 => ble_nat n1 m1 99 | end 100 | end. 101 | 102 | 103 | Lemma ble_nat_false : forall m n, m >=n -> false = ble_nat m n. 104 | double induction m n. 105 | firstorder. 106 | firstorder. 107 | firstorder. 108 | firstorder. 109 | Qed. 110 | 111 | Lemma bleq_nat_false : forall m n, m > n -> false = bleq_nat m n. 112 | double induction m n. 113 | simpl. 114 | intro; absurd (0>0). 115 | omega. 116 | assumption. 117 | 118 | intros. 119 | simpl. 120 | absurd (0> S n0). 121 | omega. 122 | assumption. 123 | 124 | intros. 125 | simpl. 126 | tauto. 127 | 128 | intros. 129 | simpl. 130 | apply H0. 131 | omega. 132 | Qed. 133 | 134 | Lemma bleq_nat_true : forall m n, m <= n -> true = bleq_nat m n. 135 | double induction m n; intros; simpl. 136 | tauto. 137 | tauto. 138 | absurd (S n <=0). 139 | omega. 140 | assumption. 141 | apply H0. 142 | omega. 143 | Qed. 144 | 145 | Lemma true_bleq_nat : forall m n, true = bleq_nat m n -> m <= n. 146 | double induction m n. 147 | simpl. 148 | omega. 149 | simpl; intros. 150 | omega. 151 | simpl; intros. 152 | discriminate H0. 153 | simpl; intros. 154 | intuition. 155 | Qed. 156 | 157 | Lemma false_bleq_nat : forall m n, false = bleq_nat m n -> m > n. 158 | intros. 159 | assert (~ m <= n). 160 | intro Habs; generalize (bleq_nat_true m n Habs); intro. 161 | rewrite <- H0 in H; discriminate H. 162 | omega. 163 | Qed. 164 | 165 | Lemma ble_nat_true : forall m n, m < n -> true = ble_nat m n. 166 | double induction m n. 167 | intro; absurd (0<0). 168 | omega. 169 | assumption. 170 | 171 | intros. 172 | simpl. 173 | tauto. 174 | 175 | intros. 176 | simpl. 177 | absurd (S n < 0). 178 | omega. 179 | assumption. 180 | 181 | intros. 182 | simpl. 183 | apply H0. 184 | omega. 185 | Qed. 186 | 187 | Lemma booleq : forall (a b : bool), ((true = a) <-> (true = b)) -> (a = b). 188 | intros; dcase a. 189 | intuition. 190 | intuition. 191 | dcase b. 192 | rewrite H0 in H2. 193 | intro; symmetry; apply H2; symmetry; assumption. 194 | tauto. 195 | Qed. 196 | 197 | 198 | Definition compose (A B C : Type) (g : B -> C) (f : A -> B) a := g(f a). 199 | 200 | Implicit Arguments compose [A B C]. 201 | 202 | Definition opt_map : forall (A B : Type), (A -> B) -> (option A -> option B) := 203 | fun A B f o => 204 | match o with 205 | | None => None 206 | | Some a => Some (f a) 207 | end. 208 | 209 | Implicit Arguments opt_map [A B]. 210 | 211 | Definition opt_lift : forall (A B : Type), (A -> option B) -> option A -> option B := 212 | fun A B f o => 213 | match o with 214 | | Some a => f a 215 | | None => None 216 | end. 217 | 218 | Implicit Arguments opt_lift [A B]. 219 | 220 | Section Lists_extras. 221 | 222 | Variable A : Type. 223 | 224 | Set Implicit Arguments. 225 | 226 | (* cons_nth *) 227 | 228 | Fixpoint cons_nth (a:A) (l:list A) (m:nat) {struct l} : list A := 229 | match l with 230 | | nil => a :: nil 231 | | b :: bs => 232 | match m with 233 | | 0 => a :: l 234 | | S m' => b :: cons_nth a bs m' 235 | end 236 | end. 237 | 238 | Lemma cons_nth_S : forall a b l m, cons_nth a (b::l) (S m) = b :: (cons_nth a l m). 239 | Proof. 240 | firstorder. 241 | Qed. 242 | 243 | Lemma nth_error_nil : forall n, nth_error (nil : list A) n = error. 244 | Proof. 245 | destruct n; trivial. 246 | Qed. 247 | 248 | Lemma nth_error_cons : forall (l:list A) n a b, (nth_error l n = Some a) = (nth_error (b::l) (S n) = Some a). 249 | simpl; auto. 250 | Qed. 251 | 252 | Lemma nth_error_cons_nth_lt : forall (l:list A) n a b m (h:n (nth_error (cons_nth b l m) n = Some a). 254 | Proof. 255 | induction l. 256 | intros; rewrite (nth_error_nil n) in H; inversion H. 257 | intros until m; case m. 258 | intro h; absurd (n < 0); [apply (lt_n_O n) | assumption]. 259 | case n ;[ trivial | simpl; intros ; apply (IHl n0 a0 b n1) ; [omega |trivial]]. 260 | Qed. 261 | 262 | Lemma nth_error_cons_nth_ltI : forall (l:list A) n a b m (h : n < m) (hn:n < length l), 263 | (nth_error (cons_nth b l m) n = Some a) -> (nth_error l n = Some a). 264 | Proof. 265 | induction l. 266 | intros. simpl in hn. destruct (lt_n_O n). auto. 267 | intros. destruct n. simpl in H. destruct m. destruct (lt_irrefl 0) ; auto. 268 | auto. simpl. simpl in H ; destruct m. destruct (lt_n_O (S n)). auto. 269 | assert (S n < S (length l)). auto. 270 | apply (IHl n a0 b m (lt_S_n _ _ h) (lt_S_n _ _ H0)). auto. 271 | Qed. 272 | 273 | Lemma nth_error_cons_nth_ge : forall (l:list A) n a b m (h:n>=m), 274 | (nth_error l n = Some a) -> (nth_error (cons_nth b l m) (S n) = Some a). 275 | Proof. 276 | induction l. 277 | intro n; rewrite (nth_error_nil n); intros; inversion H. 278 | intros until m;case m. 279 | trivial. 280 | case n. 281 | intros n0 h; absurd (0 >= S n0) ; [omega | trivial]. 282 | intros n0 n1 h; simpl; simpl in IHl. 283 | apply IHl; omega. 284 | Qed. 285 | 286 | Lemma nth_error_cons_nth_geI : forall (l:list A) n a b m (h:n>=m), 287 | (nth_error (cons_nth b l m) (S n) = Some a) -> (nth_error l n = Some a). 288 | Proof. 289 | induction l. 290 | intro n. auto. 291 | intros. destruct m. simpl in H. auto. 292 | simpl in H. destruct n. destruct (le_not_gt _ _ h). 293 | apply neq_O_lt. auto. 294 | apply (IHl n a0 b _ (le_S_n _ _ h) H). 295 | Qed. 296 | 297 | Lemma nth_error_cons_nth_eq : forall a l n, 298 | n <= length l -> nth_error (cons_nth a l n) n = value a. 299 | Proof. 300 | intros. 301 | generalize dependent l. 302 | induction n. intros. destruct l ; simpl ; reflexivity. 303 | intros. 304 | destruct l. simpl in H. 305 | destruct (absurd False H (le_Sn_O n)). 306 | simpl. 307 | assert (length (a0 :: l) = S (length l)). auto. 308 | rewrite -> H0 in H. 309 | apply (IHn _ (le_S_n _ _ H)). 310 | Qed. 311 | 312 | Lemma nth_error_length : forall l n (a:A), nth_error l n = value a -> n < length l. 313 | Proof. 314 | intros. generalize dependent l. 315 | induction n. destruct l. simpl. 316 | intros. inversion H. 317 | simpl. intros. apply gt_Sn_O. 318 | intro l. destruct l. simpl. intro. inversion H. 319 | simpl. intro. apply gt_n_S. apply (IHn _ H). 320 | Qed. 321 | 322 | Lemma nth_error_lengthI : forall (l:list A) n, n < length l -> exists a, nth_error l n = value a. 323 | Proof. 324 | intros l. induction l. simpl. intros n incon. inversion incon. 325 | intros n ; case n ; clear n. simpl. intros _. exists a. auto. 326 | intros n. simpl. intros C ; apply IHl. omega. 327 | Qed. 328 | 329 | Lemma cons_nth_length_S: forall a l n, length (cons_nth a l n) = S (length l). 330 | Proof. 331 | intros. generalize dependent l. induction n. 332 | intros. destruct l ; auto. 333 | intros. destruct l. auto. 334 | simpl. apply eq_S. apply IHn. 335 | Qed. 336 | 337 | Lemma nth_error_0 : forall (l:list A) (a:A), nth_error l 0 = Some a <-> (exists l', l = a::l'). 338 | Proof. 339 | intros; split. 340 | induction l. 341 | simpl; intro. 342 | inversion H. 343 | simpl; intro. 344 | inversion H. 345 | exists l;trivial. 346 | intro; destruct H. 347 | rewrite H. 348 | simpl;trivial. 349 | Qed. 350 | 351 | Lemma nth_error_Sn : forall (l:list A) (a:A) (n:nat), 352 | nth_error l (S n) = Some a <-> (exists l', exists a0, l = a0::l' /\ nth_error l' n = Some a). 353 | Proof. 354 | intros; split. 355 | induction l. 356 | simpl; intro. 357 | inversion H. 358 | simpl; intro. 359 | exists l;exists a0;intuition. 360 | intro; destruct H as [l' [a0 (H0,H1)]]. 361 | rewrite H0. 362 | simpl;trivial. 363 | Qed. 364 | 365 | 366 | Lemma nth_error_map: forall B (f:A -> B) (l:list A) i, 367 | nth_error (map f l) i = (match nth_error l i with Some e => Some (f e) | error => error end). 368 | Proof. 369 | intros B f l. 370 | induction l. 371 | simpl. intros i. destruct i ; simpl ; auto. 372 | intros i. destruct i. simpl. auto. 373 | simpl. apply IHl. 374 | Qed. 375 | 376 | 377 | Lemma split_list : forall (a:A) l , a :: l = (a :: nil) ++ l. 378 | Proof. 379 | trivial. 380 | Qed. 381 | 382 | Fixpoint takedrop n (l1 l2 : list A) {struct n} := match n with 383 | | 0 => (l1,l2) 384 | | S n => match l2 with 385 | | nil => (l1,l2) 386 | | h::t => takedrop n (l1++(h::nil)) t 387 | end 388 | end. 389 | 390 | Definition fuse (couple:list A * list A) := match couple with (l1,l2) => l1++l2 end. 391 | 392 | Lemma fuse_unfold : forall l1 l2, fuse (l1,l2) = l1++l2. 393 | Proof. 394 | unfold fuse; intuition. 395 | Qed. 396 | 397 | Opaque fuse. 398 | 399 | Unset Implicit Arguments. 400 | 401 | End Lists_extras. 402 | --------------------------------------------------------------------------------