# HG changeset patch # User wenzelm # Date 1635191198 -7200 # Node ID 0eb2f18b18068417932da679cdabaaf0b5ae6a6b # Parent 882de99c7c833787d04a8d113fa578b04bf27f08 more antiquotations; diff -r 882de99c7c83 -r 0eb2f18b1806 src/HOL/Library/Sum_of_Squares/positivstellensatz.ML --- a/src/HOL/Library/Sum_of_Squares/positivstellensatz.ML Mon Oct 25 21:31:35 2021 +0200 +++ b/src/HOL/Library/Sum_of_Squares/positivstellensatz.ML Mon Oct 25 21:46:38 2021 +0200 @@ -285,14 +285,16 @@ val (a, b) = Rat.dest x in if b = 1 then Numeral.mk_cnumber \<^ctyp>\real\ a - else Thm.apply (Thm.apply \<^cterm>\(/) :: real \ _\ - (Numeral.mk_cnumber \<^ctyp>\real\ a)) - (Numeral.mk_cnumber \<^ctyp>\real\ b) + else + \<^instantiate>\ + a = \Numeral.mk_cnumber \<^ctyp>\real\ a\ and + b = \Numeral.mk_cnumber \<^ctyp>\real\ b\ + in cterm \a / b\ for a b :: real\ end; fun dest_ratconst t = case Thm.term_of t of - Const(\<^const_name>\divide\, _)$a$b => Rat.make(HOLogic.dest_number a |> snd, HOLogic.dest_number b |> snd) + \<^Const_>\divide _ for a b\ => Rat.make(HOLogic.dest_number a |> snd, HOLogic.dest_number b |> snd) | _ => Rat.of_int (HOLogic.dest_number (Thm.term_of t) |> snd) fun is_ratconst t = can dest_ratconst t @@ -319,8 +321,9 @@ (* Map back polynomials to HOL. *) -fun cterm_of_varpow x k = if k = 1 then x else Thm.apply (Thm.apply \<^cterm>\(^) :: real \ _\ x) - (Numeral.mk_cnumber \<^ctyp>\nat\ k) +fun cterm_of_varpow x k = + if k = 1 then x + else \<^instantiate>\x and k = \Numeral.mk_cnumber \<^ctyp>\nat\ k\ in cterm "x ^ k" for x :: real\ fun cterm_of_monomial m = if FuncUtil.Ctermfunc.is_empty m then \<^cterm>\1::real\ @@ -328,13 +331,13 @@ let val m' = FuncUtil.dest_monomial m val vps = fold_rev (fn (x,k) => cons (cterm_of_varpow x k)) m' [] - in foldr1 (fn (s, t) => Thm.apply (Thm.apply \<^cterm>\(*) :: real \ _\ s) t) vps + in foldr1 (fn (s, t) => \<^instantiate>\s and t in cterm "s * t" for s t :: real\) vps end fun cterm_of_cmonomial (m,c) = if FuncUtil.Ctermfunc.is_empty m then cterm_of_rat c else if c = @1 then cterm_of_monomial m - else Thm.apply (Thm.apply \<^cterm>\(*)::real \ _\ (cterm_of_rat c)) (cterm_of_monomial m); + else \<^instantiate>\x = \cterm_of_rat c\ and y = \cterm_of_monomial m\ in cterm "x * y" for x y :: real\; fun cterm_of_poly p = if FuncUtil.Monomialfunc.is_empty p then \<^cterm>\0::real\ @@ -342,7 +345,7 @@ let val cms = map cterm_of_cmonomial (sort (prod_ord FuncUtil.monomial_order (K EQUAL)) (FuncUtil.Monomialfunc.dest p)) - in foldr1 (fn (t1, t2) => Thm.apply(Thm.apply \<^cterm>\(+) :: real \ _\ t1) t2) cms + in foldr1 (fn (t1, t2) => \<^instantiate>\t1 and t2 in cterm "t1 + t2" for t1 t2 :: real\) cms end; (* A general real arithmetic prover *) @@ -400,15 +403,15 @@ Axiom_eq n => nth eqs n | Axiom_le n => nth les n | Axiom_lt n => nth lts n - | Rational_eq x => eqT_elim(numeric_eq_conv(Thm.apply \<^cterm>\Trueprop\ - (Thm.apply (Thm.apply \<^cterm>\(=)::real \ _\ (mk_numeric x)) - \<^cterm>\0::real\))) - | Rational_le x => eqT_elim(numeric_ge_conv(Thm.apply \<^cterm>\Trueprop\ - (Thm.apply (Thm.apply \<^cterm>\(\)::real \ _\ - \<^cterm>\0::real\) (mk_numeric x)))) - | Rational_lt x => eqT_elim(numeric_gt_conv(Thm.apply \<^cterm>\Trueprop\ - (Thm.apply (Thm.apply \<^cterm>\(<)::real \ _\ \<^cterm>\0::real\) - (mk_numeric x)))) + | Rational_eq x => + eqT_elim (numeric_eq_conv + \<^instantiate>\x = \mk_numeric x\ in cprop "x = 0" for x :: real\) + | Rational_le x => + eqT_elim (numeric_ge_conv + \<^instantiate>\x = \mk_numeric x\ in cprop "x \ 0" for x :: real\) + | Rational_lt x => + eqT_elim (numeric_gt_conv + \<^instantiate>\x = \mk_numeric x\ in cprop "x > 0" for x :: real\) | Square pt => square_rule (cterm_of_poly pt) | Eqmul(pt,p) => emul_rule (cterm_of_poly pt) (translate p) | Sum(p1,p2) => add_rule (translate p1) (translate p2) @@ -438,8 +441,8 @@ else error "disj_cases : conclusions not alpha convertible" in Thm.implies_elim (Thm.implies_elim (Thm.implies_elim (Thm.instantiate' [] (map SOME [p,q,c]) @{thm disjE}) th) - (Thm.implies_intr (Thm.apply \<^cterm>\Trueprop\ p) th1)) - (Thm.implies_intr (Thm.apply \<^cterm>\Trueprop\ q) th2) + (Thm.implies_intr \<^instantiate>\p in cprop p\ th1)) + (Thm.implies_intr \<^instantiate>\q in cprop q\ th2) end fun overall cert_choice dun ths = case ths of @@ -481,8 +484,8 @@ val th_x = Drule.arg_cong_rule \<^cterm>\uminus :: real \ _\ th_p val th_n = fconv_rule (arg_conv poly_neg_conv) th_x val th' = Drule.binop_cong_rule \<^cterm>\HOL.disj\ - (Drule.arg_cong_rule (Thm.apply \<^cterm>\(<)::real \ _\ \<^cterm>\0::real\) th_p) - (Drule.arg_cong_rule (Thm.apply \<^cterm>\(<)::real \ _\ \<^cterm>\0::real\) th_n) + (Drule.arg_cong_rule \<^cterm>\(<) (0::real)\ th_p) + (Drule.arg_cong_rule \<^cterm>\(<) (0::real)\ th_n) in Thm.transitive th th' end fun equal_implies_1_rule PQ = @@ -495,7 +498,7 @@ let fun h (acc, t) = case Thm.term_of t of - Const(\<^const_name>\Ex\,_)$Abs(_,_,_) => + \<^Const_>\Ex _ for \Abs _\\ => h (Thm.dest_abs_global (Thm.dest_arg t) |>> (fn v => v::acc)) | _ => (acc,t) in fn t => h ([],t) @@ -508,39 +511,38 @@ fun mk_forall x th = let - val T = Thm.typ_of_cterm x - val all = Thm.cterm_of ctxt (Const (\<^const_name>\All\, (T --> \<^typ>\bool\) --> \<^typ>\bool\)) + val T = Thm.ctyp_of_cterm x + val all = \<^instantiate>\'a = T in cterm All\ in Drule.arg_cong_rule all (Thm.abstract_rule (name_of x) x th) end val specl = fold_rev (fn x => fn th => Thm.instantiate' [] [SOME x] (th RS spec)); - fun ext T = Thm.cterm_of ctxt (Const (\<^const_name>\Ex\, (T --> \<^typ>\bool\) --> \<^typ>\bool\)) - fun mk_ex v t = Thm.apply (ext (Thm.typ_of_cterm v)) (Thm.lambda v t) + fun mk_ex v t = + \<^instantiate>\'a = \Thm.ctyp_of_cterm v\ and P = \Thm.lambda v t\ + in cprop "Ex P" for P :: "'a \ bool"\ fun choose v th th' = case Thm.concl_of th of - \<^term>\Trueprop\ $ (Const(\<^const_name>\Ex\,_)$_) => - let - val p = (funpow 2 Thm.dest_arg o Thm.cprop_of) th - val T = Thm.dest_ctyp0 (Thm.ctyp_of_cterm p) - val th0 = fconv_rule (Thm.beta_conversion true) - (Thm.instantiate' [SOME T] [SOME p, (SOME o Thm.dest_arg o Thm.cprop_of) th'] exE) - val pv = (Thm.rhs_of o Thm.beta_conversion true) - (Thm.apply \<^cterm>\Trueprop\ (Thm.apply p v)) - val th1 = Thm.forall_intr v (Thm.implies_intr pv th') - in Thm.implies_elim (Thm.implies_elim th0 th) th1 end + \<^Const_>\Trueprop for \<^Const_>\Ex _ for _\\ => + let + val p = (funpow 2 Thm.dest_arg o Thm.cprop_of) th + val T = Thm.dest_ctyp0 (Thm.ctyp_of_cterm p) + val th0 = fconv_rule (Thm.beta_conversion true) + (Thm.instantiate' [SOME T] [SOME p, (SOME o Thm.dest_arg o Thm.cprop_of) th'] exE) + val pv = (Thm.rhs_of o Thm.beta_conversion true) + (Thm.apply \<^cterm>\Trueprop\ (Thm.apply p v)) + val th1 = Thm.forall_intr v (Thm.implies_intr pv th') + in Thm.implies_elim (Thm.implies_elim th0 th) th1 end | _ => raise THM ("choose",0,[th, th']) fun simple_choose v th = - choose v - (Thm.assume - ((Thm.apply \<^cterm>\Trueprop\ o mk_ex v) (Thm.dest_arg (hd (Thm.chyps_of th))))) th + choose v (Thm.assume (mk_ex v (Thm.dest_arg (hd (Thm.chyps_of th))))) th val strip_forall = let fun h (acc, t) = case Thm.term_of t of - Const(\<^const_name>\All\,_)$Abs(_,_,_) => + \<^Const_>\All _ for \Abs _\\ => h (Thm.dest_abs_global (Thm.dest_arg t) |>> (fn v => v::acc)) | _ => (acc,t) in fn t => h ([],t) @@ -558,7 +560,7 @@ fun absremover ct = (literals_conv [\<^term>\HOL.conj\, \<^term>\HOL.disj\] [] (try_conv (absconv1 then_conv binop_conv (arg_conv poly_conv))) then_conv try_conv (absconv2 then_conv nnf_norm_conv' then_conv binop_conv absremover)) ct - val nct = Thm.apply \<^cterm>\Trueprop\ (Thm.apply \<^cterm>\Not\ ct) + val nct = \<^instantiate>\ct in cprop "\ ct"\ val th0 = (init_conv then_conv arg_conv nnf_norm_conv') nct val tm0 = Thm.dest_arg (Thm.rhs_of th0) val (th, certificates) = @@ -688,8 +690,8 @@ fun is_alien ct = case Thm.term_of ct of - Const(\<^const_name>\of_nat\, _)$ n => not (can HOLogic.dest_number n) - | Const(\<^const_name>\of_int\, _)$ n => not (can HOLogic.dest_number n) + \<^Const_>\of_nat _ for n\ => not (can HOLogic.dest_number n) + | \<^Const_>\of_int _ for n\ => not (can HOLogic.dest_number n) | _ => false in fun real_linear_prover translator (eq,le,lt) =