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