src/HOL/Decision_Procs/mir_tac.ML
author blanchet
Wed, 27 Aug 2014 08:41:12 +0200
changeset 58043 a90847f03ec8
parent 57514 bdc2c6b40bf2
child 58956 a816aa3ff391
permissions -rw-r--r--
avoid 'PolyML.makestring'
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 = 
39159
0dec18004e75 more antiquotations;
wenzelm
parents: 38558
diff changeset
    14
let val ths = [@{thm "real_of_int_inject"}, @{thm "real_of_int_less_iff"}, @{thm "real_of_int_le_iff"}]
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51369
diff changeset
    15
in simpset_of (@{context} delsimps ths addsimps (map (fn th => th RS sym) ths))
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
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
    17
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
    18
val nT = HOLogic.natT;
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 45654
diff changeset
    19
  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
    20
39159
0dec18004e75 more antiquotations;
wenzelm
parents: 38558
diff changeset
    21
  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
    22
                 @{thm "add_Suc"}, @{thm add_numeral_left}, @{thm mult_numeral_left(1)}] @
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 45654
diff changeset
    23
                 (map (fn th => th RS sym) [@{thm "numeral_1_eq_1"}])
25481
aa16cd919dcc dropped legacy ml bindings
haftmann
parents: 23880
diff changeset
    24
                 @ @{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
    25
  val ths = [@{thm "mult_numeral_1"}, @{thm "mult_numeral_1_right"}, 
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 45654
diff changeset
    26
             @{thm real_of_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
    27
             @{thm "real_of_nat_Suc"}, @{thm "real_of_nat_one"}, @{thm "real_of_one"},
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
    28
             @{thm "real_of_int_zero"}, @{thm "real_of_nat_zero"},
36308
bbcfeddeafbb dequalified fact name
haftmann
parents: 35625
diff changeset
    29
             @{thm "divide_zero"}, 
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
    30
             @{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
    31
             @{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
    32
             @{thm uminus_add_conv_diff [symmetric]}, @{thm "minus_divide_left"}]
45654
cf10bde35973 more antiquotations;
wenzelm
parents: 45620
diff changeset
    33
val comp_ths = 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
    34
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
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
    36
val mod_div_equality' = @{thm "mod_div_equality'"};
30224
79136ce06bdb removed and renamed redundant lemmas
nipkow
parents: 30034
diff changeset
    37
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
    38
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 47432
diff changeset
    39
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
    40
  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
    41
    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
    42
    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
    43
    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
    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: 39159
diff changeset
    45
      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
    46
        (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
    47
      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
    48
    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
    49
      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
    50
(*    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
    51
    val np = length ps
33004
715566791eb0 always qualify NJ's old List.foldl/foldr in Isabelle/ML;
wenzelm
parents: 32740
diff changeset
    52
    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
    53
      (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
    54
    val (vs, _) = List.partition (fn t => q orelse (type_of t) = nT)
44121
44adaa6db327 old term operations are legacy;
wenzelm
parents: 43594
diff changeset
    55
      (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
    56
    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
    57
  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
    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 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
    60
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
    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
(* 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
    63
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
    64
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
    65
42368
3b8498ac2314 proper subgoal addressing via SUBGOAL/CSUBGOAL -- assuming these tactics did not handle Subscript in any special way;
wenzelm
parents: 42364
diff changeset
    66
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
    67
    Object_Logic.atomize_prems_tac ctxt
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51369
diff changeset
    68
        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
    69
          addsimps [@{thm "abs_ge_zero"}] addsimps @{thms simp_thms})
42368
3b8498ac2314 proper subgoal addressing via SUBGOAL/CSUBGOAL -- assuming these tactics did not handle Subscript in any special way;
wenzelm
parents: 42364
diff changeset
    70
        THEN' (REPEAT_DETERM o split_tac [@{thm "split_min"}, @{thm "split_max"}, @{thm "abs_split"}])
3b8498ac2314 proper subgoal addressing via SUBGOAL/CSUBGOAL -- assuming these tactics did not handle Subscript in any special way;
wenzelm
parents: 42364
diff changeset
    71
        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
    72
  let
42361
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 42083
diff changeset
    73
    val thy = Proof_Context.theory_of ctxt
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
    74
    (* Transform the term*)
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 47432
diff changeset
    75
    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
    76
    (* 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
    77
    val mod_div_simpset = put_simpset HOL_basic_ss ctxt
30224
79136ce06bdb removed and renamed redundant lemmas
nipkow
parents: 30034
diff changeset
    78
                        addsimps [refl, mod_add_eq, 
47142
d64fa2ca54b8 remove redundant lemmas
huffman
parents: 47108
diff changeset
    79
                                  @{thm mod_self},
d64fa2ca54b8 remove redundant lemmas
huffman
parents: 47108
diff changeset
    80
                                  @{thm div_0}, @{thm mod_0},
30031
bd786c37af84 Removed redundant lemmas
nipkow
parents: 29948
diff changeset
    81
                                  @{thm "div_by_1"}, @{thm "mod_by_1"}, @{thm "div_1"}, @{thm "mod_1"},
31790
05c92381363c corrected and unified thm names
nipkow
parents: 31240
diff changeset
    82
                                  @{thm "Suc_eq_plus1"}]
57514
bdc2c6b40bf2 prefer ac_simps collections over separate name bindings for add and mult
haftmann
parents: 56245
diff changeset
    83
                        addsimps @{thms add.assoc add.commute add.left_commute}
43594
ef1ddc59b825 modernized some simproc setup;
wenzelm
parents: 42368
diff changeset
    84
                        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
    85
    val simpset0 = put_simpset HOL_basic_ss ctxt
31790
05c92381363c corrected and unified thm names
nipkow
parents: 31240
diff changeset
    86
      addsimps [mod_div_equality', @{thm Suc_eq_plus1}]
23318
6d68b07ab5cf tuned tactic
chaieb
parents: 23264
diff changeset
    87
      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
    88
      |> 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
    89
          [@{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
    90
            @{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
    91
    (* 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
    92
    val simpset1 = put_simpset HOL_basic_ss ctxt
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 45654
diff changeset
    93
      addsimps [@{thm "zdvd_int"}] @ map (fn r => r RS sym)
23381
da53d861d106 Fixed Problem with ML-bindings for thm names;
chaieb
parents: 23318
diff changeset
    94
        [@{thm "int_int_eq"}, @{thm "zle_int"}, @{thm "zless_int"}, @{thm "zadd_int"}, 
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 45654
diff changeset
    95
         @{thm nat_numeral}, @{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
    96
      |> 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
    97
    (*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
    98
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51369
diff changeset
    99
    val simpset2 = put_simpset HOL_basic_ss ctxt
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 45654
diff changeset
   100
      addsimps [@{thm "nat_0_le"}, @{thm "all_nat"}, @{thm "ex_nat"}, @{thm zero_le_numeral}, 
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 45654
diff changeset
   101
                @{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
   102
      |> 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
   103
    (* simp rules for elimination of abs *)
27456
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 26939
diff changeset
   104
    val ct = cterm_of thy (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
   105
    (* 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
   106
    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
   107
      [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
   108
       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
   109
       TRY (simp_tac (put_simpset mir_ss ctxt) 1)]
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36692
diff changeset
   110
      (Thm.trivial ct))
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
   111
    fun assm_tac i = REPEAT_DETERM_N nh (assume_tac i)
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
   112
    (* The result of the quantifier elimination *)
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
    val (th, tac) = case (prop_of pre_thm) of
56245
84fc7dfa3cd4 more qualified names;
wenzelm
parents: 55506
diff changeset
   114
        Const (@{const_name Pure.imp}, _) $ (Const (@{const_name Trueprop}, _) $ t1) $ _ =>
28290
4cc2b6046258 simplified oracle interface;
wenzelm
parents: 27456
diff changeset
   115
    let val pth =
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
          (* If quick_and_dirty then run without proof generation as oracle*)
52059
2f970c7f722b proper option quick_and_dirty;
wenzelm
parents: 51717
diff changeset
   117
             if Config.get ctxt quick_and_dirty
52131
366fa32ee2a3 tuned signature;
wenzelm
parents: 52059
diff changeset
   118
             then mirfr_oracle (false, cterm_of thy (Envir.eta_long [] t1))
366fa32ee2a3 tuned signature;
wenzelm
parents: 52059
diff changeset
   119
             else mirfr_oracle (true, cterm_of thy (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
   120
    in 
55506
46f3e31c5a87 removed dead code;
wenzelm
parents: 54742
diff changeset
   121
       ((pth RS iffD2) RS pre_thm,
46f3e31c5a87 removed dead code;
wenzelm
parents: 54742
diff changeset
   122
        assm_tac (i + 1) THEN (if q then I else TRY) (rtac 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
   123
    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
   124
      | _ => (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
   125
  in rtac (((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
   126
23590
ad95084a5c63 renamed ObjectLogic.atomize_tac to ObjectLogic.atomize_prems_tac;
wenzelm
parents: 23381
diff changeset
   127
end