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