src/HOL/Decision_Procs/mir_tac.ML
author blanchet
Wed, 04 Mar 2009 11:05:29 +0100
changeset 30242 aea5d7fa7ef5
parent 30240 5b25fee0362c
parent 30208 79136ce06bdb
child 30433 57c68b3af2ea
permissions -rw-r--r--
Merge.
     1 (*  Title:      HOL/Reflection/mir_tac.ML
     2     Author:     Amine Chaieb, TU Muenchen
     3 *)
     4 
     5 structure Mir_Tac =
     6 struct
     7 
     8 val trace = ref false;
     9 fun trace_msg s = if !trace then tracing s else ();
    10 
    11 val mir_ss = 
    12 let val ths = map thm ["real_of_int_inject", "real_of_int_less_iff", "real_of_int_le_iff"]
    13 in @{simpset} delsimps ths addsimps (map (fn th => th RS sym) ths)
    14 end;
    15 
    16 val nT = HOLogic.natT;
    17   val nat_arith = map thm ["add_nat_number_of", "diff_nat_number_of", 
    18                        "mult_nat_number_of", "eq_nat_number_of", "less_nat_number_of"];
    19 
    20   val comp_arith = (map thm ["Let_def", "if_False", "if_True", "add_0", 
    21                  "add_Suc", "add_number_of_left", "mult_number_of_left", 
    22                  "Suc_eq_add_numeral_1"])@
    23                  (map (fn s => thm s RS sym) ["numeral_1_eq_1", "numeral_0_eq_0"])
    24                  @ @{thms arith_simps} @ nat_arith @ @{thms rel_simps} 
    25   val ths = [@{thm "mult_numeral_1"}, @{thm "mult_numeral_1_right"}, 
    26              @{thm "real_of_nat_number_of"},
    27              @{thm "real_of_nat_Suc"}, @{thm "real_of_nat_one"}, @{thm "real_of_one"},
    28              @{thm "real_of_int_zero"}, @{thm "real_of_nat_zero"},
    29              @{thm "Ring_and_Field.divide_zero"}, 
    30              @{thm "divide_divide_eq_left"}, @{thm "times_divide_eq_right"}, 
    31              @{thm "times_divide_eq_left"}, @{thm "divide_divide_eq_right"},
    32              @{thm "diff_def"}, @{thm "minus_divide_left"}]
    33 val comp_ths = ths @ comp_arith @ simp_thms 
    34 
    35 
    36 val zdvd_int = @{thm "zdvd_int"};
    37 val zdiff_int_split = @{thm "zdiff_int_split"};
    38 val all_nat = @{thm "all_nat"};
    39 val ex_nat = @{thm "ex_nat"};
    40 val number_of1 = @{thm "number_of1"};
    41 val number_of2 = @{thm "number_of2"};
    42 val split_zdiv = @{thm "split_zdiv"};
    43 val split_zmod = @{thm "split_zmod"};
    44 val mod_div_equality' = @{thm "mod_div_equality'"};
    45 val split_div' = @{thm "split_div'"};
    46 val Suc_plus1 = @{thm "Suc_plus1"};
    47 val imp_le_cong = @{thm "imp_le_cong"};
    48 val conj_le_cong = @{thm "conj_le_cong"};
    49 val mod_add_eq = @{thm "mod_add_eq"} RS sym;
    50 val mod_add_left_eq = @{thm "mod_add_left_eq"} RS sym;
    51 val mod_add_right_eq = @{thm "mod_add_right_eq"} RS sym;
    52 val nat_div_add_eq = @{thm "div_add1_eq"} RS sym;
    53 val int_div_add_eq = @{thm "zdiv_zadd1_eq"} RS sym;
    54 val ZDIVISION_BY_ZERO_MOD = @{thm "DIVISION_BY_ZERO"} RS conjunct2;
    55 val ZDIVISION_BY_ZERO_DIV = @{thm "DIVISION_BY_ZERO"} RS conjunct1;
    56 
    57 fun prepare_for_mir thy q fm = 
    58   let
    59     val ps = Logic.strip_params fm
    60     val hs = map HOLogic.dest_Trueprop (Logic.strip_assums_hyp fm)
    61     val c = HOLogic.dest_Trueprop (Logic.strip_assums_concl fm)
    62     fun mk_all ((s, T), (P,n)) =
    63       if 0 mem loose_bnos P then
    64         (HOLogic.all_const T $ Abs (s, T, P), n)
    65       else (incr_boundvars ~1 P, n-1)
    66     fun mk_all2 (v, t) = HOLogic.all_const (fastype_of v) $ lambda v t;
    67       val rhs = hs
    68 (*    val (rhs,irhs) = List.partition (relevant (rev ps)) hs *)
    69     val np = length ps
    70     val (fm',np) =  foldr (fn ((x, T), (fm,n)) => mk_all ((x, T), (fm,n)))
    71       (foldr HOLogic.mk_imp c rhs, np) ps
    72     val (vs, _) = List.partition (fn t => q orelse (type_of t) = nT)
    73       (OldTerm.term_frees fm' @ OldTerm.term_vars fm');
    74     val fm2 = foldr mk_all2 fm' vs
    75   in (fm2, np + length vs, length rhs) end;
    76 
    77 (*Object quantifier to meta --*)
    78 fun spec_step n th = if (n=0) then th else (spec_step (n-1) th) RS spec ;
    79 
    80 (* object implication to meta---*)
    81 fun mp_step n th = if (n=0) then th else (mp_step (n-1) th) RS mp;
    82 
    83 
    84 fun mir_tac ctxt q i = 
    85     (ObjectLogic.atomize_prems_tac i)
    86         THEN (simp_tac (HOL_basic_ss addsimps [@{thm "abs_ge_zero"}] addsimps simp_thms) i)
    87         THEN (REPEAT_DETERM (split_tac [@{thm "split_min"}, @{thm "split_max"},@{thm "abs_split"}] i))
    88         THEN (fn st =>
    89   let
    90     val g = List.nth (prems_of st, i - 1)
    91     val thy = ProofContext.theory_of ctxt
    92     (* Transform the term*)
    93     val (t,np,nh) = prepare_for_mir thy q g
    94     (* Some simpsets for dealing with mod div abs and nat*)
    95     val mod_div_simpset = HOL_basic_ss 
    96                         addsimps [refl, mod_add_eq, 
    97                                   @{thm "mod_self"}, @{thm "zmod_self"},
    98                                   @{thm "zdiv_zero"},@{thm "zmod_zero"},@{thm "div_0"}, @{thm "mod_0"},
    99                                   @{thm "div_by_1"}, @{thm "mod_by_1"}, @{thm "div_1"}, @{thm "mod_1"},
   100                                   @{thm "Suc_plus1"}]
   101                         addsimps @{thms add_ac}
   102                         addsimprocs [cancel_div_mod_proc]
   103     val simpset0 = HOL_basic_ss
   104       addsimps [mod_div_equality', Suc_plus1]
   105       addsimps comp_ths
   106       addsplits [@{thm "split_zdiv"}, @{thm "split_zmod"}, @{thm "split_div'"}, @{thm "split_min"}, @{thm "split_max"}]
   107     (* Simp rules for changing (n::int) to int n *)
   108     val simpset1 = HOL_basic_ss
   109       addsimps [@{thm "nat_number_of_def"}, @{thm "zdvd_int"}] @ map (fn r => r RS sym)
   110         [@{thm "int_int_eq"}, @{thm "zle_int"}, @{thm "zless_int"}, @{thm "zadd_int"}, 
   111          @{thm "zmult_int"}]
   112       addsplits [@{thm "zdiff_int_split"}]
   113     (*simp rules for elimination of int n*)
   114 
   115     val simpset2 = HOL_basic_ss
   116       addsimps [@{thm "nat_0_le"}, @{thm "all_nat"}, @{thm "ex_nat"}, @{thm "number_of1"}, 
   117                 @{thm "number_of2"}, @{thm "int_0"}, @{thm "int_1"}]
   118       addcongs [@{thm "conj_le_cong"}, @{thm "imp_le_cong"}]
   119     (* simp rules for elimination of abs *)
   120     val ct = cterm_of thy (HOLogic.mk_Trueprop t)
   121     (* Theorem for the nat --> int transformation *)
   122     val pre_thm = Seq.hd (EVERY
   123       [simp_tac mod_div_simpset 1, simp_tac simpset0 1,
   124        TRY (simp_tac simpset1 1), TRY (simp_tac simpset2 1), TRY (simp_tac mir_ss 1)]
   125       (trivial ct))
   126     fun assm_tac i = REPEAT_DETERM_N nh (assume_tac i)
   127     (* The result of the quantifier elimination *)
   128     val (th, tac) = case (prop_of pre_thm) of
   129         Const ("==>", _) $ (Const ("Trueprop", _) $ t1) $ _ =>
   130     let val pth =
   131           (* If quick_and_dirty then run without proof generation as oracle*)
   132              if !quick_and_dirty
   133              then mirfr_oracle (false, cterm_of thy (Pattern.eta_long [] t1))
   134              else mirfr_oracle (true, cterm_of thy (Pattern.eta_long [] t1))
   135     in 
   136           (trace_msg ("calling procedure with term:\n" ^
   137              Syntax.string_of_term ctxt t1);
   138            ((pth RS iffD2) RS pre_thm,
   139             assm_tac (i + 1) THEN (if q then I else TRY) (rtac TrueI i)))
   140     end
   141       | _ => (pre_thm, assm_tac i)
   142   in (rtac (((mp_step nh) o (spec_step np)) th) i 
   143       THEN tac) st
   144   end handle Subscript => no_tac st);
   145 
   146 fun mir_args meth =
   147  let val parse_flag = 
   148          Args.$$$ "no_quantify" >> (K (K false));
   149  in
   150    Method.simple_args 
   151   (Scan.optional (Args.$$$ "(" |-- Scan.repeat1 parse_flag --| Args.$$$ ")") [] >>
   152     curry (Library.foldl op |>) true)
   153     (fn q => fn ctxt => meth ctxt q 1)
   154   end;
   155 
   156 fun mir_method ctxt q i = Method.METHOD (fn facts =>
   157   Method.insert_tac facts 1 THEN mir_tac ctxt q i);
   158 
   159 val setup =
   160   Method.add_method ("mir",
   161      mir_args mir_method,
   162      "decision procedure for MIR arithmetic");
   163 
   164 
   165 end