haftmann@37744: (* Title: HOL/Decision_Procs/ferrante_rackoff.ML wenzelm@23466: Author: Amine Chaieb, TU Muenchen wenzelm@23466: wenzelm@23466: Ferrante and Rackoff's algorithm for quantifier elimination in dense wenzelm@23466: linear orders. Proof-synthesis and tactic. wenzelm@23466: *) wenzelm@23466: wenzelm@23567: signature FERRANTE_RACKOFF = wenzelm@23466: sig wenzelm@23567: val dlo_conv: Proof.context -> conv wenzelm@23466: val dlo_tac: Proof.context -> int -> tactic wenzelm@23466: end; wenzelm@23466: wenzelm@23466: structure FerranteRackoff: FERRANTE_RACKOFF = wenzelm@23466: struct wenzelm@23466: wenzelm@23466: open Ferrante_Rackoff_Data; wenzelm@23466: open Conv; wenzelm@23466: wenzelm@23567: type entry = {minf: thm list, pinf: thm list, nmi: thm list, npi: thm list, wenzelm@23466: ld: thm list, qe: thm, atoms : cterm list} * wenzelm@23567: {isolate_conv: cterm list -> cterm -> thm, wenzelm@23466: whatis : cterm -> cterm -> ord, wenzelm@23466: simpset : simpset}; wenzelm@23466: wenzelm@23567: fun get_p1 th = wenzelm@23567: funpow 2 (Thm.dest_arg o snd o Thm.dest_abs NONE) wenzelm@23567: (funpow 2 Thm.dest_arg (cprop_of th)) |> Thm.dest_arg wenzelm@23466: wenzelm@51717: fun ferrack_conv ctxt wenzelm@23567: (entr as ({minf = minf, pinf = pinf, nmi = nmi, npi = npi, wenzelm@23466: ld = ld, qe = qe, atoms = atoms}, wenzelm@23567: {isolate_conv = icv, whatis = wi, simpset = simpset}):entry) = wenzelm@23567: let wenzelm@23466: fun uset (vars as (x::vs)) p = case term_of p of haftmann@38795: Const(@{const_name HOL.conj}, _)$ _ $ _ => wenzelm@23567: let wenzelm@23567: val ((b,l),r) = Thm.dest_comb p |>> Thm.dest_comb wenzelm@23466: val (lS,lth) = uset vars l val (rS, rth) = uset vars r wenzelm@23567: in (lS@rS, Drule.binop_cong_rule b lth rth) end haftmann@38795: | Const(@{const_name HOL.disj}, _)$ _ $ _ => wenzelm@23567: let wenzelm@23567: val ((b,l),r) = Thm.dest_comb p |>> Thm.dest_comb wenzelm@23466: val (lS,lth) = uset vars l val (rS, rth) = uset vars r wenzelm@23567: in (lS@rS, Drule.binop_cong_rule b lth rth) end wenzelm@23567: | _ => wenzelm@23567: let wenzelm@23567: val th = icv vars p wenzelm@23466: val p' = Thm.rhs_of th wenzelm@23466: val c = wi x p' wenzelm@23567: val S = (if member (op =) [Lt, Le, Eq] c then single o Thm.dest_arg wenzelm@23567: else if member (op =) [Gt, Ge] c then single o Thm.dest_arg1 wenzelm@23567: else if c = NEq then single o Thm.dest_arg o Thm.dest_arg wenzelm@23466: else K []) p' wenzelm@23466: in (S,th) end wenzelm@23466: wenzelm@23567: val ((p1_v,p2_v),(mp1_v,mp2_v)) = wenzelm@23567: funpow 2 (Thm.dest_arg o snd o Thm.dest_abs NONE) wenzelm@23567: (funpow 4 Thm.dest_arg (cprop_of (hd minf))) wenzelm@23567: |> Thm.dest_binop |> pairself Thm.dest_binop |> apfst (pairself Thm.dest_fun) wenzelm@23567: wenzelm@23567: fun myfwd (th1, th2, th3, th4, th5) p1 p2 wenzelm@23567: [(th_1,th_2,th_3,th_4,th_5), (th_1',th_2',th_3',th_4',th_5')] = wenzelm@23466: let wenzelm@23466: val (mp1, mp2) = (get_p1 th_1, get_p1 th_1') wenzelm@23466: val (pp1, pp2) = (get_p1 th_2, get_p1 th_2') wenzelm@23567: fun fw mi th th' th'' = wenzelm@23567: let wenzelm@23567: val th0 = if mi then wenzelm@43333: Drule.instantiate_normalize ([],[(p1_v, p1),(p2_v, p2),(mp1_v, mp1), (mp2_v, mp2)]) th wenzelm@43333: else Drule.instantiate_normalize ([],[(p1_v, p1),(p2_v, p2),(mp1_v, pp1), (mp2_v, pp2)]) th wenzelm@36945: in Thm.implies_elim (Thm.implies_elim th0 th') th'' end wenzelm@23567: in (fw true th1 th_1 th_1', fw false th2 th_2 th_2', wenzelm@23567: fw true th3 th_3 th_3', fw false th4 th_4 th_4', fw true th5 th_5 th_5') wenzelm@23466: end wenzelm@23466: val U_v = (Thm.dest_arg o Thm.dest_arg o Thm.dest_arg1) (cprop_of qe) wenzelm@23567: fun main vs p = wenzelm@23567: let wenzelm@23567: val ((xn,ce),(x,fm)) = (case term_of p of haftmann@38558: Const(@{const_name Ex},_)$Abs(xn,xT,_) => wenzelm@23466: Thm.dest_comb p ||> Thm.dest_abs (SOME xn) |>> pair xn wenzelm@23567: | _ => raise CTERM ("main QE only treats existential quantifiers!", [p])) wenzelm@23466: val cT = ctyp_of_term x wenzelm@23466: val (u,nth) = uset (x::vs) fm |>> distinct (op aconvc) wenzelm@23466: val nthx = Thm.abstract_rule xn x nth wenzelm@23466: val q = Thm.rhs_of nth wenzelm@23466: val qx = Thm.rhs_of nthx wenzelm@23466: val enth = Drule.arg_cong_rule ce nthx wenzelm@23466: val [th0,th1] = map (instantiate' [SOME cT] []) @{thms "finite.intros"} wenzelm@23567: fun ins x th = wenzelm@36945: Thm.implies_elim (instantiate' [] [(SOME o Thm.dest_arg o Thm.dest_arg) wenzelm@23466: (Thm.cprop_of th), SOME x] th1) th wenzelm@23466: val fU = fold ins u th0 wenzelm@23466: val cU = funpow 2 Thm.dest_arg (Thm.cprop_of fU) wenzelm@23567: local wenzelm@23466: val insI1 = instantiate' [SOME cT] [] @{thm "insertI1"} wenzelm@23466: val insI2 = instantiate' [SOME cT] [] @{thm "insertI2"} wenzelm@23466: in wenzelm@23567: fun provein x S = wenzelm@23466: case term_of S of haftmann@32264: Const(@{const_name Orderings.bot}, _) => raise CTERM ("provein : not a member!", [S]) haftmann@30304: | Const(@{const_name insert}, _) $ y $_ => wenzelm@23466: let val (cy,S') = Thm.dest_binop S wenzelm@23466: in if term_of x aconv y then instantiate' [] [SOME x, SOME S'] insI1 wenzelm@36945: else Thm.implies_elim (instantiate' [] [SOME x, SOME S', SOME cy] insI2) wenzelm@23466: (provein x S') wenzelm@23466: end wenzelm@23466: end wenzelm@23567: val tabU = fold (fn t => fn tab => Termtab.update (term_of t, provein t cU) tab) wenzelm@23466: u Termtab.empty wenzelm@23567: val U = the o Termtab.lookup tabU o term_of wenzelm@23567: val [minf_conj, minf_disj, minf_eq, minf_neq, minf_lt, wenzelm@23466: minf_le, minf_gt, minf_ge, minf_P] = minf wenzelm@23567: val [pinf_conj, pinf_disj, pinf_eq, pinf_neq, pinf_lt, wenzelm@23466: pinf_le, pinf_gt, pinf_ge, pinf_P] = pinf wenzelm@23567: val [nmi_conj, nmi_disj, nmi_eq, nmi_neq, nmi_lt, wenzelm@43333: nmi_le, nmi_gt, nmi_ge, nmi_P] = map (Drule.instantiate_normalize ([],[(U_v,cU)])) nmi wenzelm@23567: val [npi_conj, npi_disj, npi_eq, npi_neq, npi_lt, wenzelm@43333: npi_le, npi_gt, npi_ge, npi_P] = map (Drule.instantiate_normalize ([],[(U_v,cU)])) npi wenzelm@23567: val [ld_conj, ld_disj, ld_eq, ld_neq, ld_lt, wenzelm@43333: ld_le, ld_gt, ld_ge, ld_P] = map (Drule.instantiate_normalize ([],[(U_v,cU)])) ld wenzelm@23567: wenzelm@23567: fun decomp_mpinf fm = wenzelm@23466: case term_of fm of haftmann@38795: Const(@{const_name HOL.conj},_)$_$_ => wenzelm@23567: let val (p,q) = Thm.dest_binop fm wenzelm@23567: in ([p,q], myfwd (minf_conj,pinf_conj, nmi_conj, npi_conj,ld_conj) wenzelm@46497: (Thm.lambda x p) (Thm.lambda x q)) wenzelm@23466: end haftmann@38795: | Const(@{const_name HOL.disj},_)$_$_ => wenzelm@23567: let val (p,q) = Thm.dest_binop fm wenzelm@23466: in ([p,q],myfwd (minf_disj, pinf_disj, nmi_disj, npi_disj,ld_disj) wenzelm@46497: (Thm.lambda x p) (Thm.lambda x q)) wenzelm@23466: end wenzelm@23567: | _ => wenzelm@23466: (let val c = wi x fm wenzelm@23567: val t = (if c=Nox then I wenzelm@23567: else if member (op =) [Lt, Le, Eq] c then Thm.dest_arg wenzelm@23567: else if member (op =) [Gt, Ge] c then Thm.dest_arg1 wenzelm@23567: else if c = NEq then (Thm.dest_arg o Thm.dest_arg) wenzelm@40314: else raise Fail "decomp_mpinf: Impossible case!!") fm wenzelm@23567: val [mi_th, pi_th, nmi_th, npi_th, ld_th] = wenzelm@23567: if c = Nox then map (instantiate' [] [SOME fm]) wenzelm@23466: [minf_P, pinf_P, nmi_P, npi_P, ld_P] wenzelm@23567: else wenzelm@23567: let val [mi_th,pi_th,nmi_th,npi_th,ld_th] = wenzelm@23466: map (instantiate' [] [SOME t]) wenzelm@23466: (case c of Lt => [minf_lt, pinf_lt, nmi_lt, npi_lt, ld_lt] wenzelm@23466: | Le => [minf_le, pinf_le, nmi_le, npi_le, ld_le] wenzelm@23466: | Gt => [minf_gt, pinf_gt, nmi_gt, npi_gt, ld_gt] wenzelm@23466: | Ge => [minf_ge, pinf_ge, nmi_ge, npi_ge, ld_ge] wenzelm@23466: | Eq => [minf_eq, pinf_eq, nmi_eq, npi_eq, ld_eq] wenzelm@23466: | NEq => [minf_neq, pinf_neq, nmi_neq, npi_neq, ld_neq]) wenzelm@23466: val tU = U t wenzelm@36945: fun Ufw th = Thm.implies_elim th tU wenzelm@23466: in [mi_th, pi_th, Ufw nmi_th, Ufw npi_th, Ufw ld_th] wenzelm@23466: end wenzelm@23466: in ([], K (mi_th, pi_th, nmi_th, npi_th, ld_th)) end) wenzelm@23466: val (minf_th, pinf_th, nmi_th, npi_th, ld_th) = divide_and_conquer decomp_mpinf q wenzelm@23567: val qe_th = Drule.implies_elim_list wenzelm@23567: ((fconv_rule (Thm.beta_conversion true)) wenzelm@23567: (instantiate' [] (map SOME [cU, qx, get_p1 minf_th, get_p1 pinf_th]) wenzelm@23466: qe)) wenzelm@23567: [fU, ld_th, nmi_th, npi_th, minf_th, pinf_th] wenzelm@23567: val bex_conv = wenzelm@51717: Simplifier.rewrite (put_simpset HOL_basic_ss ctxt addsimps @{thms simp_thms bex_simps(1-5)}) wenzelm@36945: val result_th = fconv_rule (arg_conv bex_conv) (Thm.transitive enth qe_th) wenzelm@23466: in result_th wenzelm@23466: end wenzelm@23466: wenzelm@23466: in main wenzelm@23466: end; wenzelm@23466: wenzelm@23567: val grab_atom_bop = wenzelm@23567: let wenzelm@23466: fun h bounds tm = wenzelm@23466: (case term_of tm of haftmann@38864: Const (@{const_name HOL.eq}, T) $ _ $ _ => wenzelm@23567: if domain_type T = HOLogic.boolT then find_args bounds tm wenzelm@23466: else Thm.dest_fun2 tm haftmann@38558: | Const (@{const_name Not}, _) $ _ => h bounds (Thm.dest_arg tm) haftmann@38558: | Const (@{const_name All}, _) $ _ => find_body bounds (Thm.dest_arg tm) haftmann@38558: | Const (@{const_name Ex}, _) $ _ => find_body bounds (Thm.dest_arg tm) haftmann@38795: | Const (@{const_name HOL.conj}, _) $ _ $ _ => find_args bounds tm haftmann@38795: | Const (@{const_name HOL.disj}, _) $ _ $ _ => find_args bounds tm haftmann@38786: | Const (@{const_name HOL.implies}, _) $ _ $ _ => find_args bounds tm wenzelm@56245: | Const (@{const_name Pure.imp}, _) $ _ $ _ => find_args bounds tm wenzelm@56245: | Const (@{const_name Pure.eq}, _) $ _ $ _ => find_args bounds tm wenzelm@56245: | Const (@{const_name Pure.all}, _) $ _ => find_body bounds (Thm.dest_arg tm) haftmann@38558: | Const (@{const_name Trueprop}, _) $ _ => h bounds (Thm.dest_arg tm) wenzelm@23466: | _ => Thm.dest_fun2 tm) wenzelm@23567: and find_args bounds tm = wenzelm@23466: (h bounds (Thm.dest_arg tm) handle CTERM _ => Thm.dest_arg1 tm) wenzelm@23466: and find_body bounds b = wenzelm@23466: let val (_, b') = Thm.dest_abs (SOME (Name.bound bounds)) b wenzelm@23466: in h (bounds + 1) b' end; wenzelm@23466: in h end; wenzelm@23466: wenzelm@51717: fun raw_ferrack_qe_conv ctxt (thy, {isolate_conv, whatis, simpset = ss}) tm = wenzelm@23567: let wenzelm@23567: val ss' = wenzelm@51717: merge_ss (simpset_of wenzelm@51717: (put_simpset HOL_basic_ss ctxt addsimps wenzelm@51717: @{thms simp_thms ex_simps all_simps not_all all_not_ex ex_disj_distrib}), ss); wenzelm@51717: val pcv = Simplifier.rewrite (put_simpset ss' ctxt); wenzelm@51717: val postcv = Simplifier.rewrite (put_simpset ss ctxt); wenzelm@51717: val nnf = K (nnf_conv ctxt then_conv postcv) wenzelm@23567: val qe_conv = Qelim.gen_qelim_conv pcv postcv pcv cons (Thm.add_cterm_frees tm []) wenzelm@23466: (isolate_conv ctxt) nnf wenzelm@51717: (fn vs => ferrack_conv ctxt (thy,{isolate_conv = isolate_conv ctxt, wenzelm@51717: whatis = whatis, simpset = ss}) vs wenzelm@23466: then_conv postcv) wenzelm@51717: in (Simplifier.rewrite (put_simpset ss ctxt) then_conv qe_conv) tm end; wenzelm@23466: wenzelm@23567: fun dlo_instance ctxt tm = wenzelm@23567: Ferrante_Rackoff_Data.match ctxt (grab_atom_bop 0 tm); wenzelm@23486: wenzelm@23567: fun dlo_conv ctxt tm = wenzelm@23567: (case dlo_instance ctxt tm of wenzelm@23567: NONE => raise CTERM ("ferrackqe_conv: no corresponding instance in context!", [tm]) wenzelm@23567: | SOME instance => raw_ferrack_qe_conv ctxt instance tm); wenzelm@23466: wenzelm@23567: fun dlo_tac ctxt = CSUBGOAL (fn (p, i) => wenzelm@23567: (case dlo_instance ctxt p of wenzelm@23567: NONE => no_tac wenzelm@23567: | SOME instance => wenzelm@54742: Object_Logic.full_atomize_tac ctxt i THEN wenzelm@51717: simp_tac (put_simpset (#simpset (snd instance)) ctxt) i THEN (* FIXME already part of raw_ferrack_qe_conv? *) wenzelm@35625: CONVERSION (Object_Logic.judgment_conv (raw_ferrack_qe_conv ctxt instance)) i THEN wenzelm@51717: simp_tac ctxt i)); (* FIXME really? *) wenzelm@23466: wenzelm@23466: end;