author | haftmann |
Thu, 23 Sep 2010 13:28:53 +0200 | |
changeset 39663 | 5096018d5359 |
parent 38715 | 6513ea67d95d |
child 41310 | 65631ca437c9 |
permissions | -rw-r--r-- |
9889 | 1 |
(* Title: FOL/simpdata.ML |
1459 | 2 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
282 | 3 |
Copyright 1994 University of Cambridge |
0 | 4 |
|
9889 | 5 |
Simplification data for FOL. |
0 | 6 |
*) |
7 |
||
282 | 8 |
(*Make meta-equalities. The operator below is Trueprop*) |
5555 | 9 |
|
282 | 10 |
fun mk_meta_eq th = case concl_of th of |
38500 | 11 |
_ $ (Const(@{const_name "op ="},_)$_$_) => th RS @{thm eq_reflection} |
12 |
| _ $ (Const(@{const_name "op <->"},_)$_$_) => th RS @{thm iff_reflection} |
|
10431 | 13 |
| _ => |
5555 | 14 |
error("conclusion must be a =-equality or <->");; |
15 |
||
16 |
fun mk_eq th = case concl_of th of |
|
394
432bb9995893
Modified mk_meta_eq to leave meta-equlities on unchanged.
nipkow
parents:
371
diff
changeset
|
17 |
Const("==",_)$_$_ => th |
38500 | 18 |
| _ $ (Const(@{const_name "op ="},_)$_$_) => mk_meta_eq th |
19 |
| _ $ (Const(@{const_name "op <->"},_)$_$_) => mk_meta_eq th |
|
20 |
| _ $ (Const(@{const_name Not},_)$_) => th RS @{thm iff_reflection_F} |
|
26288
89b9f7c18631
eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents:
22822
diff
changeset
|
21 |
| _ => th RS @{thm iff_reflection_T}; |
0 | 22 |
|
6114
45958e54d72e
congruence rules finally use == instead of = and <->
paulson
parents:
5555
diff
changeset
|
23 |
(*Replace premises x=y, X<->Y by X==Y*) |
36546 | 24 |
fun mk_meta_prems ctxt = |
25 |
rule_by_tactic ctxt |
|
26288
89b9f7c18631
eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents:
22822
diff
changeset
|
26 |
(REPEAT_FIRST (resolve_tac [@{thm meta_eq_to_obj_eq}, @{thm def_imp_iff}])); |
6114
45958e54d72e
congruence rules finally use == instead of = and <->
paulson
parents:
5555
diff
changeset
|
27 |
|
9713 | 28 |
(*Congruence rules for = or <-> (instead of ==)*) |
36546 | 29 |
fun mk_meta_cong ss rl = |
30 |
Drule.export_without_context (mk_meta_eq (mk_meta_prems (Simplifier.the_context ss) rl)) |
|
35021
c839a4c670c6
renamed old-style Drule.standard to Drule.export_without_context, to emphasize that this is in no way a standard operation;
wenzelm
parents:
32957
diff
changeset
|
31 |
handle THM _ => |
c839a4c670c6
renamed old-style Drule.standard to Drule.export_without_context, to emphasize that this is in no way a standard operation;
wenzelm
parents:
32957
diff
changeset
|
32 |
error("Premises and conclusion of congruence rules must use =-equality or <->"); |
5555 | 33 |
|
5304 | 34 |
val mksimps_pairs = |
38500 | 35 |
[(@{const_name "op -->"}, [@{thm mp}]), (@{const_name "op &"}, [@{thm conjunct1}, @{thm conjunct2}]), |
36 |
(@{const_name All}, [@{thm spec}]), (@{const_name True}, []), (@{const_name False}, [])]; |
|
5304 | 37 |
|
38 |
fun mk_atomize pairs = |
|
39 |
let fun atoms th = |
|
40 |
(case concl_of th of |
|
38500 | 41 |
Const(@{const_name Trueprop},_) $ p => |
5304 | 42 |
(case head_of p of |
43 |
Const(a,_) => |
|
17325 | 44 |
(case AList.lookup (op =) pairs a of |
32952 | 45 |
SOME(rls) => maps atoms ([th] RL rls) |
15531 | 46 |
| NONE => [th]) |
5304 | 47 |
| _ => [th]) |
48 |
| _ => [th]) |
|
49 |
in atoms end; |
|
50 |
||
36543
0e7fc5bf38de
proper context for mksimps etc. -- via simpset of the running Simplifier;
wenzelm
parents:
35232
diff
changeset
|
51 |
fun mksimps pairs (_: simpset) = map mk_eq o mk_atomize pairs o gen_all; |
981 | 52 |
|
1914 | 53 |
|
4349
50403e5a44c0
Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents:
4325
diff
changeset
|
54 |
(** make simplification procedures for quantifier elimination **) |
50403e5a44c0
Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents:
4325
diff
changeset
|
55 |
structure Quantifier1 = Quantifier1Fun( |
50403e5a44c0
Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents:
4325
diff
changeset
|
56 |
struct |
50403e5a44c0
Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents:
4325
diff
changeset
|
57 |
(*abstract syntax*) |
38500 | 58 |
fun dest_eq((c as Const(@{const_name "op ="},_)) $ s $ t) = SOME(c,s,t) |
15531 | 59 |
| dest_eq _ = NONE; |
38500 | 60 |
fun dest_conj((c as Const(@{const_name "op &"},_)) $ s $ t) = SOME(c,s,t) |
15531 | 61 |
| dest_conj _ = NONE; |
38500 | 62 |
fun dest_imp((c as Const(@{const_name "op -->"},_)) $ s $ t) = SOME(c,s,t) |
15531 | 63 |
| dest_imp _ = NONE; |
4349
50403e5a44c0
Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents:
4325
diff
changeset
|
64 |
val conj = FOLogic.conj |
50403e5a44c0
Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents:
4325
diff
changeset
|
65 |
val imp = FOLogic.imp |
50403e5a44c0
Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents:
4325
diff
changeset
|
66 |
(*rules*) |
26288
89b9f7c18631
eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents:
22822
diff
changeset
|
67 |
val iff_reflection = @{thm iff_reflection} |
89b9f7c18631
eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents:
22822
diff
changeset
|
68 |
val iffI = @{thm iffI} |
89b9f7c18631
eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents:
22822
diff
changeset
|
69 |
val iff_trans = @{thm iff_trans} |
89b9f7c18631
eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents:
22822
diff
changeset
|
70 |
val conjI= @{thm conjI} |
89b9f7c18631
eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents:
22822
diff
changeset
|
71 |
val conjE= @{thm conjE} |
89b9f7c18631
eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents:
22822
diff
changeset
|
72 |
val impI = @{thm impI} |
89b9f7c18631
eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents:
22822
diff
changeset
|
73 |
val mp = @{thm mp} |
89b9f7c18631
eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents:
22822
diff
changeset
|
74 |
val uncurry = @{thm uncurry} |
89b9f7c18631
eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents:
22822
diff
changeset
|
75 |
val exI = @{thm exI} |
89b9f7c18631
eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents:
22822
diff
changeset
|
76 |
val exE = @{thm exE} |
89b9f7c18631
eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents:
22822
diff
changeset
|
77 |
val iff_allI = @{thm iff_allI} |
89b9f7c18631
eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents:
22822
diff
changeset
|
78 |
val iff_exI = @{thm iff_exI} |
89b9f7c18631
eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents:
22822
diff
changeset
|
79 |
val all_comm = @{thm all_comm} |
89b9f7c18631
eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents:
22822
diff
changeset
|
80 |
val ex_comm = @{thm ex_comm} |
4349
50403e5a44c0
Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents:
4325
diff
changeset
|
81 |
end); |
50403e5a44c0
Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents:
4325
diff
changeset
|
82 |
|
13462 | 83 |
val defEX_regroup = |
38715
6513ea67d95d
renamed Simplifier.simproc(_i) to Simplifier.simproc_global(_i) to emphasize that this is not the real thing;
wenzelm
parents:
38500
diff
changeset
|
84 |
Simplifier.simproc_global @{theory} |
13462 | 85 |
"defined EX" ["EX x. P(x)"] Quantifier1.rearrange_ex; |
4349
50403e5a44c0
Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents:
4325
diff
changeset
|
86 |
|
50403e5a44c0
Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents:
4325
diff
changeset
|
87 |
val defALL_regroup = |
38715
6513ea67d95d
renamed Simplifier.simproc(_i) to Simplifier.simproc_global(_i) to emphasize that this is not the real thing;
wenzelm
parents:
38500
diff
changeset
|
88 |
Simplifier.simproc_global @{theory} |
13462 | 89 |
"defined ALL" ["ALL x. P(x)"] Quantifier1.rearrange_all; |
4349
50403e5a44c0
Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents:
4325
diff
changeset
|
90 |
|
50403e5a44c0
Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents:
4325
diff
changeset
|
91 |
|
50403e5a44c0
Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents:
4325
diff
changeset
|
92 |
(*** Case splitting ***) |
0 | 93 |
|
32177
bc02c5bfcb5b
renamed functor SplitterFun to Splitter, require explicit theory;
wenzelm
parents:
32155
diff
changeset
|
94 |
structure Splitter = Splitter |
bc02c5bfcb5b
renamed functor SplitterFun to Splitter, require explicit theory;
wenzelm
parents:
32155
diff
changeset
|
95 |
( |
bc02c5bfcb5b
renamed functor SplitterFun to Splitter, require explicit theory;
wenzelm
parents:
32155
diff
changeset
|
96 |
val thy = @{theory} |
bc02c5bfcb5b
renamed functor SplitterFun to Splitter, require explicit theory;
wenzelm
parents:
32155
diff
changeset
|
97 |
val mk_eq = mk_eq |
26288
89b9f7c18631
eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents:
22822
diff
changeset
|
98 |
val meta_eq_to_iff = @{thm meta_eq_to_iff} |
32177
bc02c5bfcb5b
renamed functor SplitterFun to Splitter, require explicit theory;
wenzelm
parents:
32155
diff
changeset
|
99 |
val iffD = @{thm iffD2} |
bc02c5bfcb5b
renamed functor SplitterFun to Splitter, require explicit theory;
wenzelm
parents:
32155
diff
changeset
|
100 |
val disjE = @{thm disjE} |
bc02c5bfcb5b
renamed functor SplitterFun to Splitter, require explicit theory;
wenzelm
parents:
32155
diff
changeset
|
101 |
val conjE = @{thm conjE} |
bc02c5bfcb5b
renamed functor SplitterFun to Splitter, require explicit theory;
wenzelm
parents:
32155
diff
changeset
|
102 |
val exE = @{thm exE} |
bc02c5bfcb5b
renamed functor SplitterFun to Splitter, require explicit theory;
wenzelm
parents:
32155
diff
changeset
|
103 |
val contrapos = @{thm contrapos} |
bc02c5bfcb5b
renamed functor SplitterFun to Splitter, require explicit theory;
wenzelm
parents:
32155
diff
changeset
|
104 |
val contrapos2 = @{thm contrapos2} |
bc02c5bfcb5b
renamed functor SplitterFun to Splitter, require explicit theory;
wenzelm
parents:
32155
diff
changeset
|
105 |
val notnotD = @{thm notnotD} |
bc02c5bfcb5b
renamed functor SplitterFun to Splitter, require explicit theory;
wenzelm
parents:
32155
diff
changeset
|
106 |
); |
1722 | 107 |
|
32177
bc02c5bfcb5b
renamed functor SplitterFun to Splitter, require explicit theory;
wenzelm
parents:
32155
diff
changeset
|
108 |
val split_tac = Splitter.split_tac; |
5304 | 109 |
val split_inside_tac = Splitter.split_inside_tac; |
32177
bc02c5bfcb5b
renamed functor SplitterFun to Splitter, require explicit theory;
wenzelm
parents:
32155
diff
changeset
|
110 |
val split_asm_tac = Splitter.split_asm_tac; |
bc02c5bfcb5b
renamed functor SplitterFun to Splitter, require explicit theory;
wenzelm
parents:
32155
diff
changeset
|
111 |
val op addsplits = Splitter.addsplits; |
bc02c5bfcb5b
renamed functor SplitterFun to Splitter, require explicit theory;
wenzelm
parents:
32155
diff
changeset
|
112 |
val op delsplits = Splitter.delsplits; |
4325 | 113 |
|
114 |
||
2074
30a65172e003
Added the de Morgan laws (incl quantifier versions) to basic simpset
paulson
parents:
2065
diff
changeset
|
115 |
(*** Standard simpsets ***) |
30a65172e003
Added the de Morgan laws (incl quantifier versions) to basic simpset
paulson
parents:
2065
diff
changeset
|
116 |
|
26288
89b9f7c18631
eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents:
22822
diff
changeset
|
117 |
val triv_rls = [@{thm TrueI}, @{thm refl}, reflexive_thm, @{thm iff_refl}, @{thm notFalseI}]; |
2074
30a65172e003
Added the de Morgan laws (incl quantifier versions) to basic simpset
paulson
parents:
2065
diff
changeset
|
118 |
|
26288
89b9f7c18631
eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents:
22822
diff
changeset
|
119 |
fun unsafe_solver prems = FIRST'[resolve_tac (triv_rls @ prems), |
89b9f7c18631
eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents:
22822
diff
changeset
|
120 |
atac, etac @{thm FalseE}]; |
2633
37c0b5a7ee5d
added delcongs, Delcongs, unsafe_solver, safe_solver, FOL_basic_ss,
oheimb
parents:
2601
diff
changeset
|
121 |
(*No premature instantiation of variables during simplification*) |
26288
89b9f7c18631
eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents:
22822
diff
changeset
|
122 |
fun safe_solver prems = FIRST'[match_tac (triv_rls @ prems), |
89b9f7c18631
eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents:
22822
diff
changeset
|
123 |
eq_assume_tac, ematch_tac [@{thm FalseE}]]; |
2633
37c0b5a7ee5d
added delcongs, Delcongs, unsafe_solver, safe_solver, FOL_basic_ss,
oheimb
parents:
2601
diff
changeset
|
124 |
|
3910 | 125 |
(*No simprules, but basic infastructure for simplification*) |
17892 | 126 |
val FOL_basic_ss = |
35232
f588e1169c8b
renamed Simplifier.theory_context to Simplifier.global_context to emphasize that this is not the real thing;
wenzelm
parents:
35021
diff
changeset
|
127 |
Simplifier.global_context @{theory} empty_ss |
10431 | 128 |
setsubgoaler asm_simp_tac |
129 |
setSSolver (mk_solver "FOL safe" safe_solver) |
|
130 |
setSolver (mk_solver "FOL unsafe" unsafe_solver) |
|
131 |
setmksimps (mksimps mksimps_pairs) |
|
132 |
setmkcong mk_meta_cong; |
|
5304 | 133 |
|
18324 | 134 |
fun unfold_tac ths = |
135 |
let val ss0 = Simplifier.clear_ss FOL_basic_ss addsimps ths |
|
136 |
in fn ss => ALLGOALS (full_simp_tac (Simplifier.inherit_context ss ss0)) end; |
|
17002 | 137 |
|
2633
37c0b5a7ee5d
added delcongs, Delcongs, unsafe_solver, safe_solver, FOL_basic_ss,
oheimb
parents:
2601
diff
changeset
|
138 |
|
3910 | 139 |
(*intuitionistic simprules only*) |
21539 | 140 |
val IFOL_ss = |
141 |
FOL_basic_ss |
|
26288
89b9f7c18631
eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents:
22822
diff
changeset
|
142 |
addsimps (@{thms meta_simps} @ @{thms IFOL_simps} @ @{thms int_ex_simps} @ @{thms int_all_simps}) |
10431 | 143 |
addsimprocs [defALL_regroup, defEX_regroup] |
26288
89b9f7c18631
eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents:
22822
diff
changeset
|
144 |
addcongs [@{thm imp_cong}]; |
2074
30a65172e003
Added the de Morgan laws (incl quantifier versions) to basic simpset
paulson
parents:
2065
diff
changeset
|
145 |
|
3910 | 146 |
(*classical simprules too*) |
26288
89b9f7c18631
eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents:
22822
diff
changeset
|
147 |
val FOL_ss = IFOL_ss addsimps (@{thms cla_simps} @ @{thms cla_ex_simps} @ @{thms cla_all_simps}); |
2074
30a65172e003
Added the de Morgan laws (incl quantifier versions) to basic simpset
paulson
parents:
2065
diff
changeset
|
148 |
|
26496
49ae9456eba9
purely functional setup of claset/simpset/clasimpset;
wenzelm
parents:
26288
diff
changeset
|
149 |
val simpsetup = Simplifier.map_simpset (K FOL_ss); |
2633
37c0b5a7ee5d
added delcongs, Delcongs, unsafe_solver, safe_solver, FOL_basic_ss,
oheimb
parents:
2601
diff
changeset
|
150 |
|
37c0b5a7ee5d
added delcongs, Delcongs, unsafe_solver, safe_solver, FOL_basic_ss,
oheimb
parents:
2601
diff
changeset
|
151 |
|
5219 | 152 |
(*** integration of simplifier with classical reasoner ***) |
2633
37c0b5a7ee5d
added delcongs, Delcongs, unsafe_solver, safe_solver, FOL_basic_ss,
oheimb
parents:
2601
diff
changeset
|
153 |
|
5219 | 154 |
structure Clasimp = ClasimpFun |
8472 | 155 |
(structure Simplifier = Simplifier and Splitter = Splitter |
9851 | 156 |
and Classical = Cla and Blast = Blast |
26288
89b9f7c18631
eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents:
22822
diff
changeset
|
157 |
val iffD1 = @{thm iffD1} val iffD2 = @{thm iffD2} val notE = @{thm notE}); |
4652
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
4649
diff
changeset
|
158 |
open Clasimp; |
2633
37c0b5a7ee5d
added delcongs, Delcongs, unsafe_solver, safe_solver, FOL_basic_ss,
oheimb
parents:
2601
diff
changeset
|
159 |
|
27338 | 160 |
ML_Antiquote.value "clasimpset" |
32149
ef59550a55d3
renamed simpset_of to global_simpset_of, and local_simpset_of to simpset_of -- same for claset and clasimpset;
wenzelm
parents:
32010
diff
changeset
|
161 |
(Scan.succeed "Clasimp.clasimpset_of (ML_Context.the_local_context ())"); |
22128 | 162 |
|
2633
37c0b5a7ee5d
added delcongs, Delcongs, unsafe_solver, safe_solver, FOL_basic_ss,
oheimb
parents:
2601
diff
changeset
|
163 |
val FOL_css = (FOL_cs, FOL_ss); |