| author | bulwahn | 
| Sun, 06 Aug 2017 21:49:25 +0200 | |
| changeset 66358 | fab9a53158f8 | 
| parent 63637 | 9a57baa15e1b | 
| child 69593 | 3dda49e08b9d | 
| 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 | |
| 59582 | 10 | fun mk_meta_eq th = | 
| 11 | (case Thm.concl_of th of | |
| 41310 | 12 |     _ $ (Const(@{const_name eq},_)$_$_)   => th RS @{thm eq_reflection}
 | 
| 13 |   | _ $ (Const(@{const_name iff},_)$_$_) => th RS @{thm iff_reflection}
 | |
| 59582 | 14 | | _ => error "conclusion must be a =-equality or <->"); | 
| 5555 | 15 | |
| 59582 | 16 | fun mk_eq th = | 
| 17 | (case Thm.concl_of th of | |
| 56245 | 18 |     Const(@{const_name Pure.eq},_)$_$_ => th
 | 
| 41310 | 19 |   | _ $ (Const(@{const_name eq},_)$_$_)   => mk_meta_eq th
 | 
| 20 |   | _ $ (Const(@{const_name iff},_)$_$_) => mk_meta_eq th
 | |
| 38500 | 21 |   | _ $ (Const(@{const_name Not},_)$_)      => th RS @{thm iff_reflection_F}
 | 
| 59582 | 22 |   | _  => th RS @{thm iff_reflection_T});
 | 
| 0 | 23 | |
| 6114 
45958e54d72e
congruence rules finally use == instead of = and <->
 paulson parents: 
5555diff
changeset | 24 | (*Replace premises x=y, X<->Y by X==Y*) | 
| 36546 | 25 | fun mk_meta_prems ctxt = | 
| 26 | rule_by_tactic ctxt | |
| 59498 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 wenzelm parents: 
58963diff
changeset | 27 |       (REPEAT_FIRST (resolve_tac ctxt [@{thm meta_eq_to_obj_eq}, @{thm def_imp_iff}]));
 | 
| 6114 
45958e54d72e
congruence rules finally use == instead of = and <->
 paulson parents: 
5555diff
changeset | 28 | |
| 9713 | 29 | (*Congruence rules for = or <-> (instead of ==)*) | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
45659diff
changeset | 30 | fun mk_meta_cong ctxt rl = | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
45659diff
changeset | 31 | Drule.zero_var_indexes (mk_meta_eq (mk_meta_prems ctxt 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: 
32957diff
changeset | 32 | 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: 
32957diff
changeset | 33 |       error("Premises and conclusion of congruence rules must use =-equality or <->");
 | 
| 5555 | 34 | |
| 5304 | 35 | val mksimps_pairs = | 
| 41310 | 36 |   [(@{const_name imp}, [@{thm mp}]), (@{const_name conj}, [@{thm conjunct1}, @{thm conjunct2}]),
 | 
| 38500 | 37 |    (@{const_name All}, [@{thm spec}]), (@{const_name True}, []), (@{const_name False}, [])];
 | 
| 5304 | 38 | |
| 39 | fun mk_atomize pairs = | |
| 59582 | 40 | let | 
| 41 | fun atoms th = | |
| 42 | (case Thm.concl_of th of | |
| 43 |          Const(@{const_name Trueprop},_) $ p =>
 | |
| 44 | (case head_of p of | |
| 45 | Const(a,_) => | |
| 46 | (case AList.lookup (op =) pairs a of | |
| 47 | SOME(rls) => maps atoms ([th] RL rls) | |
| 48 | | NONE => [th]) | |
| 49 | | _ => [th]) | |
| 50 | | _ => [th]) | |
| 5304 | 51 | in atoms end; | 
| 52 | ||
| 60822 | 53 | fun mksimps pairs ctxt = map mk_eq o mk_atomize pairs o Variable.gen_all ctxt; | 
| 981 | 54 | |
| 1914 | 55 | |
| 4349 
50403e5a44c0
Instantiated the one-point-rule quantifier simpprocs for FOL
 paulson parents: 
4325diff
changeset | 56 | (** make simplification procedures for quantifier elimination **) | 
| 42458 | 57 | structure Quantifier1 = Quantifier1 | 
| 58 | ( | |
| 4349 
50403e5a44c0
Instantiated the one-point-rule quantifier simpprocs for FOL
 paulson parents: 
4325diff
changeset | 59 | (*abstract syntax*) | 
| 42460 | 60 |   fun dest_eq (Const (@{const_name eq}, _) $ s $ t) = SOME (s, t)
 | 
| 61 | | dest_eq _ = NONE | |
| 62 |   fun dest_conj (Const (@{const_name conj}, _) $ s $ t) = SOME (s, t)
 | |
| 63 | | dest_conj _ = NONE | |
| 64 |   fun dest_imp (Const (@{const_name imp}, _) $ s $ t) = SOME (s, t)
 | |
| 65 | | dest_imp _ = NONE | |
| 4349 
50403e5a44c0
Instantiated the one-point-rule quantifier simpprocs for FOL
 paulson parents: 
4325diff
changeset | 66 | val conj = FOLogic.conj | 
| 
50403e5a44c0
Instantiated the one-point-rule quantifier simpprocs for FOL
 paulson parents: 
4325diff
changeset | 67 | val imp = FOLogic.imp | 
| 
50403e5a44c0
Instantiated the one-point-rule quantifier simpprocs for FOL
 paulson parents: 
4325diff
changeset | 68 | (*rules*) | 
| 26288 
89b9f7c18631
eliminated out-of-scope proofs (cf. theory IFOL and FOL);
 wenzelm parents: 
22822diff
changeset | 69 |   val iff_reflection = @{thm iff_reflection}
 | 
| 
89b9f7c18631
eliminated out-of-scope proofs (cf. theory IFOL and FOL);
 wenzelm parents: 
22822diff
changeset | 70 |   val iffI = @{thm iffI}
 | 
| 
89b9f7c18631
eliminated out-of-scope proofs (cf. theory IFOL and FOL);
 wenzelm parents: 
22822diff
changeset | 71 |   val iff_trans = @{thm iff_trans}
 | 
| 
89b9f7c18631
eliminated out-of-scope proofs (cf. theory IFOL and FOL);
 wenzelm parents: 
22822diff
changeset | 72 |   val conjI= @{thm conjI}
 | 
| 
89b9f7c18631
eliminated out-of-scope proofs (cf. theory IFOL and FOL);
 wenzelm parents: 
22822diff
changeset | 73 |   val conjE= @{thm conjE}
 | 
| 
89b9f7c18631
eliminated out-of-scope proofs (cf. theory IFOL and FOL);
 wenzelm parents: 
22822diff
changeset | 74 |   val impI = @{thm impI}
 | 
| 
89b9f7c18631
eliminated out-of-scope proofs (cf. theory IFOL and FOL);
 wenzelm parents: 
22822diff
changeset | 75 |   val mp   = @{thm mp}
 | 
| 
89b9f7c18631
eliminated out-of-scope proofs (cf. theory IFOL and FOL);
 wenzelm parents: 
22822diff
changeset | 76 |   val uncurry = @{thm uncurry}
 | 
| 
89b9f7c18631
eliminated out-of-scope proofs (cf. theory IFOL and FOL);
 wenzelm parents: 
22822diff
changeset | 77 |   val exI  = @{thm exI}
 | 
| 
89b9f7c18631
eliminated out-of-scope proofs (cf. theory IFOL and FOL);
 wenzelm parents: 
22822diff
changeset | 78 |   val exE  = @{thm exE}
 | 
| 
89b9f7c18631
eliminated out-of-scope proofs (cf. theory IFOL and FOL);
 wenzelm parents: 
22822diff
changeset | 79 |   val iff_allI = @{thm iff_allI}
 | 
| 
89b9f7c18631
eliminated out-of-scope proofs (cf. theory IFOL and FOL);
 wenzelm parents: 
22822diff
changeset | 80 |   val iff_exI = @{thm iff_exI}
 | 
| 
89b9f7c18631
eliminated out-of-scope proofs (cf. theory IFOL and FOL);
 wenzelm parents: 
22822diff
changeset | 81 |   val all_comm = @{thm all_comm}
 | 
| 
89b9f7c18631
eliminated out-of-scope proofs (cf. theory IFOL and FOL);
 wenzelm parents: 
22822diff
changeset | 82 |   val ex_comm = @{thm ex_comm}
 | 
| 42458 | 83 | ); | 
| 4349 
50403e5a44c0
Instantiated the one-point-rule quantifier simpprocs for FOL
 paulson parents: 
4325diff
changeset | 84 | |
| 
50403e5a44c0
Instantiated the one-point-rule quantifier simpprocs for FOL
 paulson parents: 
4325diff
changeset | 85 | |
| 
50403e5a44c0
Instantiated the one-point-rule quantifier simpprocs for FOL
 paulson parents: 
4325diff
changeset | 86 | (*** Case splitting ***) | 
| 0 | 87 | |
| 32177 
bc02c5bfcb5b
renamed functor SplitterFun to Splitter, require explicit theory;
 wenzelm parents: 
32155diff
changeset | 88 | structure Splitter = Splitter | 
| 
bc02c5bfcb5b
renamed functor SplitterFun to Splitter, require explicit theory;
 wenzelm parents: 
32155diff
changeset | 89 | ( | 
| 59970 | 90 |   val context = @{context}
 | 
| 32177 
bc02c5bfcb5b
renamed functor SplitterFun to Splitter, require explicit theory;
 wenzelm parents: 
32155diff
changeset | 91 | val mk_eq = mk_eq | 
| 26288 
89b9f7c18631
eliminated out-of-scope proofs (cf. theory IFOL and FOL);
 wenzelm parents: 
22822diff
changeset | 92 |   val meta_eq_to_iff = @{thm meta_eq_to_iff}
 | 
| 32177 
bc02c5bfcb5b
renamed functor SplitterFun to Splitter, require explicit theory;
 wenzelm parents: 
32155diff
changeset | 93 |   val iffD = @{thm iffD2}
 | 
| 
bc02c5bfcb5b
renamed functor SplitterFun to Splitter, require explicit theory;
 wenzelm parents: 
32155diff
changeset | 94 |   val disjE = @{thm disjE}
 | 
| 
bc02c5bfcb5b
renamed functor SplitterFun to Splitter, require explicit theory;
 wenzelm parents: 
32155diff
changeset | 95 |   val conjE = @{thm conjE}
 | 
| 
bc02c5bfcb5b
renamed functor SplitterFun to Splitter, require explicit theory;
 wenzelm parents: 
32155diff
changeset | 96 |   val exE = @{thm exE}
 | 
| 
bc02c5bfcb5b
renamed functor SplitterFun to Splitter, require explicit theory;
 wenzelm parents: 
32155diff
changeset | 97 |   val contrapos = @{thm contrapos}
 | 
| 
bc02c5bfcb5b
renamed functor SplitterFun to Splitter, require explicit theory;
 wenzelm parents: 
32155diff
changeset | 98 |   val contrapos2 = @{thm contrapos2}
 | 
| 
bc02c5bfcb5b
renamed functor SplitterFun to Splitter, require explicit theory;
 wenzelm parents: 
32155diff
changeset | 99 |   val notnotD = @{thm notnotD}
 | 
| 63637 | 100 | val safe_tac = Cla.safe_tac | 
| 32177 
bc02c5bfcb5b
renamed functor SplitterFun to Splitter, require explicit theory;
 wenzelm parents: 
32155diff
changeset | 101 | ); | 
| 1722 | 102 | |
| 32177 
bc02c5bfcb5b
renamed functor SplitterFun to Splitter, require explicit theory;
 wenzelm parents: 
32155diff
changeset | 103 | val split_tac = Splitter.split_tac; | 
| 5304 | 104 | val split_inside_tac = Splitter.split_inside_tac; | 
| 32177 
bc02c5bfcb5b
renamed functor SplitterFun to Splitter, require explicit theory;
 wenzelm parents: 
32155diff
changeset | 105 | val split_asm_tac = Splitter.split_asm_tac; | 
| 4325 | 106 | |
| 107 | ||
| 2074 
30a65172e003
Added the de Morgan laws (incl quantifier versions) to basic simpset
 paulson parents: 
2065diff
changeset | 108 | (*** Standard simpsets ***) | 
| 
30a65172e003
Added the de Morgan laws (incl quantifier versions) to basic simpset
 paulson parents: 
2065diff
changeset | 109 | |
| 26288 
89b9f7c18631
eliminated out-of-scope proofs (cf. theory IFOL and FOL);
 wenzelm parents: 
22822diff
changeset | 110 | 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: 
2065diff
changeset | 111 | |
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
45659diff
changeset | 112 | fun unsafe_solver ctxt = | 
| 59498 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 wenzelm parents: 
58963diff
changeset | 113 | FIRST' [resolve_tac ctxt (triv_rls @ Simplifier.prems_of ctxt), | 
| 58963 
26bf09b95dda
proper context for assume_tac (atac remains as fall-back without context);
 wenzelm parents: 
58957diff
changeset | 114 | assume_tac ctxt, | 
| 59498 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 wenzelm parents: 
58963diff
changeset | 115 |     eresolve_tac ctxt @{thms FalseE}];
 | 
| 43597 | 116 | |
| 2633 
37c0b5a7ee5d
added delcongs, Delcongs, unsafe_solver, safe_solver, FOL_basic_ss,
 oheimb parents: 
2601diff
changeset | 117 | (*No premature instantiation of variables during simplification*) | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
45659diff
changeset | 118 | fun safe_solver ctxt = | 
| 58957 | 119 | FIRST' [match_tac ctxt (triv_rls @ Simplifier.prems_of ctxt), | 
| 120 |     eq_assume_tac, ematch_tac ctxt @{thms FalseE}];
 | |
| 2633 
37c0b5a7ee5d
added delcongs, Delcongs, unsafe_solver, safe_solver, FOL_basic_ss,
 oheimb parents: 
2601diff
changeset | 121 | |
| 3910 | 122 | (*No simprules, but basic infastructure for simplification*) | 
| 17892 | 123 | val FOL_basic_ss = | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
45659diff
changeset | 124 |   empty_simpset @{context}
 | 
| 10431 | 125 | setSSolver (mk_solver "FOL safe" safe_solver) | 
| 126 | setSolver (mk_solver "FOL unsafe" unsafe_solver) | |
| 45625 
750c5a47400b
modernized some old-style infix operations, which were left over from the time of ML proof scripts;
 wenzelm parents: 
45620diff
changeset | 127 | |> Simplifier.set_subgoaler asm_simp_tac | 
| 
750c5a47400b
modernized some old-style infix operations, which were left over from the time of ML proof scripts;
 wenzelm parents: 
45620diff
changeset | 128 | |> Simplifier.set_mksimps (mksimps mksimps_pairs) | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
45659diff
changeset | 129 | |> Simplifier.set_mkcong mk_meta_cong | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
45659diff
changeset | 130 | |> simpset_of; | 
| 5304 | 131 | |
| 54998 | 132 | fun unfold_tac ctxt ths = | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
45659diff
changeset | 133 | ALLGOALS (full_simp_tac (clear_simpset (put_simpset FOL_basic_ss ctxt) addsimps ths)); | 
| 17002 | 134 | |
| 2633 
37c0b5a7ee5d
added delcongs, Delcongs, unsafe_solver, safe_solver, FOL_basic_ss,
 oheimb parents: 
2601diff
changeset | 135 | |
| 5219 | 136 | (*** integration of simplifier with classical reasoner ***) | 
| 2633 
37c0b5a7ee5d
added delcongs, Delcongs, unsafe_solver, safe_solver, FOL_basic_ss,
 oheimb parents: 
2601diff
changeset | 137 | |
| 42478 | 138 | structure Clasimp = Clasimp | 
| 139 | ( | |
| 140 | structure Simplifier = Simplifier | |
| 141 | and Splitter = Splitter | |
| 142 | and Classical = Cla | |
| 143 | and Blast = Blast | |
| 144 |   val iffD1 = @{thm iffD1}
 | |
| 145 |   val iffD2 = @{thm iffD2}
 | |
| 146 |   val notE = @{thm notE}
 | |
| 147 | ); | |
| 4652 
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
 oheimb parents: 
4649diff
changeset | 148 | open Clasimp; | 
| 2633 
37c0b5a7ee5d
added delcongs, Delcongs, unsafe_solver, safe_solver, FOL_basic_ss,
 oheimb parents: 
2601diff
changeset | 149 |