src/HOL/Decision_Procs/mir_tac.ML
author wenzelm
Sat, 16 Apr 2011 18:11:20 +0200
changeset 42364 8c674b3b8e44
parent 42361 23f352990944
child 42368 3b8498ac2314
permissions -rw-r--r--
eliminated old List.nth; tuned;
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
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: 30939
diff changeset
     8
  val mir_tac: Proof.context -> bool -> int -> tactic
2c20bcd70fbe proper signature constraints;
wenzelm
parents: 30939
diff changeset
     9
  val setup: theory -> theory
2c20bcd70fbe proper signature constraints;
wenzelm
parents: 30939
diff changeset
    10
end
2c20bcd70fbe proper signature constraints;
wenzelm
parents: 30939
diff changeset
    11
29788
1b80ebe713a4 established session HOL-Reflection
haftmann
parents: 29265
diff changeset
    12
structure 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
    13
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
    14
32740
9dd0a2f83429 explicit indication of Unsynchronized.ref;
wenzelm
parents: 31790
diff changeset
    15
val trace = Unsynchronized.ref false;
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
fun trace_msg s = if !trace then tracing s else ();
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 mir_ss = 
39159
0dec18004e75 more antiquotations;
wenzelm
parents: 38558
diff changeset
    19
let val ths = [@{thm "real_of_int_inject"}, @{thm "real_of_int_less_iff"}, @{thm "real_of_int_le_iff"}]
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
in @{simpset} delsimps ths addsimps (map (fn th => th RS sym) ths)
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
    21
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
    22
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
    23
val nT = HOLogic.natT;
39159
0dec18004e75 more antiquotations;
wenzelm
parents: 38558
diff changeset
    24
  val nat_arith = [@{thm "add_nat_number_of"}, @{thm "diff_nat_number_of"},
0dec18004e75 more antiquotations;
wenzelm
parents: 38558
diff changeset
    25
                       @{thm "mult_nat_number_of"}, @{thm "eq_nat_number_of"}, @{thm "less_nat_number_of"}];
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
    26
39159
0dec18004e75 more antiquotations;
wenzelm
parents: 38558
diff changeset
    27
  val comp_arith = [@{thm "Let_def"}, @{thm "if_False"}, @{thm "if_True"}, @{thm "add_0"},
0dec18004e75 more antiquotations;
wenzelm
parents: 38558
diff changeset
    28
                 @{thm "add_Suc"}, @{thm "add_number_of_left"}, @{thm "mult_number_of_left"},
0dec18004e75 more antiquotations;
wenzelm
parents: 38558
diff changeset
    29
                 @{thm "Suc_eq_plus1"}] @
0dec18004e75 more antiquotations;
wenzelm
parents: 38558
diff changeset
    30
                 (map (fn th => th RS sym) [@{thm "numeral_1_eq_1"}, @{thm "numeral_0_eq_0"}])
25481
aa16cd919dcc dropped legacy ml bindings
haftmann
parents: 23880
diff changeset
    31
                 @ @{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
    32
  val ths = [@{thm "mult_numeral_1"}, @{thm "mult_numeral_1_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
    33
             @{thm "real_of_nat_number_of"},
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
             @{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
    35
             @{thm "real_of_int_zero"}, @{thm "real_of_nat_zero"},
36308
bbcfeddeafbb dequalified fact name
haftmann
parents: 35625
diff changeset
    36
             @{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
    37
             @{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
    38
             @{thm "times_divide_eq_left"}, @{thm "divide_divide_eq_right"},
37887
2ae085b07f2f diff_minus subsumes diff_def
haftmann
parents: 36945
diff changeset
    39
             @{thm "diff_minus"}, @{thm "minus_divide_left"}]
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
val comp_ths = ths @ comp_arith @ simp_thms 
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
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
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 zdvd_int = @{thm "zdvd_int"};
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
val zdiff_int_split = @{thm "zdiff_int_split"};
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
val all_nat = @{thm "all_nat"};
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 ex_nat = @{thm "ex_nat"};
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 number_of1 = @{thm "number_of1"};
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 number_of2 = @{thm "number_of2"};
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 split_zdiv = @{thm "split_zdiv"};
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 split_zmod = @{thm "split_zmod"};
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 mod_div_equality' = @{thm "mod_div_equality'"};
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
    52
val split_div' = @{thm "split_div'"};
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
    53
val imp_le_cong = @{thm "imp_le_cong"};
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 conj_le_cong = @{thm "conj_le_cong"};
30224
79136ce06bdb removed and renamed redundant lemmas
nipkow
parents: 30034
diff changeset
    55
val mod_add_eq = @{thm "mod_add_eq"} RS sym;
30034
60f64f112174 removed redundant thms
nipkow
parents: 30031
diff changeset
    56
val mod_add_left_eq = @{thm "mod_add_left_eq"} RS sym;
60f64f112174 removed redundant thms
nipkow
parents: 30031
diff changeset
    57
val mod_add_right_eq = @{thm "mod_add_right_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
    58
val nat_div_add_eq = @{thm "div_add1_eq"} RS sym;
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
val int_div_add_eq = @{thm "zdiv_zadd1_eq"} RS sym;
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
val ZDIVISION_BY_ZERO_MOD = @{thm "DIVISION_BY_ZERO"} RS conjunct2;
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
val ZDIVISION_BY_ZERO_DIV = @{thm "DIVISION_BY_ZERO"} RS conjunct1;
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
27456
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 26939
diff changeset
    63
fun prepare_for_mir thy 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
    64
  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
    65
    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
    66
    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
    67
    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
    68
    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
    69
      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
    70
        (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
    71
      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
    72
    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
    73
      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
    74
(*    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
    75
    val np = length ps
33004
715566791eb0 always qualify NJ's old List.foldl/foldr in Isabelle/ML;
wenzelm
parents: 32740
diff changeset
    76
    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
    77
      (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
    78
    val (vs, _) = List.partition (fn t => q orelse (type_of t) = nT)
29265
5b4247055bd7 moved old add_term_vars, add_term_frees etc. to structure OldTerm;
wenzelm
parents: 28952
diff changeset
    79
      (OldTerm.term_frees fm' @ OldTerm.term_vars fm');
33004
715566791eb0 always qualify NJ's old List.foldl/foldr in Isabelle/ML;
wenzelm
parents: 32740
diff changeset
    80
    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
    81
  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
    82
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
    83
(*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
    84
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
    85
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
    86
(* 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
    87
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
    88
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
    89
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
    90
fun mir_tac ctxt q i = 
35625
9c818cab0dd0 modernized structure Object_Logic;
wenzelm
parents: 35050
diff changeset
    91
    Object_Logic.atomize_prems_tac i
31240
2c20bcd70fbe proper signature constraints;
wenzelm
parents: 30939
diff changeset
    92
        THEN simp_tac (HOL_basic_ss addsimps [@{thm "abs_ge_zero"}] addsimps simp_thms) i
2c20bcd70fbe proper signature constraints;
wenzelm
parents: 30939
diff changeset
    93
        THEN REPEAT_DETERM (split_tac [@{thm "split_min"}, @{thm "split_max"}, @{thm "abs_split"}] i)
28290
4cc2b6046258 simplified oracle interface;
wenzelm
parents: 27456
diff changeset
    94
        THEN (fn st =>
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
    95
  let
42364
8c674b3b8e44 eliminated old List.nth;
wenzelm
parents: 42361
diff changeset
    96
    val g = nth (prems_of st) (i - 1)
42361
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 42083
diff changeset
    97
    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
    98
    (* Transform the term*)
27456
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 26939
diff changeset
    99
    val (t,np,nh) = prepare_for_mir thy 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
   100
    (* Some simpsets for dealing with mod div abs and nat*)
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 mod_div_simpset = HOL_basic_ss 
30224
79136ce06bdb removed and renamed redundant lemmas
nipkow
parents: 30034
diff changeset
   102
                        addsimps [refl, mod_add_eq, 
28290
4cc2b6046258 simplified oracle interface;
wenzelm
parents: 27456
diff changeset
   103
                                  @{thm "mod_self"}, @{thm "zmod_self"},
4cc2b6046258 simplified oracle interface;
wenzelm
parents: 27456
diff changeset
   104
                                  @{thm "zdiv_zero"},@{thm "zmod_zero"},@{thm "div_0"}, @{thm "mod_0"},
30031
bd786c37af84 Removed redundant lemmas
nipkow
parents: 29948
diff changeset
   105
                                  @{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
   106
                                  @{thm "Suc_eq_plus1"}]
28290
4cc2b6046258 simplified oracle interface;
wenzelm
parents: 27456
diff changeset
   107
                        addsimps @{thms add_ac}
30939
207ec81543f6 added both cancel_div_mod_procs
haftmann
parents: 30509
diff changeset
   108
                        addsimprocs [cancel_div_mod_nat_proc, cancel_div_mod_int_proc]
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
   109
    val simpset0 = HOL_basic_ss
31790
05c92381363c corrected and unified thm names
nipkow
parents: 31240
diff changeset
   110
      addsimps [mod_div_equality', @{thm Suc_eq_plus1}]
23318
6d68b07ab5cf tuned tactic
chaieb
parents: 23264
diff changeset
   111
      addsimps comp_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
   112
      addsplits [@{thm "split_zdiv"}, @{thm "split_zmod"}, @{thm "split_div'"}, @{thm "split_min"}, @{thm "split_max"}]
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
    (* Simp rules for changing (n::int) to 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
   114
    val simpset1 = HOL_basic_ss
23381
da53d861d106 Fixed Problem with ML-bindings for thm names;
chaieb
parents: 23318
diff changeset
   115
      addsimps [@{thm "nat_number_of_def"}, @{thm "zdvd_int"}] @ map (fn r => r RS sym)
da53d861d106 Fixed Problem with ML-bindings for thm names;
chaieb
parents: 23318
diff changeset
   116
        [@{thm "int_int_eq"}, @{thm "zle_int"}, @{thm "zless_int"}, @{thm "zadd_int"}, 
da53d861d106 Fixed Problem with ML-bindings for thm names;
chaieb
parents: 23318
diff changeset
   117
         @{thm "zmult_int"}]
da53d861d106 Fixed Problem with ML-bindings for thm names;
chaieb
parents: 23318
diff changeset
   118
      addsplits [@{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
   119
    (*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
   120
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
   121
    val simpset2 = HOL_basic_ss
23381
da53d861d106 Fixed Problem with ML-bindings for thm names;
chaieb
parents: 23318
diff changeset
   122
      addsimps [@{thm "nat_0_le"}, @{thm "all_nat"}, @{thm "ex_nat"}, @{thm "number_of1"}, 
da53d861d106 Fixed Problem with ML-bindings for thm names;
chaieb
parents: 23318
diff changeset
   123
                @{thm "number_of2"}, @{thm "int_0"}, @{thm "int_1"}]
da53d861d106 Fixed Problem with ML-bindings for thm names;
chaieb
parents: 23318
diff changeset
   124
      addcongs [@{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
   125
    (* simp rules for elimination of abs *)
27456
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 26939
diff changeset
   126
    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
   127
    (* 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
   128
    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
   129
      [simp_tac mod_div_simpset 1, simp_tac simpset0 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
   130
       TRY (simp_tac simpset1 1), TRY (simp_tac simpset2 1), TRY (simp_tac mir_ss 1)]
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36692
diff changeset
   131
      (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
   132
    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
   133
    (* 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
   134
    val (th, tac) = case (prop_of pre_thm) of
38558
32ad17fe2b9c tuned quotes
haftmann
parents: 38549
diff changeset
   135
        Const ("==>", _) $ (Const (@{const_name Trueprop}, _) $ t1) $ _ =>
28290
4cc2b6046258 simplified oracle interface;
wenzelm
parents: 27456
diff changeset
   136
    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
   137
          (* If quick_and_dirty then run without proof generation as oracle*)
28290
4cc2b6046258 simplified oracle interface;
wenzelm
parents: 27456
diff changeset
   138
             if !quick_and_dirty
4cc2b6046258 simplified oracle interface;
wenzelm
parents: 27456
diff changeset
   139
             then mirfr_oracle (false, cterm_of thy (Pattern.eta_long [] t1))
4cc2b6046258 simplified oracle interface;
wenzelm
parents: 27456
diff changeset
   140
             else mirfr_oracle (true, cterm_of thy (Pattern.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
   141
    in 
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
   142
          (trace_msg ("calling procedure with term:\n" ^
26939
1035c89b4c02 moved global pretty/string_of functions from Sign to Syntax;
wenzelm
parents: 25481
diff changeset
   143
             Syntax.string_of_term ctxt 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
   144
           ((pth RS iffD2) RS pre_thm,
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
   145
            assm_tac (i + 1) THEN (if q then I else TRY) (rtac TrueI 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
   146
    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
   147
      | _ => (pre_thm, assm_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
   148
  in (rtac (((mp_step nh) o (spec_step np)) th) 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
   149
      THEN tac) st
27456
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 26939
diff changeset
   150
  end handle Subscript => no_tac st);
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
   151
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
   152
val setup =
31240
2c20bcd70fbe proper signature constraints;
wenzelm
parents: 30939
diff changeset
   153
  Method.setup @{binding mir}
2c20bcd70fbe proper signature constraints;
wenzelm
parents: 30939
diff changeset
   154
    let
2c20bcd70fbe proper signature constraints;
wenzelm
parents: 30939
diff changeset
   155
      val parse_flag = Args.$$$ "no_quantify" >> K (K false)
2c20bcd70fbe proper signature constraints;
wenzelm
parents: 30939
diff changeset
   156
    in
2c20bcd70fbe proper signature constraints;
wenzelm
parents: 30939
diff changeset
   157
      Scan.lift (Scan.optional (Args.$$$ "(" |-- Scan.repeat1 parse_flag --| Args.$$$ ")") [] >>
2c20bcd70fbe proper signature constraints;
wenzelm
parents: 30939
diff changeset
   158
        curry (Library.foldl op |>) true) >>
2c20bcd70fbe proper signature constraints;
wenzelm
parents: 30939
diff changeset
   159
      (fn q => fn ctxt => SIMPLE_METHOD' (mir_tac ctxt q))
2c20bcd70fbe proper signature constraints;
wenzelm
parents: 30939
diff changeset
   160
    end
2c20bcd70fbe proper signature constraints;
wenzelm
parents: 30939
diff changeset
   161
    "decision procedure for MIR arithmetic";
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
   162
23590
ad95084a5c63 renamed ObjectLogic.atomize_tac to ObjectLogic.atomize_prems_tac;
wenzelm
parents: 23381
diff changeset
   163
end