# HG changeset patch # User wenzelm # Date 1632860084 -7200 # Node ID 107941e8fa0140ac95aebeb3c6dc2a12802de30e # Parent 8d0294d877bd47b1890cb3acf1507b216ac24b77 clarified antiquotations; diff -r 8d0294d877bd -r 107941e8fa01 src/HOL/HOL.thy --- a/src/HOL/HOL.thy Tue Sep 28 22:14:02 2021 +0200 +++ b/src/HOL/HOL.thy Tue Sep 28 22:14:44 2021 +0200 @@ -888,7 +888,7 @@ structure Blast = Blast ( structure Classical = Classical - val Trueprop_const = dest_Const \<^const>\Trueprop\ + val Trueprop_const = dest_Const \<^Const>\Trueprop\ val equality_name = \<^const_name>\HOL.eq\ val not_name = \<^const_name>\Not\ val notE = @{thm notE} @@ -1549,7 +1549,7 @@ {lhss = [\<^term>\induct_false \ PROP P \ PROP Q\], proc = fn _ => fn _ => fn ct => (case Thm.term_of ct of - _ $ (P as _ $ \<^const>\induct_false\) $ (_ $ Q $ _) => + _ $ (P as _ $ \<^Const_>\induct_false\) $ (_ $ Q $ _) => if P <> Q then SOME Drule.swap_prems_eq else NONE | _ => NONE)}, Simplifier.make_simproc \<^context> "induct_equal_conj_curry" @@ -1558,11 +1558,11 @@ (case Thm.term_of ct of _ $ (_ $ P) $ _ => let - fun is_conj (\<^const>\induct_conj\ $ P $ Q) = + fun is_conj \<^Const_>\induct_conj for P Q\ = is_conj P andalso is_conj Q - | is_conj (Const (\<^const_name>\induct_equal\, _) $ _ $ _) = true - | is_conj \<^const>\induct_true\ = true - | is_conj \<^const>\induct_false\ = true + | is_conj \<^Const_>\induct_equal _ for _ _\ = true + | is_conj \<^Const_>\induct_true\ = true + | is_conj \<^Const_>\induct_false\ = true | is_conj _ = false in if is_conj P then SOME @{thm induct_conj_curry} else NONE end | _ => NONE)}] diff -r 8d0294d877bd -r 107941e8fa01 src/HOL/Library/Pattern_Aliases.thy --- a/src/HOL/Library/Pattern_Aliases.thy Tue Sep 28 22:14:02 2021 +0200 +++ b/src/HOL/Library/Pattern_Aliases.thy Tue Sep 28 22:14:44 2021 +0200 @@ -97,7 +97,7 @@ fun check_pattern_syntax t = case strip_all t of - (vars, \<^const>\Trueprop\ $ (Const (\<^const_name>\HOL.eq\, _) $ lhs $ rhs)) => + (vars, \<^Const>\Trueprop\ $ (Const (\<^const_name>\HOL.eq\, _) $ lhs $ rhs)) => let fun go (Const (\<^const_name>\as\, _) $ pat $ var, rhs) = let @@ -126,7 +126,7 @@ fun uncheck_pattern_syntax ctxt t = case strip_all t of - (vars, \<^const>\Trueprop\ $ (Const (\<^const_name>\HOL.eq\, _) $ lhs $ rhs)) => + (vars, \<^Const>\Trueprop\ $ (Const (\<^const_name>\HOL.eq\, _) $ lhs $ rhs)) => let (* restricted to going down abstractions; ignores eta-contracted rhs *) fun go lhs (rhs as Const (\<^const_name>\Let\, _) $ pat $ Abs (name, typ, t)) ctxt frees = diff -r 8d0294d877bd -r 107941e8fa01 src/HOL/Library/code_lazy.ML --- a/src/HOL/Library/code_lazy.ML Tue Sep 28 22:14:02 2021 +0200 +++ b/src/HOL/Library/code_lazy.ML Tue Sep 28 22:14:44 2021 +0200 @@ -378,8 +378,8 @@ ||>> fold_map (mk_name ("case_" ^ short_type_name ^ "_lazy_") "") ctr_names val mk_Lazy_delay_eq = - (#const lazy_ctr_def $ mk_delay (Bound 0), Rep_lazy $ (Bound 0 $ \<^const>\Unity\)) - |> mk_eq |> all_abs [\<^typ>\unit\ --> lazy_type] + (#const lazy_ctr_def $ mk_delay (Bound 0), Rep_lazy $ (Bound 0 $ \<^Const>\Unity\)) + |> mk_eq |> all_abs [\<^Type>\unit\ --> lazy_type] val (Lazy_delay_thm, lthy8a) = mk_thm ((Lazy_delay_eq_name, mk_Lazy_delay_eq), [#thm lazy_ctr_def, @{thm force_delay}]) lthy8 diff -r 8d0294d877bd -r 107941e8fa01 src/HOL/Library/code_test.ML --- a/src/HOL/Library/code_test.ML Tue Sep 28 22:14:02 2021 +0200 +++ b/src/HOL/Library/code_test.ML Tue Sep 28 22:14:44 2021 +0200 @@ -179,15 +179,15 @@ lhs :: rhs :: acc) | add_eval _ = I -fun mk_term_of [] = \<^term>\None :: (unit \ yxml_of_term) option\ +fun mk_term_of [] = \<^Const>\None \<^typ>\unit \ yxml_of_term\\ | mk_term_of ts = let val tuple = mk_tuples ts val T = fastype_of tuple in \<^term>\Some :: (unit \ yxml_of_term) \ (unit \ yxml_of_term) option\ $ - (absdummy \<^typ>\unit\ (\<^const>\yxml_string_of_term\ $ - (Const (\<^const_name>\Code_Evaluation.term_of\, T --> \<^typ>\term\) $ tuple))) + (absdummy \<^typ>\unit\ (\<^Const>\yxml_string_of_term\ $ + (Const (\<^const_name>\Code_Evaluation.term_of\, T --> \<^Type>\term\) $ tuple))) end fun test_terms ctxt ts target = diff -r 8d0294d877bd -r 107941e8fa01 src/HOL/Tools/Ctr_Sugar/ctr_sugar.ML --- a/src/HOL/Tools/Ctr_Sugar/ctr_sugar.ML Tue Sep 28 22:14:02 2021 +0200 +++ b/src/HOL/Tools/Ctr_Sugar/ctr_sugar.ML Tue Sep 28 22:14:44 2021 +0200 @@ -303,11 +303,11 @@ fun name_of_disc t = (case head_of t of - Abs (_, _, \<^const>\Not\ $ (t' $ Bound 0)) => + Abs (_, _, \<^Const_>\Not for \t' $ Bound 0\\) => Long_Name.map_base_name (prefix not_prefix) (name_of_disc t') - | Abs (_, _, Const (\<^const_name>\HOL.eq\, _) $ Bound 0 $ t') => + | Abs (_, _, \<^Const_>\HOL.eq _ for \Bound 0\ t'\) => Long_Name.map_base_name (prefix is_prefix) (name_of_disc t') - | Abs (_, _, \<^const>\Not\ $ (Const (\<^const_name>\HOL.eq\, _) $ Bound 0 $ t')) => + | Abs (_, _, \<^Const_>\Not for \<^Const_>\HOL.eq _ for \Bound 0\ t'\\) => Long_Name.map_base_name (prefix (not_prefix ^ is_prefix)) (name_of_disc t') | t' => name_of_const "discriminator" (perhaps (try domain_type)) t'); @@ -1037,7 +1037,7 @@ val disc_eq_case_thms = let - fun const_of_bool b = if b then \<^const>\True\ else \<^const>\False\; + fun const_of_bool b = if b then \<^Const>\True\ else \<^Const>\False\; fun mk_case_args n = map_index (fn (k, argTs) => fold_rev Term.absdummy argTs (const_of_bool (n = k))) ctr_Tss; val goals = map_index (fn (n, udisc) => diff -r 8d0294d877bd -r 107941e8fa01 src/HOL/Tools/Lifting/lifting_def_code_dt.ML --- a/src/HOL/Tools/Lifting/lifting_def_code_dt.ML Tue Sep 28 22:14:02 2021 +0200 +++ b/src/HOL/Tools/Lifting/lifting_def_code_dt.ML Tue Sep 28 22:14:44 2021 +0200 @@ -400,7 +400,7 @@ let val (sel_rhs, wits) = mk_sel_case_args lthy4 ctr_Tss ks arg in (arg |> #2, wits, list_comb (mk_sel_casex arg, sel_rhs)) end; fun mk_dis_case_args args k = map (fn (k', arg) => (if k = k' - then fold_rev Term.lambda arg \<^const>\True\ else fold_rev Term.lambda arg \<^const>\False\)) args; + then fold_rev Term.lambda arg \<^Const>\True\ else fold_rev Term.lambda arg \<^Const>\False\)) args; val sel_rhs = map (map mk_sel_rhs) sel_argss val dis_rhs = map (fn k => list_comb (dis_casex, mk_dis_case_args (ks ~~ xss) k)) ks val dis_qty = qty_isom --> HOLogic.boolT; diff -r 8d0294d877bd -r 107941e8fa01 src/HOL/Tools/Nunchaku/nunchaku.ML --- a/src/HOL/Tools/Nunchaku/nunchaku.ML Tue Sep 28 22:14:02 2021 +0200 +++ b/src/HOL/Tools/Nunchaku/nunchaku.ML Tue Sep 28 22:14:44 2021 +0200 @@ -131,7 +131,7 @@ fun none_true assigns = forall (curry (op <>) (SOME true) o snd) assigns; -fun has_lonely_bool_var (\<^const>\Pure.conjunction\ $ (\<^const>\Trueprop\ $ Free _) $ _) = true +fun has_lonely_bool_var \<^Const_>\Pure.conjunction for \<^Const_>\Trueprop for \Free _\\ _\ = true | has_lonely_bool_var _ = false; val syntactic_sorts = diff -r 8d0294d877bd -r 107941e8fa01 src/HOL/Tools/Nunchaku/nunchaku_collect.ML --- a/src/HOL/Tools/Nunchaku/nunchaku_collect.ML Tue Sep 28 22:14:02 2021 +0200 +++ b/src/HOL/Tools/Nunchaku/nunchaku_collect.ML Tue Sep 28 22:14:44 2021 +0200 @@ -240,7 +240,7 @@ fun mutual_co_datatypes_of ctxt (T_name, Ts) = (if T_name = \<^type_name>\itself\ then - (BNF_Util.Least_FP, [\<^typ>\'a itself\], [[@{const Pure.type ('a)}]]) + (BNF_Util.Least_FP, [\<^typ>\'a itself\], [[\<^Const>\Pure.type \<^typ>\'a\\]]) else let val (fp, ctr_sugars) = @@ -269,7 +269,7 @@ val A = Logic.varifyT_global \<^typ>\'a\; val absT = Type (\<^type_name>\set\, [A]); val repT = A --> HOLogic.boolT; - val pred = Abs (Name.uu, repT, \<^const>\True\); + val pred = Abs (Name.uu, repT, \<^Const>\True\); val abs = Const (\<^const_name>\Collect\, repT --> absT); val rep = Const (\<^const_name>\rmember\, absT --> repT); in @@ -518,8 +518,8 @@ fold process_spec specs NONE end; -fun lhs_of_equation (Const (\<^const_name>\Pure.eq\, _) $ t $ _) = t - | lhs_of_equation (\<^const>\Trueprop\ $ (Const (\<^const_name>\HOL.eq\, _) $ t $ _)) = t; +fun lhs_of_equation \<^Const_>\Pure.eq _ for t _\ = t + | lhs_of_equation \<^Const_>\Trueprop for \<^Const_>\HOL.eq _ for t _\\ = t; fun specialize_definition_type thy x def0 = let @@ -674,10 +674,10 @@ [cmd] :: (group :: groups) end; -fun defined_by (Const (\<^const_name>\All\, _) $ t) = defined_by t +fun defined_by \<^Const_>\All _ for t\ = defined_by t | defined_by (Abs (_, _, t)) = defined_by t - | defined_by (\<^const>\implies\ $ _ $ u) = defined_by u - | defined_by (Const (\<^const_name>\HOL.eq\, _) $ t $ _) = head_of t + | defined_by \<^Const_>\implies for _ u\ = defined_by u + | defined_by \<^Const_>\HOL.eq _ for t _\ = head_of t | defined_by t = head_of t; fun partition_props [_] props = SOME [props] @@ -1002,14 +1002,14 @@ val (poly_axioms, mono_axioms0) = orphan_axioms_of ctxt |> List.partition has_polymorphism; - fun implicit_evals_of pol (\<^const>\Not\ $ t) = implicit_evals_of (not pol) t - | implicit_evals_of pol (\<^const>\implies\ $ t $ u) = + fun implicit_evals_of pol \<^Const_>\Not for t\ = implicit_evals_of (not pol) t + | implicit_evals_of pol \<^Const_>\implies for t u\ = (case implicit_evals_of pol u of [] => implicit_evals_of (not pol) t | ts => ts) - | implicit_evals_of pol (\<^const>\conj\ $ t $ u) = + | implicit_evals_of pol \<^Const_>\conj for t u\ = union (op aconv) (implicit_evals_of pol t) (implicit_evals_of pol u) - | implicit_evals_of pol (\<^const>\disj\ $ t $ u) = + | implicit_evals_of pol \<^Const_>\disj for t u\ = union (op aconv) (implicit_evals_of pol t) (implicit_evals_of pol u) | implicit_evals_of false (Const (\<^const_name>\HOL.eq\, _) $ t $ u) = distinct (op aconv) [t, u] @@ -1049,7 +1049,7 @@ Syntax.string_of_term ctxt t ^ " : " ^ Syntax.string_of_typ ctxt (fastype_of t); fun is_triv_wrt (Abs (_, _, body)) = is_triv_wrt body - | is_triv_wrt \<^const>\True\ = true + | is_triv_wrt \<^Const>\True\ = true | is_triv_wrt _ = false; fun str_of_isa_type_spec ctxt {abs_typ, rep_typ, wrt, abs, rep} = diff -r 8d0294d877bd -r 107941e8fa01 src/HOL/Tools/Nunchaku/nunchaku_reconstruct.ML --- a/src/HOL/Tools/Nunchaku/nunchaku_reconstruct.ML Tue Sep 28 22:14:02 2021 +0200 +++ b/src/HOL/Tools/Nunchaku/nunchaku_reconstruct.ML Tue Sep 28 22:14:44 2021 +0200 @@ -166,11 +166,11 @@ else if id = nun_equals then Const (\<^const_name>\HOL.eq\, typ_of ty) else if id = nun_false then - \<^const>\False\ + \<^Const>\False\ else if id = nun_if then Const (\<^const_name>\If\, typ_of ty) else if id = nun_implies then - \<^term>\implies\ + \<^Const>\implies\ else if id = nun_not then HOLogic.Not else if id = nun_unique then @@ -178,7 +178,7 @@ else if id = nun_unique_unsafe then Const (\<^const_name>\The_unsafe\, typ_of ty) else if id = nun_true then - \<^const>\True\ + \<^Const>\True\ else if String.isPrefix nun_dollar_anon_fun_prefix id then let val j = Int.fromString (unprefix nun_dollar_anon_fun_prefix id) |> the_default ~1 in Var ((anonymousN ^ nat_subscript (j + 1), 0), typ_of ty) @@ -225,12 +225,12 @@ end | term_of _ (NMatch _) = raise Fail "unexpected match"; - fun rewrite_numbers (t as \<^const>\Suc\ $ _) = + fun rewrite_numbers (t as \<^Const>\Suc\ $ _) = (case try HOLogic.dest_nat t of - SOME n => HOLogic.mk_number \<^typ>\nat\ n + SOME n => HOLogic.mk_number \<^Type>\nat\ n | NONE => t) - | rewrite_numbers (\<^const>\Abs_Integ\ $ (@{const Pair (nat, nat)} $ t $ u)) = - HOLogic.mk_number \<^typ>\int\ (HOLogic.dest_nat t - HOLogic.dest_nat u) + | rewrite_numbers \<^Const_>\Abs_Integ for \<^Const>\Pair \<^Type>\nat\ \<^Type>\nat\ for t u\\ = + HOLogic.mk_number \<^Type>\int\ (HOLogic.dest_nat t - HOLogic.dest_nat u) | rewrite_numbers (t $ u) = rewrite_numbers t $ rewrite_numbers u | rewrite_numbers (Abs (s, T, t)) = Abs (s, T, rewrite_numbers t) | rewrite_numbers t = t; diff -r 8d0294d877bd -r 107941e8fa01 src/HOL/Tools/Quickcheck/quickcheck_common.ML --- a/src/HOL/Tools/Quickcheck/quickcheck_common.ML Tue Sep 28 22:14:02 2021 +0200 +++ b/src/HOL/Tools/Quickcheck/quickcheck_common.ML Tue Sep 28 22:14:44 2021 +0200 @@ -466,8 +466,8 @@ fun make_fun_upds T1 T2 (tps, t) = fold_rev (mk_fun_upd T1 T2) tps t fun make_set T1 [] = Const (\<^const_abbrev>\Set.empty\, T1 --> \<^typ>\bool\) - | make_set T1 ((_, \<^const>\False\) :: tps) = make_set T1 tps - | make_set T1 ((t1, \<^const>\True\) :: tps) = + | make_set T1 ((_, \<^Const_>\False\) :: tps) = make_set T1 tps + | make_set T1 ((t1, \<^Const_>\True\) :: tps) = Const (\<^const_name>\insert\, T1 --> (T1 --> \<^typ>\bool\) --> T1 --> \<^typ>\bool\) $ t1 $ (make_set T1 tps) | make_set T1 ((_, t) :: _) = raise TERM ("make_set", [t]) @@ -476,8 +476,8 @@ | make_coset T tps = let val U = T --> \<^typ>\bool\ - fun invert \<^const>\False\ = \<^const>\True\ - | invert \<^const>\True\ = \<^const>\False\ + fun invert \<^Const_>\False\ = \<^Const>\True\ + | invert \<^Const_>\True\ = \<^Const>\False\ in Const (\<^const_name>\Groups.minus_class.minus\, U --> U --> U) $ Const (\<^const_abbrev>\UNIV\, U) $ make_set T (map (apsnd invert) tps) @@ -505,8 +505,8 @@ (case T2 of \<^typ>\bool\ => (case t of - Abs(_, _, \<^const>\False\) => fst #> rev #> make_set T1 - | Abs(_, _, \<^const>\True\) => fst #> rev #> make_coset T1 + Abs(_, _, \<^Const_>\False\) => fst #> rev #> make_set T1 + | Abs(_, _, \<^Const_>\True\) => fst #> rev #> make_coset T1 | Abs(_, _, Const (\<^const_name>\undefined\, _)) => fst #> rev #> make_set T1 | _ => raise TERM ("post_process_term", [t])) | Type (\<^type_name>\option\, _) => diff -r 8d0294d877bd -r 107941e8fa01 src/HOL/Tools/Quickcheck/random_generators.ML --- a/src/HOL/Tools/Quickcheck/random_generators.ML Tue Sep 28 22:14:02 2021 +0200 +++ b/src/HOL/Tools/Quickcheck/random_generators.ML Tue Sep 28 22:14:44 2021 +0200 @@ -96,7 +96,7 @@ fun subst_v t' = map_aterms (fn t as Free (w, _) => if v = w then t' else t | t => t); val t_rhs = lambda t_k proto_t_rhs; val eqs0 = [subst_v \<^term>\0::natural\ eq, - subst_v (\<^const>\Code_Numeral.Suc\ $ t_k) eq]; + subst_v \<^Const>\Code_Numeral.Suc for t_k\ eq]; val eqs1 = map (Pattern.rewrite_term thy rew_ts []) eqs0; val ((_, (_, eqs2)), lthy') = lthy |> BNF_LFP_Compat.primrec_simple diff -r 8d0294d877bd -r 107941e8fa01 src/HOL/Tools/Transfer/transfer.ML --- a/src/HOL/Tools/Transfer/transfer.ML Tue Sep 28 22:14:02 2021 +0200 +++ b/src/HOL/Tools/Transfer/transfer.ML Tue Sep 28 22:14:44 2021 +0200 @@ -550,8 +550,8 @@ Symtab.join (K or3) (tab1, tab2) end val tab = go [] p (t, u) Symtab.empty - fun f (a, (true, false, false)) = SOME (a, \<^const>\implies\) - | f (a, (false, true, false)) = SOME (a, \<^const>\rev_implies\) + fun f (a, (true, false, false)) = SOME (a, \<^Const>\implies\) + | f (a, (false, true, false)) = SOME (a, \<^Const>\rev_implies\) | f (a, (true, true, _)) = SOME (a, HOLogic.eq_const HOLogic.boolT) | f _ = NONE in diff -r 8d0294d877bd -r 107941e8fa01 src/HOL/Tools/record.ML --- a/src/HOL/Tools/record.ML Tue Sep 28 22:14:02 2021 +0200 +++ b/src/HOL/Tools/record.ML Tue Sep 28 22:14:44 2021 +0200 @@ -1808,8 +1808,8 @@ distinct_discsss = [], exhaust_discs = [], exhaust_sels = [], collapses = [], expands = [], split_sels = [], split_sel_asms = [], case_eq_ifs = []}; -fun lhs_of_equation (Const (\<^const_name>\Pure.eq\, _) $ t $ _) = t - | lhs_of_equation (\<^const>\Trueprop\ $ (Const (\<^const_name>\HOL.eq\, _) $ t $ _)) = t; +fun lhs_of_equation \<^Const_>\Pure.eq _ for t _\ = t + | lhs_of_equation \<^Const_>\Trueprop for \<^Const_>\HOL.eq _ for t _\\ = t; fun add_spec_rule rule = let val head = head_of (lhs_of_equation (Thm.prop_of rule)) diff -r 8d0294d877bd -r 107941e8fa01 src/HOL/Tools/set_comprehension_pointfree.ML --- a/src/HOL/Tools/set_comprehension_pointfree.ML Tue Sep 28 22:14:02 2021 +0200 +++ b/src/HOL/Tools/set_comprehension_pointfree.ML Tue Sep 28 22:14:44 2021 +0200 @@ -250,8 +250,8 @@ (map Pattern pats, Un (fm1', fm2')) end; -fun mk_formula vs (\<^const>\HOL.conj\ $ t1 $ t2) = merge_inter vs (mk_formula vs t1) (mk_formula vs t2) - | mk_formula vs (\<^const>\HOL.disj\ $ t1 $ t2) = merge_union vs (mk_formula vs t1) (mk_formula vs t2) +fun mk_formula vs \<^Const_>\conj for t1 t2\ = merge_inter vs (mk_formula vs t1) (mk_formula vs t2) + | mk_formula vs \<^Const_>\disj for t1 t2\ = merge_union vs (mk_formula vs t1) (mk_formula vs t2) | mk_formula vs t = apfst single (mk_atom vs t) fun strip_Int (Int (fm1, fm2)) = fm1 :: (strip_Int fm2)