author | haftmann |
Tue, 20 Oct 2009 16:13:01 +0200 | |
changeset 33037 | b22e44496dc2 |
parent 32957 | 675c0c7e6a37 |
child 35021 | c839a4c670c6 |
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 |
26288
89b9f7c18631
eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents:
22822
diff
changeset
|
11 |
_ $ (Const("op =",_)$_$_) => th RS @{thm eq_reflection} |
89b9f7c18631
eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents:
22822
diff
changeset
|
12 |
| _ $ (Const("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 |
5555 | 18 |
| _ $ (Const("op =",_)$_$_) => mk_meta_eq th |
19 |
| _ $ (Const("op <->",_)$_$_) => mk_meta_eq th |
|
26288
89b9f7c18631
eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents:
22822
diff
changeset
|
20 |
| _ $ (Const("Not",_)$_) => th RS @{thm iff_reflection_F} |
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*) |
10431 | 24 |
val mk_meta_prems = |
25 |
rule_by_tactic |
|
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 ==)*) |
6114
45958e54d72e
congruence rules finally use == instead of = and <->
paulson
parents:
5555
diff
changeset
|
29 |
fun mk_meta_cong rl = |
32957 | 30 |
Drule.standard (mk_meta_eq (mk_meta_prems rl)) |
6114
45958e54d72e
congruence rules finally use == instead of = and <->
paulson
parents:
5555
diff
changeset
|
31 |
handle THM _ => |
45958e54d72e
congruence rules finally use == instead of = and <->
paulson
parents:
5555
diff
changeset
|
32 |
error("Premises and conclusion of congruence rules must use =-equality or <->"); |
5555 | 33 |
|
5304 | 34 |
val mksimps_pairs = |
26288
89b9f7c18631
eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents:
22822
diff
changeset
|
35 |
[("op -->", [@{thm mp}]), ("op &", [@{thm conjunct1}, @{thm conjunct2}]), |
89b9f7c18631
eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents:
22822
diff
changeset
|
36 |
("All", [@{thm spec}]), ("True", []), ("False", [])]; |
5304 | 37 |
|
16019 | 38 |
(* ###FIXME: move to simplifier.ML |
5304 | 39 |
val mk_atomize: (string * thm list) list -> thm -> thm list |
40 |
*) |
|
16019 | 41 |
(* ###FIXME: move to simplifier.ML *) |
5304 | 42 |
fun mk_atomize pairs = |
43 |
let fun atoms th = |
|
44 |
(case concl_of th of |
|
45 |
Const("Trueprop",_) $ p => |
|
46 |
(case head_of p of |
|
47 |
Const(a,_) => |
|
17325 | 48 |
(case AList.lookup (op =) pairs a of |
32952 | 49 |
SOME(rls) => maps atoms ([th] RL rls) |
15531 | 50 |
| NONE => [th]) |
5304 | 51 |
| _ => [th]) |
52 |
| _ => [th]) |
|
53 |
in atoms end; |
|
54 |
||
12725 | 55 |
fun mksimps pairs = (map mk_eq o mk_atomize pairs o gen_all); |
981 | 56 |
|
1914 | 57 |
|
4349
50403e5a44c0
Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents:
4325
diff
changeset
|
58 |
(** make simplification procedures for quantifier elimination **) |
50403e5a44c0
Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents:
4325
diff
changeset
|
59 |
structure Quantifier1 = Quantifier1Fun( |
50403e5a44c0
Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents:
4325
diff
changeset
|
60 |
struct |
50403e5a44c0
Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents:
4325
diff
changeset
|
61 |
(*abstract syntax*) |
15531 | 62 |
fun dest_eq((c as Const("op =",_)) $ s $ t) = SOME(c,s,t) |
63 |
| dest_eq _ = NONE; |
|
64 |
fun dest_conj((c as Const("op &",_)) $ s $ t) = SOME(c,s,t) |
|
65 |
| dest_conj _ = NONE; |
|
66 |
fun dest_imp((c as Const("op -->",_)) $ s $ t) = SOME(c,s,t) |
|
67 |
| dest_imp _ = NONE; |
|
4349
50403e5a44c0
Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents:
4325
diff
changeset
|
68 |
val conj = FOLogic.conj |
50403e5a44c0
Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents:
4325
diff
changeset
|
69 |
val imp = FOLogic.imp |
50403e5a44c0
Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents:
4325
diff
changeset
|
70 |
(*rules*) |
26288
89b9f7c18631
eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents:
22822
diff
changeset
|
71 |
val iff_reflection = @{thm iff_reflection} |
89b9f7c18631
eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents:
22822
diff
changeset
|
72 |
val iffI = @{thm iffI} |
89b9f7c18631
eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents:
22822
diff
changeset
|
73 |
val iff_trans = @{thm iff_trans} |
89b9f7c18631
eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents:
22822
diff
changeset
|
74 |
val conjI= @{thm conjI} |
89b9f7c18631
eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents:
22822
diff
changeset
|
75 |
val conjE= @{thm conjE} |
89b9f7c18631
eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents:
22822
diff
changeset
|
76 |
val impI = @{thm impI} |
89b9f7c18631
eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents:
22822
diff
changeset
|
77 |
val mp = @{thm mp} |
89b9f7c18631
eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents:
22822
diff
changeset
|
78 |
val uncurry = @{thm uncurry} |
89b9f7c18631
eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents:
22822
diff
changeset
|
79 |
val exI = @{thm exI} |
89b9f7c18631
eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents:
22822
diff
changeset
|
80 |
val exE = @{thm exE} |
89b9f7c18631
eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents:
22822
diff
changeset
|
81 |
val iff_allI = @{thm iff_allI} |
89b9f7c18631
eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents:
22822
diff
changeset
|
82 |
val iff_exI = @{thm iff_exI} |
89b9f7c18631
eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents:
22822
diff
changeset
|
83 |
val all_comm = @{thm all_comm} |
89b9f7c18631
eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents:
22822
diff
changeset
|
84 |
val ex_comm = @{thm ex_comm} |
4349
50403e5a44c0
Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents:
4325
diff
changeset
|
85 |
end); |
50403e5a44c0
Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents:
4325
diff
changeset
|
86 |
|
13462 | 87 |
val defEX_regroup = |
32010 | 88 |
Simplifier.simproc @{theory} |
13462 | 89 |
"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
|
90 |
|
50403e5a44c0
Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents:
4325
diff
changeset
|
91 |
val defALL_regroup = |
32010 | 92 |
Simplifier.simproc @{theory} |
13462 | 93 |
"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
|
94 |
|
50403e5a44c0
Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents:
4325
diff
changeset
|
95 |
|
50403e5a44c0
Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents:
4325
diff
changeset
|
96 |
(*** Case splitting ***) |
0 | 97 |
|
32177
bc02c5bfcb5b
renamed functor SplitterFun to Splitter, require explicit theory;
wenzelm
parents:
32155
diff
changeset
|
98 |
structure Splitter = Splitter |
bc02c5bfcb5b
renamed functor SplitterFun to Splitter, require explicit theory;
wenzelm
parents:
32155
diff
changeset
|
99 |
( |
bc02c5bfcb5b
renamed functor SplitterFun to Splitter, require explicit theory;
wenzelm
parents:
32155
diff
changeset
|
100 |
val thy = @{theory} |
bc02c5bfcb5b
renamed functor SplitterFun to Splitter, require explicit theory;
wenzelm
parents:
32155
diff
changeset
|
101 |
val mk_eq = mk_eq |
26288
89b9f7c18631
eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents:
22822
diff
changeset
|
102 |
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
|
103 |
val iffD = @{thm iffD2} |
bc02c5bfcb5b
renamed functor SplitterFun to Splitter, require explicit theory;
wenzelm
parents:
32155
diff
changeset
|
104 |
val disjE = @{thm disjE} |
bc02c5bfcb5b
renamed functor SplitterFun to Splitter, require explicit theory;
wenzelm
parents:
32155
diff
changeset
|
105 |
val conjE = @{thm conjE} |
bc02c5bfcb5b
renamed functor SplitterFun to Splitter, require explicit theory;
wenzelm
parents:
32155
diff
changeset
|
106 |
val exE = @{thm exE} |
bc02c5bfcb5b
renamed functor SplitterFun to Splitter, require explicit theory;
wenzelm
parents:
32155
diff
changeset
|
107 |
val contrapos = @{thm contrapos} |
bc02c5bfcb5b
renamed functor SplitterFun to Splitter, require explicit theory;
wenzelm
parents:
32155
diff
changeset
|
108 |
val contrapos2 = @{thm contrapos2} |
bc02c5bfcb5b
renamed functor SplitterFun to Splitter, require explicit theory;
wenzelm
parents:
32155
diff
changeset
|
109 |
val notnotD = @{thm notnotD} |
bc02c5bfcb5b
renamed functor SplitterFun to Splitter, require explicit theory;
wenzelm
parents:
32155
diff
changeset
|
110 |
); |
1722 | 111 |
|
32177
bc02c5bfcb5b
renamed functor SplitterFun to Splitter, require explicit theory;
wenzelm
parents:
32155
diff
changeset
|
112 |
val split_tac = Splitter.split_tac; |
5304 | 113 |
val split_inside_tac = Splitter.split_inside_tac; |
32177
bc02c5bfcb5b
renamed functor SplitterFun to Splitter, require explicit theory;
wenzelm
parents:
32155
diff
changeset
|
114 |
val split_asm_tac = Splitter.split_asm_tac; |
bc02c5bfcb5b
renamed functor SplitterFun to Splitter, require explicit theory;
wenzelm
parents:
32155
diff
changeset
|
115 |
val op addsplits = Splitter.addsplits; |
bc02c5bfcb5b
renamed functor SplitterFun to Splitter, require explicit theory;
wenzelm
parents:
32155
diff
changeset
|
116 |
val op delsplits = Splitter.delsplits; |
4325 | 117 |
|
118 |
||
2074
30a65172e003
Added the de Morgan laws (incl quantifier versions) to basic simpset
paulson
parents:
2065
diff
changeset
|
119 |
(*** Standard simpsets ***) |
30a65172e003
Added the de Morgan laws (incl quantifier versions) to basic simpset
paulson
parents:
2065
diff
changeset
|
120 |
|
26288
89b9f7c18631
eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents:
22822
diff
changeset
|
121 |
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
|
122 |
|
26288
89b9f7c18631
eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents:
22822
diff
changeset
|
123 |
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
|
124 |
atac, etac @{thm FalseE}]; |
2633
37c0b5a7ee5d
added delcongs, Delcongs, unsafe_solver, safe_solver, FOL_basic_ss,
oheimb
parents:
2601
diff
changeset
|
125 |
(*No premature instantiation of variables during simplification*) |
26288
89b9f7c18631
eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents:
22822
diff
changeset
|
126 |
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
|
127 |
eq_assume_tac, ematch_tac [@{thm FalseE}]]; |
2633
37c0b5a7ee5d
added delcongs, Delcongs, unsafe_solver, safe_solver, FOL_basic_ss,
oheimb
parents:
2601
diff
changeset
|
128 |
|
3910 | 129 |
(*No simprules, but basic infastructure for simplification*) |
17892 | 130 |
val FOL_basic_ss = |
32155 | 131 |
Simplifier.theory_context @{theory} empty_ss |
10431 | 132 |
setsubgoaler asm_simp_tac |
133 |
setSSolver (mk_solver "FOL safe" safe_solver) |
|
134 |
setSolver (mk_solver "FOL unsafe" unsafe_solver) |
|
135 |
setmksimps (mksimps mksimps_pairs) |
|
136 |
setmkcong mk_meta_cong; |
|
5304 | 137 |
|
18324 | 138 |
fun unfold_tac ths = |
139 |
let val ss0 = Simplifier.clear_ss FOL_basic_ss addsimps ths |
|
140 |
in fn ss => ALLGOALS (full_simp_tac (Simplifier.inherit_context ss ss0)) end; |
|
17002 | 141 |
|
2633
37c0b5a7ee5d
added delcongs, Delcongs, unsafe_solver, safe_solver, FOL_basic_ss,
oheimb
parents:
2601
diff
changeset
|
142 |
|
3910 | 143 |
(*intuitionistic simprules only*) |
21539 | 144 |
val IFOL_ss = |
145 |
FOL_basic_ss |
|
26288
89b9f7c18631
eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents:
22822
diff
changeset
|
146 |
addsimps (@{thms meta_simps} @ @{thms IFOL_simps} @ @{thms int_ex_simps} @ @{thms int_all_simps}) |
10431 | 147 |
addsimprocs [defALL_regroup, defEX_regroup] |
26288
89b9f7c18631
eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents:
22822
diff
changeset
|
148 |
addcongs [@{thm imp_cong}]; |
2074
30a65172e003
Added the de Morgan laws (incl quantifier versions) to basic simpset
paulson
parents:
2065
diff
changeset
|
149 |
|
3910 | 150 |
(*classical simprules too*) |
26288
89b9f7c18631
eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents:
22822
diff
changeset
|
151 |
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
|
152 |
|
26496
49ae9456eba9
purely functional setup of claset/simpset/clasimpset;
wenzelm
parents:
26288
diff
changeset
|
153 |
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
|
154 |
|
37c0b5a7ee5d
added delcongs, Delcongs, unsafe_solver, safe_solver, FOL_basic_ss,
oheimb
parents:
2601
diff
changeset
|
155 |
|
5219 | 156 |
(*** integration of simplifier with classical reasoner ***) |
2633
37c0b5a7ee5d
added delcongs, Delcongs, unsafe_solver, safe_solver, FOL_basic_ss,
oheimb
parents:
2601
diff
changeset
|
157 |
|
5219 | 158 |
structure Clasimp = ClasimpFun |
8472 | 159 |
(structure Simplifier = Simplifier and Splitter = Splitter |
9851 | 160 |
and Classical = Cla and Blast = Blast |
26288
89b9f7c18631
eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents:
22822
diff
changeset
|
161 |
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
|
162 |
open Clasimp; |
2633
37c0b5a7ee5d
added delcongs, Delcongs, unsafe_solver, safe_solver, FOL_basic_ss,
oheimb
parents:
2601
diff
changeset
|
163 |
|
27338 | 164 |
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
|
165 |
(Scan.succeed "Clasimp.clasimpset_of (ML_Context.the_local_context ())"); |
22128 | 166 |
|
2633
37c0b5a7ee5d
added delcongs, Delcongs, unsafe_solver, safe_solver, FOL_basic_ss,
oheimb
parents:
2601
diff
changeset
|
167 |
val FOL_css = (FOL_cs, FOL_ss); |