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