| author | aspinall | 
| Wed, 03 Jan 2007 21:11:42 +0100 | |
| changeset 21983 | 9fb029d1189b | 
| parent 16417 | 9bc16273c2d4 | 
| child 24893 | b8ef7afe3a6b | 
| permissions | -rw-r--r-- | 
| 1478 | 1 | (* Title: ZF/Coind/ECR.thy | 
| 915 | 2 | ID: $Id$ | 
| 1478 | 3 | Author: Jacob Frost, Cambridge University Computer Laboratory | 
| 915 | 4 | Copyright 1995 University of Cambridge | 
| 5 | *) | |
| 6 | ||
| 16417 | 7 | theory ECR imports Static Dynamic begin | 
| 915 | 8 | |
| 9 | (* The extended correspondence relation *) | |
| 10 | ||
| 11 | consts | |
| 1401 | 12 | HasTyRel :: i | 
| 915 | 13 | coinductive | 
| 14 | domains "HasTyRel" <= "Val * Ty" | |
| 12595 | 15 | intros | 
| 16 | htr_constI [intro!]: | |
| 17 | "[| c \<in> Const; t \<in> Ty; isof(c,t) |] ==> <v_const(c),t> \<in> HasTyRel" | |
| 18 | htr_closI [intro]: | |
| 19 | "[| x \<in> ExVar; e \<in> Exp; t \<in> Ty; ve \<in> ValEnv; te \<in> TyEnv; | |
| 20 | <te,e_fn(x,e),t> \<in> ElabRel; | |
| 1478 | 21 | ve_dom(ve) = te_dom(te); | 
| 12595 | 22 |           {<ve_app(ve,y),te_app(te,y)>.y \<in> ve_dom(ve)} \<in> Pow(HasTyRel) |] 
 | 
| 23 | ==> <v_clos(x,e,ve),t> \<in> HasTyRel" | |
| 6117 | 24 | monos Pow_mono | 
| 12595 | 25 | type_intros Val_ValEnv.intros | 
| 915 | 26 | |
| 27 | (* Pointwise extension to environments *) | |
| 28 | ||
| 12595 | 29 | constdefs | 
| 30 | hastyenv :: "[i,i] => o" | |
| 31 | "hastyenv(ve,te) == | |
| 32 | ve_dom(ve) = te_dom(te) & | |
| 33 | (\<forall>x \<in> ve_dom(ve). <ve_app(ve,x),te_app(te,x)> \<in> HasTyRel)" | |
| 34 | ||
| 35 | (* Specialised co-induction rule *) | |
| 36 | ||
| 37 | lemma htr_closCI [intro]: | |
| 38 | "[| x \<in> ExVar; e \<in> Exp; t \<in> Ty; ve \<in> ValEnv; te \<in> TyEnv; | |
| 39 | <te, e_fn(x, e), t> \<in> ElabRel; ve_dom(ve) = te_dom(te); | |
| 40 |          {<ve_app(ve,y),te_app(te,y)>.y \<in> ve_dom(ve)} \<in>
 | |
| 41 |            Pow({<v_clos(x,e,ve),t>} Un HasTyRel) |]     
 | |
| 42 | ==> <v_clos(x, e, ve),t> \<in> HasTyRel" | |
| 13339 
0f89104dd377
Fixed quantified variable name preservation for ball and bex (bounded quants)
 paulson parents: 
12610diff
changeset | 43 | apply (rule singletonI [THEN HasTyRel.coinduct], auto) | 
| 12595 | 44 | done | 
| 45 | ||
| 46 | (* Specialised elimination rules *) | |
| 47 | ||
| 12610 | 48 | inductive_cases | 
| 49 | htr_constE [elim!]: "<v_const(c),t> \<in> HasTyRel" | |
| 50 | and htr_closE [elim]: "<v_clos(x,e,ve),t> \<in> HasTyRel" | |
| 12595 | 51 | |
| 52 | ||
| 53 | (* Properties of the pointwise extension to environments *) | |
| 54 | ||
| 55 | lemmas HasTyRel_non_zero = | |
| 56 | HasTyRel.dom_subset [THEN subsetD, THEN SigmaD1, THEN ValNEE, standard] | |
| 57 | ||
| 58 | lemma hastyenv_owr: | |
| 59 | "[| ve \<in> ValEnv; te \<in> TyEnv; hastyenv(ve,te); <v,t> \<in> HasTyRel |] | |
| 60 | ==> hastyenv(ve_owr(ve,x,v),te_owr(te,x,t))" | |
| 61 | by (auto simp add: hastyenv_def ve_app_owr HasTyRel_non_zero) | |
| 62 | ||
| 63 | lemma basic_consistency_lem: | |
| 64 | "[| ve \<in> ValEnv; te \<in> TyEnv; isofenv(ve,te) |] ==> hastyenv(ve,te)" | |
| 65 | apply (unfold isofenv_def hastyenv_def) | |
| 66 | apply (force intro: te_appI ve_domI) | |
| 67 | done | |
| 68 | ||
| 69 | ||
| 70 | (* ############################################################ *) | |
| 71 | (* The Consistency theorem *) | |
| 72 | (* ############################################################ *) | |
| 73 | ||
| 74 | lemma consistency_const: | |
| 75 | "[| c \<in> Const; hastyenv(ve,te);<te,e_const(c),t> \<in> ElabRel |] | |
| 76 | ==> <v_const(c), t> \<in> HasTyRel" | |
| 77 | by blast | |
| 78 | ||
| 79 | ||
| 80 | lemma consistency_var: | |
| 81 | "[| x \<in> ve_dom(ve); hastyenv(ve,te); <te,e_var(x),t> \<in> ElabRel |] ==> | |
| 82 | <ve_app(ve,x),t> \<in> HasTyRel" | |
| 83 | by (unfold hastyenv_def, blast) | |
| 84 | ||
| 85 | lemma consistency_fn: | |
| 86 | "[| ve \<in> ValEnv; x \<in> ExVar; e \<in> Exp; hastyenv(ve,te); | |
| 87 | <te,e_fn(x,e),t> \<in> ElabRel | |
| 88 | |] ==> <v_clos(x, e, ve), t> \<in> HasTyRel" | |
| 89 | by (unfold hastyenv_def, blast) | |
| 90 | ||
| 91 | declare ElabRel.dom_subset [THEN subsetD, dest] | |
| 92 | ||
| 93 | declare Ty.intros [simp, intro!] | |
| 94 | declare TyEnv.intros [simp, intro!] | |
| 95 | declare Val_ValEnv.intros [simp, intro!] | |
| 96 | ||
| 97 | lemma consistency_fix: | |
| 98 | "[| ve \<in> ValEnv; x \<in> ExVar; e \<in> Exp; f \<in> ExVar; cl \<in> Val; | |
| 99 | v_clos(x,e,ve_owr(ve,f,cl)) = cl; | |
| 100 | hastyenv(ve,te); <te,e_fix(f,x,e),t> \<in> ElabRel |] ==> | |
| 101 | <cl,t> \<in> HasTyRel" | |
| 102 | apply (unfold hastyenv_def) | |
| 13339 
0f89104dd377
Fixed quantified variable name preservation for ball and bex (bounded quants)
 paulson parents: 
12610diff
changeset | 103 | apply (erule elab_fixE, safe) | 
| 12595 | 104 | apply (rule subst, assumption) | 
| 105 | apply (rule_tac te="te_owr(te, f, t_fun(t1, t2))" in htr_closCI) | |
| 106 | apply simp_all | |
| 107 | apply (blast intro: ve_owrI) | |
| 108 | apply (rule ElabRel.fnI) | |
| 13339 
0f89104dd377
Fixed quantified variable name preservation for ball and bex (bounded quants)
 paulson parents: 
12610diff
changeset | 109 | apply (simp_all add: ValNEE, force) | 
| 12595 | 110 | done | 
| 111 | ||
| 112 | ||
| 113 | lemma consistency_app1: | |
| 114 | "[| ve \<in> ValEnv; e1 \<in> Exp; e2 \<in> Exp; c1 \<in> Const; c2 \<in> Const; | |
| 115 | <ve,e1,v_const(c1)> \<in> EvalRel; | |
| 116 | \<forall>t te. | |
| 117 | hastyenv(ve,te) --> <te,e1,t> \<in> ElabRel --> <v_const(c1),t> \<in> HasTyRel; | |
| 118 | <ve, e2, v_const(c2)> \<in> EvalRel; | |
| 119 | \<forall>t te. | |
| 120 | hastyenv(ve,te) --> <te,e2,t> \<in> ElabRel --> <v_const(c2),t> \<in> HasTyRel; | |
| 121 | hastyenv(ve, te); | |
| 122 | <te,e_app(e1,e2),t> \<in> ElabRel |] | |
| 123 | ==> <v_const(c_app(c1, c2)),t> \<in> HasTyRel" | |
| 124 | by (blast intro!: c_appI intro: isof_app) | |
| 125 | ||
| 126 | lemma consistency_app2: | |
| 127 | "[| ve \<in> ValEnv; vem \<in> ValEnv; e1 \<in> Exp; e2 \<in> Exp; em \<in> Exp; xm \<in> ExVar; | |
| 128 | v \<in> Val; | |
| 129 | <ve,e1,v_clos(xm,em,vem)> \<in> EvalRel; | |
| 130 | \<forall>t te. hastyenv(ve,te) --> | |
| 131 | <te,e1,t> \<in> ElabRel --> <v_clos(xm,em,vem),t> \<in> HasTyRel; | |
| 132 | <ve,e2,v2> \<in> EvalRel; | |
| 133 | \<forall>t te. hastyenv(ve,te) --> <te,e2,t> \<in> ElabRel --> <v2,t> \<in> HasTyRel; | |
| 134 | <ve_owr(vem,xm,v2),em,v> \<in> EvalRel; | |
| 135 | \<forall>t te. hastyenv(ve_owr(vem,xm,v2),te) --> | |
| 136 | <te,em,t> \<in> ElabRel --> <v,t> \<in> HasTyRel; | |
| 137 | hastyenv(ve,te); <te,e_app(e1,e2),t> \<in> ElabRel |] | |
| 138 | ==> <v,t> \<in> HasTyRel" | |
| 139 | apply (erule elab_appE) | |
| 140 | apply (drule spec [THEN spec, THEN mp, THEN mp], assumption+) | |
| 141 | apply (drule spec [THEN spec, THEN mp, THEN mp], assumption+) | |
| 142 | apply (erule htr_closE) | |
| 13339 
0f89104dd377
Fixed quantified variable name preservation for ball and bex (bounded quants)
 paulson parents: 
12610diff
changeset | 143 | apply (erule elab_fnE, simp) | 
| 12595 | 144 | apply clarify | 
| 145 | apply (drule spec [THEN spec, THEN mp, THEN mp]) | |
| 146 | prefer 2 apply assumption+ | |
| 13339 
0f89104dd377
Fixed quantified variable name preservation for ball and bex (bounded quants)
 paulson parents: 
12610diff
changeset | 147 | apply (rule hastyenv_owr, assumption) | 
| 12595 | 148 | apply assumption | 
| 13339 
0f89104dd377
Fixed quantified variable name preservation for ball and bex (bounded quants)
 paulson parents: 
12610diff
changeset | 149 | apply (simp add: hastyenv_def, blast+) | 
| 12595 | 150 | done | 
| 151 | ||
| 152 | lemma consistency [rule_format]: | |
| 153 | "<ve,e,v> \<in> EvalRel | |
| 154 | ==> (\<forall>t te. hastyenv(ve,te) --> <te,e,t> \<in> ElabRel --> <v,t> \<in> HasTyRel)" | |
| 155 | apply (erule EvalRel.induct) | |
| 156 | apply (simp_all add: consistency_const consistency_var consistency_fn | |
| 157 | consistency_fix consistency_app1) | |
| 158 | apply (blast intro: consistency_app2) | |
| 159 | done | |
| 160 | ||
| 161 | lemma basic_consistency: | |
| 162 | "[| ve \<in> ValEnv; te \<in> TyEnv; isofenv(ve,te); | |
| 163 | <ve,e,v_const(c)> \<in> EvalRel; <te,e,t> \<in> ElabRel |] | |
| 164 | ==> isof(c,t)" | |
| 165 | by (blast dest: consistency intro!: basic_consistency_lem) | |
| 915 | 166 | |
| 167 | end |