| author | desharna | 
| Mon, 10 Jan 2022 21:34:09 +0100 | |
| changeset 74977 | e4fd3989554d | 
| parent 74319 | 54b2e5f771da | 
| child 82641 | d22294b20573 | 
| 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 | |
| 74293 | 12 |     _ $ \<^Const_>\<open>eq _ for _ _\<close> => th RS @{thm eq_reflection}
 | 
| 13 |   | _ $ \<^Const_>\<open>iff for _ _\<close> => 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 | |
| 74293 | 18 | \<^Const_>\<open>Pure.eq _ for _ _\<close> => th | 
| 19 | | _ $ \<^Const_>\<open>eq _ for _ _\<close> => mk_meta_eq th | |
| 20 | | _ $ \<^Const_>\<open>iff for _ _\<close> => mk_meta_eq th | |
| 21 |   | _ $ \<^Const_>\<open>Not for _\<close> => 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 = | 
| 69593 | 36 |   [(\<^const_name>\<open>imp\<close>, [@{thm mp}]), (\<^const_name>\<open>conj\<close>, [@{thm conjunct1}, @{thm conjunct2}]),
 | 
| 37 |    (\<^const_name>\<open>All\<close>, [@{thm spec}]), (\<^const_name>\<open>True\<close>, []), (\<^const_name>\<open>False\<close>, [])];
 | |
| 5304 | 38 | |
| 39 | fun mk_atomize pairs = | |
| 59582 | 40 | let | 
| 41 | fun atoms th = | |
| 42 | (case Thm.concl_of th of | |
| 74293 | 43 | \<^Const_>\<open>Trueprop for p\<close> => | 
| 59582 | 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*) | 
| 74293 | 60 | fun dest_eq \<^Const_>\<open>eq _ for s t\<close> = SOME (s, t) | 
| 42460 | 61 | | dest_eq _ = NONE | 
| 74293 | 62 | fun dest_conj \<^Const_>\<open>conj for s t\<close> = SOME (s, t) | 
| 42460 | 63 | | dest_conj _ = NONE | 
| 74293 | 64 | fun dest_imp \<^Const_>\<open>imp for s t\<close> = SOME (s, t) | 
| 42460 | 65 | | dest_imp _ = NONE | 
| 74319 
54b2e5f771da
clarified signature -- prefer antiquotations (with subtle change of exception content);
 wenzelm parents: 
74293diff
changeset | 66 | val conj = \<^Const>\<open>conj\<close> | 
| 
54b2e5f771da
clarified signature -- prefer antiquotations (with subtle change of exception content);
 wenzelm parents: 
74293diff
changeset | 67 | val imp = \<^Const>\<open>imp\<close> | 
| 4349 
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}
 | 
| 71916 
435cdc772110
specific atomization inert to later rule set modifications
 haftmann parents: 
69593diff
changeset | 83 | val atomize = | 
| 
435cdc772110
specific atomization inert to later rule set modifications
 haftmann parents: 
69593diff
changeset | 84 |     let val rules = @{thms atomize_all atomize_imp atomize_eq atomize_iff atomize_conj}
 | 
| 
435cdc772110
specific atomization inert to later rule set modifications
 haftmann parents: 
69593diff
changeset | 85 | in fn ctxt => Raw_Simplifier.rewrite ctxt true rules end | 
| 42458 | 86 | ); | 
| 4349 
50403e5a44c0
Instantiated the one-point-rule quantifier simpprocs for FOL
 paulson parents: 
4325diff
changeset | 87 | |
| 
50403e5a44c0
Instantiated the one-point-rule quantifier simpprocs for FOL
 paulson parents: 
4325diff
changeset | 88 | |
| 
50403e5a44c0
Instantiated the one-point-rule quantifier simpprocs for FOL
 paulson parents: 
4325diff
changeset | 89 | (*** Case splitting ***) | 
| 0 | 90 | |
| 32177 
bc02c5bfcb5b
renamed functor SplitterFun to Splitter, require explicit theory;
 wenzelm parents: 
32155diff
changeset | 91 | structure Splitter = Splitter | 
| 
bc02c5bfcb5b
renamed functor SplitterFun to Splitter, require explicit theory;
 wenzelm parents: 
32155diff
changeset | 92 | ( | 
| 69593 | 93 | val context = \<^context> | 
| 32177 
bc02c5bfcb5b
renamed functor SplitterFun to Splitter, require explicit theory;
 wenzelm parents: 
32155diff
changeset | 94 | val mk_eq = mk_eq | 
| 26288 
89b9f7c18631
eliminated out-of-scope proofs (cf. theory IFOL and FOL);
 wenzelm parents: 
22822diff
changeset | 95 |   val meta_eq_to_iff = @{thm meta_eq_to_iff}
 | 
| 32177 
bc02c5bfcb5b
renamed functor SplitterFun to Splitter, require explicit theory;
 wenzelm parents: 
32155diff
changeset | 96 |   val iffD = @{thm iffD2}
 | 
| 
bc02c5bfcb5b
renamed functor SplitterFun to Splitter, require explicit theory;
 wenzelm parents: 
32155diff
changeset | 97 |   val disjE = @{thm disjE}
 | 
| 
bc02c5bfcb5b
renamed functor SplitterFun to Splitter, require explicit theory;
 wenzelm parents: 
32155diff
changeset | 98 |   val conjE = @{thm conjE}
 | 
| 
bc02c5bfcb5b
renamed functor SplitterFun to Splitter, require explicit theory;
 wenzelm parents: 
32155diff
changeset | 99 |   val exE = @{thm exE}
 | 
| 
bc02c5bfcb5b
renamed functor SplitterFun to Splitter, require explicit theory;
 wenzelm parents: 
32155diff
changeset | 100 |   val contrapos = @{thm contrapos}
 | 
| 
bc02c5bfcb5b
renamed functor SplitterFun to Splitter, require explicit theory;
 wenzelm parents: 
32155diff
changeset | 101 |   val contrapos2 = @{thm contrapos2}
 | 
| 
bc02c5bfcb5b
renamed functor SplitterFun to Splitter, require explicit theory;
 wenzelm parents: 
32155diff
changeset | 102 |   val notnotD = @{thm notnotD}
 | 
| 63637 | 103 | val safe_tac = Cla.safe_tac | 
| 32177 
bc02c5bfcb5b
renamed functor SplitterFun to Splitter, require explicit theory;
 wenzelm parents: 
32155diff
changeset | 104 | ); | 
| 1722 | 105 | |
| 32177 
bc02c5bfcb5b
renamed functor SplitterFun to Splitter, require explicit theory;
 wenzelm parents: 
32155diff
changeset | 106 | val split_tac = Splitter.split_tac; | 
| 5304 | 107 | val split_inside_tac = Splitter.split_inside_tac; | 
| 32177 
bc02c5bfcb5b
renamed functor SplitterFun to Splitter, require explicit theory;
 wenzelm parents: 
32155diff
changeset | 108 | val split_asm_tac = Splitter.split_asm_tac; | 
| 4325 | 109 | |
| 110 | ||
| 2074 
30a65172e003
Added the de Morgan laws (incl quantifier versions) to basic simpset
 paulson parents: 
2065diff
changeset | 111 | (*** Standard simpsets ***) | 
| 
30a65172e003
Added the de Morgan laws (incl quantifier versions) to basic simpset
 paulson parents: 
2065diff
changeset | 112 | |
| 26288 
89b9f7c18631
eliminated out-of-scope proofs (cf. theory IFOL and FOL);
 wenzelm parents: 
22822diff
changeset | 113 | 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 | 114 | |
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
45659diff
changeset | 115 | fun unsafe_solver ctxt = | 
| 59498 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 wenzelm parents: 
58963diff
changeset | 116 | 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 | 117 | assume_tac ctxt, | 
| 59498 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 wenzelm parents: 
58963diff
changeset | 118 |     eresolve_tac ctxt @{thms FalseE}];
 | 
| 43597 | 119 | |
| 2633 
37c0b5a7ee5d
added delcongs, Delcongs, unsafe_solver, safe_solver, FOL_basic_ss,
 oheimb parents: 
2601diff
changeset | 120 | (*No premature instantiation of variables during simplification*) | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
45659diff
changeset | 121 | fun safe_solver ctxt = | 
| 58957 | 122 | FIRST' [match_tac ctxt (triv_rls @ Simplifier.prems_of ctxt), | 
| 123 |     eq_assume_tac, ematch_tac ctxt @{thms FalseE}];
 | |
| 2633 
37c0b5a7ee5d
added delcongs, Delcongs, unsafe_solver, safe_solver, FOL_basic_ss,
 oheimb parents: 
2601diff
changeset | 124 | |
| 3910 | 125 | (*No simprules, but basic infastructure for simplification*) | 
| 17892 | 126 | val FOL_basic_ss = | 
| 69593 | 127 | empty_simpset \<^context> | 
| 10431 | 128 | setSSolver (mk_solver "FOL safe" safe_solver) | 
| 129 | 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 | 130 | |> 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 | 131 | |> Simplifier.set_mksimps (mksimps mksimps_pairs) | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
45659diff
changeset | 132 | |> Simplifier.set_mkcong mk_meta_cong | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
45659diff
changeset | 133 | |> simpset_of; | 
| 5304 | 134 | |
| 54998 | 135 | fun unfold_tac ctxt ths = | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
45659diff
changeset | 136 | ALLGOALS (full_simp_tac (clear_simpset (put_simpset FOL_basic_ss ctxt) addsimps ths)); | 
| 17002 | 137 | |
| 2633 
37c0b5a7ee5d
added delcongs, Delcongs, unsafe_solver, safe_solver, FOL_basic_ss,
 oheimb parents: 
2601diff
changeset | 138 | |
| 5219 | 139 | (*** integration of simplifier with classical reasoner ***) | 
| 2633 
37c0b5a7ee5d
added delcongs, Delcongs, unsafe_solver, safe_solver, FOL_basic_ss,
 oheimb parents: 
2601diff
changeset | 140 | |
| 42478 | 141 | structure Clasimp = Clasimp | 
| 142 | ( | |
| 143 | structure Simplifier = Simplifier | |
| 144 | and Splitter = Splitter | |
| 145 | and Classical = Cla | |
| 146 | and Blast = Blast | |
| 147 |   val iffD1 = @{thm iffD1}
 | |
| 148 |   val iffD2 = @{thm iffD2}
 | |
| 149 |   val notE = @{thm notE}
 | |
| 150 | ); | |
| 4652 
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
 oheimb parents: 
4649diff
changeset | 151 | open Clasimp; | 
| 2633 
37c0b5a7ee5d
added delcongs, Delcongs, unsafe_solver, safe_solver, FOL_basic_ss,
 oheimb parents: 
2601diff
changeset | 152 |