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