| author | haftmann | 
| Wed, 21 Oct 2009 12:09:37 +0200 | |
| changeset 33049 | c38f02fdf35d | 
| parent 33040 | cffdb7b28498 | 
| child 33043 | ff71cadefb14 | 
| permissions | -rw-r--r-- | 
| 31118 
541d43bee678
Isolated decision procedure for noms and the general arithmetic solver
 chaieb parents: 
30868diff
changeset | 1 | (* Title: Library/normarith.ML | 
| 
541d43bee678
Isolated decision procedure for noms and the general arithmetic solver
 chaieb parents: 
30868diff
changeset | 2 | Author: Amine Chaieb, University of Cambridge | 
| 
541d43bee678
Isolated decision procedure for noms and the general arithmetic solver
 chaieb parents: 
30868diff
changeset | 3 | Description: A simple decision procedure for linear problems in euclidean space | 
| 
541d43bee678
Isolated decision procedure for noms and the general arithmetic solver
 chaieb parents: 
30868diff
changeset | 4 | *) | 
| 29841 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 5 | |
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 6 | (* Now the norm procedure for euclidean spaces *) | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 7 | |
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 8 | |
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 9 | signature NORM_ARITH = | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 10 | sig | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 11 | val norm_arith : Proof.context -> conv | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 12 | val norm_arith_tac : Proof.context -> int -> tactic | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 13 | end | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 14 | |
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 15 | structure NormArith : NORM_ARITH = | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 16 | struct | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 17 | |
| 32832 
4602cb6ae0b5
additional fixes in normarith.ML due to FuncFun and FuncUtil changes
 Philipp Meyer parents: 
32645diff
changeset | 18 | open Conv; | 
| 29841 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 19 | val bool_eq = op = : bool *bool -> bool | 
| 30868 | 20 | fun dest_ratconst t = case term_of t of | 
| 29841 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 21 |    Const(@{const_name divide}, _)$a$b => Rat.rat_of_quotient(HOLogic.dest_number a |> snd, HOLogic.dest_number b |> snd)
 | 
| 30868 | 22 |  | Const(@{const_name inverse}, _)$a => Rat.rat_of_quotient(1, HOLogic.dest_number a |> snd)
 | 
| 29841 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 23 | | _ => Rat.rat_of_int (HOLogic.dest_number (term_of t) |> snd) | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 24 | fun is_ratconst t = can dest_ratconst t | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 25 | fun augment_norm b t acc = case term_of t of | 
| 32832 
4602cb6ae0b5
additional fixes in normarith.ML due to FuncFun and FuncUtil changes
 Philipp Meyer parents: 
32645diff
changeset | 26 |      Const(@{const_name norm}, _) $ _ => insert (eq_pair bool_eq (op aconvc)) (b,Thm.dest_arg t) acc
 | 
| 29841 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 27 | | _ => acc | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 28 | fun find_normedterms t acc = case term_of t of | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 29 |     @{term "op + :: real => _"}$_$_ =>
 | 
| 32832 
4602cb6ae0b5
additional fixes in normarith.ML due to FuncFun and FuncUtil changes
 Philipp Meyer parents: 
32645diff
changeset | 30 | find_normedterms (Thm.dest_arg1 t) (find_normedterms (Thm.dest_arg t) acc) | 
| 29841 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 31 |       | @{term "op * :: real => _"}$_$n =>
 | 
| 32832 
4602cb6ae0b5
additional fixes in normarith.ML due to FuncFun and FuncUtil changes
 Philipp Meyer parents: 
32645diff
changeset | 32 | if not (is_ratconst (Thm.dest_arg1 t)) then acc else | 
| 
4602cb6ae0b5
additional fixes in normarith.ML due to FuncFun and FuncUtil changes
 Philipp Meyer parents: 
32645diff
changeset | 33 | augment_norm (dest_ratconst (Thm.dest_arg1 t) >=/ Rat.zero) | 
| 
4602cb6ae0b5
additional fixes in normarith.ML due to FuncFun and FuncUtil changes
 Philipp Meyer parents: 
32645diff
changeset | 34 | (Thm.dest_arg t) acc | 
| 29841 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 35 | | _ => augment_norm true t acc | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 36 | |
| 32832 
4602cb6ae0b5
additional fixes in normarith.ML due to FuncFun and FuncUtil changes
 Philipp Meyer parents: 
32645diff
changeset | 37 | val cterm_lincomb_neg = FuncUtil.Ctermfunc.map Rat.neg | 
| 29841 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 38 | fun cterm_lincomb_cmul c t = | 
| 32832 
4602cb6ae0b5
additional fixes in normarith.ML due to FuncFun and FuncUtil changes
 Philipp Meyer parents: 
32645diff
changeset | 39 | if c =/ Rat.zero then FuncUtil.Ctermfunc.empty else FuncUtil.Ctermfunc.map (fn x => x */ c) t | 
| 
4602cb6ae0b5
additional fixes in normarith.ML due to FuncFun and FuncUtil changes
 Philipp Meyer parents: 
32645diff
changeset | 40 | fun cterm_lincomb_add l r = FuncUtil.Ctermfunc.combine (curry op +/) (fn x => x =/ Rat.zero) l r | 
| 29841 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 41 | fun cterm_lincomb_sub l r = cterm_lincomb_add l (cterm_lincomb_neg r) | 
| 32832 
4602cb6ae0b5
additional fixes in normarith.ML due to FuncFun and FuncUtil changes
 Philipp Meyer parents: 
32645diff
changeset | 42 | fun cterm_lincomb_eq l r = FuncUtil.Ctermfunc.is_empty (cterm_lincomb_sub l r) | 
| 29841 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 43 | |
| 32832 
4602cb6ae0b5
additional fixes in normarith.ML due to FuncFun and FuncUtil changes
 Philipp Meyer parents: 
32645diff
changeset | 44 | val int_lincomb_neg = FuncUtil.Intfunc.map Rat.neg | 
| 29841 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 45 | fun int_lincomb_cmul c t = | 
| 32832 
4602cb6ae0b5
additional fixes in normarith.ML due to FuncFun and FuncUtil changes
 Philipp Meyer parents: 
32645diff
changeset | 46 | if c =/ Rat.zero then FuncUtil.Intfunc.empty else FuncUtil.Intfunc.map (fn x => x */ c) t | 
| 
4602cb6ae0b5
additional fixes in normarith.ML due to FuncFun and FuncUtil changes
 Philipp Meyer parents: 
32645diff
changeset | 47 | fun int_lincomb_add l r = FuncUtil.Intfunc.combine (curry op +/) (fn x => x =/ Rat.zero) l r | 
| 29841 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 48 | fun int_lincomb_sub l r = int_lincomb_add l (int_lincomb_neg r) | 
| 32832 
4602cb6ae0b5
additional fixes in normarith.ML due to FuncFun and FuncUtil changes
 Philipp Meyer parents: 
32645diff
changeset | 49 | fun int_lincomb_eq l r = FuncUtil.Intfunc.is_empty (int_lincomb_sub l r) | 
| 29841 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 50 | |
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 51 | fun vector_lincomb t = case term_of t of | 
| 31445 
c8a474a919a7
generalize norm method to work over class real_normed_vector
 huffman parents: 
31344diff
changeset | 52 |    Const(@{const_name plus}, _) $ _ $ _ =>
 | 
| 32832 
4602cb6ae0b5
additional fixes in normarith.ML due to FuncFun and FuncUtil changes
 Philipp Meyer parents: 
32645diff
changeset | 53 | cterm_lincomb_add (vector_lincomb (Thm.dest_arg1 t)) (vector_lincomb (Thm.dest_arg t)) | 
| 31445 
c8a474a919a7
generalize norm method to work over class real_normed_vector
 huffman parents: 
31344diff
changeset | 54 |  | Const(@{const_name minus}, _) $ _ $ _ =>
 | 
| 32832 
4602cb6ae0b5
additional fixes in normarith.ML due to FuncFun and FuncUtil changes
 Philipp Meyer parents: 
32645diff
changeset | 55 | cterm_lincomb_sub (vector_lincomb (Thm.dest_arg1 t)) (vector_lincomb (Thm.dest_arg t)) | 
| 31445 
c8a474a919a7
generalize norm method to work over class real_normed_vector
 huffman parents: 
31344diff
changeset | 56 |  | Const(@{const_name scaleR}, _)$_$_ =>
 | 
| 32832 
4602cb6ae0b5
additional fixes in normarith.ML due to FuncFun and FuncUtil changes
 Philipp Meyer parents: 
32645diff
changeset | 57 | cterm_lincomb_cmul (dest_ratconst (Thm.dest_arg1 t)) (vector_lincomb (Thm.dest_arg t)) | 
| 31445 
c8a474a919a7
generalize norm method to work over class real_normed_vector
 huffman parents: 
31344diff
changeset | 58 |  | Const(@{const_name uminus}, _)$_ =>
 | 
| 32832 
4602cb6ae0b5
additional fixes in normarith.ML due to FuncFun and FuncUtil changes
 Philipp Meyer parents: 
32645diff
changeset | 59 | cterm_lincomb_neg (vector_lincomb (Thm.dest_arg t)) | 
| 31445 
c8a474a919a7
generalize norm method to work over class real_normed_vector
 huffman parents: 
31344diff
changeset | 60 | (* FIXME: how should we handle numerals? | 
| 
c8a474a919a7
generalize norm method to work over class real_normed_vector
 huffman parents: 
31344diff
changeset | 61 |  | Const(@ {const_name vec},_)$_ => 
 | 
| 29841 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 62 | let | 
| 32832 
4602cb6ae0b5
additional fixes in normarith.ML due to FuncFun and FuncUtil changes
 Philipp Meyer parents: 
32645diff
changeset | 63 | val b = ((snd o HOLogic.dest_number o term_of o Thm.dest_arg) t = 0 | 
| 29841 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 64 | handle TERM _=> false) | 
| 32832 
4602cb6ae0b5
additional fixes in normarith.ML due to FuncFun and FuncUtil changes
 Philipp Meyer parents: 
32645diff
changeset | 65 | in if b then FuncUtil.Ctermfunc.onefunc (t,Rat.one) | 
| 
4602cb6ae0b5
additional fixes in normarith.ML due to FuncFun and FuncUtil changes
 Philipp Meyer parents: 
32645diff
changeset | 66 | else FuncUtil.Ctermfunc.empty | 
| 29841 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 67 | end | 
| 31445 
c8a474a919a7
generalize norm method to work over class real_normed_vector
 huffman parents: 
31344diff
changeset | 68 | *) | 
| 32832 
4602cb6ae0b5
additional fixes in normarith.ML due to FuncFun and FuncUtil changes
 Philipp Meyer parents: 
32645diff
changeset | 69 | | _ => FuncUtil.Ctermfunc.onefunc (t,Rat.one) | 
| 29841 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 70 | |
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 71 | fun vector_lincombs ts = | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 72 | fold_rev | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 73 | (fn t => fn fns => case AList.lookup (op aconvc) fns t of | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 74 | NONE => | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 75 | let val f = vector_lincomb t | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 76 | in case find_first (fn (_,f') => cterm_lincomb_eq f f') fns of | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 77 | SOME (_,f') => (t,f') :: fns | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 78 | | NONE => (t,f) :: fns | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 79 | end | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 80 | | SOME _ => fns) ts [] | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 81 | |
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 82 | fun replacenegnorms cv t = case term_of t of | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 83 |   @{term "op + :: real => _"}$_$_ => binop_conv (replacenegnorms cv) t
 | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 84 | | @{term "op * :: real => _"}$_$_ => 
 | 
| 32832 
4602cb6ae0b5
additional fixes in normarith.ML due to FuncFun and FuncUtil changes
 Philipp Meyer parents: 
32645diff
changeset | 85 | if dest_ratconst (Thm.dest_arg1 t) </ Rat.zero then arg_conv cv t else reflexive t | 
| 29841 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 86 | | _ => reflexive t | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 87 | fun flip v eq = | 
| 32832 
4602cb6ae0b5
additional fixes in normarith.ML due to FuncFun and FuncUtil changes
 Philipp Meyer parents: 
32645diff
changeset | 88 | if FuncUtil.Ctermfunc.defined eq v | 
| 
4602cb6ae0b5
additional fixes in normarith.ML due to FuncFun and FuncUtil changes
 Philipp Meyer parents: 
32645diff
changeset | 89 | then FuncUtil.Ctermfunc.update (v, Rat.neg (FuncUtil.Ctermfunc.apply eq v)) eq else eq | 
| 29841 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 90 | fun allsubsets s = case s of | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 91 | [] => [[]] | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 92 | |(a::t) => let val res = allsubsets t in | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 93 | map (cons a) res @ res end | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 94 | fun evaluate env lin = | 
| 32832 
4602cb6ae0b5
additional fixes in normarith.ML due to FuncFun and FuncUtil changes
 Philipp Meyer parents: 
32645diff
changeset | 95 | FuncUtil.Intfunc.fold (fn (x,c) => fn s => s +/ c */ (FuncUtil.Intfunc.apply env x)) | 
| 29841 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 96 | lin Rat.zero | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 97 | |
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 98 | fun solve (vs,eqs) = case (vs,eqs) of | 
| 32832 
4602cb6ae0b5
additional fixes in normarith.ML due to FuncFun and FuncUtil changes
 Philipp Meyer parents: 
32645diff
changeset | 99 | ([],[]) => SOME (FuncUtil.Intfunc.onefunc (0,Rat.one)) | 
| 29841 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 100 | |(_,eq::oeqs) => | 
| 32832 
4602cb6ae0b5
additional fixes in normarith.ML due to FuncFun and FuncUtil changes
 Philipp Meyer parents: 
32645diff
changeset | 101 | (case filter (member (op =) vs) (FuncUtil.Intfunc.dom eq) of (*FIXME use find_first here*) | 
| 29841 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 102 | [] => NONE | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 103 | | v::_ => | 
| 32832 
4602cb6ae0b5
additional fixes in normarith.ML due to FuncFun and FuncUtil changes
 Philipp Meyer parents: 
32645diff
changeset | 104 | if FuncUtil.Intfunc.defined eq v | 
| 29841 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 105 | then | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 106 | let | 
| 32832 
4602cb6ae0b5
additional fixes in normarith.ML due to FuncFun and FuncUtil changes
 Philipp Meyer parents: 
32645diff
changeset | 107 | val c = FuncUtil.Intfunc.apply eq v | 
| 29841 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 108 | val vdef = int_lincomb_cmul (Rat.neg (Rat.inv c)) eq | 
| 32832 
4602cb6ae0b5
additional fixes in normarith.ML due to FuncFun and FuncUtil changes
 Philipp Meyer parents: 
32645diff
changeset | 109 | fun eliminate eqn = if not (FuncUtil.Intfunc.defined eqn v) then eqn | 
| 
4602cb6ae0b5
additional fixes in normarith.ML due to FuncFun and FuncUtil changes
 Philipp Meyer parents: 
32645diff
changeset | 110 | else int_lincomb_add (int_lincomb_cmul (FuncUtil.Intfunc.apply eqn v) vdef) eqn | 
| 33040 | 111 | in (case solve (remove (op =) v vs, map eliminate oeqs) of | 
| 29841 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 112 | NONE => NONE | 
| 32832 
4602cb6ae0b5
additional fixes in normarith.ML due to FuncFun and FuncUtil changes
 Philipp Meyer parents: 
32645diff
changeset | 113 | | SOME soln => SOME (FuncUtil.Intfunc.update (v, evaluate soln (FuncUtil.Intfunc.delete_safe v vdef)) soln)) | 
| 29841 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 114 | end | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 115 | else NONE) | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 116 | |
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 117 | fun combinations k l = if k = 0 then [[]] else | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 118 | case l of | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 119 | [] => [] | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 120 | | h::t => map (cons h) (combinations (k - 1) t) @ combinations k t | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 121 | |
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 122 | |
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 123 | fun forall2 p l1 l2 = case (l1,l2) of | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 124 | ([],[]) => true | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 125 | | (h1::t1,h2::t2) => p h1 h2 andalso forall2 p t1 t2 | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 126 | | _ => false; | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 127 | |
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 128 | |
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 129 | fun vertices vs eqs = | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 130 | let | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 131 | fun vertex cmb = case solve(vs,cmb) of | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 132 | NONE => NONE | 
| 32832 
4602cb6ae0b5
additional fixes in normarith.ML due to FuncFun and FuncUtil changes
 Philipp Meyer parents: 
32645diff
changeset | 133 | | SOME soln => SOME (map (fn v => FuncUtil.Intfunc.tryapplyd soln v Rat.zero) vs) | 
| 29841 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 134 | val rawvs = map_filter vertex (combinations (length vs) eqs) | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 135 | val unset = filter (forall (fn c => c >=/ Rat.zero)) rawvs | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 136 | in fold_rev (insert (uncurry (forall2 (curry op =/)))) unset [] | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 137 | end | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 138 | |
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 139 | fun subsumes l m = forall2 (fn x => fn y => Rat.abs x <=/ Rat.abs y) l m | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 140 | |
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 141 | fun subsume todo dun = case todo of | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 142 | [] => dun | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 143 | |v::ovs => | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 144 | let val dun' = if exists (fn w => subsumes w v) dun then dun | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 145 | else v::(filter (fn w => not(subsumes v w)) dun) | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 146 | in subsume ovs dun' | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 147 | end; | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 148 | |
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 149 | fun match_mp PQ P = P RS PQ; | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 150 | |
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 151 | fun cterm_of_rat x = | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 152 | let val (a, b) = Rat.quotient_of_rat x | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 153 | in | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 154 |  if b = 1 then Numeral.mk_cnumber @{ctyp "real"} a
 | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 155 |   else Thm.capply (Thm.capply @{cterm "op / :: real => _"} 
 | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 156 |                    (Numeral.mk_cnumber @{ctyp "real"} a))
 | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 157 |         (Numeral.mk_cnumber @{ctyp "real"} b)
 | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 158 | end; | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 159 | |
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 160 | fun norm_cmul_rule c th = instantiate' [] [SOME (cterm_of_rat c)] (th RS @{thm norm_cmul_rule_thm});
 | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 161 | |
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 162 | fun norm_add_rule th1 th2 = [th1, th2] MRS @{thm norm_add_rule_thm};
 | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 163 | |
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 164 | (* I think here the static context should be sufficient!! *) | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 165 | fun inequality_canon_rule ctxt = | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 166 | let | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 167 | (* FIXME : Should be computed statically!! *) | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 168 | val real_poly_conv = | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 169 | Normalizer.semiring_normalize_wrapper ctxt | 
| 33035 | 170 |      (the (NormalizerData.match ctxt @{cterm "(0::real) + 1"}))
 | 
| 29841 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 171 |  in fconv_rule (arg_conv ((rewr_conv @{thm ge_iff_diff_ge_0}) then_conv arg_conv (field_comp_conv then_conv real_poly_conv)))
 | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 172 | end; | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 173 | |
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 174 | fun absc cv ct = case term_of ct of | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 175 | Abs (v,_, _) => | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 176 | let val (x,t) = Thm.dest_abs (SOME v) ct | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 177 | in Thm.abstract_rule ((fst o dest_Free o term_of) x) x (cv t) | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 178 | end | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 179 | | _ => all_conv ct; | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 180 | |
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 181 | fun sub_conv cv ct = (comb_conv cv else_conv absc cv) ct; | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 182 | fun botc1 conv ct = | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 183 | ((sub_conv (botc1 conv)) then_conv (conv else_conv all_conv)) ct; | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 184 | |
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 185 | fun rewrs_conv eqs ct = first_conv (map rewr_conv eqs) ct; | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 186 |  val apply_pth1 = rewr_conv @{thm pth_1};
 | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 187 |  val apply_pth2 = rewr_conv @{thm pth_2};
 | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 188 |  val apply_pth3 = rewr_conv @{thm pth_3};
 | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 189 |  val apply_pth4 = rewrs_conv @{thms pth_4};
 | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 190 |  val apply_pth5 = rewr_conv @{thm pth_5};
 | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 191 |  val apply_pth6 = rewr_conv @{thm pth_6};
 | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 192 |  val apply_pth7 = rewrs_conv @{thms pth_7};
 | 
| 31445 
c8a474a919a7
generalize norm method to work over class real_normed_vector
 huffman parents: 
31344diff
changeset | 193 |  val apply_pth8 = rewr_conv @{thm pth_8} then_conv arg1_conv field_comp_conv then_conv (try_conv (rewr_conv (mk_meta_eq @{thm scaleR_zero_left})));
 | 
| 29841 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 194 |  val apply_pth9 = rewrs_conv @{thms pth_9} then_conv arg1_conv (arg1_conv field_comp_conv);
 | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 195 |  val apply_ptha = rewr_conv @{thm pth_a};
 | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 196 |  val apply_pthb = rewrs_conv @{thms pth_b};
 | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 197 |  val apply_pthc = rewrs_conv @{thms pth_c};
 | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 198 |  val apply_pthd = try_conv (rewr_conv @{thm pth_d});
 | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 199 | |
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 200 | fun headvector t = case t of | 
| 31445 
c8a474a919a7
generalize norm method to work over class real_normed_vector
 huffman parents: 
31344diff
changeset | 201 |   Const(@{const_name plus}, _)$
 | 
| 
c8a474a919a7
generalize norm method to work over class real_normed_vector
 huffman parents: 
31344diff
changeset | 202 |    (Const(@{const_name scaleR}, _)$l$v)$r => v
 | 
| 
c8a474a919a7
generalize norm method to work over class real_normed_vector
 huffman parents: 
31344diff
changeset | 203 |  | Const(@{const_name scaleR}, _)$l$v => v
 | 
| 29841 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 204 | | _ => error "headvector: non-canonical term" | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 205 | |
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 206 | fun vector_cmul_conv ct = | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 207 | ((apply_pth5 then_conv arg1_conv field_comp_conv) else_conv | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 208 | (apply_pth6 then_conv binop_conv vector_cmul_conv)) ct | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 209 | |
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 210 | fun vector_add_conv ct = apply_pth7 ct | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 211 | handle CTERM _ => | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 212 | (apply_pth8 ct | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 213 | handle CTERM _ => | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 214 | (case term_of ct of | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 215 |      Const(@{const_name plus},_)$lt$rt =>
 | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 216 | let | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 217 | val l = headvector lt | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 218 | val r = headvector rt | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 219 | in (case TermOrd.fast_term_ord (l,r) of | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 220 | LESS => (apply_pthb then_conv arg_conv vector_add_conv | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 221 | then_conv apply_pthd) ct | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 222 | | GREATER => (apply_pthc then_conv arg_conv vector_add_conv | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 223 | then_conv apply_pthd) ct | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 224 | | EQUAL => (apply_pth9 then_conv | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 225 | ((apply_ptha then_conv vector_add_conv) else_conv | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 226 | arg_conv vector_add_conv then_conv apply_pthd)) ct) | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 227 | end | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 228 | | _ => reflexive ct)) | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 229 | |
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 230 | fun vector_canon_conv ct = case term_of ct of | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 231 |  Const(@{const_name plus},_)$_$_ =>
 | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 232 | let | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 233 | val ((p,l),r) = Thm.dest_comb ct |>> Thm.dest_comb | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 234 | val lth = vector_canon_conv l | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 235 | val rth = vector_canon_conv r | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 236 | val th = Drule.binop_cong_rule p lth rth | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 237 | in fconv_rule (arg_conv vector_add_conv) th end | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 238 | |
| 31445 
c8a474a919a7
generalize norm method to work over class real_normed_vector
 huffman parents: 
31344diff
changeset | 239 | | Const(@{const_name scaleR}, _)$_$_ =>
 | 
| 29841 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 240 | let | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 241 | val (p,r) = Thm.dest_comb ct | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 242 | val rth = Drule.arg_cong_rule p (vector_canon_conv r) | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 243 | in fconv_rule (arg_conv (apply_pth4 else_conv vector_cmul_conv)) rth | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 244 | end | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 245 | |
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 246 | | Const(@{const_name minus},_)$_$_ => (apply_pth2 then_conv vector_canon_conv) ct
 | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 247 | |
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 248 | | Const(@{const_name uminus},_)$_ => (apply_pth3 then_conv vector_canon_conv) ct
 | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 249 | |
| 31445 
c8a474a919a7
generalize norm method to work over class real_normed_vector
 huffman parents: 
31344diff
changeset | 250 | (* FIXME | 
| 29841 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 251 | | Const(@{const_name vec},_)$n => 
 | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 252 | let val n = Thm.dest_arg ct | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 253 | in if is_ratconst n andalso not (dest_ratconst n =/ Rat.zero) | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 254 | then reflexive ct else apply_pth1 ct | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 255 | end | 
| 31445 
c8a474a919a7
generalize norm method to work over class real_normed_vector
 huffman parents: 
31344diff
changeset | 256 | *) | 
| 29841 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 257 | | _ => apply_pth1 ct | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 258 | |
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 259 | fun norm_canon_conv ct = case term_of ct of | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 260 |   Const(@{const_name norm},_)$_ => arg_conv vector_canon_conv ct
 | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 261 |  | _ => raise CTERM ("norm_canon_conv", [ct])
 | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 262 | |
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 263 | fun fold_rev2 f [] [] z = z | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 264 | | fold_rev2 f (x::xs) (y::ys) z = f x y (fold_rev2 f xs ys z) | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 265 | | fold_rev2 f _ _ _ = raise UnequalLengths; | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 266 | |
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 267 | fun int_flip v eq = | 
| 32832 
4602cb6ae0b5
additional fixes in normarith.ML due to FuncFun and FuncUtil changes
 Philipp Meyer parents: 
32645diff
changeset | 268 | if FuncUtil.Intfunc.defined eq v | 
| 
4602cb6ae0b5
additional fixes in normarith.ML due to FuncFun and FuncUtil changes
 Philipp Meyer parents: 
32645diff
changeset | 269 | then FuncUtil.Intfunc.update (v, Rat.neg (FuncUtil.Intfunc.apply eq v)) eq else eq; | 
| 29841 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 270 | |
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 271 | local | 
| 31445 
c8a474a919a7
generalize norm method to work over class real_normed_vector
 huffman parents: 
31344diff
changeset | 272 |  val pth_zero = @{thm norm_zero}
 | 
| 32832 
4602cb6ae0b5
additional fixes in normarith.ML due to FuncFun and FuncUtil changes
 Philipp Meyer parents: 
32645diff
changeset | 273 | val tv_n = (ctyp_of_term o Thm.dest_arg o Thm.dest_arg1 o Thm.dest_arg o cprop_of) | 
| 29841 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 274 | pth_zero | 
| 32832 
4602cb6ae0b5
additional fixes in normarith.ML due to FuncFun and FuncUtil changes
 Philipp Meyer parents: 
32645diff
changeset | 275 | val concl = Thm.dest_arg o cprop_of | 
| 29841 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 276 | fun real_vector_combo_prover ctxt translator (nubs,ges,gts) = | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 277 | let | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 278 | (* FIXME: Should be computed statically!!*) | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 279 | val real_poly_conv = | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 280 | Normalizer.semiring_normalize_wrapper ctxt | 
| 33035 | 281 |        (the (NormalizerData.match ctxt @{cterm "(0::real) + 1"}))
 | 
| 32832 
4602cb6ae0b5
additional fixes in normarith.ML due to FuncFun and FuncUtil changes
 Philipp Meyer parents: 
32645diff
changeset | 282 | val sources = map (Thm.dest_arg o Thm.dest_arg1 o concl) nubs | 
| 
4602cb6ae0b5
additional fixes in normarith.ML due to FuncFun and FuncUtil changes
 Philipp Meyer parents: 
32645diff
changeset | 283 | val rawdests = fold_rev (find_normedterms o Thm.dest_arg o concl) (ges @ gts) [] | 
| 29841 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 284 | val _ = if not (forall fst rawdests) then error "real_vector_combo_prover: Sanity check" | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 285 | else () | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 286 | val dests = distinct (op aconvc) (map snd rawdests) | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 287 | val srcfuns = map vector_lincomb sources | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 288 | val destfuns = map vector_lincomb dests | 
| 33038 | 289 | val vvs = fold_rev (curry (union op aconvc) o FuncUtil.Ctermfunc.dom) (srcfuns @ destfuns) [] | 
| 29841 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 290 | val n = length srcfuns | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 291 | val nvs = 1 upto n | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 292 | val srccombs = srcfuns ~~ nvs | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 293 | fun consider d = | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 294 | let | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 295 | fun coefficients x = | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 296 | let | 
| 32832 
4602cb6ae0b5
additional fixes in normarith.ML due to FuncFun and FuncUtil changes
 Philipp Meyer parents: 
32645diff
changeset | 297 | val inp = if FuncUtil.Ctermfunc.defined d x then FuncUtil.Intfunc.onefunc (0, Rat.neg(FuncUtil.Ctermfunc.apply d x)) | 
| 
4602cb6ae0b5
additional fixes in normarith.ML due to FuncFun and FuncUtil changes
 Philipp Meyer parents: 
32645diff
changeset | 298 | else FuncUtil.Intfunc.empty | 
| 
4602cb6ae0b5
additional fixes in normarith.ML due to FuncFun and FuncUtil changes
 Philipp Meyer parents: 
32645diff
changeset | 299 | in fold_rev (fn (f,v) => fn g => if FuncUtil.Ctermfunc.defined f x then FuncUtil.Intfunc.update (v, FuncUtil.Ctermfunc.apply f x) g else g) srccombs inp | 
| 29841 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 300 | end | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 301 | val equations = map coefficients vvs | 
| 32832 
4602cb6ae0b5
additional fixes in normarith.ML due to FuncFun and FuncUtil changes
 Philipp Meyer parents: 
32645diff
changeset | 302 | val inequalities = map (fn n => FuncUtil.Intfunc.onefunc (n,Rat.one)) nvs | 
| 29841 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 303 | fun plausiblevertices f = | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 304 | let | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 305 | val flippedequations = map (fold_rev int_flip f) equations | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 306 | val constraints = flippedequations @ inequalities | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 307 | val rawverts = vertices nvs constraints | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 308 | fun check_solution v = | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 309 | let | 
| 32832 
4602cb6ae0b5
additional fixes in normarith.ML due to FuncFun and FuncUtil changes
 Philipp Meyer parents: 
32645diff
changeset | 310 | val f = fold_rev2 (curry FuncUtil.Intfunc.update) nvs v (FuncUtil.Intfunc.onefunc (0, Rat.one)) | 
| 29841 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 311 | in forall (fn e => evaluate f e =/ Rat.zero) flippedequations | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 312 | end | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 313 | val goodverts = filter check_solution rawverts | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 314 | val signfixups = map (fn n => if n mem_int f then ~1 else 1) nvs | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 315 | in map (map2 (fn s => fn c => Rat.rat_of_int s */ c) signfixups) goodverts | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 316 | end | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 317 | val allverts = fold_rev append (map plausiblevertices (allsubsets nvs)) [] | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 318 | in subsume allverts [] | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 319 | end | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 320 | fun compute_ineq v = | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 321 | let | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 322 | val ths = map_filter (fn (v,t) => if v =/ Rat.zero then NONE | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 323 | else SOME(norm_cmul_rule v t)) | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 324 | (v ~~ nubs) | 
| 32402 | 325 | fun end_itlist f xs = split_last xs |> uncurry (fold_rev f) | 
| 29841 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 326 | in inequality_canon_rule ctxt (end_itlist norm_add_rule ths) | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 327 | end | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 328 | val ges' = map_filter (try compute_ineq) (fold_rev (append o consider) destfuns []) @ | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 329 | map (inequality_canon_rule ctxt) nubs @ ges | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 330 | val zerodests = filter | 
| 32832 
4602cb6ae0b5
additional fixes in normarith.ML due to FuncFun and FuncUtil changes
 Philipp Meyer parents: 
32645diff
changeset | 331 | (fn t => null (FuncUtil.Ctermfunc.dom (vector_lincomb t))) (map snd rawdests) | 
| 29841 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 332 | |
| 32645 
1cc5b24f5a01
sos method generates and uses proof certificates
 Philipp Meyer parents: 
32402diff
changeset | 333 | in fst (RealArith.real_linear_prover translator | 
| 31445 
c8a474a919a7
generalize norm method to work over class real_normed_vector
 huffman parents: 
31344diff
changeset | 334 | (map (fn t => instantiate ([(tv_n, ctyp_of_term t)],[]) pth_zero) | 
| 29841 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 335 | zerodests, | 
| 32402 | 336 | map (fconv_rule (try_conv (More_Conv.top_sweep_conv (K norm_canon_conv) ctxt) then_conv | 
| 29841 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 337 | arg_conv (arg_conv real_poly_conv))) ges', | 
| 32402 | 338 | map (fconv_rule (try_conv (More_Conv.top_sweep_conv (K norm_canon_conv) ctxt) then_conv | 
| 32645 
1cc5b24f5a01
sos method generates and uses proof certificates
 Philipp Meyer parents: 
32402diff
changeset | 339 | arg_conv (arg_conv real_poly_conv))) gts)) | 
| 29841 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 340 | end | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 341 | in val real_vector_combo_prover = real_vector_combo_prover | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 342 | end; | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 343 | |
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 344 | local | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 345 |  val pth = @{thm norm_imp_pos_and_ge}
 | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 346 | val norm_mp = match_mp pth | 
| 32832 
4602cb6ae0b5
additional fixes in normarith.ML due to FuncFun and FuncUtil changes
 Philipp Meyer parents: 
32645diff
changeset | 347 | val concl = Thm.dest_arg o cprop_of | 
| 29841 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 348 |  fun conjunct1 th = th RS @{thm conjunct1}
 | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 349 |  fun conjunct2 th = th RS @{thm conjunct2}
 | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 350 | fun real_vector_ineq_prover ctxt translator (ges,gts) = | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 351 | let | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 352 | (* val _ = error "real_vector_ineq_prover: pause" *) | 
| 32832 
4602cb6ae0b5
additional fixes in normarith.ML due to FuncFun and FuncUtil changes
 Philipp Meyer parents: 
32645diff
changeset | 353 | val ntms = fold_rev find_normedterms (map (Thm.dest_arg o concl) (ges @ gts)) [] | 
| 29841 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 354 | val lctab = vector_lincombs (map snd (filter (not o fst) ntms)) | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 355 | val (fxns, ctxt') = Variable.variant_fixes (replicate (length lctab) "x") ctxt | 
| 32402 | 356 | fun instantiate_cterm' ty tms = Drule.cterm_rule (Drule.instantiate' ty tms) | 
| 32832 
4602cb6ae0b5
additional fixes in normarith.ML due to FuncFun and FuncUtil changes
 Philipp Meyer parents: 
32645diff
changeset | 357 |   fun mk_norm t = Thm.capply (instantiate_cterm' [SOME (ctyp_of_term t)] [] @{cpat "norm :: (?'a :: real_normed_vector) => real"}) t
 | 
| 
4602cb6ae0b5
additional fixes in normarith.ML due to FuncFun and FuncUtil changes
 Philipp Meyer parents: 
32645diff
changeset | 358 |   fun mk_equals l r = Thm.capply (Thm.capply (instantiate_cterm' [SOME (ctyp_of_term l)] [] @{cpat "op == :: ?'a =>_"}) l) r
 | 
| 29841 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 359 |   val asl = map2 (fn (t,_) => fn n => assume (mk_equals (mk_norm t) (cterm_of (ProofContext.theory_of ctxt') (Free(n,@{typ real}))))) lctab fxns
 | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 360 | val replace_conv = try_conv (rewrs_conv asl) | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 361 | val replace_rule = fconv_rule (funpow 2 arg_conv (replacenegnorms replace_conv)) | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 362 | val ges' = | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 363 | fold_rev (fn th => fn ths => conjunct1(norm_mp th)::ths) | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 364 | asl (map replace_rule ges) | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 365 | val gts' = map replace_rule gts | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 366 | val nubs = map (conjunct2 o norm_mp) asl | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 367 | val th1 = real_vector_combo_prover ctxt' translator (nubs,ges',gts') | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 368 | val shs = filter (member (fn (t,th) => t aconvc cprop_of th) asl) (#hyps (crep_thm th1)) | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 369 | val th11 = hd (Variable.export ctxt' ctxt [fold implies_intr shs th1]) | 
| 32832 
4602cb6ae0b5
additional fixes in normarith.ML due to FuncFun and FuncUtil changes
 Philipp Meyer parents: 
32645diff
changeset | 370 | val cps = map (swap o Thm.dest_equals) (cprems_of th11) | 
| 29841 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 371 | val th12 = instantiate ([], cps) th11 | 
| 32934 | 372 | val th13 = fold Thm.elim_implies (map (reflexive o snd) cps) th12; | 
| 29841 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 373 | in hd (Variable.export ctxt' ctxt [th13]) | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 374 | end | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 375 | in val real_vector_ineq_prover = real_vector_ineq_prover | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 376 | end; | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 377 | |
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 378 | local | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 379 |  val rawrule = fconv_rule (arg_conv (rewr_conv @{thm real_eq_0_iff_le_ge_0}))
 | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 380 |  fun conj_pair th = (th RS @{thm conjunct1}, th RS @{thm conjunct2})
 | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 381 | fun simple_cterm_ord t u = TermOrd.term_ord (term_of t, term_of u) = LESS; | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 382 | (* FIXME: Lookup in the context every time!!! Fix this !!!*) | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 383 | fun splitequation ctxt th acc = | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 384 | let | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 385 | val real_poly_neg_conv = #neg | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 386 | (Normalizer.semiring_normalizers_ord_wrapper ctxt | 
| 33035 | 387 |         (the (NormalizerData.match ctxt @{cterm "(0::real) + 1"})) simple_cterm_ord)
 | 
| 29841 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 388 | val (th1,th2) = conj_pair(rawrule th) | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 389 | in th1::fconv_rule (arg_conv (arg_conv real_poly_neg_conv)) th2::acc | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 390 | end | 
| 32645 
1cc5b24f5a01
sos method generates and uses proof certificates
 Philipp Meyer parents: 
32402diff
changeset | 391 | in fun real_vector_prover ctxt _ translator (eqs,ges,gts) = | 
| 
1cc5b24f5a01
sos method generates and uses proof certificates
 Philipp Meyer parents: 
32402diff
changeset | 392 | (real_vector_ineq_prover ctxt translator | 
| 
1cc5b24f5a01
sos method generates and uses proof certificates
 Philipp Meyer parents: 
32402diff
changeset | 393 | (fold_rev (splitequation ctxt) eqs ges,gts), RealArith.Trivial) | 
| 29841 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 394 | end; | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 395 | |
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 396 | fun init_conv ctxt = | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 397 | Simplifier.rewrite (Simplifier.context ctxt | 
| 31445 
c8a474a919a7
generalize norm method to work over class real_normed_vector
 huffman parents: 
31344diff
changeset | 398 |      (HOL_basic_ss addsimps ([(*@{thm vec_0}, @{thm vec_1},*) @{thm vector_dist_norm}, @{thm diff_0_right}, @{thm right_minus}, @{thm diff_self}, @{thm norm_zero}] @ @{thms arithmetic_simps} @ @{thms norm_pths})))
 | 
| 29841 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 399 | then_conv field_comp_conv | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 400 | then_conv nnf_conv | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 401 | |
| 32645 
1cc5b24f5a01
sos method generates and uses proof certificates
 Philipp Meyer parents: 
32402diff
changeset | 402 | fun pure ctxt = fst o RealArith.gen_prover_real_arith ctxt (real_vector_prover ctxt); | 
| 29841 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 403 | fun norm_arith ctxt ct = | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 404 | let | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 405 | val ctxt' = Variable.declare_term (term_of ct) ctxt | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 406 | val th = init_conv ctxt' ct | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 407 |   in equal_elim (Drule.arg_cong_rule @{cterm Trueprop} (symmetric th)) 
 | 
| 32832 
4602cb6ae0b5
additional fixes in normarith.ML due to FuncFun and FuncUtil changes
 Philipp Meyer parents: 
32645diff
changeset | 408 | (pure ctxt' (Thm.rhs_of th)) | 
| 29841 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 409 | end | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 410 | |
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 411 | fun norm_arith_tac ctxt = | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 412 | clarify_tac HOL_cs THEN' | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 413 | ObjectLogic.full_atomize_tac THEN' | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 414 | CSUBGOAL ( fn (p,i) => rtac (norm_arith ctxt (Thm.dest_arg p )) i); | 
| 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: diff
changeset | 415 | |
| 31445 
c8a474a919a7
generalize norm method to work over class real_normed_vector
 huffman parents: 
31344diff
changeset | 416 | end; |