src/HOL/Decision_Procs/cooper_tac.ML
author haftmann
Sat, 17 Dec 2016 15:22:14 +0100
changeset 64593 50c715579715
parent 64244 e7102c40783c
child 67118 ccab07d1196c
permissions -rw-r--r--
reoriented congruence rules in non-explosive direction
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
2c20bcd70fbe proper signature constraints;
wenzelm
parents: 31070
diff changeset
     7
  val linz_tac: Proof.context -> bool -> int -> tactic
2c20bcd70fbe proper signature constraints;
wenzelm
parents: 31070
diff changeset
     8
end
2c20bcd70fbe proper signature constraints;
wenzelm
parents: 31070
diff changeset
     9
2c20bcd70fbe proper signature constraints;
wenzelm
parents: 31070
diff changeset
    10
structure Cooper_Tac: COOPER_TAC =
23274
f997514ad8f4 New Reflected Presburger added to HOL/ex
chaieb
parents:
diff changeset
    11
struct
f997514ad8f4 New Reflected Presburger added to HOL/ex
chaieb
parents:
diff changeset
    12
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 47432
diff changeset
    13
val cooper_ss = simpset_of @{context};
23274
f997514ad8f4 New Reflected Presburger added to HOL/ex
chaieb
parents:
diff changeset
    14
31240
2c20bcd70fbe proper signature constraints;
wenzelm
parents: 31070
diff changeset
    15
fun prepare_for_linz q fm =
23274
f997514ad8f4 New Reflected Presburger added to HOL/ex
chaieb
parents:
diff changeset
    16
  let
f997514ad8f4 New Reflected Presburger added to HOL/ex
chaieb
parents:
diff changeset
    17
    val ps = Logic.strip_params fm
f997514ad8f4 New Reflected Presburger added to HOL/ex
chaieb
parents:
diff changeset
    18
    val hs = map HOLogic.dest_Trueprop (Logic.strip_assums_hyp fm)
f997514ad8f4 New Reflected Presburger added to HOL/ex
chaieb
parents:
diff changeset
    19
    val c = HOLogic.dest_Trueprop (Logic.strip_assums_concl fm)
f997514ad8f4 New Reflected Presburger added to HOL/ex
chaieb
parents:
diff changeset
    20
    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
    21
      if Term.is_dependent P then
23274
f997514ad8f4 New Reflected Presburger added to HOL/ex
chaieb
parents:
diff changeset
    22
        (HOLogic.all_const T $ Abs (s, T, P), n)
f997514ad8f4 New Reflected Presburger added to HOL/ex
chaieb
parents:
diff changeset
    23
      else (incr_boundvars ~1 P, n-1)
f997514ad8f4 New Reflected Presburger added to HOL/ex
chaieb
parents:
diff changeset
    24
    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
    25
    val rhs = hs
23274
f997514ad8f4 New Reflected Presburger added to HOL/ex
chaieb
parents:
diff changeset
    26
    val np = length ps
33004
715566791eb0 always qualify NJ's old List.foldl/foldr in Isabelle/ML;
wenzelm
parents: 32960
diff changeset
    27
    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
    28
      (List.foldr HOLogic.mk_imp c rhs, np) ps
55987
52c22561996d misc tuning;
wenzelm
parents: 55495
diff changeset
    29
    val (vs, _) = List.partition (fn t => q orelse (type_of t) = @{typ nat})
44121
44adaa6db327 old term operations are legacy;
wenzelm
parents: 43594
diff changeset
    30
      (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
    31
    val fm2 = List.foldr mk_all2 fm' vs
23274
f997514ad8f4 New Reflected Presburger added to HOL/ex
chaieb
parents:
diff changeset
    32
  in (fm2, np + length vs, length rhs) end;
f997514ad8f4 New Reflected Presburger added to HOL/ex
chaieb
parents:
diff changeset
    33
f997514ad8f4 New Reflected Presburger added to HOL/ex
chaieb
parents:
diff changeset
    34
(*Object quantifier to meta --*)
55987
52c22561996d misc tuning;
wenzelm
parents: 55495
diff changeset
    35
fun spec_step n th = if n = 0 then th else (spec_step (n - 1) th) RS spec;
23274
f997514ad8f4 New Reflected Presburger added to HOL/ex
chaieb
parents:
diff changeset
    36
f997514ad8f4 New Reflected Presburger added to HOL/ex
chaieb
parents:
diff changeset
    37
(* object implication to meta---*)
55987
52c22561996d misc tuning;
wenzelm
parents: 55495
diff changeset
    38
fun mp_step n th = if n = 0 then th else (mp_step (n - 1) th) RS mp;
23274
f997514ad8f4 New Reflected Presburger added to HOL/ex
chaieb
parents:
diff changeset
    39
f997514ad8f4 New Reflected Presburger added to HOL/ex
chaieb
parents:
diff changeset
    40
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 52131
diff changeset
    41
fun linz_tac ctxt q = Object_Logic.atomize_prems_tac ctxt THEN' SUBGOAL (fn (g, i) =>
23274
f997514ad8f4 New Reflected Presburger added to HOL/ex
chaieb
parents:
diff changeset
    42
  let
f997514ad8f4 New Reflected Presburger added to HOL/ex
chaieb
parents:
diff changeset
    43
    (* Transform the term*)
55987
52c22561996d misc tuning;
wenzelm
parents: 55495
diff changeset
    44
    val (t, np, nh) = prepare_for_linz q g;
23274
f997514ad8f4 New Reflected Presburger added to HOL/ex
chaieb
parents:
diff changeset
    45
    (* 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
    46
    val mod_div_simpset =
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 47432
diff changeset
    47
      put_simpset HOL_basic_ss ctxt
64593
50c715579715 reoriented congruence rules in non-explosive direction
haftmann
parents: 64244
diff changeset
    48
      addsimps @{thms refl mod_add_eq mod_add_left_eq
50c715579715 reoriented congruence rules in non-explosive direction
haftmann
parents: 64244
diff changeset
    49
          mod_add_right_eq
55987
52c22561996d misc tuning;
wenzelm
parents: 55495
diff changeset
    50
          div_add1_eq [symmetric] zdiv_zadd1_eq [symmetric]
52c22561996d misc tuning;
wenzelm
parents: 55495
diff changeset
    51
          mod_self
52c22561996d misc tuning;
wenzelm
parents: 55495
diff changeset
    52
          div_by_0 mod_by_0 div_0 mod_0
64244
e7102c40783c clarified theorem names
haftmann
parents: 64243
diff changeset
    53
          div_by_1 mod_by_1 div_by_Suc_0 mod_by_Suc_0
55987
52c22561996d misc tuning;
wenzelm
parents: 55495
diff changeset
    54
          Suc_eq_plus1}
57514
bdc2c6b40bf2 prefer ac_simps collections over separate name bindings for add and mult
haftmann
parents: 56245
diff changeset
    55
      addsimps @{thms ac_simps}
43594
ef1ddc59b825 modernized some simproc setup;
wenzelm
parents: 42368
diff changeset
    56
      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
    57
    val simpset0 =
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 47432
diff changeset
    58
      put_simpset HOL_basic_ss ctxt
64243
aee949f6642d eliminated irregular aliasses
haftmann
parents: 64242
diff changeset
    59
      addsimps @{thms minus_div_mult_eq_mod [symmetric] Suc_eq_plus1 simp_thms}
55987
52c22561996d misc tuning;
wenzelm
parents: 55495
diff changeset
    60
      |> fold Splitter.add_split @{thms split_zdiv split_zmod split_div' split_min split_max}
23274
f997514ad8f4 New Reflected Presburger added to HOL/ex
chaieb
parents:
diff changeset
    61
    (* 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
    62
    val simpset1 =
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 47432
diff changeset
    63
      put_simpset HOL_basic_ss ctxt
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 60754
diff changeset
    64
      addsimps @{thms zdvd_int} @ [@{thm "of_nat_add"}, @{thm "of_nat_mult"}] @
62348
9a5f43dac883 dropped various legacy fact bindings
haftmann
parents: 61694
diff changeset
    65
        map (fn r => r RS sym) @{thms of_nat_numeral [where ?'a = int] int_int_eq zle_int of_nat_less_iff [where ?'a = int]}
55987
52c22561996d misc tuning;
wenzelm
parents: 55495
diff changeset
    66
      |> Splitter.add_split @{thm zdiff_int_split}
23274
f997514ad8f4 New Reflected Presburger added to HOL/ex
chaieb
parents:
diff changeset
    67
    (*simp rules for elimination of int n*)
f997514ad8f4 New Reflected Presburger added to HOL/ex
chaieb
parents:
diff changeset
    68
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 47432
diff changeset
    69
    val simpset2 =
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 47432
diff changeset
    70
      put_simpset HOL_basic_ss ctxt
62348
9a5f43dac883 dropped various legacy fact bindings
haftmann
parents: 61694
diff changeset
    71
      addsimps [@{thm nat_0_le}, @{thm all_nat}, @{thm ex_nat}, @{thm zero_le_numeral}, @{thm order_refl}(* FIXME: necessary? *),
9a5f43dac883 dropped various legacy fact bindings
haftmann
parents: 61694
diff changeset
    72
        @{thm of_nat_0 [where ?'a = int]}, @{thm of_nat_1  [where ?'a = int]}]
55987
52c22561996d misc tuning;
wenzelm
parents: 55495
diff changeset
    73
      |> fold Simplifier.add_cong @{thms conj_le_cong imp_le_cong}
23274
f997514ad8f4 New Reflected Presburger added to HOL/ex
chaieb
parents:
diff changeset
    74
    (* simp rules for elimination of abs *)
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 47432
diff changeset
    75
    val simpset3 = put_simpset HOL_basic_ss ctxt |> Splitter.add_split @{thm abs_split}
59629
0d77c51b5040 clarified context;
wenzelm
parents: 59621
diff changeset
    76
    val ct = Thm.cterm_of ctxt (HOLogic.mk_Trueprop t)
23274
f997514ad8f4 New Reflected Presburger added to HOL/ex
chaieb
parents:
diff changeset
    77
    (* Theorem for the nat --> int transformation *)
f997514ad8f4 New Reflected Presburger added to HOL/ex
chaieb
parents:
diff changeset
    78
    val pre_thm = Seq.hd (EVERY
f997514ad8f4 New Reflected Presburger added to HOL/ex
chaieb
parents:
diff changeset
    79
      [simp_tac mod_div_simpset 1, simp_tac simpset0 1,
f997514ad8f4 New Reflected Presburger added to HOL/ex
chaieb
parents:
diff changeset
    80
       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
    81
       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
    82
      (Thm.trivial ct))
58963
26bf09b95dda proper context for assume_tac (atac remains as fall-back without context);
wenzelm
parents: 57514
diff changeset
    83
    fun assm_tac i = REPEAT_DETERM_N nh (assume_tac ctxt i)
23274
f997514ad8f4 New Reflected Presburger added to HOL/ex
chaieb
parents:
diff changeset
    84
    (* The result of the quantifier elimination *)
55987
52c22561996d misc tuning;
wenzelm
parents: 55495
diff changeset
    85
    val (th, tac) =
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 58963
diff changeset
    86
      (case Thm.prop_of pre_thm of
56245
84fc7dfa3cd4 more qualified names;
wenzelm
parents: 55987
diff changeset
    87
        Const (@{const_name Pure.imp}, _) $ (Const (@{const_name Trueprop}, _) $ t1) $ _ =>
55987
52c22561996d misc tuning;
wenzelm
parents: 55495
diff changeset
    88
          let
60325
6fc771cb42eb clarified context;
wenzelm
parents: 59629
diff changeset
    89
            val pth = linzqe_oracle (ctxt, Envir.eta_long [] t1)
55987
52c22561996d misc tuning;
wenzelm
parents: 55495
diff changeset
    90
          in
52c22561996d misc tuning;
wenzelm
parents: 55495
diff changeset
    91
            ((pth RS iffD2) RS pre_thm,
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60325
diff changeset
    92
              assm_tac (i + 1) THEN (if q then I else TRY) (resolve_tac ctxt [TrueI] i))
55987
52c22561996d misc tuning;
wenzelm
parents: 55495
diff changeset
    93
          end
52c22561996d misc tuning;
wenzelm
parents: 55495
diff changeset
    94
      | _ => (pre_thm, assm_tac i))
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60325
diff changeset
    95
  in resolve_tac ctxt [mp_step nh (spec_step np th)] i THEN tac end);
23274
f997514ad8f4 New Reflected Presburger added to HOL/ex
chaieb
parents:
diff changeset
    96
23590
ad95084a5c63 renamed ObjectLogic.atomize_tac to ObjectLogic.atomize_prems_tac;
wenzelm
parents: 23469
diff changeset
    97
end