src/HOL/Library/Sum_of_Squares/positivstellensatz.ML
author wenzelm
Fri, 29 Oct 2021 12:42:06 +0200
changeset 74622 9568db8f4882
parent 74621 82ff15b980e9
child 74623 9b1d33c7bbcc
permissions -rw-r--r--
clarified antiquotations: no need for Thm.beta_conversion, due to Thm.instantiate_beta;
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
58628
fd3c96a8ca60 tuned spelling;
wenzelm
parents: 52049
diff changeset
     5
--- also implements Fourier-Motzkin elimination as a special case
fd3c96a8ca60 tuned spelling;
wenzelm
parents: 52049
diff changeset
     6
Fourier-Motzkin elimination.
31120
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
     7
*)
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
     8
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
     9
(* A functor for finite mappings based on Tables *)
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
    10
46594
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
52049
156e12d5cb92 tuned signature;
wenzelm
parents: 51717
diff changeset
    47
fun choose f =
156e12d5cb92 tuned signature;
wenzelm
parents: 51717
diff changeset
    48
  (case Tab.min f of
156e12d5cb92 tuned signature;
wenzelm
parents: 51717
diff changeset
    49
    SOME entry => entry
156e12d5cb92 tuned signature;
wenzelm
parents: 51717
diff changeset
    50
  | NONE => error "FuncFun.choose : Completely empty function")
31120
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
    51
32829
671eb46eb0a3 tuned FuncFun and FuncUtil structure in positivstellensatz.ML
Philipp Meyer
parents: 32828
diff changeset
    52
fun onefunc kv = update kv empty
671eb46eb0a3 tuned FuncFun and FuncUtil structure in positivstellensatz.ML
Philipp Meyer
parents: 32828
diff changeset
    53
31120
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
    54
end;
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
    55
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
    56
(* Some standard functors and utility functions for them *)
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
    57
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
    58
structure FuncUtil =
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
    59
struct
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
    60
31120
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
    61
structure Intfunc = FuncFun(type key = int val ord = int_ord);
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
    62
structure Ratfunc = FuncFun(type key = Rat.rat val ord = Rat.ord);
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
    63
structure Intpairfunc = FuncFun(type key = int*int val ord = prod_ord int_ord int_ord);
31120
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
    64
structure Symfunc = FuncFun(type key = string val ord = fast_string_ord);
35408
b48ab741683b modernized structure Term_Ord;
wenzelm
parents: 35028
diff changeset
    65
structure Termfunc = FuncFun(type key = term val ord = Term_Ord.fast_term_ord);
67559
833d154ab189 clarified signature;
wenzelm
parents: 67399
diff changeset
    66
structure Ctermfunc = FuncFun(type key = cterm val ord = Thm.fast_term_ord);
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
    67
32829
671eb46eb0a3 tuned FuncFun and FuncUtil structure in positivstellensatz.ML
Philipp Meyer
parents: 32828
diff changeset
    68
type monomial = int Ctermfunc.table;
67559
833d154ab189 clarified signature;
wenzelm
parents: 67399
diff changeset
    69
val monomial_ord = list_ord (prod_ord Thm.fast_term_ord int_ord) o apply2 Ctermfunc.dest
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
    70
structure Monomialfunc = FuncFun(type key = monomial val ord = monomial_ord)
31120
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
    71
32829
671eb46eb0a3 tuned FuncFun and FuncUtil structure in positivstellensatz.ML
Philipp Meyer
parents: 32828
diff changeset
    72
type poly = Rat.rat Monomialfunc.table;
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
(* The ordering so we can create canonical HOL polynomials.                  *)
31120
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
    75
67559
833d154ab189 clarified signature;
wenzelm
parents: 67399
diff changeset
    76
fun dest_monomial mon = sort (Thm.fast_term_ord o apply2 fst) (Ctermfunc.dest mon);
31120
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
    77
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
    78
fun monomial_order (m1,m2) =
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
    79
  if Ctermfunc.is_empty m2 then LESS
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
    80
  else if Ctermfunc.is_empty m1 then GREATER
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
    81
  else
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
    82
    let
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
    83
      val mon1 = dest_monomial m1
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
    84
      val mon2 = dest_monomial m2
33002
f3f02f36a3e2 uniform use of Integer.add/mult/sum/prod;
wenzelm
parents: 32843
diff changeset
    85
      val deg1 = fold (Integer.add o snd) mon1 0
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
    86
      val deg2 = fold (Integer.add o snd) mon2 0
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
    87
    in if deg1 < deg2 then GREATER
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
    88
       else if deg1 > deg2 then LESS
67559
833d154ab189 clarified signature;
wenzelm
parents: 67399
diff changeset
    89
       else list_ord (prod_ord Thm.fast_term_ord int_ord) (mon1,mon2)
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
    90
    end;
31120
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
    91
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
    92
end
31120
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
    93
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
    94
(* positivstellensatz datatype and prover generation *)
31120
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
    95
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
    96
signature REAL_ARITH =
31120
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
    97
sig
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
    98
31120
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
    99
  datatype positivstellensatz =
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   100
    Axiom_eq of int
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   101
  | Axiom_le of int
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   102
  | Axiom_lt of int
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   103
  | Rational_eq of Rat.rat
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   104
  | Rational_le of Rat.rat
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   105
  | Rational_lt of Rat.rat
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   106
  | Square of FuncUtil.poly
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   107
  | Eqmul of FuncUtil.poly * positivstellensatz
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   108
  | Sum of positivstellensatz * positivstellensatz
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   109
  | Product of positivstellensatz * positivstellensatz;
31120
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   110
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   111
  datatype pss_tree = Trivial | Cert of positivstellensatz | Branch of pss_tree * pss_tree
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   112
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   113
  datatype tree_choice = Left | Right
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   114
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   115
  type prover = tree_choice list ->
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   116
    (thm list * thm list * thm list -> positivstellensatz -> thm) ->
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   117
      thm list * thm list * thm list -> thm * pss_tree
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   118
  type cert_conv = cterm -> thm * pss_tree
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   119
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   120
  val gen_gen_real_arith :
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   121
    Proof.context -> (Rat.rat -> cterm) * conv * conv * conv *
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   122
     conv * conv * conv * conv * conv * conv * prover -> cert_conv
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   123
  val real_linear_prover : (thm list * thm list * thm list -> positivstellensatz -> thm) ->
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   124
    thm list * thm list * thm list -> thm * pss_tree
31120
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   125
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   126
  val gen_real_arith : Proof.context ->
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   127
    (Rat.rat -> cterm) * conv * conv * conv * conv * conv * conv * conv * prover -> cert_conv
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   128
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   129
  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
   130
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   131
  val is_ratconst : cterm -> bool
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   132
  val dest_ratconst : cterm -> Rat.rat
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   133
  val cterm_of_rat : Rat.rat -> cterm
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   134
31120
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   135
end
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   136
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   137
structure RealArith : REAL_ARITH =
31120
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   138
struct
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   139
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   140
open Conv
31120
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   141
(* ------------------------------------------------------------------------- *)
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   142
(* Data structure for Positivstellensatz refutations.                        *)
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   143
(* ------------------------------------------------------------------------- *)
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   144
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   145
datatype positivstellensatz =
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   146
    Axiom_eq of int
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   147
  | Axiom_le of int
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   148
  | Axiom_lt of int
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   149
  | Rational_eq of Rat.rat
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   150
  | Rational_le of Rat.rat
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   151
  | Rational_lt of Rat.rat
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   152
  | Square of FuncUtil.poly
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   153
  | Eqmul of FuncUtil.poly * positivstellensatz
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   154
  | Sum of positivstellensatz * positivstellensatz
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   155
  | Product of positivstellensatz * positivstellensatz;
31120
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   156
         (* Theorems used in the procedure *)
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   157
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   158
datatype pss_tree = Trivial | Cert of positivstellensatz | Branch of pss_tree * pss_tree
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   159
datatype tree_choice = Left | Right
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   160
type prover = tree_choice list ->
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   161
  (thm list * thm list * thm list -> positivstellensatz -> thm) ->
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   162
    thm list * thm list * thm list -> thm * pss_tree
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   163
type cert_conv = cterm -> thm * pss_tree
31120
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   164
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   165
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   166
    (* Some useful derived rules *)
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   167
fun deduct_antisym_rule tha thb =
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59058
diff changeset
   168
    Thm.equal_intr (Thm.implies_intr (Thm.cprop_of thb) tha)
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59058
diff changeset
   169
     (Thm.implies_intr (Thm.cprop_of tha) thb);
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   170
44058
ae85c5d64913 misc tuning -- eliminated old-fashioned rep_thm;
wenzelm
parents: 39920
diff changeset
   171
fun prove_hyp tha thb =
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59058
diff changeset
   172
  if exists (curry op aconv (Thm.concl_of tha)) (Thm.hyps_of thb)  (* FIXME !? *)
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36753
diff changeset
   173
  then Thm.equal_elim (Thm.symmetric (deduct_antisym_rule tha thb)) tha else thb;
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   174
67091
1393c2340eec more symbols;
wenzelm
parents: 63667
diff changeset
   175
val pth = @{lemma "(((x::real) < y) \<equiv> (y - x > 0))" and "((x \<le> y) \<equiv> (y - x \<ge> 0))" and
1393c2340eec more symbols;
wenzelm
parents: 63667
diff changeset
   176
     "((x = y) \<equiv> (x - y = 0))" and "((\<not>(x < y)) \<equiv> (x - y \<ge> 0))" and
74615
9af55a51e185 clarified antiquotations;
wenzelm
parents: 74608
diff changeset
   177
     "((\<not>(x \<le> y)) \<equiv> (x - y > 0))"
33443
b9bbd0f3dcdb tuned header;
wenzelm
parents: 33063
diff changeset
   178
  by (atomize (full), auto simp add: less_diff_eq le_diff_eq not_less)};
31120
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   179
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   180
val pth_add =
67267
c5994f1fa0fa more symbols;
wenzelm
parents: 67232
diff changeset
   181
  @{lemma "(x = (0::real) \<Longrightarrow> y = 0 \<Longrightarrow> x + y = 0 )" and "( x = 0 \<Longrightarrow> y \<ge> 0 \<Longrightarrow> x + y \<ge> 0)" and
c5994f1fa0fa more symbols;
wenzelm
parents: 67232
diff changeset
   182
    "(x = 0 \<Longrightarrow> y > 0 \<Longrightarrow> x + y > 0)" and "(x \<ge> 0 \<Longrightarrow> y = 0 \<Longrightarrow> x + y \<ge> 0)" and
c5994f1fa0fa more symbols;
wenzelm
parents: 67232
diff changeset
   183
    "(x \<ge> 0 \<Longrightarrow> y \<ge> 0 \<Longrightarrow> x + y \<ge> 0)" and "(x \<ge> 0 \<Longrightarrow> y > 0 \<Longrightarrow> x + y > 0)" and
c5994f1fa0fa more symbols;
wenzelm
parents: 67232
diff changeset
   184
    "(x > 0 \<Longrightarrow> y = 0 \<Longrightarrow> x + y > 0)" and "(x > 0 \<Longrightarrow> y \<ge> 0 \<Longrightarrow> x + y > 0)" and
c5994f1fa0fa more symbols;
wenzelm
parents: 67232
diff changeset
   185
    "(x > 0 \<Longrightarrow> y > 0 \<Longrightarrow> x + y > 0)" by simp_all};
31120
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   186
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   187
val pth_mul =
67267
c5994f1fa0fa more symbols;
wenzelm
parents: 67232
diff changeset
   188
  @{lemma "(x = (0::real) \<Longrightarrow> y = 0 \<Longrightarrow> x * y = 0)" and "(x = 0 \<Longrightarrow> y \<ge> 0 \<Longrightarrow> x * y = 0)" and
c5994f1fa0fa more symbols;
wenzelm
parents: 67232
diff changeset
   189
    "(x = 0 \<Longrightarrow> y > 0 \<Longrightarrow> x * y = 0)" and "(x \<ge> 0 \<Longrightarrow> y = 0 \<Longrightarrow> x * y = 0)" and
c5994f1fa0fa more symbols;
wenzelm
parents: 67232
diff changeset
   190
    "(x \<ge> 0 \<Longrightarrow> y \<ge> 0 \<Longrightarrow> x * y \<ge> 0)" and "(x \<ge> 0 \<Longrightarrow> y > 0 \<Longrightarrow> x * y \<ge> 0)" and
c5994f1fa0fa more symbols;
wenzelm
parents: 67232
diff changeset
   191
    "(x > 0 \<Longrightarrow>  y = 0 \<Longrightarrow> x * y = 0)" and "(x > 0 \<Longrightarrow> y \<ge> 0 \<Longrightarrow> x * y \<ge> 0)" and
c5994f1fa0fa more symbols;
wenzelm
parents: 67232
diff changeset
   192
    "(x > 0 \<Longrightarrow>  y > 0 \<Longrightarrow> x * y > 0)"
31120
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   193
  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
   194
    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
   195
67267
c5994f1fa0fa more symbols;
wenzelm
parents: 67232
diff changeset
   196
val pth_emul = @{lemma "y = (0::real) \<Longrightarrow> x * y = 0"  by simp};
31120
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   197
33443
b9bbd0f3dcdb tuned header;
wenzelm
parents: 33063
diff changeset
   198
val weak_dnf_simps =
45654
cf10bde35973 more antiquotations;
wenzelm
parents: 44454
diff changeset
   199
  List.take (@{thms simp_thms}, 34) @
67091
1393c2340eec more symbols;
wenzelm
parents: 63667
diff changeset
   200
    @{lemma "((P \<and> (Q \<or> R)) = ((P\<and>Q) \<or> (P\<and>R)))" and "((Q \<or> R) \<and> P) = ((Q\<and>P) \<or> (R\<and>P))" and
1393c2340eec more symbols;
wenzelm
parents: 63667
diff changeset
   201
      "(P \<and> Q) = (Q \<and> P)" and "((P \<or> Q) = (Q \<or> P))" by blast+};
31120
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   202
44454
6f28f96a09bf avoid warnings
huffman
parents: 44058
diff changeset
   203
(*
33443
b9bbd0f3dcdb tuned header;
wenzelm
parents: 33063
diff changeset
   204
val nnfD_simps =
b9bbd0f3dcdb tuned header;
wenzelm
parents: 33063
diff changeset
   205
  @{lemma "((~(P & Q)) = (~P | ~Q))" and "((~(P | Q)) = (~P & ~Q) )" and
b9bbd0f3dcdb tuned header;
wenzelm
parents: 33063
diff changeset
   206
    "((P --> Q) = (~P | Q) )" and "((P = Q) = ((P & Q) | (~P & ~ Q)))" and
b9bbd0f3dcdb tuned header;
wenzelm
parents: 33063
diff changeset
   207
    "((~(P = Q)) = ((P & ~ Q) | (~P & Q)) )" and "((~ ~(P)) = P)" by blast+};
44454
6f28f96a09bf avoid warnings
huffman
parents: 44058
diff changeset
   208
*)
31120
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   209
67091
1393c2340eec more symbols;
wenzelm
parents: 63667
diff changeset
   210
val choice_iff = @{lemma "(\<forall>x. \<exists>y. P x y) = (\<exists>f. \<forall>x. P x (f x))" by metis};
33443
b9bbd0f3dcdb tuned header;
wenzelm
parents: 33063
diff changeset
   211
val prenex_simps =
b9bbd0f3dcdb tuned header;
wenzelm
parents: 33063
diff changeset
   212
  map (fn th => th RS sym)
b9bbd0f3dcdb tuned header;
wenzelm
parents: 33063
diff changeset
   213
    ([@{thm "all_conj_distrib"}, @{thm "ex_disj_distrib"}] @
37598
893dcabf0c04 explicit is better than implicit
haftmann
parents: 37117
diff changeset
   214
      @{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
   215
33443
b9bbd0f3dcdb tuned header;
wenzelm
parents: 33063
diff changeset
   216
val real_abs_thms1 = @{lemma
67091
1393c2340eec more symbols;
wenzelm
parents: 63667
diff changeset
   217
  "((-1 * \<bar>x::real\<bar> \<ge> r) = (-1 * x \<ge> r \<and> 1 * x \<ge> r))" and
1393c2340eec more symbols;
wenzelm
parents: 63667
diff changeset
   218
  "((-1 * \<bar>x\<bar> + a \<ge> r) = (a + -1 * x \<ge> r \<and> a + 1 * x \<ge> r))" and
1393c2340eec more symbols;
wenzelm
parents: 63667
diff changeset
   219
  "((a + -1 * \<bar>x\<bar> \<ge> r) = (a + -1 * x \<ge> r \<and> a + 1 * x \<ge> r))" and
1393c2340eec more symbols;
wenzelm
parents: 63667
diff changeset
   220
  "((a + -1 * \<bar>x\<bar> + b \<ge> r) = (a + -1 * x + b \<ge> r \<and> a + 1 * x + b \<ge> r))" and
1393c2340eec more symbols;
wenzelm
parents: 63667
diff changeset
   221
  "((a + b + -1 * \<bar>x\<bar> \<ge> r) = (a + b + -1 * x \<ge> r \<and> a + b + 1 * x \<ge> r))" and
1393c2340eec more symbols;
wenzelm
parents: 63667
diff changeset
   222
  "((a + b + -1 * \<bar>x\<bar> + c \<ge> r) = (a + b + -1 * x + c \<ge> r \<and> a + b + 1 * x + c \<ge> r))" and
1393c2340eec more symbols;
wenzelm
parents: 63667
diff changeset
   223
  "((-1 * max x y \<ge> r) = (-1 * x \<ge> r \<and> -1 * y \<ge> r))" and
1393c2340eec more symbols;
wenzelm
parents: 63667
diff changeset
   224
  "((-1 * max x y + a \<ge> r) = (a + -1 * x \<ge> r \<and> a + -1 * y \<ge> r))" and
1393c2340eec more symbols;
wenzelm
parents: 63667
diff changeset
   225
  "((a + -1 * max x y \<ge> r) = (a + -1 * x \<ge> r \<and> a + -1 * y \<ge> r))" and
1393c2340eec more symbols;
wenzelm
parents: 63667
diff changeset
   226
  "((a + -1 * max x y + b \<ge> r) = (a + -1 * x + b \<ge> r \<and> a + -1 * y  + b \<ge> r))" and
1393c2340eec more symbols;
wenzelm
parents: 63667
diff changeset
   227
  "((a + b + -1 * max x y \<ge> r) = (a + b + -1 * x \<ge> r \<and> a + b + -1 * y \<ge> r))" and
1393c2340eec more symbols;
wenzelm
parents: 63667
diff changeset
   228
  "((a + b + -1 * max x y + c \<ge> r) = (a + b + -1 * x + c \<ge> r \<and> a + b + -1 * y  + c \<ge> r))" and
1393c2340eec more symbols;
wenzelm
parents: 63667
diff changeset
   229
  "((1 * min x y \<ge> r) = (1 * x \<ge> r \<and> 1 * y \<ge> r))" and
1393c2340eec more symbols;
wenzelm
parents: 63667
diff changeset
   230
  "((1 * min x y + a \<ge> r) = (a + 1 * x \<ge> r \<and> a + 1 * y \<ge> r))" and
1393c2340eec more symbols;
wenzelm
parents: 63667
diff changeset
   231
  "((a + 1 * min x y \<ge> r) = (a + 1 * x \<ge> r \<and> a + 1 * y \<ge> r))" and
1393c2340eec more symbols;
wenzelm
parents: 63667
diff changeset
   232
  "((a + 1 * min x y + b \<ge> r) = (a + 1 * x + b \<ge> r \<and> a + 1 * y  + b \<ge> r))" and
1393c2340eec more symbols;
wenzelm
parents: 63667
diff changeset
   233
  "((a + b + 1 * min x y \<ge> r) = (a + b + 1 * x \<ge> r \<and> a + b + 1 * y \<ge> r))" and
1393c2340eec more symbols;
wenzelm
parents: 63667
diff changeset
   234
  "((a + b + 1 * min x y + c \<ge> r) = (a + b + 1 * x + c \<ge> r \<and> a + b + 1 * y  + c \<ge> r))" and
1393c2340eec more symbols;
wenzelm
parents: 63667
diff changeset
   235
  "((min x y \<ge> r) = (x \<ge> r \<and> y \<ge> r))" and
1393c2340eec more symbols;
wenzelm
parents: 63667
diff changeset
   236
  "((min x y + a \<ge> r) = (a + x \<ge> r \<and> a + y \<ge> r))" and
1393c2340eec more symbols;
wenzelm
parents: 63667
diff changeset
   237
  "((a + min x y \<ge> r) = (a + x \<ge> r \<and> a + y \<ge> r))" and
1393c2340eec more symbols;
wenzelm
parents: 63667
diff changeset
   238
  "((a + min x y + b \<ge> r) = (a + x + b \<ge> r \<and> a + y  + b \<ge> r))" and
1393c2340eec more symbols;
wenzelm
parents: 63667
diff changeset
   239
  "((a + b + min x y \<ge> r) = (a + b + x \<ge> r \<and> a + b + y \<ge> r))" and
1393c2340eec more symbols;
wenzelm
parents: 63667
diff changeset
   240
  "((a + b + min x y + c \<ge> r) = (a + b + x + c \<ge> r \<and> a + b + y + c \<ge> r))" and
1393c2340eec more symbols;
wenzelm
parents: 63667
diff changeset
   241
  "((-1 * \<bar>x\<bar> > r) = (-1 * x > r \<and> 1 * x > r))" and
1393c2340eec more symbols;
wenzelm
parents: 63667
diff changeset
   242
  "((-1 * \<bar>x\<bar> + a > r) = (a + -1 * x > r \<and> a + 1 * x > r))" and
1393c2340eec more symbols;
wenzelm
parents: 63667
diff changeset
   243
  "((a + -1 * \<bar>x\<bar> > r) = (a + -1 * x > r \<and> a + 1 * x > r))" and
1393c2340eec more symbols;
wenzelm
parents: 63667
diff changeset
   244
  "((a + -1 * \<bar>x\<bar> + b > r) = (a + -1 * x + b > r \<and> a + 1 * x + b > r))" and
1393c2340eec more symbols;
wenzelm
parents: 63667
diff changeset
   245
  "((a + b + -1 * \<bar>x\<bar> > r) = (a + b + -1 * x > r \<and> a + b + 1 * x > r))" and
1393c2340eec more symbols;
wenzelm
parents: 63667
diff changeset
   246
  "((a + b + -1 * \<bar>x\<bar> + c > r) = (a + b + -1 * x + c > r \<and> a + b + 1 * x + c > r))" and
1393c2340eec more symbols;
wenzelm
parents: 63667
diff changeset
   247
  "((-1 * max x y > r) = ((-1 * x > r) \<and> -1 * y > r))" and
1393c2340eec more symbols;
wenzelm
parents: 63667
diff changeset
   248
  "((-1 * max x y + a > r) = (a + -1 * x > r \<and> a + -1 * y > r))" and
1393c2340eec more symbols;
wenzelm
parents: 63667
diff changeset
   249
  "((a + -1 * max x y > r) = (a + -1 * x > r \<and> a + -1 * y > r))" and
1393c2340eec more symbols;
wenzelm
parents: 63667
diff changeset
   250
  "((a + -1 * max x y + b > r) = (a + -1 * x + b > r \<and> a + -1 * y  + b > r))" and
1393c2340eec more symbols;
wenzelm
parents: 63667
diff changeset
   251
  "((a + b + -1 * max x y > r) = (a + b + -1 * x > r \<and> a + b + -1 * y > r))" and
1393c2340eec more symbols;
wenzelm
parents: 63667
diff changeset
   252
  "((a + b + -1 * max x y + c > r) = (a + b + -1 * x + c > r \<and> a + b + -1 * y  + c > r))" and
1393c2340eec more symbols;
wenzelm
parents: 63667
diff changeset
   253
  "((min x y > r) = (x > r \<and> y > r))" and
1393c2340eec more symbols;
wenzelm
parents: 63667
diff changeset
   254
  "((min x y + a > r) = (a + x > r \<and> a + y > r))" and
1393c2340eec more symbols;
wenzelm
parents: 63667
diff changeset
   255
  "((a + min x y > r) = (a + x > r \<and> a + y > r))" and
1393c2340eec more symbols;
wenzelm
parents: 63667
diff changeset
   256
  "((a + min x y + b > r) = (a + x + b > r \<and> a + y  + b > r))" and
1393c2340eec more symbols;
wenzelm
parents: 63667
diff changeset
   257
  "((a + b + min x y > r) = (a + b + x > r \<and> a + b + y > r))" and
1393c2340eec more symbols;
wenzelm
parents: 63667
diff changeset
   258
  "((a + b + min x y + c > r) = (a + b + x + c > r \<and> a + b + y + c > r))"
31120
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   259
  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
   260
74617
wenzelm
parents: 74616
diff changeset
   261
val pth_abs =
wenzelm
parents: 74616
diff changeset
   262
  @{lemma "P \<bar>x\<bar> \<equiv> (x \<ge> 0 \<and> P x \<or> x < 0 \<and> P (-x))" for x :: real
wenzelm
parents: 74616
diff changeset
   263
    by (atomize (full)) (auto split: abs_split)}
31120
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   264
74617
wenzelm
parents: 74616
diff changeset
   265
val pth_max =
wenzelm
parents: 74616
diff changeset
   266
  @{lemma "P (max x y) \<equiv> (x \<le> y \<and> P y \<or> x > y \<and> P x)" for x y :: real
wenzelm
parents: 74616
diff changeset
   267
    by (atomize (full)) (auto simp add: max_def)}
31120
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   268
74617
wenzelm
parents: 74616
diff changeset
   269
val pth_min =
wenzelm
parents: 74616
diff changeset
   270
  @{lemma "P (min x y) \<equiv> (x \<le> y \<and> P x \<or> x > y \<and> P y)" for x y :: real
wenzelm
parents: 74616
diff changeset
   271
    by (atomize (full)) (auto simp add: min_def)}
31120
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   272
74617
wenzelm
parents: 74616
diff changeset
   273
(* Miscellaneous *)
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   274
fun literals_conv bops uops cv =
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   275
  let
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   276
    fun h t =
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59058
diff changeset
   277
      (case Thm.term_of t of
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   278
        b$_$_ => if member (op aconv) bops b then binop_conv h t else cv t
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   279
      | u$_ => if member (op aconv) uops u then arg_conv h t else cv t
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59058
diff changeset
   280
      | _ => cv t)
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   281
  in h end;
31120
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   282
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   283
fun cterm_of_rat x =
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   284
  let
63201
f151704c08e4 tuned signature;
wenzelm
parents: 63198
diff changeset
   285
    val (a, b) = Rat.dest x
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   286
  in
67267
c5994f1fa0fa more symbols;
wenzelm
parents: 67232
diff changeset
   287
    if b = 1 then Numeral.mk_cnumber \<^ctyp>\<open>real\<close> a
74583
0eb2f18b1806 more antiquotations;
wenzelm
parents: 74525
diff changeset
   288
    else
0eb2f18b1806 more antiquotations;
wenzelm
parents: 74525
diff changeset
   289
      \<^instantiate>\<open>
0eb2f18b1806 more antiquotations;
wenzelm
parents: 74525
diff changeset
   290
          a = \<open>Numeral.mk_cnumber \<^ctyp>\<open>real\<close> a\<close> and
0eb2f18b1806 more antiquotations;
wenzelm
parents: 74525
diff changeset
   291
          b = \<open>Numeral.mk_cnumber \<^ctyp>\<open>real\<close> b\<close>
0eb2f18b1806 more antiquotations;
wenzelm
parents: 74525
diff changeset
   292
        in cterm \<open>a / b\<close> for a b :: real\<close>
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   293
  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
   294
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   295
fun dest_ratconst t =
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59058
diff changeset
   296
  case Thm.term_of t of
74583
0eb2f18b1806 more antiquotations;
wenzelm
parents: 74525
diff changeset
   297
    \<^Const_>\<open>divide _ for a b\<close> => Rat.make(HOLogic.dest_number a |> snd, HOLogic.dest_number b |> snd)
63201
f151704c08e4 tuned signature;
wenzelm
parents: 63198
diff changeset
   298
  | _ => Rat.of_int (HOLogic.dest_number (Thm.term_of t) |> snd)
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   299
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
   300
44454
6f28f96a09bf avoid warnings
huffman
parents: 44058
diff changeset
   301
(*
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   302
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
   303
 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
   304
  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
   305
 | 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
   306
 | _ => raise TERM ("find_term",[t]);
44454
6f28f96a09bf avoid warnings
huffman
parents: 44058
diff changeset
   307
*)
31120
fc654c95c29e 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
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   309
fun find_cterm p t =
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   310
  if p t then t else
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59058
diff changeset
   311
  case Thm.term_of t of
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   312
    _$_ => (find_cterm p (Thm.dest_fun t) handle CTERM _ => find_cterm p (Thm.dest_arg t))
74525
c960bfcb91db discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
wenzelm
parents: 74518
diff changeset
   313
  | Abs (_,_,_) => find_cterm p (Thm.dest_abs_global t |> snd)
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   314
  | _ => 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
   315
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59058
diff changeset
   316
fun is_comb t = (case Thm.term_of t of _ $ _ => true | _ => false);
31120
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   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
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
   319
  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
   320
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   321
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   322
(* Map back polynomials to HOL.                         *)
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   323
74583
0eb2f18b1806 more antiquotations;
wenzelm
parents: 74525
diff changeset
   324
fun cterm_of_varpow x k =
0eb2f18b1806 more antiquotations;
wenzelm
parents: 74525
diff changeset
   325
  if k = 1 then x
74584
wenzelm
parents: 74583
diff changeset
   326
  else \<^instantiate>\<open>x and k = \<open>Numeral.mk_cnumber \<^ctyp>\<open>nat\<close> k\<close> in cterm \<open>x ^ k\<close> for x :: real\<close>
32645
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_monomial m =
67267
c5994f1fa0fa more symbols;
wenzelm
parents: 67232
diff changeset
   329
  if FuncUtil.Ctermfunc.is_empty m then \<^cterm>\<open>1::real\<close>
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   330
  else
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   331
    let
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   332
      val m' = FuncUtil.dest_monomial m
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   333
      val vps = fold_rev (fn (x,k) => cons (cterm_of_varpow x k)) m' []
74584
wenzelm
parents: 74583
diff changeset
   334
    in foldr1 (fn (s, t) => \<^instantiate>\<open>s and t in cterm \<open>s * t\<close> for s t :: real\<close>) vps
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   335
    end
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   336
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   337
fun cterm_of_cmonomial (m,c) =
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   338
  if FuncUtil.Ctermfunc.is_empty m then cterm_of_rat c
63205
97b721666890 prefer rat numberals;
wenzelm
parents: 63201
diff changeset
   339
  else if c = @1 then cterm_of_monomial m
74584
wenzelm
parents: 74583
diff changeset
   340
  else \<^instantiate>\<open>x = \<open>cterm_of_rat c\<close> and y = \<open>cterm_of_monomial m\<close> in cterm \<open>x * y\<close> for x y :: real\<close>;
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   341
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   342
fun cterm_of_poly p =
67267
c5994f1fa0fa more symbols;
wenzelm
parents: 67232
diff changeset
   343
  if FuncUtil.Monomialfunc.is_empty p then \<^cterm>\<open>0::real\<close>
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   344
  else
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   345
    let
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   346
      val cms = map cterm_of_cmonomial
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   347
        (sort (prod_ord FuncUtil.monomial_order (K EQUAL)) (FuncUtil.Monomialfunc.dest p))
74584
wenzelm
parents: 74583
diff changeset
   348
    in foldr1 (fn (t1, t2) => \<^instantiate>\<open>t1 and t2 in cterm \<open>t1 + t2\<close> for t1 t2 :: real\<close>) cms
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   349
    end;
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents: 32402
diff changeset
   350
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   351
(* 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
   352
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   353
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
   354
       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
   355
       poly_conv,poly_neg_conv,poly_add_conv,poly_mul_conv,
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   356
       absconv1,absconv2,prover) =
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   357
  let
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 46594
diff changeset
   358
    val pre_ss = put_simpset HOL_basic_ss ctxt addsimps
58628
fd3c96a8ca60 tuned spelling;
wenzelm
parents: 52049
diff changeset
   359
      @{thms simp_thms ex_simps all_simps not_all not_ex ex_disj_distrib
fd3c96a8ca60 tuned spelling;
wenzelm
parents: 52049
diff changeset
   360
          all_conj_distrib if_bool_eq_disj}
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 46594
diff changeset
   361
    val prenex_ss = put_simpset HOL_basic_ss ctxt addsimps prenex_simps
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 46594
diff changeset
   362
    val skolemize_ss = put_simpset HOL_basic_ss ctxt addsimps [choice_iff]
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 46594
diff changeset
   363
    val presimp_conv = Simplifier.rewrite pre_ss
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 46594
diff changeset
   364
    val prenex_conv = Simplifier.rewrite prenex_ss
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 46594
diff changeset
   365
    val skolemize_conv = Simplifier.rewrite skolemize_ss
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 46594
diff changeset
   366
    val weak_dnf_ss = put_simpset HOL_basic_ss ctxt addsimps weak_dnf_simps
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 46594
diff changeset
   367
    val weak_dnf_conv = Simplifier.rewrite weak_dnf_ss
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   368
    fun eqT_elim th = Thm.equal_elim (Thm.symmetric th) @{thm TrueI}
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   369
    fun oprconv cv ct =
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   370
      let val g = Thm.dest_fun2 ct
67564
wenzelm
parents: 67562
diff changeset
   371
      in if g aconvc \<^cterm>\<open>(\<le>) :: real \<Rightarrow> _\<close>
wenzelm
parents: 67562
diff changeset
   372
            orelse g aconvc \<^cterm>\<open>(<) :: real \<Rightarrow> _\<close>
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   373
         then arg_conv cv ct else arg1_conv cv ct
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   374
      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
   375
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   376
    fun real_ineq_conv th ct =
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   377
      let
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   378
        val th' = (Thm.instantiate (Thm.match (Thm.lhs_of th, ct)) th
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   379
          handle Pattern.MATCH => raise CTERM ("real_ineq_conv", [ct]))
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   380
      in Thm.transitive th' (oprconv poly_conv (Thm.rhs_of th'))
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   381
      end
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   382
    val [real_lt_conv, real_le_conv, real_eq_conv,
74615
9af55a51e185 clarified antiquotations;
wenzelm
parents: 74608
diff changeset
   383
         real_not_lt_conv, real_not_le_conv] =
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   384
         map real_ineq_conv pth
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   385
    fun match_mp_rule ths ths' =
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   386
      let
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   387
        fun f ths ths' = case ths of [] => raise THM("match_mp_rule",0,ths)
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   388
          | th::ths => (ths' MRS th handle THM _ => f ths ths')
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   389
      in f ths ths' end
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   390
    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
   391
         (match_mp_rule pth_mul [th, th'])
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   392
    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
   393
         (match_mp_rule pth_add [th, th'])
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   394
    fun emul_rule ct th = fconv_rule (arg_conv (oprconv poly_mul_conv))
60801
7664e0916eec tuned signature;
wenzelm
parents: 60642
diff changeset
   395
       (Thm.instantiate' [] [SOME ct] (th RS pth_emul))
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   396
    fun square_rule t = fconv_rule (arg_conv (oprconv poly_conv))
74615
9af55a51e185 clarified antiquotations;
wenzelm
parents: 74608
diff changeset
   397
       \<^instantiate>\<open>x = t in lemma \<open>x * x \<ge> 0\<close> for x :: real by simp\<close>
31120
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   398
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   399
    fun hol_of_positivstellensatz(eqs,les,lts) proof =
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   400
      let
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   401
        fun translate prf =
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   402
          case prf of
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   403
            Axiom_eq n => nth eqs n
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   404
          | Axiom_le n => nth les n
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   405
          | Axiom_lt n => nth lts n
74583
0eb2f18b1806 more antiquotations;
wenzelm
parents: 74525
diff changeset
   406
          | Rational_eq x =>
0eb2f18b1806 more antiquotations;
wenzelm
parents: 74525
diff changeset
   407
              eqT_elim (numeric_eq_conv
74584
wenzelm
parents: 74583
diff changeset
   408
                \<^instantiate>\<open>x = \<open>mk_numeric x\<close> in cprop \<open>x = 0\<close> for x :: real\<close>)
74583
0eb2f18b1806 more antiquotations;
wenzelm
parents: 74525
diff changeset
   409
          | Rational_le x =>
0eb2f18b1806 more antiquotations;
wenzelm
parents: 74525
diff changeset
   410
              eqT_elim (numeric_ge_conv
74584
wenzelm
parents: 74583
diff changeset
   411
                \<^instantiate>\<open>x = \<open>mk_numeric x\<close> in cprop \<open>x \<ge> 0\<close> for x :: real\<close>)
74583
0eb2f18b1806 more antiquotations;
wenzelm
parents: 74525
diff changeset
   412
          | Rational_lt x =>
0eb2f18b1806 more antiquotations;
wenzelm
parents: 74525
diff changeset
   413
              eqT_elim (numeric_gt_conv
74584
wenzelm
parents: 74583
diff changeset
   414
                \<^instantiate>\<open>x = \<open>mk_numeric x\<close> in cprop \<open>x > 0\<close> for x :: real\<close>)
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   415
          | Square pt => square_rule (cterm_of_poly pt)
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   416
          | Eqmul(pt,p) => emul_rule (cterm_of_poly pt) (translate p)
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   417
          | Sum(p1,p2) => add_rule (translate p1) (translate p2)
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   418
          | Product(p1,p2) => mul_rule (translate p1) (translate p2)
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   419
      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
   420
          (translate proof)
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   421
      end
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   422
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   423
    val init_conv = presimp_conv then_conv
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 46594
diff changeset
   424
        nnf_conv ctxt then_conv skolemize_conv then_conv prenex_conv then_conv
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   425
        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
   426
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59058
diff changeset
   427
    val concl = Thm.dest_arg o Thm.cprop_of
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   428
    fun is_binop opr ct = (Thm.dest_fun2 ct aconvc opr handle CTERM _ => false)
67564
wenzelm
parents: 67562
diff changeset
   429
    val is_req = is_binop \<^cterm>\<open>(=):: real \<Rightarrow> _\<close>
wenzelm
parents: 67562
diff changeset
   430
    val is_ge = is_binop \<^cterm>\<open>(\<le>):: real \<Rightarrow> _\<close>
wenzelm
parents: 67562
diff changeset
   431
    val is_gt = is_binop \<^cterm>\<open>(<):: real \<Rightarrow> _\<close>
67267
c5994f1fa0fa more symbols;
wenzelm
parents: 67232
diff changeset
   432
    val is_conj = is_binop \<^cterm>\<open>HOL.conj\<close>
c5994f1fa0fa more symbols;
wenzelm
parents: 67232
diff changeset
   433
    val is_disj = is_binop \<^cterm>\<open>HOL.disj\<close>
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   434
    fun conj_pair th = (th RS @{thm conjunct1}, th RS @{thm conjunct2})
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   435
    fun disj_cases th th1 th2 =
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   436
      let
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   437
        val (p,q) = Thm.dest_binop (concl th)
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   438
        val c = concl th1
58628
fd3c96a8ca60 tuned spelling;
wenzelm
parents: 52049
diff changeset
   439
        val _ =
fd3c96a8ca60 tuned spelling;
wenzelm
parents: 52049
diff changeset
   440
          if c aconvc (concl th2) then ()
fd3c96a8ca60 tuned spelling;
wenzelm
parents: 52049
diff changeset
   441
          else error "disj_cases : conclusions not alpha convertible"
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   442
      in Thm.implies_elim (Thm.implies_elim
60801
7664e0916eec tuned signature;
wenzelm
parents: 60642
diff changeset
   443
          (Thm.implies_elim (Thm.instantiate' [] (map SOME [p,q,c]) @{thm disjE}) th)
74583
0eb2f18b1806 more antiquotations;
wenzelm
parents: 74525
diff changeset
   444
          (Thm.implies_intr \<^instantiate>\<open>p in cprop p\<close> th1))
0eb2f18b1806 more antiquotations;
wenzelm
parents: 74525
diff changeset
   445
        (Thm.implies_intr \<^instantiate>\<open>q in cprop q\<close> 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
58628
fd3c96a8ca60 tuned spelling;
wenzelm
parents: 52049
diff changeset
   465
              val (th1, cert1) =
fd3c96a8ca60 tuned spelling;
wenzelm
parents: 52049
diff changeset
   466
                overall (Left::cert_choice) dun
67267
c5994f1fa0fa more symbols;
wenzelm
parents: 67232
diff changeset
   467
                  (Thm.assume (Thm.apply \<^cterm>\<open>Trueprop\<close> (Thm.dest_arg1 ct))::oths)
58628
fd3c96a8ca60 tuned spelling;
wenzelm
parents: 52049
diff changeset
   468
              val (th2, cert2) =
fd3c96a8ca60 tuned spelling;
wenzelm
parents: 52049
diff changeset
   469
                overall (Right::cert_choice) dun
67267
c5994f1fa0fa more symbols;
wenzelm
parents: 67232
diff changeset
   470
                  (Thm.assume (Thm.apply \<^cterm>\<open>Trueprop\<close> (Thm.dest_arg ct))::oths)
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   471
            in (disj_cases th th1 th2, Branch (cert1, cert2)) end
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   472
          else overall cert_choice (th::dun) oths
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   473
        end
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   474
    fun dest_binary b ct =
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   475
        if is_binop b ct then Thm.dest_binop ct
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   476
        else raise CTERM ("dest_binary",[b,ct])
67564
wenzelm
parents: 67562
diff changeset
   477
    val dest_eq = dest_binary \<^cterm>\<open>(=) :: real \<Rightarrow> _\<close>
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   478
    fun real_not_eq_conv ct =
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   479
      let
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   480
        val (l,r) = dest_eq (Thm.dest_arg ct)
74615
9af55a51e185 clarified antiquotations;
wenzelm
parents: 74608
diff changeset
   481
        val th =
9af55a51e185 clarified antiquotations;
wenzelm
parents: 74608
diff changeset
   482
          \<^instantiate>\<open>x = l and y = r in lemma \<open>x \<noteq> y \<equiv> x - y > 0 \<or> - (x - y) > 0\<close> for x y :: real
9af55a51e185 clarified antiquotations;
wenzelm
parents: 74608
diff changeset
   483
            by (atomize (full), auto simp add: less_diff_eq le_diff_eq not_less)\<close>;
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   484
        val th_p = poly_conv(Thm.dest_arg(Thm.dest_arg1(Thm.rhs_of th)))
67267
c5994f1fa0fa more symbols;
wenzelm
parents: 67232
diff changeset
   485
        val th_x = Drule.arg_cong_rule \<^cterm>\<open>uminus :: real \<Rightarrow> _\<close> th_p
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   486
        val th_n = fconv_rule (arg_conv poly_neg_conv) th_x
67267
c5994f1fa0fa more symbols;
wenzelm
parents: 67232
diff changeset
   487
        val th' = Drule.binop_cong_rule \<^cterm>\<open>HOL.disj\<close>
74583
0eb2f18b1806 more antiquotations;
wenzelm
parents: 74525
diff changeset
   488
          (Drule.arg_cong_rule \<^cterm>\<open>(<) (0::real)\<close> th_p)
0eb2f18b1806 more antiquotations;
wenzelm
parents: 74525
diff changeset
   489
          (Drule.arg_cong_rule \<^cterm>\<open>(<) (0::real)\<close> th_n)
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   490
      in Thm.transitive th th'
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   491
      end
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   492
    fun equal_implies_1_rule PQ =
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   493
      let
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   494
        val P = Thm.lhs_of PQ
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   495
      in Thm.implies_intr P (Thm.equal_elim PQ (Thm.assume P))
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   496
      end
63667
24126c564d8a tuned whitespace;
wenzelm
parents: 63648
diff changeset
   497
    (*FIXME!!! Copied from groebner.ml*)
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   498
    val strip_exists =
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   499
      let
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   500
        fun h (acc, t) =
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59058
diff changeset
   501
          case Thm.term_of t of
74583
0eb2f18b1806 more antiquotations;
wenzelm
parents: 74525
diff changeset
   502
            \<^Const_>\<open>Ex _ for \<open>Abs _\<close>\<close> =>
74525
c960bfcb91db discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
wenzelm
parents: 74518
diff changeset
   503
              h (Thm.dest_abs_global (Thm.dest_arg t) |>> (fn v => v::acc))
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   504
          | _ => (acc,t)
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   505
      in fn t => h ([],t)
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   506
      end
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   507
    fun name_of x =
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59058
diff changeset
   508
      case Thm.term_of x of
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   509
        Free(s,_) => s
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   510
      | Var ((s,_),_) => s
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   511
      | _ => "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
   512
58628
fd3c96a8ca60 tuned spelling;
wenzelm
parents: 52049
diff changeset
   513
    fun mk_forall x th =
61075
f6b0d827240e tuned -- avoid slightly odd @{cpat};
wenzelm
parents: 60949
diff changeset
   514
      let
74583
0eb2f18b1806 more antiquotations;
wenzelm
parents: 74525
diff changeset
   515
        val T = Thm.ctyp_of_cterm x
0eb2f18b1806 more antiquotations;
wenzelm
parents: 74525
diff changeset
   516
        val all = \<^instantiate>\<open>'a = T in cterm All\<close>
61075
f6b0d827240e tuned -- avoid slightly odd @{cpat};
wenzelm
parents: 60949
diff changeset
   517
      in Drule.arg_cong_rule all (Thm.abstract_rule (name_of x) x th) end
31120
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   518
60801
7664e0916eec tuned signature;
wenzelm
parents: 60642
diff changeset
   519
    val specl = fold_rev (fn x => fn th => Thm.instantiate' [] [SOME x] (th RS spec));
31120
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   520
74622
9568db8f4882 clarified antiquotations: no need for Thm.beta_conversion, due to Thm.instantiate_beta;
wenzelm
parents: 74621
diff changeset
   521
    fun mk_ex x t =
9568db8f4882 clarified antiquotations: no need for Thm.beta_conversion, due to Thm.instantiate_beta;
wenzelm
parents: 74621
diff changeset
   522
      \<^instantiate>\<open>'a = \<open>Thm.ctyp_of_cterm x\<close> and P = \<open>Thm.lambda x t\<close>
74584
wenzelm
parents: 74583
diff changeset
   523
        in cprop \<open>Ex P\<close> for P :: \<open>'a \<Rightarrow> bool\<close>\<close>
31120
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   524
74622
9568db8f4882 clarified antiquotations: no need for Thm.beta_conversion, due to Thm.instantiate_beta;
wenzelm
parents: 74621
diff changeset
   525
    fun choose x th th' =
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59058
diff changeset
   526
      case Thm.concl_of th of
74583
0eb2f18b1806 more antiquotations;
wenzelm
parents: 74525
diff changeset
   527
        \<^Const_>\<open>Trueprop for \<^Const_>\<open>Ex _ for _\<close>\<close> =>
0eb2f18b1806 more antiquotations;
wenzelm
parents: 74525
diff changeset
   528
          let
74622
9568db8f4882 clarified antiquotations: no need for Thm.beta_conversion, due to Thm.instantiate_beta;
wenzelm
parents: 74621
diff changeset
   529
            val P = Thm.dest_arg (Thm.dest_arg (Thm.cprop_of th))
9568db8f4882 clarified antiquotations: no need for Thm.beta_conversion, due to Thm.instantiate_beta;
wenzelm
parents: 74621
diff changeset
   530
            val T = Thm.dest_ctyp0 (Thm.ctyp_of_cterm P)
9568db8f4882 clarified antiquotations: no need for Thm.beta_conversion, due to Thm.instantiate_beta;
wenzelm
parents: 74621
diff changeset
   531
            val Q = Thm.dest_arg (Thm.cprop_of th')
9568db8f4882 clarified antiquotations: no need for Thm.beta_conversion, due to Thm.instantiate_beta;
wenzelm
parents: 74621
diff changeset
   532
            val th0 =
9568db8f4882 clarified antiquotations: no need for Thm.beta_conversion, due to Thm.instantiate_beta;
wenzelm
parents: 74621
diff changeset
   533
              \<^instantiate>\<open>'a = T and P and Q in
9568db8f4882 clarified antiquotations: no need for Thm.beta_conversion, due to Thm.instantiate_beta;
wenzelm
parents: 74621
diff changeset
   534
                lemma "\<exists>x::'a. P x \<Longrightarrow> (\<And>x. P x \<Longrightarrow> Q) \<Longrightarrow> Q" by (fact exE)\<close>
9568db8f4882 clarified antiquotations: no need for Thm.beta_conversion, due to Thm.instantiate_beta;
wenzelm
parents: 74621
diff changeset
   535
            val Px =
9568db8f4882 clarified antiquotations: no need for Thm.beta_conversion, due to Thm.instantiate_beta;
wenzelm
parents: 74621
diff changeset
   536
              \<^instantiate>\<open>'a = T and P and x in cprop \<open>P x\<close> for x :: 'a\<close>
9568db8f4882 clarified antiquotations: no need for Thm.beta_conversion, due to Thm.instantiate_beta;
wenzelm
parents: 74621
diff changeset
   537
            val th1 = Thm.forall_intr x (Thm.implies_intr Px th')
74583
0eb2f18b1806 more antiquotations;
wenzelm
parents: 74525
diff changeset
   538
          in Thm.implies_elim (Thm.implies_elim th0 th) th1  end
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   539
      | _ => 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
   540
74622
9568db8f4882 clarified antiquotations: no need for Thm.beta_conversion, due to Thm.instantiate_beta;
wenzelm
parents: 74621
diff changeset
   541
    fun simple_choose x th =
9568db8f4882 clarified antiquotations: no need for Thm.beta_conversion, due to Thm.instantiate_beta;
wenzelm
parents: 74621
diff changeset
   542
      choose x (Thm.assume (mk_ex x (Thm.dest_arg (hd (Thm.chyps_of th))))) th
31120
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   543
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   544
    val strip_forall =
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   545
      let
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   546
        fun h (acc, t) =
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59058
diff changeset
   547
          case Thm.term_of t of
74583
0eb2f18b1806 more antiquotations;
wenzelm
parents: 74525
diff changeset
   548
            \<^Const_>\<open>All _ for \<open>Abs _\<close>\<close> =>
74525
c960bfcb91db discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
wenzelm
parents: 74518
diff changeset
   549
              h (Thm.dest_abs_global (Thm.dest_arg t) |>> (fn v => v::acc))
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   550
          | _ => (acc,t)
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   551
      in fn t => h ([],t)
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   552
      end
31120
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   553
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   554
    fun f ct =
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   555
      let
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   556
        val nnf_norm_conv' =
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 46594
diff changeset
   557
          nnf_conv ctxt then_conv
67267
c5994f1fa0fa more symbols;
wenzelm
parents: 67232
diff changeset
   558
          literals_conv [\<^term>\<open>HOL.conj\<close>, \<^term>\<open>HOL.disj\<close>] []
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   559
          (Conv.cache_conv
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   560
            (first_conv [real_lt_conv, real_le_conv,
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   561
                         real_eq_conv, real_not_lt_conv,
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   562
                         real_not_le_conv, real_not_eq_conv, all_conv]))
67267
c5994f1fa0fa more symbols;
wenzelm
parents: 67232
diff changeset
   563
        fun absremover ct = (literals_conv [\<^term>\<open>HOL.conj\<close>, \<^term>\<open>HOL.disj\<close>] []
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   564
                  (try_conv (absconv1 then_conv binop_conv (arg_conv poly_conv))) then_conv
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   565
                  try_conv (absconv2 then_conv nnf_norm_conv' then_conv binop_conv absremover)) ct
74584
wenzelm
parents: 74583
diff changeset
   566
        val nct = \<^instantiate>\<open>A = ct in cprop \<open>\<not> A\<close>\<close>
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   567
        val th0 = (init_conv then_conv arg_conv nnf_norm_conv') nct
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   568
        val tm0 = Thm.dest_arg (Thm.rhs_of th0)
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   569
        val (th, certificates) =
67267
c5994f1fa0fa more symbols;
wenzelm
parents: 67232
diff changeset
   570
          if tm0 aconvc \<^cterm>\<open>False\<close> then (equal_implies_1_rule th0, Trivial) else
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   571
          let
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   572
            val (evs,bod) = strip_exists tm0
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   573
            val (avs,ibod) = strip_forall bod
67267
c5994f1fa0fa more symbols;
wenzelm
parents: 67232
diff changeset
   574
            val th1 = Drule.arg_cong_rule \<^cterm>\<open>Trueprop\<close> (fold mk_forall avs (absremover ibod))
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   575
            val (th2, certs) = overall [] [] [specl avs (Thm.assume (Thm.rhs_of th1))]
58628
fd3c96a8ca60 tuned spelling;
wenzelm
parents: 52049
diff changeset
   576
            val th3 =
fd3c96a8ca60 tuned spelling;
wenzelm
parents: 52049
diff changeset
   577
              fold simple_choose evs
67267
c5994f1fa0fa more symbols;
wenzelm
parents: 67232
diff changeset
   578
                (prove_hyp (Thm.equal_elim th1 (Thm.assume (Thm.apply \<^cterm>\<open>Trueprop\<close> bod))) th2)
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   579
          in (Drule.implies_intr_hyps (prove_hyp (Thm.equal_elim th0 (Thm.assume nct)) th3), certs)
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   580
          end
74615
9af55a51e185 clarified antiquotations;
wenzelm
parents: 74608
diff changeset
   581
      in
9af55a51e185 clarified antiquotations;
wenzelm
parents: 74608
diff changeset
   582
        (Thm.implies_elim \<^instantiate>\<open>A = ct in lemma \<open>(\<not> A \<Longrightarrow> False) \<Longrightarrow> A\<close> by blast\<close> th,
9af55a51e185 clarified antiquotations;
wenzelm
parents: 74608
diff changeset
   583
          certificates)
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   584
      end
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   585
  in f
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   586
  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
   587
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   588
(* 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
   589
local
63205
97b721666890 prefer rat numberals;
wenzelm
parents: 63201
diff changeset
   590
  val linear_add = FuncUtil.Ctermfunc.combine (curry op +) (fn z => z = @0)
63198
c583ca33076a ad-hoc overloading for standard operations on type Rat.rat;
wenzelm
parents: 62177
diff changeset
   591
  fun linear_cmul c = FuncUtil.Ctermfunc.map (fn _ => fn x => c * x)
67267
c5994f1fa0fa more symbols;
wenzelm
parents: 67232
diff changeset
   592
  val one_tm = \<^cterm>\<open>1::real\<close>
63205
97b721666890 prefer rat numberals;
wenzelm
parents: 63201
diff changeset
   593
  fun contradictory p (e,_) = ((FuncUtil.Ctermfunc.is_empty e) andalso not(p @0)) orelse
33038
8f9594c31de4 dropped redundant gen_ prefix
haftmann
parents: 33002
diff changeset
   594
     ((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
   595
       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
   596
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   597
  fun linear_ineqs vars (les,lts) =
63205
97b721666890 prefer rat numberals;
wenzelm
parents: 63201
diff changeset
   598
    case find_first (contradictory (fn x => x > @0)) lts of
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   599
      SOME r => r
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   600
    | NONE =>
63205
97b721666890 prefer rat numberals;
wenzelm
parents: 63201
diff changeset
   601
      (case find_first (contradictory (fn x => x > @0)) les of
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   602
         SOME r => r
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   603
       | NONE =>
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   604
         if null vars then error "linear_ineqs: no contradiction" else
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   605
         let
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   606
           val ineqs = les @ lts
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   607
           fun blowup v =
63205
97b721666890 prefer rat numberals;
wenzelm
parents: 63201
diff changeset
   608
             length(filter (fn (e,_) => FuncUtil.Ctermfunc.tryapplyd e v @0 = @0) ineqs) +
97b721666890 prefer rat numberals;
wenzelm
parents: 63201
diff changeset
   609
             length(filter (fn (e,_) => FuncUtil.Ctermfunc.tryapplyd e v @0 > @0) ineqs) *
97b721666890 prefer rat numberals;
wenzelm
parents: 63201
diff changeset
   610
             length(filter (fn (e,_) => FuncUtil.Ctermfunc.tryapplyd e v @0 < @0) ineqs)
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   611
           val v = fst(hd(sort (fn ((_,i),(_,j)) => int_ord (i,j))
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   612
             (map (fn v => (v,blowup v)) vars)))
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   613
           fun addup (e1,p1) (e2,p2) acc =
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   614
             let
63205
97b721666890 prefer rat numberals;
wenzelm
parents: 63201
diff changeset
   615
               val c1 = FuncUtil.Ctermfunc.tryapplyd e1 v @0
97b721666890 prefer rat numberals;
wenzelm
parents: 63201
diff changeset
   616
               val c2 = FuncUtil.Ctermfunc.tryapplyd e2 v @0
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   617
             in
63205
97b721666890 prefer rat numberals;
wenzelm
parents: 63201
diff changeset
   618
               if c1 * c2 >= @0 then acc else
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   619
               let
63211
0bec0d1d9998 more adhoc overloading;
wenzelm
parents: 63205
diff changeset
   620
                 val e1' = linear_cmul (abs c2) e1
0bec0d1d9998 more adhoc overloading;
wenzelm
parents: 63205
diff changeset
   621
                 val e2' = linear_cmul (abs c1) e2
0bec0d1d9998 more adhoc overloading;
wenzelm
parents: 63205
diff changeset
   622
                 val p1' = Product(Rational_lt (abs c2), p1)
0bec0d1d9998 more adhoc overloading;
wenzelm
parents: 63205
diff changeset
   623
                 val p2' = Product(Rational_lt (abs c1), p2)
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   624
               in (linear_add e1' e2',Sum(p1',p2'))::acc
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   625
               end
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   626
             end
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   627
           val (les0,les1) =
63205
97b721666890 prefer rat numberals;
wenzelm
parents: 63201
diff changeset
   628
             List.partition (fn (e,_) => FuncUtil.Ctermfunc.tryapplyd e v @0 = @0) les
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   629
           val (lts0,lts1) =
63205
97b721666890 prefer rat numberals;
wenzelm
parents: 63201
diff changeset
   630
             List.partition (fn (e,_) => FuncUtil.Ctermfunc.tryapplyd e v @0 = @0) lts
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   631
           val (lesp,lesn) =
63205
97b721666890 prefer rat numberals;
wenzelm
parents: 63201
diff changeset
   632
             List.partition (fn (e,_) => FuncUtil.Ctermfunc.tryapplyd e v @0 > @0) les1
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   633
           val (ltsp,ltsn) =
63205
97b721666890 prefer rat numberals;
wenzelm
parents: 63201
diff changeset
   634
             List.partition (fn (e,_) => FuncUtil.Ctermfunc.tryapplyd e v @0 > @0) lts1
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   635
           val les' = fold_rev (fn ep1 => fold_rev (addup ep1) lesp) lesn les0
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   636
           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
   637
                      (fold_rev (fn ep1 => fold_rev (addup ep1) (lesn@ltsn)) ltsp lts0)
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   638
         in linear_ineqs (remove (op aconvc) v vars) (les',lts')
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   639
         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
   640
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   641
  fun linear_eqs(eqs,les,lts) =
63205
97b721666890 prefer rat numberals;
wenzelm
parents: 63201
diff changeset
   642
    case find_first (contradictory (fn x => x = @0)) eqs of
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   643
      SOME r => r
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   644
    | NONE =>
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   645
      (case eqs of
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   646
         [] =>
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   647
         let val vars = remove (op aconvc) one_tm
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   648
             (fold_rev (union (op aconvc) o FuncUtil.Ctermfunc.dom o fst) (les@lts) [])
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   649
         in linear_ineqs vars (les,lts) end
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   650
       | (e,p)::es =>
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   651
         if FuncUtil.Ctermfunc.is_empty e then linear_eqs (es,les,lts) else
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   652
         let
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   653
           val (x,c) = FuncUtil.Ctermfunc.choose (FuncUtil.Ctermfunc.delete_safe one_tm e)
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   654
           fun xform (inp as (t,q)) =
63205
97b721666890 prefer rat numberals;
wenzelm
parents: 63201
diff changeset
   655
             let val d = FuncUtil.Ctermfunc.tryapplyd t x @0 in
97b721666890 prefer rat numberals;
wenzelm
parents: 63201
diff changeset
   656
               if d = @0 then inp else
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   657
               let
63211
0bec0d1d9998 more adhoc overloading;
wenzelm
parents: 63205
diff changeset
   658
                 val k = ~ d * abs c / c
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   659
                 val e' = linear_cmul k e
63211
0bec0d1d9998 more adhoc overloading;
wenzelm
parents: 63205
diff changeset
   660
                 val t' = linear_cmul (abs c) t
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   661
                 val p' = Eqmul(FuncUtil.Monomialfunc.onefunc (FuncUtil.Ctermfunc.empty, k),p)
63211
0bec0d1d9998 more adhoc overloading;
wenzelm
parents: 63205
diff changeset
   662
                 val q' = Product(Rational_lt (abs c), q)
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   663
               in (linear_add e' t',Sum(p',q'))
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   664
               end
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   665
             end
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   666
         in linear_eqs(map xform es,map xform les,map xform lts)
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   667
         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
   668
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   669
  fun linear_prover (eq,le,lt) =
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   670
    let
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   671
      val eqs = map_index (fn (n, p) => (p,Axiom_eq n)) eq
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   672
      val les = map_index (fn (n, p) => (p,Axiom_le n)) le
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   673
      val lts = map_index (fn (n, p) => (p,Axiom_lt n)) lt
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   674
    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
   675
    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
   676
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   677
  fun lin_of_hol ct =
67267
c5994f1fa0fa more symbols;
wenzelm
parents: 67232
diff changeset
   678
    if ct aconvc \<^cterm>\<open>0::real\<close> then FuncUtil.Ctermfunc.empty
63205
97b721666890 prefer rat numberals;
wenzelm
parents: 63201
diff changeset
   679
    else if not (is_comb ct) then FuncUtil.Ctermfunc.onefunc (ct, @1)
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   680
    else if is_ratconst ct then FuncUtil.Ctermfunc.onefunc (one_tm, dest_ratconst ct)
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   681
    else
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   682
      let val (lop,r) = Thm.dest_comb ct
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   683
      in
63205
97b721666890 prefer rat numberals;
wenzelm
parents: 63201
diff changeset
   684
        if not (is_comb lop) then FuncUtil.Ctermfunc.onefunc (ct, @1)
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   685
        else
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   686
          let val (opr,l) = Thm.dest_comb lop
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   687
          in
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 67271
diff changeset
   688
            if opr aconvc \<^cterm>\<open>(+) :: real \<Rightarrow> _\<close>
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   689
            then linear_add (lin_of_hol l) (lin_of_hol r)
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 67564
diff changeset
   690
            else if opr aconvc \<^cterm>\<open>(*) :: real \<Rightarrow> _\<close>
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   691
                    andalso is_ratconst l then FuncUtil.Ctermfunc.onefunc (r, dest_ratconst l)
63205
97b721666890 prefer rat numberals;
wenzelm
parents: 63201
diff changeset
   692
            else FuncUtil.Ctermfunc.onefunc (ct, @1)
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   693
          end
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   694
      end
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   695
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   696
  fun is_alien ct =
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59058
diff changeset
   697
    case Thm.term_of ct of
74583
0eb2f18b1806 more antiquotations;
wenzelm
parents: 74525
diff changeset
   698
      \<^Const_>\<open>of_nat _ for n\<close> => not (can HOLogic.dest_number n)
0eb2f18b1806 more antiquotations;
wenzelm
parents: 74525
diff changeset
   699
    | \<^Const_>\<open>of_int _ for n\<close> => not (can HOLogic.dest_number n)
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59058
diff changeset
   700
    | _ => false
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   701
in
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   702
fun real_linear_prover translator (eq,le,lt) =
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   703
  let
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59058
diff changeset
   704
    val lhs = lin_of_hol o Thm.dest_arg1 o Thm.dest_arg o Thm.cprop_of
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59058
diff changeset
   705
    val rhs = lin_of_hol o Thm.dest_arg o Thm.dest_arg o Thm.cprop_of
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   706
    val eq_pols = map lhs eq
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   707
    val le_pols = map rhs le
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   708
    val lt_pols = map rhs lt
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   709
    val aliens = filter is_alien
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   710
      (fold_rev (union (op aconvc) o FuncUtil.Ctermfunc.dom)
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   711
                (eq_pols @ le_pols @ lt_pols) [])
63205
97b721666890 prefer rat numberals;
wenzelm
parents: 63201
diff changeset
   712
    val le_pols' = le_pols @ map (fn v => FuncUtil.Ctermfunc.onefunc (v,@1)) aliens
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   713
    val (_,proof) = linear_prover (eq_pols,le_pols',lt_pols)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61075
diff changeset
   714
    val le' = le @ map (fn a => Thm.instantiate' [] [SOME (Thm.dest_arg a)] @{thm of_nat_0_le_iff}) aliens
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   715
  in ((translator (eq,le',lt) proof), Trivial)
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   716
  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
   717
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
   718
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   719
(* 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
   720
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   721
local
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 46594
diff changeset
   722
  val absmaxmin_elim_ss1 =
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69064
diff changeset
   723
    simpset_of (put_simpset HOL_basic_ss \<^context> addsimps real_abs_thms1)
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   724
  fun absmaxmin_elim_conv1 ctxt =
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 46594
diff changeset
   725
    Simplifier.rewrite (put_simpset absmaxmin_elim_ss1 ctxt)
31120
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   726
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   727
  val absmaxmin_elim_conv2 =
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   728
    let
74621
82ff15b980e9 clarified antiquotations: no need for Thm.beta_conversion, due to Thm.instantiate_beta;
wenzelm
parents: 74617
diff changeset
   729
      fun elim_construct pred conv tm =
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   730
        let
74621
82ff15b980e9 clarified antiquotations: no need for Thm.beta_conversion, due to Thm.instantiate_beta;
wenzelm
parents: 74617
diff changeset
   731
          val a = find_cterm (pred o Thm.term_of) tm
82ff15b980e9 clarified antiquotations: no need for Thm.beta_conversion, due to Thm.instantiate_beta;
wenzelm
parents: 74617
diff changeset
   732
          val P = Thm.lambda a tm
82ff15b980e9 clarified antiquotations: no need for Thm.beta_conversion, due to Thm.instantiate_beta;
wenzelm
parents: 74617
diff changeset
   733
        in conv P a 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
   734
74621
82ff15b980e9 clarified antiquotations: no need for Thm.beta_conversion, due to Thm.instantiate_beta;
wenzelm
parents: 74617
diff changeset
   735
      val elim_abs = elim_construct (fn \<^Const_>\<open>abs \<^Type>\<open>real\<close> for _\<close> => true | _ => false)
74616
wenzelm
parents: 74615
diff changeset
   736
        (fn P => fn a =>
74621
82ff15b980e9 clarified antiquotations: no need for Thm.beta_conversion, due to Thm.instantiate_beta;
wenzelm
parents: 74617
diff changeset
   737
          let val x = Thm.dest_arg a in
82ff15b980e9 clarified antiquotations: no need for Thm.beta_conversion, due to Thm.instantiate_beta;
wenzelm
parents: 74617
diff changeset
   738
            \<^instantiate>\<open>P and x in
82ff15b980e9 clarified antiquotations: no need for Thm.beta_conversion, due to Thm.instantiate_beta;
wenzelm
parents: 74617
diff changeset
   739
              lemma \<open>P \<bar>x\<bar> \<equiv> (x \<ge> 0 \<and> P x \<or> x < 0 \<and> P (- x))\<close> for x :: real
82ff15b980e9 clarified antiquotations: no need for Thm.beta_conversion, due to Thm.instantiate_beta;
wenzelm
parents: 74617
diff changeset
   740
                by (atomize (full)) (auto split: abs_split)\<close>
82ff15b980e9 clarified antiquotations: no need for Thm.beta_conversion, due to Thm.instantiate_beta;
wenzelm
parents: 74617
diff changeset
   741
          end)
82ff15b980e9 clarified antiquotations: no need for Thm.beta_conversion, due to Thm.instantiate_beta;
wenzelm
parents: 74617
diff changeset
   742
      val elim_max = elim_construct (fn \<^Const_>\<open>max \<^Type>\<open>real\<close> for _ _\<close> => true | _ => false)
74616
wenzelm
parents: 74615
diff changeset
   743
        (fn P => fn a =>
74621
82ff15b980e9 clarified antiquotations: no need for Thm.beta_conversion, due to Thm.instantiate_beta;
wenzelm
parents: 74617
diff changeset
   744
          let val (x, y) = Thm.dest_binop a in
82ff15b980e9 clarified antiquotations: no need for Thm.beta_conversion, due to Thm.instantiate_beta;
wenzelm
parents: 74617
diff changeset
   745
            \<^instantiate>\<open>P and x and y in
82ff15b980e9 clarified antiquotations: no need for Thm.beta_conversion, due to Thm.instantiate_beta;
wenzelm
parents: 74617
diff changeset
   746
              lemma \<open>P (max x y) \<equiv> (x \<le> y \<and> P y \<or> x > y \<and> P x)\<close> for x y :: real
82ff15b980e9 clarified antiquotations: no need for Thm.beta_conversion, due to Thm.instantiate_beta;
wenzelm
parents: 74617
diff changeset
   747
                by (atomize (full)) (auto simp add: max_def)\<close>
82ff15b980e9 clarified antiquotations: no need for Thm.beta_conversion, due to Thm.instantiate_beta;
wenzelm
parents: 74617
diff changeset
   748
          end)
82ff15b980e9 clarified antiquotations: no need for Thm.beta_conversion, due to Thm.instantiate_beta;
wenzelm
parents: 74617
diff changeset
   749
      val elim_min = elim_construct (fn \<^Const_>\<open>min \<^Type>\<open>real\<close> for _ _\<close> => true | _ => false)
74616
wenzelm
parents: 74615
diff changeset
   750
        (fn P => fn a =>
74621
82ff15b980e9 clarified antiquotations: no need for Thm.beta_conversion, due to Thm.instantiate_beta;
wenzelm
parents: 74617
diff changeset
   751
          let val (x, y) = Thm.dest_binop a in
82ff15b980e9 clarified antiquotations: no need for Thm.beta_conversion, due to Thm.instantiate_beta;
wenzelm
parents: 74617
diff changeset
   752
            \<^instantiate>\<open>P and x and y in
82ff15b980e9 clarified antiquotations: no need for Thm.beta_conversion, due to Thm.instantiate_beta;
wenzelm
parents: 74617
diff changeset
   753
              lemma \<open>P (min x y) \<equiv> (x \<le> y \<and> P x \<or> x > y \<and> P y)\<close> for x y :: real
82ff15b980e9 clarified antiquotations: no need for Thm.beta_conversion, due to Thm.instantiate_beta;
wenzelm
parents: 74617
diff changeset
   754
                by (atomize (full)) (auto simp add: min_def)\<close>
82ff15b980e9 clarified antiquotations: no need for Thm.beta_conversion, due to Thm.instantiate_beta;
wenzelm
parents: 74617
diff changeset
   755
          end)
74616
wenzelm
parents: 74615
diff changeset
   756
    in first_conv [elim_abs, elim_max, elim_min, all_conv] end;
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   757
in
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   758
fun gen_real_arith ctxt (mkconst,eq,ge,gt,norm,neg,add,mul,prover) =
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   759
  gen_gen_real_arith ctxt
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   760
    (mkconst,eq,ge,gt,norm,neg,add,mul,
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   761
     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
   762
end;
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   763
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   764
(* 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
   765
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   766
fun gen_prover_real_arith ctxt prover =
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   767
  let
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   768
    val {add, mul, neg, pow = _, sub = _, main} =
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   769
        Semiring_Normalizer.semiring_normalizers_ord_wrapper ctxt
67267
c5994f1fa0fa more symbols;
wenzelm
parents: 67232
diff changeset
   770
        (the (Semiring_Normalizer.match ctxt \<^cterm>\<open>(0::real) + 1\<close>))
67562
2427d3e72b6e clarified signature;
wenzelm
parents: 67560
diff changeset
   771
        Thm.term_ord
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   772
  in gen_real_arith ctxt
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 46594
diff changeset
   773
     (cterm_of_rat,
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 46594
diff changeset
   774
      Numeral_Simprocs.field_comp_conv ctxt,
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 46594
diff changeset
   775
      Numeral_Simprocs.field_comp_conv ctxt,
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 46594
diff changeset
   776
      Numeral_Simprocs.field_comp_conv ctxt,
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 46594
diff changeset
   777
      main ctxt, neg ctxt, add ctxt, mul ctxt, prover)
46594
f11f332b964f tuned whitespace
huffman
parents: 46497
diff changeset
   778
  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
   779
fc654c95c29e A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination
chaieb
parents:
diff changeset
   780
end