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