| author | wenzelm | 
| Tue, 07 Mar 2023 12:40:10 +0100 | |
| changeset 77560 | 5749ee7c45a0 | 
| parent 74448 | 2fd74a2c4e1c | 
| child 80754 | 701912f5645a | 
| 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 | |
| 74443 
dbf68dbacaff
clarified 'let' syntax: avoid conflict with existing 'let' in FOL;
 wenzelm parents: 
74441diff
changeset | 48 | definition "let1" :: "[i,i\<Rightarrow>i]\<Rightarrow>i" | 
| 
dbf68dbacaff
clarified 'let' syntax: avoid conflict with existing 'let' in FOL;
 wenzelm parents: 
74441diff
changeset | 49 | where let_def: "let1(t, f) == case(t,f(true),f(false), \<lambda>x y. f(<x,y>), \<lambda>u. f(lam x. u(x)))" | 
| 62143 | 50 | |
| 74443 
dbf68dbacaff
clarified 'let' syntax: avoid conflict with existing 'let' in FOL;
 wenzelm parents: 
74441diff
changeset | 51 | syntax "_let1" :: "[idt,i,i]\<Rightarrow>i"  ("(3let _ be _/ in _)" [0,0,60] 60)
 | 
| 
dbf68dbacaff
clarified 'let' syntax: avoid conflict with existing 'let' in FOL;
 wenzelm parents: 
74441diff
changeset | 52 | translations "let x be a in e" == "CONST let1(a, \<lambda>x. e)" | 
| 998 
91d09e262799
Gave tighter priorities to if, napply and the let-forms to
 lcp parents: 
610diff
changeset | 53 | |
| 62143 | 54 | definition letrec :: "[[i,i\<Rightarrow>i]\<Rightarrow>i,(i\<Rightarrow>i)\<Rightarrow>i]\<Rightarrow>i" | 
| 55 | where "letrec(h, b) == b(\<lambda>x. fix(\<lambda>f. lam x. h(x,\<lambda>y. f`y))`x)" | |
| 56 | ||
| 57 | definition letrec2 :: "[[i,i,i\<Rightarrow>i\<Rightarrow>i]\<Rightarrow>i,(i\<Rightarrow>i\<Rightarrow>i)\<Rightarrow>i]\<Rightarrow>i" | |
| 58 | where "letrec2 (h, f) == | |
| 59 | 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 | 60 | |
| 62143 | 61 | 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" | 
| 62 | where "letrec3 (h, f) == | |
| 63 | 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>>)))), | |
| 64 | \<lambda>g'. f(\<lambda>x y z. g'(<x,<y,z>>)))" | |
| 0 | 65 | |
| 62143 | 66 | syntax | 
| 74444 | 67 |   "_letrec" :: "[idt,idt,i,i]\<Rightarrow>i"  ("(3letrec _ _ be _/ in _)" [0,0,0,60] 60)
 | 
| 68 |   "_letrec2" :: "[idt,idt,idt,i,i]\<Rightarrow>i"  ("(3letrec _ _ _ be _/ in _)" [0,0,0,0,60] 60)
 | |
| 69 |   "_letrec3" :: "[idt,idt,idt,idt,i,i]\<Rightarrow>i"  ("(3letrec _ _ _ _ be _/ in _)" [0,0,0,0,0,60] 60)
 | |
| 70 | parse_translation \<open> | |
| 74441 | 71 | let | 
| 74444 | 72 | fun abs_tr t u = Syntax_Trans.abs_tr [t, u]; | 
| 73 | fun letrec_tr [f, x, a, b] = | |
| 74 | Syntax.const \<^const_syntax>\<open>letrec\<close> $ abs_tr x (abs_tr f a) $ abs_tr f b; | |
| 75 | fun letrec2_tr [f, x, y, a, b] = | |
| 76 | Syntax.const \<^const_syntax>\<open>letrec2\<close> $ abs_tr x (abs_tr y (abs_tr f a)) $ abs_tr f b; | |
| 77 | fun letrec3_tr [f, x, y, z, a, b] = | |
| 78 | Syntax.const \<^const_syntax>\<open>letrec3\<close> $ abs_tr x (abs_tr y (abs_tr z (abs_tr f a))) $ abs_tr f b; | |
| 79 | in | |
| 80 | [(\<^syntax_const>\<open>_letrec\<close>, K letrec_tr), | |
| 81 | (\<^syntax_const>\<open>_letrec2\<close>, K letrec2_tr), | |
| 82 | (\<^syntax_const>\<open>_letrec3\<close>, K letrec3_tr)] | |
| 83 | end | |
| 84 | \<close> | |
| 85 | print_translation \<open> | |
| 74441 | 86 | let | 
| 74444 | 87 | val bound = Syntax_Trans.mark_bound_abs; | 
| 88 | ||
| 89 | fun letrec_tr' [Abs(x,T,Abs(f,S,a)),Abs(ff,SS,b)] = | |
| 90 | let | |
| 74448 | 91 | val (f',b') = Syntax_Trans.print_abs(ff,SS,b) | 
| 92 | val (_,a'') = Syntax_Trans.print_abs(f,S,a) | |
| 93 | val (x',a') = Syntax_Trans.print_abs(x,T,a'') | |
| 74444 | 94 | in | 
| 95 | Syntax.const \<^syntax_const>\<open>_letrec\<close> $ bound(f',SS) $ bound(x',T) $ a' $ b' | |
| 96 | end; | |
| 74441 | 97 | |
| 74444 | 98 | fun letrec2_tr' [Abs(x,T,Abs(y,U,Abs(f,S,a))),Abs(ff,SS,b)] = | 
| 99 | let | |
| 74448 | 100 | val (f',b') = Syntax_Trans.print_abs(ff,SS,b) | 
| 101 | val ( _,a1) = Syntax_Trans.print_abs(f,S,a) | |
| 102 | val (y',a2) = Syntax_Trans.print_abs(y,U,a1) | |
| 103 | val (x',a') = Syntax_Trans.print_abs(x,T,a2) | |
| 74444 | 104 | in | 
| 105 | Syntax.const \<^syntax_const>\<open>_letrec2\<close> $ bound(f',SS) $ bound(x',T) $ bound(y',U) $ a' $ b' | |
| 106 | end; | |
| 107 | ||
| 108 | fun letrec3_tr' [Abs(x,T,Abs(y,U,Abs(z,V,Abs(f,S,a)))),Abs(ff,SS,b)] = | |
| 109 | let | |
| 74448 | 110 | val (f',b') = Syntax_Trans.print_abs(ff,SS,b) | 
| 111 | val ( _,a1) = Syntax_Trans.print_abs(f,S,a) | |
| 112 | val (z',a2) = Syntax_Trans.print_abs(z,V,a1) | |
| 113 | val (y',a3) = Syntax_Trans.print_abs(y,U,a2) | |
| 114 | val (x',a') = Syntax_Trans.print_abs(x,T,a3) | |
| 74444 | 115 | in | 
| 116 | Syntax.const \<^syntax_const>\<open>_letrec3\<close> $ | |
| 117 | bound(f',SS) $ bound(x',T) $ bound(y',U) $ bound(z',V) $ a' $ b' | |
| 118 | end; | |
| 74441 | 119 | in | 
| 74444 | 120 | [(\<^const_syntax>\<open>letrec\<close>, K letrec_tr'), | 
| 121 | (\<^const_syntax>\<open>letrec2\<close>, K letrec2_tr'), | |
| 122 | (\<^const_syntax>\<open>letrec3\<close>, K letrec3_tr')] | |
| 123 | end | |
| 60770 | 124 | \<close> | 
| 17456 | 125 | |
| 62143 | 126 | definition nrec :: "[i,i,[i,i]\<Rightarrow>i]\<Rightarrow>i" | 
| 127 | where "nrec(n,b,c) == letrec g x be ncase(x, b, \<lambda>y. c(y,g(y))) in g(n)" | |
| 128 | ||
| 129 | definition nil :: "i"  ("([])")
 | |
| 130 | where "[] == inl(one)" | |
| 17456 | 131 | |
| 62143 | 132 | definition cons :: "[i,i]\<Rightarrow>i" (infixr "$" 80) | 
| 133 | where "h$t == inr(<h,t>)" | |
| 134 | ||
| 135 | definition lcase :: "[i,i,[i,i]\<Rightarrow>i]\<Rightarrow>i" | |
| 136 | where "lcase(l,b,c) == when(l, \<lambda>x. b, \<lambda>y. split(y,c))" | |
| 17456 | 137 | |
| 62143 | 138 | definition lrec :: "[i,i,[i,i,i]\<Rightarrow>i]\<Rightarrow>i" | 
| 139 | where "lrec(l,b,c) == letrec g x be lcase(x, b, \<lambda>h t. c(h,t,g(t))) in g(l)" | |
| 17456 | 140 | |
| 62143 | 141 | definition napply :: "[i\<Rightarrow>i,i,i]\<Rightarrow>i"  ("(_ ^ _ ` _)" [56,56,56] 56)
 | 
| 142 | where "f ^n` a == nrec(n,a,\<lambda>x g. f(g))" | |
| 20140 | 143 | |
| 144 | lemmas simp_can_defs = one_def inl_def inr_def | |
| 145 | and simp_ncan_defs = if_def when_def split_def fst_def snd_def thd_def | |
| 146 | lemmas simp_defs = simp_can_defs simp_ncan_defs | |
| 147 | ||
| 148 | lemmas ind_can_defs = zero_def succ_def nil_def cons_def | |
| 149 | and ind_ncan_defs = ncase_def nrec_def lcase_def lrec_def | |
| 150 | lemmas ind_defs = ind_can_defs ind_ncan_defs | |
| 151 | ||
| 152 | lemmas data_defs = simp_defs ind_defs napply_def | |
| 153 | and genrec_defs = letrec_def letrec2_def letrec3_def | |
| 154 | ||
| 155 | ||
| 60770 | 156 | subsection \<open>Beta Rules, including strictness\<close> | 
| 20140 | 157 | |
| 58977 | 158 | lemma letB: "\<not> t=bot \<Longrightarrow> let x be t in f(x) = f(t)" | 
| 20140 | 159 | apply (unfold let_def) | 
| 160 | apply (erule rev_mp) | |
| 161 | apply (rule_tac t = "t" in term_case) | |
| 74445 | 162 | apply simp_all | 
| 20140 | 163 | done | 
| 164 | ||
| 165 | lemma letBabot: "let x be bot in f(x) = bot" | |
| 74445 | 166 | unfolding let_def by (rule caseBbot) | 
| 20140 | 167 | |
| 168 | lemma letBbbot: "let x be t in bot = bot" | |
| 169 | apply (unfold let_def) | |
| 170 | apply (rule_tac t = t in term_case) | |
| 171 | apply (rule caseBbot) | |
| 74445 | 172 | apply simp_all | 
| 20140 | 173 | done | 
| 174 | ||
| 175 | lemma applyB: "(lam x. b(x)) ` a = b(a)" | |
| 74445 | 176 | by (simp add: apply_def) | 
| 20140 | 177 | |
| 178 | lemma applyBbot: "bot ` a = bot" | |
| 74445 | 179 | unfolding apply_def by (rule caseBbot) | 
| 20140 | 180 | |
| 181 | lemma fixB: "fix(f) = f(fix(f))" | |
| 182 | apply (unfold fix_def) | |
| 183 | apply (rule applyB [THEN ssubst], rule refl) | |
| 184 | done | |
| 185 | ||
| 74445 | 186 | lemma letrecB: "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 | 187 | apply (unfold letrec_def) | 
| 188 | apply (rule fixB [THEN ssubst], rule applyB [THEN ssubst], rule refl) | |
| 189 | done | |
| 190 | ||
| 191 | lemmas rawBs = caseBs applyB applyBbot | |
| 192 | ||
| 60770 | 193 | method_setup beta_rl = \<open> | 
| 32153 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
32010diff
changeset | 194 | Scan.succeed (fn ctxt => | 
| 74445 | 195 | let val ctxt' = Context_Position.set_visible false ctxt in | 
| 196 | SIMPLE_METHOD' (CHANGED o | |
| 197 |         simp_tac (ctxt' addsimps @{thms rawBs} setloop (fn _ => stac ctxt @{thm letrecB})))
 | |
| 198 | end) | |
| 60770 | 199 | \<close> | 
| 20140 | 200 | |
| 32153 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
32010diff
changeset | 201 | lemma ifBtrue: "if true then t else u = t" | 
| 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
32010diff
changeset | 202 | and ifBfalse: "if false then t else u = u" | 
| 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
32010diff
changeset | 203 | and ifBbot: "if bot then t else u = bot" | 
| 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
32010diff
changeset | 204 | unfolding data_defs by beta_rl+ | 
| 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
32010diff
changeset | 205 | |
| 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
32010diff
changeset | 206 | lemma whenBinl: "when(inl(a),t,u) = t(a)" | 
| 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
32010diff
changeset | 207 | and whenBinr: "when(inr(a),t,u) = u(a)" | 
| 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
32010diff
changeset | 208 | and whenBbot: "when(bot,t,u) = bot" | 
| 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
32010diff
changeset | 209 | unfolding data_defs by beta_rl+ | 
| 20140 | 210 | |
| 32153 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
32010diff
changeset | 211 | lemma splitB: "split(<a,b>,h) = h(a,b)" | 
| 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
32010diff
changeset | 212 | and splitBbot: "split(bot,h) = bot" | 
| 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
32010diff
changeset | 213 | unfolding data_defs by beta_rl+ | 
| 17456 | 214 | |
| 32153 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
32010diff
changeset | 215 | lemma fstB: "fst(<a,b>) = a" | 
| 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
32010diff
changeset | 216 | and fstBbot: "fst(bot) = bot" | 
| 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
32010diff
changeset | 217 | unfolding data_defs by beta_rl+ | 
| 20140 | 218 | |
| 32153 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
32010diff
changeset | 219 | lemma sndB: "snd(<a,b>) = b" | 
| 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
32010diff
changeset | 220 | and sndBbot: "snd(bot) = bot" | 
| 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
32010diff
changeset | 221 | unfolding data_defs by beta_rl+ | 
| 20140 | 222 | |
| 32153 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
32010diff
changeset | 223 | lemma thdB: "thd(<a,<b,c>>) = c" | 
| 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
32010diff
changeset | 224 | and thdBbot: "thd(bot) = bot" | 
| 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
32010diff
changeset | 225 | unfolding data_defs by beta_rl+ | 
| 20140 | 226 | |
| 32153 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
32010diff
changeset | 227 | lemma ncaseBzero: "ncase(zero,t,u) = t" | 
| 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
32010diff
changeset | 228 | and ncaseBsucc: "ncase(succ(n),t,u) = u(n)" | 
| 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
32010diff
changeset | 229 | and ncaseBbot: "ncase(bot,t,u) = bot" | 
| 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
32010diff
changeset | 230 | unfolding data_defs by beta_rl+ | 
| 20140 | 231 | |
| 32153 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
32010diff
changeset | 232 | lemma nrecBzero: "nrec(zero,t,u) = t" | 
| 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
32010diff
changeset | 233 | 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 | 234 | and nrecBbot: "nrec(bot,t,u) = bot" | 
| 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
32010diff
changeset | 235 | unfolding data_defs by beta_rl+ | 
| 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
32010diff
changeset | 236 | |
| 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
32010diff
changeset | 237 | lemma lcaseBnil: "lcase([],t,u) = t" | 
| 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
32010diff
changeset | 238 | 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 | 239 | and lcaseBbot: "lcase(bot,t,u) = bot" | 
| 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
32010diff
changeset | 240 | unfolding data_defs by beta_rl+ | 
| 20140 | 241 | |
| 32153 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
32010diff
changeset | 242 | lemma lrecBnil: "lrec([],t,u) = t" | 
| 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
32010diff
changeset | 243 | 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 | 244 | and lrecBbot: "lrec(bot,t,u) = bot" | 
| 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
32010diff
changeset | 245 | unfolding data_defs by beta_rl+ | 
| 20140 | 246 | |
| 32153 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
32010diff
changeset | 247 | lemma letrec2B: | 
| 58977 | 248 | "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 | 249 | unfolding data_defs letrec2_def by beta_rl+ | 
| 20140 | 250 | |
| 32153 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
32010diff
changeset | 251 | lemma letrec3B: | 
| 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
32010diff
changeset | 252 | "letrec g x y z be h(x,y,z,g) in g(p,q,r) = | 
| 58977 | 253 | 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 | 254 | unfolding data_defs letrec3_def by beta_rl+ | 
| 20140 | 255 | |
| 32153 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
32010diff
changeset | 256 | lemma napplyBzero: "f^zero`a = a" | 
| 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
32010diff
changeset | 257 | 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 | 258 | unfolding data_defs by beta_rl+ | 
| 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
32010diff
changeset | 259 | |
| 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
32010diff
changeset | 260 | lemmas termBs = letB applyB applyBbot splitB splitBbot fstB fstBbot | 
| 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
32010diff
changeset | 261 | sndB sndBbot thdB thdBbot ifBtrue ifBfalse ifBbot whenBinl whenBinr | 
| 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
32010diff
changeset | 262 | whenBbot ncaseBzero ncaseBsucc ncaseBbot nrecBzero nrecBsucc | 
| 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
32010diff
changeset | 263 | nrecBbot lcaseBnil lcaseBcons lcaseBbot lrecBnil lrecBcons lrecBbot | 
| 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
32010diff
changeset | 264 | napplyBzero napplyBsucc | 
| 20140 | 265 | |
| 266 | ||
| 60770 | 267 | subsection \<open>Constructors are injective\<close> | 
| 20140 | 268 | |
| 32154 | 269 | lemma term_injs: | 
| 58977 | 270 | "(inl(a) = inl(a')) \<longleftrightarrow> (a=a')" | 
| 271 | "(inr(a) = inr(a')) \<longleftrightarrow> (a=a')" | |
| 272 | "(succ(a) = succ(a')) \<longleftrightarrow> (a=a')" | |
| 273 | "(a$b = a'$b') \<longleftrightarrow> (a=a' \<and> b=b')" | |
| 32154 | 274 | by (inj_rl applyB splitB whenBinl whenBinr ncaseBsucc lcaseBcons) | 
| 20140 | 275 | |
| 276 | ||
| 60770 | 277 | subsection \<open>Constructors are distinct\<close> | 
| 20140 | 278 | |
| 60770 | 279 | ML \<open> | 
| 56199 | 280 | ML_Thms.bind_thms ("term_dstncts",
 | 
| 69593 | 281 |   mkall_dstnct_thms \<^context> @{thms data_defs} (@{thms ccl_injs} @ @{thms term_injs})
 | 
| 24825 | 282 | [["bot","inl","inr"], ["bot","zero","succ"], ["bot","nil","cons"]]); | 
| 60770 | 283 | \<close> | 
| 20140 | 284 | |
| 285 | ||
| 62020 | 286 | subsection \<open>Rules for pre-order \<open>[=\<close>\<close> | 
| 20140 | 287 | |
| 32154 | 288 | lemma term_porews: | 
| 58977 | 289 | "inl(a) [= inl(a') \<longleftrightarrow> a [= a'" | 
| 290 | "inr(b) [= inr(b') \<longleftrightarrow> b [= b'" | |
| 291 | "succ(n) [= succ(n') \<longleftrightarrow> n [= n'" | |
| 292 | "x$xs [= x'$xs' \<longleftrightarrow> x [= x' \<and> xs [= xs'" | |
| 32154 | 293 | by (simp_all add: data_defs ccl_porews) | 
| 20140 | 294 | |
| 295 | ||
| 60770 | 296 | subsection \<open>Rewriting and Proving\<close> | 
| 20140 | 297 | |
| 60770 | 298 | ML \<open> | 
| 56199 | 299 |   ML_Thms.bind_thms ("term_injDs", XH_to_Ds @{thms term_injs});
 | 
| 60770 | 300 | \<close> | 
| 20140 | 301 | |
| 20917 | 302 | lemmas term_rews = termBs term_injs term_dstncts ccl_porews term_porews | 
| 303 | ||
| 20140 | 304 | lemmas [simp] = term_rews | 
| 20917 | 305 | lemmas [elim!] = term_dstncts [THEN notE] | 
| 306 | lemmas [dest!] = term_injDs | |
| 20140 | 307 | |
| 308 | end |