author | nipkow |
Wed, 10 Jan 2018 15:25:09 +0100 | |
changeset 67399 | eab6ce8368fa |
parent 67271 | 48ef58c6cf4c |
child 67559 | 833d154ab189 |
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 |
|
58628 | 5 |
--- also implements Fourier-Motzkin elimination as a special case |
6 |
Fourier-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 |
|
52049 | 47 |
fun choose f = |
48 |
(case Tab.min f of |
|
49 |
SOME entry => entry |
|
50 |
| 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
|
51 |
|
32829
671eb46eb0a3
tuned FuncFun and FuncUtil structure in positivstellensatz.ML
Philipp Meyer
parents:
32828
diff
changeset
|
52 |
fun onefunc kv = update kv empty |
671eb46eb0a3
tuned FuncFun and FuncUtil structure in positivstellensatz.ML
Philipp Meyer
parents:
32828
diff
changeset
|
53 |
|
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
|
54 |
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
|
55 |
|
32645
1cc5b24f5a01
sos method generates and uses proof certificates
Philipp Meyer
parents:
32402
diff
changeset
|
56 |
(* Some standard functors and utility functions for them *) |
1cc5b24f5a01
sos method generates and uses proof certificates
Philipp Meyer
parents:
32402
diff
changeset
|
57 |
|
1cc5b24f5a01
sos method generates and uses proof certificates
Philipp Meyer
parents:
32402
diff
changeset
|
58 |
structure FuncUtil = |
1cc5b24f5a01
sos method generates and uses proof certificates
Philipp Meyer
parents:
32402
diff
changeset
|
59 |
struct |
1cc5b24f5a01
sos method generates and uses proof certificates
Philipp Meyer
parents:
32402
diff
changeset
|
60 |
|
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
|
61 |
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
|
62 |
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
|
63 |
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
|
64 |
structure Symfunc = FuncFun(type key = string val ord = fast_string_ord); |
35408 | 65 |
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
|
66 |
|
59582 | 67 |
val cterm_ord = Term_Ord.fast_term_ord o apply2 Thm.term_of |
32645
1cc5b24f5a01
sos method generates and uses proof certificates
Philipp Meyer
parents:
32402
diff
changeset
|
68 |
|
1cc5b24f5a01
sos method generates and uses proof certificates
Philipp Meyer
parents:
32402
diff
changeset
|
69 |
structure Ctermfunc = FuncFun(type key = cterm val ord = cterm_ord); |
1cc5b24f5a01
sos method generates and uses proof certificates
Philipp Meyer
parents:
32402
diff
changeset
|
70 |
|
32829
671eb46eb0a3
tuned FuncFun and FuncUtil structure in positivstellensatz.ML
Philipp Meyer
parents:
32828
diff
changeset
|
71 |
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
|
72 |
|
59058
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
wenzelm
parents:
58628
diff
changeset
|
73 |
val monomial_ord = list_ord (prod_ord cterm_ord int_ord) o apply2 Ctermfunc.dest |
32645
1cc5b24f5a01
sos method generates and uses proof certificates
Philipp Meyer
parents:
32402
diff
changeset
|
74 |
|
1cc5b24f5a01
sos method generates and uses proof certificates
Philipp Meyer
parents:
32402
diff
changeset
|
75 |
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
|
76 |
|
32829
671eb46eb0a3
tuned FuncFun and FuncUtil structure in positivstellensatz.ML
Philipp Meyer
parents:
32828
diff
changeset
|
77 |
type poly = Rat.rat Monomialfunc.table; |
32645
1cc5b24f5a01
sos method generates and uses proof certificates
Philipp Meyer
parents:
32402
diff
changeset
|
78 |
|
1cc5b24f5a01
sos method generates and uses proof certificates
Philipp Meyer
parents:
32402
diff
changeset
|
79 |
(* 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
|
80 |
|
59058
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
wenzelm
parents:
58628
diff
changeset
|
81 |
fun dest_monomial mon = sort (cterm_ord o apply2 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
|
82 |
|
32645
1cc5b24f5a01
sos method generates and uses proof certificates
Philipp Meyer
parents:
32402
diff
changeset
|
83 |
fun monomial_order (m1,m2) = |
46594 | 84 |
if Ctermfunc.is_empty m2 then LESS |
85 |
else if Ctermfunc.is_empty m1 then GREATER |
|
86 |
else |
|
87 |
let |
|
88 |
val mon1 = dest_monomial m1 |
|
32645
1cc5b24f5a01
sos method generates and uses proof certificates
Philipp Meyer
parents:
32402
diff
changeset
|
89 |
val mon2 = dest_monomial m2 |
33002 | 90 |
val deg1 = fold (Integer.add o snd) mon1 0 |
46594 | 91 |
val deg2 = fold (Integer.add o snd) mon2 0 |
92 |
in if deg1 < deg2 then GREATER |
|
93 |
else if deg1 > deg2 then LESS |
|
94 |
else list_ord (prod_ord cterm_ord int_ord) (mon1,mon2) |
|
95 |
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
|
96 |
|
32645
1cc5b24f5a01
sos method generates and uses proof certificates
Philipp Meyer
parents:
32402
diff
changeset
|
97 |
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
|
98 |
|
32645
1cc5b24f5a01
sos method generates and uses proof certificates
Philipp Meyer
parents:
32402
diff
changeset
|
99 |
(* 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
|
100 |
|
46594 | 101 |
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
|
102 |
sig |
46594 | 103 |
|
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
|
104 |
datatype positivstellensatz = |
46594 | 105 |
Axiom_eq of int |
106 |
| Axiom_le of int |
|
107 |
| Axiom_lt of int |
|
108 |
| Rational_eq of Rat.rat |
|
109 |
| Rational_le of Rat.rat |
|
110 |
| Rational_lt of Rat.rat |
|
111 |
| Square of FuncUtil.poly |
|
112 |
| Eqmul of FuncUtil.poly * positivstellensatz |
|
113 |
| Sum of positivstellensatz * positivstellensatz |
|
114 |
| 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
|
115 |
|
46594 | 116 |
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
|
117 |
|
46594 | 118 |
datatype tree_choice = Left | Right |
32645
1cc5b24f5a01
sos method generates and uses proof certificates
Philipp Meyer
parents:
32402
diff
changeset
|
119 |
|
46594 | 120 |
type prover = tree_choice list -> |
121 |
(thm list * thm list * thm list -> positivstellensatz -> thm) -> |
|
122 |
thm list * thm list * thm list -> thm * pss_tree |
|
123 |
type cert_conv = cterm -> thm * pss_tree |
|
32645
1cc5b24f5a01
sos method generates and uses proof certificates
Philipp Meyer
parents:
32402
diff
changeset
|
124 |
|
46594 | 125 |
val gen_gen_real_arith : |
126 |
Proof.context -> (Rat.rat -> cterm) * conv * conv * conv * |
|
127 |
conv * conv * conv * conv * conv * conv * prover -> cert_conv |
|
128 |
val real_linear_prover : (thm list * thm list * thm list -> positivstellensatz -> thm) -> |
|
129 |
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
|
130 |
|
46594 | 131 |
val gen_real_arith : Proof.context -> |
132 |
(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
|
133 |
|
46594 | 134 |
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
|
135 |
|
46594 | 136 |
val is_ratconst : cterm -> bool |
137 |
val dest_ratconst : cterm -> Rat.rat |
|
138 |
val cterm_of_rat : Rat.rat -> cterm |
|
32645
1cc5b24f5a01
sos method generates and uses proof certificates
Philipp Meyer
parents:
32402
diff
changeset
|
139 |
|
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 |
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
|
141 |
|
32645
1cc5b24f5a01
sos method generates and uses proof certificates
Philipp Meyer
parents:
32402
diff
changeset
|
142 |
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
|
143 |
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
|
144 |
|
46594 | 145 |
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
|
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 |
(* 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
|
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 |
|
fc654c95c29e
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 |
datatype positivstellensatz = |
46594 | 151 |
Axiom_eq of int |
152 |
| Axiom_le of int |
|
153 |
| Axiom_lt of int |
|
154 |
| Rational_eq of Rat.rat |
|
155 |
| Rational_le of Rat.rat |
|
156 |
| Rational_lt of Rat.rat |
|
157 |
| Square of FuncUtil.poly |
|
158 |
| Eqmul of FuncUtil.poly * positivstellensatz |
|
159 |
| Sum of positivstellensatz * positivstellensatz |
|
160 |
| 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
|
161 |
(* 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
|
162 |
|
32645
1cc5b24f5a01
sos method generates and uses proof certificates
Philipp Meyer
parents:
32402
diff
changeset
|
163 |
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
|
164 |
datatype tree_choice = Left | Right |
46594 | 165 |
type prover = tree_choice list -> |
32645
1cc5b24f5a01
sos method generates and uses proof certificates
Philipp Meyer
parents:
32402
diff
changeset
|
166 |
(thm list * thm list * thm list -> positivstellensatz -> thm) -> |
46594 | 167 |
thm list * thm list * thm list -> thm * pss_tree |
32645
1cc5b24f5a01
sos method generates and uses proof certificates
Philipp Meyer
parents:
32402
diff
changeset
|
168 |
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
|
169 |
|
32645
1cc5b24f5a01
sos method generates and uses proof certificates
Philipp Meyer
parents:
32402
diff
changeset
|
170 |
|
1cc5b24f5a01
sos method generates and uses proof certificates
Philipp Meyer
parents:
32402
diff
changeset
|
171 |
(* Some useful derived rules *) |
46594 | 172 |
fun deduct_antisym_rule tha thb = |
59582 | 173 |
Thm.equal_intr (Thm.implies_intr (Thm.cprop_of thb) tha) |
174 |
(Thm.implies_intr (Thm.cprop_of tha) thb); |
|
32645
1cc5b24f5a01
sos method generates and uses proof certificates
Philipp Meyer
parents:
32402
diff
changeset
|
175 |
|
44058 | 176 |
fun prove_hyp tha thb = |
59582 | 177 |
if exists (curry op aconv (Thm.concl_of tha)) (Thm.hyps_of thb) (* FIXME !? *) |
36945 | 178 |
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
|
179 |
|
67091 | 180 |
val pth = @{lemma "(((x::real) < y) \<equiv> (y - x > 0))" and "((x \<le> y) \<equiv> (y - x \<ge> 0))" and |
181 |
"((x = y) \<equiv> (x - y = 0))" and "((\<not>(x < y)) \<equiv> (x - y \<ge> 0))" and |
|
182 |
"((\<not>(x \<le> y)) \<equiv> (x - y > 0))" and "((\<not>(x = y)) \<equiv> (x - y > 0 \<or> -(x - y) > 0))" |
|
33443 | 183 |
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
|
184 |
|
67091 | 185 |
val pth_final = @{lemma "(\<not>p \<Longrightarrow> False) \<Longrightarrow> p" by blast} |
46594 | 186 |
val pth_add = |
67267 | 187 |
@{lemma "(x = (0::real) \<Longrightarrow> y = 0 \<Longrightarrow> x + y = 0 )" and "( x = 0 \<Longrightarrow> y \<ge> 0 \<Longrightarrow> x + y \<ge> 0)" and |
188 |
"(x = 0 \<Longrightarrow> y > 0 \<Longrightarrow> x + y > 0)" and "(x \<ge> 0 \<Longrightarrow> y = 0 \<Longrightarrow> x + y \<ge> 0)" and |
|
189 |
"(x \<ge> 0 \<Longrightarrow> y \<ge> 0 \<Longrightarrow> x + y \<ge> 0)" and "(x \<ge> 0 \<Longrightarrow> y > 0 \<Longrightarrow> x + y > 0)" and |
|
190 |
"(x > 0 \<Longrightarrow> y = 0 \<Longrightarrow> x + y > 0)" and "(x > 0 \<Longrightarrow> y \<ge> 0 \<Longrightarrow> x + y > 0)" and |
|
191 |
"(x > 0 \<Longrightarrow> y > 0 \<Longrightarrow> 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
|
192 |
|
46594 | 193 |
val pth_mul = |
67267 | 194 |
@{lemma "(x = (0::real) \<Longrightarrow> y = 0 \<Longrightarrow> x * y = 0)" and "(x = 0 \<Longrightarrow> y \<ge> 0 \<Longrightarrow> x * y = 0)" and |
195 |
"(x = 0 \<Longrightarrow> y > 0 \<Longrightarrow> x * y = 0)" and "(x \<ge> 0 \<Longrightarrow> y = 0 \<Longrightarrow> x * y = 0)" and |
|
196 |
"(x \<ge> 0 \<Longrightarrow> y \<ge> 0 \<Longrightarrow> x * y \<ge> 0)" and "(x \<ge> 0 \<Longrightarrow> y > 0 \<Longrightarrow> x * y \<ge> 0)" and |
|
197 |
"(x > 0 \<Longrightarrow> y = 0 \<Longrightarrow> x * y = 0)" and "(x > 0 \<Longrightarrow> y \<ge> 0 \<Longrightarrow> x * y \<ge> 0)" and |
|
198 |
"(x > 0 \<Longrightarrow> y > 0 \<Longrightarrow> 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
|
199 |
by (auto intro: mult_mono[where a="0::real" and b="x" and d="y" and c="0", simplified] |
33443 | 200 |
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
|
201 |
|
67267 | 202 |
val pth_emul = @{lemma "y = (0::real) \<Longrightarrow> x * y = 0" by simp}; |
203 |
val pth_square = @{lemma "x * x \<ge> (0::real)" by simp}; |
|
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
|
204 |
|
33443 | 205 |
val weak_dnf_simps = |
45654 | 206 |
List.take (@{thms simp_thms}, 34) @ |
67091 | 207 |
@{lemma "((P \<and> (Q \<or> R)) = ((P\<and>Q) \<or> (P\<and>R)))" and "((Q \<or> R) \<and> P) = ((Q\<and>P) \<or> (R\<and>P))" and |
208 |
"(P \<and> Q) = (Q \<and> P)" and "((P \<or> Q) = (Q \<or> 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
|
209 |
|
44454 | 210 |
(* |
33443 | 211 |
val nnfD_simps = |
212 |
@{lemma "((~(P & Q)) = (~P | ~Q))" and "((~(P | Q)) = (~P & ~Q) )" and |
|
213 |
"((P --> Q) = (~P | Q) )" and "((P = Q) = ((P & Q) | (~P & ~ Q)))" and |
|
214 |
"((~(P = Q)) = ((P & ~ Q) | (~P & Q)) )" and "((~ ~(P)) = P)" by blast+}; |
|
44454 | 215 |
*) |
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
|
216 |
|
67091 | 217 |
val choice_iff = @{lemma "(\<forall>x. \<exists>y. P x y) = (\<exists>f. \<forall>x. P x (f x))" by metis}; |
33443 | 218 |
val prenex_simps = |
219 |
map (fn th => th RS sym) |
|
220 |
([@{thm "all_conj_distrib"}, @{thm "ex_disj_distrib"}] @ |
|
37598 | 221 |
@{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
|
222 |
|
33443 | 223 |
val real_abs_thms1 = @{lemma |
67091 | 224 |
"((-1 * \<bar>x::real\<bar> \<ge> r) = (-1 * x \<ge> r \<and> 1 * x \<ge> r))" and |
225 |
"((-1 * \<bar>x\<bar> + a \<ge> r) = (a + -1 * x \<ge> r \<and> a + 1 * x \<ge> r))" and |
|
226 |
"((a + -1 * \<bar>x\<bar> \<ge> r) = (a + -1 * x \<ge> r \<and> a + 1 * x \<ge> r))" and |
|
227 |
"((a + -1 * \<bar>x\<bar> + b \<ge> r) = (a + -1 * x + b \<ge> r \<and> a + 1 * x + b \<ge> r))" and |
|
228 |
"((a + b + -1 * \<bar>x\<bar> \<ge> r) = (a + b + -1 * x \<ge> r \<and> a + b + 1 * x \<ge> r))" and |
|
229 |
"((a + b + -1 * \<bar>x\<bar> + c \<ge> r) = (a + b + -1 * x + c \<ge> r \<and> a + b + 1 * x + c \<ge> r))" and |
|
230 |
"((-1 * max x y \<ge> r) = (-1 * x \<ge> r \<and> -1 * y \<ge> r))" and |
|
231 |
"((-1 * max x y + a \<ge> r) = (a + -1 * x \<ge> r \<and> a + -1 * y \<ge> r))" and |
|
232 |
"((a + -1 * max x y \<ge> r) = (a + -1 * x \<ge> r \<and> a + -1 * y \<ge> r))" and |
|
233 |
"((a + -1 * max x y + b \<ge> r) = (a + -1 * x + b \<ge> r \<and> a + -1 * y + b \<ge> r))" and |
|
234 |
"((a + b + -1 * max x y \<ge> r) = (a + b + -1 * x \<ge> r \<and> a + b + -1 * y \<ge> r))" and |
|
235 |
"((a + b + -1 * max x y + c \<ge> r) = (a + b + -1 * x + c \<ge> r \<and> a + b + -1 * y + c \<ge> r))" and |
|
236 |
"((1 * min x y \<ge> r) = (1 * x \<ge> r \<and> 1 * y \<ge> r))" and |
|
237 |
"((1 * min x y + a \<ge> r) = (a + 1 * x \<ge> r \<and> a + 1 * y \<ge> r))" and |
|
238 |
"((a + 1 * min x y \<ge> r) = (a + 1 * x \<ge> r \<and> a + 1 * y \<ge> r))" and |
|
239 |
"((a + 1 * min x y + b \<ge> r) = (a + 1 * x + b \<ge> r \<and> a + 1 * y + b \<ge> r))" and |
|
240 |
"((a + b + 1 * min x y \<ge> r) = (a + b + 1 * x \<ge> r \<and> a + b + 1 * y \<ge> r))" and |
|
241 |
"((a + b + 1 * min x y + c \<ge> r) = (a + b + 1 * x + c \<ge> r \<and> a + b + 1 * y + c \<ge> r))" and |
|
242 |
"((min x y \<ge> r) = (x \<ge> r \<and> y \<ge> r))" and |
|
243 |
"((min x y + a \<ge> r) = (a + x \<ge> r \<and> a + y \<ge> r))" and |
|
244 |
"((a + min x y \<ge> r) = (a + x \<ge> r \<and> a + y \<ge> r))" and |
|
245 |
"((a + min x y + b \<ge> r) = (a + x + b \<ge> r \<and> a + y + b \<ge> r))" and |
|
246 |
"((a + b + min x y \<ge> r) = (a + b + x \<ge> r \<and> a + b + y \<ge> r))" and |
|
247 |
"((a + b + min x y + c \<ge> r) = (a + b + x + c \<ge> r \<and> a + b + y + c \<ge> r))" and |
|
248 |
"((-1 * \<bar>x\<bar> > r) = (-1 * x > r \<and> 1 * x > r))" and |
|
249 |
"((-1 * \<bar>x\<bar> + a > r) = (a + -1 * x > r \<and> a + 1 * x > r))" and |
|
250 |
"((a + -1 * \<bar>x\<bar> > r) = (a + -1 * x > r \<and> a + 1 * x > r))" and |
|
251 |
"((a + -1 * \<bar>x\<bar> + b > r) = (a + -1 * x + b > r \<and> a + 1 * x + b > r))" and |
|
252 |
"((a + b + -1 * \<bar>x\<bar> > r) = (a + b + -1 * x > r \<and> a + b + 1 * x > r))" and |
|
253 |
"((a + b + -1 * \<bar>x\<bar> + c > r) = (a + b + -1 * x + c > r \<and> a + b + 1 * x + c > r))" and |
|
254 |
"((-1 * max x y > r) = ((-1 * x > r) \<and> -1 * y > r))" and |
|
255 |
"((-1 * max x y + a > r) = (a + -1 * x > r \<and> a + -1 * y > r))" and |
|
256 |
"((a + -1 * max x y > r) = (a + -1 * x > r \<and> a + -1 * y > r))" and |
|
257 |
"((a + -1 * max x y + b > r) = (a + -1 * x + b > r \<and> a + -1 * y + b > r))" and |
|
258 |
"((a + b + -1 * max x y > r) = (a + b + -1 * x > r \<and> a + b + -1 * y > r))" and |
|
259 |
"((a + b + -1 * max x y + c > r) = (a + b + -1 * x + c > r \<and> a + b + -1 * y + c > r))" and |
|
260 |
"((min x y > r) = (x > r \<and> y > r))" and |
|
261 |
"((min x y + a > r) = (a + x > r \<and> a + y > r))" and |
|
262 |
"((a + min x y > r) = (a + x > r \<and> a + y > r))" and |
|
263 |
"((a + min x y + b > r) = (a + x + b > r \<and> a + y + b > r))" and |
|
264 |
"((a + b + min x y > r) = (a + b + x > r \<and> a + b + y > r))" and |
|
265 |
"((a + b + min x y + c > r) = (a + b + x + c > r \<and> 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
|
266 |
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
|
267 |
|
67091 | 268 |
val abs_split' = @{lemma "P \<bar>x::'a::linordered_idom\<bar> == (x \<ge> 0 \<and> P x \<or> x < 0 \<and> P (-x))" |
63648 | 269 |
by (atomize (full)) (auto split: abs_split)}; |
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
|
270 |
|
67091 | 271 |
val max_split = @{lemma "P (max x y) \<equiv> ((x::'a::linorder) \<le> y \<and> P y \<or> x > y \<and> P x)" |
272 |
by (atomize (full)) (cases "x \<le> y", auto simp add: max_def)}; |
|
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
|
273 |
|
67091 | 274 |
val min_split = @{lemma "P (min x y) \<equiv> ((x::'a::linorder) \<le> y \<and> P x \<or> x > y \<and> P y)" |
275 |
by (atomize (full)) (cases "x \<le> y", auto simp add: min_def)}; |
|
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
|
276 |
|
fc654c95c29e
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 |
|
39920 | 278 |
(* Miscellaneous *) |
46594 | 279 |
fun literals_conv bops uops cv = |
280 |
let |
|
281 |
fun h t = |
|
59582 | 282 |
(case Thm.term_of t of |
46594 | 283 |
b$_$_ => if member (op aconv) bops b then binop_conv h t else cv t |
284 |
| u$_ => if member (op aconv) uops u then arg_conv h t else cv t |
|
59582 | 285 |
| _ => cv t) |
46594 | 286 |
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
|
287 |
|
46594 | 288 |
fun cterm_of_rat x = |
289 |
let |
|
63201 | 290 |
val (a, b) = Rat.dest x |
46594 | 291 |
in |
67267 | 292 |
if b = 1 then Numeral.mk_cnumber \<^ctyp>\<open>real\<close> a |
67399 | 293 |
else Thm.apply (Thm.apply \<^cterm>\<open>(/) :: real \<Rightarrow> _\<close> |
67267 | 294 |
(Numeral.mk_cnumber \<^ctyp>\<open>real\<close> a)) |
295 |
(Numeral.mk_cnumber \<^ctyp>\<open>real\<close> b) |
|
46594 | 296 |
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
|
297 |
|
46594 | 298 |
fun dest_ratconst t = |
59582 | 299 |
case Thm.term_of t of |
67267 | 300 |
Const(\<^const_name>\<open>divide\<close>, _)$a$b => Rat.make(HOLogic.dest_number a |> snd, HOLogic.dest_number b |> snd) |
63201 | 301 |
| _ => Rat.of_int (HOLogic.dest_number (Thm.term_of t) |> snd) |
46594 | 302 |
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
|
303 |
|
44454 | 304 |
(* |
46594 | 305 |
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
|
306 |
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
|
307 |
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
|
308 |
| 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
|
309 |
| _ => raise TERM ("find_term",[t]); |
44454 | 310 |
*) |
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
|
311 |
|
46594 | 312 |
fun find_cterm p t = |
313 |
if p t then t else |
|
59582 | 314 |
case Thm.term_of t of |
46594 | 315 |
_$_ => (find_cterm p (Thm.dest_fun t) handle CTERM _ => find_cterm p (Thm.dest_arg t)) |
316 |
| Abs (_,_,_) => find_cterm p (Thm.dest_abs NONE t |> snd) |
|
317 |
| _ => 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
|
318 |
|
59582 | 319 |
fun is_comb t = (case Thm.term_of t of _ $ _ => true | _ => 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
|
320 |
|
fc654c95c29e
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 |
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
|
322 |
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
|
323 |
|
32645
1cc5b24f5a01
sos method generates and uses proof certificates
Philipp Meyer
parents:
32402
diff
changeset
|
324 |
|
1cc5b24f5a01
sos method generates and uses proof certificates
Philipp Meyer
parents:
32402
diff
changeset
|
325 |
(* Map back polynomials to HOL. *) |
1cc5b24f5a01
sos method generates and uses proof certificates
Philipp Meyer
parents:
32402
diff
changeset
|
326 |
|
67399 | 327 |
fun cterm_of_varpow x k = if k = 1 then x else Thm.apply (Thm.apply \<^cterm>\<open>(^) :: real \<Rightarrow> _\<close> x) |
67267 | 328 |
(Numeral.mk_cnumber \<^ctyp>\<open>nat\<close> k) |
32645
1cc5b24f5a01
sos method generates and uses proof certificates
Philipp Meyer
parents:
32402
diff
changeset
|
329 |
|
46594 | 330 |
fun cterm_of_monomial m = |
67267 | 331 |
if FuncUtil.Ctermfunc.is_empty m then \<^cterm>\<open>1::real\<close> |
46594 | 332 |
else |
333 |
let |
|
334 |
val m' = FuncUtil.dest_monomial m |
|
335 |
val vps = fold_rev (fn (x,k) => cons (cterm_of_varpow x k)) m' [] |
|
67399 | 336 |
in foldr1 (fn (s, t) => Thm.apply (Thm.apply \<^cterm>\<open>( * ) :: real \<Rightarrow> _\<close> s) t) vps |
46594 | 337 |
end |
32645
1cc5b24f5a01
sos method generates and uses proof certificates
Philipp Meyer
parents:
32402
diff
changeset
|
338 |
|
46594 | 339 |
fun cterm_of_cmonomial (m,c) = |
340 |
if FuncUtil.Ctermfunc.is_empty m then cterm_of_rat c |
|
63205 | 341 |
else if c = @1 then cterm_of_monomial m |
67399 | 342 |
else Thm.apply (Thm.apply \<^cterm>\<open>( * )::real \<Rightarrow> _\<close> (cterm_of_rat c)) (cterm_of_monomial m); |
32645
1cc5b24f5a01
sos method generates and uses proof certificates
Philipp Meyer
parents:
32402
diff
changeset
|
343 |
|
46594 | 344 |
fun cterm_of_poly p = |
67267 | 345 |
if FuncUtil.Monomialfunc.is_empty p then \<^cterm>\<open>0::real\<close> |
46594 | 346 |
else |
347 |
let |
|
348 |
val cms = map cterm_of_cmonomial |
|
349 |
(sort (prod_ord FuncUtil.monomial_order (K EQUAL)) (FuncUtil.Monomialfunc.dest p)) |
|
67399 | 350 |
in foldr1 (fn (t1, t2) => Thm.apply(Thm.apply \<^cterm>\<open>(+) :: real \<Rightarrow> _\<close> t1) t2) cms |
46594 | 351 |
end; |
32645
1cc5b24f5a01
sos method generates and uses proof certificates
Philipp Meyer
parents:
32402
diff
changeset
|
352 |
|
46594 | 353 |
(* 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
|
354 |
|
fc654c95c29e
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 |
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
|
356 |
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
|
357 |
poly_conv,poly_neg_conv,poly_add_conv,poly_mul_conv, |
46594 | 358 |
absconv1,absconv2,prover) = |
359 |
let |
|
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
46594
diff
changeset
|
360 |
val pre_ss = put_simpset HOL_basic_ss ctxt addsimps |
58628 | 361 |
@{thms simp_thms ex_simps all_simps not_all not_ex ex_disj_distrib |
362 |
all_conj_distrib if_bool_eq_disj} |
|
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
46594
diff
changeset
|
363 |
val prenex_ss = put_simpset HOL_basic_ss ctxt addsimps prenex_simps |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
46594
diff
changeset
|
364 |
val skolemize_ss = put_simpset HOL_basic_ss ctxt addsimps [choice_iff] |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
46594
diff
changeset
|
365 |
val presimp_conv = Simplifier.rewrite pre_ss |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
46594
diff
changeset
|
366 |
val prenex_conv = Simplifier.rewrite prenex_ss |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
46594
diff
changeset
|
367 |
val skolemize_conv = Simplifier.rewrite skolemize_ss |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
46594
diff
changeset
|
368 |
val weak_dnf_ss = put_simpset HOL_basic_ss ctxt addsimps weak_dnf_simps |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
46594
diff
changeset
|
369 |
val weak_dnf_conv = Simplifier.rewrite weak_dnf_ss |
46594 | 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 |
|
67399 | 373 |
in if g aconvc \<^cterm>\<open>((\<le>)) :: real \<Rightarrow> _\<close> |
374 |
orelse g aconvc \<^cterm>\<open>((<)) :: real \<Rightarrow> _\<close> |
|
46594 | 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)) |
60801 | 397 |
(Thm.instantiate' [] [SOME ct] (th RS pth_emul)) |
46594 | 398 |
fun square_rule t = fconv_rule (arg_conv (oprconv poly_conv)) |
60801 | 399 |
(Thm.instantiate' [] [SOME t] pth_square) |
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 |
|
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 |
|
67267 | 408 |
| Rational_eq x => eqT_elim(numeric_eq_conv(Thm.apply \<^cterm>\<open>Trueprop\<close> |
67399 | 409 |
(Thm.apply (Thm.apply \<^cterm>\<open>((=))::real \<Rightarrow> _\<close> (mk_numeric x)) |
67267 | 410 |
\<^cterm>\<open>0::real\<close>))) |
411 |
| Rational_le x => eqT_elim(numeric_ge_conv(Thm.apply \<^cterm>\<open>Trueprop\<close> |
|
67399 | 412 |
(Thm.apply (Thm.apply \<^cterm>\<open>((\<le>))::real \<Rightarrow> _\<close> |
67267 | 413 |
\<^cterm>\<open>0::real\<close>) (mk_numeric x)))) |
414 |
| Rational_lt x => eqT_elim(numeric_gt_conv(Thm.apply \<^cterm>\<open>Trueprop\<close> |
|
67399 | 415 |
(Thm.apply (Thm.apply \<^cterm>\<open>((<))::real \<Rightarrow> _\<close> \<^cterm>\<open>0::real\<close>) |
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 |
|
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
46594
diff
changeset
|
426 |
nnf_conv ctxt then_conv skolemize_conv then_conv prenex_conv then_conv |
46594 | 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 |
|
59582 | 429 |
val concl = Thm.dest_arg o Thm.cprop_of |
46594 | 430 |
fun is_binop opr ct = (Thm.dest_fun2 ct aconvc opr handle CTERM _ => false) |
67399 | 431 |
val is_req = is_binop \<^cterm>\<open>((=)):: real \<Rightarrow> _\<close> |
432 |
val is_ge = is_binop \<^cterm>\<open>((\<le>)):: real \<Rightarrow> _\<close> |
|
433 |
val is_gt = is_binop \<^cterm>\<open>((<)):: real \<Rightarrow> _\<close> |
|
67267 | 434 |
val is_conj = is_binop \<^cterm>\<open>HOL.conj\<close> |
435 |
val is_disj = is_binop \<^cterm>\<open>HOL.disj\<close> |
|
46594 | 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 |
|
58628 | 441 |
val _ = |
442 |
if c aconvc (concl th2) then () |
|
443 |
else error "disj_cases : conclusions not alpha convertible" |
|
46594 | 444 |
in Thm.implies_elim (Thm.implies_elim |
60801 | 445 |
(Thm.implies_elim (Thm.instantiate' [] (map SOME [p,q,c]) @{thm disjE}) th) |
67267 | 446 |
(Thm.implies_intr (Thm.apply \<^cterm>\<open>Trueprop\<close> p) th1)) |
447 |
(Thm.implies_intr (Thm.apply \<^cterm>\<open>Trueprop\<close> q) th2) |
|
46594 | 448 |
end |
449 |
fun overall cert_choice dun ths = |
|
450 |
case ths of |
|
451 |
[] => |
|
452 |
let |
|
453 |
val (eq,ne) = List.partition (is_req o concl) dun |
|
454 |
val (le,nl) = List.partition (is_ge o concl) ne |
|
455 |
val lt = filter (is_gt o concl) nl |
|
456 |
in prover (rev cert_choice) hol_of_positivstellensatz (eq,le,lt) end |
|
457 |
| th::oths => |
|
458 |
let |
|
459 |
val ct = concl th |
|
460 |
in |
|
461 |
if is_conj ct then |
|
462 |
let |
|
463 |
val (th1,th2) = conj_pair th |
|
464 |
in overall cert_choice dun (th1::th2::oths) end |
|
465 |
else if is_disj ct then |
|
466 |
let |
|
58628 | 467 |
val (th1, cert1) = |
468 |
overall (Left::cert_choice) dun |
|
67267 | 469 |
(Thm.assume (Thm.apply \<^cterm>\<open>Trueprop\<close> (Thm.dest_arg1 ct))::oths) |
58628 | 470 |
val (th2, cert2) = |
471 |
overall (Right::cert_choice) dun |
|
67267 | 472 |
(Thm.assume (Thm.apply \<^cterm>\<open>Trueprop\<close> (Thm.dest_arg ct))::oths) |
46594 | 473 |
in (disj_cases th th1 th2, Branch (cert1, cert2)) end |
474 |
else overall cert_choice (th::dun) oths |
|
475 |
end |
|
476 |
fun dest_binary b ct = |
|
477 |
if is_binop b ct then Thm.dest_binop ct |
|
478 |
else raise CTERM ("dest_binary",[b,ct]) |
|
67399 | 479 |
val dest_eq = dest_binary \<^cterm>\<open>((=)) :: real \<Rightarrow> _\<close> |
46594 | 480 |
val neq_th = nth pth 5 |
481 |
fun real_not_eq_conv ct = |
|
482 |
let |
|
483 |
val (l,r) = dest_eq (Thm.dest_arg ct) |
|
67267 | 484 |
val th = Thm.instantiate ([],[((("x", 0), \<^typ>\<open>real\<close>),l),((("y", 0), \<^typ>\<open>real\<close>),r)]) neq_th |
46594 | 485 |
val th_p = poly_conv(Thm.dest_arg(Thm.dest_arg1(Thm.rhs_of th))) |
67267 | 486 |
val th_x = Drule.arg_cong_rule \<^cterm>\<open>uminus :: real \<Rightarrow> _\<close> th_p |
46594 | 487 |
val th_n = fconv_rule (arg_conv poly_neg_conv) th_x |
67267 | 488 |
val th' = Drule.binop_cong_rule \<^cterm>\<open>HOL.disj\<close> |
67399 | 489 |
(Drule.arg_cong_rule (Thm.apply \<^cterm>\<open>((<))::real \<Rightarrow> _\<close> \<^cterm>\<open>0::real\<close>) th_p) |
490 |
(Drule.arg_cong_rule (Thm.apply \<^cterm>\<open>((<))::real \<Rightarrow> _\<close> \<^cterm>\<open>0::real\<close>) th_n) |
|
46594 | 491 |
in Thm.transitive th th' |
492 |
end |
|
493 |
fun equal_implies_1_rule PQ = |
|
494 |
let |
|
495 |
val P = Thm.lhs_of PQ |
|
496 |
in Thm.implies_intr P (Thm.equal_elim PQ (Thm.assume P)) |
|
497 |
end |
|
63667 | 498 |
(*FIXME!!! Copied from groebner.ml*) |
46594 | 499 |
val strip_exists = |
500 |
let |
|
501 |
fun h (acc, t) = |
|
59582 | 502 |
case Thm.term_of t of |
67267 | 503 |
Const(\<^const_name>\<open>Ex\<close>,_)$Abs(_,_,_) => |
58628 | 504 |
h (Thm.dest_abs NONE (Thm.dest_arg t) |>> (fn v => v::acc)) |
46594 | 505 |
| _ => (acc,t) |
506 |
in fn t => h ([],t) |
|
507 |
end |
|
508 |
fun name_of x = |
|
59582 | 509 |
case Thm.term_of x of |
46594 | 510 |
Free(s,_) => s |
511 |
| Var ((s,_),_) => s |
|
512 |
| _ => "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
|
513 |
|
58628 | 514 |
fun mk_forall x th = |
61075 | 515 |
let |
516 |
val T = Thm.typ_of_cterm x |
|
67267 | 517 |
val all = Thm.cterm_of ctxt (Const (\<^const_name>\<open>All\<close>, (T --> \<^typ>\<open>bool\<close>) --> \<^typ>\<open>bool\<close>)) |
61075 | 518 |
in Drule.arg_cong_rule all (Thm.abstract_rule (name_of x) x th) 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
|
519 |
|
60801 | 520 |
val specl = fold_rev (fn x => fn th => Thm.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
|
521 |
|
67267 | 522 |
fun ext T = Thm.cterm_of ctxt (Const (\<^const_name>\<open>Ex\<close>, (T --> \<^typ>\<open>bool\<close>) --> \<^typ>\<open>bool\<close>)) |
61075 | 523 |
fun mk_ex v t = Thm.apply (ext (Thm.typ_of_cterm 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
|
524 |
|
46594 | 525 |
fun choose v th th' = |
59582 | 526 |
case Thm.concl_of th of |
67267 | 527 |
\<^term>\<open>Trueprop\<close> $ (Const(\<^const_name>\<open>Ex\<close>,_)$_) => |
46594 | 528 |
let |
59582 | 529 |
val p = (funpow 2 Thm.dest_arg o Thm.cprop_of) th |
59586 | 530 |
val T = (hd o Thm.dest_ctyp o Thm.ctyp_of_cterm) p |
46594 | 531 |
val th0 = fconv_rule (Thm.beta_conversion true) |
60801 | 532 |
(Thm.instantiate' [SOME T] [SOME p, (SOME o Thm.dest_arg o Thm.cprop_of) th'] exE) |
46594 | 533 |
val pv = (Thm.rhs_of o Thm.beta_conversion true) |
67267 | 534 |
(Thm.apply \<^cterm>\<open>Trueprop\<close> (Thm.apply p v)) |
46594 | 535 |
val th1 = Thm.forall_intr v (Thm.implies_intr pv th') |
536 |
in Thm.implies_elim (Thm.implies_elim th0 th) th1 end |
|
537 |
| _ => 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
|
538 |
|
46594 | 539 |
fun simple_choose v th = |
58628 | 540 |
choose v |
541 |
(Thm.assume |
|
67267 | 542 |
((Thm.apply \<^cterm>\<open>Trueprop\<close> o mk_ex v) (Thm.dest_arg (hd (Thm.chyps_of 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
|
543 |
|
46594 | 544 |
val strip_forall = |
545 |
let |
|
546 |
fun h (acc, t) = |
|
59582 | 547 |
case Thm.term_of t of |
67267 | 548 |
Const(\<^const_name>\<open>All\<close>,_)$Abs(_,_,_) => |
58628 | 549 |
h (Thm.dest_abs NONE (Thm.dest_arg t) |>> (fn v => v::acc)) |
46594 | 550 |
| _ => (acc,t) |
551 |
in fn t => h ([],t) |
|
552 |
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
|
553 |
|
46594 | 554 |
fun f ct = |
555 |
let |
|
556 |
val nnf_norm_conv' = |
|
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
46594
diff
changeset
|
557 |
nnf_conv ctxt then_conv |
67267 | 558 |
literals_conv [\<^term>\<open>HOL.conj\<close>, \<^term>\<open>HOL.disj\<close>] [] |
46594 | 559 |
(Conv.cache_conv |
560 |
(first_conv [real_lt_conv, real_le_conv, |
|
561 |
real_eq_conv, real_not_lt_conv, |
|
562 |
real_not_le_conv, real_not_eq_conv, all_conv])) |
|
67267 | 563 |
fun absremover ct = (literals_conv [\<^term>\<open>HOL.conj\<close>, \<^term>\<open>HOL.disj\<close>] [] |
46594 | 564 |
(try_conv (absconv1 then_conv binop_conv (arg_conv poly_conv))) then_conv |
565 |
try_conv (absconv2 then_conv nnf_norm_conv' then_conv binop_conv absremover)) ct |
|
67267 | 566 |
val nct = Thm.apply \<^cterm>\<open>Trueprop\<close> (Thm.apply \<^cterm>\<open>Not\<close> ct) |
46594 | 567 |
val th0 = (init_conv then_conv arg_conv nnf_norm_conv') nct |
568 |
val tm0 = Thm.dest_arg (Thm.rhs_of th0) |
|
569 |
val (th, certificates) = |
|
67267 | 570 |
if tm0 aconvc \<^cterm>\<open>False\<close> then (equal_implies_1_rule th0, Trivial) else |
46594 | 571 |
let |
572 |
val (evs,bod) = strip_exists tm0 |
|
573 |
val (avs,ibod) = strip_forall bod |
|
67267 | 574 |
val th1 = Drule.arg_cong_rule \<^cterm>\<open>Trueprop\<close> (fold mk_forall avs (absremover ibod)) |
46594 | 575 |
val (th2, certs) = overall [] [] [specl avs (Thm.assume (Thm.rhs_of th1))] |
58628 | 576 |
val th3 = |
577 |
fold simple_choose evs |
|
67267 | 578 |
(prove_hyp (Thm.equal_elim th1 (Thm.assume (Thm.apply \<^cterm>\<open>Trueprop\<close> bod))) th2) |
46594 | 579 |
in (Drule.implies_intr_hyps (prove_hyp (Thm.equal_elim th0 (Thm.assume nct)) th3), certs) |
580 |
end |
|
60801 | 581 |
in (Thm.implies_elim (Thm.instantiate' [] [SOME ct] pth_final) th, certificates) |
46594 | 582 |
end |
583 |
in f |
|
584 |
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
|
585 |
|
fc654c95c29e
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 |
(* 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
|
587 |
local |
63205 | 588 |
val linear_add = FuncUtil.Ctermfunc.combine (curry op +) (fn z => z = @0) |
63198
c583ca33076a
ad-hoc overloading for standard operations on type Rat.rat;
wenzelm
parents:
62177
diff
changeset
|
589 |
fun linear_cmul c = FuncUtil.Ctermfunc.map (fn _ => fn x => c * x) |
67267 | 590 |
val one_tm = \<^cterm>\<open>1::real\<close> |
63205 | 591 |
fun contradictory p (e,_) = ((FuncUtil.Ctermfunc.is_empty e) andalso not(p @0)) orelse |
33038 | 592 |
((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
|
593 |
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
|
594 |
|
46594 | 595 |
fun linear_ineqs vars (les,lts) = |
63205 | 596 |
case find_first (contradictory (fn x => x > @0)) lts of |
46594 | 597 |
SOME r => r |
598 |
| NONE => |
|
63205 | 599 |
(case find_first (contradictory (fn x => x > @0)) les of |
46594 | 600 |
SOME r => r |
601 |
| NONE => |
|
602 |
if null vars then error "linear_ineqs: no contradiction" else |
|
603 |
let |
|
604 |
val ineqs = les @ lts |
|
605 |
fun blowup v = |
|
63205 | 606 |
length(filter (fn (e,_) => FuncUtil.Ctermfunc.tryapplyd e v @0 = @0) ineqs) + |
607 |
length(filter (fn (e,_) => FuncUtil.Ctermfunc.tryapplyd e v @0 > @0) ineqs) * |
|
608 |
length(filter (fn (e,_) => FuncUtil.Ctermfunc.tryapplyd e v @0 < @0) ineqs) |
|
46594 | 609 |
val v = fst(hd(sort (fn ((_,i),(_,j)) => int_ord (i,j)) |
610 |
(map (fn v => (v,blowup v)) vars))) |
|
611 |
fun addup (e1,p1) (e2,p2) acc = |
|
612 |
let |
|
63205 | 613 |
val c1 = FuncUtil.Ctermfunc.tryapplyd e1 v @0 |
614 |
val c2 = FuncUtil.Ctermfunc.tryapplyd e2 v @0 |
|
46594 | 615 |
in |
63205 | 616 |
if c1 * c2 >= @0 then acc else |
46594 | 617 |
let |
63211 | 618 |
val e1' = linear_cmul (abs c2) e1 |
619 |
val e2' = linear_cmul (abs c1) e2 |
|
620 |
val p1' = Product(Rational_lt (abs c2), p1) |
|
621 |
val p2' = Product(Rational_lt (abs c1), p2) |
|
46594 | 622 |
in (linear_add e1' e2',Sum(p1',p2'))::acc |
623 |
end |
|
624 |
end |
|
625 |
val (les0,les1) = |
|
63205 | 626 |
List.partition (fn (e,_) => FuncUtil.Ctermfunc.tryapplyd e v @0 = @0) les |
46594 | 627 |
val (lts0,lts1) = |
63205 | 628 |
List.partition (fn (e,_) => FuncUtil.Ctermfunc.tryapplyd e v @0 = @0) lts |
46594 | 629 |
val (lesp,lesn) = |
63205 | 630 |
List.partition (fn (e,_) => FuncUtil.Ctermfunc.tryapplyd e v @0 > @0) les1 |
46594 | 631 |
val (ltsp,ltsn) = |
63205 | 632 |
List.partition (fn (e,_) => FuncUtil.Ctermfunc.tryapplyd e v @0 > @0) lts1 |
46594 | 633 |
val les' = fold_rev (fn ep1 => fold_rev (addup ep1) lesp) lesn les0 |
634 |
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
|
635 |
(fold_rev (fn ep1 => fold_rev (addup ep1) (lesn@ltsn)) ltsp lts0) |
46594 | 636 |
in linear_ineqs (remove (op aconvc) v vars) (les',lts') |
637 |
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
|
638 |
|
46594 | 639 |
fun linear_eqs(eqs,les,lts) = |
63205 | 640 |
case find_first (contradictory (fn x => x = @0)) eqs of |
46594 | 641 |
SOME r => r |
642 |
| NONE => |
|
643 |
(case eqs of |
|
644 |
[] => |
|
645 |
let val vars = remove (op aconvc) one_tm |
|
646 |
(fold_rev (union (op aconvc) o FuncUtil.Ctermfunc.dom o fst) (les@lts) []) |
|
647 |
in linear_ineqs vars (les,lts) end |
|
648 |
| (e,p)::es => |
|
649 |
if FuncUtil.Ctermfunc.is_empty e then linear_eqs (es,les,lts) else |
|
650 |
let |
|
651 |
val (x,c) = FuncUtil.Ctermfunc.choose (FuncUtil.Ctermfunc.delete_safe one_tm e) |
|
652 |
fun xform (inp as (t,q)) = |
|
63205 | 653 |
let val d = FuncUtil.Ctermfunc.tryapplyd t x @0 in |
654 |
if d = @0 then inp else |
|
46594 | 655 |
let |
63211 | 656 |
val k = ~ d * abs c / c |
46594 | 657 |
val e' = linear_cmul k e |
63211 | 658 |
val t' = linear_cmul (abs c) t |
46594 | 659 |
val p' = Eqmul(FuncUtil.Monomialfunc.onefunc (FuncUtil.Ctermfunc.empty, k),p) |
63211 | 660 |
val q' = Product(Rational_lt (abs c), q) |
46594 | 661 |
in (linear_add e' t',Sum(p',q')) |
662 |
end |
|
663 |
end |
|
664 |
in linear_eqs(map xform es,map xform les,map xform lts) |
|
665 |
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
|
666 |
|
46594 | 667 |
fun linear_prover (eq,le,lt) = |
668 |
let |
|
669 |
val eqs = map_index (fn (n, p) => (p,Axiom_eq n)) eq |
|
670 |
val les = map_index (fn (n, p) => (p,Axiom_le n)) le |
|
671 |
val lts = map_index (fn (n, p) => (p,Axiom_lt n)) lt |
|
672 |
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
|
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 |
|
46594 | 675 |
fun lin_of_hol ct = |
67267 | 676 |
if ct aconvc \<^cterm>\<open>0::real\<close> then FuncUtil.Ctermfunc.empty |
63205 | 677 |
else if not (is_comb ct) then FuncUtil.Ctermfunc.onefunc (ct, @1) |
46594 | 678 |
else if is_ratconst ct then FuncUtil.Ctermfunc.onefunc (one_tm, dest_ratconst ct) |
679 |
else |
|
680 |
let val (lop,r) = Thm.dest_comb ct |
|
681 |
in |
|
63205 | 682 |
if not (is_comb lop) then FuncUtil.Ctermfunc.onefunc (ct, @1) |
46594 | 683 |
else |
684 |
let val (opr,l) = Thm.dest_comb lop |
|
685 |
in |
|
67399 | 686 |
if opr aconvc \<^cterm>\<open>(+) :: real \<Rightarrow> _\<close> |
46594 | 687 |
then linear_add (lin_of_hol l) (lin_of_hol r) |
67399 | 688 |
else if opr aconvc \<^cterm>\<open>( * ) :: real \<Rightarrow> _\<close> |
46594 | 689 |
andalso is_ratconst l then FuncUtil.Ctermfunc.onefunc (r, dest_ratconst l) |
63205 | 690 |
else FuncUtil.Ctermfunc.onefunc (ct, @1) |
46594 | 691 |
end |
692 |
end |
|
693 |
||
694 |
fun is_alien ct = |
|
59582 | 695 |
case Thm.term_of ct of |
67267 | 696 |
Const(\<^const_name>\<open>of_nat\<close>, _)$ n => not (can HOLogic.dest_number n) |
697 |
| Const(\<^const_name>\<open>of_int\<close>, _)$ n => not (can HOLogic.dest_number n) |
|
59582 | 698 |
| _ => false |
46594 | 699 |
in |
700 |
fun real_linear_prover translator (eq,le,lt) = |
|
701 |
let |
|
59582 | 702 |
val lhs = lin_of_hol o Thm.dest_arg1 o Thm.dest_arg o Thm.cprop_of |
703 |
val rhs = lin_of_hol o Thm.dest_arg o Thm.dest_arg o Thm.cprop_of |
|
46594 | 704 |
val eq_pols = map lhs eq |
705 |
val le_pols = map rhs le |
|
706 |
val lt_pols = map rhs lt |
|
707 |
val aliens = filter is_alien |
|
708 |
(fold_rev (union (op aconvc) o FuncUtil.Ctermfunc.dom) |
|
709 |
(eq_pols @ le_pols @ lt_pols) []) |
|
63205 | 710 |
val le_pols' = le_pols @ map (fn v => FuncUtil.Ctermfunc.onefunc (v,@1)) aliens |
46594 | 711 |
val (_,proof) = linear_prover (eq_pols,le_pols',lt_pols) |
61609
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents:
61075
diff
changeset
|
712 |
val le' = le @ map (fn a => Thm.instantiate' [] [SOME (Thm.dest_arg a)] @{thm of_nat_0_le_iff}) aliens |
46594 | 713 |
in ((translator (eq,le',lt) proof), Trivial) |
714 |
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
|
715 |
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 |
|
fc654c95c29e
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 |
(* 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
|
718 |
|
fc654c95c29e
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 |
local |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
46594
diff
changeset
|
720 |
val absmaxmin_elim_ss1 = |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
46594
diff
changeset
|
721 |
simpset_of (put_simpset HOL_basic_ss @{context} addsimps real_abs_thms1) |
46594 | 722 |
fun absmaxmin_elim_conv1 ctxt = |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
46594
diff
changeset
|
723 |
Simplifier.rewrite (put_simpset absmaxmin_elim_ss1 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
|
724 |
|
46594 | 725 |
val absmaxmin_elim_conv2 = |
726 |
let |
|
67267 | 727 |
val pth_abs = Thm.instantiate' [SOME \<^ctyp>\<open>real\<close>] [] abs_split' |
728 |
val pth_max = Thm.instantiate' [SOME \<^ctyp>\<open>real\<close>] [] max_split |
|
729 |
val pth_min = Thm.instantiate' [SOME \<^ctyp>\<open>real\<close>] [] min_split |
|
730 |
val abs_tm = \<^cterm>\<open>abs :: real \<Rightarrow> _\<close> |
|
731 |
val p_v = (("P", 0), \<^typ>\<open>real \<Rightarrow> bool\<close>) |
|
732 |
val x_v = (("x", 0), \<^typ>\<open>real\<close>) |
|
733 |
val y_v = (("y", 0), \<^typ>\<open>real\<close>) |
|
734 |
val is_max = is_binop \<^cterm>\<open>max :: real \<Rightarrow> _\<close> |
|
735 |
val is_min = is_binop \<^cterm>\<open>min :: real \<Rightarrow> _\<close> |
|
46594 | 736 |
fun is_abs t = is_comb t andalso Thm.dest_fun t aconvc abs_tm |
737 |
fun eliminate_construct p c tm = |
|
738 |
let |
|
739 |
val t = find_cterm p tm |
|
740 |
val th0 = (Thm.symmetric o Thm.beta_conversion false) (Thm.apply (Thm.lambda t tm) t) |
|
741 |
val (p,ax) = (Thm.dest_comb o Thm.rhs_of) th0 |
|
742 |
in fconv_rule(arg_conv(binop_conv (arg_conv (Thm.beta_conversion false)))) |
|
743 |
(Thm.transitive th0 (c p ax)) |
|
744 |
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
|
745 |
|
46594 | 746 |
val elim_abs = eliminate_construct is_abs |
747 |
(fn p => fn ax => |
|
60642
48dd1cefb4ae
simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
wenzelm
parents:
59586
diff
changeset
|
748 |
Thm.instantiate ([], [(p_v,p), (x_v, Thm.dest_arg ax)]) pth_abs) |
46594 | 749 |
val elim_max = eliminate_construct is_max |
750 |
(fn p => fn ax => |
|
751 |
let val (ax,y) = Thm.dest_comb ax |
|
60642
48dd1cefb4ae
simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
wenzelm
parents:
59586
diff
changeset
|
752 |
in Thm.instantiate ([], [(p_v,p), (x_v, Thm.dest_arg ax), (y_v,y)]) |
46594 | 753 |
pth_max end) |
754 |
val elim_min = eliminate_construct is_min |
|
755 |
(fn p => fn ax => |
|
756 |
let val (ax,y) = Thm.dest_comb ax |
|
60642
48dd1cefb4ae
simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
wenzelm
parents:
59586
diff
changeset
|
757 |
in Thm.instantiate ([], [(p_v,p), (x_v, Thm.dest_arg ax), (y_v,y)]) |
46594 | 758 |
pth_min end) |
759 |
in first_conv [elim_abs, elim_max, elim_min, all_conv] |
|
760 |
end; |
|
761 |
in |
|
762 |
fun gen_real_arith ctxt (mkconst,eq,ge,gt,norm,neg,add,mul,prover) = |
|
763 |
gen_gen_real_arith ctxt |
|
764 |
(mkconst,eq,ge,gt,norm,neg,add,mul, |
|
765 |
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
|
766 |
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
|
767 |
|
46594 | 768 |
(* 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
|
769 |
|
46594 | 770 |
fun gen_prover_real_arith ctxt prover = |
771 |
let |
|
59582 | 772 |
fun simple_cterm_ord t u = Term_Ord.term_ord (Thm.term_of t, Thm.term_of u) = LESS |
46594 | 773 |
val {add, mul, neg, pow = _, sub = _, main} = |
774 |
Semiring_Normalizer.semiring_normalizers_ord_wrapper ctxt |
|
67267 | 775 |
(the (Semiring_Normalizer.match ctxt \<^cterm>\<open>(0::real) + 1\<close>)) |
46594 | 776 |
simple_cterm_ord |
777 |
in gen_real_arith ctxt |
|
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
46594
diff
changeset
|
778 |
(cterm_of_rat, |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
46594
diff
changeset
|
779 |
Numeral_Simprocs.field_comp_conv ctxt, |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
46594
diff
changeset
|
780 |
Numeral_Simprocs.field_comp_conv ctxt, |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
46594
diff
changeset
|
781 |
Numeral_Simprocs.field_comp_conv ctxt, |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
46594
diff
changeset
|
782 |
main ctxt, neg ctxt, add ctxt, mul ctxt, prover) |
46594 | 783 |
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
|
784 |
|
fc654c95c29e
A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff
changeset
|
785 |
end |