| author | kleing | 
| Wed, 26 Mar 2008 12:12:52 +0100 | |
| changeset 26402 | 441ddf3b8f02 | 
| parent 26075 | 815f3ccc0b45 | 
| child 26939 | 1035c89b4c02 | 
| permissions | -rw-r--r-- | 
| 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 | 1 | structure LinrTac = | 
| 
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 | 2 | 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 | 3 | |
| 
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 | 4 | val trace = ref false; | 
| 
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 | 5 | 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 | 6 | |
| 26075 
815f3ccc0b45
added lemma lists {normalize,succ,pred,minus,add,mult}_bin_simps
 huffman parents: 
23590diff
changeset | 7 | val ferrack_ss = let val ths = [@{thm real_of_int_inject}, @{thm real_of_int_less_iff}, 
 | 
| 
815f3ccc0b45
added lemma lists {normalize,succ,pred,minus,add,mult}_bin_simps
 huffman parents: 
23590diff
changeset | 8 | 				@{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 | 9 | 	     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 | 10 | 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 | 11 | |
| 
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 | val nT = HOLogic.natT; | 
| 26075 
815f3ccc0b45
added lemma lists {normalize,succ,pred,minus,add,mult}_bin_simps
 huffman parents: 
23590diff
changeset | 13 | val binarith = | 
| 
815f3ccc0b45
added lemma lists {normalize,succ,pred,minus,add,mult}_bin_simps
 huffman parents: 
23590diff
changeset | 14 |   @{thms normalize_bin_simps} @ @{thms pred_bin_simps} @ @{thms succ_bin_simps} @
 | 
| 
815f3ccc0b45
added lemma lists {normalize,succ,pred,minus,add,mult}_bin_simps
 huffman parents: 
23590diff
changeset | 15 |   @{thms add_bin_simps} @ @{thms minus_bin_simps} @  @{thms mult_bin_simps};
 | 
| 23318 | 16 | val comp_arith = binarith @ 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 | 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 | |
| 26075 
815f3ccc0b45
added lemma lists {normalize,succ,pred,minus,add,mult}_bin_simps
 huffman parents: 
23590diff
changeset | 19 | val zdvd_int = @{thm zdvd_int};
 | 
| 
815f3ccc0b45
added lemma lists {normalize,succ,pred,minus,add,mult}_bin_simps
 huffman parents: 
23590diff
changeset | 20 | val zdiff_int_split = @{thm zdiff_int_split};
 | 
| 
815f3ccc0b45
added lemma lists {normalize,succ,pred,minus,add,mult}_bin_simps
 huffman parents: 
23590diff
changeset | 21 | val all_nat = @{thm all_nat};
 | 
| 
815f3ccc0b45
added lemma lists {normalize,succ,pred,minus,add,mult}_bin_simps
 huffman parents: 
23590diff
changeset | 22 | val ex_nat = @{thm ex_nat};
 | 
| 
815f3ccc0b45
added lemma lists {normalize,succ,pred,minus,add,mult}_bin_simps
 huffman parents: 
23590diff
changeset | 23 | val number_of1 = @{thm number_of1};
 | 
| 
815f3ccc0b45
added lemma lists {normalize,succ,pred,minus,add,mult}_bin_simps
 huffman parents: 
23590diff
changeset | 24 | val number_of2 = @{thm number_of2};
 | 
| 
815f3ccc0b45
added lemma lists {normalize,succ,pred,minus,add,mult}_bin_simps
 huffman parents: 
23590diff
changeset | 25 | val split_zdiv = @{thm split_zdiv};
 | 
| 
815f3ccc0b45
added lemma lists {normalize,succ,pred,minus,add,mult}_bin_simps
 huffman parents: 
23590diff
changeset | 26 | val split_zmod = @{thm split_zmod};
 | 
| 
815f3ccc0b45
added lemma lists {normalize,succ,pred,minus,add,mult}_bin_simps
 huffman parents: 
23590diff
changeset | 27 | val mod_div_equality' = @{thm mod_div_equality'};
 | 
| 
815f3ccc0b45
added lemma lists {normalize,succ,pred,minus,add,mult}_bin_simps
 huffman parents: 
23590diff
changeset | 28 | val split_div' = @{thm split_div'};
 | 
| 
815f3ccc0b45
added lemma lists {normalize,succ,pred,minus,add,mult}_bin_simps
 huffman parents: 
23590diff
changeset | 29 | val Suc_plus1 = @{thm Suc_plus1};
 | 
| 
815f3ccc0b45
added lemma lists {normalize,succ,pred,minus,add,mult}_bin_simps
 huffman parents: 
23590diff
changeset | 30 | val imp_le_cong = @{thm imp_le_cong};
 | 
| 
815f3ccc0b45
added lemma lists {normalize,succ,pred,minus,add,mult}_bin_simps
 huffman parents: 
23590diff
changeset | 31 | val conj_le_cong = @{thm conj_le_cong};
 | 
| 23469 | 32 | val nat_mod_add_eq = @{thm mod_add1_eq} RS sym;
 | 
| 33 | val nat_mod_add_left_eq = @{thm mod_add_left_eq} RS sym;
 | |
| 34 | val nat_mod_add_right_eq = @{thm mod_add_right_eq} RS sym;
 | |
| 26075 
815f3ccc0b45
added lemma lists {normalize,succ,pred,minus,add,mult}_bin_simps
 huffman parents: 
23590diff
changeset | 35 | val int_mod_add_eq = @{thm zmod_zadd1_eq} RS sym;
 | 
| 
815f3ccc0b45
added lemma lists {normalize,succ,pred,minus,add,mult}_bin_simps
 huffman parents: 
23590diff
changeset | 36 | val int_mod_add_left_eq = @{thm zmod_zadd_left_eq} RS sym;
 | 
| 
815f3ccc0b45
added lemma lists {normalize,succ,pred,minus,add,mult}_bin_simps
 huffman parents: 
23590diff
changeset | 37 | val int_mod_add_right_eq = @{thm zmod_zadd_right_eq} RS sym;
 | 
| 
815f3ccc0b45
added lemma lists {normalize,succ,pred,minus,add,mult}_bin_simps
 huffman parents: 
23590diff
changeset | 38 | val nat_div_add_eq = @{thm div_add1_eq} RS sym;
 | 
| 
815f3ccc0b45
added lemma lists {normalize,succ,pred,minus,add,mult}_bin_simps
 huffman parents: 
23590diff
changeset | 39 | val int_div_add_eq = @{thm zdiv_zadd1_eq} RS sym;
 | 
| 
815f3ccc0b45
added lemma lists {normalize,succ,pred,minus,add,mult}_bin_simps
 huffman parents: 
23590diff
changeset | 40 | val ZDIVISION_BY_ZERO_MOD = @{thm DIVISION_BY_ZERO} RS conjunct2;
 | 
| 
815f3ccc0b45
added lemma lists {normalize,succ,pred,minus,add,mult}_bin_simps
 huffman parents: 
23590diff
changeset | 41 | val ZDIVISION_BY_ZERO_DIV = @{thm DIVISION_BY_ZERO} RS conjunct1;
 | 
| 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 | 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 | fun prepare_for_linr sg q 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 | 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 | 45 | 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 | 46 | 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 | 47 | 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 | 48 | fun mk_all ((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 | 49 | if 0 mem loose_bnos P then | 
| 
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 | (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 | 51 | 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 | 52 | 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 | 53 | 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 | 54 | (* 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 | 55 | val np = length ps | 
| 
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 | 56 | val (fm',np) = foldr (fn ((x, T), (fm,n)) => mk_all ((x, T), (fm,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 | 57 | (foldr HOLogic.mk_imp c rhs, np) ps | 
| 
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 (vs, _) = List.partition (fn t => q orelse (type_of t) = nT) | 
| 
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 | (term_frees fm' @ term_vars 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 | 60 | val fm2 = foldr mk_all2 fm' vs | 
| 
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 | 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 | 62 | |
| 
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 | (*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 | 64 | 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 | 65 | |
| 
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 | (* 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 | 67 | 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 | 68 | |
| 
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 | 69 | |
| 
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 | fun linr_tac ctxt q i = | 
| 23590 
ad95084a5c63
renamed ObjectLogic.atomize_tac to ObjectLogic.atomize_prems_tac;
 wenzelm parents: 
23469diff
changeset | 71 | (ObjectLogic.atomize_prems_tac i) | 
| 26075 
815f3ccc0b45
added lemma lists {normalize,succ,pred,minus,add,mult}_bin_simps
 huffman parents: 
23590diff
changeset | 72 | 	THEN (REPEAT_DETERM (split_tac [@{thm split_min}, @{thm split_max}, @{thm abs_split}] 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 | 73 | THEN (fn st => | 
| 
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 | 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 | 75 | val g = List.nth (prems_of st, i - 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 | 76 | val thy = ProofContext.theory_of ctxt | 
| 
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 | 77 | (* Transform the term*) | 
| 
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 (t,np,nh) = prepare_for_linr thy q g | 
| 
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 | 79 | (* 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 | 80 | val simpset0 = Simplifier.theory_context thy HOL_basic_ss addsimps comp_arith | 
| 
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 | val ct = cterm_of thy (HOLogic.mk_Trueprop 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 | 82 | (* 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 | 83 | 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 | 84 | [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 | 85 | TRY (simp_tac (Simplifier.theory_context thy ferrack_ss) 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 | 86 | (trivial ct)) | 
| 
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 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 | 88 | (* 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 | 89 | val (th, tac) = case (prop_of pre_thm) 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 | 90 |         Const ("==>", _) $ (Const ("Trueprop", _) $ t1) $ _ =>
 | 
| 
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 | let val pth = linr_oracle thy (Pattern.eta_long [] t1) | 
| 
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 | 92 | 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 | 93 |           (trace_msg ("calling procedure with term:\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 | 94 | Sign.string_of_term thy t1); | 
| 
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 | ((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 | 96 | 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 | 97 | 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 | 98 | | _ => (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 | 99 | 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 | 100 | THEN tac) st | 
| 
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 | end handle Subscript => no_tac st | ReflectedFerrack.LINR => no_tac st); | 
| 
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 | 102 | |
| 
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 | fun linr_meth src = | 
| 
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 | 104 | Method.syntax (Args.mode "no_quantify") src | 
| 
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 | #> (fn (q, ctxt) => Method.SIMPLE_METHOD' (linr_tac ctxt (not q))); | 
| 
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 | |
| 
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 | val setup = | 
| 
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 | 108 |   Method.add_method ("rferrack", linr_meth,
 | 
| 
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 | "decision procedure for linear real arithmetic"); | 
| 
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 | 110 | |
| 
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 | |
| 23469 | 112 | end |