src/HOL/Tools/Qelim/presburger.ML
author huffman
Thu, 21 Jun 2007 22:52:22 +0200
changeset 23469 3f309f885d0b
parent 23466 886655a150f6
child 23488 342029af73d1
permissions -rw-r--r--
add thm antiquotations
     1 
     2 (*  Title:      HOL/Tools/Presburger/presburger.ML
     3     ID:         $Id$
     4     Author:     Amine Chaieb, TU Muenchen
     5 *)
     6 
     7 signature PRESBURGER =
     8  sig
     9   val cooper_tac: bool -> thm list -> thm list -> Proof.context -> int -> Tactical.tactic
    10 end;
    11 
    12 structure Presburger : PRESBURGER = 
    13 struct
    14 
    15 open Conv;
    16 val comp_ss = HOL_ss addsimps @{thms "Groebner_Basis.comp_arith"};
    17 
    18 fun strip_imp_cprems ct = 
    19  case term_of ct of 
    20   Const ("==>", _) $ _ $ _ => Thm.dest_arg1 ct :: strip_imp_cprems (Thm.dest_arg ct)
    21 | _ => [];
    22 
    23 val cprems_of = strip_imp_cprems o cprop_of;
    24 
    25 fun strip_objimp ct = 
    26  case term_of ct of 
    27   Const ("op -->", _) $ _ $ _ => Thm.dest_arg1 ct :: strip_objimp (Thm.dest_arg ct)
    28 | _ => [ct];
    29 
    30 fun strip_objall ct = 
    31  case term_of ct of 
    32   Const ("All", _) $ Abs (xn,xT,p) => 
    33    let val (a,(v,t')) = (apsnd (Thm.dest_abs (SOME xn)) o Thm.dest_comb) ct
    34    in apfst (cons (a,v)) (strip_objall t')
    35    end
    36 | _ => ([],ct);
    37 
    38 local
    39   val all_maxscope_ss = 
    40      HOL_basic_ss addsimps map (fn th => th RS sym) @{thms "all_simps"}
    41 in
    42 fun thin_prems_tac P i =  simp_tac all_maxscope_ss i THEN
    43   (fn st => case try (nth (cprems_of st)) (i - 1) of
    44     NONE => no_tac st
    45   | SOME p' => 
    46     let
    47      val (qvs, p) = strip_objall (Thm.dest_arg p')
    48      val (ps, c) = split_last (strip_objimp p)
    49      val qs = filter P ps
    50      val q = if P c then c else @{cterm "False"}
    51      val ng = fold_rev (fn (a,v) => fn t => Thm.capply a (Thm.cabs v t)) qvs 
    52          (fold_rev (fn p => fn q => Thm.capply (Thm.capply @{cterm "op -->"} p) q) qs q)
    53      val g = Thm.capply (Thm.capply @{cterm "op ==>"} (Thm.capply @{cterm "Trueprop"} ng)) p'
    54      val ntac = (case qs of [] => q aconvc @{cterm "False"}
    55                          | _ => false)
    56     in 
    57     if ntac then no_tac st
    58       else rtac (Goal.prove_internal [] g (K (blast_tac HOL_cs 1))) i st 
    59     end)
    60 end;
    61 
    62 local
    63  fun ty cts t = 
    64  if not (typ_of (ctyp_of_term t) mem [HOLogic.intT, HOLogic.natT]) then false 
    65     else case term_of t of 
    66       c$_$_ => not (member (op aconv) cts c)
    67     | c$_ => not (member (op aconv) cts c)
    68     | c => not (member (op aconv) cts c)
    69     | _ => true
    70 
    71  val term_constants =
    72   let fun h acc t = case t of
    73     Const _ => insert (op aconv) t acc
    74   | a$b => h (h acc a) b
    75   | Abs (_,_,t) => h acc t
    76   | _ => acc
    77  in h [] end;
    78 in 
    79 fun is_relevant ctxt ct = 
    80   gen_subset (op aconv) (term_constants (term_of ct) , snd (CooperData.get ctxt))
    81  andalso forall (fn Free (_,T) => T = HOLogic.intT) (term_frees (term_of ct))
    82  andalso forall (fn Var (_,T) => T = HOLogic.intT) (term_vars (term_of ct));
    83 
    84 fun int_nat_terms ctxt ct =
    85  let 
    86   val cts = snd (CooperData.get ctxt)
    87   fun h acc t = if ty cts t then insert (op aconvc) t acc else
    88    case (term_of t) of
    89     _$_ => h (h acc (Thm.dest_arg t)) (Thm.dest_fun t)
    90   | Abs(_,_,_) => Thm.dest_abs NONE t ||> h acc |> uncurry (remove (op aconvc))
    91   | _ => acc
    92  in h [] ct end
    93 end;
    94 
    95 fun generalize_tac ctxt f i st = 
    96  case try (nth (cprems_of st)) (i - 1) of
    97     NONE => all_tac st
    98   | SOME p => 
    99    let 
   100    fun all T = Drule.cterm_rule (instantiate' [SOME T] []) @{cpat "all"}
   101    fun gen x t = Thm.capply (all (ctyp_of_term x)) (Thm.cabs x t)
   102    val ts = sort (fn (a,b) => Term.fast_term_ord (term_of a, term_of b)) (f p)
   103    val p' = fold_rev gen ts p
   104  in Seq.of_list [implies_intr p' (implies_elim st (fold forall_elim ts (assume p')))]
   105  end;
   106 
   107 local
   108 val ss1 = comp_ss
   109   addsimps simp_thms @ [@{thm "nat_number_of_def"}, @{thm "zdvd_int"}] 
   110       @ map (fn r => r RS sym) 
   111         [@{thm "int_int_eq"}, @{thm "zle_int"}, @{thm "zless_int"}, @{thm "zadd_int"}, 
   112          @{thm "zmult_int"}]
   113     addsplits [@{thm "zdiff_int_split"}]
   114 
   115 val ss2 = HOL_basic_ss
   116   addsimps [@{thm "nat_0_le"}, @{thm "int_nat_number_of"},
   117             @{thm "all_nat"}, @{thm "ex_nat"}, @{thm "number_of1"}, 
   118             @{thm "number_of2"}, @{thm "int_0"}, @{thm "int_1"}, @{thm "Suc_plus1"}]
   119   addcongs [@{thm "conj_le_cong"}, @{thm "imp_le_cong"}]
   120 val div_mod_ss = HOL_basic_ss addsimps simp_thms 
   121   @ map (symmetric o mk_meta_eq) 
   122     [@{thm "dvd_eq_mod_eq_0"}, @{thm "zdvd_iff_zmod_eq_0"}, @{thm "mod_add1_eq"}, 
   123      @{thm "mod_add_left_eq"}, @{thm "mod_add_right_eq"}, 
   124      @{thm "zmod_zadd1_eq"}, @{thm "zmod_zadd_left_eq"}, 
   125      @{thm "zmod_zadd_right_eq"}, @{thm "div_add1_eq"}, @{thm "zdiv_zadd1_eq"}]
   126   @ [@{thm "mod_self"}, @{thm "zmod_self"}, @{thm "DIVISION_BY_ZERO_MOD"}, 
   127      @{thm "DIVISION_BY_ZERO_DIV"}, @{thm "DIVISION_BY_ZERO"} RS conjunct1, 
   128      @{thm "DIVISION_BY_ZERO"} RS conjunct2, @{thm "zdiv_zero"}, @{thm "zmod_zero"}, 
   129      @{thm "div_0"}, @{thm "mod_0"}, @{thm "zdiv_1"}, @{thm "zmod_1"}, @{thm "div_1"}, 
   130      @{thm "mod_1"}, @{thm "Suc_plus1"}]
   131   @ add_ac
   132  addsimprocs [cancel_div_mod_proc]
   133  val splits_ss = comp_ss addsimps [@{thm "mod_div_equality'"}] addsplits 
   134      [@{thm "split_zdiv"}, @{thm "split_zmod"}, @{thm "split_div'"}, 
   135       @{thm "split_min"}, @{thm "split_max"}, @{thm "abs_split"}]
   136 in
   137 fun nat_to_int_tac ctxt i = 
   138   simp_tac (Simplifier.context ctxt ss1) i THEN 
   139   simp_tac (Simplifier.context ctxt ss2) i THEN 
   140   TRY (simp_tac (Simplifier.context ctxt comp_ss) i);
   141 
   142 fun div_mod_tac  ctxt i = simp_tac (Simplifier.context ctxt div_mod_ss) i;
   143 fun splits_tac ctxt i = simp_tac (Simplifier.context ctxt splits_ss) i;
   144 end;
   145 
   146 
   147 fun eta_beta_tac ctxt i st = case try (nth (cprems_of st)) (i - 1) of
   148    NONE => no_tac st
   149  | SOME p => 
   150    let
   151     val eq = (eta_conv (ProofContext.theory_of ctxt) then_conv Thm.beta_conversion true) p
   152     val p' = Thm.rhs_of eq
   153     val th = implies_intr p' (equal_elim (symmetric eq) (assume p'))
   154    in rtac th i st
   155    end;
   156 
   157 
   158 
   159 fun core_cooper_tac ctxt i st = 
   160  case try (nth (cprems_of st)) (i - 1) of
   161    NONE => all_tac st
   162  | SOME p => 
   163    let 
   164     val cpth = 
   165        if !quick_and_dirty 
   166        then linzqe_oracle (ProofContext.theory_of ctxt) 
   167              (Envir.beta_norm (Pattern.eta_long [] (term_of (Thm.dest_arg p))))
   168        else arg_conv (Cooper.cooper_conv ctxt) p
   169     val p' = Thm.rhs_of cpth
   170     val th = implies_intr p' (equal_elim (symmetric cpth) (assume p'))
   171    in rtac th i st end
   172    handle Cooper.COOPER _ => no_tac st;
   173 
   174 fun nogoal_tac i st = case try (nth (cprems_of st)) (i - 1) of
   175    NONE => no_tac st
   176  | SOME _ => all_tac st
   177 
   178 fun finish_tac q i st = case try (nth (cprems_of st)) (i - 1) of
   179    NONE => all_tac st
   180  | SOME _ => (if q then I else TRY) (rtac TrueI i) st
   181 
   182 fun cooper_tac elim add_ths del_ths ctxt i = 
   183 let val ss = fst (CooperData.get ctxt) delsimps del_ths addsimps add_ths
   184 in
   185 nogoal_tac i 
   186 THEN (EVERY o (map TRY))
   187  [ObjectLogic.full_atomize_tac i,
   188   eta_beta_tac ctxt i,
   189   simp_tac ss  i,
   190   generalize_tac ctxt (int_nat_terms ctxt) i,
   191   ObjectLogic.full_atomize_tac i,
   192   div_mod_tac ctxt i,
   193   splits_tac ctxt i,
   194   simp_tac ss i,
   195   eta_beta_tac ctxt i,
   196   nat_to_int_tac ctxt i, 
   197   thin_prems_tac (is_relevant ctxt) i]
   198 THEN core_cooper_tac ctxt i THEN finish_tac elim i
   199 end;
   200 
   201 end;