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