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