| author | wenzelm | 
| Fri, 21 May 1999 16:26:06 +0200 | |
| changeset 6699 | 1471f2bd74a0 | 
| parent 6112 | 5e4871c5136b | 
| child 7499 | 23e090051cb8 | 
| permissions | -rw-r--r-- | 
| 1461 | 1  | 
(* Title: Substitution.ML  | 
| 1048 | 2  | 
ID: $Id$  | 
| 1461 | 3  | 
Author: Ole Rasmussen  | 
| 1048 | 4  | 
Copyright 1995 University of Cambridge  | 
5  | 
Logic Image: ZF  | 
|
6  | 
*)  | 
|
7  | 
||
8  | 
(* ------------------------------------------------------------------------- *)  | 
|
9  | 
(* Arithmetic extensions *)  | 
|
10  | 
(* ------------------------------------------------------------------------- *)  | 
|
11  | 
||
12  | 
goal Arith.thy  | 
|
| 6068 | 13  | 
"!!m.[| p < n; n:nat|] ==> n~=p";  | 
| 2469 | 14  | 
by (Fast_tac 1);  | 
| 
3734
 
33f355f56f82
Much tidying including "qed" instead of result(), and even qed_spec_mp,
 
paulson 
parents: 
2469 
diff
changeset
 | 
15  | 
qed "gt_not_eq";  | 
| 1048 | 16  | 
|
17  | 
val succ_pred = prove_goal Arith.thy  | 
|
| 6068 | 18  | 
"!!i.[|j:nat; i:nat|] ==> i < j --> succ(j #- 1) = j"  | 
| 6070 | 19  | 
(fn prems =>[(induct_tac "j" 1),  | 
| 2469 | 20  | 
(Fast_tac 1),  | 
21  | 
(Asm_simp_tac 1)]);  | 
|
| 1048 | 22  | 
|
23  | 
goal Arith.thy  | 
|
| 6068 | 24  | 
"!!i.[|succ(x)<n; n:nat; x:nat|] ==> x < n #- 1 ";  | 
| 1461 | 25  | 
by (rtac succ_leE 1);  | 
| 1048 | 26  | 
by (forward_tac [nat_into_Ord RS le_refl RS lt_trans] 1 THEN assume_tac 1);  | 
| 4091 | 27  | 
by (asm_simp_tac (simpset() addsimps [succ_pred]) 1);  | 
| 
3734
 
33f355f56f82
Much tidying including "qed" instead of result(), and even qed_spec_mp,
 
paulson 
parents: 
2469 
diff
changeset
 | 
28  | 
qed "lt_pred";  | 
| 1048 | 29  | 
|
30  | 
goal Arith.thy  | 
|
| 6068 | 31  | 
"!!i.[|n < succ(x); p<n; p:nat; n:nat; x:nat|] ==> n #- 1 < x ";  | 
| 1461 | 32  | 
by (rtac succ_leE 1);  | 
| 4091 | 33  | 
by (asm_simp_tac (simpset() addsimps [succ_pred]) 1);  | 
| 
3734
 
33f355f56f82
Much tidying including "qed" instead of result(), and even qed_spec_mp,
 
paulson 
parents: 
2469 
diff
changeset
 | 
34  | 
qed "gt_pred";  | 
| 1048 | 35  | 
|
36  | 
||
| 2469 | 37  | 
Addsimps [nat_into_Ord, not_lt_iff_le, if_P, if_not_P];  | 
| 1048 | 38  | 
|
39  | 
||
40  | 
(* ------------------------------------------------------------------------- *)  | 
|
41  | 
(* lift_rec equality rules *)  | 
|
42  | 
(* ------------------------------------------------------------------------- *)  | 
|
| 6068 | 43  | 
Goal "[|n:nat; i:nat|] \  | 
44  | 
\ ==> lift_rec(Var(i),n) = (if i<n then Var(i) else Var(succ(i)))";  | 
|
| 6046 | 45  | 
by (asm_simp_tac (simpset() addsimps [lift_rec_def]) 1);  | 
| 
3734
 
33f355f56f82
Much tidying including "qed" instead of result(), and even qed_spec_mp,
 
paulson 
parents: 
2469 
diff
changeset
 | 
46  | 
qed "lift_rec_Var";  | 
| 1048 | 47  | 
|
| 6068 | 48  | 
Goal "[|n:nat; i:nat; k:nat; k le i|] ==> lift_rec(Var(i),k) = Var(succ(i))";  | 
| 6046 | 49  | 
by (asm_simp_tac (simpset() addsimps [lift_rec_def]) 1);  | 
| 
3734
 
33f355f56f82
Much tidying including "qed" instead of result(), and even qed_spec_mp,
 
paulson 
parents: 
2469 
diff
changeset
 | 
50  | 
qed "lift_rec_le";  | 
| 1048 | 51  | 
|
| 6068 | 52  | 
Goal "[|i:nat; k:nat; i<k |] ==> lift_rec(Var(i),k) = Var(i)";  | 
| 6046 | 53  | 
by (asm_simp_tac (simpset() addsimps [lift_rec_def]) 1);  | 
| 
3734
 
33f355f56f82
Much tidying including "qed" instead of result(), and even qed_spec_mp,
 
paulson 
parents: 
2469 
diff
changeset
 | 
54  | 
qed "lift_rec_gt";  | 
| 1048 | 55  | 
|
| 6068 | 56  | 
Goal "[|n:nat; k:nat|] ==> \  | 
| 1048 | 57  | 
\ lift_rec(Fun(t),k) = Fun(lift_rec(t,succ(k)))";  | 
| 6046 | 58  | 
by (asm_simp_tac (simpset() addsimps [lift_rec_def]) 1);  | 
| 
3734
 
33f355f56f82
Much tidying including "qed" instead of result(), and even qed_spec_mp,
 
paulson 
parents: 
2469 
diff
changeset
 | 
59  | 
qed "lift_rec_Fun";  | 
| 1048 | 60  | 
|
| 6068 | 61  | 
Goal "[|n:nat; k:nat|] ==> \  | 
| 1048 | 62  | 
\ lift_rec(App(b,f,a),k) = App(b,lift_rec(f,k),lift_rec(a,k))";  | 
| 6046 | 63  | 
by (asm_simp_tac (simpset() addsimps [lift_rec_def]) 1);  | 
| 
3734
 
33f355f56f82
Much tidying including "qed" instead of result(), and even qed_spec_mp,
 
paulson 
parents: 
2469 
diff
changeset
 | 
64  | 
qed "lift_rec_App";  | 
| 1048 | 65  | 
|
| 6046 | 66  | 
|
| 1048 | 67  | 
(* ------------------------------------------------------------------------- *)  | 
68  | 
(* substitution quality rules *)  | 
|
69  | 
(* ------------------------------------------------------------------------- *)  | 
|
70  | 
||
| 6068 | 71  | 
Goal "[|i:nat; k:nat; u:redexes|] \  | 
72  | 
\ ==> subst_rec(u,Var(i),k) = \  | 
|
73  | 
\ (if k<i then Var(i #- 1) else if k=i then u else Var(i))";  | 
|
| 6046 | 74  | 
by (asm_simp_tac (simpset() addsimps [subst_rec_def, gt_not_eq, leI]) 1);  | 
| 
3734
 
33f355f56f82
Much tidying including "qed" instead of result(), and even qed_spec_mp,
 
paulson 
parents: 
2469 
diff
changeset
 | 
75  | 
qed "subst_Var";  | 
| 1048 | 76  | 
|
| 6068 | 77  | 
Goal "[|n:nat; u:redexes|] ==> subst_rec(u,Var(n),n) = u";  | 
| 6046 | 78  | 
by (asm_simp_tac (simpset() addsimps [subst_rec_def]) 1);  | 
| 
3734
 
33f355f56f82
Much tidying including "qed" instead of result(), and even qed_spec_mp,
 
paulson 
parents: 
2469 
diff
changeset
 | 
79  | 
qed "subst_eq";  | 
| 1048 | 80  | 
|
| 6068 | 81  | 
Goal "[|n:nat; u:redexes; p:nat; p<n|] ==> \  | 
| 6046 | 82  | 
\ subst_rec(u,Var(n),p) = Var(n #- 1)";  | 
83  | 
by (asm_simp_tac (simpset() addsimps [subst_rec_def]) 1);  | 
|
| 
3734
 
33f355f56f82
Much tidying including "qed" instead of result(), and even qed_spec_mp,
 
paulson 
parents: 
2469 
diff
changeset
 | 
84  | 
qed "subst_gt";  | 
| 1048 | 85  | 
|
| 6068 | 86  | 
Goal "[|n:nat; u:redexes; p:nat; n<p|] ==> \  | 
| 6046 | 87  | 
\ subst_rec(u,Var(n),p) = Var(n)";  | 
88  | 
by (asm_simp_tac (simpset() addsimps [subst_rec_def, gt_not_eq, leI]) 1);  | 
|
| 
3734
 
33f355f56f82
Much tidying including "qed" instead of result(), and even qed_spec_mp,
 
paulson 
parents: 
2469 
diff
changeset
 | 
89  | 
qed "subst_lt";  | 
| 1048 | 90  | 
|
| 6068 | 91  | 
Goal "[|p:nat; u:redexes|] ==> \  | 
| 6046 | 92  | 
\ subst_rec(u,Fun(t),p) = Fun(subst_rec(lift(u),t,succ(p))) ";  | 
93  | 
by (asm_simp_tac (simpset() addsimps [subst_rec_def]) 1);  | 
|
| 
3734
 
33f355f56f82
Much tidying including "qed" instead of result(), and even qed_spec_mp,
 
paulson 
parents: 
2469 
diff
changeset
 | 
94  | 
qed "subst_Fun";  | 
| 1048 | 95  | 
|
| 6068 | 96  | 
Goal "[|p:nat; u:redexes|] ==> \  | 
| 6046 | 97  | 
\ subst_rec(u,App(b,f,a),p) = App(b,subst_rec(u,f,p),subst_rec(u,a,p))";  | 
98  | 
by (asm_simp_tac (simpset() addsimps [subst_rec_def]) 1);  | 
|
| 
3734
 
33f355f56f82
Much tidying including "qed" instead of result(), and even qed_spec_mp,
 
paulson 
parents: 
2469 
diff
changeset
 | 
99  | 
qed "subst_App";  | 
| 1048 | 100  | 
|
| 6046 | 101  | 
(*But not lift_rec_Var, subst_Var: too many case splits*)  | 
| 5527 | 102  | 
Addsimps [subst_Fun, subst_App];  | 
| 1048 | 103  | 
|
| 6046 | 104  | 
|
| 5268 | 105  | 
Goal "u:redexes ==> ALL k:nat. lift_rec(u,k):redexes";  | 
| 5527 | 106  | 
by (etac redexes.induct 1);  | 
107  | 
by (ALLGOALS  | 
|
108  | 
(asm_simp_tac (simpset() addsimps [lift_rec_Var,  | 
|
| 6046 | 109  | 
lift_rec_Fun, lift_rec_App])));  | 
| 6112 | 110  | 
qed_spec_mp "lift_rec_type";  | 
| 1048 | 111  | 
|
| 
5147
 
825877190618
More tidying and removal of "\!\!... from Goal commands
 
paulson 
parents: 
5116 
diff
changeset
 | 
112  | 
Goal "v:redexes ==> ALL n:nat. ALL u:redexes. subst_rec(u,v,n):redexes";  | 
| 5527 | 113  | 
by (etac redexes.induct 1);  | 
114  | 
by (ALLGOALS(asm_simp_tac  | 
|
115  | 
(simpset() addsimps [subst_Var, lift_rec_type])));  | 
|
| 6112 | 116  | 
qed_spec_mp "subst_type";  | 
| 1048 | 117  | 
|
118  | 
||
| 5527 | 119  | 
Addsimps [subst_eq, subst_gt, subst_lt, subst_type,  | 
| 2469 | 120  | 
lift_rec_le, lift_rec_gt, lift_rec_Fun, lift_rec_App,  | 
121  | 
lift_rec_type];  | 
|
| 1048 | 122  | 
|
123  | 
||
124  | 
(* ------------------------------------------------------------------------- *)  | 
|
125  | 
(* lift and substitution proofs *)  | 
|
126  | 
(* ------------------------------------------------------------------------- *)  | 
|
127  | 
||
| 
5147
 
825877190618
More tidying and removal of "\!\!... from Goal commands
 
paulson 
parents: 
5116 
diff
changeset
 | 
128  | 
Goal "u:redexes ==> ALL n:nat. ALL i:nat. i le n --> \  | 
| 1048 | 129  | 
\ (lift_rec(lift_rec(u,i),succ(n)) = lift_rec(lift_rec(u,n),i))";  | 
| 6046 | 130  | 
by (etac redexes.induct 1);  | 
131  | 
by (ALLGOALS Asm_simp_tac);  | 
|
| 
3734
 
33f355f56f82
Much tidying including "qed" instead of result(), and even qed_spec_mp,
 
paulson 
parents: 
2469 
diff
changeset
 | 
132  | 
by Safe_tac;  | 
| 5527 | 133  | 
by (case_tac "n < xa" 1);  | 
134  | 
by ((forward_tac [lt_trans2] 1) THEN (assume_tac 1));  | 
|
135  | 
by (ALLGOALS (asm_full_simp_tac (simpset() addsimps [lift_rec_Var, leI])));  | 
|
| 
3734
 
33f355f56f82
Much tidying including "qed" instead of result(), and even qed_spec_mp,
 
paulson 
parents: 
2469 
diff
changeset
 | 
136  | 
qed "lift_lift_rec";  | 
| 1048 | 137  | 
|
| 6068 | 138  | 
Goal "[|u:redexes; n:nat|] ==> \  | 
| 1048 | 139  | 
\ lift_rec(lift(u),succ(n)) = lift(lift_rec(u,n))";  | 
| 4091 | 140  | 
by (asm_simp_tac (simpset() addsimps [lift_lift_rec]) 1);  | 
| 
3734
 
33f355f56f82
Much tidying including "qed" instead of result(), and even qed_spec_mp,
 
paulson 
parents: 
2469 
diff
changeset
 | 
141  | 
qed "lift_lift";  | 
| 1048 | 142  | 
|
| 
5147
 
825877190618
More tidying and removal of "\!\!... from Goal commands
 
paulson 
parents: 
5116 
diff
changeset
 | 
143  | 
Goal "v:redexes ==> \  | 
| 3840 | 144  | 
\ ALL n:nat. ALL m:nat. ALL u:redexes. n le m-->\  | 
| 1048 | 145  | 
\ lift_rec(subst_rec(u,v,n),m) = \  | 
146  | 
\ subst_rec(lift_rec(u,m),lift_rec(v,succ(m)),n)";  | 
|
| 5527 | 147  | 
by ((etac redexes.induct 1)  | 
| 4091 | 148  | 
THEN (ALLGOALS(asm_simp_tac (simpset() addsimps [lift_lift]))));  | 
| 
3734
 
33f355f56f82
Much tidying including "qed" instead of result(), and even qed_spec_mp,
 
paulson 
parents: 
2469 
diff
changeset
 | 
149  | 
by Safe_tac;  | 
| 
5147
 
825877190618
More tidying and removal of "\!\!... from Goal commands
 
paulson 
parents: 
5116 
diff
changeset
 | 
150  | 
by (excluded_middle_tac "n < x" 1);  | 
| 5527 | 151  | 
by (Asm_full_simp_tac 1);  | 
| 
5147
 
825877190618
More tidying and removal of "\!\!... from Goal commands
 
paulson 
parents: 
5116 
diff
changeset
 | 
152  | 
by (eres_inst_tac [("j","n")] leE 1);
 | 
| 5527 | 153  | 
by (asm_simp_tac (simpset() setloop (split_inside_tac [split_if])  | 
154  | 
addsimps [lift_rec_Var,subst_Var,  | 
|
155  | 
leI,gt_pred,succ_pred]) 1);  | 
|
| 1048 | 156  | 
by (hyp_subst_tac 1);  | 
| 5527 | 157  | 
by (Asm_simp_tac 1);  | 
| 1048 | 158  | 
by (forw_inst_tac [("j","x")] lt_trans2 1);
 | 
159  | 
by (assume_tac 1);  | 
|
| 5527 | 160  | 
by (asm_simp_tac (simpset() addsimps [leI]) 1);  | 
| 
3734
 
33f355f56f82
Much tidying including "qed" instead of result(), and even qed_spec_mp,
 
paulson 
parents: 
2469 
diff
changeset
 | 
161  | 
qed "lift_rec_subst_rec";  | 
| 1048 | 162  | 
|
| 6068 | 163  | 
Goal "[|v:redexes; u:redexes; n:nat|] ==> \  | 
| 1048 | 164  | 
\ lift_rec(u/v,n) = lift_rec(u,n)/lift_rec(v,succ(n))";  | 
| 5527 | 165  | 
by (asm_simp_tac (simpset() addsimps [lift_rec_subst_rec]) 1);  | 
| 
3734
 
33f355f56f82
Much tidying including "qed" instead of result(), and even qed_spec_mp,
 
paulson 
parents: 
2469 
diff
changeset
 | 
166  | 
qed "lift_subst";  | 
| 1048 | 167  | 
|
168  | 
||
| 
5147
 
825877190618
More tidying and removal of "\!\!... from Goal commands
 
paulson 
parents: 
5116 
diff
changeset
 | 
169  | 
Goal "v:redexes ==> \  | 
| 3840 | 170  | 
\ ALL n:nat. ALL m:nat. ALL u:redexes. m le n-->\  | 
| 1048 | 171  | 
\ lift_rec(subst_rec(u,v,n),m) = \  | 
172  | 
\ subst_rec(lift_rec(u,m),lift_rec(v,m),succ(n))";  | 
|
| 5527 | 173  | 
by (etac redexes.induct 1  | 
174  | 
THEN ALLGOALS(asm_simp_tac (simpset() addsimps [lift_lift])));  | 
|
| 
3734
 
33f355f56f82
Much tidying including "qed" instead of result(), and even qed_spec_mp,
 
paulson 
parents: 
2469 
diff
changeset
 | 
175  | 
by Safe_tac;  | 
| 
5147
 
825877190618
More tidying and removal of "\!\!... from Goal commands
 
paulson 
parents: 
5116 
diff
changeset
 | 
176  | 
by (excluded_middle_tac "n < x" 1);  | 
| 5527 | 177  | 
by (Asm_full_simp_tac 1);  | 
| 1048 | 178  | 
by (eres_inst_tac [("i","x")] leE 1);
 | 
179  | 
by (forward_tac [lt_trans1] 1 THEN assume_tac 1);  | 
|
| 5527 | 180  | 
by (ALLGOALS(asm_simp_tac  | 
| 4091 | 181  | 
(simpset() addsimps [succ_pred,leI,gt_pred])));  | 
182  | 
by (asm_full_simp_tac (simpset() addsimps [leI]) 1);  | 
|
| 5527 | 183  | 
by (case_tac "n < xa" 1);  | 
184  | 
by (Asm_full_simp_tac 2);  | 
|
185  | 
by (asm_simp_tac (simpset() addsimps [leI]) 1);  | 
|
| 
3734
 
33f355f56f82
Much tidying including "qed" instead of result(), and even qed_spec_mp,
 
paulson 
parents: 
2469 
diff
changeset
 | 
186  | 
qed "lift_rec_subst_rec_lt";  | 
| 1048 | 187  | 
|
188  | 
||
| 
5147
 
825877190618
More tidying and removal of "\!\!... from Goal commands
 
paulson 
parents: 
5116 
diff
changeset
 | 
189  | 
Goal "u:redexes ==> \  | 
| 3840 | 190  | 
\ ALL n:nat. ALL v:redexes. subst_rec(v,lift_rec(u,n),n) = u";  | 
| 5527 | 191  | 
by (etac redexes.induct 1);  | 
192  | 
by Auto_tac;  | 
|
193  | 
by (case_tac "n < x" 1);  | 
|
194  | 
by Auto_tac;  | 
|
| 
3734
 
33f355f56f82
Much tidying including "qed" instead of result(), and even qed_spec_mp,
 
paulson 
parents: 
2469 
diff
changeset
 | 
195  | 
qed "subst_rec_lift_rec";  | 
| 1048 | 196  | 
|
| 5268 | 197  | 
Goal "v:redexes ==> \  | 
| 3840 | 198  | 
\ ALL m:nat. ALL n:nat. ALL u:redexes. ALL w:redexes. m le n --> \  | 
| 1048 | 199  | 
\ subst_rec(subst_rec(w,u,n),subst_rec(lift_rec(w,m),v,succ(n)),m)=\  | 
200  | 
\ subst_rec(w,subst_rec(u,v,m),n)";  | 
|
| 5527 | 201  | 
by (etac redexes.induct 1);  | 
202  | 
by (ALLGOALS(asm_simp_tac (simpset() addsimps  | 
|
203  | 
[lift_lift RS sym,lift_rec_subst_rec_lt])));  | 
|
| 
3734
 
33f355f56f82
Much tidying including "qed" instead of result(), and even qed_spec_mp,
 
paulson 
parents: 
2469 
diff
changeset
 | 
204  | 
by Safe_tac;  | 
| 
5147
 
825877190618
More tidying and removal of "\!\!... from Goal commands
 
paulson 
parents: 
5116 
diff
changeset
 | 
205  | 
by (excluded_middle_tac "n le succ(xa)" 1);  | 
| 5527 | 206  | 
by (Asm_full_simp_tac 1);  | 
| 1048 | 207  | 
by (forward_tac [nat_into_Ord RS le_refl RS lt_trans] 1 THEN assume_tac 1);  | 
| 1461 | 208  | 
by (etac leE 1);  | 
| 4091 | 209  | 
by (asm_simp_tac (simpset() addsimps [succ_pred,lt_pred]) 2);  | 
| 1048 | 210  | 
by (forward_tac [succ_leI RS lt_trans] 1 THEN assume_tac 1);  | 
211  | 
by (forw_inst_tac [("i","x")] 
 | 
|
212  | 
(nat_into_Ord RS le_refl RS lt_trans) 1 THEN assume_tac 1);  | 
|
| 4091 | 213  | 
by (asm_simp_tac (simpset() addsimps [succ_pred,lt_pred]) 1);  | 
| 
5147
 
825877190618
More tidying and removal of "\!\!... from Goal commands
 
paulson 
parents: 
5116 
diff
changeset
 | 
214  | 
by (eres_inst_tac [("i","n")] leE 1);
 | 
| 5527 | 215  | 
by (asm_simp_tac (simpset() addsimps [succ_pred,subst_rec_lift_rec,leI]) 2);  | 
| 
5147
 
825877190618
More tidying and removal of "\!\!... from Goal commands
 
paulson 
parents: 
5116 
diff
changeset
 | 
216  | 
by (excluded_middle_tac "n < x" 1);  | 
| 5527 | 217  | 
by (Asm_full_simp_tac 1);  | 
| 
5147
 
825877190618
More tidying and removal of "\!\!... from Goal commands
 
paulson 
parents: 
5116 
diff
changeset
 | 
218  | 
by (eres_inst_tac [("j","n")] leE 1);
 | 
| 4091 | 219  | 
by (asm_simp_tac (simpset() addsimps [gt_pred]) 1);  | 
220  | 
by (asm_simp_tac (simpset() addsimps [subst_rec_lift_rec]) 1);  | 
|
| 1048 | 221  | 
by (forward_tac [lt_trans2] 1 THEN assume_tac 1);  | 
| 4091 | 222  | 
by (asm_simp_tac (simpset() addsimps [gt_pred]) 1);  | 
| 
3734
 
33f355f56f82
Much tidying including "qed" instead of result(), and even qed_spec_mp,
 
paulson 
parents: 
2469 
diff
changeset
 | 
223  | 
qed "subst_rec_subst_rec";  | 
| 1048 | 224  | 
|
225  | 
||
| 6068 | 226  | 
Goal "[|v:redexes; u:redexes; w:redexes; n:nat|] ==> \  | 
| 1048 | 227  | 
\ subst_rec(w,u,n)/subst_rec(lift(w),v,succ(n)) = subst_rec(w,u/v,n)";  | 
| 4091 | 228  | 
by (asm_simp_tac (simpset() addsimps [subst_rec_subst_rec]) 1);  | 
| 
3734
 
33f355f56f82
Much tidying including "qed" instead of result(), and even qed_spec_mp,
 
paulson 
parents: 
2469 
diff
changeset
 | 
229  | 
qed "substitution";  | 
| 1048 | 230  | 
|
231  | 
(* ------------------------------------------------------------------------- *)  | 
|
232  | 
(* Preservation lemmas *)  | 
|
233  | 
(* Substitution preserves comp and regular *)  | 
|
234  | 
(* ------------------------------------------------------------------------- *)  | 
|
235  | 
||
236  | 
||
| 6068 | 237  | 
Goal "[|n:nat; u ~ v|] ==> ALL m:nat. lift_rec(u,m) ~ lift_rec(v,m)";  | 
| 1732 | 238  | 
by (etac Scomp.induct 1);  | 
| 4091 | 239  | 
by (ALLGOALS(asm_simp_tac (simpset() addsimps [comp_refl])));  | 
| 
3734
 
33f355f56f82
Much tidying including "qed" instead of result(), and even qed_spec_mp,
 
paulson 
parents: 
2469 
diff
changeset
 | 
240  | 
qed "lift_rec_preserve_comp";  | 
| 1048 | 241  | 
|
| 
5147
 
825877190618
More tidying and removal of "\!\!... from Goal commands
 
paulson 
parents: 
5116 
diff
changeset
 | 
242  | 
Goal "u2 ~ v2 ==> ALL m:nat. ALL u1:redexes. ALL v1:redexes.\  | 
| 1048 | 243  | 
\ u1 ~ v1--> subst_rec(u1,u2,m) ~ subst_rec(v1,v2,m)";  | 
| 1732 | 244  | 
by (etac Scomp.induct 1);  | 
| 5527 | 245  | 
by (ALLGOALS  | 
246  | 
(asm_simp_tac  | 
|
247  | 
(simpset() addsimps [subst_Var, lift_rec_preserve_comp, comp_refl])));  | 
|
| 
3734
 
33f355f56f82
Much tidying including "qed" instead of result(), and even qed_spec_mp,
 
paulson 
parents: 
2469 
diff
changeset
 | 
248  | 
qed "subst_rec_preserve_comp";  | 
| 1048 | 249  | 
|
| 
5147
 
825877190618
More tidying and removal of "\!\!... from Goal commands
 
paulson 
parents: 
5116 
diff
changeset
 | 
250  | 
Goal "regular(u) ==> ALL m:nat. regular(lift_rec(u,m))";  | 
| 5527 | 251  | 
by (etac Sreg.induct 1);  | 
252  | 
by (ALLGOALS (asm_full_simp_tac (simpset() addsimps [lift_rec_Var])));  | 
|
| 
3734
 
33f355f56f82
Much tidying including "qed" instead of result(), and even qed_spec_mp,
 
paulson 
parents: 
2469 
diff
changeset
 | 
253  | 
qed "lift_rec_preserve_reg";  | 
| 1048 | 254  | 
|
| 
5147
 
825877190618
More tidying and removal of "\!\!... from Goal commands
 
paulson 
parents: 
5116 
diff
changeset
 | 
255  | 
Goal "regular(v) ==> \  | 
| 3840 | 256  | 
\ ALL m:nat. ALL u:redexes. regular(u)-->regular(subst_rec(u,v,m))";  | 
| 5527 | 257  | 
by (etac Sreg.induct 1);  | 
258  | 
by (ALLGOALS(asm_full_simp_tac (simpset() addsimps [subst_Var,  | 
|
259  | 
lift_rec_preserve_reg])));  | 
|
| 
3734
 
33f355f56f82
Much tidying including "qed" instead of result(), and even qed_spec_mp,
 
paulson 
parents: 
2469 
diff
changeset
 | 
260  | 
qed "subst_rec_preserve_reg";  |