| author | wenzelm | 
| Tue, 09 Jan 2018 17:40:25 +0100 | |
| changeset 67389 | 7e21d19e7ad7 | 
| parent 62143 | 3c9a0985e6be | 
| child 69593 | 3dda49e08b9d | 
| permissions | -rw-r--r-- | 
| 17456 | 1 | (* Title: CCL/Term.thy | 
| 1474 | 2 | Author: Martin Coen | 
| 0 | 3 | Copyright 1993 University of Cambridge | 
| 4 | *) | |
| 5 | ||
| 60770 | 6 | section \<open>Definitions of usual program constructs in CCL\<close> | 
| 17456 | 7 | |
| 8 | theory Term | |
| 9 | imports CCL | |
| 10 | begin | |
| 0 | 11 | |
| 62143 | 12 | definition one :: "i" | 
| 13 | where "one == true" | |
| 0 | 14 | |
| 62143 | 15 | definition "if" :: "[i,i,i]\<Rightarrow>i"  ("(3if _/ then _/ else _)" [0,0,60] 60)
 | 
| 16 | where "if b then t else u == case(b, t, u, \<lambda> x y. bot, \<lambda>v. bot)" | |
| 0 | 17 | |
| 62143 | 18 | definition inl :: "i\<Rightarrow>i" | 
| 19 | where "inl(a) == <true,a>" | |
| 0 | 20 | |
| 62143 | 21 | definition inr :: "i\<Rightarrow>i" | 
| 22 | where "inr(b) == <false,b>" | |
| 23 | ||
| 24 | definition split :: "[i,[i,i]\<Rightarrow>i]\<Rightarrow>i" | |
| 25 | where "split(t,f) == case(t, bot, bot, f, \<lambda>u. bot)" | |
| 0 | 26 | |
| 62143 | 27 | definition "when" :: "[i,i\<Rightarrow>i,i\<Rightarrow>i]\<Rightarrow>i" | 
| 28 | where "when(t,f,g) == split(t, \<lambda>b x. if b then f(x) else g(x))" | |
| 29 | ||
| 30 | definition fst :: "i\<Rightarrow>i" | |
| 31 | where "fst(t) == split(t, \<lambda>x y. x)" | |
| 0 | 32 | |
| 62143 | 33 | definition snd :: "i\<Rightarrow>i" | 
| 34 | where "snd(t) == split(t, \<lambda>x y. y)" | |
| 35 | ||
| 36 | definition thd :: "i\<Rightarrow>i" | |
| 37 | where "thd(t) == split(t, \<lambda>x p. split(p, \<lambda>y z. z))" | |
| 38 | ||
| 39 | definition zero :: "i" | |
| 40 | where "zero == inl(one)" | |
| 0 | 41 | |
| 62143 | 42 | definition succ :: "i\<Rightarrow>i" | 
| 43 | where "succ(n) == inr(n)" | |
| 44 | ||
| 45 | definition ncase :: "[i,i,i\<Rightarrow>i]\<Rightarrow>i" | |
| 46 | where "ncase(n,b,c) == when(n, \<lambda>x. b, \<lambda>y. c(y))" | |
| 998 
91d09e262799
Gave tighter priorities to if, napply and the let-forms to
 lcp parents: 
610diff
changeset | 47 | |
| 62143 | 48 | |
| 49 | no_syntax | |
| 50 |   "_Let" :: "[letbinds, 'a] => 'a"  ("(let (_)/ in (_))" 10)
 | |
| 51 | ||
| 52 | definition "let" :: "[i,i\<Rightarrow>i]\<Rightarrow>i" | |
| 53 | where "let(t, f) == case(t,f(true),f(false), \<lambda>x y. f(<x,y>), \<lambda>u. f(lam x. u(x)))" | |
| 0 | 54 | |
| 14765 | 55 | syntax | 
| 62143 | 56 |   "_Let" :: "[letbinds, 'a] => 'a"  ("(let (_)/ in (_))" 10)
 | 
| 998 
91d09e262799
Gave tighter priorities to if, napply and the let-forms to
 lcp parents: 
610diff
changeset | 57 | |
| 62143 | 58 | definition letrec :: "[[i,i\<Rightarrow>i]\<Rightarrow>i,(i\<Rightarrow>i)\<Rightarrow>i]\<Rightarrow>i" | 
| 59 | where "letrec(h, b) == b(\<lambda>x. fix(\<lambda>f. lam x. h(x,\<lambda>y. f`y))`x)" | |
| 60 | ||
| 61 | definition letrec2 :: "[[i,i,i\<Rightarrow>i\<Rightarrow>i]\<Rightarrow>i,(i\<Rightarrow>i\<Rightarrow>i)\<Rightarrow>i]\<Rightarrow>i" | |
| 62 | where "letrec2 (h, f) == | |
| 63 | letrec (\<lambda>p g'. split(p,\<lambda>x y. h(x,y,\<lambda>u v. g'(<u,v>))), \<lambda>g'. f(\<lambda>x y. g'(<x,y>)))" | |
| 0 | 64 | |
| 62143 | 65 | definition letrec3 :: "[[i,i,i,i\<Rightarrow>i\<Rightarrow>i\<Rightarrow>i]\<Rightarrow>i,(i\<Rightarrow>i\<Rightarrow>i\<Rightarrow>i)\<Rightarrow>i]\<Rightarrow>i" | 
| 66 | where "letrec3 (h, f) == | |
| 67 | letrec (\<lambda>p g'. split(p,\<lambda>x xs. split(xs,\<lambda>y z. h(x,y,z,\<lambda>u v w. g'(<u,<v,w>>)))), | |
| 68 | \<lambda>g'. f(\<lambda>x y z. g'(<x,<y,z>>)))" | |
| 0 | 69 | |
| 62143 | 70 | syntax | 
| 71 |   "_let" :: "[id,i,i]\<Rightarrow>i"  ("(3let _ be _/ in _)" [0,0,60] 60)
 | |
| 72 |   "_letrec" :: "[id,id,i,i]\<Rightarrow>i"  ("(3letrec _ _ be _/ in _)" [0,0,0,60] 60)
 | |
| 73 |   "_letrec2" :: "[id,id,id,i,i]\<Rightarrow>i"  ("(3letrec _ _ _ be _/ in _)" [0,0,0,0,60] 60)
 | |
| 74 |   "_letrec3" :: "[id,id,id,id,i,i]\<Rightarrow>i"  ("(3letrec _ _ _ _ be _/ in _)" [0,0,0,0,0,60] 60)
 | |
| 998 
91d09e262799
Gave tighter priorities to if, napply and the let-forms to
 lcp parents: 
610diff
changeset | 75 | |
| 60770 | 76 | ML \<open> | 
| 0 | 77 | (** Quantifier translations: variable binding **) | 
| 78 | ||
| 17781 | 79 | (* FIXME does not handle "_idtdummy" *) | 
| 49660 
de49d9b4d7bc
more explicit Syntax_Trans.mark_bound_abs/mark_bound_body: preserve type information for show_markup;
 wenzelm parents: 
44241diff
changeset | 80 | (* FIXME should use Syntax_Trans.mark_bound, Syntax_Trans.variant_abs' *) | 
| 2709 | 81 | |
| 44241 | 82 | fun let_tr [Free x, a, b] = Const(@{const_syntax let},dummyT) $ a $ absfree x b;
 | 
| 0 | 83 | fun let_tr' [a,Abs(id,T,b)] = | 
| 42284 | 84 | let val (id',b') = Syntax_Trans.variant_abs(id,T,b) | 
| 35113 | 85 |      in Const(@{syntax_const "_let"},dummyT) $ Free(id',T) $ a $ b' end;
 | 
| 0 | 86 | |
| 44241 | 87 | fun letrec_tr [Free f, Free x, a, b] = | 
| 88 |       Const(@{const_syntax letrec}, dummyT) $ absfree x (absfree f a) $ absfree f b;
 | |
| 89 | fun letrec2_tr [Free f, Free x, Free y, a, b] = | |
| 90 |       Const(@{const_syntax letrec2}, dummyT) $ absfree x (absfree y (absfree f a)) $ absfree f b;
 | |
| 91 | fun letrec3_tr [Free f, Free x, Free y, Free z, a, b] = | |
| 92 |       Const(@{const_syntax letrec3}, dummyT) $
 | |
| 93 | absfree x (absfree y (absfree z (absfree f a))) $ absfree f b; | |
| 0 | 94 | |
| 95 | fun letrec_tr' [Abs(x,T,Abs(f,S,a)),Abs(ff,SS,b)] = | |
| 42284 | 96 | let val (f',b') = Syntax_Trans.variant_abs(ff,SS,b) | 
| 97 | val (_,a'') = Syntax_Trans.variant_abs(f,S,a) | |
| 98 | val (x',a') = Syntax_Trans.variant_abs(x,T,a'') | |
| 35113 | 99 |      in Const(@{syntax_const "_letrec"},dummyT) $ Free(f',SS) $ Free(x',T) $ a' $ b' end;
 | 
| 0 | 100 | fun letrec2_tr' [Abs(x,T,Abs(y,U,Abs(f,S,a))),Abs(ff,SS,b)] = | 
| 42284 | 101 | let val (f',b') = Syntax_Trans.variant_abs(ff,SS,b) | 
| 102 | val ( _,a1) = Syntax_Trans.variant_abs(f,S,a) | |
| 103 | val (y',a2) = Syntax_Trans.variant_abs(y,U,a1) | |
| 104 | val (x',a') = Syntax_Trans.variant_abs(x,T,a2) | |
| 35113 | 105 |      in Const(@{syntax_const "_letrec2"},dummyT) $ Free(f',SS) $ Free(x',T) $ Free(y',U) $ a' $ b'
 | 
| 0 | 106 | end; | 
| 107 | fun letrec3_tr' [Abs(x,T,Abs(y,U,Abs(z,V,Abs(f,S,a)))),Abs(ff,SS,b)] = | |
| 42284 | 108 | let val (f',b') = Syntax_Trans.variant_abs(ff,SS,b) | 
| 109 | val ( _,a1) = Syntax_Trans.variant_abs(f,S,a) | |
| 110 | val (z',a2) = Syntax_Trans.variant_abs(z,V,a1) | |
| 111 | val (y',a3) = Syntax_Trans.variant_abs(y,U,a2) | |
| 112 | val (x',a') = Syntax_Trans.variant_abs(x,T,a3) | |
| 35113 | 113 |      in Const(@{syntax_const "_letrec3"},dummyT) $ Free(f',SS) $ Free(x',T) $ Free(y',U) $ Free(z',V) $ a' $ b'
 | 
| 0 | 114 | end; | 
| 60770 | 115 | \<close> | 
| 17456 | 116 | |
| 60770 | 117 | parse_translation \<open> | 
| 52143 | 118 |  [(@{syntax_const "_let"}, K let_tr),
 | 
| 119 |   (@{syntax_const "_letrec"}, K letrec_tr),
 | |
| 120 |   (@{syntax_const "_letrec2"}, K letrec2_tr),
 | |
| 121 |   (@{syntax_const "_letrec3"}, K letrec3_tr)]
 | |
| 60770 | 122 | \<close> | 
| 17456 | 123 | |
| 60770 | 124 | print_translation \<open> | 
| 52143 | 125 |  [(@{const_syntax let}, K let_tr'),
 | 
| 126 |   (@{const_syntax letrec}, K letrec_tr'),
 | |
| 127 |   (@{const_syntax letrec2}, K letrec2_tr'),
 | |
| 128 |   (@{const_syntax letrec3}, K letrec3_tr')]
 | |
| 60770 | 129 | \<close> | 
| 17456 | 130 | |
| 62143 | 131 | |
| 132 | definition nrec :: "[i,i,[i,i]\<Rightarrow>i]\<Rightarrow>i" | |
| 133 | where "nrec(n,b,c) == letrec g x be ncase(x, b, \<lambda>y. c(y,g(y))) in g(n)" | |
| 134 | ||
| 135 | definition nil :: "i"  ("([])")
 | |
| 136 | where "[] == inl(one)" | |
| 17456 | 137 | |
| 62143 | 138 | definition cons :: "[i,i]\<Rightarrow>i" (infixr "$" 80) | 
| 139 | where "h$t == inr(<h,t>)" | |
| 140 | ||
| 141 | definition lcase :: "[i,i,[i,i]\<Rightarrow>i]\<Rightarrow>i" | |
| 142 | where "lcase(l,b,c) == when(l, \<lambda>x. b, \<lambda>y. split(y,c))" | |
| 17456 | 143 | |
| 62143 | 144 | definition lrec :: "[i,i,[i,i,i]\<Rightarrow>i]\<Rightarrow>i" | 
| 145 | where "lrec(l,b,c) == letrec g x be lcase(x, b, \<lambda>h t. c(h,t,g(t))) in g(l)" | |
| 17456 | 146 | |
| 62143 | 147 | definition napply :: "[i\<Rightarrow>i,i,i]\<Rightarrow>i"  ("(_ ^ _ ` _)" [56,56,56] 56)
 | 
| 148 | where "f ^n` a == nrec(n,a,\<lambda>x g. f(g))" | |
| 20140 | 149 | |
| 150 | lemmas simp_can_defs = one_def inl_def inr_def | |
| 151 | and simp_ncan_defs = if_def when_def split_def fst_def snd_def thd_def | |
| 152 | lemmas simp_defs = simp_can_defs simp_ncan_defs | |
| 153 | ||
| 154 | lemmas ind_can_defs = zero_def succ_def nil_def cons_def | |
| 155 | and ind_ncan_defs = ncase_def nrec_def lcase_def lrec_def | |
| 156 | lemmas ind_defs = ind_can_defs ind_ncan_defs | |
| 157 | ||
| 158 | lemmas data_defs = simp_defs ind_defs napply_def | |
| 159 | and genrec_defs = letrec_def letrec2_def letrec3_def | |
| 160 | ||
| 161 | ||
| 60770 | 162 | subsection \<open>Beta Rules, including strictness\<close> | 
| 20140 | 163 | |
| 58977 | 164 | lemma letB: "\<not> t=bot \<Longrightarrow> let x be t in f(x) = f(t)" | 
| 20140 | 165 | apply (unfold let_def) | 
| 166 | apply (erule rev_mp) | |
| 167 | apply (rule_tac t = "t" in term_case) | |
| 168 | apply (simp_all add: caseBtrue caseBfalse caseBpair caseBlam) | |
| 169 | done | |
| 170 | ||
| 171 | lemma letBabot: "let x be bot in f(x) = bot" | |
| 172 | apply (unfold let_def) | |
| 173 | apply (rule caseBbot) | |
| 174 | done | |
| 175 | ||
| 176 | lemma letBbbot: "let x be t in bot = bot" | |
| 177 | apply (unfold let_def) | |
| 178 | apply (rule_tac t = t in term_case) | |
| 179 | apply (rule caseBbot) | |
| 180 | apply (simp_all add: caseBtrue caseBfalse caseBpair caseBlam) | |
| 181 | done | |
| 182 | ||
| 183 | lemma applyB: "(lam x. b(x)) ` a = b(a)" | |
| 184 | apply (unfold apply_def) | |
| 185 | apply (simp add: caseBtrue caseBfalse caseBpair caseBlam) | |
| 186 | done | |
| 187 | ||
| 188 | lemma applyBbot: "bot ` a = bot" | |
| 189 | apply (unfold apply_def) | |
| 190 | apply (rule caseBbot) | |
| 191 | done | |
| 192 | ||
| 193 | lemma fixB: "fix(f) = f(fix(f))" | |
| 194 | apply (unfold fix_def) | |
| 195 | apply (rule applyB [THEN ssubst], rule refl) | |
| 196 | done | |
| 197 | ||
| 198 | lemma letrecB: | |
| 58977 | 199 | "letrec g x be h(x,g) in g(a) = h(a,\<lambda>y. letrec g x be h(x,g) in g(y))" | 
| 20140 | 200 | apply (unfold letrec_def) | 
| 201 | apply (rule fixB [THEN ssubst], rule applyB [THEN ssubst], rule refl) | |
| 202 | done | |
| 203 | ||
| 204 | lemmas rawBs = caseBs applyB applyBbot | |
| 205 | ||
| 60770 | 206 | method_setup beta_rl = \<open> | 
| 32153 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
32010diff
changeset | 207 | Scan.succeed (fn ctxt => | 
| 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
32010diff
changeset | 208 | SIMPLE_METHOD' (CHANGED o | 
| 59498 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 wenzelm parents: 
58977diff
changeset | 209 |       simp_tac (ctxt addsimps @{thms rawBs} setloop (fn _ => stac ctxt @{thm letrecB}))))
 | 
| 60770 | 210 | \<close> | 
| 20140 | 211 | |
| 32153 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
32010diff
changeset | 212 | lemma ifBtrue: "if true then t else u = t" | 
| 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
32010diff
changeset | 213 | and ifBfalse: "if false then t else u = u" | 
| 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
32010diff
changeset | 214 | and ifBbot: "if bot then t else u = bot" | 
| 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
32010diff
changeset | 215 | unfolding data_defs by beta_rl+ | 
| 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
32010diff
changeset | 216 | |
| 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
32010diff
changeset | 217 | lemma whenBinl: "when(inl(a),t,u) = t(a)" | 
| 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
32010diff
changeset | 218 | and whenBinr: "when(inr(a),t,u) = u(a)" | 
| 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
32010diff
changeset | 219 | and whenBbot: "when(bot,t,u) = bot" | 
| 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
32010diff
changeset | 220 | unfolding data_defs by beta_rl+ | 
| 20140 | 221 | |
| 32153 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
32010diff
changeset | 222 | lemma splitB: "split(<a,b>,h) = h(a,b)" | 
| 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
32010diff
changeset | 223 | and splitBbot: "split(bot,h) = bot" | 
| 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
32010diff
changeset | 224 | unfolding data_defs by beta_rl+ | 
| 17456 | 225 | |
| 32153 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
32010diff
changeset | 226 | lemma fstB: "fst(<a,b>) = a" | 
| 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
32010diff
changeset | 227 | and fstBbot: "fst(bot) = bot" | 
| 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
32010diff
changeset | 228 | unfolding data_defs by beta_rl+ | 
| 20140 | 229 | |
| 32153 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
32010diff
changeset | 230 | lemma sndB: "snd(<a,b>) = b" | 
| 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
32010diff
changeset | 231 | and sndBbot: "snd(bot) = bot" | 
| 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
32010diff
changeset | 232 | unfolding data_defs by beta_rl+ | 
| 20140 | 233 | |
| 32153 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
32010diff
changeset | 234 | lemma thdB: "thd(<a,<b,c>>) = c" | 
| 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
32010diff
changeset | 235 | and thdBbot: "thd(bot) = bot" | 
| 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
32010diff
changeset | 236 | unfolding data_defs by beta_rl+ | 
| 20140 | 237 | |
| 32153 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
32010diff
changeset | 238 | lemma ncaseBzero: "ncase(zero,t,u) = t" | 
| 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
32010diff
changeset | 239 | and ncaseBsucc: "ncase(succ(n),t,u) = u(n)" | 
| 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
32010diff
changeset | 240 | and ncaseBbot: "ncase(bot,t,u) = bot" | 
| 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
32010diff
changeset | 241 | unfolding data_defs by beta_rl+ | 
| 20140 | 242 | |
| 32153 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
32010diff
changeset | 243 | lemma nrecBzero: "nrec(zero,t,u) = t" | 
| 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
32010diff
changeset | 244 | and nrecBsucc: "nrec(succ(n),t,u) = u(n,nrec(n,t,u))" | 
| 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
32010diff
changeset | 245 | and nrecBbot: "nrec(bot,t,u) = bot" | 
| 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
32010diff
changeset | 246 | unfolding data_defs by beta_rl+ | 
| 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
32010diff
changeset | 247 | |
| 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
32010diff
changeset | 248 | lemma lcaseBnil: "lcase([],t,u) = t" | 
| 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
32010diff
changeset | 249 | and lcaseBcons: "lcase(x$xs,t,u) = u(x,xs)" | 
| 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
32010diff
changeset | 250 | and lcaseBbot: "lcase(bot,t,u) = bot" | 
| 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
32010diff
changeset | 251 | unfolding data_defs by beta_rl+ | 
| 20140 | 252 | |
| 32153 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
32010diff
changeset | 253 | lemma lrecBnil: "lrec([],t,u) = t" | 
| 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
32010diff
changeset | 254 | and lrecBcons: "lrec(x$xs,t,u) = u(x,xs,lrec(xs,t,u))" | 
| 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
32010diff
changeset | 255 | and lrecBbot: "lrec(bot,t,u) = bot" | 
| 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
32010diff
changeset | 256 | unfolding data_defs by beta_rl+ | 
| 20140 | 257 | |
| 32153 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
32010diff
changeset | 258 | lemma letrec2B: | 
| 58977 | 259 | "letrec g x y be h(x,y,g) in g(p,q) = h(p,q,\<lambda>u v. letrec g x y be h(x,y,g) in g(u,v))" | 
| 32153 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
32010diff
changeset | 260 | unfolding data_defs letrec2_def by beta_rl+ | 
| 20140 | 261 | |
| 32153 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
32010diff
changeset | 262 | lemma letrec3B: | 
| 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
32010diff
changeset | 263 | "letrec g x y z be h(x,y,z,g) in g(p,q,r) = | 
| 58977 | 264 | h(p,q,r,\<lambda>u v w. letrec g x y z be h(x,y,z,g) in g(u,v,w))" | 
| 32153 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
32010diff
changeset | 265 | unfolding data_defs letrec3_def by beta_rl+ | 
| 20140 | 266 | |
| 32153 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
32010diff
changeset | 267 | lemma napplyBzero: "f^zero`a = a" | 
| 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
32010diff
changeset | 268 | and napplyBsucc: "f^succ(n)`a = f(f^n`a)" | 
| 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
32010diff
changeset | 269 | unfolding data_defs by beta_rl+ | 
| 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
32010diff
changeset | 270 | |
| 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
32010diff
changeset | 271 | lemmas termBs = letB applyB applyBbot splitB splitBbot fstB fstBbot | 
| 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
32010diff
changeset | 272 | sndB sndBbot thdB thdBbot ifBtrue ifBfalse ifBbot whenBinl whenBinr | 
| 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
32010diff
changeset | 273 | whenBbot ncaseBzero ncaseBsucc ncaseBbot nrecBzero nrecBsucc | 
| 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
32010diff
changeset | 274 | nrecBbot lcaseBnil lcaseBcons lcaseBbot lrecBnil lrecBcons lrecBbot | 
| 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
32010diff
changeset | 275 | napplyBzero napplyBsucc | 
| 20140 | 276 | |
| 277 | ||
| 60770 | 278 | subsection \<open>Constructors are injective\<close> | 
| 20140 | 279 | |
| 32154 | 280 | lemma term_injs: | 
| 58977 | 281 | "(inl(a) = inl(a')) \<longleftrightarrow> (a=a')" | 
| 282 | "(inr(a) = inr(a')) \<longleftrightarrow> (a=a')" | |
| 283 | "(succ(a) = succ(a')) \<longleftrightarrow> (a=a')" | |
| 284 | "(a$b = a'$b') \<longleftrightarrow> (a=a' \<and> b=b')" | |
| 32154 | 285 | by (inj_rl applyB splitB whenBinl whenBinr ncaseBsucc lcaseBcons) | 
| 20140 | 286 | |
| 287 | ||
| 60770 | 288 | subsection \<open>Constructors are distinct\<close> | 
| 20140 | 289 | |
| 60770 | 290 | ML \<open> | 
| 56199 | 291 | ML_Thms.bind_thms ("term_dstncts",
 | 
| 51670 | 292 |   mkall_dstnct_thms @{context} @{thms data_defs} (@{thms ccl_injs} @ @{thms term_injs})
 | 
| 24825 | 293 | [["bot","inl","inr"], ["bot","zero","succ"], ["bot","nil","cons"]]); | 
| 60770 | 294 | \<close> | 
| 20140 | 295 | |
| 296 | ||
| 62020 | 297 | subsection \<open>Rules for pre-order \<open>[=\<close>\<close> | 
| 20140 | 298 | |
| 32154 | 299 | lemma term_porews: | 
| 58977 | 300 | "inl(a) [= inl(a') \<longleftrightarrow> a [= a'" | 
| 301 | "inr(b) [= inr(b') \<longleftrightarrow> b [= b'" | |
| 302 | "succ(n) [= succ(n') \<longleftrightarrow> n [= n'" | |
| 303 | "x$xs [= x'$xs' \<longleftrightarrow> x [= x' \<and> xs [= xs'" | |
| 32154 | 304 | by (simp_all add: data_defs ccl_porews) | 
| 20140 | 305 | |
| 306 | ||
| 60770 | 307 | subsection \<open>Rewriting and Proving\<close> | 
| 20140 | 308 | |
| 60770 | 309 | ML \<open> | 
| 56199 | 310 |   ML_Thms.bind_thms ("term_injDs", XH_to_Ds @{thms term_injs});
 | 
| 60770 | 311 | \<close> | 
| 20140 | 312 | |
| 20917 | 313 | lemmas term_rews = termBs term_injs term_dstncts ccl_porews term_porews | 
| 314 | ||
| 20140 | 315 | lemmas [simp] = term_rews | 
| 20917 | 316 | lemmas [elim!] = term_dstncts [THEN notE] | 
| 317 | lemmas [dest!] = term_injDs | |
| 20140 | 318 | |
| 319 | end |