| author | wenzelm | 
| Tue, 07 Nov 2006 11:46:49 +0100 | |
| changeset 21207 | cef082634be9 | 
| parent 21078 | 101aefd61aac | 
| child 21588 | cd0dc678a205 | 
| permissions | -rw-r--r-- | 
| 19640 
40ec89317425
added Ferrante and Rackoff Algorithm -- by Amine Chaieb;
 wenzelm parents: diff
changeset | 1 | (* | 
| 
40ec89317425
added Ferrante and Rackoff Algorithm -- by Amine Chaieb;
 wenzelm parents: diff
changeset | 2 | ID: $Id$ | 
| 
40ec89317425
added Ferrante and Rackoff Algorithm -- by Amine Chaieb;
 wenzelm parents: diff
changeset | 3 | Author: Amine Chaieb, TU Muenchen | 
| 
40ec89317425
added Ferrante and Rackoff Algorithm -- by Amine Chaieb;
 wenzelm parents: diff
changeset | 4 | |
| 
40ec89317425
added Ferrante and Rackoff Algorithm -- by Amine Chaieb;
 wenzelm parents: diff
changeset | 5 | Ferrante and Rackoff Algorithm. | 
| 
40ec89317425
added Ferrante and Rackoff Algorithm -- by Amine Chaieb;
 wenzelm parents: diff
changeset | 6 | *) | 
| 
40ec89317425
added Ferrante and Rackoff Algorithm -- by Amine Chaieb;
 wenzelm parents: diff
changeset | 7 | |
| 21078 | 8 | signature FERRANTE_RACKOFF = | 
| 19640 
40ec89317425
added Ferrante and Rackoff Algorithm -- by Amine Chaieb;
 wenzelm parents: diff
changeset | 9 | sig | 
| 21078 | 10 | val ferrack_tac: bool -> int -> tactic | 
| 11 | val setup: theory -> theory | |
| 12 | val trace: bool ref | |
| 13 | end; | |
| 14 | ||
| 15 | structure Ferrante_Rackoff : FERRANTE_RACKOFF = | |
| 19640 
40ec89317425
added Ferrante and Rackoff Algorithm -- by Amine Chaieb;
 wenzelm parents: diff
changeset | 16 | struct | 
| 
40ec89317425
added Ferrante and Rackoff Algorithm -- by Amine Chaieb;
 wenzelm parents: diff
changeset | 17 | |
| 
40ec89317425
added Ferrante and Rackoff Algorithm -- by Amine Chaieb;
 wenzelm parents: diff
changeset | 18 | val trace = ref false; | 
| 
40ec89317425
added Ferrante and Rackoff Algorithm -- by Amine Chaieb;
 wenzelm parents: diff
changeset | 19 | fun trace_msg s = if !trace then tracing s else (); | 
| 
40ec89317425
added Ferrante and Rackoff Algorithm -- by Amine Chaieb;
 wenzelm parents: diff
changeset | 20 | |
| 21078 | 21 | val binarith = map thm ["Pls_0_eq", "Min_1_eq", "pred_Pls", "pred_Min","pred_1","pred_0", | 
| 20485 | 22 | "succ_Pls", "succ_Min", "succ_1", "succ_0", | 
| 21078 | 23 | "add_Pls", "add_Min", "add_BIT_0", "add_BIT_10", "add_BIT_11", | 
| 24 | "minus_Pls", "minus_Min", "minus_1", "minus_0", | |
| 25 | "mult_Pls", "mult_Min", "mult_1", "mult_0", | |
| 20485 | 26 | "add_Pls_right", "add_Min_right"]; | 
| 21078 | 27 | |
| 28 | val intarithrel = | |
| 29 | map thm ["int_eq_number_of_eq", "int_neg_number_of_BIT", "int_le_number_of_eq", | |
| 30 | "int_iszero_number_of_0", "int_less_number_of_eq_neg"] | |
| 31 | @ map (fn s => thm s RS thm "lift_bool") ["int_iszero_number_of_Pls", | |
| 32 | "int_iszero_number_of_1", "int_neg_number_of_Min"] | |
| 33 | @ map (fn s => thm s RS thm "nlift_bool") ["int_nonzero_number_of_Min", | |
| 34 | "int_not_neg_number_of_Pls"]; | |
| 35 | ||
| 19640 
40ec89317425
added Ferrante and Rackoff Algorithm -- by Amine Chaieb;
 wenzelm parents: diff
changeset | 36 | val intarith = map thm ["int_number_of_add_sym", "int_number_of_minus_sym", | 
| 21078 | 37 | "int_number_of_diff_sym", "int_number_of_mult_sym"]; | 
| 38 | ||
| 19640 
40ec89317425
added Ferrante and Rackoff Algorithm -- by Amine Chaieb;
 wenzelm parents: diff
changeset | 39 | val natarith = map thm ["add_nat_number_of", "diff_nat_number_of", | 
| 21078 | 40 | "mult_nat_number_of", "eq_nat_number_of", "less_nat_number_of"]; | 
| 41 | ||
| 42 | val powerarith = | |
| 43 | map thm ["nat_number_of", "zpower_number_of_even", | |
| 44 | "zpower_Pls", "zpower_Min"] | |
| 45 | @ [Tactic.simplify true [thm "zero_eq_Numeral0_nring", thm "one_eq_Numeral1_nring"] | |
| 46 | (thm "zpower_number_of_odd")] | |
| 19640 
40ec89317425
added Ferrante and Rackoff Algorithm -- by Amine Chaieb;
 wenzelm parents: diff
changeset | 47 | |
| 
40ec89317425
added Ferrante and Rackoff Algorithm -- by Amine Chaieb;
 wenzelm parents: diff
changeset | 48 | val comp_arith = binarith @ intarith @ intarithrel @ natarith | 
| 21078 | 49 | @ powerarith @ [thm "not_false_eq_true", thm "not_true_eq_false"]; | 
| 19640 
40ec89317425
added Ferrante and Rackoff Algorithm -- by Amine Chaieb;
 wenzelm parents: diff
changeset | 50 | |
| 21078 | 51 | fun prepare_for_linr q fm = | 
| 19640 
40ec89317425
added Ferrante and Rackoff Algorithm -- by Amine Chaieb;
 wenzelm parents: diff
changeset | 52 | let | 
| 
40ec89317425
added Ferrante and Rackoff Algorithm -- by Amine Chaieb;
 wenzelm parents: diff
changeset | 53 | val ps = Logic.strip_params fm | 
| 
40ec89317425
added Ferrante and Rackoff Algorithm -- by Amine Chaieb;
 wenzelm parents: diff
changeset | 54 | val hs = map HOLogic.dest_Trueprop (Logic.strip_assums_hyp fm) | 
| 
40ec89317425
added Ferrante and Rackoff Algorithm -- by Amine Chaieb;
 wenzelm parents: diff
changeset | 55 | val c = HOLogic.dest_Trueprop (Logic.strip_assums_concl fm) | 
| 21078 | 56 | fun mk_all ((s, T), (P, n)) = | 
| 19640 
40ec89317425
added Ferrante and Rackoff Algorithm -- by Amine Chaieb;
 wenzelm parents: diff
changeset | 57 | if 0 mem loose_bnos P then | 
| 
40ec89317425
added Ferrante and Rackoff Algorithm -- by Amine Chaieb;
 wenzelm parents: diff
changeset | 58 | (HOLogic.all_const T $ Abs (s, T, P), n) | 
| 
40ec89317425
added Ferrante and Rackoff Algorithm -- by Amine Chaieb;
 wenzelm parents: diff
changeset | 59 | else (incr_boundvars ~1 P, n-1) | 
| 
40ec89317425
added Ferrante and Rackoff Algorithm -- by Amine Chaieb;
 wenzelm parents: diff
changeset | 60 | fun mk_all2 (v, t) = HOLogic.all_const (fastype_of v) $ lambda v t; | 
| 21078 | 61 | val rhs = hs; | 
| 62 | val np = length ps; | |
| 63 | val (fm', np) = Library.foldr mk_all (ps, (Library.foldr HOLogic.mk_imp (rhs, c), np)); | |
| 64 | val (vs, _) = List.partition (fn t => q orelse (type_of t) = HOLogic.natT) | |
| 19640 
40ec89317425
added Ferrante and Rackoff Algorithm -- by Amine Chaieb;
 wenzelm parents: diff
changeset | 65 | (term_frees fm' @ term_vars fm'); | 
| 21078 | 66 | val fm2 = Library.foldr mk_all2 (vs, fm'); | 
| 19640 
40ec89317425
added Ferrante and Rackoff Algorithm -- by Amine Chaieb;
 wenzelm parents: diff
changeset | 67 | in (fm2, np + length vs, length rhs) end; | 
| 
40ec89317425
added Ferrante and Rackoff Algorithm -- by Amine Chaieb;
 wenzelm parents: diff
changeset | 68 | |
| 21078 | 69 | fun spec_step n th = if n = 0 then th else spec_step (n - 1) th RS spec ; | 
| 70 | fun mp_step n th = if n = 0 then th else mp_step (n - 1) th RS mp; | |
| 19640 
40ec89317425
added Ferrante and Rackoff Algorithm -- by Amine Chaieb;
 wenzelm parents: diff
changeset | 71 | |
| 21078 | 72 | local | 
| 73 | val context_ss = simpset_of (the_context ()) | |
| 74 | in fun ferrack_tac q i = ObjectLogic.atomize_tac i | |
| 75 | THEN REPEAT_DETERM (split_tac [split_min, split_max,abs_split] i) | |
| 76 | THEN (fn st => | |
| 19640 
40ec89317425
added Ferrante and Rackoff Algorithm -- by Amine Chaieb;
 wenzelm parents: diff
changeset | 77 | let | 
| 21078 | 78 | val g = nth (prems_of st) (i - 1) | 
| 79 | val thy = sign_of_thm st | |
| 19640 
40ec89317425
added Ferrante and Rackoff Algorithm -- by Amine Chaieb;
 wenzelm parents: diff
changeset | 80 | (* Transform the term*) | 
| 21078 | 81 | val (t,np,nh) = prepare_for_linr q g | 
| 19640 
40ec89317425
added Ferrante and Rackoff Algorithm -- by Amine Chaieb;
 wenzelm parents: diff
changeset | 82 | (* Some simpsets for dealing with mod div abs and nat*) | 
| 
40ec89317425
added Ferrante and Rackoff Algorithm -- by Amine Chaieb;
 wenzelm parents: diff
changeset | 83 | val simpset0 = HOL_basic_ss addsimps comp_arith addsplits [split_min, split_max] | 
| 
40ec89317425
added Ferrante and Rackoff Algorithm -- by Amine Chaieb;
 wenzelm parents: diff
changeset | 84 | (* simp rules for elimination of abs *) | 
| 
40ec89317425
added Ferrante and Rackoff Algorithm -- by Amine Chaieb;
 wenzelm parents: diff
changeset | 85 | val simpset3 = HOL_basic_ss addsplits [abs_split] | 
| 21078 | 86 | val ct = cterm_of thy (HOLogic.mk_Trueprop t) | 
| 19640 
40ec89317425
added Ferrante and Rackoff Algorithm -- by Amine Chaieb;
 wenzelm parents: diff
changeset | 87 | (* Theorem for the nat --> int transformation *) | 
| 
40ec89317425
added Ferrante and Rackoff Algorithm -- by Amine Chaieb;
 wenzelm parents: diff
changeset | 88 | val pre_thm = Seq.hd (EVERY | 
| 19825 | 89 | [simp_tac simpset0 1, TRY (simp_tac context_ss 1)] | 
| 19640 
40ec89317425
added Ferrante and Rackoff Algorithm -- by Amine Chaieb;
 wenzelm parents: diff
changeset | 90 | (trivial ct)) | 
| 
40ec89317425
added Ferrante and Rackoff Algorithm -- by Amine Chaieb;
 wenzelm parents: diff
changeset | 91 | fun assm_tac i = REPEAT_DETERM_N nh (assume_tac i) | 
| 
40ec89317425
added Ferrante and Rackoff Algorithm -- by Amine Chaieb;
 wenzelm parents: diff
changeset | 92 | (* The result of the quantifier elimination *) | 
| 
40ec89317425
added Ferrante and Rackoff Algorithm -- by Amine Chaieb;
 wenzelm parents: diff
changeset | 93 | val (th, tac) = case (prop_of pre_thm) of | 
| 
40ec89317425
added Ferrante and Rackoff Algorithm -- by Amine Chaieb;
 wenzelm parents: diff
changeset | 94 |         Const ("==>", _) $ (Const ("Trueprop", _) $ t1) $ _ =>
 | 
| 21078 | 95 | let val pth = Ferrante_Rackoff_Proof.qelim (cterm_of thy (Pattern.eta_long [] t1)) | 
| 19640 
40ec89317425
added Ferrante and Rackoff Algorithm -- by Amine Chaieb;
 wenzelm parents: diff
changeset | 96 | in | 
| 
40ec89317425
added Ferrante and Rackoff Algorithm -- by Amine Chaieb;
 wenzelm parents: diff
changeset | 97 |           (trace_msg ("calling procedure with term:\n" ^
 | 
| 21078 | 98 | Sign.string_of_term thy t1); | 
| 19640 
40ec89317425
added Ferrante and Rackoff Algorithm -- by Amine Chaieb;
 wenzelm parents: diff
changeset | 99 | ((pth RS iffD2) RS pre_thm, | 
| 
40ec89317425
added Ferrante and Rackoff Algorithm -- by Amine Chaieb;
 wenzelm parents: diff
changeset | 100 | assm_tac (i + 1) THEN (if q then I else TRY) (rtac TrueI i))) | 
| 
40ec89317425
added Ferrante and Rackoff Algorithm -- by Amine Chaieb;
 wenzelm parents: diff
changeset | 101 | end | 
| 
40ec89317425
added Ferrante and Rackoff Algorithm -- by Amine Chaieb;
 wenzelm parents: diff
changeset | 102 | | _ => (pre_thm, assm_tac i) | 
| 
40ec89317425
added Ferrante and Rackoff Algorithm -- by Amine Chaieb;
 wenzelm parents: diff
changeset | 103 | in (rtac (((mp_step nh) o (spec_step np)) th) i | 
| 
40ec89317425
added Ferrante and Rackoff Algorithm -- by Amine Chaieb;
 wenzelm parents: diff
changeset | 104 | THEN tac) st | 
| 
40ec89317425
added Ferrante and Rackoff Algorithm -- by Amine Chaieb;
 wenzelm parents: diff
changeset | 105 | end handle Subscript => no_tac st | Ferrante_Rackoff_Proof.FAILURE _ => no_tac st); | 
| 21078 | 106 | end; (*local*) | 
| 19640 
40ec89317425
added Ferrante and Rackoff Algorithm -- by Amine Chaieb;
 wenzelm parents: diff
changeset | 107 | |
| 
40ec89317425
added Ferrante and Rackoff Algorithm -- by Amine Chaieb;
 wenzelm parents: diff
changeset | 108 | fun ferrack_args meth = | 
| 21078 | 109 | let | 
| 110 | val parse_flag = Args.$$$ "no_quantify" >> (K (K false)); | |
| 19640 
40ec89317425
added Ferrante and Rackoff Algorithm -- by Amine Chaieb;
 wenzelm parents: diff
changeset | 111 | in | 
| 
40ec89317425
added Ferrante and Rackoff Algorithm -- by Amine Chaieb;
 wenzelm parents: diff
changeset | 112 | Method.simple_args | 
| 
40ec89317425
added Ferrante and Rackoff Algorithm -- by Amine Chaieb;
 wenzelm parents: diff
changeset | 113 |   (Scan.optional (Args.$$$ "(" |-- Scan.repeat1 parse_flag --| Args.$$$ ")") [] >>
 | 
| 
40ec89317425
added Ferrante and Rackoff Algorithm -- by Amine Chaieb;
 wenzelm parents: diff
changeset | 114 | curry (Library.foldl op |>) true) | 
| 
40ec89317425
added Ferrante and Rackoff Algorithm -- by Amine Chaieb;
 wenzelm parents: diff
changeset | 115 | (fn q => fn _ => meth q 1) | 
| 
40ec89317425
added Ferrante and Rackoff Algorithm -- by Amine Chaieb;
 wenzelm parents: diff
changeset | 116 | end; | 
| 
40ec89317425
added Ferrante and Rackoff Algorithm -- by Amine Chaieb;
 wenzelm parents: diff
changeset | 117 | |
| 
40ec89317425
added Ferrante and Rackoff Algorithm -- by Amine Chaieb;
 wenzelm parents: diff
changeset | 118 | val setup = | 
| 
40ec89317425
added Ferrante and Rackoff Algorithm -- by Amine Chaieb;
 wenzelm parents: diff
changeset | 119 |   Method.add_method ("ferrack",
 | 
| 
40ec89317425
added Ferrante and Rackoff Algorithm -- by Amine Chaieb;
 wenzelm parents: diff
changeset | 120 | ferrack_args (Method.SIMPLE_METHOD oo ferrack_tac), | 
| 
40ec89317425
added Ferrante and Rackoff Algorithm -- by Amine Chaieb;
 wenzelm parents: diff
changeset | 121 | "LCF-proof-producing decision procedure for linear real arithmetic"); | 
| 
40ec89317425
added Ferrante and Rackoff Algorithm -- by Amine Chaieb;
 wenzelm parents: diff
changeset | 122 | |
| 
40ec89317425
added Ferrante and Rackoff Algorithm -- by Amine Chaieb;
 wenzelm parents: diff
changeset | 123 | end |