src/HOL/Decision_Procs/cooper_tac.ML
author wenzelm
Thu, 18 Apr 2013 17:07:01 +0200
changeset 51717 9e7d1c139569
parent 47432 e1576d13e933
child 52131 366fa32ee2a3
permissions -rw-r--r--
simplifier uses proper Proof.context instead of historic type simpset;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
30439
57c68b3af2ea Updated paths in Decision_Procs comments and NEWS
hoelzl
parents: 30242
diff changeset
     1
(*  Title:      HOL/Decision_Procs/cooper_tac.ML
29788
1b80ebe713a4 established session HOL-Reflection
haftmann
parents: 29265
diff changeset
     2
    Author:     Amine Chaieb, TU Muenchen
1b80ebe713a4 established session HOL-Reflection
haftmann
parents: 29265
diff changeset
     3
*)
1b80ebe713a4 established session HOL-Reflection
haftmann
parents: 29265
diff changeset
     4
31240
2c20bcd70fbe proper signature constraints;
wenzelm
parents: 31070
diff changeset
     5
signature COOPER_TAC =
2c20bcd70fbe proper signature constraints;
wenzelm
parents: 31070
diff changeset
     6
sig
32740
9dd0a2f83429 explicit indication of Unsynchronized.ref;
wenzelm
parents: 31790
diff changeset
     7
  val trace: bool Unsynchronized.ref
31240
2c20bcd70fbe proper signature constraints;
wenzelm
parents: 31070
diff changeset
     8
  val linz_tac: Proof.context -> bool -> int -> tactic
2c20bcd70fbe proper signature constraints;
wenzelm
parents: 31070
diff changeset
     9
end
2c20bcd70fbe proper signature constraints;
wenzelm
parents: 31070
diff changeset
    10
2c20bcd70fbe proper signature constraints;
wenzelm
parents: 31070
diff changeset
    11
structure Cooper_Tac: COOPER_TAC =
23274
f997514ad8f4 New Reflected Presburger added to HOL/ex
chaieb
parents:
diff changeset
    12
struct
f997514ad8f4 New Reflected Presburger added to HOL/ex
chaieb
parents:
diff changeset
    13
32740
9dd0a2f83429 explicit indication of Unsynchronized.ref;
wenzelm
parents: 31790
diff changeset
    14
val trace = Unsynchronized.ref false;
23274
f997514ad8f4 New Reflected Presburger added to HOL/ex
chaieb
parents:
diff changeset
    15
fun trace_msg s = if !trace then tracing s else ();
f997514ad8f4 New Reflected Presburger added to HOL/ex
chaieb
parents:
diff changeset
    16
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 47432
diff changeset
    17
val cooper_ss = simpset_of @{context};
23274
f997514ad8f4 New Reflected Presburger added to HOL/ex
chaieb
parents:
diff changeset
    18
f997514ad8f4 New Reflected Presburger added to HOL/ex
chaieb
parents:
diff changeset
    19
val nT = HOLogic.natT;
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 45654
diff changeset
    20
val comp_arith = @{thms simp_thms}
23274
f997514ad8f4 New Reflected Presburger added to HOL/ex
chaieb
parents:
diff changeset
    21
27651
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 26075
diff changeset
    22
val zdvd_int = @{thm zdvd_int};
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 26075
diff changeset
    23
val zdiff_int_split = @{thm zdiff_int_split};
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 26075
diff changeset
    24
val all_nat = @{thm all_nat};
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 26075
diff changeset
    25
val ex_nat = @{thm ex_nat};
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 26075
diff changeset
    26
val split_zdiv = @{thm split_zdiv};
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 26075
diff changeset
    27
val split_zmod = @{thm split_zmod};
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 26075
diff changeset
    28
val mod_div_equality' = @{thm mod_div_equality'};
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 26075
diff changeset
    29
val split_div' = @{thm split_div'};
31790
05c92381363c corrected and unified thm names
nipkow
parents: 31240
diff changeset
    30
val Suc_eq_plus1 = @{thm Suc_eq_plus1};
27651
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 26075
diff changeset
    31
val imp_le_cong = @{thm imp_le_cong};
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 26075
diff changeset
    32
val conj_le_cong = @{thm conj_le_cong};
30034
60f64f112174 removed redundant thms
nipkow
parents: 30031
diff changeset
    33
val mod_add_left_eq = @{thm mod_add_left_eq} RS sym;
60f64f112174 removed redundant thms
nipkow
parents: 30031
diff changeset
    34
val mod_add_right_eq = @{thm mod_add_right_eq} RS sym;
30224
79136ce06bdb removed and renamed redundant lemmas
nipkow
parents: 30034
diff changeset
    35
val mod_add_eq = @{thm mod_add_eq} RS sym;
27651
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 26075
diff changeset
    36
val nat_div_add_eq = @{thm div_add1_eq} RS sym;
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 26075
diff changeset
    37
val int_div_add_eq = @{thm zdiv_zadd1_eq} RS sym;
23274
f997514ad8f4 New Reflected Presburger added to HOL/ex
chaieb
parents:
diff changeset
    38
31240
2c20bcd70fbe proper signature constraints;
wenzelm
parents: 31070
diff changeset
    39
fun prepare_for_linz q fm =
23274
f997514ad8f4 New Reflected Presburger added to HOL/ex
chaieb
parents:
diff changeset
    40
  let
f997514ad8f4 New Reflected Presburger added to HOL/ex
chaieb
parents:
diff changeset
    41
    val ps = Logic.strip_params fm
f997514ad8f4 New Reflected Presburger added to HOL/ex
chaieb
parents:
diff changeset
    42
    val hs = map HOLogic.dest_Trueprop (Logic.strip_assums_hyp fm)
f997514ad8f4 New Reflected Presburger added to HOL/ex
chaieb
parents:
diff changeset
    43
    val c = HOLogic.dest_Trueprop (Logic.strip_assums_concl fm)
f997514ad8f4 New Reflected Presburger added to HOL/ex
chaieb
parents:
diff changeset
    44
    fun mk_all ((s, T), (P,n)) =
42083
e1209fc7ecdc added Term.is_open and Term.is_dependent convenience, to cover common situations of loose bounds;
wenzelm
parents: 38558
diff changeset
    45
      if Term.is_dependent P then
23274
f997514ad8f4 New Reflected Presburger added to HOL/ex
chaieb
parents:
diff changeset
    46
        (HOLogic.all_const T $ Abs (s, T, P), n)
f997514ad8f4 New Reflected Presburger added to HOL/ex
chaieb
parents:
diff changeset
    47
      else (incr_boundvars ~1 P, n-1)
f997514ad8f4 New Reflected Presburger added to HOL/ex
chaieb
parents:
diff changeset
    48
    fun mk_all2 (v, t) = HOLogic.all_const (fastype_of v) $ lambda v t;
27651
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 26075
diff changeset
    49
    val rhs = hs
23274
f997514ad8f4 New Reflected Presburger added to HOL/ex
chaieb
parents:
diff changeset
    50
    val np = length ps
33004
715566791eb0 always qualify NJ's old List.foldl/foldr in Isabelle/ML;
wenzelm
parents: 32960
diff changeset
    51
    val (fm',np) = List.foldr (fn ((x, T), (fm,n)) => mk_all ((x, T), (fm,n)))
715566791eb0 always qualify NJ's old List.foldl/foldr in Isabelle/ML;
wenzelm
parents: 32960
diff changeset
    52
      (List.foldr HOLogic.mk_imp c rhs, np) ps
23274
f997514ad8f4 New Reflected Presburger added to HOL/ex
chaieb
parents:
diff changeset
    53
    val (vs, _) = List.partition (fn t => q orelse (type_of t) = nT)
44121
44adaa6db327 old term operations are legacy;
wenzelm
parents: 43594
diff changeset
    54
      (Misc_Legacy.term_frees fm' @ Misc_Legacy.term_vars fm');
33004
715566791eb0 always qualify NJ's old List.foldl/foldr in Isabelle/ML;
wenzelm
parents: 32960
diff changeset
    55
    val fm2 = List.foldr mk_all2 fm' vs
23274
f997514ad8f4 New Reflected Presburger added to HOL/ex
chaieb
parents:
diff changeset
    56
  in (fm2, np + length vs, length rhs) end;
f997514ad8f4 New Reflected Presburger added to HOL/ex
chaieb
parents:
diff changeset
    57
f997514ad8f4 New Reflected Presburger added to HOL/ex
chaieb
parents:
diff changeset
    58
(*Object quantifier to meta --*)
f997514ad8f4 New Reflected Presburger added to HOL/ex
chaieb
parents:
diff changeset
    59
fun spec_step n th = if (n=0) then th else (spec_step (n-1) th) RS spec ;
f997514ad8f4 New Reflected Presburger added to HOL/ex
chaieb
parents:
diff changeset
    60
f997514ad8f4 New Reflected Presburger added to HOL/ex
chaieb
parents:
diff changeset
    61
(* object implication to meta---*)
f997514ad8f4 New Reflected Presburger added to HOL/ex
chaieb
parents:
diff changeset
    62
fun mp_step n th = if (n=0) then th else (mp_step (n-1) th) RS mp;
f997514ad8f4 New Reflected Presburger added to HOL/ex
chaieb
parents:
diff changeset
    63
f997514ad8f4 New Reflected Presburger added to HOL/ex
chaieb
parents:
diff changeset
    64
42368
3b8498ac2314 proper subgoal addressing via SUBGOAL/CSUBGOAL -- assuming these tactics did not handle Subscript in any special way;
wenzelm
parents: 42364
diff changeset
    65
fun linz_tac ctxt q = Object_Logic.atomize_prems_tac THEN' SUBGOAL (fn (g, i) =>
23274
f997514ad8f4 New Reflected Presburger added to HOL/ex
chaieb
parents:
diff changeset
    66
  let
42361
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 42083
diff changeset
    67
    val thy = Proof_Context.theory_of ctxt
23274
f997514ad8f4 New Reflected Presburger added to HOL/ex
chaieb
parents:
diff changeset
    68
    (* Transform the term*)
f997514ad8f4 New Reflected Presburger added to HOL/ex
chaieb
parents:
diff changeset
    69
    val (t,np,nh) = prepare_for_linz q g
f997514ad8f4 New Reflected Presburger added to HOL/ex
chaieb
parents:
diff changeset
    70
    (* Some simpsets for dealing with mod div abs and nat*)
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 47432
diff changeset
    71
    val mod_div_simpset =
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 47432
diff changeset
    72
      put_simpset HOL_basic_ss ctxt
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32740
diff changeset
    73
      addsimps [refl,mod_add_eq, mod_add_left_eq,
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32740
diff changeset
    74
          mod_add_right_eq,
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32740
diff changeset
    75
          nat_div_add_eq, int_div_add_eq,
47142
d64fa2ca54b8 remove redundant lemmas
huffman
parents: 47108
diff changeset
    76
          @{thm mod_self},
d64fa2ca54b8 remove redundant lemmas
huffman
parents: 47108
diff changeset
    77
          @{thm div_by_0}, @{thm mod_by_0}, @{thm div_0}, @{thm mod_0},
d64fa2ca54b8 remove redundant lemmas
huffman
parents: 47108
diff changeset
    78
          @{thm div_by_1}, @{thm mod_by_1}, @{thm div_1}, @{thm mod_1},
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32740
diff changeset
    79
          Suc_eq_plus1]
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32740
diff changeset
    80
      addsimps @{thms add_ac}
43594
ef1ddc59b825 modernized some simproc setup;
wenzelm
parents: 42368
diff changeset
    81
      addsimprocs [@{simproc cancel_div_mod_nat}, @{simproc cancel_div_mod_int}]
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 47432
diff changeset
    82
    val simpset0 =
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 47432
diff changeset
    83
      put_simpset HOL_basic_ss ctxt
31790
05c92381363c corrected and unified thm names
nipkow
parents: 31240
diff changeset
    84
      addsimps [mod_div_equality', Suc_eq_plus1]
23274
f997514ad8f4 New Reflected Presburger added to HOL/ex
chaieb
parents:
diff changeset
    85
      addsimps comp_arith
45620
f2a587696afb modernized some old-style infix operations, which were left over from the time of ML proof scripts;
wenzelm
parents: 44121
diff changeset
    86
      |> fold Splitter.add_split
f2a587696afb modernized some old-style infix operations, which were left over from the time of ML proof scripts;
wenzelm
parents: 44121
diff changeset
    87
          [split_zdiv, split_zmod, split_div', @{thm "split_min"}, @{thm "split_max"}]
23274
f997514ad8f4 New Reflected Presburger added to HOL/ex
chaieb
parents:
diff changeset
    88
    (* Simp rules for changing (n::int) to int n *)
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 47432
diff changeset
    89
    val simpset1 =
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 47432
diff changeset
    90
      put_simpset HOL_basic_ss ctxt
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 45654
diff changeset
    91
      addsimps [zdvd_int] @ map (fn r => r RS sym)
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 45654
diff changeset
    92
        [@{thm int_numeral}, @{thm int_int_eq}, @{thm zle_int}, @{thm zless_int}, @{thm zadd_int}, @{thm zmult_int}]
45620
f2a587696afb modernized some old-style infix operations, which were left over from the time of ML proof scripts;
wenzelm
parents: 44121
diff changeset
    93
      |> Splitter.add_split zdiff_int_split
23274
f997514ad8f4 New Reflected Presburger added to HOL/ex
chaieb
parents:
diff changeset
    94
    (*simp rules for elimination of int n*)
f997514ad8f4 New Reflected Presburger added to HOL/ex
chaieb
parents:
diff changeset
    95
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 47432
diff changeset
    96
    val simpset2 =
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 47432
diff changeset
    97
      put_simpset HOL_basic_ss ctxt
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 45654
diff changeset
    98
      addsimps [@{thm nat_0_le}, @{thm all_nat}, @{thm ex_nat}, @{thm zero_le_numeral}, @{thm order_refl}(* FIXME: necessary? *), @{thm int_0}, @{thm int_1}]
45620
f2a587696afb modernized some old-style infix operations, which were left over from the time of ML proof scripts;
wenzelm
parents: 44121
diff changeset
    99
      |> fold Simplifier.add_cong [@{thm conj_le_cong}, @{thm imp_le_cong}]
23274
f997514ad8f4 New Reflected Presburger added to HOL/ex
chaieb
parents:
diff changeset
   100
    (* simp rules for elimination of abs *)
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 47432
diff changeset
   101
    val simpset3 = put_simpset HOL_basic_ss ctxt |> Splitter.add_split @{thm abs_split}
23274
f997514ad8f4 New Reflected Presburger added to HOL/ex
chaieb
parents:
diff changeset
   102
    val ct = cterm_of thy (HOLogic.mk_Trueprop t)
f997514ad8f4 New Reflected Presburger added to HOL/ex
chaieb
parents:
diff changeset
   103
    (* Theorem for the nat --> int transformation *)
f997514ad8f4 New Reflected Presburger added to HOL/ex
chaieb
parents:
diff changeset
   104
    val pre_thm = Seq.hd (EVERY
f997514ad8f4 New Reflected Presburger added to HOL/ex
chaieb
parents:
diff changeset
   105
      [simp_tac mod_div_simpset 1, simp_tac simpset0 1,
f997514ad8f4 New Reflected Presburger added to HOL/ex
chaieb
parents:
diff changeset
   106
       TRY (simp_tac simpset1 1), TRY (simp_tac simpset2 1),
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 47432
diff changeset
   107
       TRY (simp_tac simpset3 1), TRY (simp_tac (put_simpset cooper_ss ctxt) 1)]
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36692
diff changeset
   108
      (Thm.trivial ct))
23274
f997514ad8f4 New Reflected Presburger added to HOL/ex
chaieb
parents:
diff changeset
   109
    fun assm_tac i = REPEAT_DETERM_N nh (assume_tac i)
f997514ad8f4 New Reflected Presburger added to HOL/ex
chaieb
parents:
diff changeset
   110
    (* The result of the quantifier elimination *)
f997514ad8f4 New Reflected Presburger added to HOL/ex
chaieb
parents:
diff changeset
   111
    val (th, tac) = case (prop_of pre_thm) of
38558
32ad17fe2b9c tuned quotes
haftmann
parents: 38549
diff changeset
   112
        Const ("==>", _) $ (Const (@{const_name Trueprop}, _) $ t1) $ _ =>
28290
4cc2b6046258 simplified oracle interface;
wenzelm
parents: 27651
diff changeset
   113
    let val pth = linzqe_oracle (cterm_of thy (Pattern.eta_long [] t1))
31240
2c20bcd70fbe proper signature constraints;
wenzelm
parents: 31070
diff changeset
   114
    in
23274
f997514ad8f4 New Reflected Presburger added to HOL/ex
chaieb
parents:
diff changeset
   115
          ((pth RS iffD2) RS pre_thm,
f997514ad8f4 New Reflected Presburger added to HOL/ex
chaieb
parents:
diff changeset
   116
            assm_tac (i + 1) THEN (if q then I else TRY) (rtac TrueI i))
f997514ad8f4 New Reflected Presburger added to HOL/ex
chaieb
parents:
diff changeset
   117
    end
f997514ad8f4 New Reflected Presburger added to HOL/ex
chaieb
parents:
diff changeset
   118
      | _ => (pre_thm, assm_tac i)
42368
3b8498ac2314 proper subgoal addressing via SUBGOAL/CSUBGOAL -- assuming these tactics did not handle Subscript in any special way;
wenzelm
parents: 42364
diff changeset
   119
  in rtac (((mp_step nh) o (spec_step np)) th) i THEN tac end);
23274
f997514ad8f4 New Reflected Presburger added to HOL/ex
chaieb
parents:
diff changeset
   120
23590
ad95084a5c63 renamed ObjectLogic.atomize_tac to ObjectLogic.atomize_prems_tac;
wenzelm
parents: 23469
diff changeset
   121
end