| author | blanchet | 
| Wed, 18 Jun 2014 17:42:24 +0200 | |
| changeset 57278 | 8f7d6f01a775 | 
| parent 46935 | 38ecb2dc3636 | 
| child 57492 | 74bf65a1910a | 
| permissions | -rw-r--r-- | 
| 1478 | 1 | (* Title: ZF/ex/LList.thy | 
| 2 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | |
| 515 | 3 | Copyright 1994 University of Cambridge | 
| 4 | ||
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
24893diff
changeset | 5 | Codatatype definition of Lazy Lists. | 
| 515 | 6 | |
| 7 | Equality for llist(A) as a greatest fixed point | |
| 8 | ||
| 9 | Functions for Lazy Lists | |
| 10 | ||
| 11 | STILL NEEDS: | |
| 12 | co_recursion for defining lconst, flip, etc. | |
| 13 | a typing rule for it, based on some notion of "productivity..." | |
| 14 | *) | |
| 15 | ||
| 16417 | 16 | theory LList imports Main begin | 
| 515 | 17 | |
| 18 | consts | |
| 12867 | 19 | llist :: "i=>i"; | 
| 515 | 20 | |
| 21 | codatatype | |
| 12867 | 22 |   "llist(A)" = LNil | LCons ("a \<in> A", "l \<in> llist(A)")
 | 
| 515 | 23 | |
| 24 | ||
| 25 | (*Coinductive definition of equality*) | |
| 26 | consts | |
| 12867 | 27 | lleq :: "i=>i" | 
| 515 | 28 | |
| 29 | (*Previously used <*> in the domain and variant pairs as elements. But | |
| 30 | standard pairs work just as well. To use variant pairs, must change prefix | |
| 31 | a q/Q to the Sigma, Pair and converse rules.*) | |
| 32 | coinductive | |
| 46822 
95f1e700b712
 mathematical symbols for Isabelle/ZF example theories
 paulson parents: 
42798diff
changeset | 33 | domains "lleq(A)" \<subseteq> "llist(A) * llist(A)" | 
| 12867 | 34 | intros | 
| 35 | LNil: "<LNil, LNil> \<in> lleq(A)" | |
| 36 | LCons: "[| a \<in> A; <l,l'> \<in> lleq(A) |] | |
| 37 | ==> <LCons(a,l), LCons(a,l')> \<in> lleq(A)" | |
| 38 | type_intros llist.intros | |
| 515 | 39 | |
| 40 | ||
| 41 | (*Lazy list functions; flip is not definitional!*) | |
| 24893 | 42 | definition | 
| 43 | lconst :: "i => i" where | |
| 44 | "lconst(a) == lfp(univ(a), %l. LCons(a,l))" | |
| 515 | 45 | |
| 41779 | 46 | axiomatization flip :: "i => i" | 
| 47 | where | |
| 48 | flip_LNil: "flip(LNil) = LNil" and | |
| 12867 | 49 | flip_LCons: "[| x \<in> bool; l \<in> llist(bool) |] | 
| 50 | ==> flip(LCons(x,l)) = LCons(not(x), flip(l))" | |
| 51 | ||
| 52 | ||
| 53 | (*These commands cause classical reasoning to regard the subset relation | |
| 54 | as primitive, not reducing it to membership*) | |
| 55 | ||
| 56 | declare subsetI [rule del] | |
| 57 | subsetCE [rule del] | |
| 58 | ||
| 59 | declare subset_refl [intro!] | |
| 60 | cons_subsetI [intro!] | |
| 61 | subset_consI [intro!] | |
| 62 | Union_least [intro!] | |
| 63 | UN_least [intro!] | |
| 64 | Un_least [intro!] | |
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
24893diff
changeset | 65 | Inter_greatest [intro!] | 
| 12867 | 66 | Int_greatest [intro!] | 
| 67 | RepFun_subset [intro!] | |
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
24893diff
changeset | 68 | Un_upper1 [intro!] | 
| 12867 | 69 | Un_upper2 [intro!] | 
| 70 | Int_lower1 [intro!] | |
| 71 | Int_lower2 [intro!] | |
| 72 | ||
| 73 | (*An elimination rule, for type-checking*) | |
| 74 | inductive_cases LConsE: "LCons(a,l) \<in> llist(A)" | |
| 75 | ||
| 76 | (*Proving freeness results*) | |
| 46822 
95f1e700b712
 mathematical symbols for Isabelle/ZF example theories
 paulson parents: 
42798diff
changeset | 77 | lemma LCons_iff: "LCons(a,l)=LCons(a',l') \<longleftrightarrow> a=a' & l=l'" | 
| 12867 | 78 | by auto | 
| 79 | ||
| 80 | lemma LNil_LCons_iff: "~ LNil=LCons(a,l)" | |
| 81 | by auto | |
| 82 | ||
| 83 | (* | |
| 84 | lemma llist_unfold: "llist(A) = {0} <+> (A <*> llist(A))";
 | |
| 85 | let open llist val rew = rewrite_rule con_defs in | |
| 86 | by (fast_tac (claset() addSIs (subsetI ::map rew intros) addEs [rew elim]) 1) | |
| 87 | end | |
| 88 | done | |
| 89 | *) | |
| 90 | ||
| 91 | (*** Lemmas to justify using "llist" in other recursive type definitions ***) | |
| 92 | ||
| 93 | lemma llist_mono: "A \<subseteq> B ==> llist(A) \<subseteq> llist(B)" | |
| 94 | apply (unfold llist.defs ) | |
| 95 | apply (rule gfp_mono) | |
| 96 | apply (rule llist.bnd_mono) | |
| 97 | apply (assumption | rule quniv_mono basic_monos)+ | |
| 98 | done | |
| 99 | ||
| 100 | (** Closure of quniv(A) under llist -- why so complex? Its a gfp... **) | |
| 101 | ||
| 102 | declare QPair_Int_Vset_subset_UN [THEN subset_trans, intro!] | |
| 103 | QPair_subset_univ [intro!] | |
| 104 | empty_subsetI [intro!] | |
| 105 | one_in_quniv [THEN qunivD, intro!] | |
| 106 | declare qunivD [dest!] | |
| 107 | declare Ord_in_Ord [elim!] | |
| 108 | ||
| 46935 | 109 | lemma llist_quniv_lemma: | 
| 110 | "Ord(i) ==> l \<in> llist(quniv(A)) \<Longrightarrow> l \<inter> Vset(i) \<subseteq> univ(eclose(A))" | |
| 111 | proof (induct i arbitrary: l rule: trans_induct) | |
| 112 | case (step i l) | |
| 113 | show ?case using `l \<in> llist(quniv(A))` | |
| 114 | proof (cases l rule: llist.cases) | |
| 115 | case LNil thus ?thesis | |
| 116 | by (simp add: QInl_def QInr_def llist.con_defs) | |
| 117 | next | |
| 118 | case (LCons a l) thus ?thesis using step.hyps | |
| 119 | proof (simp add: QInl_def QInr_def llist.con_defs) | |
| 120 | show "<1; <a; l>> \<inter> Vset(i) \<subseteq> univ(eclose(A))" using LCons `Ord(i)` | |
| 121 | by (fast intro: step Ord_trans Int_lower1 [THEN subset_trans]) | |
| 122 | qed | |
| 123 | qed | |
| 124 | qed | |
| 12867 | 125 | |
| 126 | lemma llist_quniv: "llist(quniv(A)) \<subseteq> quniv(A)" | |
| 127 | apply (rule qunivI [THEN subsetI]) | |
| 128 | apply (rule Int_Vset_subset) | |
| 129 | apply (assumption | rule llist_quniv_lemma)+ | |
| 130 | done | |
| 131 | ||
| 132 | lemmas llist_subset_quniv = | |
| 133 | subset_trans [OF llist_mono llist_quniv] | |
| 134 | ||
| 135 | ||
| 136 | (*** Lazy List Equality: lleq ***) | |
| 137 | ||
| 138 | declare QPair_Int_Vset_subset_UN [THEN subset_trans, intro!] | |
| 139 | QPair_mono [intro!] | |
| 140 | ||
| 141 | declare Ord_in_Ord [elim!] | |
| 142 | ||
| 143 | (*Lemma for proving finality. Unfold the lazy list; use induction hypothesis*) | |
| 46935 | 144 | lemma lleq_Int_Vset_subset: | 
| 145 | "Ord(i) ==> <l,l'> \<in> lleq(A) \<Longrightarrow> l \<inter> Vset(i) \<subseteq> l'" | |
| 146 | proof (induct i arbitrary: l l' rule: trans_induct) | |
| 147 | case (step i l l') | |
| 148 | show ?case using `\<langle>l, l'\<rangle> \<in> lleq(A)` | |
| 149 | proof (cases rule: lleq.cases) | |
| 150 | case LNil thus ?thesis | |
| 151 | by (auto simp add: QInr_def llist.con_defs) | |
| 152 | next | |
| 153 | case (LCons a l l') thus ?thesis using step.hyps | |
| 154 | by (auto simp add: QInr_def llist.con_defs intro: Ord_trans) | |
| 155 | qed | |
| 156 | qed | |
| 515 | 157 | |
| 12867 | 158 | (*lleq(A) is a symmetric relation because qconverse(lleq(A)) is a fixedpoint*) | 
| 159 | lemma lleq_symmetric: "<l,l'> \<in> lleq(A) ==> <l',l> \<in> lleq(A)" | |
| 160 | apply (erule lleq.coinduct [OF converseI]) | |
| 13339 
0f89104dd377
Fixed quantified variable name preservation for ball and bex (bounded quants)
 paulson parents: 
12867diff
changeset | 161 | apply (rule lleq.dom_subset [THEN converse_type], safe) | 
| 
0f89104dd377
Fixed quantified variable name preservation for ball and bex (bounded quants)
 paulson parents: 
12867diff
changeset | 162 | apply (erule lleq.cases, blast+) | 
| 12867 | 163 | done | 
| 164 | ||
| 165 | lemma lleq_implies_equal: "<l,l'> \<in> lleq(A) ==> l=l'" | |
| 166 | apply (rule equalityI) | |
| 167 | apply (assumption | rule lleq_Int_Vset_subset [THEN Int_Vset_subset] | | |
| 168 | erule lleq_symmetric)+ | |
| 169 | done | |
| 170 | ||
| 171 | lemma equal_llist_implies_leq: | |
| 172 | "[| l=l'; l \<in> llist(A) |] ==> <l,l'> \<in> lleq(A)" | |
| 173 | apply (rule_tac X = "{<l,l>. l \<in> llist (A) }" in lleq.coinduct)
 | |
| 174 | apply blast | |
| 175 | apply safe | |
| 13339 
0f89104dd377
Fixed quantified variable name preservation for ball and bex (bounded quants)
 paulson parents: 
12867diff
changeset | 176 | apply (erule_tac a=l in llist.cases, fast+) | 
| 12867 | 177 | done | 
| 178 | ||
| 179 | ||
| 180 | (*** Lazy List Functions ***) | |
| 181 | ||
| 182 | (*Examples of coinduction for type-checking and to prove llist equations*) | |
| 183 | ||
| 184 | (*** lconst -- defined directly using lfp, but equivalent to a LList_corec ***) | |
| 185 | ||
| 186 | lemma lconst_fun_bnd_mono: "bnd_mono(univ(a), %l. LCons(a,l))" | |
| 187 | apply (unfold llist.con_defs ) | |
| 188 | apply (rule bnd_monoI) | |
| 189 | apply (blast intro: A_subset_univ QInr_subset_univ) | |
| 190 | apply (blast intro: subset_refl QInr_mono QPair_mono) | |
| 191 | done | |
| 192 | ||
| 193 | (* lconst(a) = LCons(a,lconst(a)) *) | |
| 194 | lemmas lconst = def_lfp_unfold [OF lconst_def lconst_fun_bnd_mono] | |
| 195 | lemmas lconst_subset = lconst_def [THEN def_lfp_subset] | |
| 196 | lemmas member_subset_Union_eclose = arg_into_eclose [THEN Union_upper] | |
| 197 | ||
| 198 | lemma lconst_in_quniv: "a \<in> A ==> lconst(a) \<in> quniv(A)" | |
| 199 | apply (rule lconst_subset [THEN subset_trans, THEN qunivI]) | |
| 200 | apply (erule arg_into_eclose [THEN eclose_subset, THEN univ_mono]) | |
| 201 | done | |
| 515 | 202 | |
| 12867 | 203 | lemma lconst_type: "a \<in> A ==> lconst(a): llist(A)" | 
| 204 | apply (rule singletonI [THEN llist.coinduct]) | |
| 205 | apply (erule lconst_in_quniv [THEN singleton_subsetI]) | |
| 206 | apply (fast intro!: lconst) | |
| 207 | done | |
| 208 | ||
| 209 | (*** flip --- equations merely assumed; certain consequences proved ***) | |
| 210 | ||
| 211 | declare flip_LNil [simp] | |
| 212 | flip_LCons [simp] | |
| 213 | not_type [simp] | |
| 214 | ||
| 46822 
95f1e700b712
 mathematical symbols for Isabelle/ZF example theories
 paulson parents: 
42798diff
changeset | 215 | lemma bool_Int_subset_univ: "b \<in> bool ==> b \<inter> X \<subseteq> univ(eclose(A))" | 
| 12867 | 216 | by (fast intro: Int_lower1 [THEN subset_trans] elim!: boolE) | 
| 217 | ||
| 218 | declare not_type [intro!] | |
| 219 | declare bool_Int_subset_univ [intro] | |
| 220 | ||
| 221 | (*Reasoning borrowed from lleq.ML; a similar proof works for all | |
| 222 | "productive" functions -- cf Coquand's "Infinite Objects in Type Theory".*) | |
| 46935 | 223 | lemma flip_llist_quniv_lemma: | 
| 224 | "Ord(i) ==> l \<in> llist(bool) \<Longrightarrow> flip(l) \<inter> Vset(i) \<subseteq> univ(eclose(bool))" | |
| 225 | proof (induct i arbitrary: l rule: trans_induct) | |
| 226 | case (step i l) | |
| 227 | show ?case using `l \<in> llist(bool)` | |
| 228 | proof (cases l rule: llist.cases) | |
| 229 | case LNil thus ?thesis | |
| 230 | by (simp, simp add: QInl_def QInr_def llist.con_defs) | |
| 231 | next | |
| 232 | case (LCons a l) thus ?thesis using step.hyps | |
| 233 | proof (simp, simp add: QInl_def QInr_def llist.con_defs) | |
| 234 | show "<1; <not(a); flip(l)>> \<inter> Vset(i) \<subseteq> univ(eclose(bool))" using LCons step.hyps | |
| 235 | by (auto intro: Ord_trans) | |
| 236 | qed | |
| 237 | qed | |
| 238 | qed | |
| 12867 | 239 | |
| 240 | lemma flip_in_quniv: "l \<in> llist(bool) ==> flip(l) \<in> quniv(bool)" | |
| 13339 
0f89104dd377
Fixed quantified variable name preservation for ball and bex (bounded quants)
 paulson parents: 
12867diff
changeset | 241 | by (rule flip_llist_quniv_lemma [THEN Int_Vset_subset, THEN qunivI], assumption+) | 
| 12867 | 242 | |
| 243 | lemma flip_type: "l \<in> llist(bool) ==> flip(l): llist(bool)" | |
| 244 | apply (rule_tac X = "{flip (l) . l \<in> llist (bool) }" in llist.coinduct)
 | |
| 245 | apply blast | |
| 246 | apply (fast intro!: flip_in_quniv) | |
| 247 | apply (erule RepFunE) | |
| 13339 
0f89104dd377
Fixed quantified variable name preservation for ball and bex (bounded quants)
 paulson parents: 
12867diff
changeset | 248 | apply (erule_tac a=la in llist.cases, auto) | 
| 12867 | 249 | done | 
| 250 | ||
| 251 | lemma flip_flip: "l \<in> llist(bool) ==> flip(flip(l)) = l" | |
| 252 | apply (rule_tac X1 = "{<flip (flip (l)),l> . l \<in> llist (bool) }" in 
 | |
| 253 | lleq.coinduct [THEN lleq_implies_equal]) | |
| 254 | apply blast | |
| 255 | apply (fast intro!: flip_type) | |
| 256 | apply (erule RepFunE) | |
| 13339 
0f89104dd377
Fixed quantified variable name preservation for ball and bex (bounded quants)
 paulson parents: 
12867diff
changeset | 257 | apply (erule_tac a=la in llist.cases) | 
| 12867 | 258 | apply (simp_all add: flip_type not_not) | 
| 259 | done | |
| 515 | 260 | |
| 261 | end | |
| 262 |