# HG changeset patch # User wenzelm # Date 1631472759 -7200 # Node ID 1466f8a2f6dd3d271521e18b213b2a85553bde48 # Parent f7ee629b9beb0e61e9b02f99cea6c23de81d00fb more antiquotations; diff -r f7ee629b9beb -r 1466f8a2f6dd src/HOL/Hoare/hoare_tac.ML --- a/src/HOL/Hoare/hoare_tac.ML Sun Sep 12 20:40:18 2021 +0200 +++ b/src/HOL/Hoare/hoare_tac.ML Sun Sep 12 20:52:39 2021 +0200 @@ -28,17 +28,17 @@ local (** maps (%x1 ... xn. t) to [x1,...,xn] **) -fun abs2list (Const (\<^const_name>\case_prod\, _) $ Abs (x, T, t)) = Free (x, T) :: abs2list t +fun abs2list \<^Const_>\case_prod _ _ _ for \Abs (x, T, t)\\ = Free (x, T) :: abs2list t | abs2list (Abs (x, T, _)) = [Free (x, T)] | abs2list _ = []; (** maps {(x1,...,xn). t} to [x1,...,xn] **) -fun mk_vars (Const (\<^const_name>\Collect\,_) $ T) = abs2list T +fun mk_vars \<^Const_>\Collect _ for T\ = abs2list T | mk_vars _ = []; (** abstraction of body over a tuple formed from a list of free variables. Types are also built **) -fun mk_abstupleC [] body = absfree ("x", HOLogic.unitT) body +fun mk_abstupleC [] body = absfree ("x", \<^Type>\unit\) body | mk_abstupleC [v] body = absfree (dest_Free v) body | mk_abstupleC (v :: w) body = let @@ -49,13 +49,11 @@ Abs (_, T, _) => T | Const (_, Type (_, [_, Type (_, [T, _])])) $ _ => T); in - Const (\<^const_name>\case_prod\, - (T --> T2 --> HOLogic.boolT) --> HOLogic.mk_prodT (T, T2) --> HOLogic.boolT) $ - absfree (x, T) z + \<^Const>\case_prod T T2 \\<^Type>\bool\\ for \absfree (x, T) z\\ end; (** maps [x1,...,xn] to (x1,...,xn) and types**) -fun mk_bodyC [] = HOLogic.unit +fun mk_bodyC [] = \<^Const>\Unity\ | mk_bodyC [x] = x | mk_bodyC (x :: xs) = let @@ -64,8 +62,8 @@ val T2 = (case z of Free (_, T) => T - | Const (\<^const_name>\Pair\, Type ("fun", [_, Type ("fun", [_, T])])) $ _ $ _ => T); - in Const (\<^const_name>\Pair\, [T, T2] ---> HOLogic.mk_prodT (T, T2)) $ x $ z end; + | \<^Const_>\Pair A B for _ _\ => \<^Type>\prod A B\); + in \<^Const>\Pair T T2 for x z\ end; (** maps a subgoal of the form: VARS x1 ... xn {._.} _ {._.} or to [x1,...,xn] @@ -77,10 +75,10 @@ in mk_vars pre end; fun mk_CollectC tm = - let val Type ("fun",[t,_]) = fastype_of tm; - in HOLogic.Collect_const t $ tm end; + let val \<^Type>\fun t _\ = fastype_of tm; + in \<^Const>\Collect t for tm\ end; -fun inclt ty = Const (\<^const_name>\Orderings.less_eq\, [ty,ty] ---> HOLogic.boolT); +fun inclt ty = \<^Const>\less_eq ty\; in @@ -91,9 +89,9 @@ val vars = get_vars prop; val varsT = fastype_of (mk_bodyC vars); val big_Collect = - mk_CollectC (mk_abstupleC vars (Free (P, varsT --> HOLogic.boolT) $ mk_bodyC vars)); + mk_CollectC (mk_abstupleC vars (Free (P, varsT --> \<^Type>\bool\) $ mk_bodyC vars)); val small_Collect = - mk_CollectC (Abs ("x", varsT, Free (P, varsT --> HOLogic.boolT) $ Bound 0)); + mk_CollectC (Abs ("x", varsT, Free (P, varsT --> \<^Type>\bool\) $ Bound 0)); val MsetT = fastype_of big_Collect; fun Mset_incl t = HOLogic.mk_Trueprop (inclt MsetT $ Free (Mset, MsetT) $ t);