| author | desharna | 
| Mon, 10 Jun 2024 13:44:46 +0200 | |
| changeset 80321 | 31b9dfbe534c | 
| parent 77879 | dd222e2af01a | 
| child 82290 | a7216319c0bb | 
| permissions | -rw-r--r-- | 
| 37744 | 1 | (* Title: HOL/Decision_Procs/ferrante_rackoff.ML | 
| 23466 | 2 | Author: Amine Chaieb, TU Muenchen | 
| 3 | ||
| 4 | Ferrante and Rackoff's algorithm for quantifier elimination in dense | |
| 5 | linear orders. Proof-synthesis and tactic. | |
| 6 | *) | |
| 7 | ||
| 23567 | 8 | signature FERRANTE_RACKOFF = | 
| 23466 | 9 | sig | 
| 23567 | 10 | val dlo_conv: Proof.context -> conv | 
| 23466 | 11 | val dlo_tac: Proof.context -> int -> tactic | 
| 12 | end; | |
| 13 | ||
| 14 | structure FerranteRackoff: FERRANTE_RACKOFF = | |
| 15 | struct | |
| 16 | ||
| 17 | open Ferrante_Rackoff_Data; | |
| 18 | open Conv; | |
| 19 | ||
| 23567 | 20 | type entry = {minf: thm list, pinf: thm list, nmi: thm list, npi: thm list,
 | 
| 23466 | 21 | ld: thm list, qe: thm, atoms : cterm list} * | 
| 23567 | 22 |   {isolate_conv: cterm list -> cterm -> thm,
 | 
| 23466 | 23 | whatis : cterm -> cterm -> ord, | 
| 24 | simpset : simpset}; | |
| 25 | ||
| 23567 | 26 | fun get_p1 th = | 
| 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 | 27 | funpow 2 (Thm.dest_arg o snd o Thm.dest_abs_global) | 
| 59582 | 28 | (funpow 2 Thm.dest_arg (Thm.cprop_of th)) |> Thm.dest_arg | 
| 23466 | 29 | |
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
46497diff
changeset | 30 | fun ferrack_conv ctxt | 
| 23567 | 31 |    (entr as ({minf = minf, pinf = pinf, nmi = nmi, npi = npi,
 | 
| 23466 | 32 | ld = ld, qe = qe, atoms = atoms}, | 
| 23567 | 33 |              {isolate_conv = icv, whatis = wi, simpset = simpset}):entry) =
 | 
| 34 | let | |
| 59582 | 35 | fun uset (vars as (x::vs)) p = case Thm.term_of p of | 
| 74397 | 36 | \<^Const_>\<open>HOL.conj for _ _\<close> => | 
| 23567 | 37 | let | 
| 38 | val ((b,l),r) = Thm.dest_comb p |>> Thm.dest_comb | |
| 23466 | 39 | val (lS,lth) = uset vars l val (rS, rth) = uset vars r | 
| 23567 | 40 | in (lS@rS, Drule.binop_cong_rule b lth rth) end | 
| 74397 | 41 | | \<^Const_>\<open>HOL.disj for _ _\<close> => | 
| 23567 | 42 | let | 
| 43 | val ((b,l),r) = Thm.dest_comb p |>> Thm.dest_comb | |
| 23466 | 44 | val (lS,lth) = uset vars l val (rS, rth) = uset vars r | 
| 23567 | 45 | in (lS@rS, Drule.binop_cong_rule b lth rth) end | 
| 46 | | _ => | |
| 47 | let | |
| 48 | val th = icv vars p | |
| 23466 | 49 | val p' = Thm.rhs_of th | 
| 50 | val c = wi x p' | |
| 23567 | 51 | val S = (if member (op =) [Lt, Le, Eq] c then single o Thm.dest_arg | 
| 52 | else if member (op =) [Gt, Ge] c then single o Thm.dest_arg1 | |
| 53 | else if c = NEq then single o Thm.dest_arg o Thm.dest_arg | |
| 23466 | 54 | else K []) p' | 
| 55 | in (S,th) end | |
| 56 | ||
| 23567 | 57 | val ((p1_v,p2_v),(mp1_v,mp2_v)) = | 
| 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 | 58 | funpow 2 (Thm.dest_arg o snd o Thm.dest_abs_global) | 
| 59582 | 59 | (funpow 4 Thm.dest_arg (Thm.cprop_of (hd minf))) | 
| 59058 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 wenzelm parents: 
56245diff
changeset | 60 | |> Thm.dest_binop |> apply2 Thm.dest_binop |> apfst (apply2 Thm.dest_fun) | 
| 60642 
48dd1cefb4ae
simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
 wenzelm parents: 
59970diff
changeset | 61 | |> apply2 (apply2 (dest_Var o Thm.term_of)) | 
| 23567 | 62 | |
| 63 | fun myfwd (th1, th2, th3, th4, th5) p1 p2 | |
| 64 | [(th_1,th_2,th_3,th_4,th_5), (th_1',th_2',th_3',th_4',th_5')] = | |
| 23466 | 65 | let | 
| 66 | val (mp1, mp2) = (get_p1 th_1, get_p1 th_1') | |
| 67 | val (pp1, pp2) = (get_p1 th_2, get_p1 th_2') | |
| 23567 | 68 | fun fw mi th th' th'' = | 
| 69 | let | |
| 70 | val th0 = if mi then | |
| 74282 | 71 | Drule.instantiate_normalize (TVars.empty, Vars.make [(p1_v, p1),(p2_v, p2),(mp1_v, mp1), (mp2_v, mp2)]) th | 
| 72 | else Drule.instantiate_normalize (TVars.empty, Vars.make [(p1_v, p1),(p2_v, p2),(mp1_v, pp1), (mp2_v, pp2)]) th | |
| 36945 | 73 | in Thm.implies_elim (Thm.implies_elim th0 th') th'' end | 
| 23567 | 74 | in (fw true th1 th_1 th_1', fw false th2 th_2 th_2', | 
| 75 | fw true th3 th_3 th_3', fw false th4 th_4 th_4', fw true th5 th_5 th_5') | |
| 23466 | 76 | end | 
| 60642 
48dd1cefb4ae
simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
 wenzelm parents: 
59970diff
changeset | 77 | val U_v = dest_Var (Thm.term_of (Thm.dest_arg (Thm.dest_arg (Thm.dest_arg1 (Thm.cprop_of qe))))) | 
| 23567 | 78 | fun main vs p = | 
| 79 | let | |
| 59582 | 80 | val ((xn,ce),(x,fm)) = (case Thm.term_of p 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 | 81 | \<^Const_>\<open>Ex _ for \<open>Abs (xn, _, _)\<close>\<close> => | 
| 
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 | 82 | Thm.dest_comb p ||> Thm.dest_abs_global |>> pair xn | 
| 23567 | 83 |                  | _ => raise CTERM ("main QE only treats existential quantifiers!", [p]))
 | 
| 59586 | 84 | val cT = Thm.ctyp_of_cterm x | 
| 23466 | 85 | val (u,nth) = uset (x::vs) fm |>> distinct (op aconvc) | 
| 86 | val nthx = Thm.abstract_rule xn x nth | |
| 87 | val q = Thm.rhs_of nth | |
| 88 | val qx = Thm.rhs_of nthx | |
| 89 | val enth = Drule.arg_cong_rule ce nthx | |
| 60801 | 90 |    val [th0,th1] = map (Thm.instantiate' [SOME cT] []) @{thms "finite.intros"}
 | 
| 23567 | 91 | fun ins x th = | 
| 60801 | 92 | Thm.implies_elim (Thm.instantiate' [] [(SOME o Thm.dest_arg o Thm.dest_arg) | 
| 23466 | 93 | (Thm.cprop_of th), SOME x] th1) th | 
| 94 | val fU = fold ins u th0 | |
| 95 | val cU = funpow 2 Thm.dest_arg (Thm.cprop_of fU) | |
| 23567 | 96 | local | 
| 60801 | 97 |      val insI1 = Thm.instantiate' [SOME cT] [] @{thm "insertI1"}
 | 
| 98 |      val insI2 = Thm.instantiate' [SOME cT] [] @{thm "insertI2"}
 | |
| 23466 | 99 | in | 
| 23567 | 100 | fun provein x S = | 
| 59582 | 101 | case Thm.term_of S of | 
| 74397 | 102 |         \<^Const_>\<open>Orderings.bot _\<close> => raise CTERM ("provein : not a member!", [S])
 | 
| 103 | | \<^Const_>\<open>insert _ for y _\<close> => | |
| 23466 | 104 | let val (cy,S') = Thm.dest_binop S | 
| 60801 | 105 | in if Thm.term_of x aconv y then Thm.instantiate' [] [SOME x, SOME S'] insI1 | 
| 106 | else Thm.implies_elim (Thm.instantiate' [] [SOME x, SOME S', SOME cy] insI2) | |
| 23466 | 107 | (provein x S') | 
| 108 | end | |
| 109 | end | |
| 59582 | 110 | val tabU = fold (fn t => fn tab => Termtab.update (Thm.term_of t, provein t cU) tab) | 
| 23466 | 111 | u Termtab.empty | 
| 59582 | 112 | val U = the o Termtab.lookup tabU o Thm.term_of | 
| 23567 | 113 | val [minf_conj, minf_disj, minf_eq, minf_neq, minf_lt, | 
| 23466 | 114 | minf_le, minf_gt, minf_ge, minf_P] = minf | 
| 23567 | 115 | val [pinf_conj, pinf_disj, pinf_eq, pinf_neq, pinf_lt, | 
| 23466 | 116 | pinf_le, pinf_gt, pinf_ge, pinf_P] = pinf | 
| 23567 | 117 | val [nmi_conj, nmi_disj, nmi_eq, nmi_neq, nmi_lt, | 
| 77879 | 118 | nmi_le, nmi_gt, nmi_ge, nmi_P] = map (Drule.instantiate_normalize (TVars.empty, Vars.make1 (U_v,cU))) nmi | 
| 23567 | 119 | val [npi_conj, npi_disj, npi_eq, npi_neq, npi_lt, | 
| 77879 | 120 | npi_le, npi_gt, npi_ge, npi_P] = map (Drule.instantiate_normalize (TVars.empty, Vars.make1 (U_v,cU))) npi | 
| 23567 | 121 | val [ld_conj, ld_disj, ld_eq, ld_neq, ld_lt, | 
| 77879 | 122 | ld_le, ld_gt, ld_ge, ld_P] = map (Drule.instantiate_normalize (TVars.empty, Vars.make1 (U_v,cU))) ld | 
| 23567 | 123 | |
| 124 | fun decomp_mpinf fm = | |
| 59582 | 125 | case Thm.term_of fm of | 
| 74397 | 126 | \<^Const_>\<open>HOL.conj for _ _\<close> => | 
| 23567 | 127 | let val (p,q) = Thm.dest_binop fm | 
| 128 | in ([p,q], myfwd (minf_conj,pinf_conj, nmi_conj, npi_conj,ld_conj) | |
| 46497 
89ccf66aa73d
renamed Thm.capply to Thm.apply, and Thm.cabs to Thm.lambda in conformance with similar operations in structure Term and Logic;
 wenzelm parents: 
45654diff
changeset | 129 | (Thm.lambda x p) (Thm.lambda x q)) | 
| 23466 | 130 | end | 
| 74397 | 131 | | \<^Const_>\<open>HOL.disj for _ _\<close> => | 
| 23567 | 132 | let val (p,q) = Thm.dest_binop fm | 
| 23466 | 133 | in ([p,q],myfwd (minf_disj, pinf_disj, nmi_disj, npi_disj,ld_disj) | 
| 46497 
89ccf66aa73d
renamed Thm.capply to Thm.apply, and Thm.cabs to Thm.lambda in conformance with similar operations in structure Term and Logic;
 wenzelm parents: 
45654diff
changeset | 134 | (Thm.lambda x p) (Thm.lambda x q)) | 
| 23466 | 135 | end | 
| 23567 | 136 | | _ => | 
| 23466 | 137 | (let val c = wi x fm | 
| 23567 | 138 | val t = (if c=Nox then I | 
| 139 | else if member (op =) [Lt, Le, Eq] c then Thm.dest_arg | |
| 140 | else if member (op =) [Gt, Ge] c then Thm.dest_arg1 | |
| 141 | else if c = NEq then (Thm.dest_arg o Thm.dest_arg) | |
| 40314 
b5ec88d9ac03
replaced ancient sys_error by raise Fail, assuming that the latter is not handled specifically by the environment;
 wenzelm parents: 
38864diff
changeset | 142 | else raise Fail "decomp_mpinf: Impossible case!!") fm | 
| 23567 | 143 | val [mi_th, pi_th, nmi_th, npi_th, ld_th] = | 
| 60801 | 144 | if c = Nox then map (Thm.instantiate' [] [SOME fm]) | 
| 23466 | 145 | [minf_P, pinf_P, nmi_P, npi_P, ld_P] | 
| 23567 | 146 | else | 
| 147 | let val [mi_th,pi_th,nmi_th,npi_th,ld_th] = | |
| 60801 | 148 | map (Thm.instantiate' [] [SOME t]) | 
| 23466 | 149 | (case c of Lt => [minf_lt, pinf_lt, nmi_lt, npi_lt, ld_lt] | 
| 150 | | Le => [minf_le, pinf_le, nmi_le, npi_le, ld_le] | |
| 151 | | Gt => [minf_gt, pinf_gt, nmi_gt, npi_gt, ld_gt] | |
| 152 | | Ge => [minf_ge, pinf_ge, nmi_ge, npi_ge, ld_ge] | |
| 153 | | Eq => [minf_eq, pinf_eq, nmi_eq, npi_eq, ld_eq] | |
| 154 | | NEq => [minf_neq, pinf_neq, nmi_neq, npi_neq, ld_neq]) | |
| 155 | val tU = U t | |
| 36945 | 156 | fun Ufw th = Thm.implies_elim th tU | 
| 23466 | 157 | in [mi_th, pi_th, Ufw nmi_th, Ufw npi_th, Ufw ld_th] | 
| 158 | end | |
| 159 | in ([], K (mi_th, pi_th, nmi_th, npi_th, ld_th)) end) | |
| 160 | val (minf_th, pinf_th, nmi_th, npi_th, ld_th) = divide_and_conquer decomp_mpinf q | |
| 23567 | 161 | val qe_th = Drule.implies_elim_list | 
| 162 | ((fconv_rule (Thm.beta_conversion true)) | |
| 60801 | 163 | (Thm.instantiate' [] (map SOME [cU, qx, get_p1 minf_th, get_p1 pinf_th]) | 
| 23466 | 164 | qe)) | 
| 23567 | 165 | [fU, ld_th, nmi_th, npi_th, minf_th, pinf_th] | 
| 166 | val bex_conv = | |
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
46497diff
changeset | 167 |       Simplifier.rewrite (put_simpset HOL_basic_ss ctxt addsimps @{thms simp_thms bex_simps(1-5)})
 | 
| 36945 | 168 | val result_th = fconv_rule (arg_conv bex_conv) (Thm.transitive enth qe_th) | 
| 23466 | 169 | in result_th | 
| 170 | end | |
| 171 | ||
| 172 | in main | |
| 173 | end; | |
| 174 | ||
| 23567 | 175 | val grab_atom_bop = | 
| 176 | 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 | 177 | fun h ctxt tm = | 
| 59582 | 178 | (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 | 179 | \<^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 | 180 | | \<^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 | 181 | | \<^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 | 182 | | \<^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 | 183 | | \<^Const_>\<open>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 | 184 | | \<^Const_>\<open>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 | 185 | | \<^Const_>\<open>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 | 186 | | \<^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 | 187 | | \<^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 | 188 | | \<^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 | 189 | | \<^Const_>\<open>Trueprop for _\<close> => h ctxt (Thm.dest_arg tm) | 
| 23466 | 190 | | _ => 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 | 191 | 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 | 192 | (h ctxt (Thm.dest_arg tm) handle CTERM _ => 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 | 193 | 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 | 194 | 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 | 195 | in h ctxt' b' end; | 
| 23466 | 196 | in h end; | 
| 197 | ||
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
46497diff
changeset | 198 | fun raw_ferrack_qe_conv ctxt (thy, {isolate_conv, whatis, simpset = ss}) tm =
 | 
| 23567 | 199 | let | 
| 200 | val ss' = | |
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
46497diff
changeset | 201 | merge_ss (simpset_of | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
46497diff
changeset | 202 | (put_simpset HOL_basic_ss ctxt addsimps | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
46497diff
changeset | 203 |         @{thms simp_thms ex_simps all_simps not_all all_not_ex ex_disj_distrib}), ss);
 | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
46497diff
changeset | 204 | val pcv = Simplifier.rewrite (put_simpset ss' ctxt); | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
46497diff
changeset | 205 | val postcv = Simplifier.rewrite (put_simpset ss ctxt); | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
46497diff
changeset | 206 | val nnf = K (nnf_conv ctxt then_conv postcv) | 
| 74274 | 207 | val env = Cterms.list_set_rev (Cterms.build (Drule.add_frees_cterm tm)) | 
| 74269 | 208 | val qe_conv = Qelim.gen_qelim_conv ctxt pcv postcv pcv cons env | 
| 23466 | 209 | (isolate_conv ctxt) nnf | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
46497diff
changeset | 210 |                   (fn vs => ferrack_conv ctxt (thy,{isolate_conv = isolate_conv ctxt,
 | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
46497diff
changeset | 211 | whatis = whatis, simpset = ss}) vs | 
| 23466 | 212 | then_conv postcv) | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
46497diff
changeset | 213 | in (Simplifier.rewrite (put_simpset ss ctxt) then_conv qe_conv) tm end; | 
| 23466 | 214 | |
| 23567 | 215 | 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 | 216 | Ferrante_Rackoff_Data.match ctxt (grab_atom_bop ctxt tm); | 
| 23486 | 217 | |
| 23567 | 218 | fun dlo_conv ctxt tm = | 
| 219 | (case dlo_instance ctxt tm of | |
| 220 |     NONE => raise CTERM ("ferrackqe_conv: no corresponding instance in context!", [tm])
 | |
| 221 | | SOME instance => raw_ferrack_qe_conv ctxt instance tm); | |
| 23466 | 222 | |
| 23567 | 223 | fun dlo_tac ctxt = CSUBGOAL (fn (p, i) => | 
| 224 | (case dlo_instance ctxt p of | |
| 225 | NONE => no_tac | |
| 226 | | SOME instance => | |
| 54742 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 wenzelm parents: 
51717diff
changeset | 227 | Object_Logic.full_atomize_tac ctxt i THEN | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
46497diff
changeset | 228 | simp_tac (put_simpset (#simpset (snd instance)) ctxt) i THEN (* FIXME already part of raw_ferrack_qe_conv? *) | 
| 59970 | 229 | CONVERSION (Object_Logic.judgment_conv ctxt (raw_ferrack_qe_conv ctxt instance)) i THEN | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
46497diff
changeset | 230 | simp_tac ctxt i)); (* FIXME really? *) | 
| 23466 | 231 | |
| 232 | end; |