| author | wenzelm | 
| Mon, 22 May 2017 19:42:52 +0200 | |
| changeset 65904 | 8411f1a2272c | 
| parent 64927 | a5a09855e424 | 
| child 67091 | 1393c2340eec | 
| permissions | -rw-r--r-- | 
| 63992 | 1 | (* Title: HOL/Tools/Argo/argo_real.ML | 
| 63960 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 2 | Author: Sascha Boehme | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 3 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 4 | Extension of the Argo tactic for the reals. | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 5 | *) | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 6 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 7 | structure Argo_Real: sig end = | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 8 | struct | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 9 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 10 | (* translating input terms *) | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 11 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 12 | fun trans_type _ @{typ Real.real} tcx = SOME (Argo_Expr.Real, tcx)
 | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 13 | | trans_type _ _ _ = NONE | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 14 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 15 | fun trans_term f (@{const Groups.uminus_class.uminus (real)} $ t) tcx =
 | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 16 | tcx |> f t |>> Argo_Expr.mk_neg |> SOME | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 17 |   | trans_term f (@{const Groups.plus_class.plus (real)} $ t1 $ t2) tcx =
 | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 18 | tcx |> f t1 ||>> f t2 |>> uncurry Argo_Expr.mk_add2 |> SOME | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 19 |   | trans_term f (@{const Groups.minus_class.minus (real)} $ t1 $ t2) tcx =
 | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 20 | tcx |> f t1 ||>> f t2 |>> uncurry Argo_Expr.mk_sub |> SOME | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 21 |   | trans_term f (@{const Groups.times_class.times (real)} $ t1 $ t2) tcx =
 | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 22 | tcx |> f t1 ||>> f t2 |>> uncurry Argo_Expr.mk_mul |> SOME | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 23 |   | trans_term f (@{const Rings.divide_class.divide (real)} $ t1 $ t2) tcx =
 | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 24 | tcx |> f t1 ||>> f t2 |>> uncurry Argo_Expr.mk_div |> SOME | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 25 |   | trans_term f (@{const Orderings.ord_class.min (real)} $ t1 $ t2) tcx =
 | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 26 | tcx |> f t1 ||>> f t2 |>> uncurry Argo_Expr.mk_min |> SOME | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 27 |   | trans_term f (@{const Orderings.ord_class.max (real)} $ t1 $ t2) tcx =
 | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 28 | tcx |> f t1 ||>> f t2 |>> uncurry Argo_Expr.mk_max |> SOME | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 29 |   | trans_term f (@{const Groups.abs_class.abs (real)} $ t) tcx =
 | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 30 | tcx |> f t |>> Argo_Expr.mk_abs |> SOME | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 31 |   | trans_term f (@{const Orderings.ord_class.less (real)} $ t1 $ t2) tcx =
 | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 32 | tcx |> f t1 ||>> f t2 |>> uncurry Argo_Expr.mk_lt |> SOME | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 33 |   | trans_term f (@{const Orderings.ord_class.less_eq (real)} $ t1 $ t2) tcx =
 | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 34 | tcx |> f t1 ||>> f t2 |>> uncurry Argo_Expr.mk_le |> SOME | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 35 | | trans_term _ t tcx = | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 36 | (case try HOLogic.dest_number t of | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 37 |         SOME (@{typ Real.real}, n) => SOME (Argo_Expr.mk_num (Rat.of_int n), tcx)
 | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 38 | | _ => NONE) | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 39 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 40 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 41 | (* reverse translation *) | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 42 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 43 | fun mk_plus t1 t2 = @{const Groups.plus_class.plus (real)} $ t1 $ t2
 | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 44 | fun mk_sum ts = uncurry (fold_rev mk_plus) (split_last ts) | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 45 | fun mk_times t1 t2 = @{const Groups.times_class.times (real)} $ t1 $ t2
 | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 46 | fun mk_divide t1 t2 = @{const Rings.divide_class.divide (real)} $ t1 $ t2
 | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 47 | fun mk_le t1 t2 = @{const Orderings.ord_class.less_eq (real)} $ t1 $ t2
 | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 48 | fun mk_lt t1 t2 = @{const Orderings.ord_class.less (real)} $ t1 $ t2
 | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 49 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 50 | fun mk_real_num i = HOLogic.mk_number @{typ Real.real} i
 | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 51 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 52 | fun mk_number n = | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 53 | let val (p, q) = Rat.dest n | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 54 | in if q = 1 then mk_real_num p else mk_divide (mk_real_num p) (mk_real_num q) end | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 55 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 56 | fun term_of _ (Argo_Expr.E (Argo_Expr.Num n, _)) = SOME (mk_number n) | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 57 | | term_of f (Argo_Expr.E (Argo_Expr.Neg, [e])) = | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 58 |       SOME (@{const Groups.uminus_class.uminus (real)} $ f e)
 | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 59 | | term_of f (Argo_Expr.E (Argo_Expr.Add, es)) = SOME (mk_sum (map f es)) | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 60 | | term_of f (Argo_Expr.E (Argo_Expr.Sub, [e1, e2])) = | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 61 |       SOME (@{const Groups.minus_class.minus (real)} $ f e1 $ f e2)
 | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 62 | | term_of f (Argo_Expr.E (Argo_Expr.Mul, [e1, e2])) = SOME (mk_times (f e1) (f e2)) | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 63 | | term_of f (Argo_Expr.E (Argo_Expr.Div, [e1, e2])) = SOME (mk_divide (f e1) (f e2)) | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 64 | | term_of f (Argo_Expr.E (Argo_Expr.Min, [e1, e2])) = | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 65 |       SOME (@{const Orderings.ord_class.min (real)} $ f e1 $ f e2)
 | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 66 | | term_of f (Argo_Expr.E (Argo_Expr.Max, [e1, e2])) = | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 67 |       SOME (@{const Orderings.ord_class.max (real)} $ f e1 $ f e2)
 | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 68 |   | term_of f (Argo_Expr.E (Argo_Expr.Abs, [e])) = SOME (@{const Groups.abs_class.abs (real)} $ f e)
 | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 69 | | term_of f (Argo_Expr.E (Argo_Expr.Le, [e1, e2])) = SOME (mk_le (f e1) (f e2)) | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 70 | | term_of f (Argo_Expr.E (Argo_Expr.Lt, [e1, e2])) = SOME (mk_lt (f e1) (f e2)) | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 71 | | term_of _ _ = NONE | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 72 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 73 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 74 | (* proof replay for rewrite steps *) | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 75 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 76 | fun mk_rewr thm = thm RS @{thm eq_reflection}
 | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 77 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 78 | fun by_simp ctxt t = | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 79 |   let fun prove {context, ...} = HEADGOAL (Simplifier.simp_tac context)
 | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 80 | in Goal.prove ctxt [] [] (HOLogic.mk_Trueprop t) prove end | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 81 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 82 | fun prove_num_pred ctxt n = | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 83 | by_simp ctxt (uncurry mk_lt (apply2 mk_number (if @0 < n then (@0, n) else (n, @0)))) | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 84 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 85 | fun simp_conv ctxt t = Conv.rewr_conv (mk_rewr (by_simp ctxt t)) | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 86 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 87 | fun nums_conv mk f ctxt n m = | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 88 | simp_conv ctxt (HOLogic.mk_eq (mk (mk_number n) (mk_number m), mk_number (f (n, m)))) | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 89 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 90 | val add_nums_conv = nums_conv mk_plus (op +) | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 91 | val mul_nums_conv = nums_conv mk_times (op *) | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 92 | val div_nums_conv = nums_conv mk_divide (op /) | 
| 64927 
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
 boehmes parents: 
63992diff
changeset | 93 | fun inv_num_conv ctxt = nums_conv mk_divide (fn (_, n) => Rat.inv n) ctxt @1 | 
| 63960 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 94 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 95 | fun cmp_nums_conv ctxt b ct = | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 96 |   let val t = if b then @{const HOL.True} else @{const HOL.False}
 | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 97 | in simp_conv ctxt (HOLogic.mk_eq (Thm.term_of ct, t)) ct end | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 98 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 99 | local | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 100 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 101 | fun is_add2 (@{const Groups.plus_class.plus (real)} $ _ $ _) = true
 | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 102 | | is_add2 _ = false | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 103 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 104 | fun is_add3 (@{const Groups.plus_class.plus (real)} $ _ $ t) = is_add2 t
 | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 105 | | is_add3 _ = false | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 106 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 107 | val flatten_thm = mk_rewr @{lemma "(a::real) + b + c = a + (b + c)" by simp}
 | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 108 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 109 | fun flatten_conv ct = | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 110 | if is_add2 (Thm.term_of ct) then Argo_Tactic.flatten_conv flatten_conv flatten_thm ct | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 111 | else Conv.all_conv ct | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 112 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 113 | val swap_conv = Conv.rewrs_conv (map mk_rewr @{lemma 
 | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 114 | "(a::real) + (b + c) = b + (a + c)" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 115 | "(a::real) + b = b + a" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 116 | by simp_all}) | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 117 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 118 | val assoc_conv = Conv.rewr_conv (mk_rewr @{lemma "(a::real) + (b + c) = (a + b) + c" by simp})
 | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 119 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 120 | val norm_monom_thm = mk_rewr @{lemma "1 * (a::real) = a" by simp}
 | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 121 | fun norm_monom_conv n = if n = @1 then Conv.rewr_conv norm_monom_thm else Conv.all_conv | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 122 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 123 | val add2_thms = map mk_rewr @{lemma
 | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 124 | "n * (a::real) + m * a = (n + m) * a" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 125 | "n * (a::real) + a = (n + 1) * a" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 126 | "(a::real) + m * a = (1 + m) * a" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 127 | "(a::real) + a = (1 + 1) * a" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 128 | by algebra+} | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 129 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 130 | val add3_thms = map mk_rewr @{lemma
 | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 131 | "n * (a::real) + (m * a + b) = (n + m) * a + b" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 132 | "n * (a::real) + (a + b) = (n + 1) * a + b" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 133 | "(a::real) + (m * a + b) = (1 + m) * a + b" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 134 | "(a::real) + (a + b) = (1 + 1) * a + b" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 135 | by algebra+} | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 136 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 137 | fun choose_conv cv2 cv3 ct = if is_add3 (Thm.term_of ct) then cv3 ct else cv2 ct | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 138 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 139 | fun join_num_conv ctxt n m = | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 140 | let val conv = add_nums_conv ctxt n m | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 141 | in choose_conv conv (assoc_conv then_conv Conv.arg1_conv conv) end | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 142 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 143 | fun join_monom_conv ctxt n m = | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 144 | let | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 145 | val conv = Conv.arg1_conv (add_nums_conv ctxt n m) then_conv norm_monom_conv (n + m) | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 146 | fun seq_conv thms cv = Conv.rewrs_conv thms then_conv cv | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 147 | in choose_conv (seq_conv add2_thms conv) (seq_conv add3_thms (Conv.arg1_conv conv)) end | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 148 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 149 | fun join_conv NONE = join_num_conv | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 150 | | join_conv (SOME _) = join_monom_conv | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 151 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 152 | fun bubble_down_conv _ _ [] ct = Conv.no_conv ct | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 153 | | bubble_down_conv _ _ [_] ct = Conv.all_conv ct | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 154 | | bubble_down_conv ctxt i ((m1 as (n1, i1)) :: (m2 as (n2, i2)) :: ms) ct = | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 155 | if i1 = i then | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 156 | if i2 = i then | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 157 | (join_conv i ctxt n1 n2 then_conv bubble_down_conv ctxt i ((n1 + n2, i) :: ms)) ct | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 158 | else (swap_conv then_conv Conv.arg_conv (bubble_down_conv ctxt i (m1 :: ms))) ct | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 159 | else Conv.arg_conv (bubble_down_conv ctxt i (m2 :: ms)) ct | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 160 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 161 | fun drop_var i ms = filter_out (fn (_, i') => i' = i) ms | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 162 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 163 | fun permute_conv _ [] [] ct = Conv.all_conv ct | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 164 | | permute_conv ctxt (ms as ((_, i) :: _)) [] ct = | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 165 | (bubble_down_conv ctxt i ms then_conv permute_conv ctxt (drop_var i ms) []) ct | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 166 | | permute_conv ctxt ms1 ms2 ct = | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 167 | let val (ms2', (_, i)) = split_last ms2 | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 168 | in (bubble_down_conv ctxt i ms1 then_conv permute_conv ctxt (drop_var i ms1) ms2') ct end | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 169 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 170 | val no_monom_conv = Conv.rewr_conv (mk_rewr @{lemma "0 * (a::real) = 0" by simp})
 | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 171 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 172 | val norm_sum_conv = Conv.rewrs_conv (map mk_rewr @{lemma
 | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 173 | "0 * (a::real) + b = b" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 174 | "(a::real) + 0 * b = a" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 175 | "0 + (a::real) = a" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 176 | "(a::real) + 0 = a" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 177 | by simp_all}) | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 178 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 179 | fun drop0_conv ct = | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 180 | if is_add2 (Thm.term_of ct) then | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 181 | ((norm_sum_conv then_conv drop0_conv) else_conv Conv.arg_conv drop0_conv) ct | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 182 | else Conv.try_conv no_monom_conv ct | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 183 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 184 | fun full_add_conv ctxt ms1 ms2 = | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 185 | if eq_list (op =) (ms1, ms2) then flatten_conv | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 186 | else flatten_conv then_conv permute_conv ctxt ms1 ms2 then_conv drop0_conv | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 187 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 188 | in | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 189 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 190 | fun add_conv ctxt (ms1, ms2 as [(n, NONE)]) = | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 191 | if n = @0 then full_add_conv ctxt ms1 [] else full_add_conv ctxt ms1 ms2 | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 192 | | add_conv ctxt (ms1, ms2) = full_add_conv ctxt ms1 ms2 | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 193 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 194 | end | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 195 | |
| 64927 
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
 boehmes parents: 
63992diff
changeset | 196 | val mul_suml_thm = mk_rewr @{lemma "((x::real) + y) * z = x * z + y * z" by (rule ring_distribs)}
 | 
| 
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
 boehmes parents: 
63992diff
changeset | 197 | val mul_sumr_thm = mk_rewr @{lemma "(x::real) * (y + z) = x * y + x * z" by (rule ring_distribs)}
 | 
| 
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
 boehmes parents: 
63992diff
changeset | 198 | fun mul_sum_conv thm ct = | 
| 
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
 boehmes parents: 
63992diff
changeset | 199 | Conv.try_conv (Conv.rewr_conv thm then_conv Conv.binop_conv (mul_sum_conv thm)) ct | 
| 
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
 boehmes parents: 
63992diff
changeset | 200 | |
| 
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
 boehmes parents: 
63992diff
changeset | 201 | val div_sum_thm = mk_rewr @{lemma "((x::real) + y) / z = x / z + y / z"
 | 
| 
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
 boehmes parents: 
63992diff
changeset | 202 | by (rule add_divide_distrib)} | 
| 
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
 boehmes parents: 
63992diff
changeset | 203 | fun div_sum_conv ct = | 
| 
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
 boehmes parents: 
63992diff
changeset | 204 | Conv.try_conv (Conv.rewr_conv div_sum_thm then_conv Conv.binop_conv div_sum_conv) ct | 
| 63960 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 205 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 206 | fun var_of_add_cmp (_ $ _ $ (_ $ _ $ (_ $ _ $ Var v))) = v | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 207 |   | var_of_add_cmp t = raise TERM ("var_of_add_cmp", [t])
 | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 208 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 209 | fun add_cmp_conv ctxt n thm = | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 210 | let val v = var_of_add_cmp (Thm.prop_of thm) | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 211 | in Conv.rewr_conv (Thm.instantiate ([], [(v, Thm.cterm_of ctxt (mk_number n))]) thm) end | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 212 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 213 | fun mul_cmp_conv ctxt n pos_thm neg_thm = | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 214 | let val thm = if @0 < n then pos_thm else neg_thm | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 215 | in Conv.rewr_conv (prove_num_pred ctxt n RS thm) end | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 216 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 217 | val neg_thm = mk_rewr @{lemma "-(x::real) = -1 * x" by simp}
 | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 218 | val sub_thm = mk_rewr @{lemma "(x::real) - y = x + -1 * y" by simp}
 | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 219 | val mul_zero_thm = mk_rewr @{lemma "0 * (x::real) = 0" by (rule mult_zero_left)}
 | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 220 | val mul_one_thm = mk_rewr @{lemma "1 * (x::real) = x" by (rule mult_1)}
 | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 221 | val mul_comm_thm = mk_rewr @{lemma "(x::real) * y = y * x" by simp}
 | 
| 64927 
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
 boehmes parents: 
63992diff
changeset | 222 | val mul_assocl_thm = mk_rewr @{lemma "((x::real) * y) * z = x * (y * z)" by simp}
 | 
| 
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
 boehmes parents: 
63992diff
changeset | 223 | val mul_assocr_thm = mk_rewr @{lemma "(x::real) * (y * z) = (x * y) * z" by simp}
 | 
| 
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
 boehmes parents: 
63992diff
changeset | 224 | val mul_divl_thm = mk_rewr @{lemma "((x::real) / y) * z = (x * z) / y" by simp}
 | 
| 
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
 boehmes parents: 
63992diff
changeset | 225 | val mul_divr_thm = mk_rewr @{lemma "(x::real) * (y / z) = (x * y) / z" by simp}
 | 
| 
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
 boehmes parents: 
63992diff
changeset | 226 | val div_zero_thm = mk_rewr @{lemma "0 / (x::real) = 0" by (rule div_0)}
 | 
| 
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
 boehmes parents: 
63992diff
changeset | 227 | val div_one_thm = mk_rewr @{lemma "(x::real) / 1 = x" by (rule div_by_1)}
 | 
| 
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
 boehmes parents: 
63992diff
changeset | 228 | val div_numl_thm = mk_rewr @{lemma "(x::real) / y = x * (1 / y)" by simp}
 | 
| 
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
 boehmes parents: 
63992diff
changeset | 229 | val div_numr_thm = mk_rewr @{lemma "(x::real) / y = (1 / y) * x" by simp}
 | 
| 
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
 boehmes parents: 
63992diff
changeset | 230 | val div_mull_thm = mk_rewr @{lemma "((x::real) * y) / z = x * (y / z)" by simp}
 | 
| 
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
 boehmes parents: 
63992diff
changeset | 231 | val div_mulr_thm = mk_rewr @{lemma "(x::real) / (y * z) = (1 / y) * (x / z)" by simp}
 | 
| 
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
 boehmes parents: 
63992diff
changeset | 232 | val div_divl_thm = mk_rewr @{lemma "((x::real) / y) / z = x / (y * z)" by simp}
 | 
| 
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
 boehmes parents: 
63992diff
changeset | 233 | val div_divr_thm = mk_rewr @{lemma "(x::real) / (y / z) = (x * z) / y" by simp}
 | 
| 
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
 boehmes parents: 
63992diff
changeset | 234 | val min_eq_thm = mk_rewr @{lemma "min (x::real) x = x" by (simp add: min_def)}
 | 
| 
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
 boehmes parents: 
63992diff
changeset | 235 | val min_lt_thm = mk_rewr @{lemma "min (x::real) y = (if x <= y then x else y)" by (rule min_def)}
 | 
| 
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
 boehmes parents: 
63992diff
changeset | 236 | val min_gt_thm = mk_rewr @{lemma "min (x::real) y = (if y <= x then y else x)"
 | 
| 
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
 boehmes parents: 
63992diff
changeset | 237 | by (simp add: min_def)} | 
| 
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
 boehmes parents: 
63992diff
changeset | 238 | val max_eq_thm = mk_rewr @{lemma "max (x::real) x = x" by (simp add: max_def)}
 | 
| 
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
 boehmes parents: 
63992diff
changeset | 239 | val max_lt_thm = mk_rewr @{lemma "max (x::real) y = (if x <= y then y else x)" by (rule max_def)}
 | 
| 
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
 boehmes parents: 
63992diff
changeset | 240 | val max_gt_thm = mk_rewr @{lemma "max (x::real) y = (if y <= x then x else y)"
 | 
| 
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
 boehmes parents: 
63992diff
changeset | 241 | by (simp add: max_def)} | 
| 63960 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 242 | val abs_thm = mk_rewr @{lemma "abs (x::real) = (if 0 <= x then x else -x)" by simp}
 | 
| 64927 
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
 boehmes parents: 
63992diff
changeset | 243 | val sub_eq_thm = mk_rewr @{lemma "((x::real) = y) = (x - y = 0)" by simp}
 | 
| 63960 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 244 | val eq_le_thm = mk_rewr @{lemma "((x::real) = y) = ((x <= y) & (y <= x))" by (rule order_eq_iff)}
 | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 245 | val add_le_thm = mk_rewr @{lemma "((x::real) <= y) = (x + n <= y + n)" by simp}
 | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 246 | val add_lt_thm = mk_rewr @{lemma "((x::real) < y) = (x + n < y + n)" by simp}
 | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 247 | val sub_le_thm = mk_rewr @{lemma "((x::real) <= y) = (x - y <= 0)" by simp}
 | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 248 | val sub_lt_thm = mk_rewr @{lemma "((x::real) < y) = (x - y < 0)" by simp}
 | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 249 | val pos_mul_le_thm = mk_rewr @{lemma "0 < n ==> ((x::real) <= y) = (n * x <= n * y)" by simp}
 | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 250 | val neg_mul_le_thm = mk_rewr @{lemma "n < 0 ==> ((x::real) <= y) = (n * y <= n * x)" by simp}
 | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 251 | val pos_mul_lt_thm = mk_rewr @{lemma "0 < n ==> ((x::real) < y) = (n * x < n * y)" by simp}
 | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 252 | val neg_mul_lt_thm = mk_rewr @{lemma "n < 0 ==> ((x::real) < y) = (n * y < n * x)" by simp}
 | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 253 | val not_le_thm = mk_rewr @{lemma "(~((x::real) <= y)) = (y < x)" by auto}
 | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 254 | val not_lt_thm = mk_rewr @{lemma "(~((x::real) < y)) = (y <= x)" by auto}
 | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 255 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 256 | fun replay_rewr _ Argo_Proof.Rewr_Neg = Conv.rewr_conv neg_thm | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 257 | | replay_rewr ctxt (Argo_Proof.Rewr_Add ps) = add_conv ctxt ps | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 258 | | replay_rewr _ Argo_Proof.Rewr_Sub = Conv.rewr_conv sub_thm | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 259 | | replay_rewr _ Argo_Proof.Rewr_Mul_Zero = Conv.rewr_conv mul_zero_thm | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 260 | | replay_rewr _ Argo_Proof.Rewr_Mul_One = Conv.rewr_conv mul_one_thm | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 261 | | replay_rewr ctxt (Argo_Proof.Rewr_Mul_Nums (n, m)) = mul_nums_conv ctxt n m | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 262 | | replay_rewr _ Argo_Proof.Rewr_Mul_Comm = Conv.rewr_conv mul_comm_thm | 
| 64927 
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
 boehmes parents: 
63992diff
changeset | 263 | | replay_rewr _ (Argo_Proof.Rewr_Mul_Assoc Argo_Proof.Left) = Conv.rewr_conv mul_assocl_thm | 
| 
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
 boehmes parents: 
63992diff
changeset | 264 | | replay_rewr _ (Argo_Proof.Rewr_Mul_Assoc Argo_Proof.Right) = Conv.rewr_conv mul_assocr_thm | 
| 
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
 boehmes parents: 
63992diff
changeset | 265 | | replay_rewr _ (Argo_Proof.Rewr_Mul_Sum Argo_Proof.Left) = mul_sum_conv mul_suml_thm | 
| 
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
 boehmes parents: 
63992diff
changeset | 266 | | replay_rewr _ (Argo_Proof.Rewr_Mul_Sum Argo_Proof.Right) = mul_sum_conv mul_sumr_thm | 
| 
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
 boehmes parents: 
63992diff
changeset | 267 | | replay_rewr _ (Argo_Proof.Rewr_Mul_Div Argo_Proof.Left) = Conv.rewr_conv mul_divl_thm | 
| 
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
 boehmes parents: 
63992diff
changeset | 268 | | replay_rewr _ (Argo_Proof.Rewr_Mul_Div Argo_Proof.Right) = Conv.rewr_conv mul_divr_thm | 
| 63960 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 269 | | replay_rewr _ Argo_Proof.Rewr_Div_Zero = Conv.rewr_conv div_zero_thm | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 270 | | replay_rewr _ Argo_Proof.Rewr_Div_One = Conv.rewr_conv div_one_thm | 
| 64927 
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
 boehmes parents: 
63992diff
changeset | 271 | | replay_rewr ctxt (Argo_Proof.Rewr_Div_Nums (n, m)) = div_nums_conv ctxt n m | 
| 
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
 boehmes parents: 
63992diff
changeset | 272 | | replay_rewr _ (Argo_Proof.Rewr_Div_Num (Argo_Proof.Left, _)) = Conv.rewr_conv div_numl_thm | 
| 
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
 boehmes parents: 
63992diff
changeset | 273 | | replay_rewr ctxt (Argo_Proof.Rewr_Div_Num (Argo_Proof.Right, n)) = | 
| 
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
 boehmes parents: 
63992diff
changeset | 274 | Conv.rewr_conv div_numr_thm then_conv Conv.arg1_conv (inv_num_conv ctxt n) | 
| 
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
 boehmes parents: 
63992diff
changeset | 275 | | replay_rewr _ (Argo_Proof.Rewr_Div_Mul (Argo_Proof.Left, _)) = Conv.rewr_conv div_mull_thm | 
| 
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
 boehmes parents: 
63992diff
changeset | 276 | | replay_rewr ctxt (Argo_Proof.Rewr_Div_Mul (Argo_Proof.Right, n)) = | 
| 
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
 boehmes parents: 
63992diff
changeset | 277 | Conv.rewr_conv div_mulr_thm then_conv Conv.arg1_conv (inv_num_conv ctxt n) | 
| 
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
 boehmes parents: 
63992diff
changeset | 278 | | replay_rewr _ (Argo_Proof.Rewr_Div_Div Argo_Proof.Left) = Conv.rewr_conv div_divl_thm | 
| 
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
 boehmes parents: 
63992diff
changeset | 279 | | replay_rewr _ (Argo_Proof.Rewr_Div_Div Argo_Proof.Right) = Conv.rewr_conv div_divr_thm | 
| 
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
 boehmes parents: 
63992diff
changeset | 280 | | replay_rewr _ Argo_Proof.Rewr_Div_Sum = div_sum_conv | 
| 
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
 boehmes parents: 
63992diff
changeset | 281 | | replay_rewr _ Argo_Proof.Rewr_Min_Eq = Conv.rewr_conv min_eq_thm | 
| 
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
 boehmes parents: 
63992diff
changeset | 282 | | replay_rewr _ Argo_Proof.Rewr_Min_Lt = Conv.rewr_conv min_lt_thm | 
| 
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
 boehmes parents: 
63992diff
changeset | 283 | | replay_rewr _ Argo_Proof.Rewr_Min_Gt = Conv.rewr_conv min_gt_thm | 
| 
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
 boehmes parents: 
63992diff
changeset | 284 | | replay_rewr _ Argo_Proof.Rewr_Max_Eq = Conv.rewr_conv max_eq_thm | 
| 
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
 boehmes parents: 
63992diff
changeset | 285 | | replay_rewr _ Argo_Proof.Rewr_Max_Lt = Conv.rewr_conv max_lt_thm | 
| 
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
 boehmes parents: 
63992diff
changeset | 286 | | replay_rewr _ Argo_Proof.Rewr_Max_Gt = Conv.rewr_conv max_gt_thm | 
| 63960 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 287 | | replay_rewr _ Argo_Proof.Rewr_Abs = Conv.rewr_conv abs_thm | 
| 64927 
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
 boehmes parents: 
63992diff
changeset | 288 | | replay_rewr ctxt (Argo_Proof.Rewr_Eq_Nums b) = cmp_nums_conv ctxt b | 
| 
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
 boehmes parents: 
63992diff
changeset | 289 | | replay_rewr _ Argo_Proof.Rewr_Eq_Sub = Conv.rewr_conv sub_eq_thm | 
| 63960 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 290 | | replay_rewr _ Argo_Proof.Rewr_Eq_Le = Conv.rewr_conv eq_le_thm | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 291 | | replay_rewr ctxt (Argo_Proof.Rewr_Ineq_Nums (_, b)) = cmp_nums_conv ctxt b | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 292 | | replay_rewr ctxt (Argo_Proof.Rewr_Ineq_Add (Argo_Proof.Le, n)) = | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 293 | add_cmp_conv ctxt n add_le_thm | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 294 | | replay_rewr ctxt (Argo_Proof.Rewr_Ineq_Add (Argo_Proof.Lt, n)) = | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 295 | add_cmp_conv ctxt n add_lt_thm | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 296 | | replay_rewr _ (Argo_Proof.Rewr_Ineq_Sub Argo_Proof.Le) = Conv.rewr_conv sub_le_thm | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 297 | | replay_rewr _ (Argo_Proof.Rewr_Ineq_Sub Argo_Proof.Lt) = Conv.rewr_conv sub_lt_thm | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 298 | | replay_rewr ctxt (Argo_Proof.Rewr_Ineq_Mul (Argo_Proof.Le, n)) = | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 299 | mul_cmp_conv ctxt n pos_mul_le_thm neg_mul_le_thm | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 300 | | replay_rewr ctxt (Argo_Proof.Rewr_Ineq_Mul (Argo_Proof.Lt, n)) = | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 301 | mul_cmp_conv ctxt n pos_mul_lt_thm neg_mul_lt_thm | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 302 | | replay_rewr _ (Argo_Proof.Rewr_Not_Ineq Argo_Proof.Le) = Conv.rewr_conv not_le_thm | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 303 | | replay_rewr _ (Argo_Proof.Rewr_Not_Ineq Argo_Proof.Lt) = Conv.rewr_conv not_lt_thm | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 304 | | replay_rewr _ _ = Conv.no_conv | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 305 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 306 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 307 | (* proof replay *) | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 308 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 309 | val combine_thms = @{lemma
 | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 310 | "(a::real) < b ==> c < d ==> a + c < b + d" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 311 | "(a::real) < b ==> c <= d ==> a + c < b + d" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 312 | "(a::real) <= b ==> c < d ==> a + c < b + d" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 313 | "(a::real) <= b ==> c <= d ==> a + c <= b + d" | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 314 | by auto} | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 315 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 316 | fun combine thm1 thm2 = hd (Argo_Tactic.discharges thm2 (Argo_Tactic.discharges thm1 combine_thms)) | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 317 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 318 | fun replay _ _ Argo_Proof.Linear_Comb prems = SOME (uncurry (fold_rev combine) (split_last prems)) | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 319 | | replay _ _ _ _ = NONE | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 320 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 321 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 322 | (* real extension of the Argo solver *) | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 323 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 324 | val _ = Theory.setup (Argo_Tactic.add_extension {
 | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 325 | trans_type = trans_type, | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 326 | trans_term = trans_term, | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 327 | term_of = term_of, | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 328 | replay_rewr = replay_rewr, | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 329 | replay = replay}) | 
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 330 | |
| 
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
 boehmes parents: diff
changeset | 331 | end |