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