author | paulson |
Thu, 16 Oct 1997 15:23:53 +0200 | |
changeset 3904 | c0d56e4c823e |
parent 3896 | ee8ebb74ec00 |
child 3913 | 96e28b16861c |
permissions | -rw-r--r-- |
1465 | 1 |
(* Title: HOL/simpdata.ML |
923 | 2 |
ID: $Id$ |
1465 | 3 |
Author: Tobias Nipkow |
923 | 4 |
Copyright 1991 University of Cambridge |
5 |
||
6 |
Instantiation of the generic simplifier |
|
7 |
*) |
|
8 |
||
1984 | 9 |
section "Simplifier"; |
10 |
||
923 | 11 |
open Simplifier; |
12 |
||
1984 | 13 |
(*** Addition of rules to simpsets and clasets simultaneously ***) |
14 |
||
15 |
(*Takes UNCONDITIONAL theorems of the form A<->B to |
|
2031 | 16 |
the Safe Intr rule B==>A and |
17 |
the Safe Destruct rule A==>B. |
|
1984 | 18 |
Also ~A goes to the Safe Elim rule A ==> ?R |
19 |
Failing other cases, A is added as a Safe Intr rule*) |
|
20 |
local |
|
21 |
val iff_const = HOLogic.eq_const HOLogic.boolT; |
|
22 |
||
23 |
fun addIff th = |
|
24 |
(case HOLogic.dest_Trueprop (#prop(rep_thm th)) of |
|
2718 | 25 |
(Const("Not",_) $ A) => |
2031 | 26 |
AddSEs [zero_var_indexes (th RS notE)] |
27 |
| (con $ _ $ _) => |
|
28 |
if con=iff_const |
|
29 |
then (AddSIs [zero_var_indexes (th RS iffD2)]; |
|
30 |
AddSDs [zero_var_indexes (th RS iffD1)]) |
|
31 |
else AddSIs [th] |
|
32 |
| _ => AddSIs [th]; |
|
1984 | 33 |
Addsimps [th]) |
34 |
handle _ => error ("AddIffs: theorem must be unconditional\n" ^ |
|
2031 | 35 |
string_of_thm th) |
1984 | 36 |
|
37 |
fun delIff th = |
|
38 |
(case HOLogic.dest_Trueprop (#prop(rep_thm th)) of |
|
2718 | 39 |
(Const("Not",_) $ A) => |
2031 | 40 |
Delrules [zero_var_indexes (th RS notE)] |
41 |
| (con $ _ $ _) => |
|
42 |
if con=iff_const |
|
43 |
then Delrules [zero_var_indexes (th RS iffD2), |
|
3518 | 44 |
make_elim (zero_var_indexes (th RS iffD1))] |
2031 | 45 |
else Delrules [th] |
46 |
| _ => Delrules [th]; |
|
1984 | 47 |
Delsimps [th]) |
48 |
handle _ => warning("DelIffs: ignoring conditional theorem\n" ^ |
|
2031 | 49 |
string_of_thm th) |
1984 | 50 |
in |
51 |
val AddIffs = seq addIff |
|
52 |
val DelIffs = seq delIff |
|
53 |
end; |
|
54 |
||
55 |
||
923 | 56 |
local |
57 |
||
2935 | 58 |
fun prover s = prove_goal HOL.thy s (fn _ => [blast_tac HOL_cs 1]); |
923 | 59 |
|
1922 | 60 |
val P_imp_P_iff_True = prover "P --> (P = True)" RS mp; |
61 |
val P_imp_P_eq_True = P_imp_P_iff_True RS eq_reflection; |
|
923 | 62 |
|
1922 | 63 |
val not_P_imp_P_iff_F = prover "~P --> (P = False)" RS mp; |
64 |
val not_P_imp_P_eq_False = not_P_imp_P_iff_F RS eq_reflection; |
|
923 | 65 |
|
1922 | 66 |
fun atomize pairs = |
67 |
let fun atoms th = |
|
2031 | 68 |
(case concl_of th of |
69 |
Const("Trueprop",_) $ p => |
|
70 |
(case head_of p of |
|
71 |
Const(a,_) => |
|
72 |
(case assoc(pairs,a) of |
|
73 |
Some(rls) => flat (map atoms ([th] RL rls)) |
|
74 |
| None => [th]) |
|
75 |
| _ => [th]) |
|
76 |
| _ => [th]) |
|
1922 | 77 |
in atoms end; |
923 | 78 |
|
2134 | 79 |
fun gen_all th = forall_elim_vars (#maxidx(rep_thm th)+1) th; |
80 |
||
81 |
in |
|
82 |
||
3896
ee8ebb74ec00
Various new lemmas. Improved conversion of equations to rewrite rules:
nipkow
parents:
3842
diff
changeset
|
83 |
fun mk_meta_eq r = r RS eq_reflection; |
ee8ebb74ec00
Various new lemmas. Improved conversion of equations to rewrite rules:
nipkow
parents:
3842
diff
changeset
|
84 |
|
ee8ebb74ec00
Various new lemmas. Improved conversion of equations to rewrite rules:
nipkow
parents:
3842
diff
changeset
|
85 |
fun mk_meta_eq_simp r = case concl_of r of |
2031 | 86 |
Const("==",_)$_$_ => r |
3896
ee8ebb74ec00
Various new lemmas. Improved conversion of equations to rewrite rules:
nipkow
parents:
3842
diff
changeset
|
87 |
| _$(Const("op =",_)$lhs$rhs) => |
ee8ebb74ec00
Various new lemmas. Improved conversion of equations to rewrite rules:
nipkow
parents:
3842
diff
changeset
|
88 |
(case fst(Logic.loops (#sign(rep_thm r)) (prems_of r) lhs rhs) of |
ee8ebb74ec00
Various new lemmas. Improved conversion of equations to rewrite rules:
nipkow
parents:
3842
diff
changeset
|
89 |
None => mk_meta_eq r |
ee8ebb74ec00
Various new lemmas. Improved conversion of equations to rewrite rules:
nipkow
parents:
3842
diff
changeset
|
90 |
| Some _ => r RS P_imp_P_eq_True) |
2718 | 91 |
| _$(Const("Not",_)$_) => r RS not_P_imp_P_eq_False |
1922 | 92 |
| _ => r RS P_imp_P_eq_True; |
93 |
(* last 2 lines requires all formulae to be of the from Trueprop(.) *) |
|
923 | 94 |
|
2082 | 95 |
val simp_thms = map prover |
96 |
[ "(x=x) = True", |
|
97 |
"(~True) = False", "(~False) = True", "(~ ~ P) = P", |
|
98 |
"(~P) ~= P", "P ~= (~P)", "(P ~= Q) = (P = (~Q))", |
|
99 |
"(True=P) = P", "(P=True) = P", |
|
100 |
"(True --> P) = P", "(False --> P) = True", |
|
101 |
"(P --> True) = True", "(P --> P) = True", |
|
102 |
"(P --> False) = (~P)", "(P --> ~P) = (~P)", |
|
103 |
"(P & True) = P", "(True & P) = P", |
|
2800 | 104 |
"(P & False) = False", "(False & P) = False", |
105 |
"(P & P) = P", "(P & (P & Q)) = (P & Q)", |
|
2082 | 106 |
"(P | True) = True", "(True | P) = True", |
2800 | 107 |
"(P | False) = P", "(False | P) = P", |
108 |
"(P | P) = P", "(P | (P | Q)) = (P | Q)", |
|
2082 | 109 |
"((~P) = (~Q)) = (P=Q)", |
3842 | 110 |
"(!x. P) = P", "(? x. P) = P", "? x. x=t", "? x. t=x", |
3573 | 111 |
"(? x. x=t & P(x)) = P(t)", |
3568
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents:
3559
diff
changeset
|
112 |
"(! x. t=x --> P(x)) = P(t)" ]; |
923 | 113 |
|
988 | 114 |
(*Add congruence rules for = (instead of ==) *) |
2636
4b30dbe4a020
added delcongs, Delcongs, unsafe_solver, safe_solver, HOL_basic_ss,
oheimb
parents:
2595
diff
changeset
|
115 |
infix 4 addcongs delcongs; |
3559 | 116 |
fun ss addcongs congs = ss addeqcongs (map standard (congs RL [eq_reflection])); |
117 |
fun ss delcongs congs = ss deleqcongs (map standard (congs RL [eq_reflection])); |
|
923 | 118 |
|
1264 | 119 |
fun Addcongs congs = (simpset := !simpset addcongs congs); |
2636
4b30dbe4a020
added delcongs, Delcongs, unsafe_solver, safe_solver, HOL_basic_ss,
oheimb
parents:
2595
diff
changeset
|
120 |
fun Delcongs congs = (simpset := !simpset delcongs congs); |
1264 | 121 |
|
3896
ee8ebb74ec00
Various new lemmas. Improved conversion of equations to rewrite rules:
nipkow
parents:
3842
diff
changeset
|
122 |
fun mksimps pairs = map mk_meta_eq_simp o atomize pairs o gen_all; |
923 | 123 |
|
1922 | 124 |
val imp_cong = impI RSN |
125 |
(2, prove_goal HOL.thy "(P=P')--> (P'--> (Q=Q'))--> ((P-->Q) = (P'-->Q'))" |
|
2935 | 126 |
(fn _=> [blast_tac HOL_cs 1]) RS mp RS mp); |
1922 | 127 |
|
1948
78e5bfcbc1e9
Added miniscoping to the simplifier: quantifiers are now pushed in
paulson
parents:
1922
diff
changeset
|
128 |
(*Miniscoping: pushing in existential quantifiers*) |
78e5bfcbc1e9
Added miniscoping to the simplifier: quantifiers are now pushed in
paulson
parents:
1922
diff
changeset
|
129 |
val ex_simps = map prover |
3842 | 130 |
["(EX x. P x & Q) = ((EX x. P x) & Q)", |
131 |
"(EX x. P & Q x) = (P & (EX x. Q x))", |
|
132 |
"(EX x. P x | Q) = ((EX x. P x) | Q)", |
|
133 |
"(EX x. P | Q x) = (P | (EX x. Q x))", |
|
134 |
"(EX x. P x --> Q) = ((ALL x. P x) --> Q)", |
|
135 |
"(EX x. P --> Q x) = (P --> (EX x. Q x))"]; |
|
1948
78e5bfcbc1e9
Added miniscoping to the simplifier: quantifiers are now pushed in
paulson
parents:
1922
diff
changeset
|
136 |
|
78e5bfcbc1e9
Added miniscoping to the simplifier: quantifiers are now pushed in
paulson
parents:
1922
diff
changeset
|
137 |
(*Miniscoping: pushing in universal quantifiers*) |
78e5bfcbc1e9
Added miniscoping to the simplifier: quantifiers are now pushed in
paulson
parents:
1922
diff
changeset
|
138 |
val all_simps = map prover |
3842 | 139 |
["(ALL x. P x & Q) = ((ALL x. P x) & Q)", |
140 |
"(ALL x. P & Q x) = (P & (ALL x. Q x))", |
|
141 |
"(ALL x. P x | Q) = ((ALL x. P x) | Q)", |
|
142 |
"(ALL x. P | Q x) = (P | (ALL x. Q x))", |
|
143 |
"(ALL x. P x --> Q) = ((EX x. P x) --> Q)", |
|
144 |
"(ALL x. P --> Q x) = (P --> (ALL x. Q x))"]; |
|
1948
78e5bfcbc1e9
Added miniscoping to the simplifier: quantifiers are now pushed in
paulson
parents:
1922
diff
changeset
|
145 |
|
3568
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents:
3559
diff
changeset
|
146 |
(*** Simplification procedure for turning ? x. ... & x = t & ... |
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents:
3559
diff
changeset
|
147 |
into ? x. x = t & ... & ... |
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents:
3559
diff
changeset
|
148 |
where the latter can be rewritten via (? x. x = t & P(x)) = P(t) |
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents:
3559
diff
changeset
|
149 |
***) |
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents:
3559
diff
changeset
|
150 |
|
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents:
3559
diff
changeset
|
151 |
local |
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents:
3559
diff
changeset
|
152 |
|
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents:
3559
diff
changeset
|
153 |
fun def(eq as (c as Const("op =",_)) $ s $ t) = |
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents:
3559
diff
changeset
|
154 |
if s = Bound 0 andalso not(loose_bvar1(t,0)) then Some eq else |
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents:
3559
diff
changeset
|
155 |
if t = Bound 0 andalso not(loose_bvar1(s,0)) then Some(c$t$s) |
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents:
3559
diff
changeset
|
156 |
else None |
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents:
3559
diff
changeset
|
157 |
| def _ = None; |
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents:
3559
diff
changeset
|
158 |
|
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents:
3559
diff
changeset
|
159 |
fun extract(Const("op &",_) $ P $ Q) = |
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents:
3559
diff
changeset
|
160 |
(case def P of |
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents:
3559
diff
changeset
|
161 |
Some eq => Some(eq,Q) |
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents:
3559
diff
changeset
|
162 |
| None => (case def Q of |
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents:
3559
diff
changeset
|
163 |
Some eq => Some(eq,P) |
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents:
3559
diff
changeset
|
164 |
| None => |
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents:
3559
diff
changeset
|
165 |
(case extract P of |
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents:
3559
diff
changeset
|
166 |
Some(eq,P') => Some(eq, HOLogic.conj $ P' $ Q) |
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents:
3559
diff
changeset
|
167 |
| None => (case extract Q of |
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents:
3559
diff
changeset
|
168 |
Some(eq,Q') => Some(eq,HOLogic.conj $ P $ Q') |
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents:
3559
diff
changeset
|
169 |
| None => None)))) |
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents:
3559
diff
changeset
|
170 |
| extract _ = None; |
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents:
3559
diff
changeset
|
171 |
|
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents:
3559
diff
changeset
|
172 |
fun prove_eq(ceqt) = |
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents:
3559
diff
changeset
|
173 |
let val tac = rtac eq_reflection 1 THEN rtac iffI 1 THEN |
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents:
3559
diff
changeset
|
174 |
ALLGOALS(EVERY'[etac exE, REPEAT o (etac conjE), |
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents:
3559
diff
changeset
|
175 |
rtac exI, REPEAT o (ares_tac [conjI] ORELSE' etac sym)]) |
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents:
3559
diff
changeset
|
176 |
in rule_by_tactic tac (trivial ceqt) end; |
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents:
3559
diff
changeset
|
177 |
|
3577
9715b6e3ec5f
added prems argument to simplification procedures;
wenzelm
parents:
3573
diff
changeset
|
178 |
fun rearrange sg _ (F as ex $ Abs(x,T,P)) = |
3568
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents:
3559
diff
changeset
|
179 |
(case extract P of |
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents:
3559
diff
changeset
|
180 |
None => None |
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents:
3559
diff
changeset
|
181 |
| Some(eq,Q) => |
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents:
3559
diff
changeset
|
182 |
let val ceqt = cterm_of sg |
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents:
3559
diff
changeset
|
183 |
(Logic.mk_equals(F,ex $ Abs(x,T,HOLogic.conj$eq$Q))) |
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents:
3559
diff
changeset
|
184 |
in Some(prove_eq ceqt) end) |
3577
9715b6e3ec5f
added prems argument to simplification procedures;
wenzelm
parents:
3573
diff
changeset
|
185 |
| rearrange _ _ _ = None; |
3568
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents:
3559
diff
changeset
|
186 |
|
3842 | 187 |
val pattern = read_cterm (sign_of HOL.thy) ("? x. P(x) & Q(x)",HOLogic.boolT) |
3568
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents:
3559
diff
changeset
|
188 |
|
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents:
3559
diff
changeset
|
189 |
in |
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents:
3559
diff
changeset
|
190 |
val defEX_regroup = mk_simproc "defined EX" [pattern] rearrange; |
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents:
3559
diff
changeset
|
191 |
end; |
1722 | 192 |
|
923 | 193 |
|
2022 | 194 |
(* elimination of existential quantifiers in assumptions *) |
923 | 195 |
|
196 |
val ex_all_equiv = |
|
197 |
let val lemma1 = prove_goal HOL.thy |
|
198 |
"(? x. P(x) ==> PROP Q) ==> (!!x. P(x) ==> PROP Q)" |
|
199 |
(fn prems => [resolve_tac prems 1, etac exI 1]); |
|
200 |
val lemma2 = prove_goalw HOL.thy [Ex_def] |
|
201 |
"(!!x. P(x) ==> PROP Q) ==> (? x. P(x) ==> PROP Q)" |
|
202 |
(fn prems => [REPEAT(resolve_tac prems 1)]) |
|
203 |
in equal_intr lemma1 lemma2 end; |
|
204 |
||
205 |
end; |
|
206 |
||
3654 | 207 |
(* Elimination of True from asumptions: *) |
208 |
||
209 |
val True_implies_equals = prove_goal HOL.thy |
|
210 |
"(True ==> PROP P) == PROP P" |
|
211 |
(fn _ => [rtac equal_intr_rule 1, atac 2, |
|
212 |
METAHYPS (fn prems => resolve_tac prems 1) 1, |
|
213 |
rtac TrueI 1]); |
|
214 |
||
2935 | 215 |
fun prove nm thm = qed_goal nm HOL.thy thm (fn _ => [blast_tac HOL_cs 1]); |
923 | 216 |
|
217 |
prove "conj_commute" "(P&Q) = (Q&P)"; |
|
218 |
prove "conj_left_commute" "(P&(Q&R)) = (Q&(P&R))"; |
|
219 |
val conj_comms = [conj_commute, conj_left_commute]; |
|
2134 | 220 |
prove "conj_assoc" "((P&Q)&R) = (P&(Q&R))"; |
923 | 221 |
|
1922 | 222 |
prove "disj_commute" "(P|Q) = (Q|P)"; |
223 |
prove "disj_left_commute" "(P|(Q|R)) = (Q|(P|R))"; |
|
224 |
val disj_comms = [disj_commute, disj_left_commute]; |
|
2134 | 225 |
prove "disj_assoc" "((P|Q)|R) = (P|(Q|R))"; |
1922 | 226 |
|
923 | 227 |
prove "conj_disj_distribL" "(P&(Q|R)) = (P&Q | P&R)"; |
228 |
prove "conj_disj_distribR" "((P|Q)&R) = (P&R | Q&R)"; |
|
1485
240cc98b94a7
Added qed_spec_mp to avoid renaming of bound vars in 'th RS spec'
nipkow
parents:
1465
diff
changeset
|
229 |
|
1892 | 230 |
prove "disj_conj_distribL" "(P|(Q&R)) = ((P|Q) & (P|R))"; |
231 |
prove "disj_conj_distribR" "((P&Q)|R) = ((P|R) & (Q|R))"; |
|
232 |
||
2134 | 233 |
prove "imp_conjR" "(P --> (Q&R)) = ((P-->Q) & (P-->R))"; |
234 |
prove "imp_conjL" "((P&Q) -->R) = (P --> (Q --> R))"; |
|
235 |
prove "imp_disjL" "((P|Q) --> R) = ((P-->R)&(Q-->R))"; |
|
1892 | 236 |
|
3448 | 237 |
(*These two are specialized, but imp_disj_not1 is useful in Auth/Yahalom.ML*) |
238 |
prove "imp_disj_not1" "((P --> Q | R)) = (~Q --> P --> R)"; |
|
239 |
prove "imp_disj_not2" "((P --> Q | R)) = (~R --> P --> Q)"; |
|
240 |
||
3904 | 241 |
prove "imp_disj1" "((P-->Q)|R) = (P--> Q|R)"; |
242 |
prove "imp_disj2" "(Q|(P-->R)) = (P--> Q|R)"; |
|
243 |
||
1485
240cc98b94a7
Added qed_spec_mp to avoid renaming of bound vars in 'th RS spec'
nipkow
parents:
1465
diff
changeset
|
244 |
prove "de_Morgan_disj" "(~(P | Q)) = (~P & ~Q)"; |
240cc98b94a7
Added qed_spec_mp to avoid renaming of bound vars in 'th RS spec'
nipkow
parents:
1465
diff
changeset
|
245 |
prove "de_Morgan_conj" "(~(P & Q)) = (~P | ~Q)"; |
3446
a14e5451f613
Addition of not_imp (which pushes negation into implication) as a default
paulson
parents:
3282
diff
changeset
|
246 |
prove "not_imp" "(~(P --> Q)) = (P & ~Q)"; |
1922 | 247 |
prove "not_iff" "(P~=Q) = (P = (~Q))"; |
1485
240cc98b94a7
Added qed_spec_mp to avoid renaming of bound vars in 'th RS spec'
nipkow
parents:
1465
diff
changeset
|
248 |
|
2134 | 249 |
(*Avoids duplication of subgoals after expand_if, when the true and false |
250 |
cases boil down to the same thing.*) |
|
251 |
prove "cases_simp" "((P --> Q) & (~P --> Q)) = Q"; |
|
252 |
||
3842 | 253 |
prove "not_all" "(~ (! x. P(x))) = (? x.~P(x))"; |
1922 | 254 |
prove "imp_all" "((! x. P x) --> Q) = (? x. P x --> Q)"; |
3842 | 255 |
prove "not_ex" "(~ (? x. P(x))) = (! x.~P(x))"; |
1922 | 256 |
prove "imp_ex" "((? x. P x) --> Q) = (! x. P x --> Q)"; |
1660 | 257 |
|
1655 | 258 |
prove "ex_disj_distrib" "(? x. P(x) | Q(x)) = ((? x. P(x)) | (? x. Q(x)))"; |
259 |
prove "all_conj_distrib" "(!x. P(x) & Q(x)) = ((! x. P(x)) & (! x. Q(x)))"; |
|
260 |
||
2134 | 261 |
(* '&' congruence rule: not included by default! |
262 |
May slow rewrite proofs down by as much as 50% *) |
|
263 |
||
264 |
let val th = prove_goal HOL.thy |
|
265 |
"(P=P')--> (P'--> (Q=Q'))--> ((P&Q) = (P'&Q'))" |
|
2935 | 266 |
(fn _=> [blast_tac HOL_cs 1]) |
2134 | 267 |
in bind_thm("conj_cong",standard (impI RSN (2, th RS mp RS mp))) end; |
268 |
||
269 |
let val th = prove_goal HOL.thy |
|
270 |
"(Q=Q')--> (Q'--> (P=P'))--> ((P&Q) = (P'&Q'))" |
|
2935 | 271 |
(fn _=> [blast_tac HOL_cs 1]) |
2134 | 272 |
in bind_thm("rev_conj_cong",standard (impI RSN (2, th RS mp RS mp))) end; |
273 |
||
274 |
(* '|' congruence rule: not included by default! *) |
|
275 |
||
276 |
let val th = prove_goal HOL.thy |
|
277 |
"(P=P')--> (~P'--> (Q=Q'))--> ((P|Q) = (P'|Q'))" |
|
2935 | 278 |
(fn _=> [blast_tac HOL_cs 1]) |
2134 | 279 |
in bind_thm("disj_cong",standard (impI RSN (2, th RS mp RS mp))) end; |
280 |
||
281 |
prove "eq_sym_conv" "(x=y) = (y=x)"; |
|
282 |
||
283 |
qed_goalw "o_apply" HOL.thy [o_def] "(f o g) x = f (g x)" |
|
284 |
(fn _ => [rtac refl 1]); |
|
285 |
||
286 |
qed_goal "meta_eq_to_obj_eq" HOL.thy "x==y ==> x=y" |
|
287 |
(fn [prem] => [rewtac prem, rtac refl 1]); |
|
288 |
||
289 |
qed_goalw "if_True" HOL.thy [if_def] "(if True then x else y) = x" |
|
2935 | 290 |
(fn _=>[blast_tac (HOL_cs addIs [select_equality]) 1]); |
2134 | 291 |
|
292 |
qed_goalw "if_False" HOL.thy [if_def] "(if False then x else y) = y" |
|
2935 | 293 |
(fn _=>[blast_tac (HOL_cs addIs [select_equality]) 1]); |
2134 | 294 |
|
295 |
qed_goal "if_P" HOL.thy "P ==> (if P then x else y) = x" |
|
296 |
(fn [prem] => [ stac (prem RS eqTrueI) 1, rtac if_True 1 ]); |
|
297 |
(* |
|
298 |
qed_goal "if_not_P" HOL.thy "~P ==> (if P then x else y) = y" |
|
299 |
(fn [prem] => [ stac (prem RS not_P_imp_P_iff_F) 1, rtac if_False 1 ]); |
|
300 |
*) |
|
301 |
qed_goalw "if_not_P" HOL.thy [if_def] "!!P. ~P ==> (if P then x else y) = y" |
|
2935 | 302 |
(fn _ => [blast_tac (HOL_cs addIs [select_equality]) 1]); |
2134 | 303 |
|
304 |
qed_goal "expand_if" HOL.thy |
|
305 |
"P(if Q then x else y) = ((Q --> P(x)) & (~Q --> P(y)))" |
|
306 |
(fn _=> [ (res_inst_tac [("Q","Q")] (excluded_middle RS disjE) 1), |
|
307 |
stac if_P 2, |
|
308 |
stac if_not_P 1, |
|
2935 | 309 |
REPEAT(blast_tac HOL_cs 1) ]); |
2134 | 310 |
|
311 |
qed_goal "if_bool_eq" HOL.thy |
|
312 |
"(if P then Q else R) = ((P-->Q) & (~P-->R))" |
|
313 |
(fn _ => [rtac expand_if 1]); |
|
314 |
||
2263 | 315 |
local val mktac = mk_case_split_tac (meta_eq_to_obj_eq RS iffD2) |
316 |
in |
|
317 |
fun split_tac splits = mktac (map mk_meta_eq splits) |
|
318 |
end; |
|
319 |
||
320 |
local val mktac = mk_case_split_inside_tac (meta_eq_to_obj_eq RS iffD2) |
|
321 |
in |
|
322 |
fun split_inside_tac splits = mktac (map mk_meta_eq splits) |
|
323 |
end; |
|
324 |
||
325 |
||
2251 | 326 |
qed_goal "if_cancel" HOL.thy "(if c then x else x) = x" |
2935 | 327 |
(fn _ => [split_tac [expand_if] 1, blast_tac HOL_cs 1]); |
2251 | 328 |
|
2134 | 329 |
(** 'if' congruence rules: neither included by default! *) |
330 |
||
331 |
(*Simplifies x assuming c and y assuming ~c*) |
|
332 |
qed_goal "if_cong" HOL.thy |
|
333 |
"[| b=c; c ==> x=u; ~c ==> y=v |] ==>\ |
|
334 |
\ (if b then x else y) = (if c then u else v)" |
|
335 |
(fn rew::prems => |
|
336 |
[stac rew 1, stac expand_if 1, stac expand_if 1, |
|
2935 | 337 |
blast_tac (HOL_cs addDs prems) 1]); |
2134 | 338 |
|
339 |
(*Prevents simplification of x and y: much faster*) |
|
340 |
qed_goal "if_weak_cong" HOL.thy |
|
341 |
"b=c ==> (if b then x else y) = (if c then x else y)" |
|
342 |
(fn [prem] => [rtac (prem RS arg_cong) 1]); |
|
343 |
||
344 |
(*Prevents simplification of t: much faster*) |
|
345 |
qed_goal "let_weak_cong" HOL.thy |
|
346 |
"a = b ==> (let x=a in t(x)) = (let x=b in t(x))" |
|
347 |
(fn [prem] => [rtac (prem RS arg_cong) 1]); |
|
348 |
||
349 |
(*In general it seems wrong to add distributive laws by default: they |
|
350 |
might cause exponential blow-up. But imp_disjL has been in for a while |
|
351 |
and cannot be removed without affecting existing proofs. Moreover, |
|
352 |
rewriting by "(P|Q --> R) = ((P-->R)&(Q-->R))" might be justified on the |
|
353 |
grounds that it allows simplification of R in the two cases.*) |
|
354 |
||
355 |
val mksimps_pairs = |
|
356 |
[("op -->", [mp]), ("op &", [conjunct1,conjunct2]), |
|
357 |
("All", [spec]), ("True", []), ("False", []), |
|
358 |
("If", [if_bool_eq RS iffD1])]; |
|
1758 | 359 |
|
2636
4b30dbe4a020
added delcongs, Delcongs, unsafe_solver, safe_solver, HOL_basic_ss,
oheimb
parents:
2595
diff
changeset
|
360 |
fun unsafe_solver prems = FIRST'[resolve_tac (TrueI::refl::prems), |
4b30dbe4a020
added delcongs, Delcongs, unsafe_solver, safe_solver, HOL_basic_ss,
oheimb
parents:
2595
diff
changeset
|
361 |
atac, etac FalseE]; |
4b30dbe4a020
added delcongs, Delcongs, unsafe_solver, safe_solver, HOL_basic_ss,
oheimb
parents:
2595
diff
changeset
|
362 |
(*No premature instantiation of variables during simplification*) |
4b30dbe4a020
added delcongs, Delcongs, unsafe_solver, safe_solver, HOL_basic_ss,
oheimb
parents:
2595
diff
changeset
|
363 |
fun safe_solver prems = FIRST'[match_tac (TrueI::refl::prems), |
4b30dbe4a020
added delcongs, Delcongs, unsafe_solver, safe_solver, HOL_basic_ss,
oheimb
parents:
2595
diff
changeset
|
364 |
eq_assume_tac, ematch_tac [FalseE]]; |
2443
a81d4c219c3c
factored out HOL_base_ss and val HOL_min_ss, added HOL_safe_min_ss
oheimb
parents:
2263
diff
changeset
|
365 |
|
2636
4b30dbe4a020
added delcongs, Delcongs, unsafe_solver, safe_solver, HOL_basic_ss,
oheimb
parents:
2595
diff
changeset
|
366 |
val HOL_basic_ss = empty_ss setsubgoaler asm_simp_tac |
4b30dbe4a020
added delcongs, Delcongs, unsafe_solver, safe_solver, HOL_basic_ss,
oheimb
parents:
2595
diff
changeset
|
367 |
setSSolver safe_solver |
4b30dbe4a020
added delcongs, Delcongs, unsafe_solver, safe_solver, HOL_basic_ss,
oheimb
parents:
2595
diff
changeset
|
368 |
setSolver unsafe_solver |
4b30dbe4a020
added delcongs, Delcongs, unsafe_solver, safe_solver, HOL_basic_ss,
oheimb
parents:
2595
diff
changeset
|
369 |
setmksimps (mksimps mksimps_pairs); |
2443
a81d4c219c3c
factored out HOL_base_ss and val HOL_min_ss, added HOL_safe_min_ss
oheimb
parents:
2263
diff
changeset
|
370 |
|
3446
a14e5451f613
Addition of not_imp (which pushes negation into implication) as a default
paulson
parents:
3282
diff
changeset
|
371 |
val HOL_ss = |
a14e5451f613
Addition of not_imp (which pushes negation into implication) as a default
paulson
parents:
3282
diff
changeset
|
372 |
HOL_basic_ss addsimps |
a14e5451f613
Addition of not_imp (which pushes negation into implication) as a default
paulson
parents:
3282
diff
changeset
|
373 |
([triv_forall_equality, (* prunes params *) |
3654 | 374 |
True_implies_equals, (* prune asms `True' *) |
3446
a14e5451f613
Addition of not_imp (which pushes negation into implication) as a default
paulson
parents:
3282
diff
changeset
|
375 |
if_True, if_False, if_cancel, |
a14e5451f613
Addition of not_imp (which pushes negation into implication) as a default
paulson
parents:
3282
diff
changeset
|
376 |
o_apply, imp_disjL, conj_assoc, disj_assoc, |
3904 | 377 |
de_Morgan_conj, de_Morgan_disj, imp_disj1, imp_disj2, not_imp, |
3446
a14e5451f613
Addition of not_imp (which pushes negation into implication) as a default
paulson
parents:
3282
diff
changeset
|
378 |
not_all, not_ex, cases_simp] |
a14e5451f613
Addition of not_imp (which pushes negation into implication) as a default
paulson
parents:
3282
diff
changeset
|
379 |
@ ex_simps @ all_simps @ simp_thms) |
3568
36ff1ab12021
Prod.ML: Added split_paired_EX and lots of comments about failed attempts to
nipkow
parents:
3559
diff
changeset
|
380 |
addsimprocs [defEX_regroup] |
3446
a14e5451f613
Addition of not_imp (which pushes negation into implication) as a default
paulson
parents:
3282
diff
changeset
|
381 |
addcongs [imp_cong]; |
2082 | 382 |
|
1655 | 383 |
qed_goal "if_distrib" HOL.thy |
384 |
"f(if c then x else y) = (if c then f x else f y)" |
|
385 |
(fn _ => [simp_tac (HOL_ss setloop (split_tac [expand_if])) 1]); |
|
386 |
||
2097 | 387 |
qed_goalw "o_assoc" HOL.thy [o_def] "f o (g o h) = f o g o h" |
2098
2bfc0675c92f
corrected `correction` of o_assoc (of version 1.14),
oheimb
parents:
2097
diff
changeset
|
388 |
(fn _ => [rtac ext 1, rtac refl 1]); |
1984 | 389 |
|
390 |
||
2948
f18035b1d531
Moved expand_case_tac from Auth/Message.ML to simpdata.ML
paulson
parents:
2935
diff
changeset
|
391 |
val prems = goal HOL.thy "[| P ==> Q(True); ~P ==> Q(False) |] ==> Q(P)"; |
f18035b1d531
Moved expand_case_tac from Auth/Message.ML to simpdata.ML
paulson
parents:
2935
diff
changeset
|
392 |
by (case_tac "P" 1); |
f18035b1d531
Moved expand_case_tac from Auth/Message.ML to simpdata.ML
paulson
parents:
2935
diff
changeset
|
393 |
by (ALLGOALS (asm_simp_tac (HOL_ss addsimps prems))); |
f18035b1d531
Moved expand_case_tac from Auth/Message.ML to simpdata.ML
paulson
parents:
2935
diff
changeset
|
394 |
val expand_case = result(); |
f18035b1d531
Moved expand_case_tac from Auth/Message.ML to simpdata.ML
paulson
parents:
2935
diff
changeset
|
395 |
|
f18035b1d531
Moved expand_case_tac from Auth/Message.ML to simpdata.ML
paulson
parents:
2935
diff
changeset
|
396 |
fun expand_case_tac P i = |
f18035b1d531
Moved expand_case_tac from Auth/Message.ML to simpdata.ML
paulson
parents:
2935
diff
changeset
|
397 |
res_inst_tac [("P",P)] expand_case i THEN |
f18035b1d531
Moved expand_case_tac from Auth/Message.ML to simpdata.ML
paulson
parents:
2935
diff
changeset
|
398 |
Simp_tac (i+1) THEN |
f18035b1d531
Moved expand_case_tac from Auth/Message.ML to simpdata.ML
paulson
parents:
2935
diff
changeset
|
399 |
Simp_tac i; |
f18035b1d531
Moved expand_case_tac from Auth/Message.ML to simpdata.ML
paulson
parents:
2935
diff
changeset
|
400 |
|
f18035b1d531
Moved expand_case_tac from Auth/Message.ML to simpdata.ML
paulson
parents:
2935
diff
changeset
|
401 |
|
1984 | 402 |
|
403 |
||
404 |
(*** Install simpsets and datatypes in theory structure ***) |
|
405 |
||
2251 | 406 |
simpset := HOL_ss; |
1984 | 407 |
|
408 |
exception SS_DATA of simpset; |
|
409 |
||
410 |
let fun merge [] = SS_DATA empty_ss |
|
411 |
| merge ss = let val ss = map (fn SS_DATA x => x) ss; |
|
412 |
in SS_DATA (foldl merge_ss (hd ss, tl ss)) end; |
|
413 |
||
414 |
fun put (SS_DATA ss) = simpset := ss; |
|
415 |
||
416 |
fun get () = SS_DATA (!simpset); |
|
417 |
in add_thydata "HOL" |
|
418 |
("simpset", ThyMethods {merge = merge, put = put, get = get}) |
|
419 |
end; |
|
420 |
||
3615
e5322197cfea
Moved some functions which used to be part of thy_data.ML
berghofe
parents:
3577
diff
changeset
|
421 |
fun simpset_of tname = |
e5322197cfea
Moved some functions which used to be part of thy_data.ML
berghofe
parents:
3577
diff
changeset
|
422 |
case get_thydata tname "simpset" of |
e5322197cfea
Moved some functions which used to be part of thy_data.ML
berghofe
parents:
3577
diff
changeset
|
423 |
None => empty_ss |
e5322197cfea
Moved some functions which used to be part of thy_data.ML
berghofe
parents:
3577
diff
changeset
|
424 |
| Some (SS_DATA ss) => ss; |
e5322197cfea
Moved some functions which used to be part of thy_data.ML
berghofe
parents:
3577
diff
changeset
|
425 |
|
3040
7d48671753da
Introduced a generic "induct_tac" which picks up the right induction scheme
nipkow
parents:
2948
diff
changeset
|
426 |
type dtype_info = {case_const:term, |
7d48671753da
Introduced a generic "induct_tac" which picks up the right induction scheme
nipkow
parents:
2948
diff
changeset
|
427 |
case_rewrites:thm list, |
7d48671753da
Introduced a generic "induct_tac" which picks up the right induction scheme
nipkow
parents:
2948
diff
changeset
|
428 |
constructors:term list, |
7d48671753da
Introduced a generic "induct_tac" which picks up the right induction scheme
nipkow
parents:
2948
diff
changeset
|
429 |
induct_tac: string -> int -> tactic, |
3282
c31e6239d4c9
Added exhaustion thm and exhaust_tac for each datatype.
nipkow
parents:
3206
diff
changeset
|
430 |
nchotomy: thm, |
c31e6239d4c9
Added exhaustion thm and exhaust_tac for each datatype.
nipkow
parents:
3206
diff
changeset
|
431 |
exhaustion: thm, |
c31e6239d4c9
Added exhaustion thm and exhaust_tac for each datatype.
nipkow
parents:
3206
diff
changeset
|
432 |
exhaust_tac: string -> int -> tactic, |
3040
7d48671753da
Introduced a generic "induct_tac" which picks up the right induction scheme
nipkow
parents:
2948
diff
changeset
|
433 |
case_cong:thm}; |
1984 | 434 |
|
435 |
exception DT_DATA of (string * dtype_info) list; |
|
436 |
val datatypes = ref [] : (string * dtype_info) list ref; |
|
437 |
||
438 |
let fun merge [] = DT_DATA [] |
|
439 |
| merge ds = |
|
440 |
let val ds = map (fn DT_DATA x => x) ds; |
|
441 |
in DT_DATA (foldl (gen_union eq_fst) (hd ds, tl ds)) end; |
|
442 |
||
443 |
fun put (DT_DATA ds) = datatypes := ds; |
|
444 |
||
445 |
fun get () = DT_DATA (!datatypes); |
|
446 |
in add_thydata "HOL" |
|
447 |
("datatypes", ThyMethods {merge = merge, put = put, get = get}) |
|
448 |
end; |
|
449 |
||
450 |
||
2636
4b30dbe4a020
added delcongs, Delcongs, unsafe_solver, safe_solver, HOL_basic_ss,
oheimb
parents:
2595
diff
changeset
|
451 |
|
4b30dbe4a020
added delcongs, Delcongs, unsafe_solver, safe_solver, HOL_basic_ss,
oheimb
parents:
2595
diff
changeset
|
452 |
(*** Integration of simplifier with classical reasoner ***) |
4b30dbe4a020
added delcongs, Delcongs, unsafe_solver, safe_solver, HOL_basic_ss,
oheimb
parents:
2595
diff
changeset
|
453 |
|
4b30dbe4a020
added delcongs, Delcongs, unsafe_solver, safe_solver, HOL_basic_ss,
oheimb
parents:
2595
diff
changeset
|
454 |
(* rot_eq_tac rotates the first equality premise of subgoal i to the front, |
4b30dbe4a020
added delcongs, Delcongs, unsafe_solver, safe_solver, HOL_basic_ss,
oheimb
parents:
2595
diff
changeset
|
455 |
fails if there is no equaliy or if an equality is already at the front *) |
3538 | 456 |
local |
457 |
fun is_eq (Const ("Trueprop", _) $ (Const("op =" ,_) $ _ $ _)) = true |
|
458 |
| is_eq _ = false; |
|
459 |
fun find_eq n [] = None |
|
460 |
| find_eq n (t :: ts) = if (is_eq t) then Some n |
|
461 |
else find_eq (n + 1) ts; |
|
462 |
in |
|
463 |
val rot_eq_tac = |
|
464 |
SUBGOAL (fn (Bi,i) => |
|
465 |
case find_eq 0 (Logic.strip_assums_hyp Bi) of |
|
2805 | 466 |
None => no_tac |
467 |
| Some 0 => no_tac |
|
3538 | 468 |
| Some n => rotate_tac n i) |
469 |
end; |
|
2636
4b30dbe4a020
added delcongs, Delcongs, unsafe_solver, safe_solver, HOL_basic_ss,
oheimb
parents:
2595
diff
changeset
|
470 |
|
4b30dbe4a020
added delcongs, Delcongs, unsafe_solver, safe_solver, HOL_basic_ss,
oheimb
parents:
2595
diff
changeset
|
471 |
(*an unsatisfactory fix for the incomplete asm_full_simp_tac! |
4b30dbe4a020
added delcongs, Delcongs, unsafe_solver, safe_solver, HOL_basic_ss,
oheimb
parents:
2595
diff
changeset
|
472 |
better: asm_really_full_simp_tac, a yet to be implemented version of |
4b30dbe4a020
added delcongs, Delcongs, unsafe_solver, safe_solver, HOL_basic_ss,
oheimb
parents:
2595
diff
changeset
|
473 |
asm_full_simp_tac that applies all equalities in the |
4b30dbe4a020
added delcongs, Delcongs, unsafe_solver, safe_solver, HOL_basic_ss,
oheimb
parents:
2595
diff
changeset
|
474 |
premises to all the premises *) |
4b30dbe4a020
added delcongs, Delcongs, unsafe_solver, safe_solver, HOL_basic_ss,
oheimb
parents:
2595
diff
changeset
|
475 |
fun safe_asm_more_full_simp_tac ss = TRY o rot_eq_tac THEN' |
4b30dbe4a020
added delcongs, Delcongs, unsafe_solver, safe_solver, HOL_basic_ss,
oheimb
parents:
2595
diff
changeset
|
476 |
safe_asm_full_simp_tac ss; |
4b30dbe4a020
added delcongs, Delcongs, unsafe_solver, safe_solver, HOL_basic_ss,
oheimb
parents:
2595
diff
changeset
|
477 |
|
4b30dbe4a020
added delcongs, Delcongs, unsafe_solver, safe_solver, HOL_basic_ss,
oheimb
parents:
2595
diff
changeset
|
478 |
(*Add a simpset to a classical set!*) |
3206
a3de7f32728c
renamed addss to addSss, unsafe_addss to addss, extended auto_tac
oheimb
parents:
3040
diff
changeset
|
479 |
infix 4 addSss addss; |
a3de7f32728c
renamed addss to addSss, unsafe_addss to addss, extended auto_tac
oheimb
parents:
3040
diff
changeset
|
480 |
fun cs addSss ss = cs addSaltern (CHANGED o (safe_asm_more_full_simp_tac ss)); |
a3de7f32728c
renamed addss to addSss, unsafe_addss to addss, extended auto_tac
oheimb
parents:
3040
diff
changeset
|
481 |
fun cs addss ss = cs addbefore asm_full_simp_tac ss; |
2636
4b30dbe4a020
added delcongs, Delcongs, unsafe_solver, safe_solver, HOL_basic_ss,
oheimb
parents:
2595
diff
changeset
|
482 |
|
4b30dbe4a020
added delcongs, Delcongs, unsafe_solver, safe_solver, HOL_basic_ss,
oheimb
parents:
2595
diff
changeset
|
483 |
fun Addss ss = (claset := !claset addss ss); |
4b30dbe4a020
added delcongs, Delcongs, unsafe_solver, safe_solver, HOL_basic_ss,
oheimb
parents:
2595
diff
changeset
|
484 |
|
4b30dbe4a020
added delcongs, Delcongs, unsafe_solver, safe_solver, HOL_basic_ss,
oheimb
parents:
2595
diff
changeset
|
485 |
(*Designed to be idempotent, except if best_tac instantiates variables |
4b30dbe4a020
added delcongs, Delcongs, unsafe_solver, safe_solver, HOL_basic_ss,
oheimb
parents:
2595
diff
changeset
|
486 |
in some of the subgoals*) |
4b30dbe4a020
added delcongs, Delcongs, unsafe_solver, safe_solver, HOL_basic_ss,
oheimb
parents:
2595
diff
changeset
|
487 |
|
4b30dbe4a020
added delcongs, Delcongs, unsafe_solver, safe_solver, HOL_basic_ss,
oheimb
parents:
2595
diff
changeset
|
488 |
type clasimpset = (claset * simpset); |
4b30dbe4a020
added delcongs, Delcongs, unsafe_solver, safe_solver, HOL_basic_ss,
oheimb
parents:
2595
diff
changeset
|
489 |
|
4b30dbe4a020
added delcongs, Delcongs, unsafe_solver, safe_solver, HOL_basic_ss,
oheimb
parents:
2595
diff
changeset
|
490 |
val HOL_css = (HOL_cs, HOL_ss); |
4b30dbe4a020
added delcongs, Delcongs, unsafe_solver, safe_solver, HOL_basic_ss,
oheimb
parents:
2595
diff
changeset
|
491 |
|
4b30dbe4a020
added delcongs, Delcongs, unsafe_solver, safe_solver, HOL_basic_ss,
oheimb
parents:
2595
diff
changeset
|
492 |
fun pair_upd1 f ((a,b),x) = (f(a,x), b); |
4b30dbe4a020
added delcongs, Delcongs, unsafe_solver, safe_solver, HOL_basic_ss,
oheimb
parents:
2595
diff
changeset
|
493 |
fun pair_upd2 f ((a,b),x) = (a, f(b,x)); |
4b30dbe4a020
added delcongs, Delcongs, unsafe_solver, safe_solver, HOL_basic_ss,
oheimb
parents:
2595
diff
changeset
|
494 |
|
4b30dbe4a020
added delcongs, Delcongs, unsafe_solver, safe_solver, HOL_basic_ss,
oheimb
parents:
2595
diff
changeset
|
495 |
infix 4 addSIs2 addSEs2 addSDs2 addIs2 addEs2 addDs2 |
4b30dbe4a020
added delcongs, Delcongs, unsafe_solver, safe_solver, HOL_basic_ss,
oheimb
parents:
2595
diff
changeset
|
496 |
addsimps2 delsimps2 addcongs2 delcongs2; |
2748
3ae9ccdd701e
Eta-expanded some declarations for compatibility with value polymorphism
paulson
parents:
2718
diff
changeset
|
497 |
fun op addSIs2 arg = pair_upd1 (op addSIs) arg; |
3ae9ccdd701e
Eta-expanded some declarations for compatibility with value polymorphism
paulson
parents:
2718
diff
changeset
|
498 |
fun op addSEs2 arg = pair_upd1 (op addSEs) arg; |
3ae9ccdd701e
Eta-expanded some declarations for compatibility with value polymorphism
paulson
parents:
2718
diff
changeset
|
499 |
fun op addSDs2 arg = pair_upd1 (op addSDs) arg; |
3ae9ccdd701e
Eta-expanded some declarations for compatibility with value polymorphism
paulson
parents:
2718
diff
changeset
|
500 |
fun op addIs2 arg = pair_upd1 (op addIs ) arg; |
3ae9ccdd701e
Eta-expanded some declarations for compatibility with value polymorphism
paulson
parents:
2718
diff
changeset
|
501 |
fun op addEs2 arg = pair_upd1 (op addEs ) arg; |
3ae9ccdd701e
Eta-expanded some declarations for compatibility with value polymorphism
paulson
parents:
2718
diff
changeset
|
502 |
fun op addDs2 arg = pair_upd1 (op addDs ) arg; |
3ae9ccdd701e
Eta-expanded some declarations for compatibility with value polymorphism
paulson
parents:
2718
diff
changeset
|
503 |
fun op addsimps2 arg = pair_upd2 (op addsimps) arg; |
3ae9ccdd701e
Eta-expanded some declarations for compatibility with value polymorphism
paulson
parents:
2718
diff
changeset
|
504 |
fun op delsimps2 arg = pair_upd2 (op delsimps) arg; |
3ae9ccdd701e
Eta-expanded some declarations for compatibility with value polymorphism
paulson
parents:
2718
diff
changeset
|
505 |
fun op addcongs2 arg = pair_upd2 (op addcongs) arg; |
3ae9ccdd701e
Eta-expanded some declarations for compatibility with value polymorphism
paulson
parents:
2718
diff
changeset
|
506 |
fun op delcongs2 arg = pair_upd2 (op delcongs) arg; |
2636
4b30dbe4a020
added delcongs, Delcongs, unsafe_solver, safe_solver, HOL_basic_ss,
oheimb
parents:
2595
diff
changeset
|
507 |
|
2805 | 508 |
fun auto_tac (cs,ss) = |
509 |
let val cs' = cs addss ss |
|
510 |
in EVERY [TRY (safe_tac cs'), |
|
511 |
REPEAT (FIRSTGOAL (fast_tac cs')), |
|
3206
a3de7f32728c
renamed addss to addSss, unsafe_addss to addss, extended auto_tac
oheimb
parents:
3040
diff
changeset
|
512 |
TRY (safe_tac (cs addSss ss)), |
2805 | 513 |
prune_params_tac] |
514 |
end; |
|
2636
4b30dbe4a020
added delcongs, Delcongs, unsafe_solver, safe_solver, HOL_basic_ss,
oheimb
parents:
2595
diff
changeset
|
515 |
|
4b30dbe4a020
added delcongs, Delcongs, unsafe_solver, safe_solver, HOL_basic_ss,
oheimb
parents:
2595
diff
changeset
|
516 |
fun Auto_tac () = auto_tac (!claset, !simpset); |
4b30dbe4a020
added delcongs, Delcongs, unsafe_solver, safe_solver, HOL_basic_ss,
oheimb
parents:
2595
diff
changeset
|
517 |
|
4b30dbe4a020
added delcongs, Delcongs, unsafe_solver, safe_solver, HOL_basic_ss,
oheimb
parents:
2595
diff
changeset
|
518 |
fun auto () = by (Auto_tac ()); |