src/HOL/Library/normarith.ML
changeset 29841 86d94bb79226
child 29843 4bb780545478
equal deleted inserted replaced
29840:cfab6a76aa13 29841:86d94bb79226
       
     1 (* A functor for finite mappings based on Tables *)
       
     2 signature FUNC = 
       
     3 sig
       
     4  type 'a T
       
     5  type key
       
     6  val apply : 'a T -> key -> 'a
       
     7  val applyd :'a T -> (key -> 'a) -> key -> 'a
       
     8  val combine : ('a -> 'a -> 'a) -> ('a -> bool) -> 'a T -> 'a T -> 'a T
       
     9  val defined : 'a T -> key -> bool
       
    10  val dom : 'a T -> key list
       
    11  val fold : (key * 'a -> 'b -> 'b) -> 'a T -> 'b -> 'b
       
    12  val graph : 'a T -> (key * 'a) list
       
    13  val is_undefined : 'a T -> bool
       
    14  val mapf : ('a -> 'b) -> 'a T -> 'b T
       
    15  val tryapplyd : 'a T -> key -> 'a -> 'a
       
    16  val undefine :  key -> 'a T -> 'a T
       
    17  val undefined : 'a T
       
    18  val update : key * 'a -> 'a T -> 'a T
       
    19  val updatep : (key * 'a -> bool) -> key * 'a -> 'a T -> 'a T
       
    20  val choose : 'a T -> key * 'a
       
    21  val onefunc : key * 'a -> 'a T
       
    22  val get_first: (key*'a -> 'a option) -> 'a T -> 'a option
       
    23  val fns: 
       
    24    {key_ord: key*key -> order,
       
    25     apply : 'a T -> key -> 'a,
       
    26     applyd :'a T -> (key -> 'a) -> key -> 'a,
       
    27     combine : ('a -> 'a -> 'a) -> ('a -> bool) -> 'a T -> 'a T -> 'a T,
       
    28     defined : 'a T -> key -> bool,
       
    29     dom : 'a T -> key list,
       
    30     fold : (key * 'a -> 'b -> 'b) -> 'a T -> 'b -> 'b,
       
    31     graph : 'a T -> (key * 'a) list,
       
    32     is_undefined : 'a T -> bool,
       
    33     mapf : ('a -> 'b) -> 'a T -> 'b T,
       
    34     tryapplyd : 'a T -> key -> 'a -> 'a,
       
    35     undefine :  key -> 'a T -> 'a T,
       
    36     undefined : 'a T,
       
    37     update : key * 'a -> 'a T -> 'a T,
       
    38     updatep : (key * 'a -> bool) -> key * 'a -> 'a T -> 'a T,
       
    39     choose : 'a T -> key * 'a,
       
    40     onefunc : key * 'a -> 'a T,
       
    41     get_first: (key*'a -> 'a option) -> 'a T -> 'a option}
       
    42 end;
       
    43 
       
    44 functor FuncFun(Key: KEY) : FUNC=
       
    45 struct
       
    46 
       
    47 type key = Key.key;
       
    48 structure Tab = TableFun(Key);
       
    49 type 'a T = 'a Tab.table;
       
    50 
       
    51 val undefined = Tab.empty;
       
    52 val is_undefined = Tab.is_empty;
       
    53 val mapf = Tab.map;
       
    54 val fold = Tab.fold;
       
    55 val graph = Tab.dest;
       
    56 val dom = Tab.keys;
       
    57 fun applyd f d x = case Tab.lookup f x of 
       
    58    SOME y => y
       
    59  | NONE => d x;
       
    60 
       
    61 fun apply f x = applyd f (fn _ => raise Tab.UNDEF x) x;
       
    62 fun tryapplyd f a d = applyd f (K d) a;
       
    63 val defined = Tab.defined;
       
    64 fun undefine x t = (Tab.delete x t handle UNDEF => t);
       
    65 val update = Tab.update;
       
    66 fun updatep p (k,v) t = if p (k, v) then t else update (k,v) t
       
    67 fun combine f z a b = 
       
    68  let
       
    69   fun h (k,v) t = case Tab.lookup t k of
       
    70      NONE => Tab.update (k,v) t
       
    71    | SOME v' => let val w = f v v'
       
    72      in if z w then Tab.delete k t else Tab.update (k,w) t end;
       
    73   in Tab.fold h a b end;
       
    74 
       
    75 fun choose f = case Tab.max_key f of 
       
    76    SOME k => (k,valOf (Tab.lookup f k))
       
    77  | NONE => error "FuncFun.choose : Completely undefined function"
       
    78 
       
    79 fun onefunc kv = update kv undefined
       
    80 
       
    81 local
       
    82 fun  find f (k,v) NONE = f (k,v)
       
    83    | find f (k,v) r = r
       
    84 in
       
    85 fun get_first f t = fold (find f) t NONE
       
    86 end
       
    87 
       
    88 val fns = 
       
    89    {key_ord = Key.ord,
       
    90     apply = apply,
       
    91     applyd = applyd,
       
    92     combine = combine,
       
    93     defined = defined,
       
    94     dom = dom,
       
    95     fold = fold,
       
    96     graph = graph,
       
    97     is_undefined = is_undefined,
       
    98     mapf = mapf,
       
    99     tryapplyd = tryapplyd,
       
   100     undefine = undefine,
       
   101     undefined = undefined,
       
   102     update = update,
       
   103     updatep = updatep,
       
   104     choose = choose,
       
   105     onefunc = onefunc,
       
   106     get_first = get_first}
       
   107 
       
   108 end;
       
   109 
       
   110 structure Intfunc = FuncFun(type key = int val ord = int_ord);
       
   111 structure Symfunc = FuncFun(type key = string val ord = fast_string_ord);
       
   112 structure Termfunc = FuncFun(type key = term val ord = TermOrd.fast_term_ord);
       
   113 structure Ctermfunc = FuncFun(type key = cterm val ord = (fn (s,t) => TermOrd.fast_term_ord(term_of s, term_of t)));
       
   114 structure Ratfunc = FuncFun(type key = Rat.rat val ord = Rat.ord);
       
   115 
       
   116     (* Some conversions-related stuff which has been forbidden entrance into Pure/conv.ML*)
       
   117 structure Conv2 = 
       
   118 struct
       
   119  open Conv
       
   120 fun instantiate_cterm' ty tms = Drule.cterm_rule (Drule.instantiate' ty tms)
       
   121 fun is_comb t = case (term_of t) of _$_ => true | _ => false;
       
   122 fun is_abs t = case (term_of t) of Abs _ => true | _ => false;
       
   123 
       
   124 fun end_itlist f l =
       
   125  case l of 
       
   126    []     => error "end_itlist"
       
   127  | [x]    => x
       
   128  | (h::t) => f h (end_itlist f t);
       
   129 
       
   130  fun absc cv ct = case term_of ct of 
       
   131  Abs (v,_, _) => 
       
   132   let val (x,t) = Thm.dest_abs (SOME v) ct
       
   133   in Thm.abstract_rule ((fst o dest_Free o term_of) x) x (cv t)
       
   134   end
       
   135  | _ => all_conv ct;
       
   136 
       
   137 fun cache_conv conv =
       
   138  let 
       
   139   val tab = ref Termtab.empty
       
   140   fun cconv t =  
       
   141     case Termtab.lookup (!tab) (term_of t) of
       
   142      SOME th => th
       
   143    | NONE => let val th = conv t
       
   144              in ((tab := Termtab.insert Thm.eq_thm (term_of t, th) (!tab)); th) end
       
   145  in cconv end;
       
   146 fun is_binop ct ct' = ct aconvc (Thm.dest_fun (Thm.dest_fun ct'))
       
   147   handle CTERM _ => false;
       
   148 
       
   149 local
       
   150  fun thenqc conv1 conv2 tm =
       
   151    case try conv1 tm of
       
   152     SOME th1 => (case try conv2 (Thm.rhs_of th1) of SOME th2 => Thm.transitive th1 th2 | NONE => th1)
       
   153   | NONE => conv2 tm
       
   154 
       
   155  fun thencqc conv1 conv2 tm =
       
   156     let val th1 = conv1 tm 
       
   157     in (case try conv2 (Thm.rhs_of th1) of SOME th2 => Thm.transitive th1 th2 | NONE => th1)
       
   158     end
       
   159  fun comb_qconv conv tm =
       
   160    let val (l,r) = Thm.dest_comb tm 
       
   161    in (case try conv l of 
       
   162         SOME th1 => (case try conv r of SOME th2 => Thm.combination th1 th2 
       
   163                                       | NONE => Drule.fun_cong_rule th1 r)
       
   164       | NONE => Drule.arg_cong_rule l (conv r))
       
   165    end
       
   166  fun repeatqc conv tm = thencqc conv (repeatqc conv) tm 
       
   167  fun sub_qconv conv tm =  if is_abs tm then absc conv tm else comb_qconv conv tm 
       
   168  fun once_depth_qconv conv tm =
       
   169       (conv else_conv (sub_qconv (once_depth_qconv conv))) tm
       
   170  fun depth_qconv conv tm =
       
   171     thenqc (sub_qconv (depth_qconv conv))
       
   172            (repeatqc conv) tm
       
   173  fun redepth_qconv conv tm =
       
   174     thenqc (sub_qconv (redepth_qconv conv))
       
   175            (thencqc conv (redepth_qconv conv)) tm
       
   176  fun top_depth_qconv conv tm =
       
   177     thenqc (repeatqc conv)
       
   178            (thencqc (sub_qconv (top_depth_qconv conv))
       
   179                     (thencqc conv (top_depth_qconv conv))) tm
       
   180  fun top_sweep_qconv conv tm =
       
   181     thenqc (repeatqc conv)
       
   182            (sub_qconv (top_sweep_qconv conv)) tm
       
   183 in 
       
   184 val (once_depth_conv, depth_conv, rdepth_conv, top_depth_conv, top_sweep_conv) = 
       
   185   (fn c => try_conv (once_depth_qconv c),
       
   186    fn c => try_conv (depth_qconv c),
       
   187    fn c => try_conv (redepth_qconv c),
       
   188    fn c => try_conv (top_depth_qconv c),
       
   189    fn c => try_conv (top_sweep_qconv c));
       
   190 end;
       
   191 end;
       
   192 
       
   193 
       
   194     (* Some useful derived rules *)
       
   195 fun deduct_antisym_rule tha thb = 
       
   196     equal_intr (implies_intr (cprop_of thb) tha) 
       
   197      (implies_intr (cprop_of tha) thb);
       
   198 
       
   199 fun prove_hyp tha thb = 
       
   200   if exists (curry op aconv (concl_of tha)) (#hyps (rep_thm thb)) 
       
   201   then equal_elim (symmetric (deduct_antisym_rule tha thb)) tha else thb;
       
   202 
       
   203 
       
   204 
       
   205 signature REAL_ARITH = 
       
   206 sig
       
   207   datatype positivstellensatz =
       
   208    Axiom_eq of int
       
   209  | Axiom_le of int
       
   210  | Axiom_lt of int
       
   211  | Rational_eq of Rat.rat
       
   212  | Rational_le of Rat.rat
       
   213  | Rational_lt of Rat.rat
       
   214  | Square of cterm
       
   215  | Eqmul of cterm * positivstellensatz
       
   216  | Sum of positivstellensatz * positivstellensatz
       
   217  | Product of positivstellensatz * positivstellensatz;
       
   218 
       
   219 val gen_gen_real_arith :
       
   220   Proof.context -> (Rat.rat -> Thm.cterm) * conv * conv * conv * 
       
   221    conv * conv * conv * conv * conv * conv * 
       
   222     ( (thm list * thm list * thm list -> positivstellensatz -> thm) ->
       
   223         thm list * thm list * thm list -> thm) -> conv
       
   224 val real_linear_prover : 
       
   225   (thm list * thm list * thm list -> positivstellensatz -> thm) ->
       
   226    thm list * thm list * thm list -> thm
       
   227 
       
   228 val gen_real_arith : Proof.context ->
       
   229    (Rat.rat -> cterm) * conv * conv * conv * conv * conv * conv * conv *
       
   230    ( (thm list * thm list * thm list -> positivstellensatz -> thm) ->
       
   231        thm list * thm list * thm list -> thm) -> conv
       
   232 val gen_prover_real_arith : Proof.context ->
       
   233    ((thm list * thm list * thm list -> positivstellensatz -> thm) ->
       
   234      thm list * thm list * thm list -> thm) -> conv
       
   235 val real_arith : Proof.context -> conv
       
   236 end
       
   237 
       
   238 structure RealArith (* : REAL_ARITH *)=
       
   239 struct
       
   240 
       
   241  open Conv Thm Conv2;;
       
   242 (* ------------------------------------------------------------------------- *)
       
   243 (* Data structure for Positivstellensatz refutations.                        *)
       
   244 (* ------------------------------------------------------------------------- *)
       
   245 
       
   246 datatype positivstellensatz =
       
   247    Axiom_eq of int
       
   248  | Axiom_le of int
       
   249  | Axiom_lt of int
       
   250  | Rational_eq of Rat.rat
       
   251  | Rational_le of Rat.rat
       
   252  | Rational_lt of Rat.rat
       
   253  | Square of cterm
       
   254  | Eqmul of cterm * positivstellensatz
       
   255  | Sum of positivstellensatz * positivstellensatz
       
   256  | Product of positivstellensatz * positivstellensatz;
       
   257          (* Theorems used in the procedure *)
       
   258 
       
   259 fun conjunctions th = case try Conjunction.elim th of
       
   260    SOME (th1,th2) => (conjunctions th1) @ conjunctions th2
       
   261  | NONE => [th];
       
   262 
       
   263 val pth = @{lemma "(((x::real) < y) == (y - x > 0)) &&& ((x <= y) == (y - x >= 0)) 
       
   264      &&& ((x = y) == (x - y = 0)) &&& ((~(x < y)) == (x - y >= 0)) &&& ((~(x <= y)) == (x - y > 0))
       
   265      &&& ((~(x = y)) == (x - y > 0 | -(x - y) > 0))"
       
   266   by (atomize (full), auto simp add: less_diff_eq le_diff_eq not_less)} |> 
       
   267 conjunctions;
       
   268 
       
   269 val pth_final = @{lemma "(~p ==> False) ==> p" by blast}
       
   270 val pth_add = 
       
   271  @{lemma "(x = (0::real) ==> y = 0 ==> x + y = 0 ) &&& ( x = 0 ==> y >= 0 ==> x + y >= 0) 
       
   272     &&& (x = 0 ==> y > 0 ==> x + y > 0) &&& (x >= 0 ==> y = 0 ==> x + y >= 0) 
       
   273     &&& (x >= 0 ==> y >= 0 ==> x + y >= 0) &&& (x >= 0 ==> y > 0 ==> x + y > 0) 
       
   274     &&& (x > 0 ==> y = 0 ==> x + y > 0) &&& (x > 0 ==> y >= 0 ==> x + y > 0) 
       
   275     &&& (x > 0 ==> y > 0 ==> x + y > 0)"  by simp_all} |> conjunctions ;
       
   276 
       
   277 val pth_mul = 
       
   278   @{lemma "(x = (0::real) ==> y = 0 ==> x * y = 0) &&& (x = 0 ==> y >= 0 ==> x * y = 0) &&& 
       
   279            (x = 0 ==> y > 0 ==> x * y = 0) &&& (x >= 0 ==> y = 0 ==> x * y = 0) &&& 
       
   280            (x >= 0 ==> y >= 0 ==> x * y >= 0 ) &&& ( x >= 0 ==> y > 0 ==> x * y >= 0 ) &&&
       
   281            (x > 0 ==>  y = 0 ==> x * y = 0 ) &&& ( x > 0 ==> y >= 0 ==> x * y >= 0 ) &&&
       
   282            (x > 0 ==>  y > 0 ==> x * y > 0)"
       
   283   by (auto intro: mult_mono[where a="0::real" and b="x" and d="y" and c="0", simplified]
       
   284     mult_strict_mono[where b="x" and d="y" and a="0" and c="0", simplified])} |> conjunctions;
       
   285 
       
   286 val pth_emul = @{lemma "y = (0::real) ==> x * y = 0"  by simp};
       
   287 val pth_square = @{lemma "x * x >= (0::real)"  by simp};
       
   288 
       
   289 val weak_dnf_simps = List.take (simp_thms, 34) 
       
   290     @ conjunctions @{lemma "((P & (Q | R)) = ((P&Q) | (P&R))) &&& ((Q | R) & P) = ((Q&P) | (R&P)) &&& (P & Q) = (Q & P) &&& ((P | Q) = (Q | P))" by blast+};
       
   291 
       
   292 val nnfD_simps = conjunctions @{lemma "((~(P & Q)) = (~P | ~Q)) &&& ((~(P | Q)) = (~P & ~Q) ) &&& ((P --> Q) = (~P | Q) ) &&& ((P = Q) = ((P & Q) | (~P & ~ Q))) &&& ((~(P = Q)) = ((P & ~ Q) | (~P & Q)) ) &&& ((~ ~(P)) = P)" by blast+}
       
   293 
       
   294 val choice_iff = @{lemma "(ALL x. EX y. P x y) = (EX f. ALL x. P x (f x))" by metis};
       
   295 val prenex_simps = map (fn th => th RS sym) ([@{thm "all_conj_distrib"}, @{thm "ex_disj_distrib"}] @ @{thms "all_simps"(1-4)} @ @{thms "ex_simps"(1-4)});
       
   296 
       
   297 val real_abs_thms1 = conjunctions @{lemma
       
   298   "((-1 * abs(x::real) >= r) = (-1 * x >= r & 1 * x >= r)) &&&
       
   299   ((-1 * abs(x) + a >= r) = (a + -1 * x >= r & a + 1 * x >= r)) &&&
       
   300   ((a + -1 * abs(x) >= r) = (a + -1 * x >= r & a + 1 * x >= r)) &&&
       
   301   ((a + -1 * abs(x) + b >= r) = (a + -1 * x + b >= r & a + 1 * x + b >= r)) &&&
       
   302   ((a + b + -1 * abs(x) >= r) = (a + b + -1 * x >= r & a + b + 1 * x >= r)) &&&
       
   303   ((a + b + -1 * abs(x) + c >= r) = (a + b + -1 * x + c >= r & a + b + 1 * x + c >= r)) &&&
       
   304   ((-1 * max x y >= r) = (-1 * x >= r & -1 * y >= r)) &&&
       
   305   ((-1 * max x y + a >= r) = (a + -1 * x >= r & a + -1 * y >= r)) &&&
       
   306   ((a + -1 * max x y >= r) = (a + -1 * x >= r & a + -1 * y >= r)) &&&
       
   307   ((a + -1 * max x y + b >= r) = (a + -1 * x + b >= r & a + -1 * y  + b >= r)) &&&
       
   308   ((a + b + -1 * max x y >= r) = (a + b + -1 * x >= r & a + b + -1 * y >= r)) &&&
       
   309   ((a + b + -1 * max x y + c >= r) = (a + b + -1 * x + c >= r & a + b + -1 * y  + c >= r)) &&&
       
   310   ((1 * min x y >= r) = (1 * x >= r & 1 * y >= r)) &&&
       
   311   ((1 * min x y + a >= r) = (a + 1 * x >= r & a + 1 * y >= r)) &&&
       
   312   ((a + 1 * min x y >= r) = (a + 1 * x >= r & a + 1 * y >= r)) &&&
       
   313   ((a + 1 * min x y + b >= r) = (a + 1 * x + b >= r & a + 1 * y  + b >= r) )&&&
       
   314   ((a + b + 1 * min x y >= r) = (a + b + 1 * x >= r & a + b + 1 * y >= r)) &&&
       
   315   ((a + b + 1 * min x y + c >= r) = (a + b + 1 * x + c >= r & a + b + 1 * y  + c >= r)) &&&
       
   316   ((min x y >= r) = (x >= r &  y >= r)) &&&
       
   317   ((min x y + a >= r) = (a + x >= r & a + y >= r)) &&&
       
   318   ((a + min x y >= r) = (a + x >= r & a + y >= r)) &&&
       
   319   ((a + min x y + b >= r) = (a + x + b >= r & a + y  + b >= r)) &&&
       
   320   ((a + b + min x y >= r) = (a + b + x >= r & a + b + y >= r) )&&&
       
   321   ((a + b + min x y + c >= r) = (a + b + x + c >= r & a + b + y + c >= r)) &&&
       
   322   ((-1 * abs(x) > r) = (-1 * x > r & 1 * x > r)) &&&
       
   323   ((-1 * abs(x) + a > r) = (a + -1 * x > r & a + 1 * x > r)) &&&
       
   324   ((a + -1 * abs(x) > r) = (a + -1 * x > r & a + 1 * x > r)) &&&
       
   325   ((a + -1 * abs(x) + b > r) = (a + -1 * x + b > r & a + 1 * x + b > r)) &&&
       
   326   ((a + b + -1 * abs(x) > r) = (a + b + -1 * x > r & a + b + 1 * x > r)) &&&
       
   327   ((a + b + -1 * abs(x) + c > r) = (a + b + -1 * x + c > r & a + b + 1 * x + c > r)) &&&
       
   328   ((-1 * max x y > r) = ((-1 * x > r) & -1 * y > r)) &&&
       
   329   ((-1 * max x y + a > r) = (a + -1 * x > r & a + -1 * y > r)) &&&
       
   330   ((a + -1 * max x y > r) = (a + -1 * x > r & a + -1 * y > r)) &&&
       
   331   ((a + -1 * max x y + b > r) = (a + -1 * x + b > r & a + -1 * y  + b > r)) &&&
       
   332   ((a + b + -1 * max x y > r) = (a + b + -1 * x > r & a + b + -1 * y > r)) &&&
       
   333   ((a + b + -1 * max x y + c > r) = (a + b + -1 * x + c > r & a + b + -1 * y  + c > r)) &&&
       
   334   ((min x y > r) = (x > r &  y > r)) &&&
       
   335   ((min x y + a > r) = (a + x > r & a + y > r)) &&&
       
   336   ((a + min x y > r) = (a + x > r & a + y > r)) &&&
       
   337   ((a + min x y + b > r) = (a + x + b > r & a + y  + b > r)) &&&
       
   338   ((a + b + min x y > r) = (a + b + x > r & a + b + y > r)) &&&
       
   339   ((a + b + min x y + c > r) = (a + b + x + c > r & a + b + y + c > r))"
       
   340   by auto};
       
   341 
       
   342 val abs_split' = @{lemma "P (abs (x::'a::ordered_idom)) == (x >= 0 & P x | x < 0 & P (-x))"
       
   343   by (atomize (full)) (auto split add: abs_split)};
       
   344 
       
   345 val max_split = @{lemma "P (max x y) == ((x::'a::linorder) <= y & P y | x > y & P x)"
       
   346   by (atomize (full)) (cases "x <= y", auto simp add: max_def)};
       
   347 
       
   348 val min_split = @{lemma "P (min x y) == ((x::'a::linorder) <= y & P x | x > y & P y)"
       
   349   by (atomize (full)) (cases "x <= y", auto simp add: min_def)};
       
   350 
       
   351 
       
   352          (* Miscalineous *)
       
   353 fun literals_conv bops uops cv = 
       
   354  let fun h t =
       
   355   case (term_of t) of 
       
   356    b$_$_ => if member (op aconv) bops b then binop_conv h t else cv t
       
   357  | u$_ => if member (op aconv) uops u then arg_conv h t else cv t
       
   358  | _ => cv t
       
   359  in h end;
       
   360 
       
   361 fun cterm_of_rat x = 
       
   362 let val (a, b) = Rat.quotient_of_rat x
       
   363 in 
       
   364  if b = 1 then Numeral.mk_cnumber @{ctyp "real"} a
       
   365   else Thm.capply (Thm.capply @{cterm "op / :: real => _"} 
       
   366                    (Numeral.mk_cnumber @{ctyp "real"} a))
       
   367         (Numeral.mk_cnumber @{ctyp "real"} b)
       
   368 end;
       
   369 
       
   370   fun dest_ratconst t = case term_of t of
       
   371    Const(@{const_name divide}, _)$a$b => Rat.rat_of_quotient(HOLogic.dest_number a |> snd, HOLogic.dest_number b |> snd)
       
   372  | _ => Rat.rat_of_int (HOLogic.dest_number (term_of t) |> snd)
       
   373  fun is_ratconst t = can dest_ratconst t
       
   374 
       
   375 fun find_term p t = if p t then t else 
       
   376  case t of
       
   377   a$b => (find_term p a handle TERM _ => find_term p b)
       
   378  | Abs (_,_,t') => find_term p t'
       
   379  | _ => raise TERM ("find_term",[t]);
       
   380 
       
   381 fun find_cterm p t = if p t then t else 
       
   382  case term_of t of
       
   383   a$b => (find_cterm p (Thm.dest_fun t) handle CTERM _ => find_cterm p (Thm.dest_arg t))
       
   384  | Abs (_,_,t') => find_cterm p (Thm.dest_abs NONE t |> snd)
       
   385  | _ => raise CTERM ("find_cterm",[t]);
       
   386 
       
   387 
       
   388     (* A general real arithmetic prover *)
       
   389 
       
   390 fun gen_gen_real_arith ctxt (mk_numeric,
       
   391        numeric_eq_conv,numeric_ge_conv,numeric_gt_conv,
       
   392        poly_conv,poly_neg_conv,poly_add_conv,poly_mul_conv,
       
   393        absconv1,absconv2,prover) = 
       
   394 let
       
   395  open Conv Thm; 
       
   396  val pre_ss = HOL_basic_ss addsimps simp_thms@ ex_simps@ all_simps@[@{thm not_all},@{thm not_ex},ex_disj_distrib, all_conj_distrib, @{thm if_bool_eq_disj}]
       
   397  val prenex_ss = HOL_basic_ss addsimps prenex_simps
       
   398  val skolemize_ss = HOL_basic_ss addsimps [choice_iff]
       
   399  val presimp_conv = Simplifier.rewrite (Simplifier.context ctxt pre_ss)
       
   400  val prenex_conv = Simplifier.rewrite (Simplifier.context ctxt prenex_ss)
       
   401  val skolemize_conv = Simplifier.rewrite (Simplifier.context ctxt skolemize_ss)
       
   402  val weak_dnf_ss = HOL_basic_ss addsimps weak_dnf_simps
       
   403  val weak_dnf_conv = Simplifier.rewrite (Simplifier.context ctxt weak_dnf_ss)
       
   404  fun eqT_elim th = equal_elim (symmetric th) @{thm TrueI}
       
   405  fun oprconv cv ct = 
       
   406   let val g = Thm.dest_fun2 ct
       
   407   in if g aconvc @{cterm "op <= :: real => _"} 
       
   408        orelse g aconvc @{cterm "op < :: real => _"} 
       
   409      then arg_conv cv ct else arg1_conv cv ct
       
   410   end
       
   411 
       
   412  fun real_ineq_conv th ct =
       
   413   let
       
   414    val th' = (instantiate (match (lhs_of th, ct)) th 
       
   415       handle MATCH => raise CTERM ("real_ineq_conv", [ct]))
       
   416   in transitive th' (oprconv poly_conv (Thm.rhs_of th'))
       
   417   end 
       
   418   val [real_lt_conv, real_le_conv, real_eq_conv,
       
   419        real_not_lt_conv, real_not_le_conv, _] =
       
   420        map real_ineq_conv pth
       
   421   fun match_mp_rule ths ths' = 
       
   422    let
       
   423      fun f ths ths' = case ths of [] => raise THM("match_mp_rule",0,ths)
       
   424       | th::ths => (ths' MRS th handle THM _ => f ths ths')
       
   425    in f ths ths' end
       
   426   fun mul_rule th th' = fconv_rule (arg_conv (oprconv poly_mul_conv))
       
   427          (match_mp_rule pth_mul [th, th'])
       
   428   fun add_rule th th' = fconv_rule (arg_conv (oprconv poly_add_conv))
       
   429          (match_mp_rule pth_add [th, th'])
       
   430   fun emul_rule ct th = fconv_rule (arg_conv (oprconv poly_mul_conv)) 
       
   431        (instantiate' [] [SOME ct] (th RS pth_emul)) 
       
   432   fun square_rule t = fconv_rule (arg_conv (oprconv poly_mul_conv))
       
   433        (instantiate' [] [SOME t] pth_square)
       
   434 
       
   435   fun hol_of_positivstellensatz(eqs,les,lts) =
       
   436    let 
       
   437     fun translate prf = case prf of
       
   438         Axiom_eq n => nth eqs n
       
   439       | Axiom_le n => nth les n
       
   440       | Axiom_lt n => nth lts n
       
   441       | Rational_eq x => eqT_elim(numeric_eq_conv(capply @{cterm Trueprop} 
       
   442                           (capply (capply @{cterm "op =::real => _"} (mk_numeric x)) 
       
   443                                @{cterm "0::real"})))
       
   444       | Rational_le x => eqT_elim(numeric_ge_conv(capply @{cterm Trueprop} 
       
   445                           (capply (capply @{cterm "op <=::real => _"} 
       
   446                                      @{cterm "0::real"}) (mk_numeric x))))
       
   447       | Rational_lt x => eqT_elim(numeric_gt_conv(capply @{cterm Trueprop} 
       
   448                       (capply (capply @{cterm "op <::real => _"} @{cterm "0::real"})
       
   449                         (mk_numeric x))))
       
   450       | Square t => square_rule t
       
   451       | Eqmul(t,p) => emul_rule t (translate p)
       
   452       | Sum(p1,p2) => add_rule (translate p1) (translate p2)
       
   453       | Product(p1,p2) => mul_rule (translate p1) (translate p2)
       
   454    in fn prf => 
       
   455       fconv_rule (first_conv [numeric_ge_conv, numeric_gt_conv, numeric_eq_conv, all_conv]) 
       
   456           (translate prf)
       
   457    end
       
   458   
       
   459   val init_conv = presimp_conv then_conv
       
   460       nnf_conv then_conv skolemize_conv then_conv prenex_conv then_conv
       
   461       weak_dnf_conv
       
   462 
       
   463   val concl = dest_arg o cprop_of
       
   464   fun is_binop opr ct = (dest_fun2 ct aconvc opr handle CTERM _ => false)
       
   465   val is_req = is_binop @{cterm "op =:: real => _"}
       
   466   val is_ge = is_binop @{cterm "op <=:: real => _"}
       
   467   val is_gt = is_binop @{cterm "op <:: real => _"}
       
   468   val is_conj = is_binop @{cterm "op &"}
       
   469   val is_disj = is_binop @{cterm "op |"}
       
   470   fun conj_pair th = (th RS @{thm conjunct1}, th RS @{thm conjunct2})
       
   471   fun disj_cases th th1 th2 = 
       
   472    let val (p,q) = dest_binop (concl th)
       
   473        val c = concl th1
       
   474        val _ = if c aconvc (concl th2) then () else error "disj_cases : conclusions not alpha convertible"
       
   475    in implies_elim (implies_elim (implies_elim (instantiate' [] (map SOME [p,q,c]) @{thm disjE}) th) (implies_intr (capply @{cterm Trueprop} p) th1)) (implies_intr (capply @{cterm Trueprop} q) th2)
       
   476    end
       
   477  fun overall dun ths = case ths of
       
   478   [] =>
       
   479    let 
       
   480     val (eq,ne) = List.partition (is_req o concl) dun
       
   481      val (le,nl) = List.partition (is_ge o concl) ne
       
   482      val lt = filter (is_gt o concl) nl 
       
   483     in prover hol_of_positivstellensatz (eq,le,lt) end
       
   484  | th::oths =>
       
   485    let 
       
   486     val ct = concl th 
       
   487    in 
       
   488     if is_conj ct  then
       
   489      let 
       
   490       val (th1,th2) = conj_pair th in
       
   491       overall dun (th1::th2::oths) end
       
   492     else if is_disj ct then
       
   493       let 
       
   494        val th1 = overall dun (assume (capply @{cterm Trueprop} (dest_arg1 ct))::oths)
       
   495        val th2 = overall dun (assume (capply @{cterm Trueprop} (dest_arg ct))::oths)
       
   496       in disj_cases th th1 th2 end
       
   497    else overall (th::dun) oths
       
   498   end
       
   499   fun dest_binary b ct = if is_binop b ct then dest_binop ct 
       
   500                          else raise CTERM ("dest_binary",[b,ct])
       
   501   val dest_eq = dest_binary @{cterm "op = :: real => _"}
       
   502   val neq_th = nth pth 5
       
   503   fun real_not_eq_conv ct = 
       
   504    let 
       
   505     val (l,r) = dest_eq (dest_arg ct)
       
   506     val th = instantiate ([],[(@{cpat "?x::real"},l),(@{cpat "?y::real"},r)]) neq_th
       
   507     val th_p = poly_conv(dest_arg(dest_arg1(Thm.rhs_of th)))
       
   508     val th_x = Drule.arg_cong_rule @{cterm "uminus :: real => _"} th_p
       
   509     val th_n = fconv_rule (arg_conv poly_neg_conv) th_x
       
   510     val th' = Drule.binop_cong_rule @{cterm "op |"} 
       
   511      (Drule.arg_cong_rule (capply @{cterm "op <::real=>_"} @{cterm "0::real"}) th_p)
       
   512      (Drule.arg_cong_rule (capply @{cterm "op <::real=>_"} @{cterm "0::real"}) th_n)
       
   513     in transitive th th' 
       
   514   end
       
   515  fun equal_implies_1_rule PQ = 
       
   516   let 
       
   517    val P = lhs_of PQ
       
   518   in implies_intr P (equal_elim PQ (assume P))
       
   519   end
       
   520  (* FIXME!!! Copied from groebner.ml *)
       
   521  val strip_exists =
       
   522   let fun h (acc, t) =
       
   523    case (term_of t) of
       
   524     Const("Ex",_)$Abs(x,T,p) => h (dest_abs NONE (dest_arg t) |>> (fn v => v::acc))
       
   525   | _ => (acc,t)
       
   526   in fn t => h ([],t)
       
   527   end
       
   528   fun name_of x = case term_of x of
       
   529    Free(s,_) => s
       
   530  | Var ((s,_),_) => s
       
   531  | _ => "x"
       
   532 
       
   533   fun mk_forall x th = Drule.arg_cong_rule (instantiate_cterm' [SOME (ctyp_of_term x)] [] @{cpat "All :: (?'a => bool) => _" }) (abstract_rule (name_of x) x th)
       
   534 
       
   535   val specl = fold_rev (fn x => fn th => instantiate' [] [SOME x] (th RS spec));
       
   536 
       
   537  fun ext T = Drule.cterm_rule (instantiate' [SOME T] []) @{cpat Ex}
       
   538  fun mk_ex v t = Thm.capply (ext (ctyp_of_term v)) (Thm.cabs v t)
       
   539 
       
   540  fun choose v th th' = case concl_of th of 
       
   541    @{term Trueprop} $ (Const("Ex",_)$_) => 
       
   542     let
       
   543      val p = (funpow 2 Thm.dest_arg o cprop_of) th
       
   544      val T = (hd o Thm.dest_ctyp o ctyp_of_term) p
       
   545      val th0 = fconv_rule (Thm.beta_conversion true)
       
   546          (instantiate' [SOME T] [SOME p, (SOME o Thm.dest_arg o cprop_of) th'] exE)
       
   547      val pv = (Thm.rhs_of o Thm.beta_conversion true) 
       
   548            (Thm.capply @{cterm Trueprop} (Thm.capply p v))
       
   549      val th1 = forall_intr v (implies_intr pv th')
       
   550     in implies_elim (implies_elim th0 th) th1  end
       
   551  | _ => raise THM ("choose",0,[th, th'])
       
   552 
       
   553   fun simple_choose v th = 
       
   554      choose v (assume ((Thm.capply @{cterm Trueprop} o mk_ex v) ((Thm.dest_arg o hd o #hyps o Thm.crep_thm) th))) th
       
   555 
       
   556  val strip_forall =
       
   557   let fun h (acc, t) =
       
   558    case (term_of t) of
       
   559     Const("All",_)$Abs(x,T,p) => h (dest_abs NONE (dest_arg t) |>> (fn v => v::acc))
       
   560   | _ => (acc,t)
       
   561   in fn t => h ([],t)
       
   562   end
       
   563 
       
   564  fun f ct =
       
   565   let 
       
   566    val nnf_norm_conv' = 
       
   567      nnf_conv then_conv 
       
   568      literals_conv [@{term "op &"}, @{term "op |"}] [] 
       
   569      (cache_conv 
       
   570        (first_conv [real_lt_conv, real_le_conv, 
       
   571                     real_eq_conv, real_not_lt_conv, 
       
   572                     real_not_le_conv, real_not_eq_conv, all_conv]))
       
   573   fun absremover ct = (literals_conv [@{term "op &"}, @{term "op |"}] [] 
       
   574                   (try_conv (absconv1 then_conv binop_conv (arg_conv poly_conv))) then_conv 
       
   575         try_conv (absconv2 then_conv nnf_norm_conv' then_conv binop_conv absremover)) ct
       
   576   val nct = capply @{cterm Trueprop} (capply @{cterm "Not"} ct)
       
   577   val th0 = (init_conv then_conv arg_conv nnf_norm_conv') nct
       
   578   val tm0 = dest_arg (Thm.rhs_of th0)
       
   579   val th = if tm0 aconvc @{cterm False} then equal_implies_1_rule th0 else
       
   580    let 
       
   581     val (evs,bod) = strip_exists tm0
       
   582     val (avs,ibod) = strip_forall bod
       
   583     val th1 = Drule.arg_cong_rule @{cterm Trueprop} (fold mk_forall avs (absremover ibod))
       
   584     val th2 = overall [] [specl avs (assume (Thm.rhs_of th1))]
       
   585     val th3 = fold simple_choose evs (prove_hyp (equal_elim th1 (assume (capply @{cterm Trueprop} bod))) th2)
       
   586    in  Drule.implies_intr_hyps (prove_hyp (equal_elim th0 (assume nct)) th3)
       
   587    end
       
   588   in implies_elim (instantiate' [] [SOME ct] pth_final) th
       
   589  end
       
   590 in f
       
   591 end;
       
   592 
       
   593 (* A linear arithmetic prover *)
       
   594 local
       
   595   val linear_add = Ctermfunc.combine (curry op +/) (fn z => z =/ Rat.zero)
       
   596   fun linear_cmul c = Ctermfunc.mapf (fn x => c */ x)
       
   597   val one_tm = @{cterm "1::real"}
       
   598   fun contradictory p (e,_) = ((Ctermfunc.is_undefined e) andalso not(p Rat.zero)) orelse
       
   599      ((gen_eq_set (op aconvc) (Ctermfunc.dom e, [one_tm])) andalso not(p(Ctermfunc.apply e one_tm)))
       
   600 
       
   601   fun linear_ineqs vars (les,lts) = 
       
   602    case find_first (contradictory (fn x => x >/ Rat.zero)) lts of
       
   603     SOME r => r
       
   604   | NONE => 
       
   605    (case find_first (contradictory (fn x => x >/ Rat.zero)) les of
       
   606      SOME r => r
       
   607    | NONE => 
       
   608      if null vars then error "linear_ineqs: no contradiction" else
       
   609      let 
       
   610       val ineqs = les @ lts
       
   611       fun blowup v =
       
   612        length(filter (fn (e,_) => Ctermfunc.tryapplyd e v Rat.zero =/ Rat.zero) ineqs) +
       
   613        length(filter (fn (e,_) => Ctermfunc.tryapplyd e v Rat.zero >/ Rat.zero) ineqs) *
       
   614        length(filter (fn (e,_) => Ctermfunc.tryapplyd e v Rat.zero </ Rat.zero) ineqs)
       
   615       val  v = fst(hd(sort (fn ((_,i),(_,j)) => int_ord (i,j))
       
   616                  (map (fn v => (v,blowup v)) vars)))
       
   617       fun addup (e1,p1) (e2,p2) acc =
       
   618        let 
       
   619         val c1 = Ctermfunc.tryapplyd e1 v Rat.zero 
       
   620         val c2 = Ctermfunc.tryapplyd e2 v Rat.zero
       
   621        in if c1 */ c2 >=/ Rat.zero then acc else
       
   622         let 
       
   623          val e1' = linear_cmul (Rat.abs c2) e1
       
   624          val e2' = linear_cmul (Rat.abs c1) e2
       
   625          val p1' = Product(Rational_lt(Rat.abs c2),p1)
       
   626          val p2' = Product(Rational_lt(Rat.abs c1),p2)
       
   627         in (linear_add e1' e2',Sum(p1',p2'))::acc
       
   628         end
       
   629        end
       
   630       val (les0,les1) = 
       
   631          List.partition (fn (e,_) => Ctermfunc.tryapplyd e v Rat.zero =/ Rat.zero) les
       
   632       val (lts0,lts1) = 
       
   633          List.partition (fn (e,_) => Ctermfunc.tryapplyd e v Rat.zero =/ Rat.zero) lts
       
   634       val (lesp,lesn) = 
       
   635          List.partition (fn (e,_) => Ctermfunc.tryapplyd e v Rat.zero >/ Rat.zero) les1
       
   636       val (ltsp,ltsn) = 
       
   637          List.partition (fn (e,_) => Ctermfunc.tryapplyd e v Rat.zero >/ Rat.zero) lts1
       
   638       val les' = fold_rev (fn ep1 => fold_rev (addup ep1) lesp) lesn les0
       
   639       val lts' = fold_rev (fn ep1 => fold_rev (addup ep1) (lesp@ltsp)) ltsn
       
   640                       (fold_rev (fn ep1 => fold_rev (addup ep1) (lesn@ltsn)) ltsp lts0)
       
   641      in linear_ineqs (remove (op aconvc) v vars) (les',lts')
       
   642      end)
       
   643 
       
   644   fun linear_eqs(eqs,les,lts) = 
       
   645    case find_first (contradictory (fn x => x =/ Rat.zero)) eqs of
       
   646     SOME r => r
       
   647   | NONE => (case eqs of 
       
   648     [] => 
       
   649      let val vars = remove (op aconvc) one_tm 
       
   650            (fold_rev (curry (gen_union (op aconvc)) o Ctermfunc.dom o fst) (les@lts) []) 
       
   651      in linear_ineqs vars (les,lts) end
       
   652    | (e,p)::es => 
       
   653      if Ctermfunc.is_undefined e then linear_eqs (es,les,lts) else
       
   654      let 
       
   655       val (x,c) = Ctermfunc.choose (Ctermfunc.undefine one_tm e)
       
   656       fun xform (inp as (t,q)) =
       
   657        let val d = Ctermfunc.tryapplyd t x Rat.zero in
       
   658         if d =/ Rat.zero then inp else
       
   659         let 
       
   660          val k = (Rat.neg d) */ Rat.abs c // c
       
   661          val e' = linear_cmul k e
       
   662          val t' = linear_cmul (Rat.abs c) t
       
   663          val p' = Eqmul(cterm_of_rat k,p)
       
   664          val q' = Product(Rational_lt(Rat.abs c),q) 
       
   665         in (linear_add e' t',Sum(p',q')) 
       
   666         end 
       
   667       end
       
   668      in linear_eqs(map xform es,map xform les,map xform lts)
       
   669      end)
       
   670 
       
   671   fun linear_prover (eq,le,lt) = 
       
   672    let 
       
   673     val eqs = map2 (fn p => fn n => (p,Axiom_eq n)) eq (0 upto (length eq - 1))
       
   674     val les = map2 (fn p => fn n => (p,Axiom_le n)) le (0 upto (length le - 1))
       
   675     val lts = map2 (fn p => fn n => (p,Axiom_lt n)) lt (0 upto (length lt - 1))
       
   676    in linear_eqs(eqs,les,lts)
       
   677    end 
       
   678   
       
   679   fun lin_of_hol ct = 
       
   680    if ct aconvc @{cterm "0::real"} then Ctermfunc.undefined
       
   681    else if not (is_comb ct) then Ctermfunc.onefunc (ct, Rat.one)
       
   682    else if is_ratconst ct then Ctermfunc.onefunc (one_tm, dest_ratconst ct)
       
   683    else
       
   684     let val (lop,r) = Thm.dest_comb ct 
       
   685     in if not (is_comb lop) then Ctermfunc.onefunc (ct, Rat.one)
       
   686        else
       
   687         let val (opr,l) = Thm.dest_comb lop 
       
   688         in if opr aconvc @{cterm "op + :: real =>_"} 
       
   689            then linear_add (lin_of_hol l) (lin_of_hol r)
       
   690            else if opr aconvc @{cterm "op * :: real =>_"} 
       
   691                    andalso is_ratconst l then Ctermfunc.onefunc (r, dest_ratconst l)
       
   692            else Ctermfunc.onefunc (ct, Rat.one)
       
   693         end
       
   694     end
       
   695 
       
   696   fun is_alien ct = case term_of ct of 
       
   697    Const(@{const_name "real"}, _)$ n => 
       
   698      if can HOLogic.dest_number n then false else true
       
   699   | _ => false
       
   700  open Thm
       
   701 in 
       
   702 fun real_linear_prover translator (eq,le,lt) = 
       
   703  let 
       
   704   val lhs = lin_of_hol o dest_arg1 o dest_arg o cprop_of
       
   705   val rhs = lin_of_hol o dest_arg o dest_arg o cprop_of
       
   706   val eq_pols = map lhs eq
       
   707   val le_pols = map rhs le
       
   708   val lt_pols = map rhs lt 
       
   709   val aliens =  filter is_alien
       
   710       (fold_rev (curry (gen_union (op aconvc)) o Ctermfunc.dom) 
       
   711           (eq_pols @ le_pols @ lt_pols) [])
       
   712   val le_pols' = le_pols @ map (fn v => Ctermfunc.onefunc (v,Rat.one)) aliens
       
   713   val (_,proof) = linear_prover (eq_pols,le_pols',lt_pols)
       
   714   val le' = le @ map (fn a => instantiate' [] [SOME (dest_arg a)] @{thm real_of_nat_ge_zero}) aliens 
       
   715  in (translator (eq,le',lt) proof) : thm
       
   716  end
       
   717 end;
       
   718 
       
   719 (* A less general generic arithmetic prover dealing with abs,max and min*)
       
   720 
       
   721 local
       
   722  val absmaxmin_elim_ss1 = HOL_basic_ss addsimps real_abs_thms1
       
   723  fun absmaxmin_elim_conv1 ctxt = 
       
   724     Simplifier.rewrite (Simplifier.context ctxt absmaxmin_elim_ss1)
       
   725 
       
   726  val absmaxmin_elim_conv2 =
       
   727   let 
       
   728    val pth_abs = instantiate' [SOME @{ctyp real}] [] abs_split'
       
   729    val pth_max = instantiate' [SOME @{ctyp real}] [] max_split
       
   730    val pth_min = instantiate' [SOME @{ctyp real}] [] min_split
       
   731    val abs_tm = @{cterm "abs :: real => _"}
       
   732    val p_tm = @{cpat "?P :: real => bool"}
       
   733    val x_tm = @{cpat "?x :: real"}
       
   734    val y_tm = @{cpat "?y::real"}
       
   735    val is_max = is_binop @{cterm "max :: real => _"}
       
   736    val is_min = is_binop @{cterm "min :: real => _"} 
       
   737    fun is_abs t = is_comb t andalso dest_fun t aconvc abs_tm
       
   738    fun eliminate_construct p c tm =
       
   739     let 
       
   740      val t = find_cterm p tm
       
   741      val th0 = (symmetric o beta_conversion false) (capply (cabs t tm) t)
       
   742      val (p,ax) = (dest_comb o Thm.rhs_of) th0
       
   743     in fconv_rule(arg_conv(binop_conv (arg_conv (beta_conversion false))))
       
   744                (transitive th0 (c p ax))
       
   745    end
       
   746 
       
   747    val elim_abs = eliminate_construct is_abs
       
   748     (fn p => fn ax => 
       
   749        instantiate ([], [(p_tm,p), (x_tm, dest_arg ax)]) pth_abs)
       
   750    val elim_max = eliminate_construct is_max
       
   751     (fn p => fn ax => 
       
   752       let val (ax,y) = dest_comb ax 
       
   753       in  instantiate ([], [(p_tm,p), (x_tm, dest_arg ax), (y_tm,y)]) 
       
   754       pth_max end)
       
   755    val elim_min = eliminate_construct is_min
       
   756     (fn p => fn ax => 
       
   757       let val (ax,y) = dest_comb ax 
       
   758       in  instantiate ([], [(p_tm,p), (x_tm, dest_arg ax), (y_tm,y)]) 
       
   759       pth_min end)
       
   760    in first_conv [elim_abs, elim_max, elim_min, all_conv]
       
   761   end;
       
   762 in fun gen_real_arith ctxt (mkconst,eq,ge,gt,norm,neg,add,mul,prover) =
       
   763         gen_gen_real_arith ctxt (mkconst,eq,ge,gt,norm,neg,add,mul,
       
   764                        absmaxmin_elim_conv1 ctxt,absmaxmin_elim_conv2,prover)
       
   765 end;
       
   766 
       
   767 (* An instance for reals*) 
       
   768 
       
   769 fun gen_prover_real_arith ctxt prover = 
       
   770  let
       
   771   fun simple_cterm_ord t u = TermOrd.term_ord (term_of t, term_of u) = LESS
       
   772   val {add,mul,neg,pow,sub,main} = 
       
   773      Normalizer.semiring_normalizers_ord_wrapper ctxt
       
   774       (valOf (NormalizerData.match ctxt @{cterm "(0::real) + 1"})) 
       
   775      simple_cterm_ord
       
   776 in gen_real_arith ctxt
       
   777    (cterm_of_rat, field_comp_conv, field_comp_conv,field_comp_conv,
       
   778     main,neg,add,mul, prover)
       
   779 end;
       
   780 
       
   781 fun real_arith ctxt = gen_prover_real_arith ctxt real_linear_prover;
       
   782 end
       
   783 
       
   784   (* Now the norm procedure for euclidean spaces *)
       
   785 
       
   786 
       
   787 signature NORM_ARITH = 
       
   788 sig
       
   789  val norm_arith : Proof.context -> conv
       
   790  val norm_arith_tac : Proof.context -> int -> tactic
       
   791 end
       
   792 
       
   793 structure NormArith : NORM_ARITH = 
       
   794 struct
       
   795 
       
   796  open Conv Thm Conv2;
       
   797  val bool_eq = op = : bool *bool -> bool
       
   798  fun dest_ratconst t = case term_of t of
       
   799    Const(@{const_name divide}, _)$a$b => Rat.rat_of_quotient(HOLogic.dest_number a |> snd, HOLogic.dest_number b |> snd)
       
   800  | _ => Rat.rat_of_int (HOLogic.dest_number (term_of t) |> snd)
       
   801  fun is_ratconst t = can dest_ratconst t
       
   802  fun augment_norm b t acc = case term_of t of 
       
   803      Const(@{const_name norm}, _) $ _ => insert (eq_pair bool_eq (op aconvc)) (b,dest_arg t) acc
       
   804    | _ => acc
       
   805  fun find_normedterms t acc = case term_of t of
       
   806     @{term "op + :: real => _"}$_$_ =>
       
   807             find_normedterms (dest_arg1 t) (find_normedterms (dest_arg t) acc)
       
   808       | @{term "op * :: real => _"}$_$n =>
       
   809             if not (is_ratconst (dest_arg1 t)) then acc else
       
   810             augment_norm (dest_ratconst (dest_arg1 t) >=/ Rat.zero) 
       
   811                       (dest_arg t) acc
       
   812       | _ => augment_norm true t acc 
       
   813 
       
   814  val cterm_lincomb_neg = Ctermfunc.mapf Rat.neg
       
   815  fun cterm_lincomb_cmul c t = 
       
   816     if c =/ Rat.zero then Ctermfunc.undefined else Ctermfunc.mapf (fn x => x */ c) t
       
   817  fun cterm_lincomb_add l r = Ctermfunc.combine (curry op +/) (fn x => x =/ Rat.zero) l r
       
   818  fun cterm_lincomb_sub l r = cterm_lincomb_add l (cterm_lincomb_neg r)
       
   819  fun cterm_lincomb_eq l r = Ctermfunc.is_undefined (cterm_lincomb_sub l r)
       
   820 
       
   821  val int_lincomb_neg = Intfunc.mapf Rat.neg
       
   822  fun int_lincomb_cmul c t = 
       
   823     if c =/ Rat.zero then Intfunc.undefined else Intfunc.mapf (fn x => x */ c) t
       
   824  fun int_lincomb_add l r = Intfunc.combine (curry op +/) (fn x => x =/ Rat.zero) l r
       
   825  fun int_lincomb_sub l r = int_lincomb_add l (int_lincomb_neg r)
       
   826  fun int_lincomb_eq l r = Intfunc.is_undefined (int_lincomb_sub l r)
       
   827 
       
   828 fun vector_lincomb t = case term_of t of 
       
   829    Const(@{const_name plus},Type("fun",[Type("Finite_Cartesian_Product.^",_),_])) $ _ $ _ =>
       
   830     cterm_lincomb_add (vector_lincomb (dest_arg1 t)) (vector_lincomb (dest_arg t))
       
   831  | Const(@{const_name minus},Type("fun",[Type("Finite_Cartesian_Product.^",_),_])) $ _ $ _ =>
       
   832     cterm_lincomb_sub (vector_lincomb (dest_arg1 t)) (vector_lincomb (dest_arg t))
       
   833  | Const(@{const_name vector_scalar_mult},Type("fun",[Type("Finite_Cartesian_Product.^",_),_]))$_$_ =>
       
   834     cterm_lincomb_cmul (dest_ratconst (dest_arg1 t)) (vector_lincomb (dest_arg t))
       
   835  | Const(@{const_name uminus},Type("fun",[Type("Finite_Cartesian_Product.^",_),_]))$_ =>
       
   836      cterm_lincomb_neg (vector_lincomb (dest_arg t))
       
   837  | Const(@{const_name vec},_)$_ => 
       
   838    let 
       
   839      val b = ((snd o HOLogic.dest_number o term_of o dest_arg) t = 0 
       
   840                handle TERM _=> false)
       
   841    in if b then Ctermfunc.onefunc (t,Rat.one)
       
   842       else Ctermfunc.undefined
       
   843    end
       
   844  | _ => Ctermfunc.onefunc (t,Rat.one)
       
   845 
       
   846  fun vector_lincombs ts =
       
   847   fold_rev 
       
   848    (fn t => fn fns => case AList.lookup (op aconvc) fns t of
       
   849      NONE => 
       
   850        let val f = vector_lincomb t 
       
   851        in case find_first (fn (_,f') => cterm_lincomb_eq f f') fns of
       
   852            SOME (_,f') => (t,f') :: fns
       
   853          | NONE => (t,f) :: fns 
       
   854        end
       
   855    | SOME _ => fns) ts []
       
   856 
       
   857 fun replacenegnorms cv t = case term_of t of 
       
   858   @{term "op + :: real => _"}$_$_ => binop_conv (replacenegnorms cv) t
       
   859 | @{term "op * :: real => _"}$_$_ => 
       
   860     if dest_ratconst (dest_arg1 t) </ Rat.zero then arg_conv cv t else reflexive t
       
   861 | _ => reflexive t
       
   862 fun flip v eq = 
       
   863   if Ctermfunc.defined eq v 
       
   864   then Ctermfunc.update (v, Rat.neg (Ctermfunc.apply eq v)) eq else eq
       
   865 fun allsubsets s = case s of 
       
   866   [] => [[]]
       
   867 |(a::t) => let val res = allsubsets t in
       
   868                map (cons a) res @ res end
       
   869 fun evaluate env lin =
       
   870  Intfunc.fold (fn (x,c) => fn s => s +/ c */ (Intfunc.apply env x)) 
       
   871    lin Rat.zero
       
   872 
       
   873 fun solve (vs,eqs) = case (vs,eqs) of
       
   874   ([],[]) => SOME (Intfunc.onefunc (0,Rat.one))
       
   875  |(_,eq::oeqs) => 
       
   876    (case vs inter (Intfunc.dom eq) of
       
   877      [] => NONE
       
   878     | v::_ => 
       
   879        if Intfunc.defined eq v 
       
   880        then 
       
   881         let 
       
   882          val c = Intfunc.apply eq v
       
   883          val vdef = int_lincomb_cmul (Rat.neg (Rat.inv c)) eq
       
   884          fun eliminate eqn = if not (Intfunc.defined eqn v) then eqn 
       
   885                              else int_lincomb_add (int_lincomb_cmul (Intfunc.apply eqn v) vdef) eqn
       
   886         in (case solve (vs \ v,map eliminate oeqs) of
       
   887             NONE => NONE
       
   888           | SOME soln => SOME (Intfunc.update (v, evaluate soln (Intfunc.undefine v vdef)) soln))
       
   889         end
       
   890        else NONE)
       
   891 
       
   892 fun combinations k l = if k = 0 then [[]] else
       
   893  case l of 
       
   894   [] => []
       
   895 | h::t => map (cons h) (combinations (k - 1) t) @ combinations k t
       
   896 
       
   897 
       
   898 fun forall2 p l1 l2 = case (l1,l2) of 
       
   899    ([],[]) => true
       
   900  | (h1::t1,h2::t2) => p h1 h2 andalso forall2 p t1 t2
       
   901  | _ => false;
       
   902 
       
   903 
       
   904 fun vertices vs eqs =
       
   905  let 
       
   906   fun vertex cmb = case solve(vs,cmb) of
       
   907     NONE => NONE
       
   908    | SOME soln => SOME (map (fn v => Intfunc.tryapplyd soln v Rat.zero) vs)
       
   909   val rawvs = map_filter vertex (combinations (length vs) eqs)
       
   910   val unset = filter (forall (fn c => c >=/ Rat.zero)) rawvs 
       
   911  in fold_rev (insert (uncurry (forall2 (curry op =/)))) unset [] 
       
   912  end 
       
   913 
       
   914 fun subsumes l m = forall2 (fn x => fn y => Rat.abs x <=/ Rat.abs y) l m 
       
   915 
       
   916 fun subsume todo dun = case todo of
       
   917  [] => dun
       
   918 |v::ovs => 
       
   919    let val dun' = if exists (fn w => subsumes w v) dun then dun
       
   920                   else v::(filter (fn w => not(subsumes v w)) dun) 
       
   921    in subsume ovs dun' 
       
   922    end;
       
   923 
       
   924 fun match_mp PQ P = P RS PQ;
       
   925 
       
   926 fun cterm_of_rat x = 
       
   927 let val (a, b) = Rat.quotient_of_rat x
       
   928 in 
       
   929  if b = 1 then Numeral.mk_cnumber @{ctyp "real"} a
       
   930   else Thm.capply (Thm.capply @{cterm "op / :: real => _"} 
       
   931                    (Numeral.mk_cnumber @{ctyp "real"} a))
       
   932         (Numeral.mk_cnumber @{ctyp "real"} b)
       
   933 end;
       
   934 
       
   935 fun norm_cmul_rule c th = instantiate' [] [SOME (cterm_of_rat c)] (th RS @{thm norm_cmul_rule_thm});
       
   936 
       
   937 fun norm_add_rule th1 th2 = [th1, th2] MRS @{thm norm_add_rule_thm};
       
   938 
       
   939   (* I think here the static context should be sufficient!! *)
       
   940 fun inequality_canon_rule ctxt = 
       
   941  let 
       
   942   (* FIXME : Should be computed statically!! *)
       
   943   val real_poly_conv = 
       
   944     Normalizer.semiring_normalize_wrapper ctxt
       
   945      (valOf (NormalizerData.match ctxt @{cterm "(0::real) + 1"}))
       
   946  in fconv_rule (arg_conv ((rewr_conv @{thm ge_iff_diff_ge_0}) then_conv arg_conv (field_comp_conv then_conv real_poly_conv)))
       
   947 end;
       
   948 
       
   949  fun absc cv ct = case term_of ct of 
       
   950  Abs (v,_, _) => 
       
   951   let val (x,t) = Thm.dest_abs (SOME v) ct
       
   952   in Thm.abstract_rule ((fst o dest_Free o term_of) x) x (cv t)
       
   953   end
       
   954  | _ => all_conv ct;
       
   955 
       
   956 fun sub_conv cv ct = (comb_conv cv else_conv absc cv) ct;
       
   957 fun botc1 conv ct = 
       
   958   ((sub_conv (botc1 conv)) then_conv (conv else_conv all_conv)) ct;
       
   959 
       
   960  fun rewrs_conv eqs ct = first_conv (map rewr_conv eqs) ct;
       
   961  val apply_pth1 = rewr_conv @{thm pth_1};
       
   962  val apply_pth2 = rewr_conv @{thm pth_2};
       
   963  val apply_pth3 = rewr_conv @{thm pth_3};
       
   964  val apply_pth4 = rewrs_conv @{thms pth_4};
       
   965  val apply_pth5 = rewr_conv @{thm pth_5};
       
   966  val apply_pth6 = rewr_conv @{thm pth_6};
       
   967  val apply_pth7 = rewrs_conv @{thms pth_7};
       
   968  val apply_pth8 = rewr_conv @{thm pth_8} then_conv arg1_conv field_comp_conv then_conv (try_conv (rewr_conv (mk_meta_eq @{thm vector_smult_lzero})));
       
   969  val apply_pth9 = rewrs_conv @{thms pth_9} then_conv arg1_conv (arg1_conv field_comp_conv);
       
   970  val apply_ptha = rewr_conv @{thm pth_a};
       
   971  val apply_pthb = rewrs_conv @{thms pth_b};
       
   972  val apply_pthc = rewrs_conv @{thms pth_c};
       
   973  val apply_pthd = try_conv (rewr_conv @{thm pth_d});
       
   974 
       
   975 fun headvector t = case t of 
       
   976   Const(@{const_name plus}, Type("fun",[Type("Finite_Cartesian_Product.^",_),_]))$
       
   977    (Const(@{const_name vector_scalar_mult}, _)$l$v)$r => v
       
   978  | Const(@{const_name vector_scalar_mult}, _)$l$v => v
       
   979  | _ => error "headvector: non-canonical term"
       
   980 
       
   981 fun vector_cmul_conv ct =
       
   982    ((apply_pth5 then_conv arg1_conv field_comp_conv) else_conv
       
   983     (apply_pth6 then_conv binop_conv vector_cmul_conv)) ct
       
   984 
       
   985 fun vector_add_conv ct = apply_pth7 ct 
       
   986  handle CTERM _ => 
       
   987   (apply_pth8 ct 
       
   988    handle CTERM _ => 
       
   989     (case term_of ct of 
       
   990      Const(@{const_name plus},_)$lt$rt =>
       
   991       let 
       
   992        val l = headvector lt 
       
   993        val r = headvector rt
       
   994       in (case TermOrd.fast_term_ord (l,r) of
       
   995          LESS => (apply_pthb then_conv arg_conv vector_add_conv 
       
   996                   then_conv apply_pthd) ct
       
   997         | GREATER => (apply_pthc then_conv arg_conv vector_add_conv 
       
   998                      then_conv apply_pthd) ct 
       
   999         | EQUAL => (apply_pth9 then_conv 
       
  1000                 ((apply_ptha then_conv vector_add_conv) else_conv 
       
  1001               arg_conv vector_add_conv then_conv apply_pthd)) ct)
       
  1002       end
       
  1003      | _ => reflexive ct))
       
  1004 
       
  1005 fun vector_canon_conv ct = case term_of ct of
       
  1006  Const(@{const_name plus},_)$_$_ =>
       
  1007   let 
       
  1008    val ((p,l),r) = Thm.dest_comb ct |>> Thm.dest_comb
       
  1009    val lth = vector_canon_conv l 
       
  1010    val rth = vector_canon_conv r
       
  1011    val th = Drule.binop_cong_rule p lth rth 
       
  1012   in fconv_rule (arg_conv vector_add_conv) th end
       
  1013 
       
  1014 | Const(@{const_name vector_scalar_mult}, _)$_$_ =>
       
  1015   let 
       
  1016    val (p,r) = Thm.dest_comb ct
       
  1017    val rth = Drule.arg_cong_rule p (vector_canon_conv r) 
       
  1018   in fconv_rule (arg_conv (apply_pth4 else_conv vector_cmul_conv)) rth
       
  1019   end
       
  1020 
       
  1021 | Const(@{const_name minus},_)$_$_ => (apply_pth2 then_conv vector_canon_conv) ct
       
  1022 
       
  1023 | Const(@{const_name uminus},_)$_ => (apply_pth3 then_conv vector_canon_conv) ct
       
  1024 
       
  1025 | Const(@{const_name vec},_)$n => 
       
  1026   let val n = Thm.dest_arg ct
       
  1027   in if is_ratconst n andalso not (dest_ratconst n =/ Rat.zero) 
       
  1028      then reflexive ct else apply_pth1 ct
       
  1029   end
       
  1030 
       
  1031 | _ => apply_pth1 ct
       
  1032 
       
  1033 fun norm_canon_conv ct = case term_of ct of
       
  1034   Const(@{const_name norm},_)$_ => arg_conv vector_canon_conv ct
       
  1035  | _ => raise CTERM ("norm_canon_conv", [ct])
       
  1036 
       
  1037 fun fold_rev2 f [] [] z = z
       
  1038  | fold_rev2 f (x::xs) (y::ys) z = f x y (fold_rev2 f xs ys z)
       
  1039  | fold_rev2 f _ _ _ = raise UnequalLengths;
       
  1040 
       
  1041 fun int_flip v eq = 
       
  1042   if Intfunc.defined eq v 
       
  1043   then Intfunc.update (v, Rat.neg (Intfunc.apply eq v)) eq else eq;
       
  1044 
       
  1045 local
       
  1046  val pth_zero = @{thm "Vectors.norm_0"}
       
  1047  val tv_n = (hd o tl o dest_ctyp o ctyp_of_term o dest_arg o dest_arg1 o dest_arg o cprop_of)
       
  1048              pth_zero
       
  1049  val concl = dest_arg o cprop_of 
       
  1050  fun real_vector_combo_prover ctxt translator (nubs,ges,gts) = 
       
  1051   let 
       
  1052    (* FIXME: Should be computed statically!!*)
       
  1053    val real_poly_conv = 
       
  1054       Normalizer.semiring_normalize_wrapper ctxt
       
  1055        (valOf (NormalizerData.match ctxt @{cterm "(0::real) + 1"}))
       
  1056    val sources = map (dest_arg o dest_arg1 o concl) nubs
       
  1057    val rawdests = fold_rev (find_normedterms o dest_arg o concl) (ges @ gts) [] 
       
  1058    val _ = if not (forall fst rawdests) then error "real_vector_combo_prover: Sanity check" 
       
  1059            else ()
       
  1060    val dests = distinct (op aconvc) (map snd rawdests)
       
  1061    val srcfuns = map vector_lincomb sources
       
  1062    val destfuns = map vector_lincomb dests 
       
  1063    val vvs = fold_rev (curry (gen_union op aconvc) o Ctermfunc.dom) (srcfuns @ destfuns) []
       
  1064    val n = length srcfuns
       
  1065    val nvs = 1 upto n
       
  1066    val srccombs = srcfuns ~~ nvs
       
  1067    fun consider d =
       
  1068     let 
       
  1069      fun coefficients x =
       
  1070       let 
       
  1071        val inp = if Ctermfunc.defined d x then Intfunc.onefunc (0, Rat.neg(Ctermfunc.apply d x))
       
  1072                       else Intfunc.undefined 
       
  1073       in fold_rev (fn (f,v) => fn g => if Ctermfunc.defined f x then Intfunc.update (v, Ctermfunc.apply f x) g else g) srccombs inp 
       
  1074       end
       
  1075      val equations = map coefficients vvs
       
  1076      val inequalities = map (fn n => Intfunc.onefunc (n,Rat.one)) nvs
       
  1077      fun plausiblevertices f =
       
  1078       let 
       
  1079        val flippedequations = map (fold_rev int_flip f) equations
       
  1080        val constraints = flippedequations @ inequalities
       
  1081        val rawverts = vertices nvs constraints
       
  1082        fun check_solution v =
       
  1083         let 
       
  1084           val f = fold_rev2 (curry Intfunc.update) nvs v (Intfunc.onefunc (0, Rat.one))
       
  1085         in forall (fn e => evaluate f e =/ Rat.zero) flippedequations
       
  1086         end
       
  1087        val goodverts = filter check_solution rawverts
       
  1088        val signfixups = map (fn n => if n mem_int  f then ~1 else 1) nvs 
       
  1089       in map (map2 (fn s => fn c => Rat.rat_of_int s */ c) signfixups) goodverts
       
  1090       end
       
  1091      val allverts = fold_rev append (map plausiblevertices (allsubsets nvs)) [] 
       
  1092     in subsume allverts []
       
  1093     end
       
  1094    fun compute_ineq v =
       
  1095     let 
       
  1096      val ths = map_filter (fn (v,t) => if v =/ Rat.zero then NONE 
       
  1097                                      else SOME(norm_cmul_rule v t))
       
  1098                             (v ~~ nubs) 
       
  1099     in inequality_canon_rule ctxt (end_itlist norm_add_rule ths)
       
  1100     end
       
  1101    val ges' = map_filter (try compute_ineq) (fold_rev (append o consider) destfuns []) @
       
  1102                  map (inequality_canon_rule ctxt) nubs @ ges
       
  1103    val zerodests = filter
       
  1104         (fn t => null (Ctermfunc.dom (vector_lincomb t))) (map snd rawdests)
       
  1105 
       
  1106   in RealArith.real_linear_prover translator
       
  1107         (map (fn t => instantiate ([(tv_n,(hd o tl o dest_ctyp o ctyp_of_term) t)],[]) pth_zero)
       
  1108             zerodests,
       
  1109         map (fconv_rule (once_depth_conv (norm_canon_conv) then_conv
       
  1110                        arg_conv (arg_conv real_poly_conv))) ges',
       
  1111         map (fconv_rule (once_depth_conv (norm_canon_conv) then_conv 
       
  1112                        arg_conv (arg_conv real_poly_conv))) gts)
       
  1113   end
       
  1114 in val real_vector_combo_prover = real_vector_combo_prover
       
  1115 end;
       
  1116 
       
  1117 local
       
  1118  val pth = @{thm norm_imp_pos_and_ge}
       
  1119  val norm_mp = match_mp pth
       
  1120  val concl = dest_arg o cprop_of
       
  1121  fun conjunct1 th = th RS @{thm conjunct1}
       
  1122  fun conjunct2 th = th RS @{thm conjunct2}
       
  1123  fun C f x y = f y x
       
  1124 fun real_vector_ineq_prover ctxt translator (ges,gts) = 
       
  1125  let 
       
  1126 (*   val _ = error "real_vector_ineq_prover: pause" *)
       
  1127   val ntms = fold_rev find_normedterms (map (dest_arg o concl) (ges @ gts)) []
       
  1128   val lctab = vector_lincombs (map snd (filter (not o fst) ntms))
       
  1129   val (fxns, ctxt') = Variable.variant_fixes (replicate (length lctab) "x") ctxt
       
  1130   fun mk_norm t = capply (instantiate_cterm' [SOME (ctyp_of_term t)] [] @{cpat "norm :: (?'a :: norm) => real"}) t
       
  1131   fun mk_equals l r = capply (capply (instantiate_cterm' [SOME (ctyp_of_term l)] [] @{cpat "op == :: ?'a =>_"}) l) r
       
  1132   val asl = map2 (fn (t,_) => fn n => assume (mk_equals (mk_norm t) (cterm_of (ProofContext.theory_of ctxt') (Free(n,@{typ real}))))) lctab fxns
       
  1133   val replace_conv = try_conv (rewrs_conv asl)
       
  1134   val replace_rule = fconv_rule (funpow 2 arg_conv (replacenegnorms replace_conv))
       
  1135   val ges' =
       
  1136        fold_rev (fn th => fn ths => conjunct1(norm_mp th)::ths)
       
  1137               asl (map replace_rule ges)
       
  1138   val gts' = map replace_rule gts
       
  1139   val nubs = map (conjunct2 o norm_mp) asl
       
  1140   val th1 = real_vector_combo_prover ctxt' translator (nubs,ges',gts')
       
  1141   val shs = filter (member (fn (t,th) => t aconvc cprop_of th) asl) (#hyps (crep_thm th1)) 
       
  1142   val th11 = hd (Variable.export ctxt' ctxt [fold implies_intr shs th1])
       
  1143   val cps = map (swap o dest_equals) (cprems_of th11)
       
  1144   val th12 = instantiate ([], cps) th11
       
  1145   val th13 = fold (C implies_elim) (map (reflexive o snd) cps) th12;
       
  1146  in hd (Variable.export ctxt' ctxt [th13])
       
  1147  end 
       
  1148 in val real_vector_ineq_prover = real_vector_ineq_prover
       
  1149 end;
       
  1150 
       
  1151 local
       
  1152  val rawrule = fconv_rule (arg_conv (rewr_conv @{thm real_eq_0_iff_le_ge_0}))
       
  1153  fun conj_pair th = (th RS @{thm conjunct1}, th RS @{thm conjunct2})
       
  1154  fun simple_cterm_ord t u = TermOrd.term_ord (term_of t, term_of u) = LESS;
       
  1155   (* FIXME: Lookup in the context every time!!! Fix this !!!*)
       
  1156  fun splitequation ctxt th acc =
       
  1157   let 
       
  1158    val real_poly_neg_conv = #neg
       
  1159        (Normalizer.semiring_normalizers_ord_wrapper ctxt
       
  1160         (valOf (NormalizerData.match ctxt @{cterm "(0::real) + 1"})) simple_cterm_ord)
       
  1161    val (th1,th2) = conj_pair(rawrule th)
       
  1162   in th1::fconv_rule (arg_conv (arg_conv real_poly_neg_conv)) th2::acc
       
  1163   end
       
  1164 in fun real_vector_prover ctxt translator (eqs,ges,gts) =
       
  1165      real_vector_ineq_prover ctxt translator
       
  1166          (fold_rev (splitequation ctxt) eqs ges,gts)
       
  1167 end;
       
  1168 
       
  1169   fun init_conv ctxt = 
       
  1170    Simplifier.rewrite (Simplifier.context ctxt 
       
  1171      (HOL_basic_ss addsimps ([@{thm vec_0}, @{thm vec_1}, @{thm dist_def}, @{thm diff_0_right}, @{thm right_minus}, @{thm diff_self}, @{thm norm_0}] @ @{thms arithmetic_simps} @ @{thms norm_pths})))
       
  1172    then_conv field_comp_conv 
       
  1173    then_conv nnf_conv
       
  1174 
       
  1175  fun pure ctxt = RealArith.gen_prover_real_arith ctxt (real_vector_prover ctxt);
       
  1176  fun norm_arith ctxt ct = 
       
  1177   let 
       
  1178    val ctxt' = Variable.declare_term (term_of ct) ctxt
       
  1179    val th = init_conv ctxt' ct
       
  1180   in equal_elim (Drule.arg_cong_rule @{cterm Trueprop} (symmetric th)) 
       
  1181                 (pure ctxt' (rhs_of th))
       
  1182  end
       
  1183 
       
  1184  fun norm_arith_tac ctxt = 
       
  1185    clarify_tac HOL_cs THEN'
       
  1186    ObjectLogic.full_atomize_tac THEN'
       
  1187    CSUBGOAL ( fn (p,i) => rtac (norm_arith ctxt (Thm.dest_arg p )) i);
       
  1188 
       
  1189 end;