| author | haftmann | 
| Thu, 30 Sep 2010 08:50:45 +0200 | |
| changeset 39794 | 51451d73c3d4 | 
| parent 39027 | e4262f9e6a4e | 
| child 39920 | 7479334d2c90 | 
| permissions | -rw-r--r-- | 
| 33443 | 1  | 
(* Title: HOL/Library/positivstellensatz.ML  | 
2  | 
Author: Amine Chaieb, University of Cambridge  | 
|
3  | 
||
4  | 
A generic arithmetic prover based on Positivstellensatz certificates  | 
|
5  | 
--- also implements Fourrier-Motzkin elimination as a special case  | 
|
6  | 
Fourrier-Motzkin elimination.  | 
|
| 
31120
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
7  | 
*)  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
8  | 
|
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
9  | 
(* A functor for finite mappings based on Tables *)  | 
| 
32645
 
1cc5b24f5a01
sos method generates and uses proof certificates
 
Philipp Meyer 
parents: 
32402 
diff
changeset
 | 
10  | 
|
| 
31120
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
11  | 
signature FUNC =  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
12  | 
sig  | 
| 
32829
 
671eb46eb0a3
tuned FuncFun and FuncUtil structure in positivstellensatz.ML
 
Philipp Meyer 
parents: 
32828 
diff
changeset
 | 
13  | 
include TABLE  | 
| 
 
671eb46eb0a3
tuned FuncFun and FuncUtil structure in positivstellensatz.ML
 
Philipp Meyer 
parents: 
32828 
diff
changeset
 | 
14  | 
val apply : 'a table -> key -> 'a  | 
| 
 
671eb46eb0a3
tuned FuncFun and FuncUtil structure in positivstellensatz.ML
 
Philipp Meyer 
parents: 
32828 
diff
changeset
 | 
15  | 
val applyd :'a table -> (key -> 'a) -> key -> 'a  | 
| 
 
671eb46eb0a3
tuned FuncFun and FuncUtil structure in positivstellensatz.ML
 
Philipp Meyer 
parents: 
32828 
diff
changeset
 | 
16  | 
 val combine : ('a -> 'a -> 'a) -> ('a -> bool) -> 'a table -> 'a table -> 'a table
 | 
| 
 
671eb46eb0a3
tuned FuncFun and FuncUtil structure in positivstellensatz.ML
 
Philipp Meyer 
parents: 
32828 
diff
changeset
 | 
17  | 
val dom : 'a table -> key list  | 
| 
 
671eb46eb0a3
tuned FuncFun and FuncUtil structure in positivstellensatz.ML
 
Philipp Meyer 
parents: 
32828 
diff
changeset
 | 
18  | 
val tryapplyd : 'a table -> key -> 'a -> 'a  | 
| 
 
671eb46eb0a3
tuned FuncFun and FuncUtil structure in positivstellensatz.ML
 
Philipp Meyer 
parents: 
32828 
diff
changeset
 | 
19  | 
val updatep : (key * 'a -> bool) -> key * 'a -> 'a table -> 'a table  | 
| 
 
671eb46eb0a3
tuned FuncFun and FuncUtil structure in positivstellensatz.ML
 
Philipp Meyer 
parents: 
32828 
diff
changeset
 | 
20  | 
val choose : 'a table -> key * 'a  | 
| 
 
671eb46eb0a3
tuned FuncFun and FuncUtil structure in positivstellensatz.ML
 
Philipp Meyer 
parents: 
32828 
diff
changeset
 | 
21  | 
val onefunc : key * 'a -> 'a table  | 
| 
31120
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
22  | 
end;  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
23  | 
|
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
24  | 
functor FuncFun(Key: KEY) : FUNC=  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
25  | 
struct  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
26  | 
|
| 
31971
 
8c1b845ed105
renamed functor TableFun to Table, and GraphFun to Graph;
 
wenzelm 
parents: 
31120 
diff
changeset
 | 
27  | 
structure Tab = Table(Key);  | 
| 
31120
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
28  | 
|
| 
32829
 
671eb46eb0a3
tuned FuncFun and FuncUtil structure in positivstellensatz.ML
 
Philipp Meyer 
parents: 
32828 
diff
changeset
 | 
29  | 
open Tab;  | 
| 
 
671eb46eb0a3
tuned FuncFun and FuncUtil structure in positivstellensatz.ML
 
Philipp Meyer 
parents: 
32828 
diff
changeset
 | 
30  | 
|
| 
31120
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
31  | 
fun dom a = sort Key.ord (Tab.keys a);  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
32  | 
fun applyd f d x = case Tab.lookup f x of  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
33  | 
SOME y => y  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
34  | 
| NONE => d x;  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
35  | 
|
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
36  | 
fun apply f x = applyd f (fn _ => raise Tab.UNDEF x) x;  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
37  | 
fun tryapplyd f a d = applyd f (K d) a;  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
38  | 
fun updatep p (k,v) t = if p (k, v) then t else update (k,v) t  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
39  | 
fun combine f z a b =  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
40  | 
let  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
41  | 
fun h (k,v) t = case Tab.lookup t k of  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
42  | 
NONE => Tab.update (k,v) t  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
43  | 
| SOME v' => let val w = f v v'  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
44  | 
in if z w then Tab.delete k t else Tab.update (k,w) t end;  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
45  | 
in Tab.fold h a b end;  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
46  | 
|
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
47  | 
fun choose f = case Tab.min_key f of  | 
| 33035 | 48  | 
SOME k => (k, the (Tab.lookup f k))  | 
| 
32829
 
671eb46eb0a3
tuned FuncFun and FuncUtil structure in positivstellensatz.ML
 
Philipp Meyer 
parents: 
32828 
diff
changeset
 | 
49  | 
| NONE => error "FuncFun.choose : Completely empty function"  | 
| 
31120
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
50  | 
|
| 
32829
 
671eb46eb0a3
tuned FuncFun and FuncUtil structure in positivstellensatz.ML
 
Philipp Meyer 
parents: 
32828 
diff
changeset
 | 
51  | 
fun onefunc kv = update kv empty  | 
| 
 
671eb46eb0a3
tuned FuncFun and FuncUtil structure in positivstellensatz.ML
 
Philipp Meyer 
parents: 
32828 
diff
changeset
 | 
52  | 
|
| 
31120
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
53  | 
end;  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
54  | 
|
| 
32645
 
1cc5b24f5a01
sos method generates and uses proof certificates
 
Philipp Meyer 
parents: 
32402 
diff
changeset
 | 
55  | 
(* Some standard functors and utility functions for them *)  | 
| 
 
1cc5b24f5a01
sos method generates and uses proof certificates
 
Philipp Meyer 
parents: 
32402 
diff
changeset
 | 
56  | 
|
| 
 
1cc5b24f5a01
sos method generates and uses proof certificates
 
Philipp Meyer 
parents: 
32402 
diff
changeset
 | 
57  | 
structure FuncUtil =  | 
| 
 
1cc5b24f5a01
sos method generates and uses proof certificates
 
Philipp Meyer 
parents: 
32402 
diff
changeset
 | 
58  | 
struct  | 
| 
 
1cc5b24f5a01
sos method generates and uses proof certificates
 
Philipp Meyer 
parents: 
32402 
diff
changeset
 | 
59  | 
|
| 
31120
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
60  | 
structure Intfunc = FuncFun(type key = int val ord = int_ord);  | 
| 
32645
 
1cc5b24f5a01
sos method generates and uses proof certificates
 
Philipp Meyer 
parents: 
32402 
diff
changeset
 | 
61  | 
structure Ratfunc = FuncFun(type key = Rat.rat val ord = Rat.ord);  | 
| 
 
1cc5b24f5a01
sos method generates and uses proof certificates
 
Philipp Meyer 
parents: 
32402 
diff
changeset
 | 
62  | 
structure Intpairfunc = FuncFun(type key = int*int val ord = prod_ord int_ord int_ord);  | 
| 
31120
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
63  | 
structure Symfunc = FuncFun(type key = string val ord = fast_string_ord);  | 
| 35408 | 64  | 
structure Termfunc = FuncFun(type key = term val ord = Term_Ord.fast_term_ord);  | 
| 
32645
 
1cc5b24f5a01
sos method generates and uses proof certificates
 
Philipp Meyer 
parents: 
32402 
diff
changeset
 | 
65  | 
|
| 35408 | 66  | 
val cterm_ord = Term_Ord.fast_term_ord o pairself term_of  | 
| 
32645
 
1cc5b24f5a01
sos method generates and uses proof certificates
 
Philipp Meyer 
parents: 
32402 
diff
changeset
 | 
67  | 
|
| 
 
1cc5b24f5a01
sos method generates and uses proof certificates
 
Philipp Meyer 
parents: 
32402 
diff
changeset
 | 
68  | 
structure Ctermfunc = FuncFun(type key = cterm val ord = cterm_ord);  | 
| 
 
1cc5b24f5a01
sos method generates and uses proof certificates
 
Philipp Meyer 
parents: 
32402 
diff
changeset
 | 
69  | 
|
| 
32829
 
671eb46eb0a3
tuned FuncFun and FuncUtil structure in positivstellensatz.ML
 
Philipp Meyer 
parents: 
32828 
diff
changeset
 | 
70  | 
type monomial = int Ctermfunc.table;  | 
| 
31120
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
71  | 
|
| 
32829
 
671eb46eb0a3
tuned FuncFun and FuncUtil structure in positivstellensatz.ML
 
Philipp Meyer 
parents: 
32828 
diff
changeset
 | 
72  | 
val monomial_ord = list_ord (prod_ord cterm_ord int_ord) o pairself Ctermfunc.dest  | 
| 
32645
 
1cc5b24f5a01
sos method generates and uses proof certificates
 
Philipp Meyer 
parents: 
32402 
diff
changeset
 | 
73  | 
|
| 
 
1cc5b24f5a01
sos method generates and uses proof certificates
 
Philipp Meyer 
parents: 
32402 
diff
changeset
 | 
74  | 
structure Monomialfunc = FuncFun(type key = monomial val ord = monomial_ord)  | 
| 
31120
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
75  | 
|
| 
32829
 
671eb46eb0a3
tuned FuncFun and FuncUtil structure in positivstellensatz.ML
 
Philipp Meyer 
parents: 
32828 
diff
changeset
 | 
76  | 
type poly = Rat.rat Monomialfunc.table;  | 
| 
32645
 
1cc5b24f5a01
sos method generates and uses proof certificates
 
Philipp Meyer 
parents: 
32402 
diff
changeset
 | 
77  | 
|
| 
 
1cc5b24f5a01
sos method generates and uses proof certificates
 
Philipp Meyer 
parents: 
32402 
diff
changeset
 | 
78  | 
(* The ordering so we can create canonical HOL polynomials. *)  | 
| 
31120
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
79  | 
|
| 
32829
 
671eb46eb0a3
tuned FuncFun and FuncUtil structure in positivstellensatz.ML
 
Philipp Meyer 
parents: 
32828 
diff
changeset
 | 
80  | 
fun dest_monomial mon = sort (cterm_ord o pairself fst) (Ctermfunc.dest mon);  | 
| 
31120
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
81  | 
|
| 
32645
 
1cc5b24f5a01
sos method generates and uses proof certificates
 
Philipp Meyer 
parents: 
32402 
diff
changeset
 | 
82  | 
fun monomial_order (m1,m2) =  | 
| 
32829
 
671eb46eb0a3
tuned FuncFun and FuncUtil structure in positivstellensatz.ML
 
Philipp Meyer 
parents: 
32828 
diff
changeset
 | 
83  | 
if Ctermfunc.is_empty m2 then LESS  | 
| 
 
671eb46eb0a3
tuned FuncFun and FuncUtil structure in positivstellensatz.ML
 
Philipp Meyer 
parents: 
32828 
diff
changeset
 | 
84  | 
else if Ctermfunc.is_empty m1 then GREATER  | 
| 
32645
 
1cc5b24f5a01
sos method generates and uses proof certificates
 
Philipp Meyer 
parents: 
32402 
diff
changeset
 | 
85  | 
else  | 
| 
 
1cc5b24f5a01
sos method generates and uses proof certificates
 
Philipp Meyer 
parents: 
32402 
diff
changeset
 | 
86  | 
let val mon1 = dest_monomial m1  | 
| 
 
1cc5b24f5a01
sos method generates and uses proof certificates
 
Philipp Meyer 
parents: 
32402 
diff
changeset
 | 
87  | 
val mon2 = dest_monomial m2  | 
| 33002 | 88  | 
val deg1 = fold (Integer.add o snd) mon1 0  | 
89  | 
val deg2 = fold (Integer.add o snd) mon2 0  | 
|
| 
32645
 
1cc5b24f5a01
sos method generates and uses proof certificates
 
Philipp Meyer 
parents: 
32402 
diff
changeset
 | 
90  | 
in if deg1 < deg2 then GREATER else if deg1 > deg2 then LESS  | 
| 
 
1cc5b24f5a01
sos method generates and uses proof certificates
 
Philipp Meyer 
parents: 
32402 
diff
changeset
 | 
91  | 
else list_ord (prod_ord cterm_ord int_ord) (mon1,mon2)  | 
| 
 
1cc5b24f5a01
sos method generates and uses proof certificates
 
Philipp Meyer 
parents: 
32402 
diff
changeset
 | 
92  | 
end;  | 
| 
31120
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
93  | 
|
| 
32645
 
1cc5b24f5a01
sos method generates and uses proof certificates
 
Philipp Meyer 
parents: 
32402 
diff
changeset
 | 
94  | 
end  | 
| 
31120
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
95  | 
|
| 
32645
 
1cc5b24f5a01
sos method generates and uses proof certificates
 
Philipp Meyer 
parents: 
32402 
diff
changeset
 | 
96  | 
(* positivstellensatz datatype and prover generation *)  | 
| 
31120
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
97  | 
|
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
98  | 
signature REAL_ARITH =  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
99  | 
sig  | 
| 
32645
 
1cc5b24f5a01
sos method generates and uses proof certificates
 
Philipp Meyer 
parents: 
32402 
diff
changeset
 | 
100  | 
|
| 
31120
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
101  | 
datatype positivstellensatz =  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
102  | 
Axiom_eq of int  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
103  | 
| Axiom_le of int  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
104  | 
| Axiom_lt of int  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
105  | 
| Rational_eq of Rat.rat  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
106  | 
| Rational_le of Rat.rat  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
107  | 
| Rational_lt of Rat.rat  | 
| 
32645
 
1cc5b24f5a01
sos method generates and uses proof certificates
 
Philipp Meyer 
parents: 
32402 
diff
changeset
 | 
108  | 
| Square of FuncUtil.poly  | 
| 
 
1cc5b24f5a01
sos method generates and uses proof certificates
 
Philipp Meyer 
parents: 
32402 
diff
changeset
 | 
109  | 
| Eqmul of FuncUtil.poly * positivstellensatz  | 
| 
31120
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
110  | 
| Sum of positivstellensatz * positivstellensatz  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
111  | 
| Product of positivstellensatz * positivstellensatz;  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
112  | 
|
| 
32645
 
1cc5b24f5a01
sos method generates and uses proof certificates
 
Philipp Meyer 
parents: 
32402 
diff
changeset
 | 
113  | 
datatype pss_tree = Trivial | Cert of positivstellensatz | Branch of pss_tree * pss_tree  | 
| 
 
1cc5b24f5a01
sos method generates and uses proof certificates
 
Philipp Meyer 
parents: 
32402 
diff
changeset
 | 
114  | 
|
| 
 
1cc5b24f5a01
sos method generates and uses proof certificates
 
Philipp Meyer 
parents: 
32402 
diff
changeset
 | 
115  | 
datatype tree_choice = Left | Right  | 
| 
 
1cc5b24f5a01
sos method generates and uses proof certificates
 
Philipp Meyer 
parents: 
32402 
diff
changeset
 | 
116  | 
|
| 
 
1cc5b24f5a01
sos method generates and uses proof certificates
 
Philipp Meyer 
parents: 
32402 
diff
changeset
 | 
117  | 
type prover = tree_choice list ->  | 
| 
 
1cc5b24f5a01
sos method generates and uses proof certificates
 
Philipp Meyer 
parents: 
32402 
diff
changeset
 | 
118  | 
(thm list * thm list * thm list -> positivstellensatz -> thm) ->  | 
| 
 
1cc5b24f5a01
sos method generates and uses proof certificates
 
Philipp Meyer 
parents: 
32402 
diff
changeset
 | 
119  | 
thm list * thm list * thm list -> thm * pss_tree  | 
| 
 
1cc5b24f5a01
sos method generates and uses proof certificates
 
Philipp Meyer 
parents: 
32402 
diff
changeset
 | 
120  | 
type cert_conv = cterm -> thm * pss_tree  | 
| 
 
1cc5b24f5a01
sos method generates and uses proof certificates
 
Philipp Meyer 
parents: 
32402 
diff
changeset
 | 
121  | 
|
| 
31120
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
122  | 
val gen_gen_real_arith :  | 
| 
32646
 
962b4354ed90
used standard fold function and type aliases
 
Philipp Meyer 
parents: 
32645 
diff
changeset
 | 
123  | 
Proof.context -> (Rat.rat -> cterm) * conv * conv * conv *  | 
| 
32645
 
1cc5b24f5a01
sos method generates and uses proof certificates
 
Philipp Meyer 
parents: 
32402 
diff
changeset
 | 
124  | 
conv * conv * conv * conv * conv * conv * prover -> cert_conv  | 
| 
 
1cc5b24f5a01
sos method generates and uses proof certificates
 
Philipp Meyer 
parents: 
32402 
diff
changeset
 | 
125  | 
val real_linear_prover : (thm list * thm list * thm list -> positivstellensatz -> thm) ->  | 
| 
 
1cc5b24f5a01
sos method generates and uses proof certificates
 
Philipp Meyer 
parents: 
32402 
diff
changeset
 | 
126  | 
thm list * thm list * thm list -> thm * pss_tree  | 
| 
31120
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
127  | 
|
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
128  | 
val gen_real_arith : Proof.context ->  | 
| 
32646
 
962b4354ed90
used standard fold function and type aliases
 
Philipp Meyer 
parents: 
32645 
diff
changeset
 | 
129  | 
(Rat.rat -> cterm) * conv * conv * conv * conv * conv * conv * conv * prover -> cert_conv  | 
| 
32645
 
1cc5b24f5a01
sos method generates and uses proof certificates
 
Philipp Meyer 
parents: 
32402 
diff
changeset
 | 
130  | 
|
| 
 
1cc5b24f5a01
sos method generates and uses proof certificates
 
Philipp Meyer 
parents: 
32402 
diff
changeset
 | 
131  | 
val gen_prover_real_arith : Proof.context -> prover -> cert_conv  | 
| 
 
1cc5b24f5a01
sos method generates and uses proof certificates
 
Philipp Meyer 
parents: 
32402 
diff
changeset
 | 
132  | 
|
| 
32646
 
962b4354ed90
used standard fold function and type aliases
 
Philipp Meyer 
parents: 
32645 
diff
changeset
 | 
133  | 
val is_ratconst : cterm -> bool  | 
| 
 
962b4354ed90
used standard fold function and type aliases
 
Philipp Meyer 
parents: 
32645 
diff
changeset
 | 
134  | 
val dest_ratconst : cterm -> Rat.rat  | 
| 
 
962b4354ed90
used standard fold function and type aliases
 
Philipp Meyer 
parents: 
32645 
diff
changeset
 | 
135  | 
val cterm_of_rat : Rat.rat -> cterm  | 
| 
32645
 
1cc5b24f5a01
sos method generates and uses proof certificates
 
Philipp Meyer 
parents: 
32402 
diff
changeset
 | 
136  | 
|
| 
31120
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
137  | 
end  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
138  | 
|
| 
32645
 
1cc5b24f5a01
sos method generates and uses proof certificates
 
Philipp Meyer 
parents: 
32402 
diff
changeset
 | 
139  | 
structure RealArith : REAL_ARITH =  | 
| 
31120
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
140  | 
struct  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
141  | 
|
| 32828 | 142  | 
open Conv  | 
| 
31120
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
143  | 
(* ------------------------------------------------------------------------- *)  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
144  | 
(* Data structure for Positivstellensatz refutations. *)  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
145  | 
(* ------------------------------------------------------------------------- *)  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
146  | 
|
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
147  | 
datatype positivstellensatz =  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
148  | 
Axiom_eq of int  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
149  | 
| Axiom_le of int  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
150  | 
| Axiom_lt of int  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
151  | 
| Rational_eq of Rat.rat  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
152  | 
| Rational_le of Rat.rat  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
153  | 
| Rational_lt of Rat.rat  | 
| 
32645
 
1cc5b24f5a01
sos method generates and uses proof certificates
 
Philipp Meyer 
parents: 
32402 
diff
changeset
 | 
154  | 
| Square of FuncUtil.poly  | 
| 
 
1cc5b24f5a01
sos method generates and uses proof certificates
 
Philipp Meyer 
parents: 
32402 
diff
changeset
 | 
155  | 
| Eqmul of FuncUtil.poly * positivstellensatz  | 
| 
31120
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
156  | 
| Sum of positivstellensatz * positivstellensatz  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
157  | 
| Product of positivstellensatz * positivstellensatz;  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
158  | 
(* Theorems used in the procedure *)  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
159  | 
|
| 
32645
 
1cc5b24f5a01
sos method generates and uses proof certificates
 
Philipp Meyer 
parents: 
32402 
diff
changeset
 | 
160  | 
datatype pss_tree = Trivial | Cert of positivstellensatz | Branch of pss_tree * pss_tree  | 
| 
 
1cc5b24f5a01
sos method generates and uses proof certificates
 
Philipp Meyer 
parents: 
32402 
diff
changeset
 | 
161  | 
datatype tree_choice = Left | Right  | 
| 
 
1cc5b24f5a01
sos method generates and uses proof certificates
 
Philipp Meyer 
parents: 
32402 
diff
changeset
 | 
162  | 
type prover = tree_choice list ->  | 
| 
 
1cc5b24f5a01
sos method generates and uses proof certificates
 
Philipp Meyer 
parents: 
32402 
diff
changeset
 | 
163  | 
(thm list * thm list * thm list -> positivstellensatz -> thm) ->  | 
| 
 
1cc5b24f5a01
sos method generates and uses proof certificates
 
Philipp Meyer 
parents: 
32402 
diff
changeset
 | 
164  | 
thm list * thm list * thm list -> thm * pss_tree  | 
| 
 
1cc5b24f5a01
sos method generates and uses proof certificates
 
Philipp Meyer 
parents: 
32402 
diff
changeset
 | 
165  | 
type cert_conv = cterm -> thm * pss_tree  | 
| 
31120
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
166  | 
|
| 
32645
 
1cc5b24f5a01
sos method generates and uses proof certificates
 
Philipp Meyer 
parents: 
32402 
diff
changeset
 | 
167  | 
|
| 
 
1cc5b24f5a01
sos method generates and uses proof certificates
 
Philipp Meyer 
parents: 
32402 
diff
changeset
 | 
168  | 
(* Some useful derived rules *)  | 
| 
 
1cc5b24f5a01
sos method generates and uses proof certificates
 
Philipp Meyer 
parents: 
32402 
diff
changeset
 | 
169  | 
fun deduct_antisym_rule tha thb =  | 
| 36945 | 170  | 
Thm.equal_intr (Thm.implies_intr (cprop_of thb) tha)  | 
171  | 
(Thm.implies_intr (cprop_of tha) thb);  | 
|
| 
32645
 
1cc5b24f5a01
sos method generates and uses proof certificates
 
Philipp Meyer 
parents: 
32402 
diff
changeset
 | 
172  | 
|
| 
 
1cc5b24f5a01
sos method generates and uses proof certificates
 
Philipp Meyer 
parents: 
32402 
diff
changeset
 | 
173  | 
fun prove_hyp tha thb =  | 
| 
 
1cc5b24f5a01
sos method generates and uses proof certificates
 
Philipp Meyer 
parents: 
32402 
diff
changeset
 | 
174  | 
if exists (curry op aconv (concl_of tha)) (#hyps (rep_thm thb))  | 
| 36945 | 175  | 
then Thm.equal_elim (Thm.symmetric (deduct_antisym_rule tha thb)) tha else thb;  | 
| 
32645
 
1cc5b24f5a01
sos method generates and uses proof certificates
 
Philipp Meyer 
parents: 
32402 
diff
changeset
 | 
176  | 
|
| 33443 | 177  | 
val pth = @{lemma "(((x::real) < y) == (y - x > 0))" and "((x <= y) == (y - x >= 0))" and
 | 
178  | 
"((x = y) == (x - y = 0))" and "((~(x < y)) == (x - y >= 0))" and  | 
|
179  | 
"((~(x <= y)) == (x - y > 0))" and "((~(x = y)) == (x - y > 0 | -(x - y) > 0))"  | 
|
180  | 
by (atomize (full), auto simp add: less_diff_eq le_diff_eq not_less)};  | 
|
| 
31120
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
181  | 
|
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
182  | 
val pth_final = @{lemma "(~p ==> False) ==> p" by blast}
 | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
183  | 
val pth_add =  | 
| 33443 | 184  | 
  @{lemma "(x = (0::real) ==> y = 0 ==> x + y = 0 )" and "( x = 0 ==> y >= 0 ==> x + y >= 0)" and
 | 
185  | 
"(x = 0 ==> y > 0 ==> x + y > 0)" and "(x >= 0 ==> y = 0 ==> x + y >= 0)" and  | 
|
186  | 
"(x >= 0 ==> y >= 0 ==> x + y >= 0)" and "(x >= 0 ==> y > 0 ==> x + y > 0)" and  | 
|
187  | 
"(x > 0 ==> y = 0 ==> x + y > 0)" and "(x > 0 ==> y >= 0 ==> x + y > 0)" and  | 
|
188  | 
"(x > 0 ==> y > 0 ==> x + y > 0)" by simp_all};  | 
|
| 
31120
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
189  | 
|
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
190  | 
val pth_mul =  | 
| 33443 | 191  | 
  @{lemma "(x = (0::real) ==> y = 0 ==> x * y = 0)" and "(x = 0 ==> y >= 0 ==> x * y = 0)" and
 | 
192  | 
"(x = 0 ==> y > 0 ==> x * y = 0)" and "(x >= 0 ==> y = 0 ==> x * y = 0)" and  | 
|
193  | 
"(x >= 0 ==> y >= 0 ==> x * y >= 0)" and "(x >= 0 ==> y > 0 ==> x * y >= 0)" and  | 
|
194  | 
"(x > 0 ==> y = 0 ==> x * y = 0)" and "(x > 0 ==> y >= 0 ==> x * y >= 0)" and  | 
|
195  | 
"(x > 0 ==> y > 0 ==> x * y > 0)"  | 
|
| 
31120
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
196  | 
by (auto intro: mult_mono[where a="0::real" and b="x" and d="y" and c="0", simplified]  | 
| 33443 | 197  | 
mult_strict_mono[where b="x" and d="y" and a="0" and c="0", simplified])};  | 
| 
31120
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
198  | 
|
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
199  | 
val pth_emul = @{lemma "y = (0::real) ==> x * y = 0"  by simp};
 | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
200  | 
val pth_square = @{lemma "x * x >= (0::real)"  by simp};
 | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
201  | 
|
| 33443 | 202  | 
val weak_dnf_simps =  | 
203  | 
List.take (simp_thms, 34) @  | 
|
204  | 
    @{lemma "((P & (Q | R)) = ((P&Q) | (P&R)))" and "((Q | R) & P) = ((Q&P) | (R&P))" and
 | 
|
205  | 
"(P & Q) = (Q & P)" and "((P | Q) = (Q | P))" by blast+};  | 
|
| 
31120
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
206  | 
|
| 33443 | 207  | 
val nnfD_simps =  | 
208  | 
  @{lemma "((~(P & Q)) = (~P | ~Q))" and "((~(P | Q)) = (~P & ~Q) )" and
 | 
|
209  | 
"((P --> Q) = (~P | Q) )" and "((P = Q) = ((P & Q) | (~P & ~ Q)))" and  | 
|
210  | 
"((~(P = Q)) = ((P & ~ Q) | (~P & Q)) )" and "((~ ~(P)) = P)" by blast+};  | 
|
| 
31120
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
211  | 
|
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
212  | 
val choice_iff = @{lemma "(ALL x. EX y. P x y) = (EX f. ALL x. P x (f x))" by metis};
 | 
| 33443 | 213  | 
val prenex_simps =  | 
214  | 
map (fn th => th RS sym)  | 
|
215  | 
    ([@{thm "all_conj_distrib"}, @{thm "ex_disj_distrib"}] @
 | 
|
| 37598 | 216  | 
      @{thms "HOL.all_simps"(1-4)} @ @{thms "ex_simps"(1-4)});
 | 
| 
31120
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
217  | 
|
| 33443 | 218  | 
val real_abs_thms1 = @{lemma
 | 
219  | 
"((-1 * abs(x::real) >= r) = (-1 * x >= r & 1 * x >= r))" and  | 
|
220  | 
"((-1 * abs(x) + a >= r) = (a + -1 * x >= r & a + 1 * x >= r))" and  | 
|
221  | 
"((a + -1 * abs(x) >= r) = (a + -1 * x >= r & a + 1 * x >= r))" and  | 
|
222  | 
"((a + -1 * abs(x) + b >= r) = (a + -1 * x + b >= r & a + 1 * x + b >= r))" and  | 
|
223  | 
"((a + b + -1 * abs(x) >= r) = (a + b + -1 * x >= r & a + b + 1 * x >= r))" and  | 
|
224  | 
"((a + b + -1 * abs(x) + c >= r) = (a + b + -1 * x + c >= r & a + b + 1 * x + c >= r))" and  | 
|
225  | 
"((-1 * max x y >= r) = (-1 * x >= r & -1 * y >= r))" and  | 
|
226  | 
"((-1 * max x y + a >= r) = (a + -1 * x >= r & a + -1 * y >= r))" and  | 
|
227  | 
"((a + -1 * max x y >= r) = (a + -1 * x >= r & a + -1 * y >= r))" and  | 
|
228  | 
"((a + -1 * max x y + b >= r) = (a + -1 * x + b >= r & a + -1 * y + b >= r))" and  | 
|
229  | 
"((a + b + -1 * max x y >= r) = (a + b + -1 * x >= r & a + b + -1 * y >= r))" and  | 
|
230  | 
"((a + b + -1 * max x y + c >= r) = (a + b + -1 * x + c >= r & a + b + -1 * y + c >= r))" and  | 
|
231  | 
"((1 * min x y >= r) = (1 * x >= r & 1 * y >= r))" and  | 
|
232  | 
"((1 * min x y + a >= r) = (a + 1 * x >= r & a + 1 * y >= r))" and  | 
|
233  | 
"((a + 1 * min x y >= r) = (a + 1 * x >= r & a + 1 * y >= r))" and  | 
|
234  | 
"((a + 1 * min x y + b >= r) = (a + 1 * x + b >= r & a + 1 * y + b >= r))" and  | 
|
235  | 
"((a + b + 1 * min x y >= r) = (a + b + 1 * x >= r & a + b + 1 * y >= r))" and  | 
|
236  | 
"((a + b + 1 * min x y + c >= r) = (a + b + 1 * x + c >= r & a + b + 1 * y + c >= r))" and  | 
|
237  | 
"((min x y >= r) = (x >= r & y >= r))" and  | 
|
238  | 
"((min x y + a >= r) = (a + x >= r & a + y >= r))" and  | 
|
239  | 
"((a + min x y >= r) = (a + x >= r & a + y >= r))" and  | 
|
240  | 
"((a + min x y + b >= r) = (a + x + b >= r & a + y + b >= r))" and  | 
|
241  | 
"((a + b + min x y >= r) = (a + b + x >= r & a + b + y >= r))" and  | 
|
242  | 
"((a + b + min x y + c >= r) = (a + b + x + c >= r & a + b + y + c >= r))" and  | 
|
243  | 
"((-1 * abs(x) > r) = (-1 * x > r & 1 * x > r))" and  | 
|
244  | 
"((-1 * abs(x) + a > r) = (a + -1 * x > r & a + 1 * x > r))" and  | 
|
245  | 
"((a + -1 * abs(x) > r) = (a + -1 * x > r & a + 1 * x > r))" and  | 
|
246  | 
"((a + -1 * abs(x) + b > r) = (a + -1 * x + b > r & a + 1 * x + b > r))" and  | 
|
247  | 
"((a + b + -1 * abs(x) > r) = (a + b + -1 * x > r & a + b + 1 * x > r))" and  | 
|
248  | 
"((a + b + -1 * abs(x) + c > r) = (a + b + -1 * x + c > r & a + b + 1 * x + c > r))" and  | 
|
249  | 
"((-1 * max x y > r) = ((-1 * x > r) & -1 * y > r))" and  | 
|
250  | 
"((-1 * max x y + a > r) = (a + -1 * x > r & a + -1 * y > r))" and  | 
|
251  | 
"((a + -1 * max x y > r) = (a + -1 * x > r & a + -1 * y > r))" and  | 
|
252  | 
"((a + -1 * max x y + b > r) = (a + -1 * x + b > r & a + -1 * y + b > r))" and  | 
|
253  | 
"((a + b + -1 * max x y > r) = (a + b + -1 * x > r & a + b + -1 * y > r))" and  | 
|
254  | 
"((a + b + -1 * max x y + c > r) = (a + b + -1 * x + c > r & a + b + -1 * y + c > r))" and  | 
|
255  | 
"((min x y > r) = (x > r & y > r))" and  | 
|
256  | 
"((min x y + a > r) = (a + x > r & a + y > r))" and  | 
|
257  | 
"((a + min x y > r) = (a + x > r & a + y > r))" and  | 
|
258  | 
"((a + min x y + b > r) = (a + x + b > r & a + y + b > r))" and  | 
|
259  | 
"((a + b + min x y > r) = (a + b + x > r & a + b + y > r))" and  | 
|
260  | 
"((a + b + min x y + c > r) = (a + b + x + c > r & a + b + y + c > r))"  | 
|
| 
31120
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
261  | 
by auto};  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
262  | 
|
| 
35028
 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 
haftmann 
parents: 
33443 
diff
changeset
 | 
263  | 
val abs_split' = @{lemma "P (abs (x::'a::linordered_idom)) == (x >= 0 & P x | x < 0 & P (-x))"
 | 
| 
31120
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
264  | 
by (atomize (full)) (auto split add: abs_split)};  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
265  | 
|
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
266  | 
val max_split = @{lemma "P (max x y) == ((x::'a::linorder) <= y & P y | x > y & P x)"
 | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
267  | 
by (atomize (full)) (cases "x <= y", auto simp add: max_def)};  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
268  | 
|
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
269  | 
val min_split = @{lemma "P (min x y) == ((x::'a::linorder) <= y & P x | x > y & P y)"
 | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
270  | 
by (atomize (full)) (cases "x <= y", auto simp add: min_def)};  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
271  | 
|
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
272  | 
|
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
273  | 
(* Miscalineous *)  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
274  | 
fun literals_conv bops uops cv =  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
275  | 
let fun h t =  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
276  | 
case (term_of t) of  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
277  | 
b$_$_ => if member (op aconv) bops b then binop_conv h t else cv t  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
278  | 
| u$_ => if member (op aconv) uops u then arg_conv h t else cv t  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
279  | 
| _ => cv t  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
280  | 
in h end;  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
281  | 
|
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
282  | 
fun cterm_of_rat x =  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
283  | 
let val (a, b) = Rat.quotient_of_rat x  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
284  | 
in  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
285  | 
 if b = 1 then Numeral.mk_cnumber @{ctyp "real"} a
 | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
286  | 
  else Thm.capply (Thm.capply @{cterm "op / :: real => _"} 
 | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
287  | 
                   (Numeral.mk_cnumber @{ctyp "real"} a))
 | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
288  | 
        (Numeral.mk_cnumber @{ctyp "real"} b)
 | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
289  | 
end;  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
290  | 
|
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
291  | 
fun dest_ratconst t = case term_of t of  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
292  | 
   Const(@{const_name divide}, _)$a$b => Rat.rat_of_quotient(HOLogic.dest_number a |> snd, HOLogic.dest_number b |> snd)
 | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
293  | 
| _ => Rat.rat_of_int (HOLogic.dest_number (term_of t) |> snd)  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
294  | 
fun is_ratconst t = can dest_ratconst t  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
295  | 
|
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
296  | 
fun find_term p t = if p t then t else  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
297  | 
case t of  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
298  | 
a$b => (find_term p a handle TERM _ => find_term p b)  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
299  | 
| Abs (_,_,t') => find_term p t'  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
300  | 
 | _ => raise TERM ("find_term",[t]);
 | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
301  | 
|
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
302  | 
fun find_cterm p t = if p t then t else  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
303  | 
case term_of t of  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
304  | 
a$b => (find_cterm p (Thm.dest_fun t) handle CTERM _ => find_cterm p (Thm.dest_arg t))  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
305  | 
| Abs (_,_,t') => find_cterm p (Thm.dest_abs NONE t |> snd)  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
306  | 
 | _ => raise CTERM ("find_cterm",[t]);
 | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
307  | 
|
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
308  | 
(* Some conversions-related stuff which has been forbidden entrance into Pure/conv.ML*)  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
309  | 
fun instantiate_cterm' ty tms = Drule.cterm_rule (Drule.instantiate' ty tms)  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
310  | 
fun is_comb t = case (term_of t) of _$_ => true | _ => false;  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
311  | 
|
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
312  | 
fun is_binop ct ct' = ct aconvc (Thm.dest_fun (Thm.dest_fun ct'))  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
313  | 
handle CTERM _ => false;  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
314  | 
|
| 
32645
 
1cc5b24f5a01
sos method generates and uses proof certificates
 
Philipp Meyer 
parents: 
32402 
diff
changeset
 | 
315  | 
|
| 
 
1cc5b24f5a01
sos method generates and uses proof certificates
 
Philipp Meyer 
parents: 
32402 
diff
changeset
 | 
316  | 
(* Map back polynomials to HOL. *)  | 
| 
 
1cc5b24f5a01
sos method generates and uses proof certificates
 
Philipp Meyer 
parents: 
32402 
diff
changeset
 | 
317  | 
|
| 32828 | 318  | 
fun cterm_of_varpow x k = if k = 1 then x else Thm.capply (Thm.capply @{cterm "op ^ :: real => _"} x) 
 | 
319  | 
  (Numeral.mk_cnumber @{ctyp nat} k)
 | 
|
| 
32645
 
1cc5b24f5a01
sos method generates and uses proof certificates
 
Philipp Meyer 
parents: 
32402 
diff
changeset
 | 
320  | 
|
| 
 
1cc5b24f5a01
sos method generates and uses proof certificates
 
Philipp Meyer 
parents: 
32402 
diff
changeset
 | 
321  | 
fun cterm_of_monomial m =  | 
| 
32829
 
671eb46eb0a3
tuned FuncFun and FuncUtil structure in positivstellensatz.ML
 
Philipp Meyer 
parents: 
32828 
diff
changeset
 | 
322  | 
 if FuncUtil.Ctermfunc.is_empty m then @{cterm "1::real"} 
 | 
| 
32645
 
1cc5b24f5a01
sos method generates and uses proof certificates
 
Philipp Meyer 
parents: 
32402 
diff
changeset
 | 
323  | 
else  | 
| 
 
1cc5b24f5a01
sos method generates and uses proof certificates
 
Philipp Meyer 
parents: 
32402 
diff
changeset
 | 
324  | 
let  | 
| 32828 | 325  | 
val m' = FuncUtil.dest_monomial m  | 
| 
32645
 
1cc5b24f5a01
sos method generates and uses proof certificates
 
Philipp Meyer 
parents: 
32402 
diff
changeset
 | 
326  | 
val vps = fold_rev (fn (x,k) => cons (cterm_of_varpow x k)) m' []  | 
| 32828 | 327  | 
  in foldr1 (fn (s, t) => Thm.capply (Thm.capply @{cterm "op * :: real => _"} s) t) vps
 | 
| 
32645
 
1cc5b24f5a01
sos method generates and uses proof certificates
 
Philipp Meyer 
parents: 
32402 
diff
changeset
 | 
328  | 
end  | 
| 
 
1cc5b24f5a01
sos method generates and uses proof certificates
 
Philipp Meyer 
parents: 
32402 
diff
changeset
 | 
329  | 
|
| 
32829
 
671eb46eb0a3
tuned FuncFun and FuncUtil structure in positivstellensatz.ML
 
Philipp Meyer 
parents: 
32828 
diff
changeset
 | 
330  | 
fun cterm_of_cmonomial (m,c) = if FuncUtil.Ctermfunc.is_empty m then cterm_of_rat c  | 
| 
32645
 
1cc5b24f5a01
sos method generates and uses proof certificates
 
Philipp Meyer 
parents: 
32402 
diff
changeset
 | 
331  | 
else if c = Rat.one then cterm_of_monomial m  | 
| 32828 | 332  | 
    else Thm.capply (Thm.capply @{cterm "op *::real => _"} (cterm_of_rat c)) (cterm_of_monomial m);
 | 
| 
32645
 
1cc5b24f5a01
sos method generates and uses proof certificates
 
Philipp Meyer 
parents: 
32402 
diff
changeset
 | 
333  | 
|
| 
 
1cc5b24f5a01
sos method generates and uses proof certificates
 
Philipp Meyer 
parents: 
32402 
diff
changeset
 | 
334  | 
fun cterm_of_poly p =  | 
| 
32829
 
671eb46eb0a3
tuned FuncFun and FuncUtil structure in positivstellensatz.ML
 
Philipp Meyer 
parents: 
32828 
diff
changeset
 | 
335  | 
 if FuncUtil.Monomialfunc.is_empty p then @{cterm "0::real"} 
 | 
| 
32645
 
1cc5b24f5a01
sos method generates and uses proof certificates
 
Philipp Meyer 
parents: 
32402 
diff
changeset
 | 
336  | 
else  | 
| 
 
1cc5b24f5a01
sos method generates and uses proof certificates
 
Philipp Meyer 
parents: 
32402 
diff
changeset
 | 
337  | 
let  | 
| 
 
1cc5b24f5a01
sos method generates and uses proof certificates
 
Philipp Meyer 
parents: 
32402 
diff
changeset
 | 
338  | 
val cms = map cterm_of_cmonomial  | 
| 
32829
 
671eb46eb0a3
tuned FuncFun and FuncUtil structure in positivstellensatz.ML
 
Philipp Meyer 
parents: 
32828 
diff
changeset
 | 
339  | 
(sort (prod_ord FuncUtil.monomial_order (K EQUAL)) (FuncUtil.Monomialfunc.dest p))  | 
| 32828 | 340  | 
  in foldr1 (fn (t1, t2) => Thm.capply(Thm.capply @{cterm "op + :: real => _"} t1) t2) cms
 | 
| 
32645
 
1cc5b24f5a01
sos method generates and uses proof certificates
 
Philipp Meyer 
parents: 
32402 
diff
changeset
 | 
341  | 
end;  | 
| 
 
1cc5b24f5a01
sos method generates and uses proof certificates
 
Philipp Meyer 
parents: 
32402 
diff
changeset
 | 
342  | 
|
| 
31120
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
343  | 
(* A general real arithmetic prover *)  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
344  | 
|
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
345  | 
fun gen_gen_real_arith ctxt (mk_numeric,  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
346  | 
numeric_eq_conv,numeric_ge_conv,numeric_gt_conv,  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
347  | 
poly_conv,poly_neg_conv,poly_add_conv,poly_mul_conv,  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
348  | 
absconv1,absconv2,prover) =  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
349  | 
let  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
350  | 
 val pre_ss = HOL_basic_ss addsimps simp_thms@ ex_simps@ all_simps@[@{thm not_all},@{thm not_ex},ex_disj_distrib, all_conj_distrib, @{thm if_bool_eq_disj}]
 | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
351  | 
val prenex_ss = HOL_basic_ss addsimps prenex_simps  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
352  | 
val skolemize_ss = HOL_basic_ss addsimps [choice_iff]  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
353  | 
val presimp_conv = Simplifier.rewrite (Simplifier.context ctxt pre_ss)  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
354  | 
val prenex_conv = Simplifier.rewrite (Simplifier.context ctxt prenex_ss)  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
355  | 
val skolemize_conv = Simplifier.rewrite (Simplifier.context ctxt skolemize_ss)  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
356  | 
val weak_dnf_ss = HOL_basic_ss addsimps weak_dnf_simps  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
357  | 
val weak_dnf_conv = Simplifier.rewrite (Simplifier.context ctxt weak_dnf_ss)  | 
| 36945 | 358  | 
 fun eqT_elim th = Thm.equal_elim (Thm.symmetric th) @{thm TrueI}
 | 
| 
31120
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
359  | 
fun oprconv cv ct =  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
360  | 
let val g = Thm.dest_fun2 ct  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
361  | 
  in if g aconvc @{cterm "op <= :: real => _"} 
 | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
362  | 
       orelse g aconvc @{cterm "op < :: real => _"} 
 | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
363  | 
then arg_conv cv ct else arg1_conv cv ct  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
364  | 
end  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
365  | 
|
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
366  | 
fun real_ineq_conv th ct =  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
367  | 
let  | 
| 32828 | 368  | 
val th' = (Thm.instantiate (Thm.match (Thm.lhs_of th, ct)) th  | 
| 
37117
 
59cee8807c29
eliminated various catch-all exception patterns, guessing at the concrete exeptions that are intended here;
 
wenzelm 
parents: 
36945 
diff
changeset
 | 
369  | 
      handle Pattern.MATCH => raise CTERM ("real_ineq_conv", [ct]))
 | 
| 36945 | 370  | 
in Thm.transitive th' (oprconv poly_conv (Thm.rhs_of th'))  | 
| 
31120
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
371  | 
end  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
372  | 
val [real_lt_conv, real_le_conv, real_eq_conv,  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
373  | 
real_not_lt_conv, real_not_le_conv, _] =  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
374  | 
map real_ineq_conv pth  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
375  | 
fun match_mp_rule ths ths' =  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
376  | 
let  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
377  | 
     fun f ths ths' = case ths of [] => raise THM("match_mp_rule",0,ths)
 | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
378  | 
| th::ths => (ths' MRS th handle THM _ => f ths ths')  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
379  | 
in f ths ths' end  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
380  | 
fun mul_rule th th' = fconv_rule (arg_conv (oprconv poly_mul_conv))  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
381  | 
(match_mp_rule pth_mul [th, th'])  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
382  | 
fun add_rule th th' = fconv_rule (arg_conv (oprconv poly_add_conv))  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
383  | 
(match_mp_rule pth_add [th, th'])  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
384  | 
fun emul_rule ct th = fconv_rule (arg_conv (oprconv poly_mul_conv))  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
385  | 
(instantiate' [] [SOME ct] (th RS pth_emul))  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
386  | 
fun square_rule t = fconv_rule (arg_conv (oprconv poly_conv))  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
387  | 
(instantiate' [] [SOME t] pth_square)  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
388  | 
|
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
389  | 
fun hol_of_positivstellensatz(eqs,les,lts) proof =  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
390  | 
let  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
391  | 
fun translate prf = case prf of  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
392  | 
Axiom_eq n => nth eqs n  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
393  | 
| Axiom_le n => nth les n  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
394  | 
| Axiom_lt n => nth lts n  | 
| 32828 | 395  | 
      | Rational_eq x => eqT_elim(numeric_eq_conv(Thm.capply @{cterm Trueprop} 
 | 
396  | 
                          (Thm.capply (Thm.capply @{cterm "op =::real => _"} (mk_numeric x)) 
 | 
|
| 
31120
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
397  | 
                               @{cterm "0::real"})))
 | 
| 32828 | 398  | 
      | Rational_le x => eqT_elim(numeric_ge_conv(Thm.capply @{cterm Trueprop} 
 | 
399  | 
                          (Thm.capply (Thm.capply @{cterm "op <=::real => _"} 
 | 
|
| 
31120
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
400  | 
                                     @{cterm "0::real"}) (mk_numeric x))))
 | 
| 32828 | 401  | 
      | Rational_lt x => eqT_elim(numeric_gt_conv(Thm.capply @{cterm Trueprop} 
 | 
402  | 
                      (Thm.capply (Thm.capply @{cterm "op <::real => _"} @{cterm "0::real"})
 | 
|
| 
31120
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
403  | 
(mk_numeric x))))  | 
| 
32645
 
1cc5b24f5a01
sos method generates and uses proof certificates
 
Philipp Meyer 
parents: 
32402 
diff
changeset
 | 
404  | 
| Square pt => square_rule (cterm_of_poly pt)  | 
| 
 
1cc5b24f5a01
sos method generates and uses proof certificates
 
Philipp Meyer 
parents: 
32402 
diff
changeset
 | 
405  | 
| Eqmul(pt,p) => emul_rule (cterm_of_poly pt) (translate p)  | 
| 
31120
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
406  | 
| Sum(p1,p2) => add_rule (translate p1) (translate p2)  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
407  | 
| Product(p1,p2) => mul_rule (translate p1) (translate p2)  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
408  | 
in fconv_rule (first_conv [numeric_ge_conv, numeric_gt_conv, numeric_eq_conv, all_conv])  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
409  | 
(translate proof)  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
410  | 
end  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
411  | 
|
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
412  | 
val init_conv = presimp_conv then_conv  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
413  | 
nnf_conv then_conv skolemize_conv then_conv prenex_conv then_conv  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
414  | 
weak_dnf_conv  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
415  | 
|
| 32828 | 416  | 
val concl = Thm.dest_arg o cprop_of  | 
417  | 
fun is_binop opr ct = (Thm.dest_fun2 ct aconvc opr handle CTERM _ => false)  | 
|
| 
31120
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
418  | 
  val is_req = is_binop @{cterm "op =:: real => _"}
 | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
419  | 
  val is_ge = is_binop @{cterm "op <=:: real => _"}
 | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
420  | 
  val is_gt = is_binop @{cterm "op <:: real => _"}
 | 
| 
38795
 
848be46708dc
formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
 
haftmann 
parents: 
38558 
diff
changeset
 | 
421  | 
  val is_conj = is_binop @{cterm HOL.conj}
 | 
| 
 
848be46708dc
formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
 
haftmann 
parents: 
38558 
diff
changeset
 | 
422  | 
  val is_disj = is_binop @{cterm HOL.disj}
 | 
| 
31120
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
423  | 
  fun conj_pair th = (th RS @{thm conjunct1}, th RS @{thm conjunct2})
 | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
424  | 
fun disj_cases th th1 th2 =  | 
| 32828 | 425  | 
let val (p,q) = Thm.dest_binop (concl th)  | 
| 
31120
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
426  | 
val c = concl th1  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
427  | 
val _ = if c aconvc (concl th2) then () else error "disj_cases : conclusions not alpha convertible"  | 
| 36945 | 428  | 
in Thm.implies_elim (Thm.implies_elim  | 
429  | 
          (Thm.implies_elim (instantiate' [] (map SOME [p,q,c]) @{thm disjE}) th)
 | 
|
430  | 
          (Thm.implies_intr (Thm.capply @{cterm Trueprop} p) th1))
 | 
|
431  | 
        (Thm.implies_intr (Thm.capply @{cterm Trueprop} q) th2)
 | 
|
| 
31120
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
432  | 
end  | 
| 
32645
 
1cc5b24f5a01
sos method generates and uses proof certificates
 
Philipp Meyer 
parents: 
32402 
diff
changeset
 | 
433  | 
fun overall cert_choice dun ths = case ths of  | 
| 
31120
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
434  | 
[] =>  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
435  | 
let  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
436  | 
val (eq,ne) = List.partition (is_req o concl) dun  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
437  | 
val (le,nl) = List.partition (is_ge o concl) ne  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
438  | 
val lt = filter (is_gt o concl) nl  | 
| 
32645
 
1cc5b24f5a01
sos method generates and uses proof certificates
 
Philipp Meyer 
parents: 
32402 
diff
changeset
 | 
439  | 
in prover (rev cert_choice) hol_of_positivstellensatz (eq,le,lt) end  | 
| 
31120
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
440  | 
| th::oths =>  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
441  | 
let  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
442  | 
val ct = concl th  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
443  | 
in  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
444  | 
if is_conj ct then  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
445  | 
let  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
446  | 
val (th1,th2) = conj_pair th in  | 
| 
32645
 
1cc5b24f5a01
sos method generates and uses proof certificates
 
Philipp Meyer 
parents: 
32402 
diff
changeset
 | 
447  | 
overall cert_choice dun (th1::th2::oths) end  | 
| 
31120
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
448  | 
else if is_disj ct then  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
449  | 
let  | 
| 36945 | 450  | 
       val (th1, cert1) = overall (Left::cert_choice) dun (Thm.assume (Thm.capply @{cterm Trueprop} (Thm.dest_arg1 ct))::oths)
 | 
451  | 
       val (th2, cert2) = overall (Right::cert_choice) dun (Thm.assume (Thm.capply @{cterm Trueprop} (Thm.dest_arg ct))::oths)
 | 
|
| 
32645
 
1cc5b24f5a01
sos method generates and uses proof certificates
 
Philipp Meyer 
parents: 
32402 
diff
changeset
 | 
452  | 
in (disj_cases th th1 th2, Branch (cert1, cert2)) end  | 
| 
 
1cc5b24f5a01
sos method generates and uses proof certificates
 
Philipp Meyer 
parents: 
32402 
diff
changeset
 | 
453  | 
else overall cert_choice (th::dun) oths  | 
| 
31120
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
454  | 
end  | 
| 32828 | 455  | 
fun dest_binary b ct = if is_binop b ct then Thm.dest_binop ct  | 
| 
31120
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
456  | 
                         else raise CTERM ("dest_binary",[b,ct])
 | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
457  | 
  val dest_eq = dest_binary @{cterm "op = :: real => _"}
 | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
458  | 
val neq_th = nth pth 5  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
459  | 
fun real_not_eq_conv ct =  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
460  | 
let  | 
| 32828 | 461  | 
val (l,r) = dest_eq (Thm.dest_arg ct)  | 
462  | 
    val th = Thm.instantiate ([],[(@{cpat "?x::real"},l),(@{cpat "?y::real"},r)]) neq_th
 | 
|
463  | 
val th_p = poly_conv(Thm.dest_arg(Thm.dest_arg1(Thm.rhs_of th)))  | 
|
| 
31120
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
464  | 
    val th_x = Drule.arg_cong_rule @{cterm "uminus :: real => _"} th_p
 | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
465  | 
val th_n = fconv_rule (arg_conv poly_neg_conv) th_x  | 
| 
38795
 
848be46708dc
formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
 
haftmann 
parents: 
38558 
diff
changeset
 | 
466  | 
    val th' = Drule.binop_cong_rule @{cterm HOL.disj} 
 | 
| 32828 | 467  | 
     (Drule.arg_cong_rule (Thm.capply @{cterm "op <::real=>_"} @{cterm "0::real"}) th_p)
 | 
468  | 
     (Drule.arg_cong_rule (Thm.capply @{cterm "op <::real=>_"} @{cterm "0::real"}) th_n)
 | 
|
| 36945 | 469  | 
in Thm.transitive th th'  | 
| 
31120
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
470  | 
end  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
471  | 
fun equal_implies_1_rule PQ =  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
472  | 
let  | 
| 32828 | 473  | 
val P = Thm.lhs_of PQ  | 
| 36945 | 474  | 
in Thm.implies_intr P (Thm.equal_elim PQ (Thm.assume P))  | 
| 
31120
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
475  | 
end  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
476  | 
(* FIXME!!! Copied from groebner.ml *)  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
477  | 
val strip_exists =  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
478  | 
let fun h (acc, t) =  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
479  | 
case (term_of t) of  | 
| 38558 | 480  | 
    Const(@{const_name Ex},_)$Abs(x,T,p) => h (Thm.dest_abs NONE (Thm.dest_arg t) |>> (fn v => v::acc))
 | 
| 
31120
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
481  | 
| _ => (acc,t)  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
482  | 
in fn t => h ([],t)  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
483  | 
end  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
484  | 
fun name_of x = case term_of x of  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
485  | 
Free(s,_) => s  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
486  | 
| Var ((s,_),_) => s  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
487  | 
| _ => "x"  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
488  | 
|
| 36945 | 489  | 
  fun mk_forall x th = Drule.arg_cong_rule (instantiate_cterm' [SOME (ctyp_of_term x)] [] @{cpat "All :: (?'a => bool) => _" }) (Thm.abstract_rule (name_of x) x th)
 | 
| 
31120
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
490  | 
|
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
491  | 
val specl = fold_rev (fn x => fn th => instantiate' [] [SOME x] (th RS spec));  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
492  | 
|
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
493  | 
 fun ext T = Drule.cterm_rule (instantiate' [SOME T] []) @{cpat Ex}
 | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
494  | 
fun mk_ex v t = Thm.capply (ext (ctyp_of_term v)) (Thm.cabs v t)  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
495  | 
|
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
496  | 
fun choose v th th' = case concl_of th of  | 
| 38558 | 497  | 
   @{term Trueprop} $ (Const(@{const_name Ex},_)$_) => 
 | 
| 
31120
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
498  | 
let  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
499  | 
val p = (funpow 2 Thm.dest_arg o cprop_of) th  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
500  | 
val T = (hd o Thm.dest_ctyp o ctyp_of_term) p  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
501  | 
val th0 = fconv_rule (Thm.beta_conversion true)  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
502  | 
(instantiate' [SOME T] [SOME p, (SOME o Thm.dest_arg o cprop_of) th'] exE)  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
503  | 
val pv = (Thm.rhs_of o Thm.beta_conversion true)  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
504  | 
           (Thm.capply @{cterm Trueprop} (Thm.capply p v))
 | 
| 36945 | 505  | 
val th1 = Thm.forall_intr v (Thm.implies_intr pv th')  | 
506  | 
in Thm.implies_elim (Thm.implies_elim th0 th) th1 end  | 
|
| 
31120
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
507  | 
 | _ => raise THM ("choose",0,[th, th'])
 | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
508  | 
|
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
509  | 
fun simple_choose v th =  | 
| 36945 | 510  | 
     choose v (Thm.assume ((Thm.capply @{cterm Trueprop} o mk_ex v) ((Thm.dest_arg o hd o #hyps o Thm.crep_thm) th))) th
 | 
| 
31120
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
511  | 
|
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
512  | 
val strip_forall =  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
513  | 
let fun h (acc, t) =  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
514  | 
case (term_of t) of  | 
| 38558 | 515  | 
    Const(@{const_name All},_)$Abs(x,T,p) => h (Thm.dest_abs NONE (Thm.dest_arg t) |>> (fn v => v::acc))
 | 
| 
31120
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
516  | 
| _ => (acc,t)  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
517  | 
in fn t => h ([],t)  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
518  | 
end  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
519  | 
|
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
520  | 
fun f ct =  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
521  | 
let  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
522  | 
val nnf_norm_conv' =  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
523  | 
nnf_conv then_conv  | 
| 
38795
 
848be46708dc
formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
 
haftmann 
parents: 
38558 
diff
changeset
 | 
524  | 
     literals_conv [@{term HOL.conj}, @{term HOL.disj}] [] 
 | 
| 
32843
 
c8f5a7c8353f
moved cache_conv to src/Pure/conv.ML, use Thm.cterm_cache;
 
wenzelm 
parents: 
32829 
diff
changeset
 | 
525  | 
(Conv.cache_conv  | 
| 
31120
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
526  | 
(first_conv [real_lt_conv, real_le_conv,  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
527  | 
real_eq_conv, real_not_lt_conv,  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
528  | 
real_not_le_conv, real_not_eq_conv, all_conv]))  | 
| 
38795
 
848be46708dc
formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
 
haftmann 
parents: 
38558 
diff
changeset
 | 
529  | 
  fun absremover ct = (literals_conv [@{term HOL.conj}, @{term HOL.disj}] [] 
 | 
| 
31120
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
530  | 
(try_conv (absconv1 then_conv binop_conv (arg_conv poly_conv))) then_conv  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
531  | 
try_conv (absconv2 then_conv nnf_norm_conv' then_conv binop_conv absremover)) ct  | 
| 32828 | 532  | 
  val nct = Thm.capply @{cterm Trueprop} (Thm.capply @{cterm "Not"} ct)
 | 
| 
31120
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
533  | 
val th0 = (init_conv then_conv arg_conv nnf_norm_conv') nct  | 
| 32828 | 534  | 
val tm0 = Thm.dest_arg (Thm.rhs_of th0)  | 
| 
32645
 
1cc5b24f5a01
sos method generates and uses proof certificates
 
Philipp Meyer 
parents: 
32402 
diff
changeset
 | 
535  | 
  val (th, certificates) = if tm0 aconvc @{cterm False} then (equal_implies_1_rule th0, Trivial) else
 | 
| 
31120
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
536  | 
let  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
537  | 
val (evs,bod) = strip_exists tm0  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
538  | 
val (avs,ibod) = strip_forall bod  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
539  | 
    val th1 = Drule.arg_cong_rule @{cterm Trueprop} (fold mk_forall avs (absremover ibod))
 | 
| 36945 | 540  | 
val (th2, certs) = overall [] [] [specl avs (Thm.assume (Thm.rhs_of th1))]  | 
541  | 
    val th3 = fold simple_choose evs (prove_hyp (Thm.equal_elim th1 (Thm.assume (Thm.capply @{cterm Trueprop} bod))) th2)
 | 
|
542  | 
in (Drule.implies_intr_hyps (prove_hyp (Thm.equal_elim th0 (Thm.assume nct)) th3), certs)  | 
|
| 
31120
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
543  | 
end  | 
| 36945 | 544  | 
in (Thm.implies_elim (instantiate' [] [SOME ct] pth_final) th, certificates)  | 
| 
31120
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
545  | 
end  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
546  | 
in f  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
547  | 
end;  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
548  | 
|
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
549  | 
(* A linear arithmetic prover *)  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
550  | 
local  | 
| 32828 | 551  | 
val linear_add = FuncUtil.Ctermfunc.combine (curry op +/) (fn z => z =/ Rat.zero)  | 
| 39027 | 552  | 
fun linear_cmul c = FuncUtil.Ctermfunc.map (fn _ => fn x => c */ x)  | 
| 
31120
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
553  | 
  val one_tm = @{cterm "1::real"}
 | 
| 
32829
 
671eb46eb0a3
tuned FuncFun and FuncUtil structure in positivstellensatz.ML
 
Philipp Meyer 
parents: 
32828 
diff
changeset
 | 
554  | 
fun contradictory p (e,_) = ((FuncUtil.Ctermfunc.is_empty e) andalso not(p Rat.zero)) orelse  | 
| 33038 | 555  | 
((eq_set (op aconvc) (FuncUtil.Ctermfunc.dom e, [one_tm])) andalso  | 
| 
32829
 
671eb46eb0a3
tuned FuncFun and FuncUtil structure in positivstellensatz.ML
 
Philipp Meyer 
parents: 
32828 
diff
changeset
 | 
556  | 
not(p(FuncUtil.Ctermfunc.apply e one_tm)))  | 
| 
31120
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
557  | 
|
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
558  | 
fun linear_ineqs vars (les,lts) =  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
559  | 
case find_first (contradictory (fn x => x >/ Rat.zero)) lts of  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
560  | 
SOME r => r  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
561  | 
| NONE =>  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
562  | 
(case find_first (contradictory (fn x => x >/ Rat.zero)) les of  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
563  | 
SOME r => r  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
564  | 
| NONE =>  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
565  | 
if null vars then error "linear_ineqs: no contradiction" else  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
566  | 
let  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
567  | 
val ineqs = les @ lts  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
568  | 
fun blowup v =  | 
| 32828 | 569  | 
length(filter (fn (e,_) => FuncUtil.Ctermfunc.tryapplyd e v Rat.zero =/ Rat.zero) ineqs) +  | 
570  | 
length(filter (fn (e,_) => FuncUtil.Ctermfunc.tryapplyd e v Rat.zero >/ Rat.zero) ineqs) *  | 
|
571  | 
length(filter (fn (e,_) => FuncUtil.Ctermfunc.tryapplyd e v Rat.zero </ Rat.zero) ineqs)  | 
|
| 
31120
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
572  | 
val v = fst(hd(sort (fn ((_,i),(_,j)) => int_ord (i,j))  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
573  | 
(map (fn v => (v,blowup v)) vars)))  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
574  | 
fun addup (e1,p1) (e2,p2) acc =  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
575  | 
let  | 
| 32828 | 576  | 
val c1 = FuncUtil.Ctermfunc.tryapplyd e1 v Rat.zero  | 
577  | 
val c2 = FuncUtil.Ctermfunc.tryapplyd e2 v Rat.zero  | 
|
| 
31120
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
578  | 
in if c1 */ c2 >=/ Rat.zero then acc else  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
579  | 
let  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
580  | 
val e1' = linear_cmul (Rat.abs c2) e1  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
581  | 
val e2' = linear_cmul (Rat.abs c1) e2  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
582  | 
val p1' = Product(Rational_lt(Rat.abs c2),p1)  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
583  | 
val p2' = Product(Rational_lt(Rat.abs c1),p2)  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
584  | 
in (linear_add e1' e2',Sum(p1',p2'))::acc  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
585  | 
end  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
586  | 
end  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
587  | 
val (les0,les1) =  | 
| 32828 | 588  | 
List.partition (fn (e,_) => FuncUtil.Ctermfunc.tryapplyd e v Rat.zero =/ Rat.zero) les  | 
| 
31120
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
589  | 
val (lts0,lts1) =  | 
| 32828 | 590  | 
List.partition (fn (e,_) => FuncUtil.Ctermfunc.tryapplyd e v Rat.zero =/ Rat.zero) lts  | 
| 
31120
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
591  | 
val (lesp,lesn) =  | 
| 32828 | 592  | 
List.partition (fn (e,_) => FuncUtil.Ctermfunc.tryapplyd e v Rat.zero >/ Rat.zero) les1  | 
| 
31120
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
593  | 
val (ltsp,ltsn) =  | 
| 32828 | 594  | 
List.partition (fn (e,_) => FuncUtil.Ctermfunc.tryapplyd e v Rat.zero >/ Rat.zero) lts1  | 
| 
31120
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
595  | 
val les' = fold_rev (fn ep1 => fold_rev (addup ep1) lesp) lesn les0  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
596  | 
val lts' = fold_rev (fn ep1 => fold_rev (addup ep1) (lesp@ltsp)) ltsn  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
597  | 
(fold_rev (fn ep1 => fold_rev (addup ep1) (lesn@ltsn)) ltsp lts0)  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
598  | 
in linear_ineqs (remove (op aconvc) v vars) (les',lts')  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
599  | 
end)  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
600  | 
|
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
601  | 
fun linear_eqs(eqs,les,lts) =  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
602  | 
case find_first (contradictory (fn x => x =/ Rat.zero)) eqs of  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
603  | 
SOME r => r  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
604  | 
| NONE => (case eqs of  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
605  | 
[] =>  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
606  | 
let val vars = remove (op aconvc) one_tm  | 
| 33042 | 607  | 
(fold_rev (union (op aconvc) o FuncUtil.Ctermfunc.dom o fst) (les@lts) [])  | 
| 
31120
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
608  | 
in linear_ineqs vars (les,lts) end  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
609  | 
| (e,p)::es =>  | 
| 
32829
 
671eb46eb0a3
tuned FuncFun and FuncUtil structure in positivstellensatz.ML
 
Philipp Meyer 
parents: 
32828 
diff
changeset
 | 
610  | 
if FuncUtil.Ctermfunc.is_empty e then linear_eqs (es,les,lts) else  | 
| 
31120
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
611  | 
let  | 
| 
32829
 
671eb46eb0a3
tuned FuncFun and FuncUtil structure in positivstellensatz.ML
 
Philipp Meyer 
parents: 
32828 
diff
changeset
 | 
612  | 
val (x,c) = FuncUtil.Ctermfunc.choose (FuncUtil.Ctermfunc.delete_safe one_tm e)  | 
| 
31120
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
613  | 
fun xform (inp as (t,q)) =  | 
| 32828 | 614  | 
let val d = FuncUtil.Ctermfunc.tryapplyd t x Rat.zero in  | 
| 
31120
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
615  | 
if d =/ Rat.zero then inp else  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
616  | 
let  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
617  | 
val k = (Rat.neg d) */ Rat.abs c // c  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
618  | 
val e' = linear_cmul k e  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
619  | 
val t' = linear_cmul (Rat.abs c) t  | 
| 
32829
 
671eb46eb0a3
tuned FuncFun and FuncUtil structure in positivstellensatz.ML
 
Philipp Meyer 
parents: 
32828 
diff
changeset
 | 
620  | 
val p' = Eqmul(FuncUtil.Monomialfunc.onefunc (FuncUtil.Ctermfunc.empty, k),p)  | 
| 
31120
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
621  | 
val q' = Product(Rational_lt(Rat.abs c),q)  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
622  | 
in (linear_add e' t',Sum(p',q'))  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
623  | 
end  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
624  | 
end  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
625  | 
in linear_eqs(map xform es,map xform les,map xform lts)  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
626  | 
end)  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
627  | 
|
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
628  | 
fun linear_prover (eq,le,lt) =  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
629  | 
let  | 
| 33063 | 630  | 
val eqs = map_index (fn (n, p) => (p,Axiom_eq n)) eq  | 
631  | 
val les = map_index (fn (n, p) => (p,Axiom_le n)) le  | 
|
632  | 
val lts = map_index (fn (n, p) => (p,Axiom_lt n)) lt  | 
|
| 
31120
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
633  | 
in linear_eqs(eqs,les,lts)  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
634  | 
end  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
635  | 
|
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
636  | 
fun lin_of_hol ct =  | 
| 
32829
 
671eb46eb0a3
tuned FuncFun and FuncUtil structure in positivstellensatz.ML
 
Philipp Meyer 
parents: 
32828 
diff
changeset
 | 
637  | 
   if ct aconvc @{cterm "0::real"} then FuncUtil.Ctermfunc.empty
 | 
| 32828 | 638  | 
else if not (is_comb ct) then FuncUtil.Ctermfunc.onefunc (ct, Rat.one)  | 
639  | 
else if is_ratconst ct then FuncUtil.Ctermfunc.onefunc (one_tm, dest_ratconst ct)  | 
|
| 
31120
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
640  | 
else  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
641  | 
let val (lop,r) = Thm.dest_comb ct  | 
| 32828 | 642  | 
in if not (is_comb lop) then FuncUtil.Ctermfunc.onefunc (ct, Rat.one)  | 
| 
31120
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
643  | 
else  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
644  | 
let val (opr,l) = Thm.dest_comb lop  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
645  | 
        in if opr aconvc @{cterm "op + :: real =>_"} 
 | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
646  | 
then linear_add (lin_of_hol l) (lin_of_hol r)  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
647  | 
           else if opr aconvc @{cterm "op * :: real =>_"} 
 | 
| 32828 | 648  | 
andalso is_ratconst l then FuncUtil.Ctermfunc.onefunc (r, dest_ratconst l)  | 
649  | 
else FuncUtil.Ctermfunc.onefunc (ct, Rat.one)  | 
|
| 
31120
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
650  | 
end  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
651  | 
end  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
652  | 
|
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
653  | 
fun is_alien ct = case term_of ct of  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
654  | 
   Const(@{const_name "real"}, _)$ n => 
 | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
655  | 
if can HOLogic.dest_number n then false else true  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
656  | 
| _ => false  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
657  | 
in  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
658  | 
fun real_linear_prover translator (eq,le,lt) =  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
659  | 
let  | 
| 32828 | 660  | 
val lhs = lin_of_hol o Thm.dest_arg1 o Thm.dest_arg o cprop_of  | 
661  | 
val rhs = lin_of_hol o Thm.dest_arg o Thm.dest_arg o cprop_of  | 
|
| 
31120
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
662  | 
val eq_pols = map lhs eq  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
663  | 
val le_pols = map rhs le  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
664  | 
val lt_pols = map rhs lt  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
665  | 
val aliens = filter is_alien  | 
| 33042 | 666  | 
(fold_rev (union (op aconvc) o FuncUtil.Ctermfunc.dom)  | 
| 
31120
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
667  | 
(eq_pols @ le_pols @ lt_pols) [])  | 
| 32828 | 668  | 
val le_pols' = le_pols @ map (fn v => FuncUtil.Ctermfunc.onefunc (v,Rat.one)) aliens  | 
| 
31120
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
669  | 
val (_,proof) = linear_prover (eq_pols,le_pols',lt_pols)  | 
| 32828 | 670  | 
  val le' = le @ map (fn a => instantiate' [] [SOME (Thm.dest_arg a)] @{thm real_of_nat_ge_zero}) aliens 
 | 
| 
32645
 
1cc5b24f5a01
sos method generates and uses proof certificates
 
Philipp Meyer 
parents: 
32402 
diff
changeset
 | 
671  | 
in ((translator (eq,le',lt) proof), Trivial)  | 
| 
31120
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
672  | 
end  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
673  | 
end;  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
674  | 
|
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
675  | 
(* A less general generic arithmetic prover dealing with abs,max and min*)  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
676  | 
|
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
677  | 
local  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
678  | 
val absmaxmin_elim_ss1 = HOL_basic_ss addsimps real_abs_thms1  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
679  | 
fun absmaxmin_elim_conv1 ctxt =  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
680  | 
Simplifier.rewrite (Simplifier.context ctxt absmaxmin_elim_ss1)  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
681  | 
|
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
682  | 
val absmaxmin_elim_conv2 =  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
683  | 
let  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
684  | 
   val pth_abs = instantiate' [SOME @{ctyp real}] [] abs_split'
 | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
685  | 
   val pth_max = instantiate' [SOME @{ctyp real}] [] max_split
 | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
686  | 
   val pth_min = instantiate' [SOME @{ctyp real}] [] min_split
 | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
687  | 
   val abs_tm = @{cterm "abs :: real => _"}
 | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
688  | 
   val p_tm = @{cpat "?P :: real => bool"}
 | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
689  | 
   val x_tm = @{cpat "?x :: real"}
 | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
690  | 
   val y_tm = @{cpat "?y::real"}
 | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
691  | 
   val is_max = is_binop @{cterm "max :: real => _"}
 | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
692  | 
   val is_min = is_binop @{cterm "min :: real => _"} 
 | 
| 32828 | 693  | 
fun is_abs t = is_comb t andalso Thm.dest_fun t aconvc abs_tm  | 
| 
31120
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
694  | 
fun eliminate_construct p c tm =  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
695  | 
let  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
696  | 
val t = find_cterm p tm  | 
| 36945 | 697  | 
val th0 = (Thm.symmetric o Thm.beta_conversion false) (Thm.capply (Thm.cabs t tm) t)  | 
| 32828 | 698  | 
val (p,ax) = (Thm.dest_comb o Thm.rhs_of) th0  | 
| 36945 | 699  | 
in fconv_rule(arg_conv(binop_conv (arg_conv (Thm.beta_conversion false))))  | 
700  | 
(Thm.transitive th0 (c p ax))  | 
|
| 
31120
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
701  | 
end  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
702  | 
|
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
703  | 
val elim_abs = eliminate_construct is_abs  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
704  | 
(fn p => fn ax =>  | 
| 32828 | 705  | 
Thm.instantiate ([], [(p_tm,p), (x_tm, Thm.dest_arg ax)]) pth_abs)  | 
| 
31120
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
706  | 
val elim_max = eliminate_construct is_max  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
707  | 
(fn p => fn ax =>  | 
| 32828 | 708  | 
let val (ax,y) = Thm.dest_comb ax  | 
709  | 
in Thm.instantiate ([], [(p_tm,p), (x_tm, Thm.dest_arg ax), (y_tm,y)])  | 
|
| 
31120
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
710  | 
pth_max end)  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
711  | 
val elim_min = eliminate_construct is_min  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
712  | 
(fn p => fn ax =>  | 
| 32828 | 713  | 
let val (ax,y) = Thm.dest_comb ax  | 
714  | 
in Thm.instantiate ([], [(p_tm,p), (x_tm, Thm.dest_arg ax), (y_tm,y)])  | 
|
| 
31120
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
715  | 
pth_min end)  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
716  | 
in first_conv [elim_abs, elim_max, elim_min, all_conv]  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
717  | 
end;  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
718  | 
in fun gen_real_arith ctxt (mkconst,eq,ge,gt,norm,neg,add,mul,prover) =  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
719  | 
gen_gen_real_arith ctxt (mkconst,eq,ge,gt,norm,neg,add,mul,  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
720  | 
absmaxmin_elim_conv1 ctxt,absmaxmin_elim_conv2,prover)  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
721  | 
end;  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
722  | 
|
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
723  | 
(* An instance for reals*)  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
724  | 
|
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
725  | 
fun gen_prover_real_arith ctxt prover =  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
726  | 
let  | 
| 35408 | 727  | 
fun simple_cterm_ord t u = Term_Ord.term_ord (term_of t, term_of u) = LESS  | 
| 
31120
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
728  | 
  val {add,mul,neg,pow,sub,main} = 
 | 
| 
36753
 
5cf4e9128f22
renamed Normalizer to the more specific Semiring_Normalizer
 
haftmann 
parents: 
36751 
diff
changeset
 | 
729  | 
Semiring_Normalizer.semiring_normalizers_ord_wrapper ctxt  | 
| 
 
5cf4e9128f22
renamed Normalizer to the more specific Semiring_Normalizer
 
haftmann 
parents: 
36751 
diff
changeset
 | 
730  | 
      (the (Semiring_Normalizer.match ctxt @{cterm "(0::real) + 1"})) 
 | 
| 
31120
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
731  | 
simple_cterm_ord  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
732  | 
in gen_real_arith ctxt  | 
| 
36751
 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 
haftmann 
parents: 
36721 
diff
changeset
 | 
733  | 
(cterm_of_rat, Numeral_Simprocs.field_comp_conv, Numeral_Simprocs.field_comp_conv, Numeral_Simprocs.field_comp_conv,  | 
| 
31120
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
734  | 
main,neg,add,mul, prover)  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
735  | 
end;  | 
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
736  | 
|
| 
 
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
 
chaieb 
parents:  
diff
changeset
 | 
737  | 
end  |