src/HOL/Library/positivstellensatz.ML
author wenzelm
Tue, 29 Sep 2009 16:24:36 +0200
changeset 32740 9dd0a2f83429
parent 32646 962b4354ed90
child 32828 ad76967c703d
permissions -rw-r--r--
explicit indication of Unsynchronized.ref;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
     1
(* Title:      Library/Sum_Of_Squares/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
     2
   Author:     Amine Chaieb, University of Cambridge
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
     3
   Description: A generic arithmetic prover based on Positivstellensatz certificates --- 
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
     4
    also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination.
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
     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
     6
fc654c95c29e 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
(* A functor for finite mappings based on Tables *)
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
     8
31120
fc654c95c29e 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
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
    10
sig
fc654c95c29e 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
 type 'a 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
    12
 type key
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
    13
 val apply : 'a T -> key -> '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
    14
 val applyd :'a T -> (key -> 'a) -> key -> '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
    15
 val combine : ('a -> 'a -> 'a) -> ('a -> bool) -> 'a T -> 'a T -> 'a 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
    16
 val defined : 'a T -> key -> 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
    17
 val dom : 'a T -> key list
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
    18
 val fold : (key * 'a -> 'b -> 'b) -> 'a T -> 'b -> '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
    19
 val fold_rev : (key * 'a -> 'b -> 'b) -> 'a T -> 'b -> '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
    20
 val graph : 'a T -> (key * 'a) list
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
    21
 val is_undefined : 'a T -> 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
    22
 val mapf : ('a -> 'b) -> 'a T -> 'b 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
    23
 val tryapplyd : 'a T -> key -> 'a -> '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
    24
 val undefine :  key -> 'a T -> 'a 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
    25
 val undefined : 'a 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
    26
 val update : key * 'a -> 'a T -> 'a 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
    27
 val updatep : (key * 'a -> bool) -> key * 'a -> 'a T -> 'a 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
    28
 val choose : 'a T -> key * '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
    29
 val onefunc : key * 'a -> 'a 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
    30
 val get_first: (key*'a -> 'a option) -> 'a T -> 'a option
fc654c95c29e 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
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
    32
fc654c95c29e 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
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
    34
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
    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
type key = Key.key;
31971
8c1b845ed105 renamed functor TableFun to Table, and GraphFun to Graph;
wenzelm
parents: 31120
diff changeset
    37
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
    38
type 'a T = 'a Tab.table;
fc654c95c29e 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
fc654c95c29e 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
val undefined = Tab.empty;
fc654c95c29e 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
val is_undefined = Tab.is_empty;
fc654c95c29e 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
val mapf = Tab.map;
fc654c95c29e 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
val fold = Tab.fold;
fc654c95c29e 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
val fold_rev = Tab.fold_rev;
fc654c95c29e 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
val graph = Tab.dest;
fc654c95c29e 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
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
    47
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
    48
   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
    49
 | 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
    50
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
    51
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
    52
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
    53
val defined = Tab.defined;
fc654c95c29e 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
fun undefine x t = (Tab.delete x t handle UNDEF => 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
    55
val update = Tab.update;
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
    56
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
    57
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
    58
 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
    59
  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
    60
     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
    61
   | 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
    62
     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
    63
  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
    64
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
    65
fun choose f = case Tab.min_key f 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
    66
   SOME k => (k,valOf (Tab.lookup f k))
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
    67
 | NONE => error "FuncFun.choose : Completely undefined function"
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
    68
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
    69
fun onefunc kv = update kv undefined
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
    70
fc654c95c29e 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
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
    72
fun  find f (k,v) NONE = f (k,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
    73
   | find f (k,v) 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
    74
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
    75
fun get_first f t = fold (find f) t 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
    76
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
    77
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
    78
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
    79
(* Some standard functors and utility functions for them *)
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
    80
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
    81
structure FuncUtil =
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
    82
struct
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
    83
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
    84
fun increasing f ord (x,y) = ord (f x, f y);
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
    85
31120
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
    86
structure Intfunc = FuncFun(type key = int val ord = int_ord);
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
    87
structure Ratfunc = FuncFun(type key = Rat.rat val ord = Rat.ord);
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
    88
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
    89
structure Symfunc = FuncFun(type key = string val ord = fast_string_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
    90
structure Termfunc = FuncFun(type key = term val ord = TermOrd.fast_term_ord);
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
    91
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
    92
val cterm_ord = (fn (s,t) => TermOrd.fast_term_ord(term_of s, term_of t))
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
    93
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
    94
structure Ctermfunc = FuncFun(type key = cterm val ord = cterm_ord);
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
    95
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
    96
type monomial = int Ctermfunc.T;
31120
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
    97
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
    98
fun monomial_ord (m1,m2) = list_ord (prod_ord cterm_ord int_ord) (Ctermfunc.graph m1, Ctermfunc.graph m2)
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
    99
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   100
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
   101
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   102
type poly = Rat.rat Monomialfunc.T;
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   103
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   104
(* 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
   105
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   106
fun dest_monomial mon = sort (increasing fst cterm_ord) (Ctermfunc.graph 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
   107
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   108
fun monomial_order (m1,m2) =
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   109
 if Ctermfunc.is_undefined m2 then LESS 
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   110
 else if Ctermfunc.is_undefined m1 then GREATER 
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   111
 else
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   112
  let val mon1 = dest_monomial m1 
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   113
      val mon2 = dest_monomial m2
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   114
      val deg1 = fold (curry op + o snd) mon1 0
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   115
      val deg2 = fold (curry op + o snd) mon2 0 
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   116
  in if deg1 < deg2 then GREATER else if deg1 > deg2 then LESS
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   117
     else list_ord (prod_ord cterm_ord int_ord) (mon1,mon2)
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   118
  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
   119
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   120
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
   121
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   122
(* 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
   123
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   124
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
   125
sig
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   126
  
31120
fc654c95c29e 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
  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
   128
   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
   129
 | 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
   130
 | 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
   131
 | 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
   132
 | 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
   133
 | Rational_lt of Rat.rat
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   134
 | Square of FuncUtil.poly
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   135
 | 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
   136
 | 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
   137
 | 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
   138
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   139
datatype pss_tree = Trivial | Cert of positivstellensatz | Branch of pss_tree * pss_tree
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   140
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   141
datatype tree_choice = Left | Right
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   142
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   143
type prover = tree_choice list -> 
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   144
  (thm list * thm list * thm list -> positivstellensatz -> thm) ->
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   145
  thm list * thm list * thm list -> thm * pss_tree
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   146
type cert_conv = cterm -> thm * pss_tree
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   147
31120
fc654c95c29e 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
val gen_gen_real_arith :
32646
962b4354ed90 used standard fold function and type aliases
Philipp Meyer
parents: 32645
diff changeset
   149
  Proof.context -> (Rat.rat -> cterm) * conv * conv * conv *
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   150
   conv * conv * conv * conv * conv * conv * prover -> cert_conv
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   151
val real_linear_prover : (thm list * thm list * thm list -> positivstellensatz -> thm) ->
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   152
  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
   153
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   154
val gen_real_arith : Proof.context ->
32646
962b4354ed90 used standard fold function and type aliases
Philipp Meyer
parents: 32645
diff changeset
   155
  (Rat.rat -> cterm) * conv * conv * conv * conv * conv * conv * conv * prover -> cert_conv
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   156
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   157
val gen_prover_real_arith : Proof.context -> prover -> cert_conv
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   158
32646
962b4354ed90 used standard fold function and type aliases
Philipp Meyer
parents: 32645
diff changeset
   159
val is_ratconst : cterm -> bool
962b4354ed90 used standard fold function and type aliases
Philipp Meyer
parents: 32645
diff changeset
   160
val dest_ratconst : cterm -> Rat.rat
962b4354ed90 used standard fold function and type aliases
Philipp Meyer
parents: 32645
diff changeset
   161
val cterm_of_rat : Rat.rat -> cterm
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   162
31120
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   163
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
   164
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   165
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
   166
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
   167
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   168
 open Conv Thm FuncUtil;;
31120
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   169
(* ------------------------------------------------------------------------- *)
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   170
(* 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
   171
(* ------------------------------------------------------------------------- *)
fc654c95c29e 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
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   173
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
   174
   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
   175
 | 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
   176
 | 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
   177
 | 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
   178
 | 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
   179
 | Rational_lt of Rat.rat
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   180
 | Square of FuncUtil.poly
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   181
 | 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
   182
 | 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
   183
 | 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
   184
         (* 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
   185
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   186
datatype pss_tree = Trivial | Cert of positivstellensatz | Branch of pss_tree * pss_tree
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   187
datatype tree_choice = Left | Right
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   188
type prover = tree_choice list -> 
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   189
  (thm list * thm list * thm list -> positivstellensatz -> thm) ->
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   190
  thm list * thm list * thm list -> thm * pss_tree
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   191
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
   192
32740
9dd0a2f83429 explicit indication of Unsynchronized.ref;
wenzelm
parents: 32646
diff changeset
   193
val my_eqs = Unsynchronized.ref ([] : thm list);
9dd0a2f83429 explicit indication of Unsynchronized.ref;
wenzelm
parents: 32646
diff changeset
   194
val my_les = Unsynchronized.ref ([] : thm list);
9dd0a2f83429 explicit indication of Unsynchronized.ref;
wenzelm
parents: 32646
diff changeset
   195
val my_lts = Unsynchronized.ref ([] : thm list);
9dd0a2f83429 explicit indication of Unsynchronized.ref;
wenzelm
parents: 32646
diff changeset
   196
val my_proof = Unsynchronized.ref (Axiom_eq 0);
9dd0a2f83429 explicit indication of Unsynchronized.ref;
wenzelm
parents: 32646
diff changeset
   197
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
   198
32740
9dd0a2f83429 explicit indication of Unsynchronized.ref;
wenzelm
parents: 32646
diff changeset
   199
val my_mk_numeric = Unsynchronized.ref ((K @{cterm True}) :Rat.rat -> cterm);
9dd0a2f83429 explicit indication of Unsynchronized.ref;
wenzelm
parents: 32646
diff changeset
   200
val my_numeric_eq_conv = Unsynchronized.ref no_conv;
9dd0a2f83429 explicit indication of Unsynchronized.ref;
wenzelm
parents: 32646
diff changeset
   201
val my_numeric_ge_conv = Unsynchronized.ref no_conv;
9dd0a2f83429 explicit indication of Unsynchronized.ref;
wenzelm
parents: 32646
diff changeset
   202
val my_numeric_gt_conv = Unsynchronized.ref no_conv;
9dd0a2f83429 explicit indication of Unsynchronized.ref;
wenzelm
parents: 32646
diff changeset
   203
val my_poly_conv = Unsynchronized.ref no_conv;
9dd0a2f83429 explicit indication of Unsynchronized.ref;
wenzelm
parents: 32646
diff changeset
   204
val my_poly_neg_conv = Unsynchronized.ref no_conv;
9dd0a2f83429 explicit indication of Unsynchronized.ref;
wenzelm
parents: 32646
diff changeset
   205
val my_poly_add_conv = Unsynchronized.ref no_conv;
9dd0a2f83429 explicit indication of Unsynchronized.ref;
wenzelm
parents: 32646
diff changeset
   206
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
   207
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   208
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   209
    (* Some useful derived rules *)
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   210
fun deduct_antisym_rule tha thb = 
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   211
    equal_intr (implies_intr (cprop_of thb) tha) 
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   212
     (implies_intr (cprop_of tha) thb);
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   213
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   214
fun prove_hyp tha thb = 
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   215
  if exists (curry op aconv (concl_of tha)) (#hyps (rep_thm thb)) 
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   216
  then equal_elim (symmetric (deduct_antisym_rule tha thb)) tha else thb;
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   217
31120
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   218
fun conjunctions th = case try Conjunction.elim 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
   219
   SOME (th1,th2) => (conjunctions th1) @ conjunctions th2
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   220
 | NONE => [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
   221
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   222
val pth = @{lemma "(((x::real) < y) == (y - x > 0)) &&& ((x <= y) == (y - x >= 0)) 
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   223
     &&& ((x = y) == (x - y = 0)) &&& ((~(x < y)) == (x - y >= 0)) &&& ((~(x <= y)) == (x - y > 0))
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   224
     &&& ((~(x = y)) == (x - y > 0 | -(x - y) > 0))"
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   225
  by (atomize (full), auto simp add: less_diff_eq le_diff_eq not_less)} |> 
fc654c95c29e 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
conjunctions;
fc654c95c29e 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
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   228
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
   229
val pth_add = 
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   230
 @{lemma "(x = (0::real) ==> y = 0 ==> x + y = 0 ) &&& ( x = 0 ==> y >= 0 ==> x + y >= 0) 
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   231
    &&& (x = 0 ==> y > 0 ==> x + y > 0) &&& (x >= 0 ==> y = 0 ==> x + y >= 0) 
fc654c95c29e 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
    &&& (x >= 0 ==> y >= 0 ==> x + y >= 0) &&& (x >= 0 ==> y > 0 ==> x + y > 0) 
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   233
    &&& (x > 0 ==> y = 0 ==> x + y > 0) &&& (x > 0 ==> y >= 0 ==> x + y > 0) 
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   234
    &&& (x > 0 ==> y > 0 ==> x + y > 0)"  by simp_all} |> conjunctions ;
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   235
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   236
val pth_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
   237
  @{lemma "(x = (0::real) ==> y = 0 ==> x * y = 0) &&& (x = 0 ==> y >= 0 ==> x * y = 0) &&& 
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   238
           (x = 0 ==> y > 0 ==> x * y = 0) &&& (x >= 0 ==> y = 0 ==> x * y = 0) &&& 
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   239
           (x >= 0 ==> y >= 0 ==> x * y >= 0 ) &&& ( x >= 0 ==> y > 0 ==> x * y >= 0 ) &&&
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   240
           (x > 0 ==>  y = 0 ==> x * y = 0 ) &&& ( x > 0 ==> y >= 0 ==> x * y >= 0 ) &&&
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   241
           (x > 0 ==>  y > 0 ==> x * y > 0)"
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   242
  by (auto intro: mult_mono[where a="0::real" and b="x" and d="y" and c="0", simplified]
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   243
    mult_strict_mono[where b="x" and d="y" and a="0" and c="0", simplified])} |> conjunctions;
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   244
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   245
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
   246
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
   247
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   248
val weak_dnf_simps = List.take (simp_thms, 34) 
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   249
    @ conjunctions @{lemma "((P & (Q | R)) = ((P&Q) | (P&R))) &&& ((Q | R) & P) = ((Q&P) | (R&P)) &&& (P & Q) = (Q & P) &&& ((P | Q) = (Q | 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
   250
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   251
val nnfD_simps = conjunctions @{lemma "((~(P & Q)) = (~P | ~Q)) &&& ((~(P | Q)) = (~P & ~Q) ) &&& ((P --> Q) = (~P | Q) ) &&& ((P = Q) = ((P & Q) | (~P & ~ Q))) &&& ((~(P = Q)) = ((P & ~ Q) | (~P & Q)) ) &&& ((~ ~(P)) = 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
   252
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   253
val choice_iff = @{lemma "(ALL x. EX y. P x y) = (EX f. ALL x. P x (f x))" by metis};
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   254
val prenex_simps = map (fn th => th RS sym) ([@{thm "all_conj_distrib"}, @{thm "ex_disj_distrib"}] @ @{thms "all_simps"(1-4)} @ @{thms "ex_simps"(1-4)});
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   255
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   256
val real_abs_thms1 = conjunctions @{lemma
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   257
  "((-1 * abs(x::real) >= r) = (-1 * x >= r & 1 * x >= 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
   258
  ((-1 * abs(x) + a >= r) = (a + -1 * x >= r & a + 1 * x >= 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
   259
  ((a + -1 * abs(x) >= r) = (a + -1 * x >= r & a + 1 * x >= 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
   260
  ((a + -1 * abs(x) + b >= r) = (a + -1 * x + b >= r & a + 1 * x + b >= 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
   261
  ((a + b + -1 * abs(x) >= r) = (a + b + -1 * x >= r & a + b + 1 * x >= 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
   262
  ((a + b + -1 * abs(x) + c >= r) = (a + b + -1 * x + c >= r & a + b + 1 * x + c >= 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
   263
  ((-1 * max x y >= r) = (-1 * x >= r & -1 * y >= 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
   264
  ((-1 * max x y + a >= r) = (a + -1 * x >= r & a + -1 * y >= 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
   265
  ((a + -1 * max x y >= r) = (a + -1 * x >= r & a + -1 * y >= 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
   266
  ((a + -1 * max x y + b >= r) = (a + -1 * x + b >= r & a + -1 * y  + b >= 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
   267
  ((a + b + -1 * max x y >= r) = (a + b + -1 * x >= r & a + b + -1 * y >= 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
   268
  ((a + b + -1 * max x y + c >= r) = (a + b + -1 * x + c >= r & a + b + -1 * y  + c >= 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
   269
  ((1 * min x y >= r) = (1 * x >= r & 1 * y >= 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
   270
  ((1 * min x y + a >= r) = (a + 1 * x >= r & a + 1 * y >= 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
   271
  ((a + 1 * min x y >= r) = (a + 1 * x >= r & a + 1 * y >= 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
   272
  ((a + 1 * min x y + b >= r) = (a + 1 * x + b >= r & a + 1 * y  + b >= 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
   273
  ((a + b + 1 * min x y >= r) = (a + b + 1 * x >= r & a + b + 1 * y >= 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
   274
  ((a + b + 1 * min x y + c >= r) = (a + b + 1 * x + c >= r & a + b + 1 * y  + c >= 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
   275
  ((min x y >= r) = (x >= r &  y >= 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
   276
  ((min x y + a >= r) = (a + x >= r & a + y >= 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
   277
  ((a + min x y >= r) = (a + x >= r & a + y >= 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
   278
  ((a + min x y + b >= r) = (a + x + b >= r & a + y  + b >= 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
   279
  ((a + b + min x y >= r) = (a + b + x >= r & a + b + y >= 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
   280
  ((a + b + min x y + c >= r) = (a + b + x + c >= r & a + b + y + c >= 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
   281
  ((-1 * abs(x) > r) = (-1 * x > r & 1 * x > 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
   282
  ((-1 * abs(x) + a > r) = (a + -1 * x > r & a + 1 * x > 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
   283
  ((a + -1 * abs(x) > r) = (a + -1 * x > r & a + 1 * x > 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
   284
  ((a + -1 * abs(x) + b > r) = (a + -1 * x + b > r & a + 1 * x + b > 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
   285
  ((a + b + -1 * abs(x) > r) = (a + b + -1 * x > r & a + b + 1 * x > 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
   286
  ((a + b + -1 * abs(x) + c > r) = (a + b + -1 * x + c > r & a + b + 1 * x + c > 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
   287
  ((-1 * max x y > r) = ((-1 * x > r) & -1 * y > 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
   288
  ((-1 * max x y + a > r) = (a + -1 * x > r & a + -1 * y > 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
   289
  ((a + -1 * max x y > r) = (a + -1 * x > r & a + -1 * y > 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
   290
  ((a + -1 * max x y + b > r) = (a + -1 * x + b > r & a + -1 * y  + b > 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
   291
  ((a + b + -1 * max x y > r) = (a + b + -1 * x > r & a + b + -1 * y > 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
   292
  ((a + b + -1 * max x y + c > r) = (a + b + -1 * x + c > r & a + b + -1 * y  + c > 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
   293
  ((min x y > r) = (x > r &  y > 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
   294
  ((min x y + a > r) = (a + x > r & a + y > 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
   295
  ((a + min x y > r) = (a + x > r & a + y > 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
   296
  ((a + min x y + b > r) = (a + x + b > r & a + y  + b > 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
   297
  ((a + b + min x y > r) = (a + b + x > r & a + b + y > 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
   298
  ((a + b + min x y + c > r) = (a + b + x + c > r & a + b + y + c > 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
   299
  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
   300
fc654c95c29e 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
val abs_split' = @{lemma "P (abs (x::'a::ordered_idom)) == (x >= 0 & P x | x < 0 & 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
   302
  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
   303
fc654c95c29e 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
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
   305
  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
   306
fc654c95c29e 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
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
   308
  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
   309
fc654c95c29e 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
         (* 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
   312
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
   313
 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
   314
  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
   315
   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
   316
 | 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
   317
 | _ => 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
   318
 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
   319
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   320
fun 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
   321
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
   322
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
   323
 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
   324
  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
   325
                   (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
   326
        (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
   327
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
   328
fc654c95c29e 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
  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
   330
   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
   331
 | _ => 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
   332
 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
   333
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   334
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
   335
 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
   336
  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
   337
 | 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
   338
 | _ => 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
   339
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   340
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
   341
 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
   342
  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
   343
 | 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
   344
 | _ => 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
   345
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   346
    (* 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
   347
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
   348
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
   349
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   350
fun 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
   351
  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
   352
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   353
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   354
(* Map back polynomials to HOL.                         *)
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   355
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   356
local
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   357
 open Thm Numeral
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   358
in
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   359
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   360
fun cterm_of_varpow x k = if k = 1 then x else capply (capply @{cterm "op ^ :: real => _"} x) 
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   361
  (mk_cnumber @{ctyp nat} k)
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   362
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   363
fun cterm_of_monomial m = 
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   364
 if Ctermfunc.is_undefined m then @{cterm "1::real"} 
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   365
 else 
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   366
  let 
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   367
   val m' = dest_monomial m
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   368
   val vps = fold_rev (fn (x,k) => cons (cterm_of_varpow x k)) m' [] 
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   369
  in foldr1 (fn (s, t) => capply (capply @{cterm "op * :: real => _"} s) t) vps
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   370
  end
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   371
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   372
fun cterm_of_cmonomial (m,c) = if Ctermfunc.is_undefined m then cterm_of_rat c
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   373
    else if c = Rat.one then cterm_of_monomial m
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   374
    else capply (capply @{cterm "op *::real => _"} (cterm_of_rat c)) (cterm_of_monomial m);
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   375
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   376
fun cterm_of_poly p = 
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   377
 if Monomialfunc.is_undefined p then @{cterm "0::real"} 
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   378
 else
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   379
  let 
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   380
   val cms = map cterm_of_cmonomial
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   381
     (sort (prod_ord monomial_order (K EQUAL)) (Monomialfunc.graph p))
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   382
  in foldr1 (fn (t1, t2) => capply(capply @{cterm "op + :: real => _"} t1) t2) cms
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   383
  end;
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   384
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   385
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
   386
    (* 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
   387
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   388
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
   389
       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
   390
       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
   391
       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
   392
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
   393
 open Conv Thm;
fc654c95c29e 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
 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
   395
 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
   396
          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
   397
          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
   398
          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
   399
 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
   400
 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
   401
 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
   402
 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
   403
 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
   404
 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
   405
 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
   406
 val weak_dnf_conv = Simplifier.rewrite (Simplifier.context ctxt weak_dnf_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
   407
 fun eqT_elim th = equal_elim (symmetric th) @{thm TrueI}
fc654c95c29e 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
 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
   409
  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
   410
  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
   411
       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
   412
     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
   413
  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
   414
fc654c95c29e 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
 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
   416
  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
   417
   val th' = (instantiate (match (lhs_of th, ct)) 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
   418
      handle MATCH => raise CTERM ("real_ineq_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
   419
  in transitive th' (oprconv poly_conv (Thm.rhs_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
   420
  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
   421
  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
   422
       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
   423
       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
   424
  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
   425
   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
   426
     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
   427
      | 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
   428
   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
   429
  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
   430
         (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
   431
  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
   432
         (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
   433
  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
   434
       (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
   435
  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
   436
       (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
   437
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   438
  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
   439
   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
   440
    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
   441
    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
   442
        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
   443
      | 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
   444
      | Axiom_lt n => nth lts 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
   445
      | Rational_eq x => eqT_elim(numeric_eq_conv(capply @{cterm Trueprop} 
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   446
                          (capply (capply @{cterm "op =::real => _"} (mk_numeric 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
   447
                               @{cterm "0::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
   448
      | Rational_le x => eqT_elim(numeric_ge_conv(capply @{cterm Trueprop} 
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   449
                          (capply (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
   450
                                     @{cterm "0::real"}) (mk_numeric 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
   451
      | Rational_lt x => eqT_elim(numeric_gt_conv(capply @{cterm Trueprop} 
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   452
                      (capply (capply @{cterm "op <::real => _"} @{cterm "0::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
   453
                        (mk_numeric x))))
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   454
      | Square pt => square_rule (cterm_of_poly pt)
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   455
      | 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
   456
      | 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
   457
      | 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
   458
   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
   459
          (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
   460
   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
   461
  
fc654c95c29e 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
  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
   463
      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
   464
      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
   465
fc654c95c29e 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
  val concl = dest_arg o cprop_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
   467
  fun is_binop opr ct = (dest_fun2 ct aconvc opr 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
   468
  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
   469
  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
   470
  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
   471
  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
   472
  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
   473
  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
   474
  fun disj_cases th th1 th2 = 
fc654c95c29e 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
   let val (p,q) = dest_binop (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
   476
       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
   477
       val _ = if c aconvc (concl th2) then () else error "disj_cases : conclusions not alpha convertible"
fc654c95c29e 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
   in implies_elim (implies_elim (implies_elim (instantiate' [] (map SOME [p,q,c]) @{thm disjE}) th) (implies_intr (capply @{cterm Trueprop} p) th1)) (implies_intr (capply @{cterm Trueprop} q) th2)
fc654c95c29e 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
   end
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   480
 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
   481
  [] =>
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   482
   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
   483
    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
   484
     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
   485
     val lt = filter (is_gt o concl) nl 
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   486
    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
   487
 | 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
   488
   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
   489
    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
   490
   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
   491
    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
   492
     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
   493
      val (th1,th2) = conj_pair th in
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   494
      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
   495
    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
   496
      let 
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   497
       val (th1, cert1) = overall (Left::cert_choice) dun (assume (capply @{cterm Trueprop} (dest_arg1 ct))::oths)
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   498
       val (th2, cert2) = overall (Right::cert_choice) dun (assume (capply @{cterm Trueprop} (dest_arg ct))::oths)
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   499
      in (disj_cases th th1 th2, Branch (cert1, cert2)) end
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   500
   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
   501
  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
   502
  fun dest_binary b ct = if is_binop b ct then dest_binop 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
   503
                         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
   504
  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
   505
  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
   506
  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
   507
   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
   508
    val (l,r) = dest_eq (dest_arg 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
   509
    val th = instantiate ([],[(@{cpat "?x::real"},l),(@{cpat "?y::real"},r)]) neq_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
   510
    val th_p = poly_conv(dest_arg(dest_arg1(rhs_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
   511
    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
   512
    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
   513
    val th' = Drule.binop_cong_rule @{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
   514
     (Drule.arg_cong_rule (capply @{cterm "op <::real=>_"} @{cterm "0::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
   515
     (Drule.arg_cong_rule (capply @{cterm "op <::real=>_"} @{cterm "0::real"}) th_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
   516
    in transitive 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
   517
  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
   518
 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
   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 = lhs_of 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
   521
  in implies_intr P (equal_elim PQ (assume 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
  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
   523
 (* 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
   524
 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
   525
  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
   526
   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
   527
    Const("Ex",_)$Abs(x,T,p) => h (dest_abs NONE (dest_arg t) |>> (fn v => v::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
   528
  | _ => (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
   529
  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
   530
  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
   531
  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
   532
   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
   533
 | 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
   534
 | _ => "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
   535
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   536
  fun mk_forall x th = Drule.arg_cong_rule (instantiate_cterm' [SOME (ctyp_of_term x)] [] @{cpat "All :: (?'a => bool) => _" }) (abstract_rule (name_of x) x 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
   537
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   538
  val 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
   539
fc654c95c29e 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
 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
   541
 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
   542
fc654c95c29e 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
 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
   544
   @{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
   545
    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
   546
     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
   547
     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
   548
     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
   549
         (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
   550
     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
   551
           (Thm.capply @{cterm Trueprop} (Thm.capply p 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
   552
     val th1 = forall_intr v (implies_intr pv 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
   553
    in implies_elim (implies_elim th0 th) th1  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
   554
 | _ => 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
   555
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   556
  fun simple_choose v 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
   557
     choose v (assume ((Thm.capply @{cterm Trueprop} o mk_ex v) ((Thm.dest_arg o hd o #hyps o Thm.crep_thm) 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
   558
fc654c95c29e 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 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
   560
  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
   561
   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
   562
    Const("All",_)$Abs(x,T,p) => h (dest_abs NONE (dest_arg t) |>> (fn v => v::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
   563
  | _ => (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
   564
  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
   565
  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
   566
fc654c95c29e 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
 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
   568
  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
   569
   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
   570
     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
   571
     literals_conv [@{term "op &"}, @{term "op |"}] [] 
32402
5731300da417 added further conversions and conversionals
boehmes
parents: 31971
diff changeset
   572
     (More_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
   573
       (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
   574
                    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
   575
                    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
   576
  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
   577
                  (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
   578
        try_conv (absconv2 then_conv nnf_norm_conv' then_conv binop_conv absremover)) 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
   579
  val nct = capply @{cterm Trueprop} (capply @{cterm "Not"} 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
   580
  val th0 = (init_conv then_conv arg_conv nnf_norm_conv') nct
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   581
  val tm0 = dest_arg (rhs_of th0)
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   582
  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
   583
   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
   584
    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
   585
    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
   586
    val th1 = Drule.arg_cong_rule @{cterm Trueprop} (fold mk_forall avs (absremover ibod))
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   587
    val (th2, certs) = overall [] [] [specl avs (assume (rhs_of th1))]
31120
fc654c95c29e 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 th3 = fold simple_choose evs (prove_hyp (equal_elim th1 (assume (capply @{cterm Trueprop} bod))) th2)
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   589
   in (Drule.implies_intr_hyps (prove_hyp (equal_elim th0 (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
   590
   end
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   591
  in (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
   592
 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
   593
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
   594
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
   595
fc654c95c29e 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
(* 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
   597
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
   598
  val linear_add = Ctermfunc.combine (curry op +/) (fn z => z =/ Rat.zero)
fc654c95c29e 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
  fun linear_cmul c = Ctermfunc.mapf (fn x => c */ 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
   600
  val one_tm = @{cterm "1::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
   601
  fun contradictory p (e,_) = ((Ctermfunc.is_undefined e) andalso not(p Rat.zero)) orelse
fc654c95c29e 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
     ((gen_eq_set (op aconvc) (Ctermfunc.dom e, [one_tm])) andalso not(p(Ctermfunc.apply e one_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
   603
fc654c95c29e 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
  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
   605
   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
   606
    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
   607
  | 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
   608
   (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
   609
     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
   610
   | 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
   611
     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
   612
     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
   613
      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
   614
      fun blowup 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
   615
       length(filter (fn (e,_) => Ctermfunc.tryapplyd e v Rat.zero =/ Rat.zero) ineqs) +
fc654c95c29e 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
       length(filter (fn (e,_) => Ctermfunc.tryapplyd e v Rat.zero >/ Rat.zero) ineqs) *
fc654c95c29e 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
       length(filter (fn (e,_) => Ctermfunc.tryapplyd e v Rat.zero </ Rat.zero) ineqs)
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   618
      val  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
   619
                 (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
   620
      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
   621
       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
   622
        val c1 = Ctermfunc.tryapplyd e1 v Rat.zero 
fc654c95c29e 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
        val c2 = Ctermfunc.tryapplyd e2 v Rat.zero
fc654c95c29e 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
       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
   625
        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
   626
         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
   627
         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
   628
         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
   629
         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
   630
        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
   631
        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
   632
       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
   633
      val (les0,les1) = 
fc654c95c29e 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
         List.partition (fn (e,_) => Ctermfunc.tryapplyd e v Rat.zero =/ Rat.zero) les
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   635
      val (lts0,lts1) = 
fc654c95c29e 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
         List.partition (fn (e,_) => Ctermfunc.tryapplyd e v Rat.zero =/ Rat.zero) 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
   637
      val (lesp,lesn) = 
fc654c95c29e 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
         List.partition (fn (e,_) => Ctermfunc.tryapplyd e v Rat.zero >/ Rat.zero) les1
fc654c95c29e 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 (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
   640
         List.partition (fn (e,_) => Ctermfunc.tryapplyd e v Rat.zero >/ Rat.zero) lts1
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   641
      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
   642
      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
   643
                      (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
   644
     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
   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
fc654c95c29e 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
  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
   648
   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
   649
    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
   650
  | 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
   651
    [] => 
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   652
     let val vars = remove (op aconvc) one_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
   653
           (fold_rev (curry (gen_union (op aconvc)) o Ctermfunc.dom o fst) (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
   654
     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
   655
   | (e,p)::es => 
fc654c95c29e 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
     if Ctermfunc.is_undefined e then linear_eqs (es,les,lts) 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
   657
     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
   658
      val (x,c) = Ctermfunc.choose (Ctermfunc.undefine one_tm 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
   659
      fun xform (inp as (t,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
   660
       let val d = Ctermfunc.tryapplyd t x Rat.zero 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
   661
        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
   662
        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
   663
         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
   664
         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
   665
         val t' = linear_cmul (Rat.abs c) t
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   666
         val p' = Eqmul(Monomialfunc.onefunc (Ctermfunc.undefined, 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
   667
         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
   668
        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
   669
        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
   670
      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
   671
     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
   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 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
   675
   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
   676
    val eqs = map2 (fn p => fn n => (p,Axiom_eq n)) eq (0 upto (length eq - 1))
fc654c95c29e 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
    val les = map2 (fn p => fn n => (p,Axiom_le n)) le (0 upto (length le - 1))
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   678
    val lts = map2 (fn p => fn n => (p,Axiom_lt n)) lt (0 upto (length lt - 1))
fc654c95c29e 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
   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
   680
   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
   681
  
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   682
  fun lin_of_hol 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
   683
   if ct aconvc @{cterm "0::real"} then Ctermfunc.undefined
fc654c95c29e 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
   else if not (is_comb ct) then Ctermfunc.onefunc (ct, Rat.one)
fc654c95c29e 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
   else if is_ratconst ct then Ctermfunc.onefunc (one_tm, dest_ratconst 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
   686
   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
   687
    let val (lop,r) = Thm.dest_comb 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
   688
    in if not (is_comb lop) then Ctermfunc.onefunc (ct, Rat.one)
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   689
       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
   690
        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
   691
        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
   692
           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
   693
           else 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
   694
                   andalso is_ratconst l then Ctermfunc.onefunc (r, dest_ratconst l)
fc654c95c29e 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
           else Ctermfunc.onefunc (ct, Rat.one)
fc654c95c29e 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
        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
   697
    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
   698
fc654c95c29e 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
  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
   700
   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
   701
     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
   702
  | _ => 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
   703
 open Thm
fc654c95c29e 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
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
   705
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
   706
 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
   707
  val lhs = lin_of_hol o dest_arg1 o dest_arg o cprop_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
   708
  val rhs = lin_of_hol o dest_arg o dest_arg o cprop_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
   709
  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
   710
  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
   711
  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
   712
  val aliens =  filter is_alien
fc654c95c29e 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
      (fold_rev (curry (gen_union (op aconvc)) o Ctermfunc.dom) 
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   714
          (eq_pols @ le_pols @ lt_pols) [])
fc654c95c29e 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
  val le_pols' = le_pols @ map (fn v => Ctermfunc.onefunc (v,Rat.one)) aliens
fc654c95c29e 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
  val (_,proof) = linear_prover (eq_pols,le_pols',lt_pols)
fc654c95c29e 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 le' = le @ map (fn a => instantiate' [] [SOME (dest_arg a)] @{thm real_of_nat_ge_zero}) aliens 
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   718
 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
   719
 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
   720
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
   721
fc654c95c29e 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
(* 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
   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
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
   725
 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
   726
 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
   727
    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
   728
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   729
 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
   730
  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
   731
   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
   732
   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
   733
   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
   734
   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
   735
   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
   736
   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
   737
   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
   738
   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
   739
   val is_min = is_binop @{cterm "min :: 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
   740
   fun is_abs t = is_comb t andalso dest_fun t aconvc abs_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
   741
   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
   742
    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
   743
     val t = find_cterm p 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
   744
     val th0 = (symmetric o beta_conversion false) (capply (cabs t tm) 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
   745
     val (p,ax) = (dest_comb o rhs_of) th0
fc654c95c29e 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
    in fconv_rule(arg_conv(binop_conv (arg_conv (beta_conversion 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
   747
               (transitive th0 (c p ax))
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   748
   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
   749
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   750
   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
   751
    (fn p => fn ax => 
fc654c95c29e 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
       instantiate ([], [(p_tm,p), (x_tm, dest_arg ax)]) pth_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
   753
   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
   754
    (fn p => fn ax => 
fc654c95c29e 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
      let val (ax,y) = dest_comb ax 
fc654c95c29e 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
      in  instantiate ([], [(p_tm,p), (x_tm, dest_arg ax), (y_tm,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
   757
      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
   758
   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
   759
    (fn p => fn ax => 
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   760
      let val (ax,y) = dest_comb ax 
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   761
      in  instantiate ([], [(p_tm,p), (x_tm, dest_arg ax), (y_tm,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
   762
      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
   763
   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
   764
  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
   765
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
   766
        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
   767
                       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
   768
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
   769
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   770
(* 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
   771
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   772
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
   773
 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
   774
  fun simple_cterm_ord t u = TermOrd.term_ord (term_of t, term_of u) = LESS
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   775
  val {add,mul,neg,pow,sub,main} = 
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   776
     Normalizer.semiring_normalizers_ord_wrapper 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
   777
      (valOf (NormalizerData.match ctxt @{cterm "(0::real) + 1"})) 
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   778
     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
   779
in gen_real_arith 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
   780
   (cterm_of_rat, field_comp_conv, field_comp_conv,field_comp_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
   781
    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
   782
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
   783
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   784
end