| author | Fabian Huch <huch@in.tum.de> | 
| Wed, 18 Oct 2023 20:12:07 +0200 | |
| changeset 78843 | fc3ba0a1c82f | 
| parent 78808 | 64973b03b778 | 
| child 80703 | cc4ecaa8e96e | 
| permissions | -rw-r--r-- | 
| 37744 | 1 | (* Title: HOL/Decision_Procs/langford.ML | 
| 29269 
5c25a2012975
moved term order operations to structure TermOrd (cf. Pure/term_ord.ML);
 wenzelm parents: 
29265diff
changeset | 2 | Author: Amine Chaieb, TU Muenchen | 
| 
5c25a2012975
moved term order operations to structure TermOrd (cf. Pure/term_ord.ML);
 wenzelm parents: 
29265diff
changeset | 3 | *) | 
| 
5c25a2012975
moved term order operations to structure TermOrd (cf. Pure/term_ord.ML);
 wenzelm parents: 
29265diff
changeset | 4 | |
| 55848 | 5 | signature LANGFORD = | 
| 23906 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 6 | sig | 
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 7 | val dlo_tac : Proof.context -> int -> tactic | 
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 8 | val dlo_conv : Proof.context -> cterm -> thm | 
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 9 | end | 
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 10 | |
| 55848 | 11 | structure Langford: LANGFORD = | 
| 23906 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 12 | struct | 
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 13 | |
| 30452 | 14 | val dest_set = | 
| 55792 | 15 | let | 
| 16 | fun h acc ct = | |
| 59582 | 17 | (case Thm.term_of ct of | 
| 74397 | 18 | \<^Const_>\<open>bot _\<close> => acc | 
| 19 | | \<^Const_>\<open>insert _ for _ _\<close> => h (Thm.dest_arg1 ct :: acc) (Thm.dest_arg ct)); | |
| 55792 | 20 | in h [] end; | 
| 30452 | 21 | |
| 55506 | 22 | fun prove_finite cT u = | 
| 55792 | 23 | let | 
| 60801 | 24 |     val [th0, th1] = map (Thm.instantiate' [SOME cT] []) @{thms finite.intros}
 | 
| 23906 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 25 | fun ins x th = | 
| 55792 | 26 | Thm.implies_elim | 
| 60801 | 27 | (Thm.instantiate' [] [(SOME o Thm.dest_arg o Thm.dest_arg) (Thm.cprop_of th), SOME x] th1) th | 
| 55792 | 28 | in fold ins u th0 end; | 
| 23906 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 29 | |
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
46497diff
changeset | 30 | fun simp_rule ctxt = | 
| 45654 | 31 | Conv.fconv_rule | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
46497diff
changeset | 32 | (Conv.arg_conv | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
46497diff
changeset | 33 |       (Simplifier.rewrite (put_simpset HOL_basic_ss ctxt addsimps @{thms ball_simps simp_thms})));
 | 
| 23906 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 34 | |
| 55506 | 35 | fun basic_dloqe ctxt stupid dlo_qeth dlo_qeth_nolb dlo_qeth_noub gather ep = | 
| 59582 | 36 | (case Thm.term_of ep of | 
| 74397 | 37 | \<^Const_>\<open>Ex _ for _\<close> => | 
| 55846 | 38 | let | 
| 39 | val p = Thm.dest_arg ep | |
| 40 | val ths = | |
| 41 | simplify (put_simpset HOL_basic_ss ctxt addsimps gather) | |
| 60801 | 42 | (Thm.instantiate' [] [SOME p] stupid) | 
| 55846 | 43 | val (L, U) = | 
| 74525 
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
 wenzelm parents: 
74518diff
changeset | 44 | let val (_, q) = Thm.dest_abs_global (Thm.dest_arg (Thm.rhs_of ths)) | 
| 55846 | 45 | in (Thm.dest_arg1 q |> Thm.dest_arg1, Thm.dest_arg q |> Thm.dest_arg1) end | 
| 46 | fun proveneF S = | |
| 47 | let | |
| 48 | val (a, A) = Thm.dest_comb S |>> Thm.dest_arg | |
| 59586 | 49 | val cT = Thm.ctyp_of_cterm a | 
| 74614 | 50 |             val ne = \<^instantiate>\<open>'a = cT and a and A in lemma \<open>insert a A \<noteq> {}\<close> by simp\<close>
 | 
| 55846 | 51 | val f = prove_finite cT (dest_set S) | 
| 52 | in (ne, f) end | |
| 23906 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 53 | |
| 55846 | 54 | val qe = | 
| 59582 | 55 | (case (Thm.term_of L, Thm.term_of U) of | 
| 74397 | 56 | (\<^Const_>\<open>bot _\<close>, _) => | 
| 55846 | 57 | let val (neU, fU) = proveneF U | 
| 58 | in simp_rule ctxt (Thm.transitive ths (dlo_qeth_nolb OF [neU, fU])) end | |
| 74397 | 59 | | (_, \<^Const_>\<open>bot _\<close>) => | 
| 55846 | 60 | let val (neL,fL) = proveneF L | 
| 61 | in simp_rule ctxt (Thm.transitive ths (dlo_qeth_noub OF [neL, fL])) end | |
| 62 | | _ => | |
| 63 | let | |
| 64 | val (neL, fL) = proveneF L | |
| 65 | val (neU, fU) = proveneF U | |
| 66 | in simp_rule ctxt (Thm.transitive ths (dlo_qeth OF [neL, neU, fL, fU])) end) | |
| 67 | in qe end | |
| 68 | | _ => error "dlo_qe : Not an existential formula"); | |
| 23906 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 69 | |
| 55506 | 70 | val all_conjuncts = | 
| 55846 | 71 | let | 
| 72 | fun h acc ct = | |
| 59582 | 73 | (case Thm.term_of ct of | 
| 74397 | 74 | \<^Const>\<open>HOL.conj for _ _\<close> => h (h acc (Thm.dest_arg ct)) (Thm.dest_arg1 ct) | 
| 55846 | 75 | | _ => ct :: acc) | 
| 76 | in h [] end; | |
| 23906 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 77 | |
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 78 | fun conjuncts ct = | 
| 59582 | 79 | (case Thm.term_of ct of | 
| 74397 | 80 | \<^Const>\<open>HOL.conj for _ _\<close> => Thm.dest_arg1 ct :: conjuncts (Thm.dest_arg ct) | 
| 55846 | 81 | | _ => [ct]); | 
| 23906 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 82 | |
| 55846 | 83 | val list_conj = | 
| 74570 | 84 | foldr1 (fn (A, B) => \<^instantiate>\<open>A and B in cterm \<open>A \<and> B\<close>\<close>); | 
| 23906 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 85 | |
| 55506 | 86 | fun mk_conj_tab th = | 
| 55846 | 87 | let | 
| 88 | fun h acc th = | |
| 59582 | 89 | (case Thm.prop_of th of | 
| 74397 | 90 | \<^Const_>\<open>Trueprop for \<^Const_>\<open>HOL.conj for p q\<close>\<close> => | 
| 55846 | 91 | h (h acc (th RS conjunct2)) (th RS conjunct1) | 
| 74397 | 92 | | \<^Const_>\<open>Trueprop for p\<close> => (p, th) :: acc) | 
| 55846 | 93 | in fold (Termtab.insert Thm.eq_thm) (h [] th) Termtab.empty end; | 
| 23906 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 94 | |
| 74397 | 95 | fun is_conj \<^Const_>\<open>HOL.conj for _ _\<close> = true | 
| 23906 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 96 | | is_conj _ = false; | 
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 97 | |
| 55506 | 98 | fun prove_conj tab cjs = | 
| 55846 | 99 | (case cjs of | 
| 100 | [c] => | |
| 59582 | 101 | if is_conj (Thm.term_of c) | 
| 55846 | 102 | then prove_conj tab (conjuncts c) | 
| 103 | else tab c | |
| 104 | | c :: cs => conjI OF [prove_conj tab [c], prove_conj tab cs]); | |
| 23906 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 105 | |
| 55506 | 106 | fun conj_aci_rule eq = | 
| 55846 | 107 | let | 
| 108 | val (l, r) = Thm.dest_equals eq | |
| 59582 | 109 | fun tabl c = the (Termtab.lookup (mk_conj_tab (Thm.assume l)) (Thm.term_of c)) | 
| 110 | fun tabr c = the (Termtab.lookup (mk_conj_tab (Thm.assume r)) (Thm.term_of c)) | |
| 55846 | 111 | val ll = Thm.dest_arg l | 
| 112 | val rr = Thm.dest_arg r | |
| 113 | val thl = prove_conj tabl (conjuncts rr) |> Drule.implies_intr_hyps | |
| 114 | val thr = prove_conj tabr (conjuncts ll) |> Drule.implies_intr_hyps | |
| 74614 | 115 | val eqI = | 
| 116 | \<^instantiate>\<open>P = ll and Q = rr in lemma \<open>(P \<Longrightarrow> Q) \<Longrightarrow> (Q \<Longrightarrow> P) \<Longrightarrow> P \<longleftrightarrow> Q\<close> by (rule iffI)\<close> | |
| 55846 | 117 | in Thm.implies_elim (Thm.implies_elim eqI thl) thr |> mk_meta_eq end; | 
| 55506 | 118 | |
| 55846 | 119 | fun contains x ct = | 
| 59582 | 120 | member (op aconv) (Misc_Legacy.term_frees (Thm.term_of ct)) (Thm.term_of x); | 
| 23906 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 121 | |
| 55846 | 122 | fun is_eqx x eq = | 
| 59582 | 123 | (case Thm.term_of eq of | 
| 74397 | 124 | \<^Const_>\<open>HOL.eq _ for l r\<close> => | 
| 59582 | 125 | l aconv Thm.term_of x orelse r aconv Thm.term_of x | 
| 55846 | 126 | | _ => false); | 
| 23906 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 127 | |
| 55506 | 128 | local | 
| 55846 | 129 | |
| 78808 
64973b03b778
more standard simproc_setup using ML antiquotation;
 wenzelm parents: 
74614diff
changeset | 130 | fun reduce_ex_proc ctxt ct = | 
| 59582 | 131 | (case Thm.term_of ct of | 
| 74525 
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
 wenzelm parents: 
74518diff
changeset | 132 | \<^Const_>\<open>Ex _ for \<open>Abs _\<close>\<close> => | 
| 55846 | 133 | let | 
| 134 | val e = Thm.dest_fun ct | |
| 74525 
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
 wenzelm parents: 
74518diff
changeset | 135 | val (x,p) = Thm.dest_abs_global (Thm.dest_arg ct) | 
| 
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
 wenzelm parents: 
74518diff
changeset | 136 | val Free (xn, _) = Thm.term_of x | 
| 55846 | 137 | val (eqs,neqs) = List.partition (is_eqx x) (all_conjuncts p) | 
| 138 | in | |
| 139 | (case eqs of | |
| 140 | [] => | |
| 141 | let | |
| 142 | val (dx, ndx) = List.partition (contains x) neqs | |
| 143 | in | |
| 144 | case ndx of | |
| 145 | [] => NONE | |
| 146 | | _ => | |
| 74569 | 147 | conj_aci_rule | 
| 74570 | 148 | \<^instantiate>\<open>A = p and B = \<open>list_conj (ndx @ dx)\<close> in cterm \<open>Trueprop A \<equiv> Trueprop B\<close>\<close> | 
| 55846 | 149 | |> Thm.abstract_rule xn x | 
| 150 | |> Drule.arg_cong_rule e | |
| 151 | |> Conv.fconv_rule | |
| 152 | (Conv.arg_conv | |
| 153 | (Simplifier.rewrite | |
| 154 |                       (put_simpset HOL_basic_ss ctxt addsimps @{thms simp_thms ex_simps})))
 | |
| 155 | |> SOME | |
| 156 | end | |
| 157 | | _ => | |
| 74569 | 158 | conj_aci_rule | 
| 74570 | 159 | \<^instantiate>\<open>A = p and B = \<open>list_conj (eqs @ neqs)\<close> in cterm \<open>Trueprop A \<equiv> Trueprop B\<close>\<close> | 
| 55846 | 160 | |> Thm.abstract_rule xn x |> Drule.arg_cong_rule e | 
| 161 | |> Conv.fconv_rule | |
| 162 | (Conv.arg_conv | |
| 163 | (Simplifier.rewrite | |
| 164 |                     (put_simpset HOL_basic_ss ctxt addsimps @{thms simp_thms ex_simps})))
 | |
| 165 | |> SOME) | |
| 166 | end | |
| 167 | | _ => NONE); | |
| 168 | ||
| 169 | in | |
| 170 | ||
| 78808 
64973b03b778
more standard simproc_setup using ML antiquotation;
 wenzelm parents: 
74614diff
changeset | 171 | val reduce_ex_simproc = \<^simproc_setup>\<open>passive reduce_ex ("\<exists>x. P x") = \<open>K reduce_ex_proc\<close>\<close>;
 | 
| 55846 | 172 | |
| 23906 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 173 | end; | 
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 174 | |
| 55506 | 175 | fun raw_dlo_conv ctxt dlo_ss ({qe_bnds, qe_nolb, qe_noub, gst, gs, ...}: Langford_Data.entry) =
 | 
| 55846 | 176 | let | 
| 55847 
c38ad094e5bf
silence warning due to addsimps @{thms dnf_simps}: duplicate not_not rule via simp_thms and nnf_simps;
 wenzelm parents: 
55846diff
changeset | 177 | val ctxt' = | 
| 
c38ad094e5bf
silence warning due to addsimps @{thms dnf_simps}: duplicate not_not rule via simp_thms and nnf_simps;
 wenzelm parents: 
55846diff
changeset | 178 | Context_Position.set_visible false (put_simpset dlo_ss ctxt) | 
| 
c38ad094e5bf
silence warning due to addsimps @{thms dnf_simps}: duplicate not_not rule via simp_thms and nnf_simps;
 wenzelm parents: 
55846diff
changeset | 179 |         addsimps @{thms dnf_simps} addsimprocs [reduce_ex_simproc]
 | 
| 55846 | 180 | val dnfex_conv = Simplifier.rewrite ctxt' | 
| 181 | val pcv = | |
| 182 | Simplifier.rewrite | |
| 183 | (put_simpset dlo_ss ctxt | |
| 184 |           addsimps @{thms simp_thms ex_simps all_simps all_not_ex not_all ex_disj_distrib})
 | |
| 74274 | 185 | val mk_env = Cterms.list_set_rev o Cterms.build o Drule.add_frees_cterm | 
| 55846 | 186 | in | 
| 187 | fn p => | |
| 61075 | 188 | Qelim.gen_qelim_conv ctxt pcv pcv dnfex_conv cons | 
| 74269 | 189 | (mk_env p) (K Thm.reflexive) (K Thm.reflexive) | 
| 55846 | 190 | (K (basic_dloqe ctxt gst qe_bnds qe_nolb qe_noub gs)) p | 
| 191 | end; | |
| 23906 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 192 | |
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 193 | val grab_atom_bop = | 
| 55846 | 194 | let | 
| 74525 
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
 wenzelm parents: 
74518diff
changeset | 195 | fun h ctxt tm = | 
| 59582 | 196 | (case Thm.term_of tm of | 
| 74525 
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
 wenzelm parents: 
74518diff
changeset | 197 | \<^Const_>\<open>HOL.eq \<^Type>\<open>bool\<close> for _ _\<close> => find_args ctxt tm | 
| 
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
 wenzelm parents: 
74518diff
changeset | 198 | | \<^Const_>\<open>Not for _\<close> => h ctxt (Thm.dest_arg tm) | 
| 
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
 wenzelm parents: 
74518diff
changeset | 199 | | \<^Const_>\<open>All _ for _\<close> => find_body ctxt (Thm.dest_arg tm) | 
| 
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
 wenzelm parents: 
74518diff
changeset | 200 | | \<^Const_>\<open>Pure.all _ for _\<close> => find_body ctxt (Thm.dest_arg tm) | 
| 
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
 wenzelm parents: 
74518diff
changeset | 201 | | \<^Const_>\<open>Ex _ for _\<close> => find_body ctxt (Thm.dest_arg tm) | 
| 
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
 wenzelm parents: 
74518diff
changeset | 202 | | \<^Const_>\<open>HOL.conj for _ _\<close> => find_args ctxt tm | 
| 
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
 wenzelm parents: 
74518diff
changeset | 203 | | \<^Const_>\<open>HOL.disj for _ _\<close> => find_args ctxt tm | 
| 
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
 wenzelm parents: 
74518diff
changeset | 204 | | \<^Const_>\<open>HOL.implies for _ _\<close> => find_args ctxt tm | 
| 
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
 wenzelm parents: 
74518diff
changeset | 205 | | \<^Const_>\<open>Pure.imp for _ _\<close> => find_args ctxt tm | 
| 
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
 wenzelm parents: 
74518diff
changeset | 206 | | \<^Const_>\<open>Pure.eq _ for _ _\<close> => find_args ctxt tm | 
| 
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
 wenzelm parents: 
74518diff
changeset | 207 | | \<^Const_>\<open>Trueprop for _\<close> => h ctxt (Thm.dest_arg tm) | 
| 55846 | 208 | | _ => Thm.dest_fun2 tm) | 
| 74525 
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
 wenzelm parents: 
74518diff
changeset | 209 | and find_args ctxt tm = | 
| 
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
 wenzelm parents: 
74518diff
changeset | 210 | (h ctxt (Thm.dest_arg tm) handle CTERM _ => h ctxt (Thm.dest_arg1 tm)) | 
| 
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
 wenzelm parents: 
74518diff
changeset | 211 | and find_body ctxt b = | 
| 
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
 wenzelm parents: 
74518diff
changeset | 212 | let val ((_, b'), ctxt') = Variable.dest_abs_cterm b ctxt | 
| 
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
 wenzelm parents: 
74518diff
changeset | 213 | in h ctxt' b' end; | 
| 55846 | 214 | in h end; | 
| 23906 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 215 | |
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 216 | fun dlo_instance ctxt tm = | 
| 74525 
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
 wenzelm parents: 
74518diff
changeset | 217 | (fst (Langford_Data.get ctxt), Langford_Data.match ctxt (grab_atom_bop ctxt tm)); | 
| 23906 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 218 | |
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 219 | fun dlo_conv ctxt tm = | 
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 220 | (case dlo_instance ctxt tm of | 
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 221 |     (_, NONE) => raise CTERM ("dlo_conv (langford): no corresponding instance in context!", [tm])
 | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
46497diff
changeset | 222 | | (ss, SOME instance) => raw_dlo_conv ctxt ss instance tm); | 
| 23906 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 223 | |
| 61075 | 224 | fun generalize_tac ctxt f = CSUBGOAL (fn (p, _) => PRIMITIVE (fn st => | 
| 55846 | 225 | let | 
| 61075 | 226 | fun all x t = | 
| 227 | Thm.apply (Thm.cterm_of ctxt (Logic.all_const (Thm.typ_of_cterm x))) (Thm.lambda x t) | |
| 67559 | 228 | val ts = sort Thm.fast_term_ord (f p) | 
| 61075 | 229 | val p' = fold_rev all ts p | 
| 55846 | 230 | in Thm.implies_intr p' (Thm.implies_elim st (fold Thm.forall_elim ts (Thm.assume p'))) end)); | 
| 24083 
4ea3656380b1
find_body goes under meta-quantifier ; tactic generalizes free variables;
 chaieb parents: 
23906diff
changeset | 231 | |
| 
4ea3656380b1
find_body goes under meta-quantifier ; tactic generalizes free variables;
 chaieb parents: 
23906diff
changeset | 232 | fun cfrees ats ct = | 
| 55846 | 233 | let | 
| 234 | val ins = insert (op aconvc) | |
| 235 | fun h acc t = | |
| 59582 | 236 | (case Thm.term_of t of | 
| 55846 | 237 | _ $ _ $ _ => | 
| 238 | if member (op aconvc) ats (Thm.dest_fun2 t) | |
| 239 | then ins (Thm.dest_arg t) (ins (Thm.dest_arg1 t) acc) | |
| 240 | else h (h acc (Thm.dest_arg t)) (Thm.dest_fun t) | |
| 241 | | _ $ _ => h (h acc (Thm.dest_arg t)) (Thm.dest_fun t) | |
| 74525 
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
 wenzelm parents: 
74518diff
changeset | 242 | | Abs _ => Thm.dest_abs_global t ||> h acc |> uncurry (remove (op aconvc)) | 
| 55846 | 243 | | Free _ => if member (op aconvc) ats t then acc else ins t acc | 
| 244 | | Var _ => if member (op aconvc) ats t then acc else ins t acc | |
| 245 | | _ => acc) | |
| 246 | in h [] ct end | |
| 24083 
4ea3656380b1
find_body goes under meta-quantifier ; tactic generalizes free variables;
 chaieb parents: 
23906diff
changeset | 247 | |
| 23906 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 248 | fun dlo_tac ctxt = CSUBGOAL (fn (p, i) => | 
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 249 | (case dlo_instance ctxt p of | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
46497diff
changeset | 250 | (ss, NONE) => simp_tac (put_simpset ss ctxt) i | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
46497diff
changeset | 251 | | (ss, SOME instance) => | 
| 54742 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 wenzelm parents: 
51717diff
changeset | 252 | Object_Logic.full_atomize_tac ctxt i THEN | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
46497diff
changeset | 253 | simp_tac (put_simpset ss ctxt) i | 
| 24083 
4ea3656380b1
find_body goes under meta-quantifier ; tactic generalizes free variables;
 chaieb parents: 
23906diff
changeset | 254 | THEN (CONVERSION Thm.eta_long_conversion) i | 
| 61075 | 255 | THEN (TRY o generalize_tac ctxt (cfrees (#atoms instance))) i | 
| 54742 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 wenzelm parents: 
51717diff
changeset | 256 | THEN Object_Logic.full_atomize_tac ctxt i | 
| 59970 | 257 | THEN CONVERSION (Object_Logic.judgment_conv ctxt (raw_dlo_conv ctxt ss instance)) i | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
46497diff
changeset | 258 | THEN (simp_tac (put_simpset ss ctxt) i))); | 
| 62391 | 259 | end; |