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