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