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