| 
24584
 | 
     1  | 
(* Title:      HOL/Tools/Qelim/ferrante_rackoff.ML
  | 
| 
23466
 | 
     2  | 
   ID:         $Id$
  | 
| 
 | 
     3  | 
   Author:     Amine Chaieb, TU Muenchen
  | 
| 
 | 
     4  | 
  | 
| 
 | 
     5  | 
Ferrante and Rackoff's algorithm for quantifier elimination in dense
  | 
| 
 | 
     6  | 
linear orders.  Proof-synthesis and tactic.
  | 
| 
 | 
     7  | 
*)
  | 
| 
 | 
     8  | 
  | 
| 
23567
 | 
     9  | 
signature FERRANTE_RACKOFF =
  | 
| 
23466
 | 
    10  | 
sig
  | 
| 
23567
 | 
    11  | 
  val dlo_conv: Proof.context -> conv
  | 
| 
23466
 | 
    12  | 
  val dlo_tac: Proof.context -> int -> tactic
  | 
| 
 | 
    13  | 
end;
  | 
| 
 | 
    14  | 
  | 
| 
 | 
    15  | 
structure FerranteRackoff: FERRANTE_RACKOFF =
  | 
| 
 | 
    16  | 
struct
  | 
| 
 | 
    17  | 
  | 
| 
 | 
    18  | 
open Ferrante_Rackoff_Data;
  | 
| 
 | 
    19  | 
open Conv;
  | 
| 
 | 
    20  | 
  | 
| 
23567
 | 
    21  | 
type entry = {minf: thm list, pinf: thm list, nmi: thm list, npi: thm list,
 | 
| 
23466
 | 
    22  | 
   ld: thm list, qe: thm, atoms : cterm list} *
  | 
| 
23567
 | 
    23  | 
  {isolate_conv: cterm list -> cterm -> thm,
 | 
| 
23466
 | 
    24  | 
                 whatis : cterm -> cterm -> ord,
  | 
| 
 | 
    25  | 
                 simpset : simpset};
  | 
| 
 | 
    26  | 
  | 
| 
23567
 | 
    27  | 
fun get_p1 th =
  | 
| 
 | 
    28  | 
  funpow 2 (Thm.dest_arg o snd o Thm.dest_abs NONE)
  | 
| 
 | 
    29  | 
    (funpow 2 Thm.dest_arg (cprop_of th)) |> Thm.dest_arg
  | 
| 
23466
 | 
    30  | 
  | 
| 
 | 
    31  | 
fun ferrack_conv
  | 
| 
23567
 | 
    32  | 
   (entr as ({minf = minf, pinf = pinf, nmi = nmi, npi = npi,
 | 
| 
23466
 | 
    33  | 
              ld = ld, qe = qe, atoms = atoms},
  | 
| 
23567
 | 
    34  | 
             {isolate_conv = icv, whatis = wi, simpset = simpset}):entry) =
 | 
| 
 | 
    35  | 
let
  | 
| 
23466
 | 
    36  | 
 fun uset (vars as (x::vs)) p = case term_of p of
  | 
| 
23567
 | 
    37  | 
   Const("op &", _)$ _ $ _ =>
 | 
| 
 | 
    38  | 
     let
  | 
| 
 | 
    39  | 
       val ((b,l),r) = Thm.dest_comb p |>> Thm.dest_comb
  | 
| 
23466
 | 
    40  | 
       val (lS,lth) = uset vars l  val (rS, rth) = uset vars r
  | 
| 
23567
 | 
    41  | 
     in (lS@rS, Drule.binop_cong_rule b lth rth) end
  | 
| 
 | 
    42  | 
 |  Const("op |", _)$ _ $ _ =>
 | 
| 
 | 
    43  | 
     let
  | 
| 
 | 
    44  | 
       val ((b,l),r) = Thm.dest_comb p |>> Thm.dest_comb
  | 
| 
23466
 | 
    45  | 
       val (lS,lth) = uset vars l  val (rS, rth) = uset vars r
  | 
| 
23567
 | 
    46  | 
     in (lS@rS, Drule.binop_cong_rule b lth rth) end
  | 
| 
 | 
    47  | 
 | _ =>
  | 
| 
 | 
    48  | 
    let
  | 
| 
 | 
    49  | 
      val th = icv vars p
  | 
| 
23466
 | 
    50  | 
      val p' = Thm.rhs_of th
  | 
| 
 | 
    51  | 
      val c = wi x p'
  | 
| 
23567
 | 
    52  | 
      val S = (if member (op =) [Lt, Le, Eq] c then single o Thm.dest_arg
  | 
| 
 | 
    53  | 
               else if member (op =) [Gt, Ge] c then single o Thm.dest_arg1
  | 
| 
 | 
    54  | 
               else if c = NEq then single o Thm.dest_arg o Thm.dest_arg
  | 
| 
23466
 | 
    55  | 
               else K []) p'
  | 
| 
 | 
    56  | 
    in (S,th) end
  | 
| 
 | 
    57  | 
  | 
| 
23567
 | 
    58  | 
 val ((p1_v,p2_v),(mp1_v,mp2_v)) =
  | 
| 
 | 
    59  | 
   funpow 2 (Thm.dest_arg o snd o Thm.dest_abs NONE)
  | 
| 
 | 
    60  | 
     (funpow 4 Thm.dest_arg (cprop_of (hd minf)))
  | 
| 
 | 
    61  | 
   |> Thm.dest_binop |> pairself Thm.dest_binop |> apfst (pairself Thm.dest_fun)
  | 
| 
 | 
    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
  | 
| 
23466
 | 
    71  | 
           instantiate ([],[(p1_v, p1),(p2_v, p2),(mp1_v, mp1), (mp2_v, mp2)]) th
  | 
| 
 | 
    72  | 
        else instantiate ([],[(p1_v, p1),(p2_v, p2),(mp1_v, pp1), (mp2_v, pp2)]) th
  | 
| 
 | 
    73  | 
     in implies_elim (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
  | 
| 
 | 
    77  | 
 val U_v = (Thm.dest_arg o Thm.dest_arg o Thm.dest_arg1) (cprop_of qe)
  | 
| 
23567
 | 
    78  | 
 fun main vs p =
  | 
| 
 | 
    79  | 
  let
  | 
| 
 | 
    80  | 
   val ((xn,ce),(x,fm)) = (case term_of p of
  | 
| 
 | 
    81  | 
                   Const("Ex",_)$Abs(xn,xT,_) =>
 | 
| 
23466
 | 
    82  | 
                        Thm.dest_comb p ||> Thm.dest_abs (SOME xn) |>> pair xn
  | 
| 
23567
 | 
    83  | 
                 | _ => raise CTERM ("main QE only treats existential quantifiers!", [p]))
 | 
| 
23466
 | 
    84  | 
   val cT = ctyp_of_term x
  | 
| 
 | 
    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
  | 
| 
 | 
    90  | 
   val [th0,th1] = map (instantiate' [SOME cT] []) @{thms "finite.intros"}
 | 
| 
23567
 | 
    91  | 
   fun ins x th =
  | 
| 
 | 
    92  | 
      implies_elim (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
  | 
| 
23466
 | 
    97  | 
     val insI1 = instantiate' [SOME cT] [] @{thm "insertI1"}
 | 
| 
 | 
    98  | 
     val insI2 = instantiate' [SOME cT] [] @{thm "insertI2"}
 | 
| 
 | 
    99  | 
   in
  | 
| 
23567
 | 
   100  | 
    fun provein x S =
  | 
| 
23466
 | 
   101  | 
     case term_of S of
  | 
| 
23567
 | 
   102  | 
        Const("{}",_) => raise CTERM ("provein : not a member!", [S])
 | 
| 
 | 
   103  | 
      | Const("insert",_)$y$_ =>
 | 
| 
23466
 | 
   104  | 
         let val (cy,S') = Thm.dest_binop S
  | 
| 
 | 
   105  | 
         in if term_of x aconv y then instantiate' [] [SOME x, SOME S'] insI1
  | 
| 
23567
 | 
   106  | 
         else implies_elim (instantiate' [] [SOME x, SOME S', SOME cy] insI2)
  | 
| 
23466
 | 
   107  | 
                           (provein x S')
  | 
| 
 | 
   108  | 
         end
  | 
| 
 | 
   109  | 
   end
  | 
| 
23567
 | 
   110  | 
   val tabU = fold (fn t => fn tab => Termtab.update (term_of t, provein t cU) tab)
  | 
| 
23466
 | 
   111  | 
                   u Termtab.empty
  | 
| 
23567
 | 
   112  | 
   val U = the o Termtab.lookup tabU o term_of
  | 
| 
 | 
   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,
  | 
| 
23466
 | 
   118  | 
        nmi_le, nmi_gt, nmi_ge, nmi_P] = map (instantiate ([],[(U_v,cU)])) nmi
  | 
| 
23567
 | 
   119  | 
   val [npi_conj, npi_disj, npi_eq, npi_neq, npi_lt,
  | 
| 
23466
 | 
   120  | 
        npi_le, npi_gt, npi_ge, npi_P] = map (instantiate ([],[(U_v,cU)])) npi
  | 
| 
23567
 | 
   121  | 
   val [ld_conj, ld_disj, ld_eq, ld_neq, ld_lt,
  | 
| 
23466
 | 
   122  | 
        ld_le, ld_gt, ld_ge, ld_P] = map (instantiate ([],[(U_v,cU)])) ld
  | 
| 
23567
 | 
   123  | 
  | 
| 
 | 
   124  | 
   fun decomp_mpinf fm =
  | 
| 
23466
 | 
   125  | 
     case term_of fm of
  | 
| 
23567
 | 
   126  | 
       Const("op &",_)$_$_ =>
 | 
| 
 | 
   127  | 
        let val (p,q) = Thm.dest_binop fm
  | 
| 
 | 
   128  | 
        in ([p,q], myfwd (minf_conj,pinf_conj, nmi_conj, npi_conj,ld_conj)
  | 
| 
23466
 | 
   129  | 
                         (Thm.cabs x p) (Thm.cabs x q))
  | 
| 
 | 
   130  | 
        end
  | 
| 
23567
 | 
   131  | 
     | Const("op |",_)$_$_ =>
 | 
| 
 | 
   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)
  | 
| 
 | 
   134  | 
                         (Thm.cabs x p) (Thm.cabs x q))
  | 
| 
 | 
   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)
  | 
| 
 | 
   142  | 
                      else sys_error "decomp_mpinf: Impossible case!!") fm
  | 
| 
 | 
   143  | 
             val [mi_th, pi_th, nmi_th, npi_th, ld_th] =
  | 
| 
 | 
   144  | 
               if c = Nox then map (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] =
  | 
| 
23466
 | 
   148  | 
                 map (instantiate' [] [SOME t])
  | 
| 
 | 
   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
  | 
| 
 | 
   156  | 
                    fun Ufw th = implies_elim th tU
  | 
| 
 | 
   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))
  | 
| 
 | 
   163  | 
                   (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 =
  | 
| 
23466
 | 
   167  | 
      Simplifier.rewrite (HOL_basic_ss addsimps simp_thms@(@{thms "bex_simps" (1-5)}))
 | 
| 
 | 
   168  | 
    val result_th = fconv_rule (arg_conv bex_conv) (transitive enth qe_th)
  | 
| 
 | 
   169  | 
   in result_th
  | 
| 
 | 
   170  | 
   end
  | 
| 
 | 
   171  | 
  | 
| 
 | 
   172  | 
in main
  | 
| 
 | 
   173  | 
end;
  | 
| 
 | 
   174  | 
  | 
| 
23567
 | 
   175  | 
val grab_atom_bop =
  | 
| 
 | 
   176  | 
 let
  | 
| 
23466
 | 
   177  | 
  fun h bounds tm =
  | 
| 
 | 
   178  | 
   (case term_of tm of
  | 
| 
 | 
   179  | 
     Const ("op =", T) $ _ $ _ =>
 | 
| 
23567
 | 
   180  | 
       if domain_type T = HOLogic.boolT then find_args bounds tm
  | 
| 
23466
 | 
   181  | 
       else Thm.dest_fun2 tm
  | 
| 
 | 
   182  | 
   | Const ("Not", _) $ _ => h bounds (Thm.dest_arg tm)
 | 
| 
 | 
   183  | 
   | Const ("All", _) $ _ => find_body bounds (Thm.dest_arg tm)
 | 
| 
 | 
   184  | 
   | Const ("Ex", _) $ _ => find_body bounds (Thm.dest_arg tm)
 | 
| 
 | 
   185  | 
   | Const ("op &", _) $ _ $ _ => find_args bounds tm
 | 
| 
 | 
   186  | 
   | Const ("op |", _) $ _ $ _ => find_args bounds tm
 | 
| 
 | 
   187  | 
   | Const ("op -->", _) $ _ $ _ => find_args bounds tm
 | 
| 
 | 
   188  | 
   | Const ("==>", _) $ _ $ _ => find_args bounds tm
 | 
| 
 | 
   189  | 
   | Const ("==", _) $ _ $ _ => find_args bounds tm
 | 
| 
 | 
   190  | 
   | Const ("Trueprop", _) $ _ => h bounds (Thm.dest_arg tm)
 | 
| 
 | 
   191  | 
   | _ => Thm.dest_fun2 tm)
  | 
| 
23567
 | 
   192  | 
  and find_args bounds tm =
  | 
| 
23466
 | 
   193  | 
           (h bounds (Thm.dest_arg tm) handle CTERM _ => Thm.dest_arg1 tm)
  | 
| 
 | 
   194  | 
 and find_body bounds b =
  | 
| 
 | 
   195  | 
   let val (_, b') = Thm.dest_abs (SOME (Name.bound bounds)) b
  | 
| 
 | 
   196  | 
   in h (bounds + 1) b' end;
  | 
| 
 | 
   197  | 
in h end;
  | 
| 
 | 
   198  | 
  | 
| 
23567
 | 
   199  | 
fun raw_ferrack_qe_conv ctxt (thy, {isolate_conv, whatis, simpset}) tm =
 | 
| 
 | 
   200  | 
 let
  | 
| 
23466
 | 
   201  | 
   val ss = simpset
  | 
| 
23567
 | 
   202  | 
   val ss' =
  | 
| 
 | 
   203  | 
     merge_ss (HOL_basic_ss addsimps (simp_thms @ ex_simps @ all_simps)
  | 
| 
 | 
   204  | 
              @ [not_all,@{thm "all_not_ex"}, ex_disj_distrib], ss)
 | 
| 
 | 
   205  | 
     |> Simplifier.inherit_context ss
  | 
| 
 | 
   206  | 
   val pcv = Simplifier.rewrite ss'     
  | 
| 
 | 
   207  | 
   val postcv = Simplifier.rewrite ss
  | 
| 
 | 
   208  | 
   val nnf = K (nnf_conv then_conv postcv)
  | 
| 
 | 
   209  | 
   val qe_conv = Qelim.gen_qelim_conv pcv postcv pcv cons (Thm.add_cterm_frees tm [])
  | 
| 
23466
 | 
   210  | 
                  (isolate_conv ctxt) nnf
  | 
| 
23567
 | 
   211  | 
                  (fn vs => ferrack_conv (thy,{isolate_conv = isolate_conv ctxt,
 | 
| 
23466
 | 
   212  | 
                                               whatis = whatis, simpset = simpset}) vs
  | 
| 
 | 
   213  | 
                   then_conv postcv)
  | 
| 
23486
 | 
   214  | 
 in (Simplifier.rewrite ss then_conv qe_conv) tm end;
  | 
| 
23466
 | 
   215  | 
  | 
| 
23567
 | 
   216  | 
fun dlo_instance ctxt tm =
  | 
| 
 | 
   217  | 
  Ferrante_Rackoff_Data.match ctxt (grab_atom_bop 0 tm);
  | 
| 
23486
 | 
   218  | 
  | 
| 
23567
 | 
   219  | 
fun dlo_conv ctxt tm =
  | 
| 
 | 
   220  | 
  (case dlo_instance ctxt tm of
  | 
| 
 | 
   221  | 
    NONE => raise CTERM ("ferrackqe_conv: no corresponding instance in context!", [tm])
 | 
| 
 | 
   222  | 
  | SOME instance => raw_ferrack_qe_conv ctxt instance tm);
  | 
| 
23466
 | 
   223  | 
  | 
| 
23567
 | 
   224  | 
fun dlo_tac ctxt = CSUBGOAL (fn (p, i) =>
  | 
| 
 | 
   225  | 
  (case dlo_instance ctxt p of
  | 
| 
 | 
   226  | 
    NONE => no_tac
  | 
| 
 | 
   227  | 
  | SOME instance =>
  | 
| 
 | 
   228  | 
      ObjectLogic.full_atomize_tac i THEN
  | 
| 
 | 
   229  | 
      simp_tac (#simpset (snd instance)) i THEN  (* FIXME already part of raw_ferrack_qe_conv? *)
  | 
| 
 | 
   230  | 
      CONVERSION (ObjectLogic.judgment_conv (raw_ferrack_qe_conv ctxt instance)) i THEN
  | 
| 
 | 
   231  | 
      simp_tac (Simplifier.local_simpset_of ctxt) i));  (* FIXME really? *)
  | 
| 
23466
 | 
   232  | 
  | 
| 
 | 
   233  | 
end;
  |