author | oheimb |
Thu, 12 Mar 1998 13:17:13 +0100 | |
changeset 4743 | b3bfcbd9fb93 |
parent 4686 | 74a12e86b20b |
child 4832 | bc11b5b06f87 |
permissions | -rw-r--r-- |
1465 | 1 |
(* Title: HOL/Lex/AutoChopper.ML |
1344 | 2 |
ID: $Id$ |
1465 | 3 |
Author: Richard Mayr & Tobias Nipkow |
1344 | 4 |
Copyright 1995 TUM |
5 |
||
6 |
Main result: auto_chopper satisfies the is_auto_chopper specification. |
|
7 |
*) |
|
8 |
||
1950
97f1c6bf3ace
Miniscoping rules are deleted, as these brittle proofs
paulson
parents:
1894
diff
changeset
|
9 |
Delsimps (ex_simps @ all_simps); |
97f1c6bf3ace
Miniscoping rules are deleted, as these brittle proofs
paulson
parents:
1894
diff
changeset
|
10 |
|
1344 | 11 |
open AutoChopper; |
12 |
||
13 |
infix repeat_RS; |
|
14 |
fun th1 repeat_RS th2 = ((th1 RS th2) repeat_RS th2) handle _ => th1; |
|
15 |
||
16 |
Addsimps [Let_def]; |
|
17 |
||
18 |
goal AutoChopper.thy "!st us p y ys. acc xs st (ys@[y]) us p A ~= ([],zs)"; |
|
19 |
by (list.induct_tac "xs" 1); |
|
20 |
by (Simp_tac 1); |
|
4686 | 21 |
by (Asm_simp_tac 1); |
1344 | 22 |
val accept_not_Nil = result() repeat_RS spec; |
23 |
Addsimps [accept_not_Nil]; |
|
24 |
||
25 |
goal AutoChopper.thy |
|
26 |
"!st us. acc xs st [] us ([],ys) A = ([], zs) --> \ |
|
27 |
\ zs = ys & (!ys. ys ~= [] & ys<=xs --> ~fin A (nexts A st ys))"; |
|
28 |
by (list.induct_tac "xs" 1); |
|
4089 | 29 |
by (simp_tac (simpset() addcongs [conj_cong]) 1); |
4686 | 30 |
by (Simp_tac 1); |
1344 | 31 |
by (strip_tac 1); |
1465 | 32 |
by (rtac conjI 1); |
1894 | 33 |
by (Fast_tac 1); |
4089 | 34 |
by (simp_tac (simpset() addsimps [prefix_Cons] addcongs [conj_cong]) 1); |
1344 | 35 |
by (strip_tac 1); |
2056
93c093620c28
Removed commands made redundant by new one-point rules
paulson
parents:
1950
diff
changeset
|
36 |
by (REPEAT(eresolve_tac [conjE,exE] 1)); |
93c093620c28
Removed commands made redundant by new one-point rules
paulson
parents:
1950
diff
changeset
|
37 |
by (hyp_subst_tac 1); |
1344 | 38 |
by (Simp_tac 1); |
39 |
by (case_tac "zsa = []" 1); |
|
40 |
by (Asm_simp_tac 1); |
|
1894 | 41 |
by (Fast_tac 1); |
1344 | 42 |
bind_thm("no_acc", result() RS spec RS spec RS mp); |
43 |
||
44 |
||
3842 | 45 |
val [prem] = goal HOL.thy "? x. P(f(x)) ==> ? y. P(y)"; |
1344 | 46 |
by (cut_facts_tac [prem] 1); |
1894 | 47 |
by (Fast_tac 1); |
1344 | 48 |
val ex_special = result(); |
49 |
||
50 |
||
51 |
goal AutoChopper.thy |
|
52 |
"! r erk l rst st ys yss zs::'a list. \ |
|
53 |
\ acc xs st erk r (l,rst) A = (ys#yss, zs) --> \ |
|
2609 | 54 |
\ ys@concat(yss)@zs = (if acc_prefix A st xs then r@xs else erk@concat(l)@rst)"; |
1344 | 55 |
by (list.induct_tac "xs" 1); |
4686 | 56 |
by (simp_tac (simpset() addcongs [conj_cong]) 1); |
57 |
by (Asm_simp_tac 1); |
|
2056
93c093620c28
Removed commands made redundant by new one-point rules
paulson
parents:
1950
diff
changeset
|
58 |
by (res_inst_tac [("p","acc list (start A) [] [] ([],list) A")] PairE 1); |
93c093620c28
Removed commands made redundant by new one-point rules
paulson
parents:
1950
diff
changeset
|
59 |
by (rename_tac "vss lrst" 1); |
4686 | 60 |
by (Asm_simp_tac 1); |
1344 | 61 |
by (res_inst_tac[("xs","vss")] list_eq_cases 1); |
2056
93c093620c28
Removed commands made redundant by new one-point rules
paulson
parents:
1950
diff
changeset
|
62 |
by (hyp_subst_tac 1); |
93c093620c28
Removed commands made redundant by new one-point rules
paulson
parents:
1950
diff
changeset
|
63 |
by (Simp_tac 1); |
4089 | 64 |
by (fast_tac (claset() addSDs [no_acc]) 1); |
2056
93c093620c28
Removed commands made redundant by new one-point rules
paulson
parents:
1950
diff
changeset
|
65 |
by (hyp_subst_tac 1); |
4686 | 66 |
by (Asm_simp_tac 1); |
1344 | 67 |
val step2_a = (result() repeat_RS spec) RS mp; |
68 |
||
69 |
||
70 |
goal AutoChopper.thy |
|
71 |
"! st erk r l rest ys yss zs.\ |
|
72 |
\ acc xs st erk r (l,rest) A = (ys#yss, zs) --> \ |
|
73 |
\ (if acc_prefix A st xs \ |
|
74 |
\ then ys ~= [] \ |
|
75 |
\ else ys ~= [] | (erk=[] & (l,rest) = (ys#yss,zs)))"; |
|
4686 | 76 |
by (Simp_tac 1); |
1344 | 77 |
by (list.induct_tac "xs" 1); |
4686 | 78 |
by (simp_tac (simpset() addcongs [conj_cong]) 1); |
1894 | 79 |
by (Fast_tac 1); |
4686 | 80 |
by (asm_simp_tac (simpset() addcongs [conj_cong]) 1); |
2056
93c093620c28
Removed commands made redundant by new one-point rules
paulson
parents:
1950
diff
changeset
|
81 |
by (res_inst_tac [("p","acc list (start A) [] [] ([],list) A")] PairE 1); |
93c093620c28
Removed commands made redundant by new one-point rules
paulson
parents:
1950
diff
changeset
|
82 |
by (rename_tac "vss lrst" 1); |
93c093620c28
Removed commands made redundant by new one-point rules
paulson
parents:
1950
diff
changeset
|
83 |
by (Asm_simp_tac 1); |
1344 | 84 |
by (strip_tac 1); |
85 |
by (case_tac "acc_prefix A (next A st a) list" 1); |
|
2056
93c093620c28
Removed commands made redundant by new one-point rules
paulson
parents:
1950
diff
changeset
|
86 |
by (Asm_simp_tac 1); |
1344 | 87 |
by (subgoal_tac "r @ [a] ~= []" 1); |
1894 | 88 |
by (Fast_tac 1); |
1344 | 89 |
by (Simp_tac 1); |
90 |
val step2_b = (result() repeat_RS spec) RS mp; |
|
91 |
||
92 |
||
93 |
goal AutoChopper.thy |
|
94 |
"! st erk r l rest ys yss zs. \ |
|
95 |
\ acc xs st erk r (l,rest) A = (ys#yss, zs) --> \ |
|
96 |
\ (if acc_prefix A st xs \ |
|
97 |
\ then ? g. ys=r@g & fin A (nexts A st g) \ |
|
98 |
\ else (erk~=[] & erk=ys) | (erk=[] & (l,rest) = (ys#yss,zs)))"; |
|
4686 | 99 |
by (Simp_tac 1); |
1344 | 100 |
by (list.induct_tac "xs" 1); |
4686 | 101 |
by (simp_tac (simpset() addcongs [conj_cong]) 1); |
1894 | 102 |
by (Fast_tac 1); |
4686 | 103 |
by (asm_simp_tac (simpset() addcongs [conj_cong]) 1); |
1344 | 104 |
by (strip_tac 1); |
1465 | 105 |
by (rtac conjI 1); |
2056
93c093620c28
Removed commands made redundant by new one-point rules
paulson
parents:
1950
diff
changeset
|
106 |
by (res_inst_tac [("p","acc list (start A) [] [] ([],list) A")] PairE 1); |
93c093620c28
Removed commands made redundant by new one-point rules
paulson
parents:
1950
diff
changeset
|
107 |
by (rename_tac "vss lrst" 1); |
93c093620c28
Removed commands made redundant by new one-point rules
paulson
parents:
1950
diff
changeset
|
108 |
by (Asm_simp_tac 1); |
1344 | 109 |
by (case_tac "acc_prefix A (next A st a) list" 1); |
110 |
by (strip_tac 1); |
|
3842 | 111 |
by (res_inst_tac [("f","%k. a#k")] ex_special 1); |
1344 | 112 |
by (Simp_tac 1); |
3842 | 113 |
by (res_inst_tac [("t","%k. ys=r@a#k"),("s","%k. ys=(r@[a])@k")] subst 1); |
1344 | 114 |
by (Simp_tac 1); |
1894 | 115 |
by (Fast_tac 1); |
1344 | 116 |
by (strip_tac 1); |
117 |
by (res_inst_tac [("x","[a]")] exI 1); |
|
118 |
by (Asm_simp_tac 1); |
|
119 |
by (subgoal_tac "r @ [a] ~= []" 1); |
|
2056
93c093620c28
Removed commands made redundant by new one-point rules
paulson
parents:
1950
diff
changeset
|
120 |
by (rtac sym 1); |
1894 | 121 |
by (Fast_tac 1); |
1344 | 122 |
by (Simp_tac 1); |
123 |
by (strip_tac 1); |
|
3842 | 124 |
by (res_inst_tac [("f","%k. a#k")] ex_special 1); |
1344 | 125 |
by (Simp_tac 1); |
3842 | 126 |
by (res_inst_tac [("t","%k. ys=r@a#k"),("s","%k. ys=(r@[a])@k")] subst 1); |
1344 | 127 |
by (Simp_tac 1); |
1894 | 128 |
by (Fast_tac 1); |
1344 | 129 |
val step2_c = (result() repeat_RS spec) RS mp; |
130 |
||
131 |
||
132 |
goal AutoChopper.thy |
|
133 |
"! st erk r l rest ys yss zs. \ |
|
134 |
\ acc xs st erk r (l,rest) A = (ys#yss, zs) --> \ |
|
135 |
\ (if acc_prefix A st xs \ |
|
2609 | 136 |
\ then acc(concat(yss)@zs)(start A) [] [] ([],concat(yss)@zs) A = (yss,zs) \ |
1344 | 137 |
\ else (erk~=[] & (l,rest)=(yss,zs)) | (erk=[] & (l, rest)=(ys#yss,zs)))"; |
4686 | 138 |
by (Simp_tac 1); |
1344 | 139 |
by (list.induct_tac "xs" 1); |
4686 | 140 |
by (simp_tac (simpset() addcongs [conj_cong]) 1); |
1894 | 141 |
by (Fast_tac 1); |
4686 | 142 |
by (asm_simp_tac (simpset() addcongs [conj_cong]) 1); |
2056
93c093620c28
Removed commands made redundant by new one-point rules
paulson
parents:
1950
diff
changeset
|
143 |
by (res_inst_tac [("p","acc list (start A) [] [] ([],list) A")] PairE 1); |
93c093620c28
Removed commands made redundant by new one-point rules
paulson
parents:
1950
diff
changeset
|
144 |
by (rename_tac "vss lrst" 1); |
93c093620c28
Removed commands made redundant by new one-point rules
paulson
parents:
1950
diff
changeset
|
145 |
by (Asm_simp_tac 1); |
1344 | 146 |
by (strip_tac 1); |
147 |
by (case_tac "acc_prefix A (next A st a) list" 1); |
|
148 |
by (Asm_simp_tac 1); |
|
149 |
by (subgoal_tac "acc list (start A) [] [] ([],list) A = (yss,zs)" 1); |
|
150 |
by (Asm_simp_tac 2); |
|
151 |
by (subgoal_tac "r@[a] ~= []" 2); |
|
1894 | 152 |
by (Fast_tac 2); |
1344 | 153 |
by (Simp_tac 2); |
2609 | 154 |
by (subgoal_tac "concat(yss) @ zs = list" 1); |
2056
93c093620c28
Removed commands made redundant by new one-point rules
paulson
parents:
1950
diff
changeset
|
155 |
by (hyp_subst_tac 1); |
93c093620c28
Removed commands made redundant by new one-point rules
paulson
parents:
1950
diff
changeset
|
156 |
by (atac 1); |
1344 | 157 |
by (case_tac "yss = []" 1); |
158 |
by (Asm_simp_tac 1); |
|
159 |
by (hyp_subst_tac 1); |
|
4089 | 160 |
by (fast_tac (claset() addSDs [no_acc]) 1); |
1465 | 161 |
by (etac ((neq_Nil_conv RS iffD1) RS exE) 1); |
162 |
by (etac exE 1); |
|
1344 | 163 |
by (hyp_subst_tac 1); |
164 |
by (Simp_tac 1); |
|
1465 | 165 |
by (rtac trans 1); |
2056
93c093620c28
Removed commands made redundant by new one-point rules
paulson
parents:
1950
diff
changeset
|
166 |
by (etac step2_a 1); |
4686 | 167 |
by (Simp_tac 1); |
1344 | 168 |
val step2_d = (result() repeat_RS spec) RS mp; |
169 |
||
170 |
Delsimps [split_paired_All]; |
|
171 |
goal AutoChopper.thy |
|
172 |
"! st erk r p ys yss zs. \ |
|
173 |
\ acc xs st erk r p A = (ys#yss, zs) --> \ |
|
174 |
\ (if acc_prefix A st xs \ |
|
3842 | 175 |
\ then ? g. ys=r@g & (!as. as<=xs & g<=as & g~=as --> ~fin A (nexts A st as))\ |
1344 | 176 |
\ else (erk~=[] & ys=erk) | (erk=[] & (ys#yss,zs)=p))"; |
4686 | 177 |
by (Simp_tac 1); |
1344 | 178 |
by (list.induct_tac "xs" 1); |
4686 | 179 |
by (simp_tac (simpset() addcongs [conj_cong]) 1); |
1894 | 180 |
by (Fast_tac 1); |
4686 | 181 |
by (asm_simp_tac (simpset() addcongs [conj_cong]) 1); |
1344 | 182 |
by (strip_tac 1); |
183 |
by (case_tac "acc_prefix A (next A st a) list" 1); |
|
2056
93c093620c28
Removed commands made redundant by new one-point rules
paulson
parents:
1950
diff
changeset
|
184 |
by (rtac conjI 1); |
1344 | 185 |
by (strip_tac 1); |
3842 | 186 |
by (res_inst_tac [("f","%k. a#k")] ex_special 1); |
187 |
by (res_inst_tac [("t","%k. ys=r@a#k"),("s","%k. ys=(r@[a])@k")] subst 1); |
|
1344 | 188 |
by (Simp_tac 1); |
3842 | 189 |
by (res_inst_tac [("P","%k. ys = (r@[a])@k & (!as. as<=list & k<=as & k ~= as --> ~ fin A (nexts A (next A st a) as))")] exE 1); |
1344 | 190 |
by (asm_simp_tac HOL_ss 1); |
191 |
by (res_inst_tac [("x","x")] exI 1); |
|
192 |
by (Asm_simp_tac 1); |
|
2056
93c093620c28
Removed commands made redundant by new one-point rules
paulson
parents:
1950
diff
changeset
|
193 |
by (rtac list_cases 1); |
1344 | 194 |
by (Simp_tac 1); |
4089 | 195 |
by (asm_simp_tac (simpset() addcongs[conj_cong]) 1); |
1344 | 196 |
by (strip_tac 1); |
3842 | 197 |
by (res_inst_tac [("f","%k. a#k")] ex_special 1); |
198 |
by (res_inst_tac [("t","%k. ys=r@a#k"),("s","%k. ys=(r@[a])@k")] subst 1); |
|
1344 | 199 |
by (Simp_tac 1); |
3842 | 200 |
by (res_inst_tac [("P","%k. ys=(r@[a])@k & (!as. as<=list & k<=as & k~=as --> ~ fin A (nexts A (next A st a) as))")] exE 1); |
1344 | 201 |
by (asm_simp_tac HOL_ss 1); |
202 |
by (res_inst_tac [("x","x")] exI 1); |
|
203 |
by (Asm_simp_tac 1); |
|
2056
93c093620c28
Removed commands made redundant by new one-point rules
paulson
parents:
1950
diff
changeset
|
204 |
by (rtac list_cases 1); |
1344 | 205 |
by (Simp_tac 1); |
4089 | 206 |
by (asm_simp_tac (simpset() addcongs[conj_cong]) 1); |
1344 | 207 |
by (Asm_simp_tac 1); |
208 |
by (strip_tac 1); |
|
209 |
by (res_inst_tac [("x","[a]")] exI 1); |
|
1465 | 210 |
by (rtac conjI 1); |
1344 | 211 |
by (subgoal_tac "r @ [a] ~= []" 1); |
1894 | 212 |
by (Fast_tac 1); |
1344 | 213 |
by (Simp_tac 1); |
1465 | 214 |
by (rtac list_cases 1); |
1344 | 215 |
by (Simp_tac 1); |
4089 | 216 |
by (asm_full_simp_tac (simpset() addsimps [acc_prefix_def] addcongs[conj_cong]) 1); |
1465 | 217 |
by (etac thin_rl 1); (* speed up *) |
1894 | 218 |
by (Fast_tac 1); |
1344 | 219 |
val step2_e = (result() repeat_RS spec) RS mp; |
220 |
Addsimps[split_paired_All]; |
|
221 |
||
222 |
goalw AutoChopper.thy [accepts_def, is_auto_chopper_def, auto_chopper_def, |
|
223 |
Chopper.is_longest_prefix_chopper_def] |
|
224 |
"is_auto_chopper(auto_chopper)"; |
|
225 |
by (REPEAT(ares_tac [no_acc,allI,impI,conjI] 1)); |
|
2056
93c093620c28
Removed commands made redundant by new one-point rules
paulson
parents:
1950
diff
changeset
|
226 |
by (rtac mp 1); |
93c093620c28
Removed commands made redundant by new one-point rules
paulson
parents:
1950
diff
changeset
|
227 |
by (etac step2_b 2); |
4686 | 228 |
by (Simp_tac 1); |
1465 | 229 |
by (rtac conjI 1); |
2056
93c093620c28
Removed commands made redundant by new one-point rules
paulson
parents:
1950
diff
changeset
|
230 |
by (rtac mp 1); |
93c093620c28
Removed commands made redundant by new one-point rules
paulson
parents:
1950
diff
changeset
|
231 |
by (etac step2_c 2); |
4686 | 232 |
by (Simp_tac 1); |
1465 | 233 |
by (rtac conjI 1); |
4686 | 234 |
by (asm_simp_tac (simpset() addsimps [step2_a]) 1); |
1465 | 235 |
by (rtac conjI 1); |
2056
93c093620c28
Removed commands made redundant by new one-point rules
paulson
parents:
1950
diff
changeset
|
236 |
by (rtac mp 1); |
93c093620c28
Removed commands made redundant by new one-point rules
paulson
parents:
1950
diff
changeset
|
237 |
by (etac step2_d 2); |
4686 | 238 |
by (Simp_tac 1); |
1465 | 239 |
by (rtac mp 1); |
2056
93c093620c28
Removed commands made redundant by new one-point rules
paulson
parents:
1950
diff
changeset
|
240 |
by (etac step2_e 2); |
4686 | 241 |
by (Simp_tac 1); |
1344 | 242 |
qed"auto_chopper_is_auto_chopper"; |