src/HOL/Decision_Procs/mir_tac.ML
author haftmann
Sun, 16 Oct 2016 09:31:05 +0200
changeset 64244 e7102c40783c
parent 64243 aee949f6642d
child 64593 50c715579715
permissions -rw-r--r--
clarified theorem names
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/mir_tac.ML
23858
5500610fe1e5 adapted to new code generator framework
haftmann
parents: 23590
diff changeset
     2
    Author:     Amine Chaieb, TU Muenchen
5500610fe1e5 adapted to new code generator framework
haftmann
parents: 23590
diff changeset
     3
*)
5500610fe1e5 adapted to new code generator framework
haftmann
parents: 23590
diff changeset
     4
31240
2c20bcd70fbe proper signature constraints;
wenzelm
parents: 30939
diff changeset
     5
signature MIR_TAC =
2c20bcd70fbe proper signature constraints;
wenzelm
parents: 30939
diff changeset
     6
sig
2c20bcd70fbe proper signature constraints;
wenzelm
parents: 30939
diff changeset
     7
  val mir_tac: Proof.context -> bool -> int -> tactic
2c20bcd70fbe proper signature constraints;
wenzelm
parents: 30939
diff changeset
     8
end
2c20bcd70fbe proper signature constraints;
wenzelm
parents: 30939
diff changeset
     9
55506
46f3e31c5a87 removed dead code;
wenzelm
parents: 54742
diff changeset
    10
structure Mir_Tac: MIR_TAC =
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    11
struct
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    12
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    13
val mir_ss = 
61652
90c65a811257 MIR decision procedure again working
paulson <lp15@cam.ac.uk>
parents: 61609
diff changeset
    14
  simpset_of (@{context} delsimps [@{thm "of_int_eq_iff"}, @{thm "of_int_less_iff"}, @{thm "of_int_le_iff"}] 
90c65a811257 MIR decision procedure again working
paulson <lp15@cam.ac.uk>
parents: 61609
diff changeset
    15
               addsimps @{thms "iff_real_of_int"});
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    16
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    17
val nT = HOLogic.natT;
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 45654
diff changeset
    18
  val nat_arith = [@{thm diff_nat_numeral}];
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    19
39159
0dec18004e75 more antiquotations;
wenzelm
parents: 38558
diff changeset
    20
  val comp_arith = [@{thm "Let_def"}, @{thm "if_False"}, @{thm "if_True"}, @{thm "add_0"},
57514
bdc2c6b40bf2 prefer ac_simps collections over separate name bindings for add and mult
haftmann
parents: 56245
diff changeset
    21
                 @{thm "add_Suc"}, @{thm add_numeral_left}, @{thm mult_numeral_left(1)}] @
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61652
diff changeset
    22
                 (map (fn th => th RS sym) [@{thm "numeral_One"}])
25481
aa16cd919dcc dropped legacy ml bindings
haftmann
parents: 23880
diff changeset
    23
                 @ @{thms arith_simps} @ nat_arith @ @{thms rel_simps} 
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    24
  val ths = [@{thm "mult_numeral_1"}, @{thm "mult_numeral_1_right"}, 
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60754
diff changeset
    25
             @{thm of_nat_numeral},
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60754
diff changeset
    26
             @{thm "of_nat_Suc"}, @{thm "of_nat_1"},
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60754
diff changeset
    27
             @{thm "of_int_0"}, @{thm "of_nat_0"},
64240
eabf80376aab more standardized names
haftmann
parents: 62348
diff changeset
    28
             @{thm "div_by_0"}, 
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    29
             @{thm "divide_divide_eq_left"}, @{thm "times_divide_eq_right"}, 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    30
             @{thm "times_divide_eq_left"}, @{thm "divide_divide_eq_right"},
54230
b1d955791529 more simplification rules on unary and binary minus
haftmann
parents: 52131
diff changeset
    31
             @{thm uminus_add_conv_diff [symmetric]}, @{thm "minus_divide_left"}]
61652
90c65a811257 MIR decision procedure again working
paulson <lp15@cam.ac.uk>
parents: 61609
diff changeset
    32
val comp_ths = distinct Thm.eq_thm (ths @ comp_arith @ @{thms simp_thms});
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    33
30224
79136ce06bdb removed and renamed redundant lemmas
nipkow
parents: 30034
diff changeset
    34
val mod_add_eq = @{thm "mod_add_eq"} RS sym;
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    35
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 47432
diff changeset
    36
fun prepare_for_mir q fm = 
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    37
  let
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    38
    val ps = Logic.strip_params fm
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    39
    val hs = map HOLogic.dest_Trueprop (Logic.strip_assums_hyp fm)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    40
    val c = HOLogic.dest_Trueprop (Logic.strip_assums_concl fm)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    41
    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: 39159
diff changeset
    42
      if Term.is_dependent P then
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    43
        (HOLogic.all_const T $ Abs (s, T, P), n)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    44
      else (incr_boundvars ~1 P, n-1)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    45
    fun mk_all2 (v, t) = HOLogic.all_const (fastype_of v) $ lambda v t;
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    46
      val rhs = hs
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    47
(*    val (rhs,irhs) = List.partition (relevant (rev ps)) hs *)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    48
    val np = length ps
33004
715566791eb0 always qualify NJ's old List.foldl/foldr in Isabelle/ML;
wenzelm
parents: 32740
diff changeset
    49
    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: 32740
diff changeset
    50
      (List.foldr HOLogic.mk_imp c rhs, np) ps
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    51
    val (vs, _) = List.partition (fn t => q orelse (type_of t) = nT)
44121
44adaa6db327 old term operations are legacy;
wenzelm
parents: 43594
diff changeset
    52
      (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: 32740
diff changeset
    53
    val fm2 = List.foldr mk_all2 fm' vs
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    54
  in (fm2, np + length vs, length rhs) end;
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    55
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    56
(*Object quantifier to meta --*)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    57
fun spec_step n th = if (n=0) then th else (spec_step (n-1) th) RS spec ;
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    58
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    59
(* object implication to meta---*)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    60
fun mp_step n th = if (n=0) then th else (mp_step (n-1) th) RS mp;
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    61
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    62
42368
3b8498ac2314 proper subgoal addressing via SUBGOAL/CSUBGOAL -- assuming these tactics did not handle Subscript in any special way;
wenzelm
parents: 42364
diff changeset
    63
fun mir_tac ctxt q = 
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 54230
diff changeset
    64
    Object_Logic.atomize_prems_tac ctxt
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51369
diff changeset
    65
        THEN' simp_tac (put_simpset HOL_basic_ss ctxt
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51369
diff changeset
    66
          addsimps [@{thm "abs_ge_zero"}] addsimps @{thms simp_thms})
58956
a816aa3ff391 proper context for compose_tac, Splitter.split_tac (relevant for unify trace options);
wenzelm
parents: 57514
diff changeset
    67
        THEN' (REPEAT_DETERM o split_tac ctxt [@{thm "split_min"}, @{thm "split_max"}, @{thm "abs_split"}])
42368
3b8498ac2314 proper subgoal addressing via SUBGOAL/CSUBGOAL -- assuming these tactics did not handle Subscript in any special way;
wenzelm
parents: 42364
diff changeset
    68
        THEN' SUBGOAL (fn (g, i) =>
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    69
  let
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    70
    (* Transform the term*)
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 47432
diff changeset
    71
    val (t,np,nh) = prepare_for_mir q g
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    72
    (* Some simpsets for dealing with mod div abs and nat*)
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51369
diff changeset
    73
    val mod_div_simpset = put_simpset HOL_basic_ss ctxt
30224
79136ce06bdb removed and renamed redundant lemmas
nipkow
parents: 30034
diff changeset
    74
                        addsimps [refl, mod_add_eq, 
47142
d64fa2ca54b8 remove redundant lemmas
huffman
parents: 47108
diff changeset
    75
                                  @{thm mod_self},
d64fa2ca54b8 remove redundant lemmas
huffman
parents: 47108
diff changeset
    76
                                  @{thm div_0}, @{thm mod_0},
64244
e7102c40783c clarified theorem names
haftmann
parents: 64243
diff changeset
    77
                                  @{thm div_by_1}, @{thm mod_by_1}, @{thm div_by_Suc_0}, @{thm mod_by_Suc_0},
31790
05c92381363c corrected and unified thm names
nipkow
parents: 31240
diff changeset
    78
                                  @{thm "Suc_eq_plus1"}]
57514
bdc2c6b40bf2 prefer ac_simps collections over separate name bindings for add and mult
haftmann
parents: 56245
diff changeset
    79
                        addsimps @{thms add.assoc add.commute add.left_commute}
43594
ef1ddc59b825 modernized some simproc setup;
wenzelm
parents: 42368
diff changeset
    80
                        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: 51369
diff changeset
    81
    val simpset0 = put_simpset HOL_basic_ss ctxt
64243
aee949f6642d eliminated irregular aliasses
haftmann
parents: 64242
diff changeset
    82
      addsimps @{thms minus_div_mult_eq_mod [symmetric] Suc_eq_plus1}
23318
6d68b07ab5cf tuned tactic
chaieb
parents: 23264
diff changeset
    83
      addsimps comp_ths
45620
f2a587696afb modernized some old-style infix operations, which were left over from the time of ML proof scripts;
wenzelm
parents: 44121
diff changeset
    84
      |> 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
    85
          [@{thm "split_zdiv"}, @{thm "split_zmod"}, @{thm "split_div'"},
f2a587696afb modernized some old-style infix operations, which were left over from the time of ML proof scripts;
wenzelm
parents: 44121
diff changeset
    86
            @{thm "split_min"}, @{thm "split_max"}]
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    87
    (* Simp rules for changing (n::int) to int n *)
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51369
diff changeset
    88
    val simpset1 = 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: 61652
diff changeset
    89
      addsimps [@{thm "zdvd_int"}, @{thm "of_nat_add"}, @{thm "of_nat_mult"}] @ 
62348
9a5f43dac883 dropped various legacy fact bindings
haftmann
parents: 61694
diff changeset
    90
          map (fn r => r RS sym) [@{thm "int_int_eq"}, @{thm "zle_int"}, @{thm "of_nat_less_iff"}, @{thm nat_numeral}]
45620
f2a587696afb modernized some old-style infix operations, which were left over from the time of ML proof scripts;
wenzelm
parents: 44121
diff changeset
    91
      |> Splitter.add_split @{thm "zdiff_int_split"}
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    92
    (*simp rules for elimination of int n*)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    93
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51369
diff changeset
    94
    val simpset2 = put_simpset HOL_basic_ss ctxt
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 45654
diff changeset
    95
      addsimps [@{thm "nat_0_le"}, @{thm "all_nat"}, @{thm "ex_nat"}, @{thm zero_le_numeral}, 
62348
9a5f43dac883 dropped various legacy fact bindings
haftmann
parents: 61694
diff changeset
    96
                @{thm "of_nat_0"}, @{thm "of_nat_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
    97
      |> fold Simplifier.add_cong [@{thm "conj_le_cong"}, @{thm "imp_le_cong"}]
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    98
    (* simp rules for elimination of abs *)
59629
0d77c51b5040 clarified context;
wenzelm
parents: 59621
diff changeset
    99
    val ct = Thm.cterm_of ctxt (HOLogic.mk_Trueprop t)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   100
    (* Theorem for the nat --> int transformation *)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   101
    val pre_thm = Seq.hd (EVERY
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   102
      [simp_tac mod_div_simpset 1, simp_tac simpset0 1,
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51369
diff changeset
   103
       TRY (simp_tac simpset1 1), TRY (simp_tac simpset2 1),
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51369
diff changeset
   104
       TRY (simp_tac (put_simpset mir_ss ctxt) 1)]
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36692
diff changeset
   105
      (Thm.trivial ct))
58963
26bf09b95dda proper context for assume_tac (atac remains as fall-back without context);
wenzelm
parents: 58956
diff changeset
   106
    fun assm_tac i = REPEAT_DETERM_N nh (assume_tac ctxt i)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   107
    (* The result of the quantifier elimination *)
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 58963
diff changeset
   108
    val (th, tac) =
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 58963
diff changeset
   109
      case Thm.prop_of pre_thm of
56245
84fc7dfa3cd4 more qualified names;
wenzelm
parents: 55506
diff changeset
   110
        Const (@{const_name Pure.imp}, _) $ (Const (@{const_name Trueprop}, _) $ t1) $ _ =>
60325
6fc771cb42eb clarified context;
wenzelm
parents: 59629
diff changeset
   111
    let
6fc771cb42eb clarified context;
wenzelm
parents: 59629
diff changeset
   112
      val pth = mirfr_oracle (ctxt, Envir.eta_long [] t1)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   113
    in 
55506
46f3e31c5a87 removed dead code;
wenzelm
parents: 54742
diff changeset
   114
       ((pth RS iffD2) RS pre_thm,
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60325
diff changeset
   115
        assm_tac (i + 1) THEN (if q then I else TRY) (resolve_tac ctxt [TrueI] i))
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   116
    end
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   117
      | _ => (pre_thm, assm_tac i)
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60325
diff changeset
   118
  in resolve_tac ctxt [((mp_step nh) o (spec_step np)) th] i THEN tac end);
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   119
23590
ad95084a5c63 renamed ObjectLogic.atomize_tac to ObjectLogic.atomize_prems_tac;
wenzelm
parents: 23381
diff changeset
   120
end