# HG changeset patch # User wenzelm # Date 1512589389 -3600 # Node ID e615578847998153bd84f8ccd1565e850abd5241 # Parent d24dcac5eb4c0a3ddefb842ac410c106c209b8f2 prefer control symbol antiquotations; diff -r d24dcac5eb4c -r e61557884799 src/HOL/Deriv.thy --- a/src/HOL/Deriv.thy Wed Dec 06 19:34:59 2017 +0100 +++ b/src/HOL/Deriv.thy Wed Dec 06 20:43:09 2017 +0100 @@ -51,7 +51,7 @@ fun eq_rule thm = get_first (try (fn eq_thm => eq_thm OF [thm])) eq_thms in Global_Theory.add_thms_dynamic - (@{binding derivative_eq_intros}, + (\<^binding>\derivative_eq_intros\, fn context => Named_Theorems.get (Context.proof_of context) @{named_theorems derivative_intros} |> map_filter eq_rule) diff -r d24dcac5eb4c -r e61557884799 src/HOL/HOL.thy --- a/src/HOL/HOL.thy Wed Dec 06 19:34:59 2017 +0100 +++ b/src/HOL/HOL.thy Wed Dec 06 20:43:09 2017 +0100 @@ -28,23 +28,23 @@ ML_file "~~/src/Provers/quantifier1.ML" ML_file "~~/src/Tools/atomize_elim.ML" ML_file "~~/src/Tools/cong_tac.ML" -ML_file "~~/src/Tools/intuitionistic.ML" setup \Intuitionistic.method_setup @{binding iprover}\ +ML_file "~~/src/Tools/intuitionistic.ML" setup \Intuitionistic.method_setup \<^binding>\iprover\\ ML_file "~~/src/Tools/project_rule.ML" ML_file "~~/src/Tools/subtyping.ML" ML_file "~~/src/Tools/case_product.ML" -ML \Plugin_Name.declare_setup @{binding extraction}\ +ML \Plugin_Name.declare_setup \<^binding>\extraction\\ ML \ - Plugin_Name.declare_setup @{binding quickcheck_random}; - Plugin_Name.declare_setup @{binding quickcheck_exhaustive}; - Plugin_Name.declare_setup @{binding quickcheck_bounded_forall}; - Plugin_Name.declare_setup @{binding quickcheck_full_exhaustive}; - Plugin_Name.declare_setup @{binding quickcheck_narrowing}; + Plugin_Name.declare_setup \<^binding>\quickcheck_random\; + Plugin_Name.declare_setup \<^binding>\quickcheck_exhaustive\; + Plugin_Name.declare_setup \<^binding>\quickcheck_bounded_forall\; + Plugin_Name.declare_setup \<^binding>\quickcheck_full_exhaustive\; + Plugin_Name.declare_setup \<^binding>\quickcheck_narrowing\; \ ML \ - Plugin_Name.define_setup @{binding quickcheck} + Plugin_Name.define_setup \<^binding>\quickcheck\ [@{plugin quickcheck_exhaustive}, @{plugin quickcheck_random}, @{plugin quickcheck_bounded_forall}, @@ -66,7 +66,7 @@ subsubsection \Core syntax\ -setup \Axclass.class_axiomatization (@{binding type}, [])\ +setup \Axclass.class_axiomatization (\<^binding>\type\, [])\ default_sort type setup \Object_Logic.add_base_sort @{sort type}\ diff -r d24dcac5eb4c -r e61557884799 src/HOL/Tools/Argo/argo_real.ML --- a/src/HOL/Tools/Argo/argo_real.ML Wed Dec 06 19:34:59 2017 +0100 +++ b/src/HOL/Tools/Argo/argo_real.ML Wed Dec 06 20:43:09 2017 +0100 @@ -9,7 +9,7 @@ (* translating input terms *) -fun trans_type _ @{typ Real.real} tcx = SOME (Argo_Expr.Real, tcx) +fun trans_type _ \<^typ>\Real.real\ tcx = SOME (Argo_Expr.Real, tcx) | trans_type _ _ _ = NONE fun trans_term f (@{const Groups.uminus_class.uminus (real)} $ t) tcx = @@ -34,7 +34,7 @@ tcx |> f t1 ||>> f t2 |>> uncurry Argo_Expr.mk_le |> SOME | trans_term _ t tcx = (case try HOLogic.dest_number t of - SOME (@{typ Real.real}, n) => SOME (Argo_Expr.mk_num (Rat.of_int n), tcx) + SOME (\<^typ>\Real.real\, n) => SOME (Argo_Expr.mk_num (Rat.of_int n), tcx) | _ => NONE) @@ -47,7 +47,7 @@ fun mk_le t1 t2 = @{const Orderings.ord_class.less_eq (real)} $ t1 $ t2 fun mk_lt t1 t2 = @{const Orderings.ord_class.less (real)} $ t1 $ t2 -fun mk_real_num i = HOLogic.mk_number @{typ Real.real} i +fun mk_real_num i = HOLogic.mk_number \<^typ>\Real.real\ i fun mk_number n = let val (p, q) = Rat.dest n @@ -93,7 +93,7 @@ fun inv_num_conv ctxt = nums_conv mk_divide (fn (_, n) => Rat.inv n) ctxt @1 fun cmp_nums_conv ctxt b ct = - let val t = if b then @{const HOL.True} else @{const HOL.False} + let val t = if b then \<^const>\HOL.True\ else \<^const>\HOL.False\ in simp_conv ctxt (HOLogic.mk_eq (Thm.term_of ct, t)) ct end local diff -r d24dcac5eb4c -r e61557884799 src/HOL/Tools/Argo/argo_tactic.ML --- a/src/HOL/Tools/Argo/argo_tactic.ML Wed Dec 06 19:34:59 2017 +0100 +++ b/src/HOL/Tools/Argo/argo_tactic.ML Wed Dec 06 20:43:09 2017 +0100 @@ -62,7 +62,7 @@ | requires_mode Basic = [Basic, Full] | requires_mode Full = [Full] -val trace = Attrib.setup_config_string @{binding argo_trace} (K (string_of_mode None)) +val trace = Attrib.setup_config_string \<^binding>\argo_trace\ (K (string_of_mode None)) fun allows_mode ctxt = exists (equal (Config.get ctxt trace) o string_of_mode) o requires_mode @@ -77,7 +77,7 @@ (* timeout *) -val timeout = Attrib.setup_config_real @{binding argo_timeout} (K 10.0) +val timeout = Attrib.setup_config_real \<^binding>\argo_timeout\ (K 10.0) fun time_of_timeout ctxt = Time.fromReal (Config.get ctxt timeout) @@ -144,8 +144,8 @@ SOME ty => (ty, tcx) | NONE => add_new_type T tcx) -fun trans_type _ @{typ HOL.bool} = pair Argo_Expr.Bool - | trans_type ctxt (Type (@{type_name "fun"}, [T1, T2])) = +fun trans_type _ \<^typ>\HOL.bool\ = pair Argo_Expr.Bool + | trans_type ctxt (Type (\<^type_name>\fun\, [T1, T2])) = trans_type ctxt T1 ##>> trans_type ctxt T2 #>> Argo_Expr.Func | trans_type ctxt T = (fn tcx => (case ext_trans_type ctxt (trans_type ctxt) T tcx of @@ -164,22 +164,22 @@ SOME c => (Argo_Expr.mk_con c, tcx) | NONE => add_new_term ctxt t (Term.fastype_of t) tcx) -fun mk_eq @{typ HOL.bool} = Argo_Expr.mk_iff +fun mk_eq \<^typ>\HOL.bool\ = Argo_Expr.mk_iff | mk_eq _ = Argo_Expr.mk_eq -fun trans_term _ @{const HOL.True} = pair Argo_Expr.true_expr - | trans_term _ @{const HOL.False} = pair Argo_Expr.false_expr - | trans_term ctxt (@{const HOL.Not} $ t) = trans_term ctxt t #>> Argo_Expr.mk_not - | trans_term ctxt (@{const HOL.conj} $ t1 $ t2) = +fun trans_term _ \<^const>\HOL.True\ = pair Argo_Expr.true_expr + | trans_term _ \<^const>\HOL.False\ = pair Argo_Expr.false_expr + | trans_term ctxt (\<^const>\HOL.Not\ $ t) = trans_term ctxt t #>> Argo_Expr.mk_not + | trans_term ctxt (\<^const>\HOL.conj\ $ t1 $ t2) = trans_term ctxt t1 ##>> trans_term ctxt t2 #>> uncurry Argo_Expr.mk_and2 - | trans_term ctxt (@{const HOL.disj} $ t1 $ t2) = + | trans_term ctxt (\<^const>\HOL.disj\ $ t1 $ t2) = trans_term ctxt t1 ##>> trans_term ctxt t2 #>> uncurry Argo_Expr.mk_or2 - | trans_term ctxt (@{const HOL.implies} $ t1 $ t2) = + | trans_term ctxt (\<^const>\HOL.implies\ $ t1 $ t2) = trans_term ctxt t1 ##>> trans_term ctxt t2 #>> uncurry Argo_Expr.mk_imp - | trans_term ctxt (Const (@{const_name HOL.If}, _) $ t1 $ t2 $ t3) = + | trans_term ctxt (Const (\<^const_name>\HOL.If\, _) $ t1 $ t2 $ t3) = trans_term ctxt t1 ##>> trans_term ctxt t2 ##>> trans_term ctxt t3 #>> (fn ((u1, u2), u3) => Argo_Expr.mk_ite u1 u2 u3) - | trans_term ctxt (Const (@{const_name HOL.eq}, T) $ t1 $ t2) = + | trans_term ctxt (Const (\<^const_name>\HOL.eq\, T) $ t1 $ t2) = trans_term ctxt t1 ##>> trans_term ctxt t2 #>> uncurry (mk_eq (Term.domain_type T)) | trans_term ctxt (t as (t1 $ t2)) = (fn tcx => (case ext_trans_term ctxt (trans_term ctxt) t tcx of @@ -260,16 +260,16 @@ fun mk_ite t1 t2 t3 = let val T = Term.fastype_of t2 - val ite = Const (@{const_name HOL.If}, [@{typ HOL.bool}, T, T] ---> T) + val ite = Const (\<^const_name>\HOL.If\, [\<^typ>\HOL.bool\, T, T] ---> T) in ite $ t1 $ t2 $ t3 end -fun term_of _ (Argo_Expr.E (Argo_Expr.True, _)) = @{const HOL.True} - | term_of _ (Argo_Expr.E (Argo_Expr.False, _)) = @{const HOL.False} +fun term_of _ (Argo_Expr.E (Argo_Expr.True, _)) = \<^const>\HOL.True\ + | term_of _ (Argo_Expr.E (Argo_Expr.False, _)) = \<^const>\HOL.False\ | term_of cx (Argo_Expr.E (Argo_Expr.Not, [e])) = HOLogic.mk_not (term_of cx e) | term_of cx (Argo_Expr.E (Argo_Expr.And, es)) = - mk_nary' @{const HOL.True} HOLogic.mk_conj (map (term_of cx) es) + mk_nary' \<^const>\HOL.True\ HOLogic.mk_conj (map (term_of cx) es) | term_of cx (Argo_Expr.E (Argo_Expr.Or, es)) = - mk_nary' @{const HOL.False} HOLogic.mk_disj (map (term_of cx) es) + mk_nary' \<^const>\HOL.False\ HOLogic.mk_disj (map (term_of cx) es) | term_of cx (Argo_Expr.E (Argo_Expr.Imp, [e1, e2])) = HOLogic.mk_imp (term_of cx e1, term_of cx e2) | term_of cx (Argo_Expr.E (Argo_Expr.Iff, [e1, e2])) = @@ -289,7 +289,7 @@ SOME t => t | NONE => raise Fail "bad expression") -fun as_prop ct = Thm.apply @{cterm HOL.Trueprop} ct +fun as_prop ct = Thm.apply \<^cterm>\HOL.Trueprop\ ct fun cterm_of ctxt cons e = Thm.cterm_of ctxt (term_of (ctxt, cons) e) fun cprop_of ctxt cons e = as_prop (cterm_of ctxt cons e) @@ -316,7 +316,7 @@ fun with_frees ctxt n mk = let val ns = map (fn i => "P" ^ string_of_int i) (0 upto (n - 1)) - val ts = map (Free o rpair @{typ bool}) ns + val ts = map (Free o rpair \<^typ>\bool\) ns val t = mk_nary HOLogic.mk_disj (mk ts) in prove_taut ctxt ns t end @@ -373,12 +373,12 @@ fun flat_conj_conv ct = (case Thm.term_of ct of - @{const HOL.conj} $ _ $ _ => flatten_conv flat_conj_conv flatten_and_thm ct + \<^const>\HOL.conj\ $ _ $ _ => flatten_conv flat_conj_conv flatten_and_thm ct | _ => Conv.all_conv ct) fun flat_disj_conv ct = (case Thm.term_of ct of - @{const HOL.disj} $ _ $ _ => flatten_conv flat_disj_conv flatten_or_thm ct + \<^const>\HOL.disj\ $ _ $ _ => flatten_conv flat_disj_conv flatten_or_thm ct | _ => Conv.all_conv ct) fun explode rule1 rule2 thm = @@ -414,7 +414,7 @@ in mk_rewr (discharge2 lhs rhs rule) end fun with_conj f g ct = iff_intro iff_conj_thm (f o explode_conj) g ct -fun with_ndis f g ct = iff_intro iff_ndis_thm (f o explode_ndis) g (Thm.apply @{cterm HOL.Not} ct) +fun with_ndis f g ct = iff_intro iff_ndis_thm (f o explode_ndis) g (Thm.apply \<^cterm>\HOL.Not\ ct) fun swap_indices n iss = map (fn i => find_index (fn is => member (op =) is i) iss) (0 upto (n - 1)) fun sort_nary w f g (n, iss) = w (f (map hd iss)) (under_assumption (f (swap_indices n iss) o g)) @@ -484,7 +484,7 @@ | replay_args_conv ctxt [c1, c2] ct = binop_conv (replay_conv ctxt c1) (replay_conv ctxt c2) ct | replay_args_conv ctxt (c :: cs) ct = (case Term.head_of (Thm.term_of ct) of - Const (@{const_name HOL.If}, _) => + Const (\<^const_name>\HOL.If\, _) => let val (cs', c') = split_last cs in Conv.combination_conv (replay_args_conv ctxt (c :: cs')) (replay_conv ctxt c') ct end | _ => binop_conv (replay_conv ctxt c) (replay_args_conv ctxt cs) ct) @@ -520,7 +520,7 @@ val unit_rule = @{lemma "(P \ False) \ (\P \ False) \ False" by fast} val unit_rule_var = Thm.dest_arg (Thm.dest_arg1 (Thm.cprem_of unit_rule 1)) -val bogus_ct = @{cterm HOL.True} +val bogus_ct = \<^cterm>\HOL.True\ fun replay_unit_res lit (pthm, plits) (nthm, nlits) = let @@ -543,7 +543,7 @@ fun replay_hyp i ct = if i < 0 then (Thm.assume ct, [(~i, ct)]) else - let val cu = as_prop (Thm.apply @{cterm HOL.Not} (Thm.apply @{cterm HOL.Not} (Thm.dest_arg ct))) + let val cu = as_prop (Thm.apply \<^cterm>\HOL.Not\ (Thm.apply \<^cterm>\HOL.Not\ (Thm.dest_arg ct))) in (discharge (Thm.assume cu) dneg_rule, [(~i, cu)]) end @@ -601,7 +601,7 @@ fun unclausify (thm, lits) ls = (case (Thm.prop_of thm, lits) of - (@{const HOL.Trueprop} $ @{const HOL.False}, [(_, ct)]) => + (\<^const>\HOL.Trueprop\ $ \<^const>\HOL.False\, [(_, ct)]) => let val thm = Thm.implies_intr ct thm in (discharge thm unclausify_rule1 handle THM _ => discharge thm unclausify_rule2, ls) end | _ => (thm, Ord_List.union lit_ord lits ls)) @@ -668,22 +668,22 @@ let val ct = Drule.strip_imp_concl (Thm.cprop_of thm) in (case Thm.term_of ct of - @{const HOL.Trueprop} $ Var (_, @{typ HOL.bool}) => - instantiate (Thm.dest_arg ct) @{cterm HOL.False} thm - | Var _ => instantiate ct @{cprop HOL.False} thm + \<^const>\HOL.Trueprop\ $ Var (_, \<^typ>\HOL.bool\) => + instantiate (Thm.dest_arg ct) \<^cterm>\HOL.False\ thm + | Var _ => instantiate ct \<^cprop>\HOL.False\ thm | _ => thm) end fun atomize_conv ctxt ct = (case Thm.term_of ct of - @{const HOL.Trueprop} $ _ => Conv.all_conv - | @{const Pure.imp} $ _ $ _ => + \<^const>\HOL.Trueprop\ $ _ => Conv.all_conv + | \<^const>\Pure.imp\ $ _ $ _ => Conv.binop_conv (atomize_conv ctxt) then_conv Conv.rewr_conv @{thm atomize_imp} - | Const (@{const_name Pure.eq}, _) $ _ $ _ => + | Const (\<^const_name>\Pure.eq\, _) $ _ $ _ => Conv.binop_conv (atomize_conv ctxt) then_conv Conv.rewr_conv @{thm atomize_eq} - | Const (@{const_name Pure.all}, _) $ Abs _ => + | Const (\<^const_name>\Pure.all\, _) $ Abs _ => Conv.binder_conv (atomize_conv o snd) ctxt then_conv Conv.rewr_conv @{thm atomize_all} | _ => Conv.all_conv) ct @@ -723,7 +723,7 @@ | (NONE, _) => Tactical.no_tac)) ctxt val _ = - Theory.setup (Method.setup @{binding argo} + Theory.setup (Method.setup \<^binding>\argo\ (Scan.optional Attrib.thms [] >> (fn thms => fn ctxt => METHOD (fn facts => HEADGOAL (argo_tac ctxt (thms @ facts))))) diff -r d24dcac5eb4c -r e61557884799 src/HOL/Tools/Function/fun.ML --- a/src/HOL/Tools/Function/fun.ML Wed Dec 06 19:34:59 2017 +0100 +++ b/src/HOL/Tools/Function/fun.ML Wed Dec 06 20:43:09 2017 +0100 @@ -70,7 +70,7 @@ val qs = map Free (Name.invent Name.context "a" n ~~ argTs) in HOLogic.mk_eq(list_comb (Free (fname, fT), qs), - Const (@{const_name undefined}, rT)) + Const (\<^const_name>\undefined\, rT)) |> HOLogic.mk_Trueprop |> fold_rev Logic.all qs end @@ -171,7 +171,7 @@ val _ = - Outer_Syntax.local_theory' @{command_keyword fun} + Outer_Syntax.local_theory' \<^command_keyword>\fun\ "define general recursive functions (short version)" (function_parser fun_config >> (fn (config, (fixes, specs)) => add_fun_cmd fixes specs config)) diff -r d24dcac5eb4c -r e61557884799 src/HOL/Tools/Function/fun_cases.ML --- a/src/HOL/Tools/Function/fun_cases.ML Wed Dec 06 19:34:59 2017 +0100 +++ b/src/HOL/Tools/Function/fun_cases.ML Wed Dec 06 20:43:09 2017 +0100 @@ -54,7 +54,7 @@ val fun_cases_cmd = gen_fun_cases Attrib.check_src Syntax.read_prop; val _ = - Outer_Syntax.local_theory @{command_keyword fun_cases} + Outer_Syntax.local_theory \<^command_keyword>\fun_cases\ "create simplified instances of elimination rules for function equations" (Parse.and_list1 Parse_Spec.simple_specs >> (snd oo fun_cases_cmd)); diff -r d24dcac5eb4c -r e61557884799 src/HOL/Tools/Function/function.ML --- a/src/HOL/Tools/Function/function.ML Wed Dec 06 19:34:59 2017 +0100 +++ b/src/HOL/Tools/Function/function.ML Wed Dec 06 20:43:09 2017 +0100 @@ -273,13 +273,13 @@ (* outer syntax *) val _ = - Outer_Syntax.local_theory_to_proof' @{command_keyword function} + Outer_Syntax.local_theory_to_proof' \<^command_keyword>\function\ "define general recursive functions" (function_parser default_config >> (fn (config, (fixes, specs)) => function_cmd fixes specs config)) val _ = - Outer_Syntax.local_theory_to_proof @{command_keyword termination} + Outer_Syntax.local_theory_to_proof \<^command_keyword>\termination\ "prove termination of a recursive function" (Scan.option Parse.term >> termination_cmd) diff -r d24dcac5eb4c -r e61557884799 src/HOL/Tools/Function/function_common.ML --- a/src/HOL/Tools/Function/function_common.ML Wed Dec 06 19:34:59 2017 +0100 +++ b/src/HOL/Tools/Function/function_common.ML Wed Dec 06 20:43:09 2017 +0100 @@ -118,7 +118,7 @@ fun PROFILE msg = if !profile then timeap_msg msg else I -val acc_const_name = @{const_name Wellfounded.accp} +val acc_const_name = \<^const_name>\Wellfounded.accp\ fun mk_acc domT R = Const (acc_const_name, (domT --> domT --> HOLogic.boolT) --> domT --> HOLogic.boolT) $ R @@ -128,8 +128,8 @@ fun split_def ctxt check_head geq = let fun input_error msg = cat_lines [msg, Syntax.string_of_term ctxt geq] - val qs = Term.strip_qnt_vars @{const_name Pure.all} geq - val imp = Term.strip_qnt_body @{const_name Pure.all} geq + val qs = Term.strip_qnt_vars \<^const_name>\Pure.all\ geq + val imp = Term.strip_qnt_body \<^const_name>\Pure.all\ geq val (gs, eq) = Logic.strip_horn imp val (f_args, rhs) = HOLogic.dest_eq (HOLogic.dest_Trueprop eq) @@ -372,7 +372,7 @@ || (Parse.reserved "no_partials" >> K No_Partials)) fun config_parser default = - (Scan.optional (@{keyword "("} |-- Parse.!!! (Parse.list1 option_parser) --| @{keyword ")"}) []) + (Scan.optional (\<^keyword>\(\ |-- Parse.!!! (Parse.list1 option_parser) --| \<^keyword>\)\) []) >> (fn opts => fold apply_opt opts default) in fun function_parser default_cfg = diff -r d24dcac5eb4c -r e61557884799 src/HOL/Tools/Function/function_core.ML --- a/src/HOL/Tools/Function/function_core.ML Wed Dec 06 19:34:59 2017 +0100 +++ b/src/HOL/Tools/Function/function_core.ML Wed Dec 06 20:43:09 2017 +0100 @@ -394,7 +394,7 @@ (* Inductive Hypothesis: !!z. (z,x):R ==> EX!y. (z,y):G *) val ihyp = Logic.all_const domT $ Abs ("z", domT, Logic.mk_implies (HOLogic.mk_Trueprop (R $ Bound 0 $ x), - HOLogic.mk_Trueprop (Const (@{const_name Ex1}, (ranT --> boolT) --> boolT) $ + HOLogic.mk_Trueprop (Const (\<^const_name>\Ex1\, (ranT --> boolT) --> boolT) $ Abs ("y", ranT, G $ Bound 1 $ Bound 0)))) |> Thm.cterm_of ctxt @@ -494,7 +494,7 @@ Thm.make_def_binding (Config.get lthy function_internals) (derived_name_suffix defname "_sumC") val f_def = - Abs ("x", domT, Const (@{const_name Fun_Def.THE_default}, ranT --> (ranT --> boolT) --> ranT) + Abs ("x", domT, Const (\<^const_name>\Fun_Def.THE_default\, ranT --> (ranT --> boolT) --> ranT) $ (default $ Bound 0) $ Abs ("y", ranT, G $ Bound 1 $ Bound 0)) |> Syntax.check_term lthy in @@ -777,10 +777,10 @@ val R' = Free (Rn, fastype_of R) val Rrel = Free (Rn, HOLogic.mk_setT (HOLogic.mk_prodT (domT, domT))) - val inrel_R = Const (@{const_name Fun_Def.in_rel}, + val inrel_R = Const (\<^const_name>\Fun_Def.in_rel\, HOLogic.mk_setT (HOLogic.mk_prodT (domT, domT)) --> fastype_of R) $ Rrel - val wfR' = HOLogic.mk_Trueprop (Const (@{const_name Wellfounded.wfP}, + val wfR' = HOLogic.mk_Trueprop (Const (\<^const_name>\Wellfounded.wfP\, (domT --> domT --> boolT) --> boolT) $ R') |> Thm.cterm_of ctxt' (* "wf R'" *) diff -r d24dcac5eb4c -r e61557884799 src/HOL/Tools/Function/function_elims.ML --- a/src/HOL/Tools/Function/function_elims.ML Wed Dec 06 19:34:59 2017 +0100 +++ b/src/HOL/Tools/Function/function_elims.ML Wed Dec 06 20:43:09 2017 +0100 @@ -24,9 +24,9 @@ (* Extract a function and its arguments from a proposition that is either of the form "f x y z = ..." or, in case of function that returns a boolean, "f x y z" *) -fun dest_funprop (Const (@{const_name HOL.eq}, _) $ lhs $ rhs) = (strip_comb lhs, rhs) - | dest_funprop (Const (@{const_name Not}, _) $ t) = (strip_comb t, @{term False}) - | dest_funprop t = (strip_comb t, @{term True}); +fun dest_funprop (Const (\<^const_name>\HOL.eq\, _) $ lhs $ rhs) = (strip_comb lhs, rhs) + | dest_funprop (Const (\<^const_name>\Not\, _) $ t) = (strip_comb t, \<^term>\False\) + | dest_funprop t = (strip_comb t, \<^term>\True\); local @@ -34,9 +34,9 @@ let fun inspect eq = (case eq of - Const (@{const_name Trueprop}, _) $ (Const (@{const_name HOL.eq}, _) $ Free x $ t) => + Const (\<^const_name>\Trueprop\, _) $ (Const (\<^const_name>\HOL.eq\, _) $ Free x $ t) => if Logic.occs (Free x, t) then raise Match else true - | Const (@{const_name Trueprop}, _) $ (Const (@{const_name HOL.eq}, _) $ t $ Free x) => + | Const (\<^const_name>\Trueprop\, _) $ (Const (\<^const_name>\HOL.eq\, _) $ t $ Free x) => if Logic.occs (Free x, t) then raise Match else false | _ => raise Match); fun mk_eq thm = @@ -68,7 +68,7 @@ |> Tactic.rule_by_tactic ctxt (distinct_subgoals_tac THEN TRY (resolve_tac ctxt eq_boolI 1)) |> Tactic.rule_by_tactic ctxt (distinct_subgoals_tac THEN ALLGOALS (bool_subst_tac ctxt)); in - map mk_bool_elim [@{cterm True}, @{cterm False}] + map mk_bool_elim [\<^cterm>\True\, \<^cterm>\False\] end; in @@ -86,7 +86,7 @@ fun mk_funeq 0 T (acc_args, acc_lhs) = let val y = variant_free acc_lhs ("y", T) in (y, rev acc_args, HOLogic.mk_Trueprop (HOLogic.mk_eq (acc_lhs, y))) end - | mk_funeq n (Type (@{type_name fun}, [S, T])) (acc_args, acc_lhs) = + | mk_funeq n (Type (\<^type_name>\fun\, [S, T])) (acc_args, acc_lhs) = let val x = variant_free acc_lhs ("x", S) in mk_funeq (n - 1) T (x :: acc_args, acc_lhs $ x) end | mk_funeq _ _ _ = raise TERM ("Not a function", [f]); @@ -111,7 +111,7 @@ val args = HOLogic.mk_tuple arg_vars; val domT = R |> dest_Free |> snd |> hd o snd o dest_Type; - val P = Thm.cterm_of ctxt (variant_free prop ("P", @{typ bool})); + val P = Thm.cterm_of ctxt (variant_free prop ("P", \<^typ>\bool\)); val sumtree_inj = Sum_Tree.mk_inj domT n_fs (idx + 1) args; val cprop = Thm.cterm_of ctxt prop; @@ -137,7 +137,7 @@ |> Thm.forall_intr (Thm.cterm_of ctxt rhs_var); val bool_elims = - if fastype_of rhs_var = @{typ bool} + if fastype_of rhs_var = \<^typ>\bool\ then mk_bool_elims ctxt elim_stripped else []; diff -r d24dcac5eb4c -r e61557884799 src/HOL/Tools/Function/function_lib.ML --- a/src/HOL/Tools/Function/function_lib.ML Wed Dec 06 19:34:59 2017 +0100 +++ b/src/HOL/Tools/Function/function_lib.ML Wed Dec 06 20:43:09 2017 +0100 @@ -50,7 +50,7 @@ (* Removes all quantifiers from a term, replacing bound variables by frees. *) -fun dest_all_all (t as (Const (@{const_name Pure.all},_) $ _)) = +fun dest_all_all (t as (Const (\<^const_name>\Pure.all\,_) $ _)) = let val (v,b) = Logic.dest_all t |> apfst Free val (vs, b') = dest_all_all b @@ -129,7 +129,7 @@ (* instance for unions *) fun regroup_union_conv ctxt = - regroup_conv ctxt @{const_abbrev Set.empty} @{const_name Lattices.sup} + regroup_conv ctxt \<^const_abbrev>\Set.empty\ \<^const_name>\Lattices.sup\ (map (fn t => t RS eq_reflection) (@{thms Un_ac} @ @{thms Un_empty_right} @ @{thms Un_empty_left})) diff -r d24dcac5eb4c -r e61557884799 src/HOL/Tools/Function/induction_schema.ML --- a/src/HOL/Tools/Function/induction_schema.ML Wed Dec 06 19:34:59 2017 +0100 +++ b/src/HOL/Tools/Function/induction_schema.ML Wed Dec 06 20:43:09 2017 +0100 @@ -140,7 +140,7 @@ end fun mk_wf R (IndScheme {T, ...}) = - HOLogic.Trueprop $ (Const (@{const_name wf}, mk_relT T --> HOLogic.boolT) $ R) + HOLogic.Trueprop $ (Const (\<^const_name>\wf\, mk_relT T --> HOLogic.boolT) $ R) fun mk_ineqs R thesisn (IndScheme {T, cases, branches}) = let @@ -215,7 +215,7 @@ val ihyp = Logic.all_const T $ Abs ("z", T, Logic.mk_implies (HOLogic.mk_Trueprop ( - Const (@{const_name Set.member}, HOLogic.mk_prodT (T, T) --> mk_relT T --> HOLogic.boolT) + Const (\<^const_name>\Set.member\, HOLogic.mk_prodT (T, T) --> mk_relT T --> HOLogic.boolT) $ (HOLogic.pair_const T T $ Bound 0 $ x) $ R), HOLogic.mk_Trueprop (P_comp $ Bound 0))) diff -r d24dcac5eb4c -r e61557884799 src/HOL/Tools/Function/lexicographic_order.ML --- a/src/HOL/Tools/Function/lexicographic_order.ML Wed Dec 06 19:34:59 2017 +0100 +++ b/src/HOL/Tools/Function/lexicographic_order.ML Wed Dec 06 20:43:09 2017 +0100 @@ -22,9 +22,9 @@ let val relT = HOLogic.mk_setT (HOLogic.mk_prodT (domT, domT)) val mlexT = (domT --> HOLogic.natT) --> relT --> relT - fun mk_ms [] = Const (@{const_abbrev Set.empty}, relT) + fun mk_ms [] = Const (\<^const_abbrev>\Set.empty\, relT) | mk_ms (f::fs) = - Const (@{const_name mlex_prod}, mlexT) $ f $ mk_ms fs + Const (\<^const_name>\mlex_prod\, mlexT) $ f $ mk_ms fs in mk_ms mfuns end @@ -70,13 +70,13 @@ let val goals = Thm.cterm_of ctxt o mk_goal (vars, prems, mfun $ lhs, mfun $ rhs) in - (case try_proof ctxt (goals @{const_name Orderings.less}) solve_tac of + (case try_proof ctxt (goals \<^const_name>\Orderings.less\) solve_tac of Solved thm => Less thm | Stuck thm => - (case try_proof ctxt (goals @{const_name Orderings.less_eq}) solve_tac of + (case try_proof ctxt (goals \<^const_name>\Orderings.less_eq\) solve_tac of Solved thm2 => LessEq (thm2, thm) | Stuck thm2 => - if Thm.prems_of thm2 = [HOLogic.Trueprop $ @{term False}] then False thm2 + if Thm.prems_of thm2 = [HOLogic.Trueprop $ \<^term>\False\] then False thm2 else None (thm2, thm) | _ => raise Match) (* FIXME *) | _ => raise Match) diff -r d24dcac5eb4c -r e61557884799 src/HOL/Tools/Function/measure_functions.ML --- a/src/HOL/Tools/Function/measure_functions.ML Wed Dec 06 19:34:59 2017 +0100 +++ b/src/HOL/Tools/Function/measure_functions.ML Wed Dec 06 20:43:09 2017 +0100 @@ -15,10 +15,10 @@ (** User-declared size functions **) fun mk_is_measure t = - Const (@{const_name is_measure}, fastype_of t --> HOLogic.boolT) $ t + Const (\<^const_name>\is_measure\, fastype_of t --> HOLogic.boolT) $ t fun find_measures ctxt T = - DEPTH_SOLVE (resolve_tac ctxt (rev (Named_Theorems.get ctxt @{named_theorems measure_function})) 1) + DEPTH_SOLVE (resolve_tac ctxt (rev (Named_Theorems.get ctxt \<^named_theorems>\measure_function\)) 1) (HOLogic.mk_Trueprop (mk_is_measure (Var (("f",0), T --> HOLogic.natT))) |> Thm.cterm_of ctxt |> Goal.init) |> Seq.map (Thm.prop_of #> (fn _ $ (_ $ (_ $ f)) => f)) @@ -30,17 +30,17 @@ fun constant_0 T = Abs ("x", T, HOLogic.zero) fun constant_1 T = Abs ("x", T, HOLogic.Suc_zero) -fun mk_funorder_funs (Type (@{type_name Sum_Type.sum}, [fT, sT])) = +fun mk_funorder_funs (Type (\<^type_name>\Sum_Type.sum\, [fT, sT])) = map (fn m => Sum_Tree.mk_sumcase fT sT HOLogic.natT m (constant_0 sT)) (mk_funorder_funs fT) @ map (fn m => Sum_Tree.mk_sumcase fT sT HOLogic.natT (constant_0 fT) m) (mk_funorder_funs sT) | mk_funorder_funs T = [ constant_1 T ] -fun mk_ext_base_funs ctxt (Type (@{type_name Sum_Type.sum}, [fT, sT])) = +fun mk_ext_base_funs ctxt (Type (\<^type_name>\Sum_Type.sum\, [fT, sT])) = map_product (Sum_Tree.mk_sumcase fT sT HOLogic.natT) (mk_ext_base_funs ctxt fT) (mk_ext_base_funs ctxt sT) | mk_ext_base_funs ctxt T = find_measures ctxt T -fun mk_all_measure_funs ctxt (T as Type (@{type_name Sum_Type.sum}, _)) = +fun mk_all_measure_funs ctxt (T as Type (\<^type_name>\Sum_Type.sum\, _)) = mk_ext_base_funs ctxt T @ mk_funorder_funs T | mk_all_measure_funs ctxt T = find_measures ctxt T diff -r d24dcac5eb4c -r e61557884799 src/HOL/Tools/Function/mutual.ML --- a/src/HOL/Tools/Function/mutual.ML Wed Dec 06 19:34:59 2017 +0100 +++ b/src/HOL/Tools/Function/mutual.ML Wed Dec 06 20:43:09 2017 +0100 @@ -247,7 +247,7 @@ fun mutual_cases_rule ctxt cases_rule n ST (MutualPart {i, cargTs = Ts, ...}) = let val [P, x] = - Variable.variant_frees ctxt [] [("P", @{typ bool}), ("x", HOLogic.mk_tupleT Ts)] + Variable.variant_frees ctxt [] [("P", \<^typ>\bool\), ("x", HOLogic.mk_tupleT Ts)] |> map (Thm.cterm_of ctxt o Free); val sumtree_inj = Thm.cterm_of ctxt (Sum_Tree.mk_inj ST n i (Thm.term_of x)); diff -r d24dcac5eb4c -r e61557884799 src/HOL/Tools/Function/partial_function.ML --- a/src/HOL/Tools/Function/partial_function.ML Wed Dec 06 19:34:59 2017 +0100 +++ b/src/HOL/Tools/Function/partial_function.ML Wed Dec 06 20:43:09 2017 +0100 @@ -109,7 +109,7 @@ fun mono_tac ctxt = K (Local_Defs.unfold0_tac ctxt [@{thm curry_def}]) THEN' (TRY o REPEAT_ALL_NEW - (resolve_tac ctxt (rev (Named_Theorems.get ctxt @{named_theorems partial_function_mono})) + (resolve_tac ctxt (rev (Named_Theorems.get ctxt \<^named_theorems>\partial_function_mono\)) ORELSE' split_cases_tac ctxt)); @@ -135,7 +135,7 @@ (*** currying transformation ***) fun curry_const (A, B, C) = - Const (@{const_name Product_Type.curry}, + Const (\<^const_name>\Product_Type.curry\, [HOLogic.mk_prodT (A, B) --> C, A, B] ---> C); fun mk_curry f = @@ -310,10 +310,10 @@ val add_partial_function = gen_add_partial_function Specification.check_multi_specs; val add_partial_function_cmd = gen_add_partial_function Specification.read_multi_specs; -val mode = @{keyword "("} |-- Parse.name --| @{keyword ")"}; +val mode = \<^keyword>\(\ |-- Parse.name --| \<^keyword>\)\; val _ = - Outer_Syntax.local_theory @{command_keyword partial_function} "define partial function" + Outer_Syntax.local_theory \<^command_keyword>\partial_function\ "define partial function" ((mode -- (Parse.vars -- (Parse.where_ |-- Parse_Spec.simple_spec))) >> (fn (mode, (vars, spec)) => add_partial_function_cmd mode vars spec #> #2)); diff -r d24dcac5eb4c -r e61557884799 src/HOL/Tools/Function/sum_tree.ML --- a/src/HOL/Tools/Function/sum_tree.ML Wed Dec 06 19:34:59 2017 +0100 +++ b/src/HOL/Tools/Function/sum_tree.ML Wed Dec 06 20:43:09 2017 +0100 @@ -30,28 +30,28 @@ {left = (fn f => f o left), right = (fn f => f o right), init = I} len i init (* Sum types *) -fun mk_sumT LT RT = Type (@{type_name Sum_Type.sum}, [LT, RT]) +fun mk_sumT LT RT = Type (\<^type_name>\Sum_Type.sum\, [LT, RT]) fun mk_sumcase TL TR T l r = - Const (@{const_name sum.case_sum}, (TL --> T) --> (TR --> T) --> mk_sumT TL TR --> T) $ l $ r + Const (\<^const_name>\sum.case_sum\, (TL --> T) --> (TR --> T) --> mk_sumT TL TR --> T) $ l $ r val App = curry op $ fun mk_inj ST n i = access_top_down { init = (ST, I : term -> term), - left = (fn (T as Type (@{type_name Sum_Type.sum}, [LT, RT]), inj) => - (LT, inj o App (Const (@{const_name Inl}, LT --> T)))), - right =(fn (T as Type (@{type_name Sum_Type.sum}, [LT, RT]), inj) => - (RT, inj o App (Const (@{const_name Inr}, RT --> T))))} n i + left = (fn (T as Type (\<^type_name>\Sum_Type.sum\, [LT, RT]), inj) => + (LT, inj o App (Const (\<^const_name>\Inl\, LT --> T)))), + right =(fn (T as Type (\<^type_name>\Sum_Type.sum\, [LT, RT]), inj) => + (RT, inj o App (Const (\<^const_name>\Inr\, RT --> T))))} n i |> snd fun mk_proj ST n i = access_top_down { init = (ST, I : term -> term), - left = (fn (T as Type (@{type_name Sum_Type.sum}, [LT, RT]), proj) => - (LT, App (Const (@{const_name Sum_Type.projl}, T --> LT)) o proj)), - right =(fn (T as Type (@{type_name Sum_Type.sum}, [LT, RT]), proj) => - (RT, App (Const (@{const_name Sum_Type.projr}, T --> RT)) o proj))} n i + left = (fn (T as Type (\<^type_name>\Sum_Type.sum\, [LT, RT]), proj) => + (LT, App (Const (\<^const_name>\Sum_Type.projl\, T --> LT)) o proj)), + right =(fn (T as Type (\<^type_name>\Sum_Type.sum\, [LT, RT]), proj) => + (RT, App (Const (\<^const_name>\Sum_Type.projr\, T --> RT)) o proj))} n i |> snd fun mk_sumcases T fs = diff -r d24dcac5eb4c -r e61557884799 src/HOL/Tools/Function/termination.ML --- a/src/HOL/Tools/Function/termination.ML Wed Dec 06 19:34:59 2017 +0100 +++ b/src/HOL/Tools/Function/termination.ML Wed Dec 06 20:43:09 2017 +0100 @@ -67,14 +67,14 @@ (* concrete versions for sum types *) -fun is_inj (Const (@{const_name Sum_Type.Inl}, _) $ _) = true - | is_inj (Const (@{const_name Sum_Type.Inr}, _) $ _) = true +fun is_inj (Const (\<^const_name>\Sum_Type.Inl\, _) $ _) = true + | is_inj (Const (\<^const_name>\Sum_Type.Inr\, _) $ _) = true | is_inj _ = false -fun dest_inl (Const (@{const_name Sum_Type.Inl}, _) $ t) = SOME t +fun dest_inl (Const (\<^const_name>\Sum_Type.Inl\, _) $ t) = SOME t | dest_inl _ = NONE -fun dest_inr (Const (@{const_name Sum_Type.Inr}, _) $ t) = SOME t +fun dest_inr (Const (\<^const_name>\Sum_Type.Inr\, _) $ t) = SOME t | dest_inr _ = NONE @@ -92,7 +92,7 @@ end (* compute list of types for nodes *) -fun node_types sk T = dest_tree (fn Type (@{type_name Sum_Type.sum}, [LT, RT]) => (LT, RT)) sk T |> map snd +fun node_types sk T = dest_tree (fn Type (\<^type_name>\Sum_Type.sum\, [LT, RT]) => (LT, RT)) sk T |> map snd (* find index and raw term *) fun dest_inj (SLeaf i) trm = (i, trm) @@ -125,11 +125,11 @@ fun mk_sum_skel rel = let - val cs = Function_Lib.dest_binop_list @{const_name Lattices.sup} rel - fun collect_pats (Const (@{const_name Collect}, _) $ Abs (_, _, c)) = + val cs = Function_Lib.dest_binop_list \<^const_name>\Lattices.sup\ rel + fun collect_pats (Const (\<^const_name>\Collect\, _) $ Abs (_, _, c)) = let - val (Const (@{const_name HOL.conj}, _) $ (Const (@{const_name HOL.eq}, _) $ _ $ (Const (@{const_name Pair}, _) $ r $ l)) $ _) - = Term.strip_qnt_body @{const_name Ex} c + val (Const (\<^const_name>\HOL.conj\, _) $ (Const (\<^const_name>\HOL.eq\, _) $ _ $ (Const (\<^const_name>\Pair\, _) $ r $ l)) $ _) + = Term.strip_qnt_body \<^const_name>\Ex\ c in cons r o cons l end in mk_skel (fold collect_pats cs []) @@ -138,8 +138,8 @@ fun prove_chain ctxt chain_tac (c1, c2) = let val goal = - HOLogic.mk_eq (HOLogic.mk_binop @{const_name Relation.relcomp} (c1, c2), - Const (@{const_abbrev Set.empty}, fastype_of c1)) + HOLogic.mk_eq (HOLogic.mk_binop \<^const_name>\Relation.relcomp\ (c1, c2), + Const (\<^const_abbrev>\Set.empty\, fastype_of c1)) |> HOLogic.mk_Trueprop (* "C1 O C2 = {}" *) in (case Function_Lib.try_proof ctxt (Thm.cterm_of ctxt goal) chain_tac of @@ -148,13 +148,13 @@ end -fun dest_call' sk (Const (@{const_name Collect}, _) $ Abs (_, _, c)) = +fun dest_call' sk (Const (\<^const_name>\Collect\, _) $ Abs (_, _, c)) = let - val vs = Term.strip_qnt_vars @{const_name Ex} c + val vs = Term.strip_qnt_vars \<^const_name>\Ex\ c (* FIXME: throw error "dest_call" for malformed terms *) - val (Const (@{const_name HOL.conj}, _) $ (Const (@{const_name HOL.eq}, _) $ _ $ (Const (@{const_name Pair}, _) $ r $ l)) $ Gam) - = Term.strip_qnt_body @{const_name Ex} c + val (Const (\<^const_name>\HOL.conj\, _) $ (Const (\<^const_name>\HOL.eq\, _) $ _ $ (Const (\<^const_name>\Pair\, _) $ r $ l)) $ Gam) + = Term.strip_qnt_body \<^const_name>\Ex\ c val (p, l') = dest_inj sk l val (q, r') = dest_inj sk r in @@ -170,16 +170,16 @@ try_proof ctxt (Thm.cterm_of ctxt (Logic.list_all (vs, Logic.mk_implies (HOLogic.mk_Trueprop Gam, - HOLogic.mk_Trueprop (Const (rel, @{typ "nat \ nat \ bool"}) + HOLogic.mk_Trueprop (Const (rel, \<^typ>\nat \ nat \ bool\) $ (m2 $ r) $ (m1 $ l)))))) tac in - (case try @{const_name Orderings.less} of + (case try \<^const_name>\Orderings.less\ of Solved thm => Less thm | Stuck thm => - (case try @{const_name Orderings.less_eq} of + (case try \<^const_name>\Orderings.less_eq\ of Solved thm2 => LessEq (thm2, thm) | Stuck thm2 => - if Thm.prems_of thm2 = [HOLogic.Trueprop $ @{term False}] + if Thm.prems_of thm2 = [HOLogic.Trueprop $ \<^term>\False\] then False thm2 else None (thm2, thm) | _ => raise Match) (* FIXME *) | _ => raise Match) @@ -225,13 +225,13 @@ fun CALLS tac i st = if Thm.no_prems st then all_tac st else case Thm.term_of (Thm.cprem_of st i) of - (_ $ (_ $ rel)) => tac (Function_Lib.dest_binop_list @{const_name Lattices.sup} rel, i) st + (_ $ (_ $ rel)) => tac (Function_Lib.dest_binop_list \<^const_name>\Lattices.sup\ rel, i) st |_ => no_tac st type ttac = data -> int -> tactic fun TERMINATION ctxt atac tac = - SUBGOAL (fn (_ $ (Const (@{const_name wf}, wfT) $ rel), i) => + SUBGOAL (fn (_ $ (Const (\<^const_name>\wf\, wfT) $ rel), i) => let val (T, _) = HOLogic.dest_prodT (HOLogic.dest_setT (domain_type wfT)) in @@ -258,7 +258,7 @@ val pT = HOLogic.mk_prodT (T, T) val n = length qs val peq = HOLogic.eq_const pT $ Bound n $ (HOLogic.pair_const T T $ l $ r) - val conds' = if null conds then [@{term True}] else conds + val conds' = if null conds then [\<^term>\True\] else conds in HOLogic.Collect_const pT $ Abs ("uu_", pT, @@ -284,9 +284,9 @@ val relation = if null ineqs - then Const (@{const_abbrev Set.empty}, fastype_of rel) + then Const (\<^const_abbrev>\Set.empty\, fastype_of rel) else map mk_compr ineqs - |> foldr1 (HOLogic.mk_binop @{const_name Lattices.sup}) + |> foldr1 (HOLogic.mk_binop \<^const_name>\Lattices.sup\) fun solve_membership_tac i = (EVERY' (replicate (i - 2) (resolve_tac ctxt @{thms UnI2})) (* pick the right component of the union *) diff -r d24dcac5eb4c -r e61557884799 src/HOL/Tools/Meson/meson.ML --- a/src/HOL/Tools/Meson/meson.ML Wed Dec 06 19:34:59 2017 +0100 +++ b/src/HOL/Tools/Meson/meson.ML Wed Dec 06 20:43:09 2017 +0100 @@ -48,11 +48,11 @@ structure Meson : MESON = struct -val trace = Attrib.setup_config_bool @{binding meson_trace} (K false) +val trace = Attrib.setup_config_bool \<^binding>\meson_trace\ (K false) fun trace_msg ctxt msg = if Config.get ctxt trace then tracing (msg ()) else () -val max_clauses = Attrib.setup_config_int @{binding meson_max_clauses} (K 60) +val max_clauses = Attrib.setup_config_int \<^binding>\meson_max_clauses\ (K 60) (*No known example (on 1-5-2007) needs even thirty*) val iter_deepen_limit = 50; @@ -100,7 +100,7 @@ try (fn () => let val thy = Proof_Context.theory_of ctxt val tmA = Thm.concl_of thA - val Const(@{const_name Pure.imp},_) $ tmB $ _ = Thm.prop_of thB + val Const(\<^const_name>\Pure.imp\,_) $ tmB $ _ = Thm.prop_of thB val tenv = Pattern.first_order_match thy (tmB, tmA) (Vartab.empty, Vartab.empty) |> snd @@ -121,18 +121,18 @@ fun fix_bound_var_names old_t new_t = let - fun quant_of @{const_name All} = SOME true - | quant_of @{const_name Ball} = SOME true - | quant_of @{const_name Ex} = SOME false - | quant_of @{const_name Bex} = SOME false + fun quant_of \<^const_name>\All\ = SOME true + | quant_of \<^const_name>\Ball\ = SOME true + | quant_of \<^const_name>\Ex\ = SOME false + | quant_of \<^const_name>\Bex\ = SOME false | quant_of _ = NONE val flip_quant = Option.map not fun some_eq (SOME x) (SOME y) = x = y | some_eq _ _ = false fun add_names quant (Const (quant_s, _) $ Abs (s, _, t')) = add_names quant t' #> some_eq quant (quant_of quant_s) ? cons s - | add_names quant (@{const Not} $ t) = add_names (flip_quant quant) t - | add_names quant (@{const implies} $ t1 $ t2) = + | add_names quant (\<^const>\Not\ $ t) = add_names (flip_quant quant) t + | add_names quant (\<^const>\implies\ $ t1 $ t2) = add_names (flip_quant quant) t1 #> add_names quant t2 | add_names quant (t1 $ t2) = fold (add_names quant) [t1, t2] | add_names _ _ = I @@ -169,7 +169,7 @@ workaround is to instantiate "?P := (%c. ... c ... c ... c ...)" manually. *) fun quant_resolve_tac ctxt th i st = case (Thm.concl_of st, Thm.prop_of th) of - (@{const Trueprop} $ (Var _ $ (c as Free _)), @{const Trueprop} $ _) => + (\<^const>\Trueprop\ $ (Var _ $ (c as Free _)), \<^const>\Trueprop\ $ _) => let val cc = Thm.cterm_of ctxt c val ct = Thm.dest_arg (Thm.cprop_of th) @@ -197,21 +197,21 @@ (*Are any of the logical connectives in "bs" present in the term?*) fun has_conns bs = let fun has (Const _) = false - | has (Const(@{const_name Trueprop},_) $ p) = has p - | has (Const(@{const_name Not},_) $ p) = has p - | has (Const(@{const_name HOL.disj},_) $ p $ q) = member (op =) bs @{const_name HOL.disj} orelse has p orelse has q - | has (Const(@{const_name HOL.conj},_) $ p $ q) = member (op =) bs @{const_name HOL.conj} orelse has p orelse has q - | has (Const(@{const_name All},_) $ Abs(_,_,p)) = member (op =) bs @{const_name All} orelse has p - | has (Const(@{const_name Ex},_) $ Abs(_,_,p)) = member (op =) bs @{const_name Ex} orelse has p + | has (Const(\<^const_name>\Trueprop\,_) $ p) = has p + | has (Const(\<^const_name>\Not\,_) $ p) = has p + | has (Const(\<^const_name>\HOL.disj\,_) $ p $ q) = member (op =) bs \<^const_name>\HOL.disj\ orelse has p orelse has q + | has (Const(\<^const_name>\HOL.conj\,_) $ p $ q) = member (op =) bs \<^const_name>\HOL.conj\ orelse has p orelse has q + | has (Const(\<^const_name>\All\,_) $ Abs(_,_,p)) = member (op =) bs \<^const_name>\All\ orelse has p + | has (Const(\<^const_name>\Ex\,_) $ Abs(_,_,p)) = member (op =) bs \<^const_name>\Ex\ orelse has p | has _ = false in has end; (**** Clause handling ****) -fun literals (Const(@{const_name Trueprop},_) $ P) = literals P - | literals (Const(@{const_name HOL.disj},_) $ P $ Q) = literals P @ literals Q - | literals (Const(@{const_name Not},_) $ P) = [(false,P)] +fun literals (Const(\<^const_name>\Trueprop\,_) $ P) = literals P + | literals (Const(\<^const_name>\HOL.disj\,_) $ P $ Q) = literals P @ literals Q + | literals (Const(\<^const_name>\Not\,_) $ P) = [(false,P)] | literals P = [(true,P)]; (*number of literals in a term*) @@ -220,23 +220,23 @@ (*** Tautology Checking ***) -fun signed_lits_aux (Const (@{const_name HOL.disj}, _) $ P $ Q) (poslits, neglits) = +fun signed_lits_aux (Const (\<^const_name>\HOL.disj\, _) $ P $ Q) (poslits, neglits) = signed_lits_aux Q (signed_lits_aux P (poslits, neglits)) - | signed_lits_aux (Const(@{const_name Not},_) $ P) (poslits, neglits) = (poslits, P::neglits) + | signed_lits_aux (Const(\<^const_name>\Not\,_) $ P) (poslits, neglits) = (poslits, P::neglits) | signed_lits_aux P (poslits, neglits) = (P::poslits, neglits); fun signed_lits th = signed_lits_aux (HOLogic.dest_Trueprop (Thm.concl_of th)) ([],[]); (*Literals like X=X are tautologous*) -fun taut_poslit (Const(@{const_name HOL.eq},_) $ t $ u) = t aconv u - | taut_poslit (Const(@{const_name True},_)) = true +fun taut_poslit (Const(\<^const_name>\HOL.eq\,_) $ t $ u) = t aconv u + | taut_poslit (Const(\<^const_name>\True\,_)) = true | taut_poslit _ = false; fun is_taut th = let val (poslits,neglits) = signed_lits th in exists taut_poslit poslits orelse - exists (member (op aconv) neglits) (@{term False} :: poslits) + exists (member (op aconv) neglits) (\<^term>\False\ :: poslits) end handle TERM _ => false; (*probably dest_Trueprop on a weird theorem*) @@ -256,18 +256,18 @@ fun refl_clause_aux 0 th = th | refl_clause_aux n th = case HOLogic.dest_Trueprop (Thm.concl_of th) of - (Const (@{const_name HOL.disj}, _) $ (Const (@{const_name HOL.disj}, _) $ _ $ _) $ _) => + (Const (\<^const_name>\HOL.disj\, _) $ (Const (\<^const_name>\HOL.disj\, _) $ _ $ _) $ _) => refl_clause_aux n (th RS disj_assoc) (*isolate an atom as first disjunct*) - | (Const (@{const_name HOL.disj}, _) $ (Const(@{const_name Not},_) $ (Const(@{const_name HOL.eq},_) $ t $ u)) $ _) => + | (Const (\<^const_name>\HOL.disj\, _) $ (Const(\<^const_name>\Not\,_) $ (Const(\<^const_name>\HOL.eq\,_) $ t $ u)) $ _) => if eliminable(t,u) then refl_clause_aux (n-1) (th RS not_refl_disj_D) (*Var inequation: delete*) else refl_clause_aux (n-1) (th RS disj_comm) (*not between Vars: ignore*) - | (Const (@{const_name HOL.disj}, _) $ _ $ _) => refl_clause_aux n (th RS disj_comm) + | (Const (\<^const_name>\HOL.disj\, _) $ _ $ _) => refl_clause_aux n (th RS disj_comm) | _ => (*not a disjunction*) th; -fun notequal_lits_count (Const (@{const_name HOL.disj}, _) $ P $ Q) = +fun notequal_lits_count (Const (\<^const_name>\HOL.disj\, _) $ P $ Q) = notequal_lits_count P + notequal_lits_count Q - | notequal_lits_count (Const(@{const_name Not},_) $ (Const(@{const_name HOL.eq},_) $ _ $ _)) = 1 + | notequal_lits_count (Const(\<^const_name>\Not\,_) $ (Const(\<^const_name>\HOL.eq\,_) $ _ $ _)) = 1 | notequal_lits_count _ = 0; (*Simplify a clause by applying reflexivity to its negated equality literals*) @@ -312,26 +312,26 @@ fun prod x y = if x < bound andalso y < bound then x*y else bound (*Estimate the number of clauses in order to detect infeasible theorems*) - fun signed_nclauses b (Const(@{const_name Trueprop},_) $ t) = signed_nclauses b t - | signed_nclauses b (Const(@{const_name Not},_) $ t) = signed_nclauses (not b) t - | signed_nclauses b (Const(@{const_name HOL.conj},_) $ t $ u) = + fun signed_nclauses b (Const(\<^const_name>\Trueprop\,_) $ t) = signed_nclauses b t + | signed_nclauses b (Const(\<^const_name>\Not\,_) $ t) = signed_nclauses (not b) t + | signed_nclauses b (Const(\<^const_name>\HOL.conj\,_) $ t $ u) = if b then sum (signed_nclauses b t) (signed_nclauses b u) else prod (signed_nclauses b t) (signed_nclauses b u) - | signed_nclauses b (Const(@{const_name HOL.disj},_) $ t $ u) = + | signed_nclauses b (Const(\<^const_name>\HOL.disj\,_) $ t $ u) = if b then prod (signed_nclauses b t) (signed_nclauses b u) else sum (signed_nclauses b t) (signed_nclauses b u) - | signed_nclauses b (Const(@{const_name HOL.implies},_) $ t $ u) = + | signed_nclauses b (Const(\<^const_name>\HOL.implies\,_) $ t $ u) = if b then prod (signed_nclauses (not b) t) (signed_nclauses b u) else sum (signed_nclauses (not b) t) (signed_nclauses b u) - | signed_nclauses b (Const(@{const_name HOL.eq}, Type ("fun", [T, _])) $ t $ u) = + | signed_nclauses b (Const(\<^const_name>\HOL.eq\, Type ("fun", [T, _])) $ t $ u) = if T = HOLogic.boolT then (*Boolean equality is if-and-only-if*) if b then sum (prod (signed_nclauses (not b) t) (signed_nclauses b u)) (prod (signed_nclauses (not b) u) (signed_nclauses b t)) else sum (prod (signed_nclauses b t) (signed_nclauses b u)) (prod (signed_nclauses (not b) t) (signed_nclauses (not b) u)) else 1 - | signed_nclauses b (Const(@{const_name Ex}, _) $ Abs (_,_,t)) = signed_nclauses b t - | signed_nclauses b (Const(@{const_name All},_) $ Abs (_,_,t)) = signed_nclauses b t + | signed_nclauses b (Const(\<^const_name>\Ex\, _) $ Abs (_,_,t)) = signed_nclauses b t + | signed_nclauses b (Const(\<^const_name>\All\,_) $ Abs (_,_,t)) = signed_nclauses b t | signed_nclauses _ _ = 1; (* literal *) in signed_nclauses true t end @@ -346,7 +346,7 @@ val spec_var = Thm.dest_arg (Thm.dest_arg (#2 (Thm.dest_implies (Thm.cprop_of spec)))) |> Thm.term_of |> dest_Var; - fun name_of (Const (@{const_name All}, _) $ Abs(x, _, _)) = x | name_of _ = Name.uu; + fun name_of (Const (\<^const_name>\All\, _) $ Abs(x, _, _)) = x | name_of _ = Name.uu; in fun freeze_spec th ctxt = let @@ -370,18 +370,18 @@ let val ctxt_ref = Unsynchronized.ref ctxt (* FIXME ??? *) fun cnf_aux (th,ths) = if not (can HOLogic.dest_Trueprop (Thm.prop_of th)) then ths (*meta-level: ignore*) - else if not (has_conns [@{const_name All}, @{const_name Ex}, @{const_name HOL.conj}] (Thm.prop_of th)) + else if not (has_conns [\<^const_name>\All\, \<^const_name>\Ex\, \<^const_name>\HOL.conj\] (Thm.prop_of th)) then nodups ctxt th :: ths (*no work to do, terminate*) else case head_of (HOLogic.dest_Trueprop (Thm.concl_of th)) of - Const (@{const_name HOL.conj}, _) => (*conjunction*) + Const (\<^const_name>\HOL.conj\, _) => (*conjunction*) cnf_aux (th RS conjunct1, cnf_aux (th RS conjunct2, ths)) - | Const (@{const_name All}, _) => (*universal quantifier*) + | Const (\<^const_name>\All\, _) => (*universal quantifier*) let val (th', ctxt') = freeze_spec th (! ctxt_ref) in ctxt_ref := ctxt'; cnf_aux (th', ths) end - | Const (@{const_name Ex}, _) => + | Const (\<^const_name>\Ex\, _) => (*existential quantifier: Insert Skolem functions*) cnf_aux (apply_skolem_theorem (! ctxt_ref) (th, old_skolem_ths), ths) - | Const (@{const_name HOL.disj}, _) => + | Const (\<^const_name>\HOL.disj\, _) => (*Disjunction of P, Q: Create new goal of proving ?P | ?Q and solve it using all combinations of converting P, Q to CNF.*) (*There is one assumption, which gets bound to prem and then normalized via @@ -409,8 +409,8 @@ (**** Generation of contrapositives ****) -fun is_left (Const (@{const_name Trueprop}, _) $ - (Const (@{const_name HOL.disj}, _) $ (Const (@{const_name HOL.disj}, _) $ _ $ _) $ _)) = true +fun is_left (Const (\<^const_name>\Trueprop\, _) $ + (Const (\<^const_name>\HOL.disj\, _) $ (Const (\<^const_name>\HOL.disj\, _) $ _ $ _) $ _)) = true | is_left _ = false; (*Associate disjuctions to right -- make leftmost disjunct a LITERAL*) @@ -431,8 +431,8 @@ fun rigid t = not (is_Var (head_of t)); -fun ok4horn (Const (@{const_name Trueprop},_) $ (Const (@{const_name HOL.disj}, _) $ t $ _)) = rigid t - | ok4horn (Const (@{const_name Trueprop},_) $ t) = rigid t +fun ok4horn (Const (\<^const_name>\Trueprop\,_) $ (Const (\<^const_name>\HOL.disj\, _) $ t $ _)) = rigid t + | ok4horn (Const (\<^const_name>\Trueprop\,_) $ t) = rigid t | ok4horn _ = false; (*Create a meta-level Horn clause*) @@ -466,7 +466,7 @@ (***** MESON PROOF PROCEDURE *****) -fun rhyps (Const(@{const_name Pure.imp},_) $ (Const(@{const_name Trueprop},_) $ A) $ phi, +fun rhyps (Const(\<^const_name>\Pure.imp\,_) $ (Const(\<^const_name>\Trueprop\,_) $ A) $ phi, As) = rhyps(phi, A::As) | rhyps (_, As) = As; @@ -511,8 +511,8 @@ val nnf_rls = [imp_to_disjD, iff_to_disjD, not_conjD, not_disjD, not_impD, not_iffD, not_allD, not_exD, not_notD]; -fun ok4nnf (Const (@{const_name Trueprop},_) $ (Const (@{const_name Not}, _) $ t)) = rigid t - | ok4nnf (Const (@{const_name Trueprop},_) $ t) = rigid t +fun ok4nnf (Const (\<^const_name>\Trueprop\,_) $ (Const (\<^const_name>\Not\, _) $ t)) = rigid t + | ok4nnf (Const (\<^const_name>\Trueprop\,_) $ t) = rigid t | ok4nnf _ = false; fun make_nnf1 ctxt th = @@ -537,13 +537,12 @@ val nnf_ss = simpset_of (put_simpset HOL_basic_ss @{context} addsimps nnf_extra_simps - addsimprocs [@{simproc defined_All}, @{simproc defined_Ex}, @{simproc neq}, - @{simproc let_simp}]) + addsimprocs [\<^simproc>\defined_All\, \<^simproc>\defined_Ex\, \<^simproc>\neq\, \<^simproc>\let_simp\]) val presimplified_consts = - [@{const_name simp_implies}, @{const_name False}, @{const_name True}, - @{const_name Ex1}, @{const_name Ball}, @{const_name Bex}, @{const_name If}, - @{const_name Let}] + [\<^const_name>\simp_implies\, \<^const_name>\False\, \<^const_name>\True\, + \<^const_name>\Ex1\, \<^const_name>\Ball\, \<^const_name>\Bex\, \<^const_name>\If\, + \<^const_name>\Let\] fun presimplify ctxt = rewrite_rule ctxt (map safe_mk_meta_eq nnf_simps) @@ -563,7 +562,7 @@ fun skolemize_with_choice_theorems ctxt choice_ths = let fun aux th = - if not (has_conns [@{const_name Ex}] (Thm.prop_of th)) then + if not (has_conns [\<^const_name>\Ex\] (Thm.prop_of th)) then th else tryres (th, choice_ths @ @@ -604,7 +603,7 @@ end end in - if T = @{typ bool} then + if T = \<^typ>\bool\ then NONE else case pat t u of (SOME T, p as _ $ _) => SOME (Abs (Name.uu, T, p)) @@ -617,8 +616,8 @@ (* Strengthens "f g ~= f h" to "f g ~= f h & (EX x. g x ~= h x)". *) fun cong_extensionalize_thm ctxt th = (case Thm.concl_of th of - @{const Trueprop} $ (@{const Not} - $ (Const (@{const_name HOL.eq}, Type (_, [T, _])) + \<^const>\Trueprop\ $ (\<^const>\Not\ + $ (Const (\<^const_name>\HOL.eq\, Type (_, [T, _])) $ (t as _ $ _) $ (u as _ $ _))) => (case get_F_pattern T t u of SOME p => th RS infer_instantiate ctxt [(("F", 0), Thm.cterm_of ctxt p)] ext_cong_neq @@ -630,7 +629,7 @@ proof in "Tarski" that relies on the current behavior. *) fun abs_extensionalize_conv ctxt ct = (case Thm.term_of ct of - Const (@{const_name HOL.eq}, _) $ _ $ Abs _ => + Const (\<^const_name>\HOL.eq\, _) $ _ $ Abs _ => ct |> (Conv.rewr_conv @{thm fun_eq_iff [THEN eq_reflection]} then_conv abs_extensionalize_conv ctxt) | _ $ _ => Conv.comb_conv (abs_extensionalize_conv ctxt) ct diff -r d24dcac5eb4c -r e61557884799 src/HOL/Tools/Meson/meson_clausify.ML --- a/src/HOL/Tools/Meson/meson_clausify.ML Wed Dec 06 19:34:59 2017 +0100 +++ b/src/HOL/Tools/Meson/meson_clausify.ML Wed Dec 06 20:43:09 2017 +0100 @@ -35,8 +35,8 @@ (**** Transformation of Elimination Rules into First-Order Formulas****) -val cfalse = Thm.cterm_of @{theory_context HOL} @{term False}; -val ctp_false = Thm.cterm_of @{theory_context HOL} (HOLogic.mk_Trueprop @{term False}); +val cfalse = Thm.cterm_of \<^theory_context>\HOL\ \<^term>\False\; +val ctp_false = Thm.cterm_of \<^theory_context>\HOL\ (HOLogic.mk_Trueprop \<^term>\False\); (* Converts an elim-rule into an equivalent theorem that does not have the predicate variable. Leaves other theorems unchanged. We simply instantiate @@ -44,9 +44,9 @@ "Sledgehammer_Util".) *) fun transform_elim_theorem th = (case Thm.concl_of th of (*conclusion variable*) - @{const Trueprop} $ (Var (v as (_, @{typ bool}))) => + \<^const>\Trueprop\ $ (Var (v as (_, \<^typ>\bool\))) => Thm.instantiate ([], [(v, cfalse)]) th - | Var (v as (_, @{typ prop})) => + | Var (v as (_, \<^typ>\prop\)) => Thm.instantiate ([], [(v, ctp_false)]) th | _ => th) @@ -55,7 +55,7 @@ fun mk_old_skolem_term_wrapper t = let val T = fastype_of t in - Const (@{const_name Meson.skolem}, T --> T) $ t + Const (\<^const_name>\Meson.skolem\, T --> T) $ t end fun beta_eta_in_abs_body (Abs (s, T, t')) = Abs (s, T, beta_eta_in_abs_body t') @@ -64,7 +64,7 @@ (*Traverse a theorem, accumulating Skolem function definitions.*) fun old_skolem_defs th = let - fun dec_sko (Const (@{const_name Ex}, _) $ (body as Abs (_, T, p))) rhss = + fun dec_sko (Const (\<^const_name>\Ex\, _) $ (body as Abs (_, T, p))) rhss = (*Existential: declare a Skolem function, then insert into body and continue*) let val args = Misc_Legacy.term_frees body @@ -75,20 +75,20 @@ |> mk_old_skolem_term_wrapper val comb = list_comb (rhs, args) in dec_sko (subst_bound (comb, p)) (rhs :: rhss) end - | dec_sko (Const (@{const_name All},_) $ Abs (a, T, p)) rhss = + | dec_sko (Const (\<^const_name>\All\,_) $ Abs (a, T, p)) rhss = (*Universal quant: insert a free variable into body and continue*) let val fname = singleton (Name.variant_list (Misc_Legacy.add_term_names (p, []))) a in dec_sko (subst_bound (Free(fname,T), p)) rhss end - | dec_sko (@{const conj} $ p $ q) rhss = rhss |> dec_sko p |> dec_sko q - | dec_sko (@{const disj} $ p $ q) rhss = rhss |> dec_sko p |> dec_sko q - | dec_sko (@{const Trueprop} $ p) rhss = dec_sko p rhss + | dec_sko (\<^const>\conj\ $ p $ q) rhss = rhss |> dec_sko p |> dec_sko q + | dec_sko (\<^const>\disj\ $ p $ q) rhss = rhss |> dec_sko p |> dec_sko q + | dec_sko (\<^const>\Trueprop\ $ p) rhss = dec_sko p rhss | dec_sko _ rhss = rhss in dec_sko (Thm.prop_of th) [] end; (**** REPLACING ABSTRACTIONS BY COMBINATORS ****) -fun is_quasi_lambda_free (Const (@{const_name Meson.skolem}, _) $ _) = true +fun is_quasi_lambda_free (Const (\<^const_name>\Meson.skolem\, _) $ _) = true | is_quasi_lambda_free (t1 $ t2) = is_quasi_lambda_free t1 andalso is_quasi_lambda_free t2 | is_quasi_lambda_free (Abs _) = false @@ -98,7 +98,7 @@ fun abstract ctxt ct = let val Abs(x,_,body) = Thm.term_of ct - val Type (@{type_name fun}, [xT,bodyT]) = Thm.typ_of_cterm ct + val Type (\<^type_name>\fun\, [xT,bodyT]) = Thm.typ_of_cterm ct val cxT = Thm.ctyp_of ctxt xT val cbodyT = Thm.ctyp_of ctxt bodyT fun makeK () = @@ -196,7 +196,7 @@ val (hilbert, cabs) = ch |> Thm.dest_comb |>> Thm.term_of val T = case hilbert of - Const (_, Type (@{type_name fun}, [_, T])) => T + Const (_, Type (\<^type_name>\fun\, [_, T])) => T | _ => raise TERM ("old_skolem_theorem_of_def: expected \"Eps\"", [hilbert]) val cex = Thm.cterm_of ctxt (HOLogic.exists_const T) val ex_tm = Thm.apply cTrueprop (Thm.apply cex cabs) @@ -238,10 +238,10 @@ fun aux (cluster as (cluster_no, cluster_skolem)) index_no pos t = case t of (t1 as Const (s, _)) $ Abs (s', T, t') => - if s = @{const_name Pure.all} orelse s = @{const_name All} orelse - s = @{const_name Ex} then + if s = \<^const_name>\Pure.all\ orelse s = \<^const_name>\All\ orelse + s = \<^const_name>\Ex\ then let - val skolem = (pos = (s = @{const_name Ex})) + val skolem = (pos = (s = \<^const_name>\Ex\)) val (cluster, index_no) = if skolem = cluster_skolem then (cluster, index_no) else ((cluster_no ||> cluster_skolem ? Integer.add 1, skolem), 0) @@ -250,17 +250,17 @@ else t | (t1 as Const (s, _)) $ t2 $ t3 => - if s = @{const_name Pure.imp} orelse s = @{const_name HOL.implies} then + if s = \<^const_name>\Pure.imp\ orelse s = \<^const_name>\HOL.implies\ then t1 $ aux cluster index_no (not pos) t2 $ aux cluster index_no pos t3 - else if s = @{const_name HOL.conj} orelse - s = @{const_name HOL.disj} then + else if s = \<^const_name>\HOL.conj\ orelse + s = \<^const_name>\HOL.disj\ then t1 $ aux cluster index_no pos t2 $ aux cluster index_no pos t3 else t | (t1 as Const (s, _)) $ t2 => - if s = @{const_name Trueprop} then + if s = \<^const_name>\Trueprop\ then t1 $ aux cluster index_no pos t2 - else if s = @{const_name Not} then + else if s = \<^const_name>\Not\ then t1 $ aux cluster index_no (not pos) t2 else t @@ -271,28 +271,28 @@ ct |> (case Thm.term_of ct of Const (s, _) $ Abs (s', _, _) => - if s = @{const_name Pure.all} orelse s = @{const_name All} orelse - s = @{const_name Ex} then + if s = \<^const_name>\Pure.all\ orelse s = \<^const_name>\All\ orelse + s = \<^const_name>\Ex\ then Thm.dest_comb #> snd #> Thm.dest_abs (SOME s') #> snd #> zap pos else Conv.all_conv | Const (s, _) $ _ $ _ => - if s = @{const_name Pure.imp} orelse s = @{const_name implies} then + if s = \<^const_name>\Pure.imp\ orelse s = \<^const_name>\implies\ then Conv.combination_conv (Conv.arg_conv (zap (not pos))) (zap pos) - else if s = @{const_name conj} orelse s = @{const_name disj} then + else if s = \<^const_name>\conj\ orelse s = \<^const_name>\disj\ then Conv.combination_conv (Conv.arg_conv (zap pos)) (zap pos) else Conv.all_conv | Const (s, _) $ _ => - if s = @{const_name Trueprop} then Conv.arg_conv (zap pos) - else if s = @{const_name Not} then Conv.arg_conv (zap (not pos)) + if s = \<^const_name>\Trueprop\ then Conv.arg_conv (zap pos) + else if s = \<^const_name>\Not\ then Conv.arg_conv (zap (not pos)) else Conv.all_conv | _ => Conv.all_conv) fun ss_only ths ctxt = clear_simpset (put_simpset HOL_basic_ss ctxt) addsimps ths val cheat_choice = - @{prop "\x. \y. Q x y \ \f. \x. Q x (f x)"} + \<^prop>\\x. \y. Q x y \ \f. \x. Q x (f x)\ |> Logic.varify_global |> Skip_Proof.make_thm @{theory} @@ -374,7 +374,7 @@ th ctxt val (cnf_ths, ctxt) = clausify nnf_th fun intr_imp ct th = - Thm.instantiate ([], [((("i", 0), @{typ nat}), Thm.cterm_of ctxt (HOLogic.mk_nat ax_no))]) + Thm.instantiate ([], [((("i", 0), \<^typ>\nat\), Thm.cterm_of ctxt (HOLogic.mk_nat ax_no))]) (zero_var_indexes @{thm skolem_COMBK_D}) RS Thm.implies_intr ct th in diff -r d24dcac5eb4c -r e61557884799 src/HOL/Tools/Meson/meson_tactic.ML --- a/src/HOL/Tools/Meson/meson_tactic.ML Wed Dec 06 19:34:59 2017 +0100 +++ b/src/HOL/Tools/Meson/meson_tactic.ML Wed Dec 06 20:43:09 2017 +0100 @@ -19,7 +19,7 @@ val _ = Theory.setup - (Method.setup @{binding meson} (Attrib.thms >> (fn ths => fn ctxt => + (Method.setup \<^binding>\meson\ (Attrib.thms >> (fn ths => fn ctxt => SIMPLE_METHOD' (CHANGED_PROP o meson_general_tac ctxt ths))) "MESON resolution proof procedure") diff -r d24dcac5eb4c -r e61557884799 src/HOL/Tools/Quickcheck/abstract_generators.ML --- a/src/HOL/Tools/Quickcheck/abstract_generators.ML Wed Dec 06 19:34:59 2017 +0100 +++ b/src/HOL/Tools/Quickcheck/abstract_generators.ML Wed Dec 06 20:43:09 2017 +0100 @@ -45,14 +45,14 @@ in (descr, vs, [tyco], name, ([name], []), ([Type (tyco, map TFree vs)], [])) end fun ensure_sort (sort, instantiate) = Quickcheck_Common.ensure_sort - (((@{sort typerep}, @{sort term_of}), sort), (the_descr, instantiate)) + (((\<^sort>\typerep\, \<^sort>\term_of\), sort), (the_descr, instantiate)) Old_Datatype_Aux.default_config [tyco] in thy |> ensure_sort - (@{sort full_exhaustive}, Exhaustive_Generators.instantiate_full_exhaustive_datatype) - |> ensure_sort (@{sort exhaustive}, Exhaustive_Generators.instantiate_exhaustive_datatype) - |> ensure_sort (@{sort random}, Random_Generators.instantiate_random_datatype) + (\<^sort>\full_exhaustive\, Exhaustive_Generators.instantiate_full_exhaustive_datatype) + |> ensure_sort (\<^sort>\exhaustive\, Exhaustive_Generators.instantiate_exhaustive_datatype) + |> ensure_sort (\<^sort>\random\, Random_Generators.instantiate_random_datatype) |> (case opt_pred of NONE => I | SOME t => Context.theory_map (Quickcheck_Common.register_predicate (t, tyco))) end @@ -65,11 +65,11 @@ Syntax.read_term val _ = - Outer_Syntax.command @{command_keyword quickcheck_generator} + Outer_Syntax.command \<^command_keyword>\quickcheck_generator\ "define quickcheck generators for abstract types" ((Parse.type_const -- - Scan.option (Args.$$$ "predicate" |-- @{keyword ":"} |-- Parse.term)) -- - (Args.$$$ "constructors" |-- @{keyword ":"} |-- Parse.list1 Parse.term) + Scan.option (Args.$$$ "predicate" |-- \<^keyword>\:\ |-- Parse.term)) -- + (Args.$$$ "constructors" |-- \<^keyword>\:\ |-- Parse.list1 Parse.term) >> (fn ((tyco, opt_pred), constrs) => Toplevel.theory (quickcheck_generator_cmd tyco opt_pred constrs))) diff -r d24dcac5eb4c -r e61557884799 src/HOL/Tools/Quickcheck/exhaustive_generators.ML --- a/src/HOL/Tools/Quickcheck/exhaustive_generators.ML Wed Dec 06 19:34:59 2017 +0100 +++ b/src/HOL/Tools/Quickcheck/exhaustive_generators.ML Wed Dec 06 20:43:09 2017 +0100 @@ -38,30 +38,30 @@ (** dynamic options **) -val smart_quantifier = Attrib.setup_config_bool @{binding quickcheck_smart_quantifier} (K true) -val optimise_equality = Attrib.setup_config_bool @{binding quickcheck_optimise_equality} (K true) +val smart_quantifier = Attrib.setup_config_bool \<^binding>\quickcheck_smart_quantifier\ (K true) +val optimise_equality = Attrib.setup_config_bool \<^binding>\quickcheck_optimise_equality\ (K true) -val fast = Attrib.setup_config_bool @{binding quickcheck_fast} (K false) -val bounded_forall = Attrib.setup_config_bool @{binding quickcheck_bounded_forall} (K false) -val full_support = Attrib.setup_config_bool @{binding quickcheck_full_support} (K true) -val quickcheck_pretty = Attrib.setup_config_bool @{binding quickcheck_pretty} (K true) +val fast = Attrib.setup_config_bool \<^binding>\quickcheck_fast\ (K false) +val bounded_forall = Attrib.setup_config_bool \<^binding>\quickcheck_bounded_forall\ (K false) +val full_support = Attrib.setup_config_bool \<^binding>\quickcheck_full_support\ (K true) +val quickcheck_pretty = Attrib.setup_config_bool \<^binding>\quickcheck_pretty\ (K true) (** abstract syntax **) -fun termifyT T = HOLogic.mk_prodT (T, @{typ "unit \ Code_Evaluation.term"}) +fun termifyT T = HOLogic.mk_prodT (T, \<^typ>\unit \ Code_Evaluation.term\) -val size = @{term "i :: natural"} -val size_pred = @{term "(i :: natural) - 1"} -val size_ge_zero = @{term "(i :: natural) > 0"} +val size = \<^term>\i :: natural\ +val size_pred = \<^term>\(i :: natural) - 1\ +val size_ge_zero = \<^term>\(i :: natural) > 0\ fun mk_none_continuation (x, y) = - let val (T as Type (@{type_name option}, _)) = fastype_of x - in Const (@{const_name orelse}, T --> T --> T) $ x $ y end + let val (T as Type (\<^type_name>\option\, _)) = fastype_of x + in Const (\<^const_name>\orelse\, T --> T --> T) $ x $ y end fun mk_if (b, t, e) = let val T = fastype_of t - in Const (@{const_name If}, @{typ bool} --> T --> T --> T) $ b $ t $ e end + in Const (\<^const_name>\If\, \<^typ>\bool\ --> T --> T --> T) $ b $ t $ e end (* handling inductive datatypes *) @@ -72,19 +72,19 @@ exception Counterexample of term list -val resultT = @{typ "(bool * term list) option"} +val resultT = \<^typ>\(bool \ term list) option\ val exhaustiveN = "exhaustive" val full_exhaustiveN = "full_exhaustive" val bounded_forallN = "bounded_forall" -fun fast_exhaustiveT T = (T --> @{typ unit}) --> @{typ natural} --> @{typ unit} +fun fast_exhaustiveT T = (T --> \<^typ>\unit\) --> \<^typ>\natural\ --> \<^typ>\unit\ -fun exhaustiveT T = (T --> resultT) --> @{typ natural} --> resultT +fun exhaustiveT T = (T --> resultT) --> \<^typ>\natural\ --> resultT -fun bounded_forallT T = (T --> @{typ bool}) --> @{typ natural} --> @{typ bool} +fun bounded_forallT T = (T --> \<^typ>\bool\) --> \<^typ>\natural\ --> \<^typ>\bool\ -fun full_exhaustiveT T = (termifyT T --> resultT) --> @{typ natural} --> resultT +fun full_exhaustiveT T = (termifyT T --> resultT) --> \<^typ>\natural\ --> resultT fun check_allT T = (termifyT T --> resultT) --> resultT @@ -119,35 +119,35 @@ fun mk_equations functerms = let fun test_function T = Free ("f", T --> resultT) - val mk_call = gen_mk_call (fn T => Const (@{const_name exhaustive}, exhaustiveT T)) + val mk_call = gen_mk_call (fn T => Const (\<^const_name>\exhaustive\, exhaustiveT T)) val mk_aux_call = gen_mk_aux_call functerms val mk_consexpr = gen_mk_consexpr test_function fun mk_rhs exprs = - mk_if (size_ge_zero, foldr1 mk_none_continuation exprs, Const (@{const_name None}, resultT)) + mk_if (size_ge_zero, foldr1 mk_none_continuation exprs, Const (\<^const_name>\None\, resultT)) in mk_equation_terms (mk_call, mk_aux_call, mk_consexpr, mk_rhs, test_function, functerms) end fun mk_bounded_forall_equations functerms = let - fun test_function T = Free ("P", T --> @{typ bool}) - val mk_call = gen_mk_call (fn T => Const (@{const_name bounded_forall}, bounded_forallT T)) + fun test_function T = Free ("P", T --> \<^typ>\bool\) + val mk_call = gen_mk_call (fn T => Const (\<^const_name>\bounded_forall\, bounded_forallT T)) val mk_aux_call = gen_mk_aux_call functerms val mk_consexpr = gen_mk_consexpr test_function - fun mk_rhs exprs = mk_if (size_ge_zero, foldr1 HOLogic.mk_conj exprs, @{term True}) + fun mk_rhs exprs = mk_if (size_ge_zero, foldr1 HOLogic.mk_conj exprs, \<^term>\True\) in mk_equation_terms (mk_call, mk_aux_call, mk_consexpr, mk_rhs, test_function, functerms) end fun mk_full_equations functerms = let fun test_function T = Free ("f", termifyT T --> resultT) fun case_prod_const T = - HOLogic.case_prod_const (T, @{typ "unit \ Code_Evaluation.term"}, resultT) + HOLogic.case_prod_const (T, \<^typ>\unit \ Code_Evaluation.term\, resultT) fun mk_call T = let - val full_exhaustive = Const (@{const_name full_exhaustive}, full_exhaustiveT T) + val full_exhaustive = Const (\<^const_name>\full_exhaustive\, full_exhaustiveT T) in (T, fn t => full_exhaustive $ - (case_prod_const T $ absdummy T (absdummy @{typ "unit \ Code_Evaluation.term"} t)) $ + (case_prod_const T $ absdummy T (absdummy \<^typ>\unit \ Code_Evaluation.term\ t)) $ size_pred) end fun mk_aux_call fTs (k, _) (tyco, Ts) = @@ -158,7 +158,7 @@ (T, fn t => nth functerms k $ - (case_prod_const T $ absdummy T (absdummy @{typ "unit \ Code_Evaluation.term"} t)) $ + (case_prod_const T $ absdummy T (absdummy \<^typ>\unit \ Code_Evaluation.term\ t)) $ size_pred) end fun mk_consexpr simpleT (c, xs) = @@ -167,21 +167,21 @@ val constr = Const (c, Ts ---> simpleT) val bounds = map (fn x => Bound (2 * x + 1)) (((length xs) - 1) downto 0) val Eval_App = - Const (@{const_name Code_Evaluation.App}, + Const (\<^const_name>\Code_Evaluation.App\, HOLogic.termT --> HOLogic.termT --> HOLogic.termT) val Eval_Const = - Const (@{const_name Code_Evaluation.Const}, - HOLogic.literalT --> @{typ typerep} --> HOLogic.termT) + Const (\<^const_name>\Code_Evaluation.Const\, + HOLogic.literalT --> \<^typ>\typerep\ --> HOLogic.termT) val term = - fold (fn u => fn t => Eval_App $ t $ (u $ @{term "()"})) + fold (fn u => fn t => Eval_App $ t $ (u $ \<^term>\()\)) bounds (Eval_Const $ HOLogic.mk_literal c $ HOLogic.mk_typerep (Ts ---> simpleT)) val start_term = test_function simpleT $ - (HOLogic.pair_const simpleT @{typ "unit \ Code_Evaluation.term"} $ - (list_comb (constr, bounds)) $ absdummy @{typ unit} term) + (HOLogic.pair_const simpleT \<^typ>\unit \ Code_Evaluation.term\ $ + (list_comb (constr, bounds)) $ absdummy \<^typ>\unit\ term) in fold_rev (fn f => fn t => f t) fns start_term end fun mk_rhs exprs = - mk_if (size_ge_zero, foldr1 mk_none_continuation exprs, Const (@{const_name None}, resultT)) + mk_if (size_ge_zero, foldr1 mk_none_continuation exprs, Const (\<^const_name>\None\, resultT)) in mk_equation_terms (mk_call, mk_aux_call, mk_consexpr, mk_rhs, test_function, functerms) end @@ -212,17 +212,17 @@ val instantiate_bounded_forall_datatype = instantiate_datatype - ("bounded universal quantifiers", bounded_forallN, @{sort bounded_forall}, + ("bounded universal quantifiers", bounded_forallN, \<^sort>\bounded_forall\, mk_bounded_forall_equations, bounded_forallT, ["P", "i"]) val instantiate_exhaustive_datatype = instantiate_datatype - ("exhaustive generators", exhaustiveN, @{sort exhaustive}, + ("exhaustive generators", exhaustiveN, \<^sort>\exhaustive\, mk_equations, exhaustiveT, ["f", "i"]) val instantiate_full_exhaustive_datatype = instantiate_datatype - ("full exhaustive generators", full_exhaustiveN, @{sort full_exhaustive}, + ("full exhaustive generators", full_exhaustiveN, \<^sort>\full_exhaustive\, mk_full_equations, full_exhaustiveT, ["f", "i"]) @@ -230,15 +230,15 @@ fun mk_let_expr (x, t, e) genuine = let val (T1, T2) = (fastype_of x, fastype_of (e genuine)) - in Const (@{const_name Let}, T1 --> (T1 --> T2) --> T2) $ t $ lambda x (e genuine) end + in Const (\<^const_name>\Let\, T1 --> (T1 --> T2) --> T2) $ t $ lambda x (e genuine) end fun mk_safe_let_expr genuine_only none safe (x, t, e) genuine = let val (T1, T2) = (fastype_of x, fastype_of (e genuine)) - val if_t = Const (@{const_name If}, @{typ bool} --> T2 --> T2 --> T2) + val if_t = Const (\<^const_name>\If\, \<^typ>\bool\ --> T2 --> T2 --> T2) in - Const (@{const_name Quickcheck_Random.catch_match}, T2 --> T2 --> T2) $ - (Const (@{const_name Let}, T1 --> (T1 --> T2) --> T2) $ t $ lambda x (e genuine)) $ + Const (\<^const_name>\Quickcheck_Random.catch_match\, T2 --> T2 --> T2) $ + (Const (\<^const_name>\Let\, T1 --> (T1 --> T2) --> T2) $ t $ lambda x (e genuine)) $ (if_t $ genuine_only $ none $ safe false) end @@ -323,30 +323,30 @@ val frees = map Free (Term.add_frees t []) fun lookup v = the (AList.lookup (op =) (names ~~ frees) v) val ([depth_name], _) = Variable.variant_fixes ["depth"] ctxt' - val depth = Free (depth_name, @{typ natural}) + val depth = Free (depth_name, \<^typ>\natural\) fun return _ = - @{term "throw_Counterexample :: term list \ unit"} $ - (HOLogic.mk_list @{typ term} + \<^term>\throw_Counterexample :: term list \ unit\ $ + (HOLogic.mk_list \<^typ>\term\ (map (fn t => HOLogic.mk_term_of (fastype_of t) t) (frees @ eval_terms))) fun mk_exhaustive_closure (free as Free (_, T)) t = - Const (@{const_name fast_exhaustive}, fast_exhaustiveT T) $ lambda free t $ depth - val none_t = @{term "()"} + Const (\<^const_name>\fast_exhaustive\, fast_exhaustiveT T) $ lambda free t $ depth + val none_t = \<^term>\()\ fun mk_safe_if (cond, then_t, else_t) genuine = mk_if (cond, then_t, else_t genuine) fun mk_let _ def v_opt t e = mk_let_expr (the_default def v_opt, t, e) val mk_test_term = mk_test_term lookup mk_exhaustive_closure mk_safe_if mk_let none_t return ctxt - in lambda depth (@{term "catch_Counterexample :: unit => term list option"} $ mk_test_term t) end + in lambda depth (\<^term>\catch_Counterexample :: unit => term list option\ $ mk_test_term t) end fun mk_unknown_term T = - HOLogic.reflect_term (Const (@{const_name unknown}, T)) + HOLogic.reflect_term (Const (\<^const_name>\unknown\, T)) fun mk_safe_term t = - @{term "Quickcheck_Random.catch_match :: term \ term \ term"} $ + \<^term>\Quickcheck_Random.catch_match :: term \ term \ term\ $ (HOLogic.mk_term_of (fastype_of t) t) $ mk_unknown_term (fastype_of t) fun mk_return t genuine = - @{term "Some :: bool \ term list \ (bool \ term list) option"} $ - (HOLogic.pair_const @{typ bool} @{typ "term list"} $ + \<^term>\Some :: bool \ term list \ (bool \ term list) option\ $ + (HOLogic.pair_const \<^typ>\bool\ \<^typ>\term list\ $ Quickcheck_Common.reflect_bool genuine $ t) fun mk_generator_expr ctxt (t, eval_terms) = @@ -357,14 +357,14 @@ fun lookup v = the (AList.lookup (op =) (names ~~ frees) v) val ([depth_name, genuine_only_name], _) = Variable.variant_fixes ["depth", "genuine_only"] ctxt' - val depth = Free (depth_name, @{typ natural}) - val genuine_only = Free (genuine_only_name, @{typ bool}) + val depth = Free (depth_name, \<^typ>\natural\) + val genuine_only = Free (genuine_only_name, \<^typ>\bool\) val return = - mk_return (HOLogic.mk_list @{typ term} + mk_return (HOLogic.mk_list \<^typ>\term\ (map (fn t => HOLogic.mk_term_of (fastype_of t) t) frees @ map mk_safe_term eval_terms)) fun mk_exhaustive_closure (free as Free (_, T)) t = - Const (@{const_name exhaustive}, exhaustiveT T) $ lambda free t $ depth - val none_t = Const (@{const_name None}, resultT) + Const (\<^const_name>\exhaustive\, exhaustiveT T) $ lambda free t $ depth + val none_t = Const (\<^const_name>\None\, resultT) val mk_if = Quickcheck_Common.mk_safe_if genuine_only none_t fun mk_let safe def v_opt t e = mk_safe_let_expr genuine_only none_t safe (the_default def v_opt, t, e) @@ -380,28 +380,28 @@ val ([depth_name, genuine_only_name], ctxt'') = Variable.variant_fixes ["depth", "genuine_only"] ctxt' val (term_names, _) = Variable.variant_fixes (map (prefix "t_") names) ctxt'' - val depth = Free (depth_name, @{typ natural}) - val genuine_only = Free (genuine_only_name, @{typ bool}) - val term_vars = map (fn n => Free (n, @{typ "unit \ term"})) term_names + val depth = Free (depth_name, \<^typ>\natural\) + val genuine_only = Free (genuine_only_name, \<^typ>\bool\) + val term_vars = map (fn n => Free (n, \<^typ>\unit \ term\)) term_names fun lookup v = the (AList.lookup (op =) (names ~~ (frees ~~ term_vars)) v) val return = mk_return - (HOLogic.mk_list @{typ term} - (map (fn v => v $ @{term "()"}) term_vars @ map mk_safe_term eval_terms)) + (HOLogic.mk_list \<^typ>\term\ + (map (fn v => v $ \<^term>\()\) term_vars @ map mk_safe_term eval_terms)) fun mk_exhaustive_closure (free as Free (_, T), term_var) t = - if Sign.of_sort thy (T, @{sort check_all}) then - Const (@{const_name check_all}, check_allT T) $ - (HOLogic.case_prod_const (T, @{typ "unit \ term"}, resultT) $ + if Sign.of_sort thy (T, \<^sort>\check_all\) then + Const (\<^const_name>\check_all\, check_allT T) $ + (HOLogic.case_prod_const (T, \<^typ>\unit \ term\, resultT) $ lambda free (lambda term_var t)) else - Const (@{const_name full_exhaustive}, full_exhaustiveT T) $ - (HOLogic.case_prod_const (T, @{typ "unit \ term"}, resultT) $ + Const (\<^const_name>\full_exhaustive\, full_exhaustiveT T) $ + (HOLogic.case_prod_const (T, \<^typ>\unit \ term\, resultT) $ lambda free (lambda term_var t)) $ depth - val none_t = Const (@{const_name None}, resultT) + val none_t = Const (\<^const_name>\None\, resultT) val mk_if = Quickcheck_Common.mk_safe_if genuine_only none_t fun mk_let safe _ (SOME (v, term_var)) t e = mk_safe_let_expr genuine_only none_t safe (v, t, - e #> subst_free [(term_var, absdummy @{typ unit} (mk_safe_term t))]) + e #> subst_free [(term_var, absdummy \<^typ>\unit\ (mk_safe_term t))]) | mk_let safe v NONE t e = mk_safe_let_expr genuine_only none_t safe (v, t, e) val mk_test_term = mk_test_term lookup mk_exhaustive_closure mk_if mk_let none_t return ctxt in lambda genuine_only (lambda depth (mk_test_term t)) end @@ -409,37 +409,37 @@ fun mk_parametric_generator_expr mk_generator_expr = Quickcheck_Common.gen_mk_parametric_generator_expr ((mk_generator_expr, - absdummy @{typ bool} (absdummy @{typ natural} (Const (@{const_name None}, resultT)))), - @{typ bool} --> @{typ natural} --> resultT) + absdummy \<^typ>\bool\ (absdummy \<^typ>\natural\ (Const (\<^const_name>\None\, resultT)))), + \<^typ>\bool\ --> \<^typ>\natural\ --> resultT) fun mk_validator_expr ctxt t = let - fun bounded_forallT T = (T --> @{typ bool}) --> @{typ natural} --> @{typ bool} + fun bounded_forallT T = (T --> \<^typ>\bool\) --> \<^typ>\natural\ --> \<^typ>\bool\ val ctxt' = Variable.auto_fixes t ctxt val names = Term.add_free_names t [] val frees = map Free (Term.add_frees t []) fun lookup v = the (AList.lookup (op =) (names ~~ frees) v) val ([depth_name], _) = Variable.variant_fixes ["depth"] ctxt' - val depth = Free (depth_name, @{typ natural}) + val depth = Free (depth_name, \<^typ>\natural\) fun mk_bounded_forall (Free (s, T)) t = - Const (@{const_name bounded_forall}, bounded_forallT T) $ lambda (Free (s, T)) t $ depth + Const (\<^const_name>\bounded_forall\, bounded_forallT T) $ lambda (Free (s, T)) t $ depth fun mk_safe_if (cond, then_t, else_t) genuine = mk_if (cond, then_t, else_t genuine) fun mk_let _ def v_opt t e = mk_let_expr (the_default def v_opt, t, e) val mk_test_term = - mk_test_term lookup mk_bounded_forall mk_safe_if mk_let @{term True} (K @{term False}) ctxt + mk_test_term lookup mk_bounded_forall mk_safe_if mk_let \<^term>\True\ (K \<^term>\False\) ctxt in lambda depth (mk_test_term t) end fun mk_bounded_forall_generator_expr ctxt (t, eval_terms) = let val frees = Term.add_free_names t [] val dummy_term = - @{term "Code_Evaluation.Const (STR ''Pure.dummy_pattern'') (Typerep.Typerep (STR ''dummy'') [])"} + \<^term>\Code_Evaluation.Const (STR ''Pure.dummy_pattern'') (Typerep.Typerep (STR ''dummy'') [])\ val return = - @{term "Some :: term list => term list option"} $ - (HOLogic.mk_list @{typ term} (replicate (length frees + length eval_terms) dummy_term)) - val wrap = absdummy @{typ bool} - (@{term "If :: bool \ term list option \ term list option \ term list option"} $ - Bound 0 $ @{term "None :: term list option"} $ return) + \<^term>\Some :: term list => term list option\ $ + (HOLogic.mk_list \<^typ>\term\ (replicate (length frees + length eval_terms) dummy_term)) + val wrap = absdummy \<^typ>\bool\ + (\<^term>\If :: bool \ term list option \ term list option \ term list option\ $ + Bound 0 $ \<^term>\None :: term list option\ $ return) in HOLogic.mk_comp (wrap, mk_validator_expr ctxt t) end @@ -504,7 +504,7 @@ (get_counterexample_batch, put_counterexample_batch, "Exhaustive_Generators.put_counterexample_batch") ctxt (SOME target) (fn proc => map (fn g => g #> (Option.map o map) proc)) - (HOLogic.mk_list @{typ "natural => term list option"} ts') [] + (HOLogic.mk_list \<^typ>\natural \ term list option\ ts') [] in map (fn compile => fn size => compile size |> (Option.map o map) Quickcheck_Common.post_process_term) compiles @@ -518,14 +518,14 @@ let val ts' = map (mk_validator_expr ctxt) ts in Code_Runtime.dynamic_value_strict (get_validator_batch, put_validator_batch, "Exhaustive_Generators.put_validator_batch") - ctxt (SOME target) (K I) (HOLogic.mk_list @{typ "natural \ bool"} ts') [] + ctxt (SOME target) (K I) (HOLogic.mk_list \<^typ>\natural \ bool\ ts') [] end fun compile_validator_exprs ctxt ts = compile_validator_exprs_raw ctxt ts |> map (fn f => fn size => f (Code_Numeral.natural_of_integer size)) -fun size_matters_for thy Ts = not (forall (fn T => Sign.of_sort thy (T, @{sort check_all})) Ts) +fun size_matters_for thy Ts = not (forall (fn T => Sign.of_sort thy (T, \<^sort>\check_all\)) Ts) val test_goals = Quickcheck_Common.generator_test_goal_terms @@ -535,22 +535,22 @@ (* setup *) val setup_exhaustive_datatype_interpretation = - Quickcheck_Common.datatype_interpretation @{plugin quickcheck_exhaustive} - (@{sort exhaustive}, instantiate_exhaustive_datatype) + Quickcheck_Common.datatype_interpretation \<^plugin>\quickcheck_exhaustive\ + (\<^sort>\exhaustive\, instantiate_exhaustive_datatype) val setup_bounded_forall_datatype_interpretation = - BNF_LFP_Compat.interpretation @{plugin quickcheck_bounded_forall} Quickcheck_Common.compat_prefs + BNF_LFP_Compat.interpretation \<^plugin>\quickcheck_bounded_forall\ Quickcheck_Common.compat_prefs (Quickcheck_Common.ensure_sort - (((@{sort type}, @{sort type}), @{sort bounded_forall}), + (((\<^sort>\type\, \<^sort>\type\), \<^sort>\bounded_forall\), (fn thy => BNF_LFP_Compat.the_descr thy Quickcheck_Common.compat_prefs, instantiate_bounded_forall_datatype))) -val active = Attrib.setup_config_bool @{binding quickcheck_exhaustive_active} (K true) +val active = Attrib.setup_config_bool \<^binding>\quickcheck_exhaustive_active\ (K true) val _ = Theory.setup - (Quickcheck_Common.datatype_interpretation @{plugin quickcheck_full_exhaustive} - (@{sort full_exhaustive}, instantiate_full_exhaustive_datatype) + (Quickcheck_Common.datatype_interpretation \<^plugin>\quickcheck_full_exhaustive\ + (\<^sort>\full_exhaustive\, instantiate_full_exhaustive_datatype) #> Context.theory_map (Quickcheck.add_tester ("exhaustive", (active, test_goals))) #> Context.theory_map (Quickcheck.add_batch_generator ("exhaustive", compile_generator_exprs)) #> Context.theory_map (Quickcheck.add_batch_validator ("exhaustive", compile_validator_exprs))) diff -r d24dcac5eb4c -r e61557884799 src/HOL/Tools/Quickcheck/find_unused_assms.ML --- a/src/HOL/Tools/Quickcheck/find_unused_assms.ML Wed Dec 06 19:34:59 2017 +0100 +++ b/src/HOL/Tools/Quickcheck/find_unused_assms.ML Wed Dec 06 20:43:09 2017 +0100 @@ -113,7 +113,7 @@ end val _ = - Outer_Syntax.command @{command_keyword find_unused_assms} + Outer_Syntax.command \<^command_keyword>\find_unused_assms\ "find theorems with (potentially) superfluous assumptions" (Scan.option Parse.name >> (fn name => Toplevel.keep (fn state => print_unused_assms (Toplevel.context_of state) name))) diff -r d24dcac5eb4c -r e61557884799 src/HOL/Tools/Quickcheck/narrowing_generators.ML --- a/src/HOL/Tools/Quickcheck/narrowing_generators.ML Wed Dec 06 19:34:59 2017 +0100 +++ b/src/HOL/Tools/Quickcheck/narrowing_generators.ML Wed Dec 06 20:43:09 2017 +0100 @@ -24,37 +24,37 @@ (* configurations *) -val allow_existentials = Attrib.setup_config_bool @{binding quickcheck_allow_existentials} (K true) -val finite_functions = Attrib.setup_config_bool @{binding quickcheck_finite_functions} (K true) -val overlord = Attrib.setup_config_bool @{binding quickcheck_narrowing_overlord} (K false) -val ghc_options = Attrib.setup_config_string @{binding quickcheck_narrowing_ghc_options} (K "") +val allow_existentials = Attrib.setup_config_bool \<^binding>\quickcheck_allow_existentials\ (K true) +val finite_functions = Attrib.setup_config_bool \<^binding>\quickcheck_finite_functions\ (K true) +val overlord = Attrib.setup_config_bool \<^binding>\quickcheck_narrowing_overlord\ (K false) +val ghc_options = Attrib.setup_config_string \<^binding>\quickcheck_narrowing_ghc_options\ (K "") (* partial_term_of instances *) fun mk_partial_term_of (x, T) = - Const (@{const_name Quickcheck_Narrowing.partial_term_of_class.partial_term_of}, - Term.itselfT T --> @{typ narrowing_term} --> @{typ Code_Evaluation.term}) $ Logic.mk_type T $ x + Const (\<^const_name>\Quickcheck_Narrowing.partial_term_of_class.partial_term_of\, + Term.itselfT T --> \<^typ>\narrowing_term\ --> \<^typ>\Code_Evaluation.term\) $ Logic.mk_type T $ x (** formal definition **) fun add_partial_term_of tyco raw_vs thy = let - val vs = map (fn (v, _) => (v, @{sort typerep})) raw_vs + val vs = map (fn (v, _) => (v, \<^sort>\typerep\)) raw_vs val ty = Type (tyco, map TFree vs) val lhs = - Const (@{const_name partial_term_of}, - Term.itselfT ty --> @{typ narrowing_term} --> @{typ Code_Evaluation.term}) $ - Free ("x", Term.itselfT ty) $ Free ("t", @{typ narrowing_term}) - val rhs = @{term "undefined :: Code_Evaluation.term"} + Const (\<^const_name>\partial_term_of\, + Term.itselfT ty --> \<^typ>\narrowing_term\ --> \<^typ>\Code_Evaluation.term\) $ + Free ("x", Term.itselfT ty) $ Free ("t", \<^typ>\narrowing_term\) + val rhs = \<^term>\undefined :: Code_Evaluation.term\ val eq = HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs)) fun triv_name_of t = (fst o dest_Free o fst o strip_comb o fst o HOLogic.dest_eq o HOLogic.dest_Trueprop) t ^ "_triv" in thy - |> Class.instantiation ([tyco], vs, @{sort partial_term_of}) + |> Class.instantiation ([tyco], vs, \<^sort>\partial_term_of\) |> `(fn lthy => Syntax.check_term lthy eq) |-> (fn eq => Specification.definition NONE [] [] ((Binding.name (triv_name_of eq), []), eq)) |> snd @@ -63,8 +63,8 @@ fun ensure_partial_term_of (tyco, (raw_vs, _)) thy = let - val need_inst = not (Sorts.has_instance (Sign.classes_of thy) tyco @{sort partial_term_of}) - andalso Sorts.has_instance (Sign.classes_of thy) tyco @{sort typerep} + val need_inst = not (Sorts.has_instance (Sign.classes_of thy) tyco \<^sort>\partial_term_of\) + andalso Sorts.has_instance (Sign.classes_of thy) tyco \<^sort>\typerep\ in if need_inst then add_partial_term_of tyco raw_vs thy else thy end @@ -72,14 +72,14 @@ fun mk_partial_term_of_eq thy ty (i, (c, (_, tys))) = let - val frees = map Free (Name.invent_names Name.context "a" (map (K @{typ narrowing_term}) tys)) + val frees = map Free (Name.invent_names Name.context "a" (map (K \<^typ>\narrowing_term\) tys)) val narrowing_term = - @{term Quickcheck_Narrowing.Narrowing_constructor} $ HOLogic.mk_number @{typ integer} i $ - HOLogic.mk_list @{typ narrowing_term} (rev frees) + \<^term>\Quickcheck_Narrowing.Narrowing_constructor\ $ HOLogic.mk_number \<^typ>\integer\ i $ + HOLogic.mk_list \<^typ>\narrowing_term\ (rev frees) val rhs = - fold (fn u => fn t => @{term "Code_Evaluation.App"} $ t $ u) + fold (fn u => fn t => \<^term>\Code_Evaluation.App\ $ t $ u) (map mk_partial_term_of (frees ~~ tys)) - (@{term "Code_Evaluation.Const"} $ HOLogic.mk_literal c $ HOLogic.mk_typerep (tys ---> ty)) + (\<^term>\Code_Evaluation.Const\ $ HOLogic.mk_literal c $ HOLogic.mk_typerep (tys ---> ty)) val insts = map (SOME o Thm.global_cterm_of thy o Logic.unvarify_types_global o Logic.varify_global) [Free ("ty", Term.itselfT ty), narrowing_term, rhs] @@ -93,16 +93,16 @@ fun add_partial_term_of_code tyco raw_vs raw_cs thy = let val algebra = Sign.classes_of thy - val vs = map (fn (v, sort) => (v, curry (Sorts.inter_sort algebra) @{sort typerep} sort)) raw_vs + val vs = map (fn (v, sort) => (v, curry (Sorts.inter_sort algebra) \<^sort>\typerep\ sort)) raw_vs val ty = Type (tyco, map TFree vs) val cs = (map o apsnd o apsnd o map o map_atyps) (fn TFree (v, _) => TFree (v, (the o AList.lookup (op =) vs) v)) raw_cs - val const = Axclass.param_of_inst thy (@{const_name partial_term_of}, tyco) + val const = Axclass.param_of_inst thy (\<^const_name>\partial_term_of\, tyco) val var_insts = map (SOME o Thm.global_cterm_of thy o Logic.unvarify_types_global o Logic.varify_global) - [Free ("ty", Term.itselfT ty), @{term "Quickcheck_Narrowing.Narrowing_variable p tt"}, - @{term "Code_Evaluation.Free (STR ''_'')"} $ HOLogic.mk_typerep ty] + [Free ("ty", Term.itselfT ty), \<^term>\Quickcheck_Narrowing.Narrowing_variable p tt\, + \<^term>\Code_Evaluation.Free (STR ''_'')\ $ HOLogic.mk_typerep ty] val var_eq = @{thm partial_term_of_anything} |> Thm.instantiate' [SOME (Thm.global_ctyp_of thy ty)] var_insts @@ -114,7 +114,7 @@ end fun ensure_partial_term_of_code (tyco, (raw_vs, cs)) thy = - let val has_inst = Sorts.has_instance (Sign.classes_of thy) tyco @{sort partial_term_of} + let val has_inst = Sorts.has_instance (Sign.classes_of thy) tyco \<^sort>\partial_term_of\ in if has_inst then add_partial_term_of_code tyco raw_vs cs thy else thy end @@ -126,21 +126,21 @@ val narrowingN = "narrowing" -fun narrowingT T = @{typ integer} --> Type (@{type_name Quickcheck_Narrowing.narrowing_cons}, [T]) +fun narrowingT T = \<^typ>\integer\ --> Type (\<^type_name>\Quickcheck_Narrowing.narrowing_cons\, [T]) -fun mk_cons c T = Const (@{const_name Quickcheck_Narrowing.cons}, T --> narrowingT T) $ Const (c, T) +fun mk_cons c T = Const (\<^const_name>\Quickcheck_Narrowing.cons\, T --> narrowingT T) $ Const (c, T) fun mk_apply (T, t) (U, u) = let val (_, U') = dest_funT U in - (U', Const (@{const_name Quickcheck_Narrowing.apply}, + (U', Const (\<^const_name>\Quickcheck_Narrowing.apply\, narrowingT U --> narrowingT T --> narrowingT U') $ u $ t) end fun mk_sum (t, u) = let val T = fastype_of t - in Const (@{const_name Quickcheck_Narrowing.sum}, T --> T --> T) $ t $ u end + in Const (\<^const_name>\Quickcheck_Narrowing.sum\, T --> T --> T) $ t $ u end (** deriving narrowing instances **) @@ -148,7 +148,7 @@ fun mk_equations descr vs narrowings = let fun mk_call T = - (T, Const (@{const_name "Quickcheck_Narrowing.narrowing_class.narrowing"}, narrowingT T)) + (T, Const (\<^const_name>\Quickcheck_Narrowing.narrowing_class.narrowing\, narrowingT T)) fun mk_aux_call fTs (k, _) (tyco, Ts) = let val T = Type (tyco, Ts) @@ -181,7 +181,7 @@ in if not (contains_recursive_type_under_function_types descr) then thy - |> Class.instantiation (tycos, vs, @{sort narrowing}) + |> Class.instantiation (tycos, vs, \<^sort>\narrowing\) |> Quickcheck_Common.define_functions (fn narrowings => mk_equations descr vs narrowings, NONE) prfx [] narrowingsN (map narrowingT (Ts @ Us)) @@ -348,29 +348,29 @@ let val (names, boundTs) = split_list xTs fun mk_eval_ffun dT rT = - Const (@{const_name "Quickcheck_Narrowing.eval_ffun"}, - Type (@{type_name "Quickcheck_Narrowing.ffun"}, [dT, rT]) --> dT --> rT) + Const (\<^const_name>\Quickcheck_Narrowing.eval_ffun\, + Type (\<^type_name>\Quickcheck_Narrowing.ffun\, [dT, rT]) --> dT --> rT) fun mk_eval_cfun dT rT = - Const (@{const_name "Quickcheck_Narrowing.eval_cfun"}, - Type (@{type_name "Quickcheck_Narrowing.cfun"}, [rT]) --> dT --> rT) - fun eval_function (Type (@{type_name fun}, [dT, rT])) = + Const (\<^const_name>\Quickcheck_Narrowing.eval_cfun\, + Type (\<^type_name>\Quickcheck_Narrowing.cfun\, [rT]) --> dT --> rT) + fun eval_function (Type (\<^type_name>\fun\, [dT, rT])) = let val (rt', rT') = eval_function rT in (case dT of - Type (@{type_name fun}, _) => + Type (\<^type_name>\fun\, _) => (fn t => absdummy dT (rt' (mk_eval_cfun dT rT' $ incr_boundvars 1 t $ Bound 0)), - Type (@{type_name "Quickcheck_Narrowing.cfun"}, [rT'])) + Type (\<^type_name>\Quickcheck_Narrowing.cfun\, [rT'])) | _ => (fn t => absdummy dT (rt' (mk_eval_ffun dT rT' $ incr_boundvars 1 t $ Bound 0)), - Type (@{type_name "Quickcheck_Narrowing.ffun"}, [dT, rT']))) + Type (\<^type_name>\Quickcheck_Narrowing.ffun\, [dT, rT']))) end - | eval_function (T as Type (@{type_name prod}, [fT, sT])) = + | eval_function (T as Type (\<^type_name>\prod\, [fT, sT])) = let val (ft', fT') = eval_function fT val (st', sT') = eval_function sT - val T' = Type (@{type_name prod}, [fT', sT']) - val map_const = Const (@{const_name map_prod}, (fT' --> fT) --> (sT' --> sT) --> T' --> T) + val T' = Type (\<^type_name>\prod\, [fT', sT']) + val map_const = Const (\<^const_name>\map_prod\, (fT' --> fT) --> (sT' --> sT) --> T' --> T) fun apply_dummy T t = absdummy T (t (Bound 0)) in (fn t => list_comb (map_const, [apply_dummy fT' ft', apply_dummy sT' st', t]), T') @@ -382,11 +382,11 @@ (names ~~ boundTs', t') end -fun dest_ffun (Type (@{type_name "Quickcheck_Narrowing.ffun"}, [dT, rT])) = (dT, rT) +fun dest_ffun (Type (\<^type_name>\Quickcheck_Narrowing.ffun\, [dT, rT])) = (dT, rT) -fun eval_finite_functions (Const (@{const_name "Quickcheck_Narrowing.ffun.Constant"}, T) $ value) = +fun eval_finite_functions (Const (\<^const_name>\Quickcheck_Narrowing.ffun.Constant\, T) $ value) = absdummy (fst (dest_ffun (body_type T))) (eval_finite_functions value) - | eval_finite_functions (Const (@{const_name "Quickcheck_Narrowing.ffun.Update"}, T) $ a $ b $ f) = + | eval_finite_functions (Const (\<^const_name>\Quickcheck_Narrowing.ffun.Update\, T) $ a $ b $ f) = let val (T1, T2) = dest_ffun (body_type T) in @@ -407,29 +407,29 @@ fun make_pnf_term thy t = Pattern.rewrite_term thy rewrs [] t -fun strip_quantifiers (Const (@{const_name Ex}, _) $ Abs (x, T, t)) = - apfst (cons (@{const_name Ex}, (x, T))) (strip_quantifiers t) - | strip_quantifiers (Const (@{const_name All}, _) $ Abs (x, T, t)) = - apfst (cons (@{const_name All}, (x, T))) (strip_quantifiers t) +fun strip_quantifiers (Const (\<^const_name>\Ex\, _) $ Abs (x, T, t)) = + apfst (cons (\<^const_name>\Ex\, (x, T))) (strip_quantifiers t) + | strip_quantifiers (Const (\<^const_name>\All\, _) $ Abs (x, T, t)) = + apfst (cons (\<^const_name>\All\, (x, T))) (strip_quantifiers t) | strip_quantifiers t = ([], t) fun contains_existentials t = - exists (fn (Q, _) => Q = @{const_name Ex}) (fst (strip_quantifiers t)) + exists (fn (Q, _) => Q = \<^const_name>\Ex\) (fst (strip_quantifiers t)) fun mk_property qs t = let - fun enclose (@{const_name Ex}, (x, T)) t = - Const (@{const_name Quickcheck_Narrowing.exists}, - (T --> @{typ property}) --> @{typ property}) $ Abs (x, T, t) - | enclose (@{const_name All}, (x, T)) t = - Const (@{const_name Quickcheck_Narrowing.all}, - (T --> @{typ property}) --> @{typ property}) $ Abs (x, T, t) - in fold_rev enclose qs (@{term Quickcheck_Narrowing.Property} $ t) end + fun enclose (\<^const_name>\Ex\, (x, T)) t = + Const (\<^const_name>\Quickcheck_Narrowing.exists\, + (T --> \<^typ>\property\) --> \<^typ>\property\) $ Abs (x, T, t) + | enclose (\<^const_name>\All\, (x, T)) t = + Const (\<^const_name>\Quickcheck_Narrowing.all\, + (T --> \<^typ>\property\) --> \<^typ>\property\) $ Abs (x, T, t) + in fold_rev enclose qs (\<^term>\Quickcheck_Narrowing.Property\ $ t) end -fun mk_case_term ctxt p ((@{const_name Ex}, (x, T)) :: qs') (Existential_Counterexample cs) = +fun mk_case_term ctxt p ((\<^const_name>\Ex\, (x, T)) :: qs') (Existential_Counterexample cs) = Case_Translation.make_case ctxt Case_Translation.Quiet Name.context (Free (x, T)) (map (fn (t, c) => (t, mk_case_term ctxt (p - 1) qs' c)) cs) - | mk_case_term ctxt p ((@{const_name All}, _) :: qs') (Universal_Counterexample (t, c)) = + | mk_case_term ctxt p ((\<^const_name>\All\, _) :: qs') (Universal_Counterexample (t, c)) = if p = 0 then t else mk_case_term ctxt (p - 1) qs' c val post_process = @@ -437,7 +437,7 @@ fun mk_terms ctxt qs result = let - val ps = filter (fn (_, (@{const_name All}, _)) => true | _ => false) (map_index I qs) + val ps = filter (fn (_, (\<^const_name>\All\, _)) => true | _ => false) (map_index I qs) in map (fn (p, (_, (x, _))) => (x, mk_case_term ctxt p qs result)) ps |> map (apsnd post_process) @@ -485,7 +485,7 @@ fun wrap f t = uncurry (fold_rev Term.abs) (f (strip_abs t)) val finitize = if Config.get ctxt finite_functions then wrap finitize_functions else I fun ensure_testable t = - Const (@{const_name Quickcheck_Narrowing.ensure_testable}, + Const (\<^const_name>\Quickcheck_Narrowing.ensure_testable\, fastype_of t --> fastype_of t) $ t fun is_genuine (SOME (true, _)) = true | is_genuine _ = false @@ -531,14 +531,14 @@ (* setup *) -val active = Attrib.setup_config_bool @{binding quickcheck_narrowing_active} (K false) +val active = Attrib.setup_config_bool \<^binding>\quickcheck_narrowing_active\ (K false) val _ = Theory.setup (Code.datatype_interpretation ensure_partial_term_of #> Code.datatype_interpretation ensure_partial_term_of_code - #> Quickcheck_Common.datatype_interpretation @{plugin quickcheck_narrowing} - (@{sort narrowing}, instantiate_narrowing_datatype) + #> Quickcheck_Common.datatype_interpretation \<^plugin>\quickcheck_narrowing\ + (\<^sort>\narrowing\, instantiate_narrowing_datatype) #> Context.theory_map (Quickcheck.add_tester ("narrowing", (active, test_goals)))) end diff -r d24dcac5eb4c -r e61557884799 src/HOL/Tools/Quickcheck/quickcheck_common.ML --- a/src/HOL/Tools/Quickcheck/quickcheck_common.ML Wed Dec 06 19:34:59 2017 +0100 +++ b/src/HOL/Tools/Quickcheck/quickcheck_common.ML Wed Dec 06 20:43:09 2017 +0100 @@ -56,12 +56,12 @@ (* HOLogic's term functions *) -fun strip_imp (Const(@{const_name HOL.implies}, _) $ A $ B) = apfst (cons A) (strip_imp B) +fun strip_imp (Const(\<^const_name>\HOL.implies\, _) $ A $ B) = apfst (cons A) (strip_imp B) | strip_imp A = ([], A) -fun reflect_bool b = if b then @{term True} else @{term False} +fun reflect_bool b = if b then \<^term>\True\ else \<^term>\False\ -fun mk_undefined T = Const (@{const_name undefined}, T) +fun mk_undefined T = Const (\<^const_name>\undefined\, T) (* testing functions: testing with increasing sizes (and cardinalities) *) @@ -210,8 +210,8 @@ fun get_finite_types ctxt = fst (chop (Config.get ctxt Quickcheck.finite_type_size) - [@{typ Enum.finite_1}, @{typ Enum.finite_2}, @{typ Enum.finite_3}, - @{typ Enum.finite_4}, @{typ Enum.finite_5}]) + [\<^typ>\Enum.finite_1\, \<^typ>\Enum.finite_2\, \<^typ>\Enum.finite_3\, + \<^typ>\Enum.finite_4\, \<^typ>\Enum.finite_5\]) exception WELLSORTED of string @@ -233,7 +233,7 @@ (* minimalistic preprocessing *) -fun strip_all (Const (@{const_name HOL.All}, _) $ Abs (a, T, t)) = +fun strip_all (Const (\<^const_name>\HOL.All\, _) $ Abs (a, T, t)) = let val (a', t') = strip_all t in ((a, T) :: a', t') end | strip_all t = ([], t) @@ -315,9 +315,9 @@ fun mk_safe_if genuine_only none (cond, then_t, else_t) genuine = let val T = fastype_of then_t - val if_t = Const (@{const_name If}, @{typ bool} --> T --> T --> T) + val if_t = Const (\<^const_name>\If\, \<^typ>\bool\ --> T --> T --> T) in - Const (@{const_name "Quickcheck_Random.catch_match"}, T --> T --> T) $ + Const (\<^const_name>\Quickcheck_Random.catch_match\, T --> T --> T) $ (if_t $ cond $ then_t $ else_t genuine) $ (if_t $ genuine_only $ none $ else_t false) end @@ -429,7 +429,7 @@ end) fun ensure_common_sort_datatype (sort, instantiate) = - ensure_sort (((@{sort typerep}, @{sort term_of}), sort), + ensure_sort (((\<^sort>\typerep\, \<^sort>\term_of\), sort), (fn thy => BNF_LFP_Compat.the_descr thy compat_prefs, instantiate)) fun datatype_interpretation name = @@ -440,20 +440,20 @@ fun gen_mk_parametric_generator_expr ((mk_generator_expr, out_of_bounds), T) ctxt ts = let - val if_t = Const (@{const_name If}, @{typ bool} --> T --> T --> T) + val if_t = Const (\<^const_name>\If\, \<^typ>\bool\ --> T --> T --> T) fun mk_if (index, (t, eval_terms)) else_t = - if_t $ (HOLogic.eq_const @{typ natural} $ Bound 0 $ HOLogic.mk_number @{typ natural} index) $ + if_t $ (HOLogic.eq_const \<^typ>\natural\ $ Bound 0 $ HOLogic.mk_number \<^typ>\natural\ index) $ (mk_generator_expr ctxt (t, eval_terms)) $ else_t - in absdummy @{typ natural} (fold_rev mk_if (1 upto (length ts) ~~ ts) out_of_bounds) end + in absdummy \<^typ>\natural\ (fold_rev mk_if (1 upto (length ts) ~~ ts) out_of_bounds) end (** post-processing of function terms **) -fun dest_fun_upd (Const (@{const_name fun_upd}, _) $ t0 $ t1 $ t2) = (t0, (t1, t2)) +fun dest_fun_upd (Const (\<^const_name>\fun_upd\, _) $ t0 $ t1 $ t2) = (t0, (t1, t2)) | dest_fun_upd t = raise TERM ("dest_fun_upd", [t]) fun mk_fun_upd T1 T2 (t1, t2) t = - Const (@{const_name fun_upd}, (T1 --> T2) --> T1 --> T2 --> T1 --> T2) $ t $ t1 $ t2 + Const (\<^const_name>\fun_upd\, (T1 --> T2) --> T1 --> T2 --> T1 --> T2) $ t $ t1 $ t2 fun dest_fun_upds t = (case try dest_fun_upd t of @@ -465,26 +465,26 @@ 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) = - Const (@{const_name insert}, T1 --> (T1 --> @{typ bool}) --> T1 --> @{typ bool}) $ +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) = + 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]) -fun make_coset T [] = Const (@{const_abbrev UNIV}, T --> @{typ bool}) +fun make_coset T [] = Const (\<^const_abbrev>\UNIV\, T --> \<^typ>\bool\) | make_coset T tps = let - val U = T --> @{typ bool} - fun invert @{const False} = @{const True} - | invert @{const True} = @{const False} + val U = T --> \<^typ>\bool\ + 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) + Const (\<^const_name>\Groups.minus_class.minus\, U --> U --> U) $ + Const (\<^const_abbrev>\UNIV\, U) $ make_set T (map (apsnd invert) tps) end -fun make_map T1 T2 [] = Const (@{const_abbrev Map.empty}, T1 --> T2) - | make_map T1 T2 ((_, Const (@{const_name None}, _)) :: tps) = make_map T1 T2 tps +fun make_map T1 T2 [] = Const (\<^const_abbrev>\Map.empty\, T1 --> T2) + | make_map T1 T2 ((_, Const (\<^const_name>\None\, _)) :: tps) = make_map T1 T2 tps | make_map T1 T2 ((t1, t2) :: tps) = mk_fun_upd T1 T2 (t1, t2) (make_map T1 T2 tps) fun post_process_term t = @@ -498,21 +498,21 @@ (c as Const (_, _), ts) => list_comb (c, map post_process_term ts)) in (case fastype_of t of - Type (@{type_name fun}, [T1, T2]) => + Type (\<^type_name>\fun\, [T1, T2]) => (case try dest_fun_upds t of SOME (tps, t) => (map (apply2 post_process_term) tps, map_Abs post_process_term t) |> (case T2 of - @{typ bool} => + \<^typ>\bool\ => (case t of - 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 + 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}, _) => + | Type (\<^type_name>\option\, _) => (case t of - Abs(_, _, Const (@{const_name None}, _)) => fst #> make_map T1 T2 - | Abs(_, _, Const (@{const_name undefined}, _)) => fst #> make_map T1 T2 + Abs(_, _, Const (\<^const_name>\None\, _)) => fst #> make_map T1 T2 + | Abs(_, _, Const (\<^const_name>\undefined\, _)) => fst #> make_map T1 T2 | _ => make_fun_upds T1 T2) | _ => make_fun_upds T1 T2) | NONE => process_args t) diff -r d24dcac5eb4c -r e61557884799 src/HOL/Tools/Quickcheck/random_generators.ML --- a/src/HOL/Tools/Quickcheck/random_generators.ML Wed Dec 06 19:34:59 2017 +0100 +++ b/src/HOL/Tools/Quickcheck/random_generators.ML Wed Dec 06 20:43:09 2017 +0100 @@ -27,13 +27,13 @@ (** abstract syntax **) -fun termifyT T = HOLogic.mk_prodT (T, @{typ "unit \ term"}) -val size = @{term "i::natural"}; -val size_pred = @{term "(i::natural) - 1"}; -val size' = @{term "j::natural"}; -val seed = @{term "s::Random.seed"}; +fun termifyT T = HOLogic.mk_prodT (T, \<^typ>\unit \ term\) +val size = \<^term>\i::natural\; +val size_pred = \<^term>\(i::natural) - 1\; +val size' = \<^term>\j::natural\; +val seed = \<^term>\s::Random.seed\; -val resultT = @{typ "(bool \ term list) option"}; +val resultT = \<^typ>\(bool \ term list) option\; (** typ "'a \ 'b" **) @@ -42,7 +42,7 @@ fun random_fun T1 T2 eq term_of random random_split seed = let - val fun_upd = Const (@{const_name fun_upd}, (T1 --> T2) --> T1 --> T2 --> T1 --> T2); + val fun_upd = Const (\<^const_name>\fun_upd\, (T1 --> T2) --> T1 --> T2 --> T1 --> T2); val ((_, t2), seed') = random seed; val (seed'', seed''') = random_split seed'; @@ -96,8 +96,8 @@ val inst = Thm.instantiate_cterm ([(a_v, icT)], []); 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]; + val eqs0 = [subst_v \<^term>\0::natural\ eq, + subst_v (\<^const>\Code_Numeral.Suc\ $ t_k) eq]; val eqs1 = map (Pattern.rewrite_term thy rew_ts []) eqs0; val ((_, (_, eqs2)), lthy') = lthy |> BNF_LFP_Compat.primrec_simple @@ -134,8 +134,8 @@ (aux_lhs, foldr1 HOLogic.mk_prod rhss); fun mk_proj t [T] = [t] | mk_proj t (Ts as T :: (Ts' as _ :: _)) = - Const (@{const_name fst}, foldr1 HOLogic.mk_prodT Ts --> T) $ t - :: mk_proj (Const (@{const_name snd}, + Const (\<^const_name>\fst\, foldr1 HOLogic.mk_prodT Ts --> T) $ t + :: mk_proj (Const (\<^const_name>\snd\, foldr1 HOLogic.mk_prodT Ts --> foldr1 HOLogic.mk_prodT Ts') $ t) Ts'; val projs = mk_proj (aux_lhs) Ts; val proj_eqs = map2 (fn v => fn proj => (v, lambda arg proj)) vs projs; @@ -193,9 +193,9 @@ val mk_const = curry (Sign.mk_const thy); val random_auxsN = map (prefix (random_auxN ^ "_")) (names @ auxnames); val rTs = Ts @ Us; - fun random_resultT T = @{typ Random.seed} - --> HOLogic.mk_prodT (termifyT T,@{typ Random.seed}); - fun sizeT T = @{typ natural} --> @{typ natural} --> T; + fun random_resultT T = \<^typ>\Random.seed\ + --> HOLogic.mk_prodT (termifyT T,\<^typ>\Random.seed\); + fun sizeT T = \<^typ>\natural\ --> \<^typ>\natural\ --> T; val random_auxT = sizeT o random_resultT; val random_auxs = map2 (fn s => fn rT => Free (s, random_auxT rT)) random_auxsN rTs; @@ -205,7 +205,7 @@ val T = Type (tyco, Ts); fun mk_random_fun_lift [] t = t | mk_random_fun_lift (fT :: fTs) t = - mk_const @{const_name random_fun_lift} [fTs ---> T, fT] $ + mk_const \<^const_name>\random_fun_lift\ [fTs ---> T, fT] $ mk_random_fun_lift fTs t; val t = mk_random_fun_lift fTs (nth random_auxs k $ size_pred $ size'); val size = Option.map snd (Old_Datatype_Aux.find_shortest_path descr k) @@ -221,16 +221,16 @@ val is_rec = exists is_some ks; val k = fold (fn NONE => I | SOME k => Integer.max k) ks 0; val vs = Name.invent_names Name.context "x" (map snd simple_tTs); - val tc = HOLogic.mk_return T @{typ Random.seed} + val tc = HOLogic.mk_return T \<^typ>\Random.seed\ (HOLogic.mk_valtermify_app c vs simpleT); val t = HOLogic.mk_ST - (map2 (fn (t, _) => fn (v, T') => ((t, @{typ Random.seed}), SOME ((v, termifyT T')))) tTs vs) - tc @{typ Random.seed} (SOME T, @{typ Random.seed}); + (map2 (fn (t, _) => fn (v, T') => ((t, \<^typ>\Random.seed\), SOME ((v, termifyT T')))) tTs vs) + tc \<^typ>\Random.seed\ (SOME T, \<^typ>\Random.seed\); val tk = if is_rec then if k = 0 then size - else @{term "Quickcheck_Random.beyond :: natural \ natural \ natural"} - $ HOLogic.mk_number @{typ natural} k $ size - else @{term "1::natural"} + else \<^term>\Quickcheck_Random.beyond :: natural \ natural \ natural\ + $ HOLogic.mk_number \<^typ>\natural\ k $ size + else \<^term>\1::natural\ in (is_rec, HOLogic.mk_prod (tk, t)) end; fun sort_rec xs = map_filter (fn (true, t) => SOME t | _ => NONE) xs @@ -239,9 +239,9 @@ |> (map o apfst) Type |> map (fn (T, cs) => (T, (sort_rec o map (mk_consexpr T)) cs)); fun mk_select (rT, xs) = - mk_const @{const_name Quickcheck_Random.collapse} [@{typ Random.seed}, termifyT rT] - $ (mk_const @{const_name Random.select_weight} [random_resultT rT] - $ HOLogic.mk_list (HOLogic.mk_prodT (@{typ natural}, random_resultT rT)) xs) + mk_const \<^const_name>\Quickcheck_Random.collapse\ [\<^typ>\Random.seed\, termifyT rT] + $ (mk_const \<^const_name>\Random.select_weight\ [random_resultT rT] + $ HOLogic.mk_list (HOLogic.mk_prodT (\<^typ>\natural\, random_resultT rT)) xs) $ seed; val auxs_lhss = map (fn t => t $ size $ size' $ seed) random_auxs; val auxs_rhss = map mk_select gen_exprss; @@ -253,8 +253,8 @@ val mk_prop_eq = HOLogic.mk_Trueprop o HOLogic.mk_eq; fun mk_size_arg k = case Old_Datatype_Aux.find_shortest_path descr k of SOME (_, l) => if l = 0 then size - else @{term "max :: natural \ natural \ natural"} - $ HOLogic.mk_number @{typ natural} l $ size + else \<^term>\max :: natural \ natural \ natural\ + $ HOLogic.mk_number \<^typ>\natural\ l $ size | NONE => size; val (random_auxs, auxs_eqs) = (apsnd o map) mk_prop_eq (mk_random_aux_eqs thy descr vs (names, auxnames) (Ts, Us)); @@ -262,7 +262,7 @@ (HOLogic.mk_random T size, nth random_auxs k $ mk_size_arg k $ size)) Ts; in thy - |> Class.instantiation (tycos, vs, @{sort random}) + |> Class.instantiation (tycos, vs, \<^sort>\random\) |> random_aux_specification prfx random_auxN auxs_eqs |> `(fn lthy => map (Syntax.check_term lthy) random_defs) |-> (fn random_defs' => fold_map (fn random_def => @@ -305,88 +305,88 @@ val bounds = map_index (fn (i, ty) => (2 * (bound_max - i) + 1, 2 * (bound_max - i), 2 * i, ty)) Ts; val result = list_comb (prop, map (fn (i, _, _, _) => Bound i) bounds); - val terms = HOLogic.mk_list @{typ term} (map (fn (_, i, _, _) => Bound i $ @{term "()"}) bounds); + val terms = HOLogic.mk_list \<^typ>\term\ (map (fn (_, i, _, _) => Bound i $ \<^term>\()\) bounds); val ([genuine_only_name], _) = Variable.variant_fixes ["genuine_only"] ctxt - val genuine_only = Free (genuine_only_name, @{typ bool}) - val none_t = Const (@{const_name "None"}, resultT) + val genuine_only = Free (genuine_only_name, \<^typ>\bool\) + val none_t = Const (\<^const_name>\None\, resultT) val check = Quickcheck_Common.mk_safe_if genuine_only none_t (result, none_t, - fn genuine => @{term "Some :: bool * term list => (bool * term list) option"} $ + fn genuine => \<^term>\Some :: bool \ term list => (bool \ term list) option\ $ HOLogic.mk_prod (Quickcheck_Common.reflect_bool genuine, terms)) - val return = HOLogic.pair_const resultT @{typ Random.seed}; + val return = HOLogic.pair_const resultT \<^typ>\Random.seed\; fun liftT T sT = sT --> HOLogic.mk_prodT (T, sT); - fun mk_termtyp T = HOLogic.mk_prodT (T, @{typ "unit => term"}); - fun mk_scomp T1 T2 sT f g = Const (@{const_name scomp}, + fun mk_termtyp T = HOLogic.mk_prodT (T, \<^typ>\unit \ term\); + fun mk_scomp T1 T2 sT f g = Const (\<^const_name>\scomp\, liftT T1 sT --> (T1 --> liftT T2 sT) --> liftT T2 sT) $ f $ g; fun mk_case_prod T = Sign.mk_const thy - (@{const_name case_prod}, [T, @{typ "unit => term"}, liftT resultT @{typ Random.seed}]); + (\<^const_name>\case_prod\, [T, \<^typ>\unit \ term\, liftT resultT \<^typ>\Random.seed\]); fun mk_scomp_split T t t' = - mk_scomp (mk_termtyp T) resultT @{typ Random.seed} t - (mk_case_prod T $ Abs ("", T, Abs ("", @{typ "unit => term"}, t'))); + mk_scomp (mk_termtyp T) resultT \<^typ>\Random.seed\ t + (mk_case_prod T $ Abs ("", T, Abs ("", \<^typ>\unit => term\, t'))); fun mk_bindclause (_, _, i, T) = mk_scomp_split T - (Sign.mk_const thy (@{const_name Quickcheck_Random.random}, [T]) $ Bound i); + (Sign.mk_const thy (\<^const_name>\Quickcheck_Random.random\, [T]) $ Bound i); in lambda genuine_only - (Abs ("n", @{typ natural}, fold_rev mk_bindclause bounds (return $ check true))) + (Abs ("n", \<^typ>\natural\, fold_rev mk_bindclause bounds (return $ check true))) end; fun mk_reporting_generator_expr ctxt (t, _) = let val thy = Proof_Context.theory_of ctxt - val resultT = @{typ "(bool * term list) option * (bool list * bool)"} + val resultT = \<^typ>\(bool \ term list) option \ (bool list \ bool)\ val prop = fold_rev absfree (Term.add_frees t []) t val Ts = (map snd o fst o strip_abs) prop val bound_max = length Ts - 1 val bounds = map_index (fn (i, ty) => (2 * (bound_max - i) + 1, 2 * (bound_max - i), 2 * i, ty)) Ts; val prop' = betapplys (prop, map (fn (i, _, _, _) => Bound i) bounds); - val terms = HOLogic.mk_list @{typ term} (map (fn (_, i, _, _) => Bound i $ @{term "()"}) bounds) + val terms = HOLogic.mk_list \<^typ>\term\ (map (fn (_, i, _, _) => Bound i $ \<^term>\()\) bounds) val (assms, concl) = Quickcheck_Common.strip_imp prop' - val return = HOLogic.pair_const resultT @{typ Random.seed}; + val return = HOLogic.pair_const resultT \<^typ>\Random.seed\; fun mk_assms_report i = - HOLogic.mk_prod (@{term "None :: (bool * term list) option"}, + HOLogic.mk_prod (\<^term>\None :: (bool \ term list) option\, HOLogic.mk_prod (HOLogic.mk_list HOLogic.boolT - (replicate i @{term True} @ replicate (length assms - i) @{term False}), - @{term False})) + (replicate i \<^term>\True\ @ replicate (length assms - i) \<^term>\False\), + \<^term>\False\)) fun mk_concl_report b = - HOLogic.mk_prod (HOLogic.mk_list HOLogic.boolT (replicate (length assms) @{term True}), + HOLogic.mk_prod (HOLogic.mk_list HOLogic.boolT (replicate (length assms) \<^term>\True\), Quickcheck_Common.reflect_bool b) val ([genuine_only_name], _) = Variable.variant_fixes ["genuine_only"] ctxt - val genuine_only = Free (genuine_only_name, @{typ bool}) - val none_t = HOLogic.mk_prod (@{term "None :: (bool * term list) option"}, mk_concl_report true) + val genuine_only = Free (genuine_only_name, \<^typ>\bool\) + val none_t = HOLogic.mk_prod (\<^term>\None :: (bool \ term list) option\, mk_concl_report true) val concl_check = Quickcheck_Common.mk_safe_if genuine_only none_t (concl, none_t, - fn genuine => HOLogic.mk_prod (@{term "Some :: bool * term list => (bool * term list) option"} $ + fn genuine => HOLogic.mk_prod (\<^term>\Some :: bool \ term list => (bool \ term list) option\ $ HOLogic.mk_prod (Quickcheck_Common.reflect_bool genuine, terms), mk_concl_report false)) val check = fold_rev (fn (i, assm) => fn t => Quickcheck_Common.mk_safe_if genuine_only (mk_assms_report i) (HOLogic.mk_not assm, mk_assms_report i, t)) (map_index I assms) concl_check fun liftT T sT = sT --> HOLogic.mk_prodT (T, sT); - fun mk_termtyp T = HOLogic.mk_prodT (T, @{typ "unit => term"}); - fun mk_scomp T1 T2 sT f g = Const (@{const_name scomp}, + fun mk_termtyp T = HOLogic.mk_prodT (T, \<^typ>\unit \ term\); + fun mk_scomp T1 T2 sT f g = Const (\<^const_name>\scomp\, liftT T1 sT --> (T1 --> liftT T2 sT) --> liftT T2 sT) $ f $ g; fun mk_case_prod T = Sign.mk_const thy - (@{const_name case_prod}, [T, @{typ "unit => term"}, liftT resultT @{typ Random.seed}]); + (\<^const_name>\case_prod\, [T, \<^typ>\unit \ term\, liftT resultT \<^typ>\Random.seed\]); fun mk_scomp_split T t t' = - mk_scomp (mk_termtyp T) resultT @{typ Random.seed} t - (mk_case_prod T $ Abs ("", T, Abs ("", @{typ "unit => term"}, t'))); + mk_scomp (mk_termtyp T) resultT \<^typ>\Random.seed\ t + (mk_case_prod T $ Abs ("", T, Abs ("", \<^typ>\unit \ term\, t'))); fun mk_bindclause (_, _, i, T) = mk_scomp_split T - (Sign.mk_const thy (@{const_name Quickcheck_Random.random}, [T]) $ Bound i); + (Sign.mk_const thy (\<^const_name>\Quickcheck_Random.random\, [T]) $ Bound i); in lambda genuine_only - (Abs ("n", @{typ natural}, fold_rev mk_bindclause bounds (return $ check true))) + (Abs ("n", \<^typ>\natural\, fold_rev mk_bindclause bounds (return $ check true))) end val mk_parametric_generator_expr = Quickcheck_Common.gen_mk_parametric_generator_expr ((mk_generator_expr, - absdummy @{typ bool} (absdummy @{typ natural} - @{term "Pair None :: Random.seed => (bool * term list) option * Random.seed"})), - @{typ "bool => natural => Random.seed => (bool * term list) option * Random.seed"}) + absdummy \<^typ>\bool\ (absdummy \<^typ>\natural\ + \<^term>\Pair None :: Random.seed \ (bool \ term list) option \ Random.seed\)), + \<^typ>\bool \ natural \ Random.seed \ (bool \ term list) option \ Random.seed\) val mk_parametric_reporting_generator_expr = Quickcheck_Common.gen_mk_parametric_generator_expr ((mk_reporting_generator_expr, - absdummy @{typ bool} (absdummy @{typ natural} - @{term "Pair (None, ([], False)) :: Random.seed => - ((bool * term list) option * (bool list * bool)) * Random.seed"})), - @{typ "bool => natural => Random.seed => ((bool * term list) option * (bool list * bool)) * Random.seed"}) + absdummy \<^typ>\bool\ (absdummy \<^typ>\natural\ + \<^term>\Pair (None, ([], False)) :: Random.seed \ + ((bool \ term list) option \ (bool list \ bool)) \ Random.seed\)), + \<^typ>\bool \ natural \ Random.seed \ ((bool \ term list) option \ (bool list \ bool)) \ Random.seed\) (* single quickcheck report *) @@ -468,8 +468,8 @@ compiled genuine_only [Code_Numeral.natural_of_integer card, Code_Numeral.natural_of_integer size] end; -val size_types = [@{type_name Enum.finite_1}, @{type_name Enum.finite_2}, - @{type_name Enum.finite_3}, @{type_name Enum.finite_4}, @{type_name Enum.finite_5}]; +val size_types = [\<^type_name>\Enum.finite_1\, \<^type_name>\Enum.finite_2\, + \<^type_name>\Enum.finite_3\, \<^type_name>\Enum.finite_4\, \<^type_name>\Enum.finite_5\]; fun size_matters_for _ Ts = not (forall (fn Type (tyco, []) => member (op =) size_types tyco | _ => false) Ts); @@ -480,12 +480,12 @@ (** setup **) -val active = Attrib.setup_config_bool @{binding quickcheck_random_active} (K false); +val active = Attrib.setup_config_bool \<^binding>\quickcheck_random_active\ (K false); val _ = Theory.setup - (Quickcheck_Common.datatype_interpretation @{plugin quickcheck_random} - (@{sort random}, instantiate_random_datatype) #> + (Quickcheck_Common.datatype_interpretation \<^plugin>\quickcheck_random\ + (\<^sort>\random\, instantiate_random_datatype) #> Context.theory_map (Quickcheck.add_tester ("random", (active, test_goals)))); end; diff -r d24dcac5eb4c -r e61557884799 src/HOL/Tools/SMT/conj_disj_perm.ML --- a/src/HOL/Tools/SMT/conj_disj_perm.ML Wed Dec 06 19:34:59 2017 +0100 +++ b/src/HOL/Tools/SMT/conj_disj_perm.ML Wed Dec 06 20:43:09 2017 +0100 @@ -13,7 +13,7 @@ struct fun with_assumption ct f = - let val ct' = Thm.apply @{cterm HOL.Trueprop} ct + let val ct' = Thm.apply \<^cterm>\HOL.Trueprop\ ct in Thm.implies_intr ct' (f (Thm.assume ct')) end fun eq_from_impls thm1 thm2 = thm2 INCR_COMP (thm1 INCR_COMP @{thm iffI}) @@ -25,9 +25,9 @@ fun explode_thm thm = (case HOLogic.dest_Trueprop (Thm.prop_of thm) of - @{const HOL.conj} $ _ $ _ => explode_conj_thm @{thm conjunct1} @{thm conjunct2} thm - | @{const HOL.Not} $ (@{const HOL.disj} $ _ $ _) => explode_conj_thm ndisj1_rule ndisj2_rule thm - | @{const HOL.Not} $ (@{const HOL.Not} $ _) => explode_thm (thm RS @{thm notnotD}) + \<^const>\HOL.conj\ $ _ $ _ => explode_conj_thm @{thm conjunct1} @{thm conjunct2} thm + | \<^const>\HOL.Not\ $ (\<^const>\HOL.disj\ $ _ $ _) => explode_conj_thm ndisj1_rule ndisj2_rule thm + | \<^const>\HOL.Not\ $ (\<^const>\HOL.Not\ $ _) => explode_thm (thm RS @{thm notnotD}) | _ => add_lit thm) and explode_conj_thm rule1 rule2 thm lits = @@ -36,7 +36,7 @@ val not_false_rule = @{lemma "\False" by auto} fun explode thm = explode_thm thm (add_lit not_false_rule (add_lit @{thm TrueI} Termtab.empty)) -fun find_dual_lit lits (@{const HOL.Not} $ t, thm) = Termtab.lookup lits t |> Option.map (pair thm) +fun find_dual_lit lits (\<^const>\HOL.Not\ $ t, thm) = Termtab.lookup lits t |> Option.map (pair thm) | find_dual_lit _ _ = NONE fun find_dual_lits lits = Termtab.get_first (find_dual_lit lits) lits @@ -49,10 +49,10 @@ SOME thm => thm | NONE => join_term lits t) -and join_term lits (@{const HOL.conj} $ t $ u) = @{thm conjI} OF (map (join lits) [t, u]) - | join_term lits (@{const HOL.Not} $ (@{const HOL.disj} $ t $ u)) = +and join_term lits (\<^const>\HOL.conj\ $ t $ u) = @{thm conjI} OF (map (join lits) [t, u]) + | join_term lits (\<^const>\HOL.Not\ $ (\<^const>\HOL.disj\ $ t $ u)) = ndisj_rule OF (map (join lits o HOLogic.mk_not) [t, u]) - | join_term lits (@{const HOL.Not} $ (@{const HOL.Not} $ t)) = join lits t RS not_not_rule + | join_term lits (\<^const>\HOL.Not\ $ (\<^const>\HOL.Not\ $ t)) = join lits t RS not_not_rule | join_term _ t = raise TERM ("join_term", [t]) fun prove_conj_disj_imp ct cu = with_assumption ct (fn thm => join (explode thm) (Thm.term_of cu)) @@ -68,19 +68,19 @@ fun prove_any_imp ct = (case Thm.term_of ct of - @{const HOL.False} => @{thm FalseE} - | @{const HOL.Not} $ (@{const HOL.Not} $ @{const HOL.False}) => not_not_false_rule - | @{const HOL.Not} $ @{const HOL.True} => not_true_rule + \<^const>\HOL.False\ => @{thm FalseE} + | \<^const>\HOL.Not\ $ (\<^const>\HOL.Not\ $ \<^const>\HOL.False\) => not_not_false_rule + | \<^const>\HOL.Not\ $ \<^const>\HOL.True\ => not_true_rule | _ => raise CTERM ("prove_any_imp", [ct])) fun prove_contradiction_imp ct = with_assumption ct (fn thm => let val lits = explode thm in - (case Termtab.lookup lits @{const HOL.False} of + (case Termtab.lookup lits \<^const>\HOL.False\ of SOME thm' => thm' RS @{thm FalseE} | NONE => - (case Termtab.lookup lits (@{const HOL.Not} $ @{const HOL.True}) of + (case Termtab.lookup lits (\<^const>\HOL.Not\ $ \<^const>\HOL.True\) of SOME thm' => thm' RS not_true_rule | NONE => (case find_dual_lits lits of @@ -99,13 +99,13 @@ datatype kind = True | False | Conj | Disj | Other -fun choose t _ _ _ @{const HOL.True} = t - | choose _ f _ _ @{const HOL.False} = f - | choose _ _ c _ (@{const HOL.conj} $ _ $ _) = c - | choose _ _ _ d (@{const HOL.disj} $ _ $ _) = d +fun choose t _ _ _ \<^const>\HOL.True\ = t + | choose _ f _ _ \<^const>\HOL.False\ = f + | choose _ _ c _ (\<^const>\HOL.conj\ $ _ $ _) = c + | choose _ _ _ d (\<^const>\HOL.disj\ $ _ $ _) = d | choose _ _ _ _ _ = Other -fun kind_of (@{const HOL.Not} $ t) = choose False True Disj Conj t +fun kind_of (\<^const>\HOL.Not\ $ t) = choose False True Disj Conj t | kind_of t = choose True False Conj Disj t fun prove_conj_disj_perm ct cp = @@ -120,7 +120,7 @@ fun conj_disj_perm_tac ctxt = CSUBGOAL (fn (ct, i) => (case Thm.term_of ct of - @{const HOL.Trueprop} $ (@{const HOL.eq(bool)} $ _ $ _) => + \<^const>\HOL.Trueprop\ $ (@{const HOL.eq(bool)} $ _ $ _) => resolve_tac ctxt [prove_conj_disj_perm ct (Thm.dest_binop (Thm.dest_arg ct))] i | _ => no_tac)) diff -r d24dcac5eb4c -r e61557884799 src/HOL/Tools/SMT/smt_builtin.ML --- a/src/HOL/Tools/SMT/smt_builtin.ML Wed Dec 06 19:34:59 2017 +0100 +++ b/src/HOL/Tools/SMT/smt_builtin.ML Wed Dec 06 20:43:09 2017 +0100 @@ -180,8 +180,8 @@ fun dest_builtin_eq ctxt t u = let - val aT = TFree (Name.aT, @{sort type}) - val c = (@{const_name HOL.eq}, aT --> aT --> @{typ bool}) + val aT = TFree (Name.aT, \<^sort>\type\) + val c = (\<^const_name>\HOL.eq\, aT --> aT --> \<^typ>\bool\) fun mk ts = Term.list_comb (HOLogic.eq_const (Term.fastype_of (hd ts)), ts) in dest_builtin_fun ctxt c [] @@ -193,10 +193,10 @@ dest_builtin_fun ctxt c ts else NONE -fun dest_builtin_pred ctxt = special_builtin_fun (equal @{typ bool} o fst) ctxt +fun dest_builtin_pred ctxt = special_builtin_fun (equal \<^typ>\bool\ o fst) ctxt fun dest_builtin_conn ctxt = - special_builtin_fun (forall (equal @{typ bool}) o (op ::)) ctxt + special_builtin_fun (forall (equal \<^typ>\bool\) o (op ::)) ctxt fun dest_builtin ctxt c ts = let val t = Term.list_comb (Const c, ts) diff -r d24dcac5eb4c -r e61557884799 src/HOL/Tools/SMT/smt_config.ML --- a/src/HOL/Tools/SMT/smt_config.ML Wed Dec 06 19:34:59 2017 +0100 +++ b/src/HOL/Tools/SMT/smt_config.ML Wed Dec 06 20:43:09 2017 +0100 @@ -106,7 +106,7 @@ context |> Data.map (map_solvers (Symtab.update (name, (info, [])))) |> Context.map_theory (Attrib.setup (Binding.name (name ^ "_options")) - (Scan.lift (@{keyword "="} |-- Args.name) >> + (Scan.lift (\<^keyword>\=\ |-- Args.name) >> (Thm.declaration_attribute o K o set_solver_options o pair name)) ("additional command line options for SMT solver " ^ quote name)) @@ -164,30 +164,30 @@ in solver_info_of (K []) all_options ctxt end val setup_solver = - Attrib.setup @{binding smt_solver} - (Scan.lift (@{keyword "="} |-- Args.name) >> + Attrib.setup \<^binding>\smt_solver\ + (Scan.lift (\<^keyword>\=\ |-- Args.name) >> (Thm.declaration_attribute o K o select_solver)) "SMT solver configuration" (* options *) -val oracle = Attrib.setup_config_bool @{binding smt_oracle} (K true) -val timeout = Attrib.setup_config_real @{binding smt_timeout} (K 30.0) -val reconstruction_step_timeout = Attrib.setup_config_real @{binding smt_reconstruction_step_timeout} (K 10.0) -val random_seed = Attrib.setup_config_int @{binding smt_random_seed} (K 1) -val read_only_certificates = Attrib.setup_config_bool @{binding smt_read_only_certificates} (K false) -val verbose = Attrib.setup_config_bool @{binding smt_verbose} (K true) -val trace = Attrib.setup_config_bool @{binding smt_trace} (K false) -val statistics = Attrib.setup_config_bool @{binding smt_statistics} (K false) -val monomorph_limit = Attrib.setup_config_int @{binding smt_monomorph_limit} (K 10) -val monomorph_instances = Attrib.setup_config_int @{binding smt_monomorph_instances} (K 500) -val explicit_application = Attrib.setup_config_int @{binding smt_explicit_application} (K 1) -val higher_order = Attrib.setup_config_bool @{binding smt_higher_order} (K false) -val nat_as_int = Attrib.setup_config_bool @{binding smt_nat_as_int} (K false) -val infer_triggers = Attrib.setup_config_bool @{binding smt_infer_triggers} (K false) -val debug_files = Attrib.setup_config_string @{binding smt_debug_files} (K "") -val sat_solver = Attrib.setup_config_string @{binding smt_sat_solver} (K "cdclite") +val oracle = Attrib.setup_config_bool \<^binding>\smt_oracle\ (K true) +val timeout = Attrib.setup_config_real \<^binding>\smt_timeout\ (K 30.0) +val reconstruction_step_timeout = Attrib.setup_config_real \<^binding>\smt_reconstruction_step_timeout\ (K 10.0) +val random_seed = Attrib.setup_config_int \<^binding>\smt_random_seed\ (K 1) +val read_only_certificates = Attrib.setup_config_bool \<^binding>\smt_read_only_certificates\ (K false) +val verbose = Attrib.setup_config_bool \<^binding>\smt_verbose\ (K true) +val trace = Attrib.setup_config_bool \<^binding>\smt_trace\ (K false) +val statistics = Attrib.setup_config_bool \<^binding>\smt_statistics\ (K false) +val monomorph_limit = Attrib.setup_config_int \<^binding>\smt_monomorph_limit\ (K 10) +val monomorph_instances = Attrib.setup_config_int \<^binding>\smt_monomorph_instances\ (K 500) +val explicit_application = Attrib.setup_config_int \<^binding>\smt_explicit_application\ (K 1) +val higher_order = Attrib.setup_config_bool \<^binding>\smt_higher_order\ (K false) +val nat_as_int = Attrib.setup_config_bool \<^binding>\smt_nat_as_int\ (K false) +val infer_triggers = Attrib.setup_config_bool \<^binding>\smt_infer_triggers\ (K false) +val debug_files = Attrib.setup_config_string \<^binding>\smt_debug_files\ (K "") +val sat_solver = Attrib.setup_config_string \<^binding>\smt_sat_solver\ (K "cdclite") (* diagnostics *) @@ -222,8 +222,8 @@ |> SOME o Cache_IO.unsynchronized_init)) val setup_certificates = - Attrib.setup @{binding smt_certificates} - (Scan.lift (@{keyword "="} |-- Args.name) >> + Attrib.setup \<^binding>\smt_certificates\ + (Scan.lift (\<^keyword>\=\ |-- Args.name) >> (Thm.declaration_attribute o K o select_certificates)) "SMT certificates configuration" @@ -263,7 +263,7 @@ end val _ = - Outer_Syntax.command @{command_keyword smt_status} + Outer_Syntax.command \<^command_keyword>\smt_status\ "show the available SMT solvers, the currently selected SMT solver, \ \and the values of SMT configuration options" (Scan.succeed (Toplevel.keep (print_setup o Toplevel.context_of))) diff -r d24dcac5eb4c -r e61557884799 src/HOL/Tools/SMT/smt_datatypes.ML --- a/src/HOL/Tools/SMT/smt_datatypes.ML Wed Dec 06 19:34:59 2017 +0100 +++ b/src/HOL/Tools/SMT/smt_datatypes.ML Wed Dec 06 20:43:09 2017 +0100 @@ -100,7 +100,7 @@ (http://church.cims.nyu.edu/bugzilla3/show_bug.cgi?id=597). It should be removed once the bug is fixed. *) if forall (forall (forall (is_homogenously_nested_co_recursive))) ctrXs_Tsss andalso - forall (forall (forall (curry (op <>) @{typ bool}))) + forall (forall (forall (curry (op <>) \<^typ>\bool\))) (map (map (map substAs)) ctrXs_Tsss) then get_ctr_sugar_decl ctr_sugar T Ts ctxt |>> map (pair fp) else @@ -127,9 +127,9 @@ fun add (TFree _) = I | add (TVar _) = I - | add (T as Type (@{type_name fun}, _)) = + | add (T as Type (\<^type_name>\fun\, _)) = fold add (Term.body_type T :: Term.binder_types T) - | add @{typ bool} = I + | add \<^typ>\bool\ = I | add (T as Type (n, Ts)) = (fn (dss, ctxt1) => if declared T declss orelse declared' T dss then (dss, ctxt1) diff -r d24dcac5eb4c -r e61557884799 src/HOL/Tools/SMT/smt_normalize.ML --- a/src/HOL/Tools/SMT/smt_normalize.ML Wed Dec 06 19:34:59 2017 +0100 +++ b/src/HOL/Tools/SMT/smt_normalize.ML Wed Dec 06 20:43:09 2017 +0100 @@ -31,7 +31,7 @@ (** instantiate elimination rules **) local - val (cpfalse, cfalse) = `SMT_Util.mk_cprop (Thm.cterm_of @{context} @{const False}) + val (cpfalse, cfalse) = `SMT_Util.mk_cprop (Thm.cterm_of @{context} \<^const>\False\) fun inst f ct thm = let val cv = f (Drule.strip_imp_concl (Thm.cprop_of thm)) @@ -40,7 +40,7 @@ fun instantiate_elim thm = (case Thm.concl_of thm of - @{const Trueprop} $ Var (_, @{typ bool}) => inst Thm.dest_arg cfalse thm + \<^const>\Trueprop\ $ Var (_, \<^typ>\bool\) => inst Thm.dest_arg cfalse thm | Var _ => inst I cpfalse thm | _ => thm) @@ -51,9 +51,9 @@ fun norm_def thm = (case Thm.prop_of thm of - @{const Trueprop} $ (Const (@{const_name HOL.eq}, _) $ _ $ Abs _) => + \<^const>\Trueprop\ $ (Const (\<^const_name>\HOL.eq\, _) $ _ $ Abs _) => norm_def (thm RS @{thm fun_cong}) - | Const (@{const_name Pure.eq}, _) $ _ $ Abs _ => norm_def (thm RS @{thm meta_eq_to_obj_eq}) + | Const (\<^const_name>\Pure.eq\, _) $ _ $ Abs _ => norm_def (thm RS @{thm meta_eq_to_obj_eq}) | _ => thm) @@ -61,26 +61,26 @@ fun atomize_conv ctxt ct = (case Thm.term_of ct of - @{const Pure.imp} $ _ $ _ => + \<^const>\Pure.imp\ $ _ $ _ => Conv.binop_conv (atomize_conv ctxt) then_conv Conv.rewr_conv @{thm atomize_imp} - | Const (@{const_name Pure.eq}, _) $ _ $ _ => + | Const (\<^const_name>\Pure.eq\, _) $ _ $ _ => Conv.binop_conv (atomize_conv ctxt) then_conv Conv.rewr_conv @{thm atomize_eq} - | Const (@{const_name Pure.all}, _) $ Abs _ => + | Const (\<^const_name>\Pure.all\, _) $ Abs _ => Conv.binder_conv (atomize_conv o snd) ctxt then_conv Conv.rewr_conv @{thm atomize_all} | _ => Conv.all_conv) ct handle CTERM _ => Conv.all_conv ct val setup_atomize = - fold SMT_Builtin.add_builtin_fun_ext'' [@{const_name Pure.imp}, @{const_name Pure.eq}, - @{const_name Pure.all}, @{const_name Trueprop}] + fold SMT_Builtin.add_builtin_fun_ext'' [\<^const_name>\Pure.imp\, \<^const_name>\Pure.eq\, + \<^const_name>\Pure.all\, \<^const_name>\Trueprop\] (** unfold special quantifiers **) val special_quant_table = [ - (@{const_name Ex1}, @{thm Ex1_def_raw}), - (@{const_name Ball}, @{thm Ball_def_raw}), - (@{const_name Bex}, @{thm Bex_def_raw})] + (\<^const_name>\Ex1\, @{thm Ex1_def_raw}), + (\<^const_name>\Ball\, @{thm Ball_def_raw}), + (\<^const_name>\Bex\, @{thm Bex_def_raw})] local fun special_quant (Const (n, _)) = AList.lookup (op =) special_quant_table n @@ -105,8 +105,8 @@ local (*** check trigger syntax ***) - fun dest_trigger (Const (@{const_name pat}, _) $ _) = SOME true - | dest_trigger (Const (@{const_name nopat}, _) $ _) = SOME false + fun dest_trigger (Const (\<^const_name>\pat\, _) $ _) = SOME true + | dest_trigger (Const (\<^const_name>\nopat\, _) $ _) = SOME false | dest_trigger _ = NONE fun eq_list [] = false @@ -120,9 +120,9 @@ fun proper_quant inside f t = (case t of - Const (@{const_name All}, _) $ Abs (_, _, u) => proper_quant true f u - | Const (@{const_name Ex}, _) $ Abs (_, _, u) => proper_quant true f u - | @{const trigger} $ p $ u => + Const (\<^const_name>\All\, _) $ Abs (_, _, u) => proper_quant true f u + | Const (\<^const_name>\Ex\, _) $ Abs (_, _, u) => proper_quant true f u + | \<^const>\trigger\ $ p $ u => (if inside then f p else false) andalso proper_quant false f u | Abs (_, _, u) => proper_quant false f u | u1 $ u2 => proper_quant false f u1 andalso proper_quant false f u2 @@ -141,8 +141,8 @@ fun dest_cond_eq ct = (case Thm.term_of ct of - Const (@{const_name HOL.eq}, _) $ _ $ _ => Thm.dest_binop ct - | @{const HOL.implies} $ _ $ _ => dest_cond_eq (Thm.dest_arg ct) + Const (\<^const_name>\HOL.eq\, _) $ _ $ _ => Thm.dest_binop ct + | \<^const>\HOL.implies\ $ _ $ _ => dest_cond_eq (Thm.dest_arg ct) | _ => raise CTERM ("no equation", [ct])) fun get_constrs thy (Type (n, _)) = these (BNF_LFP_Compat.get_constrs thy n) @@ -186,14 +186,14 @@ Pattern.matches thy (t', u) andalso not (t aconv u)) in not (Term.exists_subterm some_match u) end - val pat = SMT_Util.mk_const_pat @{theory} @{const_name pat} SMT_Util.destT1 + val pat = SMT_Util.mk_const_pat @{theory} \<^const_name>\pat\ SMT_Util.destT1 fun mk_pat ct = Thm.apply (SMT_Util.instT' ct pat) ct fun mk_clist T = apply2 (Thm.cterm_of @{context}) (SMT_Util.symb_cons_const T, SMT_Util.symb_nil_const T) fun mk_list (ccons, cnil) f cts = fold_rev (Thm.mk_binop ccons o f) cts cnil - val mk_pat_list = mk_list (mk_clist @{typ pattern}) - val mk_mpat_list = mk_list (mk_clist @{typ "pattern symb_list"}) + val mk_pat_list = mk_list (mk_clist \<^typ>\pattern\) + val mk_mpat_list = mk_list (mk_clist \<^typ>\pattern symb_list\) fun mk_trigger ctss = mk_mpat_list (mk_pat_list mk_pat) ctss val trigger_eq = mk_meta_eq @{lemma "p = trigger t p" by (simp add: trigger_def)} @@ -220,7 +220,7 @@ in insert_trigger_conv (filter_mpats (get_mpats lhs)) ct end - fun has_trigger (@{const trigger} $ _ $ _) = true + fun has_trigger (\<^const>\trigger\ $ _ $ _) = true | has_trigger _ = false fun try_trigger_conv cv ct = @@ -238,7 +238,7 @@ val setup_trigger = fold SMT_Builtin.add_builtin_fun_ext'' - [@{const_name pat}, @{const_name nopat}, @{const_name trigger}] + [\<^const_name>\pat\, \<^const_name>\nopat\, \<^const_name>\trigger\] end @@ -280,10 +280,10 @@ (** rewrite bool case expressions as if expressions **) -val case_bool_entry = (@{const_name "bool.case_bool"}, @{thm case_bool_if}) +val case_bool_entry = (\<^const_name>\bool.case_bool\, @{thm case_bool_if}) local - fun is_case_bool (Const (@{const_name "bool.case_bool"}, _)) = true + fun is_case_bool (Const (\<^const_name>\bool.case_bool\, _)) = true | is_case_bool _ = false fun unfold_conv _ = @@ -294,7 +294,7 @@ fun rewrite_case_bool_conv ctxt = SMT_Util.if_exists_conv is_case_bool (Conv.top_conv unfold_conv ctxt) -val setup_case_bool = SMT_Builtin.add_builtin_fun_ext'' @{const_name "bool.case_bool"} +val setup_case_bool = SMT_Builtin.add_builtin_fun_ext'' \<^const_name>\bool.case_bool\ end @@ -302,12 +302,12 @@ (** unfold abs, min and max **) val abs_min_max_table = [ - (@{const_name min}, @{thm min_def_raw}), - (@{const_name max}, @{thm max_def_raw}), - (@{const_name abs}, @{thm abs_if_raw})] + (\<^const_name>\min\, @{thm min_def_raw}), + (\<^const_name>\max\, @{thm max_def_raw}), + (\<^const_name>\abs\, @{thm abs_if_raw})] local - fun abs_min_max ctxt (Const (n, Type (@{type_name fun}, [T, _]))) = + fun abs_min_max ctxt (Const (n, Type (\<^type_name>\fun\, [T, _]))) = (case AList.lookup (op =) abs_min_max_table n of NONE => NONE | SOME thm => if SMT_Builtin.is_builtin_typ_ext ctxt T then SOME thm else NONE) @@ -334,7 +334,7 @@ val simple_nat_ops = [ @{const less (nat)}, @{const less_eq (nat)}, - @{const Suc}, @{const plus (nat)}, @{const minus (nat)}] + \<^const>\Suc\, @{const plus (nat)}, @{const minus (nat)}] val mult_nat_ops = [@{const times (nat)}, @{const divide (nat)}, @{const modulo (nat)}] @@ -344,7 +344,7 @@ val nat_consts = nat_ops @ [@{const numeral (nat)}, @{const zero_class.zero (nat)}, @{const one_class.one (nat)}] - val nat_int_coercions = [@{const of_nat (int)}, @{const nat}] + val nat_int_coercions = [@{const of_nat (int)}, \<^const>\nat\] val builtin_nat_ops = nat_int_coercions @ simple_nat_ops @@ -399,7 +399,7 @@ if exists (uses_nat_int o Thm.prop_of) thms then (thms, nat_embedding) else (thms, []) val setup_nat_as_int = - SMT_Builtin.add_builtin_typ_ext (@{typ nat}, + SMT_Builtin.add_builtin_typ_ext (\<^typ>\nat\, fn ctxt => K (Config.get ctxt SMT_Config.nat_as_int)) #> fold (SMT_Builtin.add_builtin_fun_ext' o Term.dest_Const) builtin_nat_ops @@ -414,9 +414,9 @@ rewrite - 0 into 0 *) - fun is_irregular_number (Const (@{const_name numeral}, _) $ Const (@{const_name num.One}, _)) = + fun is_irregular_number (Const (\<^const_name>\numeral\, _) $ Const (\<^const_name>\num.One\, _)) = true - | is_irregular_number (Const (@{const_name uminus}, _) $ Const (@{const_name Groups.zero}, _)) = + | is_irregular_number (Const (\<^const_name>\uminus\, _) $ Const (\<^const_name>\Groups.zero\, _)) = true | is_irregular_number _ = false @@ -478,12 +478,12 @@ in burrow_ids (fold (fn e => e ctxt) es o rpair []) ithms end local - val ignored = member (op =) [@{const_name All}, @{const_name Ex}, - @{const_name Let}, @{const_name If}, @{const_name HOL.eq}] + val ignored = member (op =) [\<^const_name>\All\, \<^const_name>\Ex\, + \<^const_name>\Let\, \<^const_name>\If\, \<^const_name>\HOL.eq\] val schematic_consts_of = let - fun collect (@{const trigger} $ p $ t) = collect_trigger p #> collect t + fun collect (\<^const>\trigger\ $ p $ t) = collect_trigger p #> collect t | collect (t $ u) = collect t #> collect u | collect (Abs (_, _, t)) = collect t | collect (t as Const (n, _)) = @@ -492,8 +492,8 @@ and collect_trigger t = let val dest = these o try SMT_Util.dest_symb_list in fold (fold collect_pat o dest) (dest t) end - and collect_pat (Const (@{const_name pat}, _) $ t) = collect t - | collect_pat (Const (@{const_name nopat}, _) $ t) = collect t + and collect_pat (Const (\<^const_name>\pat\, _) $ t) = collect t + | collect_pat (Const (\<^const_name>\nopat\, _) $ t) = collect t | collect_pat _ = I in (fn t => collect t Symtab.empty) end in diff -r d24dcac5eb4c -r e61557884799 src/HOL/Tools/SMT/smt_translate.ML --- a/src/HOL/Tools/SMT/smt_translate.ML Wed Dec 06 19:34:59 2017 +0100 +++ b/src/HOL/Tools/SMT/smt_translate.ML Wed Dec 06 20:43:09 2017 +0100 @@ -199,21 +199,21 @@ SOME k => Term.list_comb (t, ts) |> k <> length ts ? expf k (length ts) T | NONE => Term.list_comb (t, ts)) - fun expand ((q as Const (@{const_name All}, _)) $ Abs a) = q $ abs_expand a - | expand ((q as Const (@{const_name All}, T)) $ t) = q $ exp expand T t - | expand (q as Const (@{const_name All}, T)) = exp2 T q - | expand ((q as Const (@{const_name Ex}, _)) $ Abs a) = q $ abs_expand a - | expand ((q as Const (@{const_name Ex}, T)) $ t) = q $ exp expand T t - | expand (q as Const (@{const_name Ex}, T)) = exp2 T q - | expand (Const (@{const_name Let}, T) $ t) = + fun expand ((q as Const (\<^const_name>\All\, _)) $ Abs a) = q $ abs_expand a + | expand ((q as Const (\<^const_name>\All\, T)) $ t) = q $ exp expand T t + | expand (q as Const (\<^const_name>\All\, T)) = exp2 T q + | expand ((q as Const (\<^const_name>\Ex\, _)) $ Abs a) = q $ abs_expand a + | expand ((q as Const (\<^const_name>\Ex\, T)) $ t) = q $ exp expand T t + | expand (q as Const (\<^const_name>\Ex\, T)) = exp2 T q + | expand (Const (\<^const_name>\Let\, T) $ t) = let val U = Term.domain_type (Term.range_type T) in Abs (Name.uu, U, Bound 0 $ Term.incr_boundvars 1 t) end - | expand (Const (@{const_name Let}, T)) = + | expand (Const (\<^const_name>\Let\, T)) = let val U = Term.domain_type (Term.range_type T) in Abs (Name.uu, Term.domain_type T, Abs (Name.uu, U, Bound 0 $ Bound 1)) end | expand t = (case Term.strip_comb t of - (Const (@{const_name Let}, _), t1 :: t2 :: ts) => + (Const (\<^const_name>\Let\, _), t1 :: t2 :: ts) => Term.betapplys (Term.betapply (expand t2, expand t1), map expand ts) | (u as Const (c as (_, T)), ts) => (case SMT_Builtin.dest_builtin ctxt c ts of @@ -252,12 +252,12 @@ fun take_vars_into_account types t i = let - fun find_min j (T as Type (@{type_name fun}, [_, T'])) = + fun find_min j (T as Type (\<^type_name>\fun\, [_, T'])) = if j = i orelse Typtab.defined types T then j else find_min (j + 1) T' | find_min j _ = j in find_min 0 (Term.type_of t) end - fun app u (t, T) = (Const (@{const_name fun_app}, T --> T) $ t $ u, Term.range_type T) + fun app u (t, T) = (Const (\<^const_name>\fun_app\, T --> T) $ t $ u, Term.range_type T) fun apply i t T ts = let @@ -288,11 +288,11 @@ fun traverse Ts t = (case Term.strip_comb t of - (q as Const (@{const_name All}, _), [Abs (x, T, u)]) => + (q as Const (\<^const_name>\All\, _), [Abs (x, T, u)]) => q $ Abs (x, T, in_trigger (T :: Ts) u) - | (q as Const (@{const_name Ex}, _), [Abs (x, T, u)]) => + | (q as Const (\<^const_name>\Ex\, _), [Abs (x, T, u)]) => q $ Abs (x, T, in_trigger (T :: Ts) u) - | (q as Const (@{const_name Let}, _), [u1, u2 as Abs _]) => + | (q as Const (\<^const_name>\Let\, _), [u1, u2 as Abs _]) => q $ traverse Ts u1 $ traverse Ts u2 | (u as Const (c as (_, T)), ts) => (case SMT_Builtin.dest_builtin ctxt c ts of @@ -306,12 +306,12 @@ | (u as Bound i, ts) => apply 0 u (nth Ts i) (map (traverse Ts) ts) | (Abs (n, T, u), ts) => traverses Ts (Abs (n, T, traverse (T::Ts) u)) ts | (u, ts) => traverses Ts u ts) - and in_trigger Ts ((c as @{const trigger}) $ p $ t) = c $ in_pats Ts p $ traverse Ts t + and in_trigger Ts ((c as \<^const>\trigger\) $ p $ t) = c $ in_pats Ts p $ traverse Ts t | in_trigger Ts t = traverse Ts t and in_pats Ts ps = - in_list @{typ "pattern symb_list"} (in_list @{typ pattern} (in_pat Ts)) ps - and in_pat Ts ((p as Const (@{const_name pat}, _)) $ t) = p $ traverse Ts t - | in_pat Ts ((p as Const (@{const_name nopat}, _)) $ t) = p $ traverse Ts t + in_list \<^typ>\pattern symb_list\ (in_list \<^typ>\pattern\ (in_pat Ts)) ps + and in_pat Ts ((p as Const (\<^const_name>\pat\, _)) $ t) = p $ traverse Ts t + | in_pat Ts ((p as Const (\<^const_name>\nopat\, _)) $ t) = p $ traverse Ts t | in_pat _ t = raise TERM ("bad pattern", [t]) and traverses Ts t ts = Term.list_comb (t, map (traverse Ts) ts) in map (traverse []) ts end @@ -324,7 +324,7 @@ (** map HOL formulas to FOL formulas (i.e., separate formulas froms terms) **) local - val is_quant = member (op =) [@{const_name All}, @{const_name Ex}] + val is_quant = member (op =) [\<^const_name>\All\, \<^const_name>\Ex\] val fol_rules = [ Let_def, @@ -343,9 +343,9 @@ fun in_term pat t = (case Term.strip_comb t of - (@{const True}, []) => t - | (@{const False}, []) => t - | (u as Const (@{const_name If}, _), [t1, t2, t3]) => + (\<^const>\True\, []) => t + | (\<^const>\False\, []) => t + | (u as Const (\<^const_name>\If\, _), [t1, t2, t3]) => if pat then raise BAD_PATTERN () else u $ in_form t1 $ in_term pat t2 $ in_term pat t3 | (Const (c as (n, _)), ts) => if is_builtin_conn_or_pred ctxt c ts orelse is_quant n then @@ -355,16 +355,16 @@ | (Free c, ts) => Term.list_comb (Free c, map (in_term pat) ts) | _ => t) - and in_pat ((p as Const (@{const_name pat}, _)) $ t) = + and in_pat ((p as Const (\<^const_name>\pat\, _)) $ t) = p $ in_term true t - | in_pat ((p as Const (@{const_name nopat}, _)) $ t) = + | in_pat ((p as Const (\<^const_name>\nopat\, _)) $ t) = p $ in_term true t | in_pat t = raise TERM ("bad pattern", [t]) and in_pats ps = - in_list @{typ "pattern symb_list"} (SOME o in_list @{typ pattern} (try in_pat)) ps + in_list \<^typ>\pattern symb_list\ (SOME o in_list \<^typ>\pattern\ (try in_pat)) ps - and in_trigger ((c as @{const trigger}) $ p $ t) = c $ in_pats p $ in_form t + and in_trigger ((c as \<^const>\trigger\) $ p $ t) = c $ in_pats p $ in_form t | in_trigger t = in_form t and in_form t = @@ -393,16 +393,16 @@ (** utility functions **) val quantifier = (fn - @{const_name All} => SOME SForall - | @{const_name Ex} => SOME SExists + \<^const_name>\All\ => SOME SForall + | \<^const_name>\Ex\ => SOME SExists | _ => NONE) fun group_quant qname Ts (t as Const (q, _) $ Abs (_, T, u)) = if q = qname then group_quant qname (T :: Ts) u else (Ts, t) | group_quant _ Ts t = (Ts, t) -fun dest_pat (Const (@{const_name pat}, _) $ t) = (t, true) - | dest_pat (Const (@{const_name nopat}, _) $ t) = (t, false) +fun dest_pat (Const (\<^const_name>\pat\, _) $ t) = (t, true) + | dest_pat (Const (\<^const_name>\nopat\, _) $ t) = (t, false) | dest_pat t = raise TERM ("bad pattern", [t]) fun dest_pats [] = I @@ -412,7 +412,7 @@ | (ps, [false]) => cons (SNoPat ps) | _ => raise TERM ("bad multi-pattern", ts)) -fun dest_trigger (@{const trigger} $ tl $ t) = +fun dest_trigger (\<^const>\trigger\ $ tl $ t) = (rev (fold (dest_pats o SMT_Util.dest_symb_list) (SMT_Util.dest_symb_list tl) []), t) | dest_trigger t = ([], t) @@ -501,17 +501,17 @@ ((empty_tr_context, ctxt), ts1) |-> (if null fp_kinds then no_dtyps else collect_co_datatypes fp_kinds) - fun is_binder (Const (@{const_name Let}, _) $ _) = true + fun is_binder (Const (\<^const_name>\Let\, _) $ _) = true | is_binder t = Lambda_Lifting.is_quantifier t - fun mk_trigger ((q as Const (@{const_name All}, _)) $ Abs (n, T, t)) = + fun mk_trigger ((q as Const (\<^const_name>\All\, _)) $ Abs (n, T, t)) = q $ Abs (n, T, mk_trigger t) - | mk_trigger (eq as (Const (@{const_name HOL.eq}, T) $ lhs $ _)) = - Term.domain_type T --> @{typ pattern} - |> (fn T => Const (@{const_name pat}, T) $ lhs) - |> SMT_Util.mk_symb_list @{typ pattern} o single - |> SMT_Util.mk_symb_list @{typ "pattern symb_list"} o single - |> (fn t => @{const trigger} $ t $ eq) + | mk_trigger (eq as (Const (\<^const_name>\HOL.eq\, T) $ lhs $ _)) = + Term.domain_type T --> \<^typ>\pattern\ + |> (fn T => Const (\<^const_name>\pat\, T) $ lhs) + |> SMT_Util.mk_symb_list \<^typ>\pattern\ o single + |> SMT_Util.mk_symb_list \<^typ>\pattern symb_list\ o single + |> (fn t => \<^const>\trigger\ $ t $ eq) | mk_trigger t = t val (ctxt2, (ts3, ll_defs)) = diff -r d24dcac5eb4c -r e61557884799 src/HOL/Tools/arith_data.ML --- a/src/HOL/Tools/arith_data.ML Wed Dec 06 19:34:59 2017 +0100 +++ b/src/HOL/Tools/arith_data.ML Wed Dec 06 20:43:09 2017 +0100 @@ -44,11 +44,11 @@ val _ = Theory.setup - (Method.setup @{binding arith} + (Method.setup \<^binding>\arith\ (Scan.succeed (fn ctxt => METHOD (fn facts => HEADGOAL - (Method.insert_tac ctxt (rev (Named_Theorems.get ctxt @{named_theorems arith}) @ facts) + (Method.insert_tac ctxt (rev (Named_Theorems.get ctxt \<^named_theorems>\arith\) @ facts) THEN' arith_tac ctxt)))) "various arithmetic decision procedures"); @@ -58,11 +58,11 @@ fun mk_number T 1 = HOLogic.numeral_const T $ HOLogic.one_const | mk_number T n = HOLogic.mk_number T n; -val mk_plus = HOLogic.mk_binop @{const_name Groups.plus}; +val mk_plus = HOLogic.mk_binop \<^const_name>\Groups.plus\; fun mk_minus t = let val T = Term.fastype_of t - in Const (@{const_name Groups.uminus}, T --> T) $ t end; + in Const (\<^const_name>\Groups.uminus\, T --> T) $ t end; (*Thus mk_sum[t] yields t+0; longer sums don't have a trailing zero*) fun mk_sum T [] = mk_number T 0 @@ -74,9 +74,9 @@ | long_mk_sum T (t :: ts) = mk_plus (t, long_mk_sum T ts); (*decompose additions AND subtractions as a sum*) -fun dest_summing (pos, Const (@{const_name Groups.plus}, _) $ t $ u, ts) = +fun dest_summing (pos, Const (\<^const_name>\Groups.plus\, _) $ t $ u, ts) = dest_summing (pos, t, dest_summing (pos, u, ts)) - | dest_summing (pos, Const (@{const_name Groups.minus}, _) $ t $ u, ts) = + | dest_summing (pos, Const (\<^const_name>\Groups.minus\, _) $ t $ u, ts) = dest_summing (pos, t, dest_summing (not pos, u, ts)) | dest_summing (pos, t, ts) = (if pos then t else mk_minus t) :: ts; diff -r d24dcac5eb4c -r e61557884799 src/HOL/Tools/boolean_algebra_cancel.ML --- a/src/HOL/Tools/boolean_algebra_cancel.ML Wed Dec 06 19:34:59 2017 +0100 +++ b/src/HOL/Tools/boolean_algebra_cancel.ML Wed Dec 06 20:43:09 2017 +0100 @@ -27,17 +27,17 @@ fun move_to_front rule path = Conv.rewr_conv (Library.foldl (op RS) (rule, path)) -fun add_atoms sup pos path (t as Const (@{const_name Lattices.sup}, _) $ x $ y) = +fun add_atoms sup pos path (t as Const (\<^const_name>\Lattices.sup\, _) $ x $ y) = if sup then add_atoms sup pos (sup1::path) x #> add_atoms sup pos (sup2::path) y else cons ((pos, t), path) - | add_atoms sup pos path (t as Const (@{const_name Lattices.inf}, _) $ x $ y) = + | add_atoms sup pos path (t as Const (\<^const_name>\Lattices.inf\, _) $ x $ y) = if not sup then add_atoms sup pos (inf1::path) x #> add_atoms sup pos (inf2::path) y else cons ((pos, t), path) - | add_atoms _ _ _ (Const (@{const_name Orderings.bot}, _)) = I - | add_atoms _ _ _ (Const (@{const_name Orderings.top}, _)) = I - | add_atoms _ pos path (Const (@{const_name Groups.uminus}, _) $ x) = cons ((not pos, x), path) + | add_atoms _ _ _ (Const (\<^const_name>\Orderings.bot\, _)) = I + | add_atoms _ _ _ (Const (\<^const_name>\Orderings.top\, _)) = I + | add_atoms _ pos path (Const (\<^const_name>\Groups.uminus\, _) $ x) = cons ((not pos, x), path) | add_atoms _ pos path x = cons ((pos, x), path); fun atoms sup pos t = add_atoms sup pos [] t [] diff -r d24dcac5eb4c -r e61557884799 src/HOL/Tools/choice_specification.ML --- a/src/HOL/Tools/choice_specification.ML Wed Dec 06 19:34:59 2017 +0100 +++ b/src/HOL/Tools/choice_specification.ML Wed Dec 06 20:43:09 2017 +0100 @@ -18,7 +18,7 @@ fun mk_definitional [] arg = arg | mk_definitional ((thname, cname, covld) :: cos) (thy, thm) = (case HOLogic.dest_Trueprop (Thm.concl_of thm) of - Const (@{const_name Ex}, _) $ P => + Const (\<^const_name>\Ex\, _) $ P => let val ctype = domain_type (type_of P) val cname_full = Sign.intern_const thy cname @@ -90,7 +90,7 @@ fun proc_single prop = let val frees = Misc_Legacy.term_frees prop - val _ = forall (fn v => Sign.of_sort thy (type_of v, @{sort type})) frees + val _ = forall (fn v => Sign.of_sort thy (type_of v, \<^sort>\type\)) frees orelse error "Specificaton: Only free variables of sort 'type' allowed" val prop_closed = close_form prop in (prop_closed, frees) end @@ -194,12 +194,12 @@ (* outer syntax *) -val opt_name = Scan.optional (Parse.name --| @{keyword ":"}) "" +val opt_name = Scan.optional (Parse.name --| \<^keyword>\:\) "" val opt_overloaded = Parse.opt_keyword "overloaded" val _ = - Outer_Syntax.command @{command_keyword specification} "define constants by specification" - (@{keyword "("} |-- Scan.repeat1 (opt_name -- Parse.term -- opt_overloaded) --| @{keyword ")"} -- + Outer_Syntax.command \<^command_keyword>\specification\ "define constants by specification" + (\<^keyword>\(\ |-- Scan.repeat1 (opt_name -- Parse.term -- opt_overloaded) --| \<^keyword>\)\ -- Scan.repeat1 ((Parse_Spec.opt_thm_name ":" >> apfst Binding.name_of) -- Parse.prop) >> (fn (cos, alt_props) => Toplevel.theory_to_proof (process_spec cos alt_props))) diff -r d24dcac5eb4c -r e61557884799 src/HOL/Tools/cnf.ML --- a/src/HOL/Tools/cnf.ML Wed Dec 06 19:34:59 2017 +0100 +++ b/src/HOL/Tools/cnf.ML Wed Dec 06 20:43:09 2017 +0100 @@ -94,19 +94,19 @@ val cnftac_eq_imp = @{lemma "[| P = Q; P |] ==> Q" by auto}; -fun is_atom (Const (@{const_name False}, _)) = false - | is_atom (Const (@{const_name True}, _)) = false - | is_atom (Const (@{const_name HOL.conj}, _) $ _ $ _) = false - | is_atom (Const (@{const_name HOL.disj}, _) $ _ $ _) = false - | is_atom (Const (@{const_name HOL.implies}, _) $ _ $ _) = false - | is_atom (Const (@{const_name HOL.eq}, Type ("fun", @{typ bool} :: _)) $ _ $ _) = false - | is_atom (Const (@{const_name Not}, _) $ _) = false +fun is_atom (Const (\<^const_name>\False\, _)) = false + | is_atom (Const (\<^const_name>\True\, _)) = false + | is_atom (Const (\<^const_name>\HOL.conj\, _) $ _ $ _) = false + | is_atom (Const (\<^const_name>\HOL.disj\, _) $ _ $ _) = false + | is_atom (Const (\<^const_name>\HOL.implies\, _) $ _ $ _) = false + | is_atom (Const (\<^const_name>\HOL.eq\, Type ("fun", \<^typ>\bool\ :: _)) $ _ $ _) = false + | is_atom (Const (\<^const_name>\Not\, _) $ _) = false | is_atom _ = true; -fun is_literal (Const (@{const_name Not}, _) $ x) = is_atom x +fun is_literal (Const (\<^const_name>\Not\, _) $ x) = is_atom x | is_literal x = is_atom x; -fun is_clause (Const (@{const_name HOL.disj}, _) $ x $ y) = is_clause x andalso is_clause y +fun is_clause (Const (\<^const_name>\HOL.disj\, _) $ x $ y) = is_clause x andalso is_clause y | is_clause x = is_literal x; (* ------------------------------------------------------------------------- *) @@ -116,7 +116,7 @@ fun clause_is_trivial c = let - fun dual (Const (@{const_name Not}, _) $ x) = x + fun dual (Const (\<^const_name>\Not\, _) $ x) = x | dual x = HOLogic.Not $ x fun has_duals [] = false | has_duals (x::xs) = member (op =) xs (dual x) orelse has_duals xs @@ -178,28 +178,28 @@ (* eliminated (possibly causing an exponential blowup) *) (* ------------------------------------------------------------------------- *) -fun make_nnf_thm thy (Const (@{const_name HOL.conj}, _) $ x $ y) = +fun make_nnf_thm thy (Const (\<^const_name>\HOL.conj\, _) $ x $ y) = let val thm1 = make_nnf_thm thy x val thm2 = make_nnf_thm thy y in conj_cong OF [thm1, thm2] end - | make_nnf_thm thy (Const (@{const_name HOL.disj}, _) $ x $ y) = + | make_nnf_thm thy (Const (\<^const_name>\HOL.disj\, _) $ x $ y) = let val thm1 = make_nnf_thm thy x val thm2 = make_nnf_thm thy y in disj_cong OF [thm1, thm2] end - | make_nnf_thm thy (Const (@{const_name HOL.implies}, _) $ x $ y) = + | make_nnf_thm thy (Const (\<^const_name>\HOL.implies\, _) $ x $ y) = let val thm1 = make_nnf_thm thy (HOLogic.Not $ x) val thm2 = make_nnf_thm thy y in make_nnf_imp OF [thm1, thm2] end - | make_nnf_thm thy (Const (@{const_name HOL.eq}, Type ("fun", @{typ bool} :: _)) $ x $ y) = + | make_nnf_thm thy (Const (\<^const_name>\HOL.eq\, Type ("fun", \<^typ>\bool\ :: _)) $ x $ y) = let val thm1 = make_nnf_thm thy x val thm2 = make_nnf_thm thy (HOLogic.Not $ x) @@ -208,18 +208,18 @@ in make_nnf_iff OF [thm1, thm2, thm3, thm4] end - | make_nnf_thm _ (Const (@{const_name Not}, _) $ Const (@{const_name False}, _)) = + | make_nnf_thm _ (Const (\<^const_name>\Not\, _) $ Const (\<^const_name>\False\, _)) = make_nnf_not_false - | make_nnf_thm _ (Const (@{const_name Not}, _) $ Const (@{const_name True}, _)) = + | make_nnf_thm _ (Const (\<^const_name>\Not\, _) $ Const (\<^const_name>\True\, _)) = make_nnf_not_true - | make_nnf_thm thy (Const (@{const_name Not}, _) $ (Const (@{const_name HOL.conj}, _) $ x $ y)) = + | make_nnf_thm thy (Const (\<^const_name>\Not\, _) $ (Const (\<^const_name>\HOL.conj\, _) $ x $ y)) = let val thm1 = make_nnf_thm thy (HOLogic.Not $ x) val thm2 = make_nnf_thm thy (HOLogic.Not $ y) in make_nnf_not_conj OF [thm1, thm2] end - | make_nnf_thm thy (Const (@{const_name Not}, _) $ (Const (@{const_name HOL.disj}, _) $ x $ y)) = + | make_nnf_thm thy (Const (\<^const_name>\Not\, _) $ (Const (\<^const_name>\HOL.disj\, _) $ x $ y)) = let val thm1 = make_nnf_thm thy (HOLogic.Not $ x) val thm2 = make_nnf_thm thy (HOLogic.Not $ y) @@ -227,8 +227,7 @@ make_nnf_not_disj OF [thm1, thm2] end | make_nnf_thm thy - (Const (@{const_name Not}, _) $ - (Const (@{const_name HOL.implies}, _) $ x $ y)) = + (Const (\<^const_name>\Not\, _) $ (Const (\<^const_name>\HOL.implies\, _) $ x $ y)) = let val thm1 = make_nnf_thm thy x val thm2 = make_nnf_thm thy (HOLogic.Not $ y) @@ -236,8 +235,8 @@ make_nnf_not_imp OF [thm1, thm2] end | make_nnf_thm thy - (Const (@{const_name Not}, _) $ - (Const (@{const_name HOL.eq}, Type ("fun", @{typ bool} :: _)) $ x $ y)) = + (Const (\<^const_name>\Not\, _) $ + (Const (\<^const_name>\HOL.eq\, Type ("fun", \<^typ>\bool\ :: _)) $ x $ y)) = let val thm1 = make_nnf_thm thy x val thm2 = make_nnf_thm thy (HOLogic.Not $ x) @@ -246,7 +245,7 @@ in make_nnf_not_iff OF [thm1, thm2, thm3, thm4] end - | make_nnf_thm thy (Const (@{const_name Not}, _) $ (Const (@{const_name Not}, _) $ x)) = + | make_nnf_thm thy (Const (\<^const_name>\Not\, _) $ (Const (\<^const_name>\Not\, _) $ x)) = let val thm1 = make_nnf_thm thy x in @@ -289,45 +288,45 @@ (* Theory.theory -> Term.term -> Thm.thm *) -fun simp_True_False_thm thy (Const (@{const_name HOL.conj}, _) $ x $ y) = +fun simp_True_False_thm thy (Const (\<^const_name>\HOL.conj\, _) $ x $ y) = let val thm1 = simp_True_False_thm thy x val x'= (snd o HOLogic.dest_eq o HOLogic.dest_Trueprop o Thm.prop_of) thm1 in - if x' = @{term False} then + if x' = \<^term>\False\ then simp_TF_conj_False_l OF [thm1] (* (x & y) = False *) else let val thm2 = simp_True_False_thm thy y val y' = (snd o HOLogic.dest_eq o HOLogic.dest_Trueprop o Thm.prop_of) thm2 in - if x' = @{term True} then + if x' = \<^term>\True\ then simp_TF_conj_True_l OF [thm1, thm2] (* (x & y) = y' *) - else if y' = @{term False} then + else if y' = \<^term>\False\ then simp_TF_conj_False_r OF [thm2] (* (x & y) = False *) - else if y' = @{term True} then + else if y' = \<^term>\True\ then simp_TF_conj_True_r OF [thm1, thm2] (* (x & y) = x' *) else conj_cong OF [thm1, thm2] (* (x & y) = (x' & y') *) end end - | simp_True_False_thm thy (Const (@{const_name HOL.disj}, _) $ x $ y) = + | simp_True_False_thm thy (Const (\<^const_name>\HOL.disj\, _) $ x $ y) = let val thm1 = simp_True_False_thm thy x val x' = (snd o HOLogic.dest_eq o HOLogic.dest_Trueprop o Thm.prop_of) thm1 in - if x' = @{term True} then + if x' = \<^term>\True\ then simp_TF_disj_True_l OF [thm1] (* (x | y) = True *) else let val thm2 = simp_True_False_thm thy y val y' = (snd o HOLogic.dest_eq o HOLogic.dest_Trueprop o Thm.prop_of) thm2 in - if x' = @{term False} then + if x' = \<^term>\False\ then simp_TF_disj_False_l OF [thm1, thm2] (* (x | y) = y' *) - else if y' = @{term True} then + else if y' = \<^term>\True\ then simp_TF_disj_True_r OF [thm2] (* (x | y) = True *) - else if y' = @{term False} then + else if y' = \<^term>\False\ then simp_TF_disj_False_r OF [thm1, thm2] (* (x | y) = x' *) else disj_cong OF [thm1, thm2] (* (x | y) = (x' | y') *) @@ -344,24 +343,24 @@ fun make_cnf_thm ctxt t = let val thy = Proof_Context.theory_of ctxt - fun make_cnf_thm_from_nnf (Const (@{const_name HOL.conj}, _) $ x $ y) = + fun make_cnf_thm_from_nnf (Const (\<^const_name>\HOL.conj\, _) $ x $ y) = let val thm1 = make_cnf_thm_from_nnf x val thm2 = make_cnf_thm_from_nnf y in conj_cong OF [thm1, thm2] end - | make_cnf_thm_from_nnf (Const (@{const_name HOL.disj}, _) $ x $ y) = + | make_cnf_thm_from_nnf (Const (\<^const_name>\HOL.disj\, _) $ x $ y) = let (* produces a theorem "(x' | y') = t'", where x', y', and t' are in CNF *) - fun make_cnf_disj_thm (Const (@{const_name HOL.conj}, _) $ x1 $ x2) y' = + fun make_cnf_disj_thm (Const (\<^const_name>\HOL.conj\, _) $ x1 $ x2) y' = let val thm1 = make_cnf_disj_thm x1 y' val thm2 = make_cnf_disj_thm x2 y' in make_cnf_disj_conj_l OF [thm1, thm2] (* ((x1 & x2) | y') = ((x1 | y')' & (x2 | y')') *) end - | make_cnf_disj_thm x' (Const (@{const_name HOL.conj}, _) $ y1 $ y2) = + | make_cnf_disj_thm x' (Const (\<^const_name>\HOL.conj\, _) $ y1 $ y2) = let val thm1 = make_cnf_disj_thm x' y1 val thm2 = make_cnf_disj_thm x' y2 @@ -417,35 +416,35 @@ val var_id = Unsynchronized.ref 0 (* properly initialized below *) fun new_free () = Free ("cnfx_" ^ string_of_int (Unsynchronized.inc var_id), HOLogic.boolT) - fun make_cnfx_thm_from_nnf (Const (@{const_name HOL.conj}, _) $ x $ y) : thm = + fun make_cnfx_thm_from_nnf (Const (\<^const_name>\HOL.conj\, _) $ x $ y) : thm = let val thm1 = make_cnfx_thm_from_nnf x val thm2 = make_cnfx_thm_from_nnf y in conj_cong OF [thm1, thm2] end - | make_cnfx_thm_from_nnf (Const (@{const_name HOL.disj}, _) $ x $ y) = + | make_cnfx_thm_from_nnf (Const (\<^const_name>\HOL.disj\, _) $ x $ y) = if is_clause x andalso is_clause y then inst_thm thy [HOLogic.mk_disj (x, y)] iff_refl else if is_literal y orelse is_literal x then let (* produces a theorem "(x' | y') = t'", where x', y', and t' are *) (* almost in CNF, and x' or y' is a literal *) - fun make_cnfx_disj_thm (Const (@{const_name HOL.conj}, _) $ x1 $ x2) y' = + fun make_cnfx_disj_thm (Const (\<^const_name>\HOL.conj\, _) $ x1 $ x2) y' = let val thm1 = make_cnfx_disj_thm x1 y' val thm2 = make_cnfx_disj_thm x2 y' in make_cnf_disj_conj_l OF [thm1, thm2] (* ((x1 & x2) | y') = ((x1 | y')' & (x2 | y')') *) end - | make_cnfx_disj_thm x' (Const (@{const_name HOL.conj}, _) $ y1 $ y2) = + | make_cnfx_disj_thm x' (Const (\<^const_name>\HOL.conj\, _) $ y1 $ y2) = let val thm1 = make_cnfx_disj_thm x' y1 val thm2 = make_cnfx_disj_thm x' y2 in make_cnf_disj_conj_r OF [thm1, thm2] (* (x' | (y1 & y2)) = ((x' | y1)' & (x' | y2)') *) end - | make_cnfx_disj_thm (@{term "Ex::(bool => bool) => bool"} $ x') y' = + | make_cnfx_disj_thm (\<^term>\Ex :: (bool \ bool) \ bool\ $ x') y' = let val thm1 = inst_thm thy [x', y'] make_cnfx_disj_ex_l (* ((Ex x') | y') = (Ex (x' | y')) *) val var = new_free () @@ -456,7 +455,7 @@ in iff_trans OF [thm1, thm5] (* ((Ex x') | y') = (Ex v. body') *) end - | make_cnfx_disj_thm x' (@{term "Ex::(bool => bool) => bool"} $ y') = + | make_cnfx_disj_thm x' (\<^term>\Ex :: (bool \ bool) \ bool\ $ y') = let val thm1 = inst_thm thy [x', y'] make_cnfx_disj_ex_r (* (x' | (Ex y')) = (Ex (x' | y')) *) val var = new_free () diff -r d24dcac5eb4c -r e61557884799 src/HOL/Tools/code_evaluation.ML --- a/src/HOL/Tools/code_evaluation.ML Wed Dec 06 19:34:59 2017 +0100 +++ b/src/HOL/Tools/code_evaluation.ML Wed Dec 06 20:43:09 2017 +0100 @@ -23,17 +23,16 @@ fun add_term_of_inst tyco thy = let val ((raw_vs, _), _) = Code.get_type thy tyco; - val vs = map (fn (v, _) => (v, @{sort typerep})) raw_vs; + val vs = map (fn (v, _) => (v, \<^sort>\typerep\)) raw_vs; val ty = Type (tyco, map TFree vs); - val lhs = Const (@{const_name term_of}, ty --> @{typ term}) - $ Free ("x", ty); - val rhs = @{term "undefined :: term"}; + val lhs = Const (\<^const_name>\term_of\, ty --> \<^typ>\term\) $ Free ("x", ty); + val rhs = \<^term>\undefined :: term\; val eq = HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs)); fun triv_name_of t = (fst o dest_Free o fst o strip_comb o fst o HOLogic.dest_eq o HOLogic.dest_Trueprop) t ^ "_triv"; in thy - |> Class.instantiation ([tyco], vs, @{sort term_of}) + |> Class.instantiation ([tyco], vs, \<^sort>\term_of\) |> `(fn lthy => Syntax.check_term lthy eq) |-> (fn eq => Specification.definition NONE [] [] ((Binding.name (triv_name_of eq), []), eq)) |> snd @@ -42,15 +41,15 @@ fun ensure_term_of_inst tyco thy = let - val need_inst = not (Sorts.has_instance (Sign.classes_of thy) tyco @{sort term_of}) - andalso Sorts.has_instance (Sign.classes_of thy) tyco @{sort typerep}; + val need_inst = not (Sorts.has_instance (Sign.classes_of thy) tyco \<^sort>\term_of\) + andalso Sorts.has_instance (Sign.classes_of thy) tyco \<^sort>\typerep\; in if need_inst then add_term_of_inst tyco thy else thy end; fun for_term_of_instance tyco vs f thy = let val algebra = Sign.classes_of thy; in - case try (Sorts.mg_domain algebra tyco) @{sort term_of} of + case try (Sorts.mg_domain algebra tyco) \<^sort>\term_of\ of NONE => thy | SOME sorts => f tyco (map2 (fn (v, sort) => fn sort' => (v, Sorts.inter_sort algebra (sort, sort'))) vs sorts) thy @@ -95,7 +94,7 @@ fun mk_abs_term_of_eq thy ty abs ty_rep proj = let val arg = Var (("x", 0), ty); - val rhs = Abs ("y", @{typ term}, HOLogic.reflect_term (Const (abs, ty_rep --> ty) $ Bound 0)) $ + val rhs = Abs ("y", \<^typ>\term\, HOLogic.reflect_term (Const (abs, ty_rep --> ty) $ Bound 0)) $ (HOLogic.mk_term_of ty_rep (Const (proj, ty --> ty_rep) $ arg)) |> Thm.global_cterm_of thy; val cty = Thm.global_ctyp_of thy ty; @@ -139,7 +138,7 @@ else NONE end; -fun subst_termify_app (Const (@{const_name termify}, _), [t]) = +fun subst_termify_app (Const (\<^const_name>\termify\, _), [t]) = if not (Term.exists_subterm (fn Abs _ => true | _ => false) t) then if fold_aterms (fn Const _ => I | _ => K false) t true then SOME (HOLogic.reflect_term t) diff -r d24dcac5eb4c -r e61557884799 src/HOL/Tools/coinduction.ML --- a/src/HOL/Tools/coinduction.ML Wed Dec 06 19:34:59 2017 +0100 +++ b/src/HOL/Tools/coinduction.ML Wed Dec 06 20:43:09 2017 +0100 @@ -82,7 +82,7 @@ names Ts raw_eqs; val phi = eqs @ map (HOLogic.dest_Trueprop o Thm.prop_of) prems |> try (Library.foldr1 HOLogic.mk_conj) - |> the_default @{term True} + |> the_default \<^term>\True\ |> Ctr_Sugar_Util.list_exists_free vars |> Term.map_abs_vars (Variable.revert_fixed ctxt) |> fold_rev Term.absfree (names ~~ Ts) @@ -159,7 +159,7 @@ val _ = Theory.setup - (Method.setup @{binding coinduction} + (Method.setup \<^binding>\coinduction\ (arbitrary -- Scan.option coinduct_rules >> (fn (arbitrary, opt_rules) => fn _ => fn facts => Seq.DETERM (coinduction_context_tactic arbitrary opt_rules facts 1))) diff -r d24dcac5eb4c -r e61557884799 src/HOL/Tools/functor.ML --- a/src/HOL/Tools/functor.ML Wed Dec 06 19:34:59 2017 +0100 +++ b/src/HOL/Tools/functor.ML Wed Dec 06 20:43:09 2017 +0100 @@ -275,8 +275,8 @@ val functor_cmd = gen_functor Syntax.read_term; val _ = - Outer_Syntax.local_theory_to_proof @{command_keyword functor} + Outer_Syntax.local_theory_to_proof \<^command_keyword>\functor\ "register operations managing the functorial structure of a type" - (Scan.option (Parse.name --| @{keyword ":"}) -- Parse.term >> uncurry functor_cmd); + (Scan.option (Parse.name --| \<^keyword>\:\) -- Parse.term >> uncurry functor_cmd); end; diff -r d24dcac5eb4c -r e61557884799 src/HOL/Tools/groebner.ML --- a/src/HOL/Tools/groebner.ML Wed Dec 06 19:34:59 2017 +0100 +++ b/src/HOL/Tools/groebner.ML Wed Dec 06 20:43:09 2017 +0100 @@ -327,7 +327,7 @@ (* Produce Strong Nullstellensatz certificate for a power of pol. *) fun grobner_strong vars pols pol = - let val vars' = @{cterm "True"}::vars + let val vars' = \<^cterm>\True\::vars val grob_z = [(@1, 1::(map (K 0) vars))] val grob_1 = [(@1, (map (K 0) vars'))] fun augment p= map (fn (c,m) => (c,0::m)) p @@ -349,7 +349,7 @@ fun refute_disj rfn tm = case Thm.term_of tm of - Const(@{const_name HOL.disj},_)$_$_ => + Const(\<^const_name>\HOL.disj\,_)$_$_ => Drule.compose (refute_disj rfn (Thm.dest_arg tm), 2, Drule.compose (refute_disj rfn (Thm.dest_arg1 tm), 2, disjE)) @@ -360,11 +360,11 @@ fun is_neg t = case Thm.term_of t of - (Const(@{const_name Not},_)$_) => true + (Const(\<^const_name>\Not\,_)$_) => true | _ => false; fun is_eq t = case Thm.term_of t of - (Const(@{const_name HOL.eq},_)$_$_) => true + (Const(\<^const_name>\HOL.eq\,_)$_$_) => true | _ => false; fun end_itlist f l = @@ -385,7 +385,7 @@ val strip_exists = let fun h (acc, t) = case Thm.term_of t of - Const (@{const_name Ex}, _) $ Abs _ => + Const (\<^const_name>\Ex\, _) $ Abs _ => h (Thm.dest_abs NONE (Thm.dest_arg t) |>> (fn v => v::acc)) | _ => (acc,t) in fn t => h ([],t) @@ -393,7 +393,7 @@ fun is_forall t = case Thm.term_of t of - (Const(@{const_name All},_)$Abs(_,_,_)) => true + (Const(\<^const_name>\All\,_)$Abs(_,_,_)) => true | _ => false; val mk_object_eq = fn th => th COMP meta_eq_to_obj_eq; @@ -412,9 +412,9 @@ val specl = fold_rev (fn x => fn th => Thm.instantiate' [] [SOME x] (th RS spec)); -val cTrp = @{cterm "Trueprop"}; -val cConj = @{cterm HOL.conj}; -val (cNot,false_tm) = (@{cterm "Not"}, @{cterm "False"}); +val cTrp = \<^cterm>\Trueprop\; +val cConj = \<^cterm>\HOL.conj\; +val (cNot,false_tm) = (\<^cterm>\Not\, \<^cterm>\False\); val assume_Trueprop = Thm.apply cTrp #> Thm.assume; val list_mk_conj = list_mk_binop cConj; val conjs = list_dest_binop cConj; @@ -438,7 +438,7 @@ (* FIXME : copied from cqe.ML -- complex QE*) fun conjuncts ct = case Thm.term_of ct of - @{term HOL.conj}$_$_ => (Thm.dest_arg1 ct)::(conjuncts (Thm.dest_arg ct)) + \<^term>\HOL.conj\$_$_ => (Thm.dest_arg1 ct)::(conjuncts (Thm.dest_arg ct)) | _ => [ct]; fun fold1 f = foldr1 (uncurry f); @@ -446,12 +446,12 @@ fun mk_conj_tab th = let fun h acc th = case Thm.prop_of th of - @{term "Trueprop"}$(@{term HOL.conj}$_$_) => + \<^term>\Trueprop\$(\<^term>\HOL.conj\$_$_) => h (h acc (th RS conjunct2)) (th RS conjunct1) - | @{term "Trueprop"}$p => (p,th)::acc + | \<^term>\Trueprop\$p => (p,th)::acc in fold (Termtab.insert Thm.eq_thm) (h [] th) Termtab.empty end; -fun is_conj (@{term HOL.conj}$_$_) = true +fun is_conj (\<^term>\HOL.conj\$_$_) = true | is_conj _ = false; fun prove_conj tab cjs = @@ -462,8 +462,8 @@ fun conj_ac_rule eq = let val (l,r) = Thm.dest_equals eq - val ctabl = mk_conj_tab (Thm.assume (Thm.apply @{cterm Trueprop} l)) - val ctabr = mk_conj_tab (Thm.assume (Thm.apply @{cterm Trueprop} r)) + val ctabl = mk_conj_tab (Thm.assume (Thm.apply \<^cterm>\Trueprop\ l)) + val ctabr = mk_conj_tab (Thm.assume (Thm.apply \<^cterm>\Trueprop\ r)) fun tabl c = the (Termtab.lookup ctabl (Thm.term_of c)) fun tabr c = the (Termtab.lookup ctabr (Thm.term_of c)) val thl = prove_conj tabl (conjuncts r) |> implies_intr_hyps @@ -475,24 +475,24 @@ (* Conversion for the equivalence of existential statements where EX quantifiers are rearranged differently *) -fun ext ctxt T = Thm.cterm_of ctxt (Const (@{const_name Ex}, (T --> @{typ bool}) --> @{typ bool})) +fun ext ctxt T = Thm.cterm_of ctxt (Const (\<^const_name>\Ex\, (T --> \<^typ>\bool\) --> \<^typ>\bool\)) fun mk_ex ctxt v t = Thm.apply (ext ctxt (Thm.typ_of_cterm v)) (Thm.lambda v t) fun choose v th th' = case Thm.concl_of th of - @{term Trueprop} $ (Const(@{const_name Ex},_)$_) => + \<^term>\Trueprop\ $ (Const(\<^const_name>\Ex\,_)$_) => let val p = (funpow 2 Thm.dest_arg o Thm.cprop_of) th val T = (hd o Thm.dest_ctyp o Thm.ctyp_of_cterm) p val th0 = Conv.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)) + (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 | _ => error "" (* FIXME ? *) fun simple_choose ctxt v th = - choose v (Thm.assume ((Thm.apply @{cterm Trueprop} o mk_ex ctxt v) + choose v (Thm.assume ((Thm.apply \<^cterm>\Trueprop\ o mk_ex ctxt v) (Thm.dest_arg (hd (Thm.chyps_of th))))) th @@ -509,7 +509,7 @@ val (p0,q0) = Thm.dest_binop t val (vs',P) = strip_exists p0 val (vs,_) = strip_exists q0 - val th = Thm.assume (Thm.apply @{cterm Trueprop} P) + val th = Thm.assume (Thm.apply \<^cterm>\Trueprop\ P) val th1 = implies_intr_hyps (fold (simple_choose ctxt) vs' (fold mkexi vs th)) val th2 = implies_intr_hyps (fold (simple_choose ctxt) vs (fold mkexi vs' th)) val p = (Thm.dest_arg o Thm.dest_arg1 o Thm.cprop_of) th1 @@ -523,7 +523,7 @@ Free(s,_) => s | Var ((s,_),_) => s | _ => "x" - fun mk_eq s t = Thm.apply (Thm.apply @{cterm "op == :: bool => _"} s) t + fun mk_eq s t = Thm.apply (Thm.apply \<^cterm>\op \ :: bool \ _\ s) t fun mk_exists ctxt v th = Drule.arg_cong_rule (ext ctxt (Thm.typ_of_cterm v)) (Thm.abstract_rule (getname v) v th) fun simp_ex_conv ctxt = @@ -552,12 +552,12 @@ val (ring_sub_tm, ring_neg_tm) = (case r_ops of [sub_pat, neg_pat] => (Thm.dest_fun2 sub_pat, Thm.dest_fun neg_pat) - |_ => (@{cterm "True"}, @{cterm "True"})); + |_ => (\<^cterm>\True\, \<^cterm>\True\)); val (field_div_tm, field_inv_tm) = (case f_ops of [div_pat, inv_pat] => (Thm.dest_fun2 div_pat, Thm.dest_fun inv_pat) - | _ => (@{cterm "True"}, @{cterm "True"})); + | _ => (\<^cterm>\True\, \<^cterm>\True\)); val [idom_thm, neq_thm] = idom; val [idl_sub, idl_add0] = @@ -653,7 +653,7 @@ val holify_polynomial = let fun holify_varpow (v,n) = - if n = 1 then v else ring_mk_pow v (Numeral.mk_cnumber @{ctyp nat} n) (* FIXME *) + if n = 1 then v else ring_mk_pow v (Numeral.mk_cnumber \<^ctyp>\nat\ n) (* FIXME *) fun holify_monomial vars (c,m) = let val xps = map holify_varpow (filter (fn (_,n) => n <> 0) (vars ~~ m)) in end_itlist ring_mk_mul (mk_const c :: xps) @@ -737,7 +737,7 @@ fun mk_forall x p = let val T = Thm.typ_of_cterm x; - val all = Thm.cterm_of ctxt (Const (@{const_name All}, (T --> @{typ bool}) --> @{typ bool})) + val all = Thm.cterm_of ctxt (Const (\<^const_name>\All\, (T --> \<^typ>\bool\) --> \<^typ>\bool\)) in Thm.apply all (Thm.lambda x p) end val avs = Drule.cterm_add_frees tm [] val P' = fold mk_forall avs tm @@ -819,15 +819,15 @@ fun unwind_polys_conv ctxt tm = let val (vars,bod) = strip_exists tm - val cjs = striplist (dest_binary @{cterm HOL.conj}) bod + val cjs = striplist (dest_binary \<^cterm>\HOL.conj\) bod val th1 = (the (get_first (try (isolate_variable vars)) cjs) handle Option.Option => raise CTERM ("unwind_polys_conv",[tm])) val eq = Thm.lhs_of th1 - val bod' = list_mk_binop @{cterm HOL.conj} (eq::(remove op aconvc eq cjs)) + val bod' = list_mk_binop \<^cterm>\HOL.conj\ (eq::(remove op aconvc eq cjs)) val th2 = conj_ac_rule (mk_eq bod bod') val th3 = Thm.transitive th2 - (Drule.binop_cong_rule @{cterm HOL.conj} th1 + (Drule.binop_cong_rule \<^cterm>\HOL.conj\ th1 (Thm.reflexive (Thm.dest_arg (Thm.rhs_of th2)))) val v = Thm.dest_arg1(Thm.dest_arg1(Thm.rhs_of th3)) val th4 = Conv.fconv_rule (Conv.arg_conv (simp_ex_conv ctxt)) (mk_exists ctxt v th3) @@ -890,18 +890,18 @@ fun find_term bounds tm = (case Thm.term_of tm of - Const (@{const_name HOL.eq}, T) $ _ $ _ => + Const (\<^const_name>\HOL.eq\, T) $ _ $ _ => if domain_type T = HOLogic.boolT then find_args bounds tm else Thm.dest_arg tm - | Const (@{const_name Not}, _) $ _ => find_term bounds (Thm.dest_arg tm) - | Const (@{const_name All}, _) $ _ => find_body bounds (Thm.dest_arg tm) - | Const (@{const_name Ex}, _) $ _ => find_body bounds (Thm.dest_arg tm) - | Const (@{const_name HOL.conj}, _) $ _ $ _ => find_args bounds tm - | Const (@{const_name HOL.disj}, _) $ _ $ _ => find_args bounds tm - | Const (@{const_name HOL.implies}, _) $ _ $ _ => find_args bounds tm - | @{term "op ==>"} $_$_ => find_args bounds tm - | Const("op ==",_)$_$_ => find_args bounds tm - | @{term Trueprop}$_ => find_term bounds (Thm.dest_arg tm) + | Const (\<^const_name>\Not\, _) $ _ => find_term bounds (Thm.dest_arg tm) + | Const (\<^const_name>\All\, _) $ _ => find_body bounds (Thm.dest_arg tm) + | Const (\<^const_name>\Ex\, _) $ _ => find_body bounds (Thm.dest_arg tm) + | Const (\<^const_name>\HOL.conj\, _) $ _ $ _ => find_args bounds tm + | Const (\<^const_name>\HOL.disj\, _) $ _ $ _ => find_args bounds tm + | Const (\<^const_name>\HOL.implies\, _) $ _ $ _ => find_args bounds tm + | \<^term>\op \\ $_$_ => find_args bounds tm + | Const("op ==",_)$_$_ => find_args bounds tm (* FIXME proper const name *) + | \<^term>\Trueprop\$_ => find_term bounds (Thm.dest_arg tm) | _ => raise TERM ("find_term", [])) and find_args bounds tm = let val (t, u) = Thm.dest_binop tm @@ -925,7 +925,7 @@ fun presimplify ctxt add_thms del_thms = asm_full_simp_tac (put_simpset HOL_basic_ss ctxt - addsimps (Named_Theorems.get ctxt @{named_theorems algebra}) + addsimps (Named_Theorems.get ctxt \<^named_theorems>\algebra\) delsimps del_thms addsimps add_thms); fun ring_tac add_ths del_ths ctxt = @@ -943,7 +943,7 @@ local fun lhs t = case Thm.term_of t of - Const(@{const_name HOL.eq},_)$_$_ => Thm.dest_arg1 t + Const(\<^const_name>\HOL.eq\,_)$_$_ => Thm.dest_arg1 t | _=> raise CTERM ("ideal_tac - lhs",[t]) fun exitac _ NONE = no_tac | exitac ctxt (SOME y) = diff -r d24dcac5eb4c -r e61557884799 src/HOL/Tools/group_cancel.ML --- a/src/HOL/Tools/group_cancel.ML Wed Dec 06 19:34:59 2017 +0100 +++ b/src/HOL/Tools/group_cancel.ML Wed Dec 06 20:43:09 2017 +0100 @@ -36,13 +36,13 @@ [Conv.rewr_conv (Library.foldl (op RS) (rule0, path)), Conv.arg1_conv (Conv.repeat_conv (Conv.rewr_conv minus_minus))] -fun add_atoms pos path (Const (@{const_name Groups.plus}, _) $ x $ y) = +fun add_atoms pos path (Const (\<^const_name>\Groups.plus\, _) $ x $ y) = add_atoms pos (add1::path) x #> add_atoms pos (add2::path) y - | add_atoms pos path (Const (@{const_name Groups.minus}, _) $ x $ y) = + | add_atoms pos path (Const (\<^const_name>\Groups.minus\, _) $ x $ y) = add_atoms pos (sub1::path) x #> add_atoms (not pos) (sub2::path) y - | add_atoms pos path (Const (@{const_name Groups.uminus}, _) $ x) = + | add_atoms pos path (Const (\<^const_name>\Groups.uminus\, _) $ x) = add_atoms (not pos) (neg1::path) x - | add_atoms _ _ (Const (@{const_name Groups.zero}, _)) = I + | add_atoms _ _ (Const (\<^const_name>\Groups.zero\, _)) = I | add_atoms pos path x = cons ((pos, x), path) fun atoms t = add_atoms true [] t [] diff -r d24dcac5eb4c -r e61557884799 src/HOL/Tools/hologic.ML --- a/src/HOL/Tools/hologic.ML Wed Dec 06 19:34:59 2017 +0100 +++ b/src/HOL/Tools/hologic.ML Wed Dec 06 20:43:09 2017 +0100 @@ -187,16 +187,16 @@ (* logic *) -val Trueprop = Const (@{const_name Trueprop}, boolT --> propT); +val Trueprop = Const (\<^const_name>\Trueprop\, boolT --> propT); fun mk_Trueprop P = Trueprop $ P; -fun dest_Trueprop (Const (@{const_name Trueprop}, _) $ P) = P +fun dest_Trueprop (Const (\<^const_name>\Trueprop\, _) $ P) = P | dest_Trueprop t = raise TERM ("dest_Trueprop", [t]); fun Trueprop_conv cv ct = (case Thm.term_of ct of - Const (@{const_name Trueprop}, _) $ _ => Conv.arg_conv cv ct + Const (\<^const_name>\Trueprop\, _) $ _ => Conv.arg_conv cv ct | _ => raise CTERM ("Trueprop_conv", [ct])) @@ -220,10 +220,10 @@ let val (th1, th2) = conj_elim ctxt th in conj_elims ctxt th1 @ conj_elims ctxt th2 end handle THM _ => [th]; -val conj = @{term HOL.conj} -and disj = @{term HOL.disj} -and imp = @{term implies} -and Not = @{term Not}; +val conj = \<^term>\HOL.conj\ +and disj = \<^term>\HOL.disj\ +and imp = \<^term>\implies\ +and Not = \<^term>\Not\; fun mk_conj (t1, t2) = conj $ t1 $ t2 and mk_disj (t1, t2) = disj $ t1 $ t2 @@ -257,21 +257,21 @@ fun conj_conv cv1 cv2 ct = (case Thm.term_of ct of - Const (@{const_name HOL.conj}, _) $ _ $ _ => + Const (\<^const_name>\HOL.conj\, _) $ _ $ _ => Conv.combination_conv (Conv.arg_conv cv1) cv2 ct | _ => raise CTERM ("conj_conv", [ct])); -fun eq_const T = Const (@{const_name HOL.eq}, T --> T --> boolT); +fun eq_const T = Const (\<^const_name>\HOL.eq\, T --> T --> boolT); fun mk_eq (t, u) = eq_const (fastype_of t) $ t $ u; -fun dest_eq (Const (@{const_name HOL.eq}, _) $ lhs $ rhs) = (lhs, rhs) +fun dest_eq (Const (\<^const_name>\HOL.eq\, _) $ lhs $ rhs) = (lhs, rhs) | dest_eq t = raise TERM ("dest_eq", [t]) fun eq_conv cv1 cv2 ct = (case Thm.term_of ct of - Const (@{const_name HOL.eq}, _) $ _ $ _ => Conv.combination_conv (Conv.arg_conv cv1) cv2 ct + Const (\<^const_name>\HOL.eq\, _) $ _ $ _ => Conv.combination_conv (Conv.arg_conv cv1) cv2 ct | _ => raise CTERM ("eq_conv", [ct])); diff -r d24dcac5eb4c -r e61557884799 src/HOL/Tools/inductive.ML --- a/src/HOL/Tools/inductive.ML Wed Dec 06 19:34:59 2017 +0100 +++ b/src/HOL/Tools/inductive.ML Wed Dec 06 20:43:09 2017 +0100 @@ -125,7 +125,7 @@ (** misc utilities **) -val inductive_internals = Attrib.setup_config_bool @{binding inductive_internals} (K false); +val inductive_internals = Attrib.setup_config_bool \<^binding>\inductive_internals\ (K false); fun message quiet_mode s = if quiet_mode then () else writeln s; @@ -142,7 +142,7 @@ (if i mod 2 = 0 then f x else g x) :: make_bool_args f g xs (i div 2); fun make_bool_args' xs = - make_bool_args (K @{term False}) (K @{term True}) xs; + make_bool_args (K \<^term>\False\) (K \<^term>\True\) xs; fun arg_types_of k c = drop k (binder_types (fastype_of c)); @@ -154,7 +154,7 @@ else apsnd (cons p) (find_arg T x ps); fun make_args Ts xs = - map (fn (T, (NONE, ())) => Const (@{const_name undefined}, T) | (_, (SOME t, ())) => t) + map (fn (T, (NONE, ())) => Const (\<^const_name>\undefined\, T) | (_, (SOME t, ())) => t) (fold (fn (t, T) => snd o find_arg T t) xs (map (rpair (NONE, ())) Ts)); fun make_args' Ts xs Us = @@ -280,9 +280,9 @@ handle THM _ => thm RS @{thm le_boolD} in (case Thm.concl_of thm of - Const (@{const_name Pure.eq}, _) $ _ $ _ => eq_to_mono (thm RS meta_eq_to_obj_eq) - | _ $ (Const (@{const_name HOL.eq}, _) $ _ $ _) => eq_to_mono thm - | _ $ (Const (@{const_name Orderings.less_eq}, _) $ _ $ _) => + Const (\<^const_name>\Pure.eq\, _) $ _ $ _ => eq_to_mono (thm RS meta_eq_to_obj_eq) + | _ $ (Const (\<^const_name>\HOL.eq\, _) $ _ $ _) => eq_to_mono thm + | _ $ (Const (\<^const_name>\Orderings.less_eq\, _) $ _ $ _) => dest_less_concl (Seq.hd (REPEAT (FIRSTGOAL (resolve_tac ctxt [@{thm le_funI}, @{thm le_boolI'}])) thm)) | _ => thm) @@ -300,7 +300,7 @@ val _ = Theory.setup - (Attrib.setup @{binding mono} (Attrib.add_del mono_add mono_del) + (Attrib.setup \<^binding>\mono\ (Attrib.add_del mono_add mono_del) "declaration of monotonicity rule"); @@ -370,7 +370,7 @@ val _ = (case concl of - Const (@{const_name Trueprop}, _) $ t => + Const (\<^const_name>\Trueprop\, _) $ t => if member (op =) cs (head_of t) then (check_ind (err_in_rule ctxt binding rule') t; List.app check_prem (prems ~~ aprems)) @@ -400,7 +400,7 @@ (if skip_mono then Goal.prove_sorry else Goal.prove_future) ctxt [] [] (HOLogic.mk_Trueprop - (Const (@{const_name Orderings.mono}, (predT --> predT) --> HOLogic.boolT) $ fp_fun)) + (Const (\<^const_name>\Orderings.mono\, (predT --> predT) --> HOLogic.boolT) $ fp_fun)) (fn _ => EVERY [resolve_tac ctxt @{thms monoI} 1, REPEAT (resolve_tac ctxt [@{thm le_funI}, @{thm le_boolI'}] 1), REPEAT (FIRST @@ -510,11 +510,11 @@ HOLogic.exists_const T $ Abs (a, T, list_ex (vars, t)); val conjs = map2 (curry HOLogic.mk_eq) frees us @ map HOLogic.dest_Trueprop ts; in - list_ex (params', if null conjs then @{term True} else foldr1 HOLogic.mk_conj conjs) + list_ex (params', if null conjs then \<^term>\True\ else foldr1 HOLogic.mk_conj conjs) end; val lhs = list_comb (c, params @ frees); val rhs = - if null c_intrs then @{term False} + if null c_intrs then \<^term>\False\ else foldr1 HOLogic.mk_disj (map mk_intr_conj c_intrs); val eq = HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs)); fun prove_intr1 (i, _) = Subgoal.FOCUS_PREMS (fn {context = ctxt'', params, prems, ...} => @@ -567,7 +567,7 @@ local (*delete needless equality assumptions*) -val refl_thin = Goal.prove_global @{theory HOL} [] [] @{prop "!!P. a = a ==> P ==> P"} +val refl_thin = Goal.prove_global @{theory HOL} [] [] \<^prop>\\P. a = a \ P \ P\ (fn {context = ctxt, ...} => assume_tac ctxt 1); val elim_rls = [asm_rl, FalseE, refl_thin, conjE, exE]; fun elim_tac ctxt = REPEAT o eresolve_tac ctxt elim_rls; @@ -630,7 +630,7 @@ val _ = Theory.setup - (Method.setup @{binding ind_cases} + (Method.setup \<^binding>\ind_cases\ (Scan.lift (Scan.repeat1 Parse.prop -- Parse.for_fixes) >> (fn (props, fixes) => fn ctxt => Method.erule ctxt 0 (ind_cases_rules ctxt props fixes))) @@ -705,7 +705,7 @@ val P = list_comb (nth preds i, map (incr_boundvars k) ys @ bs); val Q = fold_rev Term.abs (mk_names "x" k ~~ Ts) - (HOLogic.mk_binop @{const_name HOL.induct_conj} + (HOLogic.mk_binop \<^const_name>\HOL.induct_conj\ (list_comb (incr_boundvars k s, bs), P)); in (Q, case Ts of [] => SOME (s, P) | _ => NONE) end | NONE => @@ -753,7 +753,7 @@ val ind_concl = HOLogic.mk_Trueprop - (HOLogic.mk_binrel @{const_name Orderings.less_eq} (rec_const, ind_pred)); + (HOLogic.mk_binrel \<^const_name>\Orderings.less_eq\ (rec_const, ind_pred)); val raw_fp_induct = mono RS (fp_def RS @{thm def_lfp_induct}); @@ -791,7 +791,7 @@ (* prove coinduction rule *) -fun If_const T = Const (@{const_name If}, HOLogic.boolT --> T --> T --> T); +fun If_const T = Const (\<^const_name>\If\, HOLogic.boolT --> T --> T --> T); fun mk_If p t f = let val T = fastype_of t in If_const T $ p $ t $ f end; fun prove_coindrule quiet_mode preds cs argTs bs xs params intr_ts mono @@ -803,20 +803,20 @@ make_args' argTs xs (arg_types_of (length params) pred) |> `length) preds; val xTss = map (map fastype_of) xss; val (Rs_names, names_ctxt) = Variable.variant_fixes (mk_names "X" n) ctxt; - val Rs = map2 (fn name => fn Ts => Free (name, Ts ---> @{typ bool})) Rs_names xTss; + val Rs = map2 (fn name => fn Ts => Free (name, Ts ---> \<^typ>\bool\)) Rs_names xTss; val Rs_applied = map2 (curry list_comb) Rs xss; val preds_applied = map2 (curry list_comb) (map (fn p => list_comb (p, params)) preds) xss; val abstract_list = fold_rev (absfree o dest_Free); val bss = map (make_bool_args - (fn b => HOLogic.mk_eq (b, @{term False})) - (fn b => HOLogic.mk_eq (b, @{term True})) bs) (0 upto n - 1); + (fn b => HOLogic.mk_eq (b, \<^term>\False\)) + (fn b => HOLogic.mk_eq (b, \<^term>\True\)) bs) (0 upto n - 1); val eq_undefinedss = map (fn ys => map (fn x => - HOLogic.mk_eq (x, Const (@{const_name undefined}, fastype_of x))) + HOLogic.mk_eq (x, Const (\<^const_name>\undefined\, fastype_of x))) (subtract (op =) ys xs)) xss; val R = @{fold 3} (fn bs => fn eqs => fn R => fn t => if null bs andalso null eqs then R else mk_If (Library.foldr1 HOLogic.mk_conj (bs @ eqs)) R t) - bss eq_undefinedss Rs_applied @{term False} + bss eq_undefinedss Rs_applied \<^term>\False\ |> abstract_list (bs @ xs); fun subst t = @@ -847,7 +847,7 @@ in (i, fold_rev (fn (x, T) => fn P => HOLogic.exists_const T $ Abs (x, T, P)) (Logic.strip_params r) - (if null ps then @{term True} else foldr1 HOLogic.mk_conj ps)) + (if null ps then \<^term>\True\ else foldr1 HOLogic.mk_conj ps)) end; fun mk_prem i Ps = Logic.mk_implies @@ -908,7 +908,7 @@ fun mk_ind_def quiet_mode skip_mono alt_name coind cs intr_ts monos params cnames_syn lthy = let - val fp_name = if coind then @{const_name Inductive.gfp} else @{const_name Inductive.lfp}; + val fp_name = if coind then \<^const_name>\Inductive.gfp\ else \<^const_name>\Inductive.lfp\; val argTs = fold (combine (op =) o arg_types_of (length params)) cs []; val k = log 2 1 (length cs); @@ -954,14 +954,14 @@ in fold_rev (fn (x, T) => fn P => HOLogic.exists_const T $ Abs (x, T, P)) (Logic.strip_params r) - (if null ps then @{term True} else foldr1 HOLogic.mk_conj ps) + (if null ps then \<^term>\True\ else foldr1 HOLogic.mk_conj ps) end; (* make a disjunction of all introduction rules *) val fp_fun = fold_rev lambda (p :: bs @ xs) - (if null intr_ts then @{term False} + (if null intr_ts then \<^term>\False\ else foldr1 HOLogic.mk_disj (map transform_rule intr_ts)); (* add definition of recursive predicates to theory *) @@ -1296,32 +1296,32 @@ fun gen_ind_decl mk_def coind = Parse.vars -- Parse.for_fixes -- Scan.optional Parse_Spec.where_multi_specs [] -- - Scan.optional (@{keyword "monos"} |-- Parse.!!! Parse.thms1) [] + Scan.optional (\<^keyword>\monos\ |-- Parse.!!! Parse.thms1) [] >> (fn (((preds, params), specs), monos) => (snd o gen_add_inductive mk_def true coind preds params specs monos)); val ind_decl = gen_ind_decl add_ind_def; val _ = - Outer_Syntax.local_theory @{command_keyword inductive} "define inductive predicates" + Outer_Syntax.local_theory \<^command_keyword>\inductive\ "define inductive predicates" (ind_decl false); val _ = - Outer_Syntax.local_theory @{command_keyword coinductive} "define coinductive predicates" + Outer_Syntax.local_theory \<^command_keyword>\coinductive\ "define coinductive predicates" (ind_decl true); val _ = - Outer_Syntax.local_theory @{command_keyword inductive_cases} + Outer_Syntax.local_theory \<^command_keyword>\inductive_cases\ "create simplified instances of elimination rules" (Parse.and_list1 Parse_Spec.simple_specs >> (snd oo inductive_cases)); val _ = - Outer_Syntax.local_theory @{command_keyword inductive_simps} + Outer_Syntax.local_theory \<^command_keyword>\inductive_simps\ "create simplification rules for inductive predicates" (Parse.and_list1 Parse_Spec.simple_specs >> (snd oo inductive_simps)); val _ = - Outer_Syntax.command @{command_keyword print_inductives} + Outer_Syntax.command \<^command_keyword>\print_inductives\ "print (co)inductive definitions and monotonicity rules" (Parse.opt_bang >> (fn b => Toplevel.keep (print_inductives b o Toplevel.context_of))); diff -r d24dcac5eb4c -r e61557884799 src/HOL/Tools/inductive_realizer.ML --- a/src/HOL/Tools/inductive_realizer.ML Wed Dec 06 19:34:59 2017 +0100 +++ b/src/HOL/Tools/inductive_realizer.ML Wed Dec 06 20:43:09 2017 +0100 @@ -515,7 +515,7 @@ | SOME (SOME sets') => subtract (op =) sets' sets) end I); -val _ = Theory.setup (Attrib.setup @{binding ind_realizer} +val _ = Theory.setup (Attrib.setup \<^binding>\ind_realizer\ ((Scan.option (Scan.lift (Args.$$$ "irrelevant") |-- Scan.option (Scan.lift (Args.colon) |-- Scan.repeat1 (Args.const {proper = true, strict = true})))) >> rlz_attrib) diff -r d24dcac5eb4c -r e61557884799 src/HOL/Tools/inductive_set.ML --- a/src/HOL/Tools/inductive_set.ML Wed Dec 06 19:34:59 2017 +0100 +++ b/src/HOL/Tools/inductive_set.ML Wed Dec 06 20:43:09 2017 +0100 @@ -39,7 +39,7 @@ fun strong_ind_simproc tab = Simplifier.make_simproc @{context} "strong_ind" - {lhss = [@{term "x::'a::{}"}], + {lhss = [\<^term>\x::'a::{}\], proc = fn _ => fn ctxt => fn ct => let fun close p t f = @@ -47,20 +47,20 @@ in Thm.instantiate' [] (rev (map (SOME o Thm.cterm_of ctxt o Var) vs)) (p (fold (Logic.all o Var) vs t) f) end; - fun mkop @{const_name HOL.conj} T x = - SOME (Const (@{const_name Lattices.inf}, T --> T --> T), x) - | mkop @{const_name HOL.disj} T x = - SOME (Const (@{const_name Lattices.sup}, T --> T --> T), x) + fun mkop \<^const_name>\HOL.conj\ T x = + SOME (Const (\<^const_name>\Lattices.inf\, T --> T --> T), x) + | mkop \<^const_name>\HOL.disj\ T x = + SOME (Const (\<^const_name>\Lattices.sup\, T --> T --> T), x) | mkop _ _ _ = NONE; fun mk_collect p T t = let val U = HOLogic.dest_setT T in HOLogic.Collect_const U $ HOLogic.mk_ptupleabs (HOLogic.flat_tuple_paths p) U HOLogic.boolT t end; - fun decomp (Const (s, _) $ ((m as Const (@{const_name Set.member}, + fun decomp (Const (s, _) $ ((m as Const (\<^const_name>\Set.member\, Type (_, [_, Type (_, [T, _])]))) $ p $ S) $ u) = mkop s T (m, p, S, mk_collect p T (head_of u)) - | decomp (Const (s, _) $ u $ ((m as Const (@{const_name Set.member}, + | decomp (Const (s, _) $ u $ ((m as Const (\<^const_name>\Set.member\, Type (_, [_, Type (_, [T, _])]))) $ p $ S)) = mkop s T (m, p, mk_collect p T (head_of u), S) | decomp _ = NONE; @@ -241,7 +241,7 @@ in Simplifier.simplify (put_simpset HOL_basic_ss ctxt addsimps @{thms mem_Collect_eq case_prod_conv} - addsimprocs [@{simproc Collect_mem}]) thm'' + addsimprocs [\<^simproc>\Collect_mem\]) thm'' |> zero_var_indexes |> eta_contract_thm ctxt (equal p) end; @@ -252,14 +252,14 @@ fun add context thm (tab as {to_set_simps, to_pred_simps, set_arities, pred_arities}) = (case Thm.prop_of thm of - Const (@{const_name Trueprop}, _) $ (Const (@{const_name HOL.eq}, Type (_, [T, _])) $ lhs $ rhs) => + Const (\<^const_name>\Trueprop\, _) $ (Const (\<^const_name>\HOL.eq\, Type (_, [T, _])) $ lhs $ rhs) => (case body_type T of - @{typ bool} => + \<^typ>\bool\ => let val thy = Context.theory_of context; val ctxt = Context.proof_of context; fun factors_of t fs = case strip_abs_body t of - Const (@{const_name Set.member}, _) $ u $ S => + Const (\<^const_name>\Set.member\, _) $ u $ S => if is_Free S orelse is_Var S then let val ps = HOLogic.flat_tuple_paths u in (SOME ps, (S, ps) :: fs) end @@ -269,7 +269,7 @@ val (pfs, fs) = fold_map factors_of ts []; val ((h', ts'), fs') = (case rhs of Abs _ => (case strip_abs_body rhs of - Const (@{const_name Set.member}, _) $ u $ S => + Const (\<^const_name>\Set.member\, _) $ u $ S => (strip_comb S, SOME (HOLogic.flat_tuple_paths u)) | _ => raise Malformed "member symbol on right-hand side expected") | _ => (strip_comb rhs, NONE)) @@ -384,7 +384,7 @@ thm |> Thm.instantiate ([], insts) |> Simplifier.full_simplify (put_simpset HOL_basic_ss ctxt addsimps to_set_simps - addsimprocs [strong_ind_simproc pred_arities, @{simproc Collect_mem}]) |> + addsimprocs [strong_ind_simproc pred_arities, \<^simproc>\Collect_mem\]) |> Rule_Cases.save thm end; @@ -401,7 +401,7 @@ val {set_arities, pred_arities, to_pred_simps, ...} = Data.get (Context.Proof lthy); fun infer (Abs (_, _, t)) = infer t - | infer (Const (@{const_name Set.member}, _) $ t $ u) = + | infer (Const (\<^const_name>\Set.member\, _) $ t $ u) = infer_arities thy set_arities (SOME (HOLogic.flat_tuple_paths t), u) | infer (t $ u) = infer t #> infer u | infer _ = I; @@ -534,13 +534,13 @@ val _ = Theory.setup - (Attrib.setup @{binding pred_set_conv} (Scan.succeed pred_set_conv_att) + (Attrib.setup \<^binding>\pred_set_conv\ (Scan.succeed pred_set_conv_att) "declare rules for converting between predicate and set notation" #> - Attrib.setup @{binding to_set} (Attrib.thms >> to_set_att) + Attrib.setup \<^binding>\to_set\ (Attrib.thms >> to_set_att) "convert rule to set notation" #> - Attrib.setup @{binding to_pred} (Attrib.thms >> to_pred_att) + Attrib.setup \<^binding>\to_pred\ (Attrib.thms >> to_pred_att) "convert rule to predicate notation" #> - Attrib.setup @{binding mono_set} (Attrib.add_del mono_add mono_del) + Attrib.setup \<^binding>\mono_set\ (Attrib.add_del mono_add mono_del) "declare of monotonicity rule for set operators"); @@ -549,11 +549,11 @@ val ind_set_decl = Inductive.gen_ind_decl add_ind_set_def; val _ = - Outer_Syntax.local_theory @{command_keyword inductive_set} "define inductive sets" + Outer_Syntax.local_theory \<^command_keyword>\inductive_set\ "define inductive sets" (ind_set_decl false); val _ = - Outer_Syntax.local_theory @{command_keyword coinductive_set} "define coinductive sets" + Outer_Syntax.local_theory \<^command_keyword>\coinductive_set\ "define coinductive sets" (ind_set_decl true); end; diff -r d24dcac5eb4c -r e61557884799 src/HOL/Tools/int_arith.ML --- a/src/HOL/Tools/int_arith.ML Wed Dec 06 19:34:59 2017 +0100 +++ b/src/HOL/Tools/int_arith.ML Wed Dec 06 20:43:09 2017 +0100 @@ -24,10 +24,10 @@ val zero_to_of_int_zero_simproc = Simplifier.make_simproc @{context} "zero_to_of_int_zero_simproc" - {lhss = [@{term "0::'a::ring"}], + {lhss = [\<^term>\0::'a::ring\], proc = fn _ => fn ctxt => fn ct => let val T = Thm.ctyp_of_cterm ct in - if Thm.typ_of T = @{typ int} then NONE + if Thm.typ_of T = \<^typ>\int\ then NONE else SOME (Thm.instantiate' [SOME T] [] zeroth) end}; @@ -35,20 +35,20 @@ val one_to_of_int_one_simproc = Simplifier.make_simproc @{context} "one_to_of_int_one_simproc" - {lhss = [@{term "1::'a::ring_1"}], + {lhss = [\<^term>\1::'a::ring_1\], proc = fn _ => fn ctxt => fn ct => let val T = Thm.ctyp_of_cterm ct in - if Thm.typ_of T = @{typ int} then NONE + if Thm.typ_of T = \<^typ>\int\ then NONE else SOME (Thm.instantiate' [SOME T] [] oneth) end}; -fun check (Const (@{const_name Groups.one}, @{typ int})) = false - | check (Const (@{const_name Groups.one}, _)) = true - | check (Const (s, _)) = member (op =) [@{const_name HOL.eq}, - @{const_name Groups.times}, @{const_name Groups.uminus}, - @{const_name Groups.minus}, @{const_name Groups.plus}, - @{const_name Groups.zero}, - @{const_name Orderings.less}, @{const_name Orderings.less_eq}] s +fun check (Const (\<^const_name>\Groups.one\, \<^typ>\int\)) = false + | check (Const (\<^const_name>\Groups.one\, _)) = true + | check (Const (s, _)) = member (op =) [\<^const_name>\HOL.eq\, + \<^const_name>\Groups.times\, \<^const_name>\Groups.uminus\, + \<^const_name>\Groups.minus\, \<^const_name>\Groups.plus\, + \<^const_name>\Groups.zero\, + \<^const_name>\Orderings.less\, \<^const_name>\Orderings.less_eq\] s | check (a $ b) = check a andalso check b | check _ = false; @@ -63,9 +63,9 @@ val zero_one_idom_simproc = Simplifier.make_simproc @{context} "zero_one_idom_simproc" {lhss = - [@{term "(x::'a::ring_char_0) = y"}, - @{term "(x::'a::linordered_idom) < y"}, - @{term "(x::'a::linordered_idom) \ y"}], + [\<^term>\(x::'a::ring_char_0) = y\, + \<^term>\(x::'a::linordered_idom) < y\, + \<^term>\(x::'a::linordered_idom) \ y\], proc = fn _ => fn ctxt => fn ct => if check (Thm.term_of ct) then SOME (Simplifier.rewrite (put_simpset conv_ss ctxt) ct) @@ -73,7 +73,7 @@ fun number_of ctxt T n = - if not (Sign.of_sort (Proof_Context.theory_of ctxt) (T, @{sort numeral})) + if not (Sign.of_sort (Proof_Context.theory_of ctxt) (T, \<^sort>\numeral\)) then raise CTERM ("number_of", []) else Numeral.mk_cnumber (Thm.ctyp_of ctxt T) n; @@ -85,7 +85,7 @@ [@{thm of_int_numeral}, @{thm nat_0}, @{thm nat_1}, @{thm diff_nat_numeral}, @{thm nat_numeral}] #> Lin_Arith.add_simprocs [zero_one_idom_simproc] #> Lin_Arith.set_number_of number_of - #> Lin_Arith.add_inj_const (@{const_name of_nat}, HOLogic.natT --> HOLogic.intT) - #> Lin_Arith.add_discrete_type @{type_name Int.int} + #> Lin_Arith.add_inj_const (\<^const_name>\of_nat\, HOLogic.natT --> HOLogic.intT) + #> Lin_Arith.add_discrete_type \<^type_name>\Int.int\ end; diff -r d24dcac5eb4c -r e61557884799 src/HOL/Tools/lin_arith.ML --- a/src/HOL/Tools/lin_arith.ML Wed Dec 06 19:34:59 2017 +0100 +++ b/src/HOL/Tools/lin_arith.ML Wed Dec 06 20:43:09 2017 +0100 @@ -45,17 +45,17 @@ val mk_Trueprop = HOLogic.mk_Trueprop; fun atomize thm = case Thm.prop_of thm of - Const (@{const_name Trueprop}, _) $ (Const (@{const_name HOL.conj}, _) $ _ $ _) => + Const (\<^const_name>\Trueprop\, _) $ (Const (\<^const_name>\HOL.conj\, _) $ _ $ _) => atomize (thm RS conjunct1) @ atomize (thm RS conjunct2) | _ => [thm]; -fun neg_prop ((TP as Const(@{const_name Trueprop}, _)) $ (Const (@{const_name Not}, _) $ t)) = TP $ t - | neg_prop ((TP as Const(@{const_name Trueprop}, _)) $ t) = TP $ (HOLogic.Not $t) +fun neg_prop ((TP as Const(\<^const_name>\Trueprop\, _)) $ (Const (\<^const_name>\Not\, _) $ t)) = TP $ t + | neg_prop ((TP as Const(\<^const_name>\Trueprop\, _)) $ t) = TP $ (HOLogic.Not $t) | neg_prop t = raise TERM ("neg_prop", [t]); fun is_False thm = let val _ $ t = Thm.prop_of thm - in t = @{term False} end; + in t = \<^term>\False\ end; fun is_nat t = (fastype_of1 t = HOLogic.natT); @@ -97,9 +97,9 @@ {splits = splits, inj_consts = update (op =) c inj_consts, discrete = discrete}); -val split_limit = Attrib.setup_config_int @{binding linarith_split_limit} (K 9); -val neq_limit = Attrib.setup_config_int @{binding linarith_neq_limit} (K 9); -val trace = Attrib.setup_config_bool @{binding linarith_trace} (K false); +val split_limit = Attrib.setup_config_int \<^binding>\linarith_split_limit\ (K 9); +val neq_limit = Attrib.setup_config_int \<^binding>\linarith_neq_limit\ (K 9); +val trace = Attrib.setup_config_bool \<^binding>\linarith_trace\ (K false); structure LA_Data: LIN_ARITH_DATA = @@ -134,12 +134,12 @@ returns either (SOME term, associated multiplicity) or (NONE, constant) *) -fun of_field_sort thy U = Sign.of_sort thy (U, @{sort inverse}); +fun of_field_sort thy U = Sign.of_sort thy (U, \<^sort>\inverse\); fun demult thy (inj_consts : (string * typ) list) : term * Rat.rat -> term option * Rat.rat = let - fun demult ((mC as Const (@{const_name Groups.times}, _)) $ s $ t, m) = - (case s of Const (@{const_name Groups.times}, _) $ s1 $ s2 => + fun demult ((mC as Const (\<^const_name>\Groups.times\, _)) $ s $ t, m) = + (case s of Const (\<^const_name>\Groups.times\, _) $ s1 $ s2 => (* bracketing to the right: '(s1 * s2) * t' becomes 's1 * (s2 * t)' *) demult (mC $ s1 $ (mC $ s2 $ t), m) | _ => @@ -150,7 +150,7 @@ (SOME t', m'') => (SOME (mC $ s' $ t'), m'') | (NONE, m'') => (SOME s', m'')) | (NONE, m') => demult (t, m'))) - | demult (atom as (mC as Const (@{const_name Rings.divide}, T)) $ s $ t, m) = + | demult (atom as (mC as Const (\<^const_name>\Rings.divide\, T)) $ s $ t, m) = (* FIXME: Shouldn't we simplify nested quotients, e.g. '(s/t)/u' could become 's/(t*u)', and '(s*t)/u' could become 's*(t/u)' ? Note that if we choose to do so here, the simpset used by arith must be able to @@ -165,15 +165,15 @@ (SOME s', SOME t') => SOME (mC $ s' $ t') | (SOME s', NONE) => SOME s' | (NONE, SOME t') => - SOME (mC $ Const (@{const_name Groups.one}, domain_type (snd (dest_Const mC))) $ t') + SOME (mC $ Const (\<^const_name>\Groups.one\, domain_type (snd (dest_Const mC))) $ t') | (NONE, NONE) => NONE, Rat.mult m' (Rat.inv p)) end else (SOME atom, m) (* terms that evaluate to numeric constants *) - | demult (Const (@{const_name Groups.uminus}, _) $ t, m) = demult (t, ~ m) - | demult (Const (@{const_name Groups.zero}, _), _) = (NONE, @0) - | demult (Const (@{const_name Groups.one}, _), m) = (NONE, m) + | demult (Const (\<^const_name>\Groups.uminus\, _) $ t, m) = demult (t, ~ m) + | demult (Const (\<^const_name>\Groups.zero\, _), _) = (NONE, @0) + | demult (Const (\<^const_name>\Groups.one\, _), m) = (NONE, m) (*Warning: in rare cases (neg_)numeral encloses a non-numeral, in which case dest_numeral raises TERM; hence all the handles below. Same for Suc-terms that turn out not to be numerals - @@ -181,7 +181,7 @@ | demult (t as Const ("Num.numeral_class.numeral", _) (*DYNAMIC BINDING!*) $ n, m) = ((NONE, Rat.mult m (Rat.of_int (HOLogic.dest_numeral n))) handle TERM _ => (SOME t, m)) - | demult (t as Const (@{const_name Suc}, _) $ _, m) = + | demult (t as Const (\<^const_name>\Suc\, _) $ _, m) = ((NONE, Rat.mult m (Rat.of_int (HOLogic.dest_nat t))) handle TERM _ => (SOME t, m)) (* injection constants are ignored *) @@ -197,27 +197,27 @@ (* Turns a term 'all' and associated multiplicity 'm' into a list 'p' of summands and associated multiplicities, plus a constant 'i' (with implicit multiplicity 1) *) - fun poly (Const (@{const_name Groups.plus}, _) $ s $ t, + fun poly (Const (\<^const_name>\Groups.plus\, _) $ s $ t, m : Rat.rat, pi : (term * Rat.rat) list * Rat.rat) = poly (s, m, poly (t, m, pi)) - | poly (all as Const (@{const_name Groups.minus}, T) $ s $ t, m, pi) = + | poly (all as Const (\<^const_name>\Groups.minus\, T) $ s $ t, m, pi) = if nT T then add_atom all m pi else poly (s, m, poly (t, ~ m, pi)) - | poly (all as Const (@{const_name Groups.uminus}, T) $ t, m, pi) = + | poly (all as Const (\<^const_name>\Groups.uminus\, T) $ t, m, pi) = if nT T then add_atom all m pi else poly (t, ~ m, pi) - | poly (Const (@{const_name Groups.zero}, _), _, pi) = + | poly (Const (\<^const_name>\Groups.zero\, _), _, pi) = pi - | poly (Const (@{const_name Groups.one}, _), m, (p, i)) = + | poly (Const (\<^const_name>\Groups.one\, _), m, (p, i)) = (p, Rat.add i m) | poly (all as Const ("Num.numeral_class.numeral", _) (*DYNAMIC BINDING!*) $ t, m, pi as (p, i)) = (let val k = HOLogic.dest_numeral t in (p, Rat.add i (Rat.mult m (Rat.of_int k))) end handle TERM _ => add_atom all m pi) - | poly (Const (@{const_name Suc}, _) $ t, m, (p, i)) = + | poly (Const (\<^const_name>\Suc\, _) $ t, m, (p, i)) = poly (t, m, (p, Rat.add i m)) - | poly (all as Const (@{const_name Groups.times}, _) $ _ $ _, m, pi as (p, i)) = + | poly (all as Const (\<^const_name>\Groups.times\, _) $ _ $ _, m, pi as (p, i)) = (case demult thy inj_consts (all, m) of (NONE, m') => (p, Rat.add i m') | (SOME u, m') => add_atom u m' pi) - | poly (all as Const (@{const_name Rings.divide}, T) $ _ $ _, m, pi as (p, i)) = + | poly (all as Const (\<^const_name>\Rings.divide\, T) $ _ $ _, m, pi as (p, i)) = if of_field_sort thy (domain_type T) then (case demult thy inj_consts (all, m) of (NONE, m') => (p, Rat.add i m') @@ -231,14 +231,14 @@ val (q, j) = poly (rhs, @1, ([], @0)) in case rel of - @{const_name Orderings.less} => SOME (p, i, "<", q, j) - | @{const_name Orderings.less_eq} => SOME (p, i, "<=", q, j) - | @{const_name HOL.eq} => SOME (p, i, "=", q, j) + \<^const_name>\Orderings.less\ => SOME (p, i, "<", q, j) + | \<^const_name>\Orderings.less_eq\ => SOME (p, i, "<=", q, j) + | \<^const_name>\HOL.eq\ => SOME (p, i, "=", q, j) | _ => NONE end handle General.Div => NONE; fun of_lin_arith_sort thy U = - Sign.of_sort thy (U, @{sort Rings.linordered_idom}); + Sign.of_sort thy (U, \<^sort>\Rings.linordered_idom\); fun allows_lin_arith thy (discrete : string list) (U as Type (D, [])) : bool * bool = if of_lin_arith_sort thy U then (true, member (op =) discrete D) @@ -261,10 +261,10 @@ | negate NONE = NONE; fun decomp_negation thy data - ((Const (@{const_name Trueprop}, _)) $ (Const (rel, T) $ lhs $ rhs)) : decomp option = + ((Const (\<^const_name>\Trueprop\, _)) $ (Const (rel, T) $ lhs $ rhs)) : decomp option = decomp_typecheck thy data (T, (rel, lhs, rhs)) | decomp_negation thy data - ((Const (@{const_name Trueprop}, _)) $ (Const (@{const_name Not}, _) $ (Const (rel, T) $ lhs $ rhs))) = + ((Const (\<^const_name>\Trueprop\, _)) $ (Const (\<^const_name>\Not\, _) $ (Const (rel, T) $ lhs $ rhs))) = negate (decomp_typecheck thy data (T, (rel, lhs, rhs))) | decomp_negation _ _ _ = NONE; @@ -276,7 +276,7 @@ in decomp_negation thy (discrete, inj_consts) end; fun domain_is_nat (_ $ (Const (_, T) $ _ $ _)) = nT T - | domain_is_nat (_ $ (Const (@{const_name Not}, _) $ (Const (_, T) $ _ $ _))) = nT T + | domain_is_nat (_ $ (Const (\<^const_name>\Not\, _) $ (Const (_, T) $ _ $ _))) = nT T | domain_is_nat _ = false; @@ -313,24 +313,24 @@ let val {inj_consts, ...} = get_arith_data ctxt in member (op =) inj_consts f end -fun abstract_arith ((c as Const (@{const_name Groups.plus}, _)) $ u1 $ u2) cx = +fun abstract_arith ((c as Const (\<^const_name>\Groups.plus\, _)) $ u1 $ u2) cx = with2 abstract_arith c u1 u2 cx - | abstract_arith (t as (c as Const (@{const_name Groups.minus}, T)) $ u1 $ u2) cx = + | abstract_arith (t as (c as Const (\<^const_name>\Groups.minus\, T)) $ u1 $ u2) cx = if nT T then abstract_atom t cx else with2 abstract_arith c u1 u2 cx - | abstract_arith (t as (c as Const (@{const_name Groups.uminus}, T)) $ u) cx = + | abstract_arith (t as (c as Const (\<^const_name>\Groups.uminus\, T)) $ u) cx = if nT T then abstract_atom t cx else abstract_arith u cx |>> apply c - | abstract_arith ((c as Const (@{const_name Suc}, _)) $ u) cx = abstract_arith u cx |>> apply c - | abstract_arith ((c as Const (@{const_name Groups.times}, _)) $ u1 $ u2) cx = + | abstract_arith ((c as Const (\<^const_name>\Suc\, _)) $ u) cx = abstract_arith u cx |>> apply c + | abstract_arith ((c as Const (\<^const_name>\Groups.times\, _)) $ u1 $ u2) cx = with2 abstract_arith c u1 u2 cx - | abstract_arith (t as (c as Const (@{const_name Rings.divide}, T)) $ u1 $ u2) cx = + | abstract_arith (t as (c as Const (\<^const_name>\Rings.divide\, T)) $ u1 $ u2) cx = if is_field_sort cx T then with2 abstract_arith c u1 u2 cx else abstract_atom t cx | abstract_arith (t as (c as Const f) $ u) cx = if is_inj_const cx f then abstract_arith u cx |>> apply c else abstract_num t cx | abstract_arith t cx = abstract_num t cx -fun is_lin_arith_rel @{const_name Orderings.less} = true - | is_lin_arith_rel @{const_name Orderings.less_eq} = true - | is_lin_arith_rel @{const_name HOL.eq} = true +fun is_lin_arith_rel \<^const_name>\Orderings.less\ = true + | is_lin_arith_rel \<^const_name>\Orderings.less_eq\ = true + | is_lin_arith_rel \<^const_name>\HOL.eq\ = true | is_lin_arith_rel _ = false fun is_lin_arith_type (_, ctxt) T = @@ -342,10 +342,10 @@ else abstract_atom t cx | abstract_rel t cx = abstract_atom t cx -fun abstract_neg ((c as Const (@{const_name Not}, _)) $ t) cx = abstract_rel t cx |>> apply c +fun abstract_neg ((c as Const (\<^const_name>\Not\, _)) $ t) cx = abstract_rel t cx |>> apply c | abstract_neg t cx = abstract_rel t cx -fun abstract ((c as Const (@{const_name Trueprop}, _)) $ t) cx = abstract_neg t cx |>> apply c +fun abstract ((c as Const (\<^const_name>\Trueprop\, _)) $ t) cx = abstract_neg t cx |>> apply c | abstract t cx = abstract_atom t cx @@ -363,13 +363,13 @@ (case head_of lhs of Const (a, _) => member (op =) - [@{const_name Orderings.max}, - @{const_name Orderings.min}, - @{const_name Groups.abs}, - @{const_name Groups.minus}, + [\<^const_name>\Orderings.max\, + \<^const_name>\Orderings.min\, + \<^const_name>\Groups.abs\, + \<^const_name>\Groups.minus\, "Int.nat" (*DYNAMIC BINDING!*), - @{const_name Rings.modulo}, - @{const_name Rings.divide}] a + \<^const_name>\Rings.modulo\, + \<^const_name>\Rings.divide\] a | _ => (if Context_Position.is_visible ctxt then warning ("Lin. Arith.: wrong format for split rule " ^ Thm.string_of_thm ctxt thm) @@ -417,8 +417,7 @@ (* tn' --> ... --> t1' --> False , *) (* where ti' = HOLogic.dest_Trueprop ti *) fun REPEAT_DETERM_etac_rev_mp tms = - fold (curry HOLogic.mk_imp) (map HOLogic.dest_Trueprop tms) - @{term False} + fold (curry HOLogic.mk_imp) (map HOLogic.dest_Trueprop tms) \<^term>\False\ val split_thms = filter (is_split_thm ctxt) (#splits (get_arith_data ctxt)) val cmap = Splitter.cmap_of_split_thms split_thms val goal_tm = REPEAT_DETERM_etac_rev_mp terms @@ -441,72 +440,72 @@ (* ignore all but the first possible split *) (case strip_comb split_term of (* ?P (max ?i ?j) = ((?i <= ?j --> ?P ?j) & (~ ?i <= ?j --> ?P ?i)) *) - (Const (@{const_name Orderings.max}, _), [t1, t2]) => + (Const (\<^const_name>\Orderings.max\, _), [t1, t2]) => let val rev_terms = rev terms val terms1 = map (subst_term [(split_term, t1)]) rev_terms val terms2 = map (subst_term [(split_term, t2)]) rev_terms - val t1_leq_t2 = Const (@{const_name Orderings.less_eq}, + val t1_leq_t2 = Const (\<^const_name>\Orderings.less_eq\, split_type --> split_type --> HOLogic.boolT) $ t1 $ t2 val not_t1_leq_t2 = HOLogic.Not $ t1_leq_t2 - val not_false = HOLogic.mk_Trueprop (HOLogic.Not $ @{term False}) + val not_false = HOLogic.mk_Trueprop (HOLogic.Not $ \<^term>\False\) val subgoal1 = (HOLogic.mk_Trueprop t1_leq_t2) :: terms2 @ [not_false] val subgoal2 = (HOLogic.mk_Trueprop not_t1_leq_t2) :: terms1 @ [not_false] in SOME [(Ts, subgoal1), (Ts, subgoal2)] end (* ?P (min ?i ?j) = ((?i <= ?j --> ?P ?i) & (~ ?i <= ?j --> ?P ?j)) *) - | (Const (@{const_name Orderings.min}, _), [t1, t2]) => + | (Const (\<^const_name>\Orderings.min\, _), [t1, t2]) => let val rev_terms = rev terms val terms1 = map (subst_term [(split_term, t1)]) rev_terms val terms2 = map (subst_term [(split_term, t2)]) rev_terms - val t1_leq_t2 = Const (@{const_name Orderings.less_eq}, + val t1_leq_t2 = Const (\<^const_name>\Orderings.less_eq\, split_type --> split_type --> HOLogic.boolT) $ t1 $ t2 val not_t1_leq_t2 = HOLogic.Not $ t1_leq_t2 - val not_false = HOLogic.mk_Trueprop (HOLogic.Not $ @{term False}) + val not_false = HOLogic.mk_Trueprop (HOLogic.Not $ \<^term>\False\) val subgoal1 = (HOLogic.mk_Trueprop t1_leq_t2) :: terms1 @ [not_false] val subgoal2 = (HOLogic.mk_Trueprop not_t1_leq_t2) :: terms2 @ [not_false] in SOME [(Ts, subgoal1), (Ts, subgoal2)] end (* ?P (abs ?a) = ((0 <= ?a --> ?P ?a) & (?a < 0 --> ?P (- ?a))) *) - | (Const (@{const_name Groups.abs}, _), [t1]) => + | (Const (\<^const_name>\Groups.abs\, _), [t1]) => let val rev_terms = rev terms val terms1 = map (subst_term [(split_term, t1)]) rev_terms - val terms2 = map (subst_term [(split_term, Const (@{const_name Groups.uminus}, + val terms2 = map (subst_term [(split_term, Const (\<^const_name>\Groups.uminus\, split_type --> split_type) $ t1)]) rev_terms - val zero = Const (@{const_name Groups.zero}, split_type) - val zero_leq_t1 = Const (@{const_name Orderings.less_eq}, + val zero = Const (\<^const_name>\Groups.zero\, split_type) + val zero_leq_t1 = Const (\<^const_name>\Orderings.less_eq\, split_type --> split_type --> HOLogic.boolT) $ zero $ t1 - val t1_lt_zero = Const (@{const_name Orderings.less}, + val t1_lt_zero = Const (\<^const_name>\Orderings.less\, split_type --> split_type --> HOLogic.boolT) $ t1 $ zero - val not_false = HOLogic.mk_Trueprop (HOLogic.Not $ @{term False}) + val not_false = HOLogic.mk_Trueprop (HOLogic.Not $ \<^term>\False\) val subgoal1 = (HOLogic.mk_Trueprop zero_leq_t1) :: terms1 @ [not_false] val subgoal2 = (HOLogic.mk_Trueprop t1_lt_zero) :: terms2 @ [not_false] in SOME [(Ts, subgoal1), (Ts, subgoal2)] end (* ?P (?a - ?b) = ((?a < ?b --> ?P 0) & (ALL d. ?a = ?b + d --> ?P d)) *) - | (Const (@{const_name Groups.minus}, _), [t1, t2]) => + | (Const (\<^const_name>\Groups.minus\, _), [t1, t2]) => let (* "d" in the above theorem becomes a new bound variable after NNF *) (* transformation, therefore some adjustment of indices is necessary *) val rev_terms = rev terms - val zero = Const (@{const_name Groups.zero}, split_type) + val zero = Const (\<^const_name>\Groups.zero\, split_type) val d = Bound 0 val terms1 = map (subst_term [(split_term, zero)]) rev_terms val terms2 = map (subst_term [(incr_boundvars 1 split_term, d)]) (map (incr_boundvars 1) rev_terms) val t1' = incr_boundvars 1 t1 val t2' = incr_boundvars 1 t2 - val t1_lt_t2 = Const (@{const_name Orderings.less}, + val t1_lt_t2 = Const (\<^const_name>\Orderings.less\, split_type --> split_type --> HOLogic.boolT) $ t1 $ t2 - val t1_eq_t2_plus_d = Const (@{const_name HOL.eq}, split_type --> split_type --> HOLogic.boolT) $ t1' $ - (Const (@{const_name Groups.plus}, + val t1_eq_t2_plus_d = Const (\<^const_name>\HOL.eq\, split_type --> split_type --> HOLogic.boolT) $ t1' $ + (Const (\<^const_name>\Groups.plus\, split_type --> split_type --> split_type) $ t2' $ d) - val not_false = HOLogic.mk_Trueprop (HOLogic.Not $ @{term False}) + val not_false = HOLogic.mk_Trueprop (HOLogic.Not $ \<^term>\False\) val subgoal1 = (HOLogic.mk_Trueprop t1_lt_t2) :: terms1 @ [not_false] val subgoal2 = (HOLogic.mk_Trueprop t1_eq_t2_plus_d) :: terms2 @ [not_false] in @@ -516,18 +515,18 @@ | (Const ("Int.nat", _), (*DYNAMIC BINDING!*) [t1]) => let val rev_terms = rev terms - val zero_int = Const (@{const_name Groups.zero}, HOLogic.intT) - val zero_nat = Const (@{const_name Groups.zero}, HOLogic.natT) + val zero_int = Const (\<^const_name>\Groups.zero\, HOLogic.intT) + val zero_nat = Const (\<^const_name>\Groups.zero\, HOLogic.natT) val n = Bound 0 val terms1 = map (subst_term [(incr_boundvars 1 split_term, n)]) (map (incr_boundvars 1) rev_terms) val terms2 = map (subst_term [(split_term, zero_nat)]) rev_terms val t1' = incr_boundvars 1 t1 - val t1_eq_nat_n = Const (@{const_name HOL.eq}, HOLogic.intT --> HOLogic.intT --> HOLogic.boolT) $ t1' $ - (Const (@{const_name of_nat}, HOLogic.natT --> HOLogic.intT) $ n) - val t1_lt_zero = Const (@{const_name Orderings.less}, + val t1_eq_nat_n = Const (\<^const_name>\HOL.eq\, HOLogic.intT --> HOLogic.intT --> HOLogic.boolT) $ t1' $ + (Const (\<^const_name>\of_nat\, HOLogic.natT --> HOLogic.intT) $ n) + val t1_lt_zero = Const (\<^const_name>\Orderings.less\, HOLogic.intT --> HOLogic.intT --> HOLogic.boolT) $ t1 $ zero_int - val not_false = HOLogic.mk_Trueprop (HOLogic.Not $ @{term False}) + val not_false = HOLogic.mk_Trueprop (HOLogic.Not $ \<^term>\False\) val subgoal1 = (HOLogic.mk_Trueprop t1_eq_nat_n) :: terms1 @ [not_false] val subgoal2 = (HOLogic.mk_Trueprop t1_lt_zero) :: terms2 @ [not_false] in @@ -536,10 +535,10 @@ (* ?P ((?n::nat) mod (numeral ?k)) = ((numeral ?k = 0 --> ?P ?n) & (~ (numeral ?k = 0) --> (ALL i j. j < numeral ?k --> ?n = numeral ?k * i + j --> ?P j))) *) - | (Const (@{const_name Rings.modulo}, Type ("fun", [@{typ nat}, _])), [t1, t2]) => + | (Const (\<^const_name>\Rings.modulo\, Type ("fun", [\<^typ>\nat\, _])), [t1, t2]) => let val rev_terms = rev terms - val zero = Const (@{const_name Groups.zero}, split_type) + val zero = Const (\<^const_name>\Groups.zero\, split_type) val i = Bound 1 val j = Bound 0 val terms1 = map (subst_term [(split_term, t1)]) rev_terms @@ -547,17 +546,17 @@ (map (incr_boundvars 2) rev_terms) val t1' = incr_boundvars 2 t1 val t2' = incr_boundvars 2 t2 - val t2_eq_zero = Const (@{const_name HOL.eq}, + val t2_eq_zero = Const (\<^const_name>\HOL.eq\, split_type --> split_type --> HOLogic.boolT) $ t2 $ zero - val t2_neq_zero = HOLogic.mk_not (Const (@{const_name HOL.eq}, + val t2_neq_zero = HOLogic.mk_not (Const (\<^const_name>\HOL.eq\, split_type --> split_type --> HOLogic.boolT) $ t2' $ zero) - val j_lt_t2 = Const (@{const_name Orderings.less}, + val j_lt_t2 = Const (\<^const_name>\Orderings.less\, split_type --> split_type--> HOLogic.boolT) $ j $ t2' - val t1_eq_t2_times_i_plus_j = Const (@{const_name HOL.eq}, split_type --> split_type --> HOLogic.boolT) $ t1' $ - (Const (@{const_name Groups.plus}, split_type --> split_type --> split_type) $ - (Const (@{const_name Groups.times}, + val t1_eq_t2_times_i_plus_j = Const (\<^const_name>\HOL.eq\, split_type --> split_type --> HOLogic.boolT) $ t1' $ + (Const (\<^const_name>\Groups.plus\, split_type --> split_type --> split_type) $ + (Const (\<^const_name>\Groups.times\, split_type --> split_type --> split_type) $ t2' $ i) $ j) - val not_false = HOLogic.mk_Trueprop (HOLogic.Not $ @{term False}) + val not_false = HOLogic.mk_Trueprop (HOLogic.Not $ \<^term>\False\) val subgoal1 = (HOLogic.mk_Trueprop t2_eq_zero) :: terms1 @ [not_false] val subgoal2 = (map HOLogic.mk_Trueprop [t2_neq_zero, j_lt_t2, t1_eq_t2_times_i_plus_j]) @@ -568,10 +567,10 @@ (* ?P ((?n::nat) div (numeral ?k)) = ((numeral ?k = 0 --> ?P 0) & (~ (numeral ?k = 0) --> (ALL i j. j < numeral ?k --> ?n = numeral ?k * i + j --> ?P i))) *) - | (Const (@{const_name Rings.divide}, Type ("fun", [@{typ nat}, _])), [t1, t2]) => + | (Const (\<^const_name>\Rings.divide\, Type ("fun", [\<^typ>\nat\, _])), [t1, t2]) => let val rev_terms = rev terms - val zero = Const (@{const_name Groups.zero}, split_type) + val zero = Const (\<^const_name>\Groups.zero\, split_type) val i = Bound 1 val j = Bound 0 val terms1 = map (subst_term [(split_term, zero)]) rev_terms @@ -579,17 +578,17 @@ (map (incr_boundvars 2) rev_terms) val t1' = incr_boundvars 2 t1 val t2' = incr_boundvars 2 t2 - val t2_eq_zero = Const (@{const_name HOL.eq}, + val t2_eq_zero = Const (\<^const_name>\HOL.eq\, split_type --> split_type --> HOLogic.boolT) $ t2 $ zero - val t2_neq_zero = HOLogic.mk_not (Const (@{const_name HOL.eq}, + val t2_neq_zero = HOLogic.mk_not (Const (\<^const_name>\HOL.eq\, split_type --> split_type --> HOLogic.boolT) $ t2' $ zero) - val j_lt_t2 = Const (@{const_name Orderings.less}, + val j_lt_t2 = Const (\<^const_name>\Orderings.less\, split_type --> split_type--> HOLogic.boolT) $ j $ t2' - val t1_eq_t2_times_i_plus_j = Const (@{const_name HOL.eq}, split_type --> split_type --> HOLogic.boolT) $ t1' $ - (Const (@{const_name Groups.plus}, split_type --> split_type --> split_type) $ - (Const (@{const_name Groups.times}, + val t1_eq_t2_times_i_plus_j = Const (\<^const_name>\HOL.eq\, split_type --> split_type --> HOLogic.boolT) $ t1' $ + (Const (\<^const_name>\Groups.plus\, split_type --> split_type --> split_type) $ + (Const (\<^const_name>\Groups.times\, split_type --> split_type --> split_type) $ t2' $ i) $ j) - val not_false = HOLogic.mk_Trueprop (HOLogic.Not $ @{term False}) + val not_false = HOLogic.mk_Trueprop (HOLogic.Not $ \<^term>\False\) val subgoal1 = (HOLogic.mk_Trueprop t2_eq_zero) :: terms1 @ [not_false] val subgoal2 = (map HOLogic.mk_Trueprop [t2_neq_zero, j_lt_t2, t1_eq_t2_times_i_plus_j]) @@ -605,11 +604,11 @@ (numeral ?k < 0 --> (ALL i j. numeral ?k < j & j <= 0 & ?n = numeral ?k * i + j --> ?P j))) *) - | (Const (@{const_name Rings.modulo}, + | (Const (\<^const_name>\Rings.modulo\, Type ("fun", [Type ("Int.int", []), _])), (*DYNAMIC BINDING!*) [t1, t2]) => let val rev_terms = rev terms - val zero = Const (@{const_name Groups.zero}, split_type) + val zero = Const (\<^const_name>\Groups.zero\, split_type) val i = Bound 1 val j = Bound 0 val terms1 = map (subst_term [(split_term, t1)]) rev_terms @@ -617,25 +616,25 @@ (map (incr_boundvars 2) rev_terms) val t1' = incr_boundvars 2 t1 val t2' = incr_boundvars 2 t2 - val t2_eq_zero = Const (@{const_name HOL.eq}, + val t2_eq_zero = Const (\<^const_name>\HOL.eq\, split_type --> split_type --> HOLogic.boolT) $ t2 $ zero - val zero_lt_t2 = Const (@{const_name Orderings.less}, + val zero_lt_t2 = Const (\<^const_name>\Orderings.less\, split_type --> split_type --> HOLogic.boolT) $ zero $ t2' - val t2_lt_zero = Const (@{const_name Orderings.less}, + val t2_lt_zero = Const (\<^const_name>\Orderings.less\, split_type --> split_type --> HOLogic.boolT) $ t2' $ zero - val zero_leq_j = Const (@{const_name Orderings.less_eq}, + val zero_leq_j = Const (\<^const_name>\Orderings.less_eq\, split_type --> split_type --> HOLogic.boolT) $ zero $ j - val j_leq_zero = Const (@{const_name Orderings.less_eq}, + val j_leq_zero = Const (\<^const_name>\Orderings.less_eq\, split_type --> split_type --> HOLogic.boolT) $ j $ zero - val j_lt_t2 = Const (@{const_name Orderings.less}, + val j_lt_t2 = Const (\<^const_name>\Orderings.less\, split_type --> split_type--> HOLogic.boolT) $ j $ t2' - val t2_lt_j = Const (@{const_name Orderings.less}, + val t2_lt_j = Const (\<^const_name>\Orderings.less\, split_type --> split_type--> HOLogic.boolT) $ t2' $ j - val t1_eq_t2_times_i_plus_j = Const (@{const_name HOL.eq}, split_type --> split_type --> HOLogic.boolT) $ t1' $ - (Const (@{const_name Groups.plus}, split_type --> split_type --> split_type) $ - (Const (@{const_name Groups.times}, + val t1_eq_t2_times_i_plus_j = Const (\<^const_name>\HOL.eq\, split_type --> split_type --> HOLogic.boolT) $ t1' $ + (Const (\<^const_name>\Groups.plus\, split_type --> split_type --> split_type) $ + (Const (\<^const_name>\Groups.times\, split_type --> split_type --> split_type) $ t2' $ i) $ j) - val not_false = HOLogic.mk_Trueprop (HOLogic.Not $ @{term False}) + val not_false = HOLogic.mk_Trueprop (HOLogic.Not $ \<^term>\False\) val subgoal1 = (HOLogic.mk_Trueprop t2_eq_zero) :: terms1 @ [not_false] val subgoal2 = (map HOLogic.mk_Trueprop [zero_lt_t2, zero_leq_j]) @ hd terms2_3 @@ -659,11 +658,11 @@ (numeral ?k < 0 --> (ALL i j. numeral ?k < j & j <= 0 & ?n = numeral ?k * i + j --> ?P i))) *) - | (Const (@{const_name Rings.divide}, + | (Const (\<^const_name>\Rings.divide\, Type ("fun", [Type ("Int.int", []), _])), (*DYNAMIC BINDING!*) [t1, t2]) => let val rev_terms = rev terms - val zero = Const (@{const_name Groups.zero}, split_type) + val zero = Const (\<^const_name>\Groups.zero\, split_type) val i = Bound 1 val j = Bound 0 val terms1 = map (subst_term [(split_term, zero)]) rev_terms @@ -671,25 +670,25 @@ (map (incr_boundvars 2) rev_terms) val t1' = incr_boundvars 2 t1 val t2' = incr_boundvars 2 t2 - val t2_eq_zero = Const (@{const_name HOL.eq}, + val t2_eq_zero = Const (\<^const_name>\HOL.eq\, split_type --> split_type --> HOLogic.boolT) $ t2 $ zero - val zero_lt_t2 = Const (@{const_name Orderings.less}, + val zero_lt_t2 = Const (\<^const_name>\Orderings.less\, split_type --> split_type --> HOLogic.boolT) $ zero $ t2' - val t2_lt_zero = Const (@{const_name Orderings.less}, + val t2_lt_zero = Const (\<^const_name>\Orderings.less\, split_type --> split_type --> HOLogic.boolT) $ t2' $ zero - val zero_leq_j = Const (@{const_name Orderings.less_eq}, + val zero_leq_j = Const (\<^const_name>\Orderings.less_eq\, split_type --> split_type --> HOLogic.boolT) $ zero $ j - val j_leq_zero = Const (@{const_name Orderings.less_eq}, + val j_leq_zero = Const (\<^const_name>\Orderings.less_eq\, split_type --> split_type --> HOLogic.boolT) $ j $ zero - val j_lt_t2 = Const (@{const_name Orderings.less}, + val j_lt_t2 = Const (\<^const_name>\Orderings.less\, split_type --> split_type--> HOLogic.boolT) $ j $ t2' - val t2_lt_j = Const (@{const_name Orderings.less}, + val t2_lt_j = Const (\<^const_name>\Orderings.less\, split_type --> split_type--> HOLogic.boolT) $ t2' $ j - val t1_eq_t2_times_i_plus_j = Const (@{const_name HOL.eq}, split_type --> split_type --> HOLogic.boolT) $ t1' $ - (Const (@{const_name Groups.plus}, split_type --> split_type --> split_type) $ - (Const (@{const_name Groups.times}, + val t1_eq_t2_times_i_plus_j = Const (\<^const_name>\HOL.eq\, split_type --> split_type --> HOLogic.boolT) $ t1' $ + (Const (\<^const_name>\Groups.plus\, split_type --> split_type --> split_type) $ + (Const (\<^const_name>\Groups.times\, split_type --> split_type --> split_type) $ t2' $ i) $ j) - val not_false = HOLogic.mk_Trueprop (HOLogic.Not $ @{term False}) + val not_false = HOLogic.mk_Trueprop (HOLogic.Not $ \<^term>\False\) val subgoal1 = (HOLogic.mk_Trueprop t2_eq_zero) :: terms1 @ [not_false] val subgoal2 = (map HOLogic.mk_Trueprop [zero_lt_t2, zero_leq_j]) @ hd terms2_3 @@ -734,7 +733,7 @@ fun negated_term_occurs_positively (terms : term list) : bool = exists - (fn (Trueprop $ (Const (@{const_name Not}, _) $ t)) => + (fn (Trueprop $ (Const (\<^const_name>\Not\, _) $ t)) => member Envir.aeconv terms (Trueprop $ t) | _ => false) terms; @@ -959,9 +958,9 @@ val global_setup = map_theory_simpset (fn ctxt => ctxt addSolver (mk_solver "lin_arith" (add_arith_facts #> Fast_Arith.prems_lin_arith_tac))) #> - Attrib.setup @{binding arith_split} (Scan.succeed (Thm.declaration_attribute add_split)) + Attrib.setup \<^binding>\arith_split\ (Scan.succeed (Thm.declaration_attribute add_split)) "declaration of split rules for arithmetic procedure" #> - Method.setup @{binding linarith} + Method.setup \<^binding>\linarith\ (Scan.succeed (fn ctxt => METHOD (fn facts => HEADGOAL @@ -972,7 +971,7 @@ val setup = init_arith_data - #> add_discrete_type @{type_name nat} + #> add_discrete_type \<^type_name>\nat\ #> add_lessD @{thm Suc_leI} #> add_simps (@{thms simp_thms} @ @{thms ring_distribs} @ [@{thm if_True}, @{thm if_False}, @{thm minus_diff_eq}, @@ -982,11 +981,11 @@ #> add_simps [@{thm add_Suc}, @{thm add_Suc_right}, @{thm nat.inject}, @{thm Suc_le_mono}, @{thm Suc_less_eq}, @{thm Zero_not_Suc}, @{thm Suc_not_Zero}, @{thm le_0_eq}, @{thm One_nat_def}] - #> add_simprocs [@{simproc group_cancel_add}, @{simproc group_cancel_diff}, - @{simproc group_cancel_eq}, @{simproc group_cancel_le}, - @{simproc group_cancel_less}] + #> add_simprocs [\<^simproc>\group_cancel_add\, \<^simproc>\group_cancel_diff\, + \<^simproc>\group_cancel_eq\, \<^simproc>\group_cancel_le\, + \<^simproc>\group_cancel_less\] (*abel_cancel helps it work in abstract algebraic domains*) - #> add_simprocs [@{simproc nateq_cancel_sums},@{simproc natless_cancel_sums}, - @{simproc natle_cancel_sums}]; + #> add_simprocs [\<^simproc>\nateq_cancel_sums\,\<^simproc>\natless_cancel_sums\, + \<^simproc>\natle_cancel_sums\]; end; diff -r d24dcac5eb4c -r e61557884799 src/HOL/Tools/monomorph.ML --- a/src/HOL/Tools/monomorph.ML Wed Dec 06 19:34:59 2017 +0100 +++ b/src/HOL/Tools/monomorph.ML Wed Dec 06 20:43:09 2017 +0100 @@ -63,16 +63,16 @@ (* configuration options *) -val max_rounds = Attrib.setup_config_int @{binding monomorph_max_rounds} (K 5) +val max_rounds = Attrib.setup_config_int \<^binding>\monomorph_max_rounds\ (K 5) val max_new_instances = - Attrib.setup_config_int @{binding monomorph_max_new_instances} (K 500) + Attrib.setup_config_int \<^binding>\monomorph_max_new_instances\ (K 500) val max_thm_instances = - Attrib.setup_config_int @{binding monomorph_max_thm_instances} (K 20) + Attrib.setup_config_int \<^binding>\monomorph_max_thm_instances\ (K 20) val max_new_const_instances_per_round = - Attrib.setup_config_int @{binding monomorph_max_new_const_instances_per_round} (K 5) + Attrib.setup_config_int \<^binding>\monomorph_max_new_const_instances_per_round\ (K 5) fun limit_rounds ctxt f = let diff -r d24dcac5eb4c -r e61557884799 src/HOL/Tools/nat_arith.ML --- a/src/HOL/Tools/nat_arith.ML Wed Dec 06 19:34:59 2017 +0100 +++ b/src/HOL/Tools/nat_arith.ML Wed Dec 06 20:43:09 2017 +0100 @@ -30,11 +30,11 @@ [Conv.rewr_conv (Library.foldl (op RS) (rule0, path)), Conv.arg_conv (Raw_Simplifier.rewrite ctxt false norm_rules)] -fun add_atoms path (Const (@{const_name Groups.plus}, _) $ x $ y) = +fun add_atoms path (Const (\<^const_name>\Groups.plus\, _) $ x $ y) = add_atoms (add1::path) x #> add_atoms (add2::path) y - | add_atoms path (Const (@{const_name Nat.Suc}, _) $ x) = + | add_atoms path (Const (\<^const_name>\Nat.Suc\, _) $ x) = add_atoms (suc1::path) x - | add_atoms _ (Const (@{const_name Groups.zero}, _)) = I + | add_atoms _ (Const (\<^const_name>\Groups.zero\, _)) = I | add_atoms path x = cons (x, path) fun atoms t = add_atoms [] t [] diff -r d24dcac5eb4c -r e61557884799 src/HOL/Tools/record.ML --- a/src/HOL/Tools/record.ML Wed Dec 06 19:34:59 2017 +0100 +++ b/src/HOL/Tools/record.ML Wed Dec 06 20:43:09 2017 +0100 @@ -73,7 +73,7 @@ val iso_tuple_intro = @{thm isomorphic_tuple_intro}; val iso_tuple_intros = Tactic.build_net @{thms isomorphic_tuple.intros}; -val tuple_iso_tuple = (@{const_name Record.tuple_iso_tuple}, @{thm tuple_iso_tuple}); +val tuple_iso_tuple = (\<^const_name>\Record.tuple_iso_tuple\, @{thm tuple_iso_tuple}); structure Iso_Tuple_Thms = Theory_Data ( @@ -111,13 +111,13 @@ let val (leftT, rightT) = (fastype_of left, fastype_of right); val prodT = HOLogic.mk_prodT (leftT, rightT); - val isomT = Type (@{type_name tuple_isomorphism}, [prodT, leftT, rightT]); + val isomT = Type (\<^type_name>\tuple_isomorphism\, [prodT, leftT, rightT]); in - Const (@{const_name Record.iso_tuple_cons}, isomT --> leftT --> rightT --> prodT) $ + Const (\<^const_name>\Record.iso_tuple_cons\, isomT --> leftT --> rightT --> prodT) $ Const (fst tuple_iso_tuple, isomT) $ left $ right end; -fun dest_cons_tuple (Const (@{const_name Record.iso_tuple_cons}, _) $ Const _ $ t $ u) = (t, u) +fun dest_cons_tuple (Const (\<^const_name>\Record.iso_tuple_cons\, _) $ Const _ $ t $ u) = (t, u) | dest_cons_tuple t = raise TERM ("dest_cons_tuple", [t]); fun add_iso_tuple_type overloaded (b, alphas) (leftT, rightT) thy = @@ -149,7 +149,7 @@ [((Binding.concealed (Thm.def_binding isom_binding), Logic.mk_equals (isom, body)), [])]; val iso_tuple = isom_def RS (abs_inverse RS (rep_inject RS iso_tuple_intro)); - val cons = Const (@{const_name Record.iso_tuple_cons}, isomT --> leftT --> rightT --> absT); + val cons = Const (\<^const_name>\Record.iso_tuple_cons\, isomT --> leftT --> rightT --> absT); val thm_thy = cdef_thy @@ -171,8 +171,8 @@ val goal' = Envir.beta_eta_contract goal; val is = (case goal' of - Const (@{const_name Trueprop}, _) $ - (Const (@{const_name isomorphic_tuple}, _) $ Const is) => is + Const (\<^const_name>\Trueprop\, _) $ + (Const (\<^const_name>\isomorphic_tuple\, _) $ Const is) => is | _ => err "unexpected goal format" goal'); val isthm = (case Symtab.lookup isthms (#1 is) of @@ -202,7 +202,7 @@ val updacc_cong_triv = @{thm update_accessor_cong_assist_triv}; val updacc_cong_from_eq = @{thm iso_tuple_update_accessor_cong_from_eq}; -val codegen = Attrib.setup_config_bool @{binding record_codegen} (K true); +val codegen = Attrib.setup_config_bool \<^binding>\record_codegen\ (K true); (** name components **) @@ -229,7 +229,7 @@ (* timing *) -val timing = Attrib.setup_config_bool @{binding record_timing} (K false); +val timing = Attrib.setup_config_bool \<^binding>\record_timing\ (K false); fun timeit_msg ctxt s x = if Config.get ctxt timing then (warning s; timeit x) else x (); fun timing_msg ctxt s = if Config.get ctxt timing then warning s else (); @@ -628,11 +628,11 @@ | split_args (_ :: _) [] = raise Fail "expecting more fields" | split_args _ _ = ([], []); -fun field_type_tr ((Const (@{syntax_const "_field_type"}, _) $ Const (name, _) $ arg)) = +fun field_type_tr ((Const (\<^syntax_const>\_field_type\, _) $ Const (name, _) $ arg)) = (name, arg) | field_type_tr t = raise TERM ("field_type_tr", [t]); -fun field_types_tr (Const (@{syntax_const "_field_types"}, _) $ t $ u) = +fun field_types_tr (Const (\<^syntax_const>\_field_types\, _) $ t $ u) = field_type_tr t :: field_types_tr u | field_types_tr t = [field_type_tr t]; @@ -673,17 +673,17 @@ mk_ext (field_types_tr t) end; -fun record_type_tr ctxt [t] = record_field_types_tr (Syntax.const @{type_syntax unit}) ctxt t +fun record_type_tr ctxt [t] = record_field_types_tr (Syntax.const \<^type_syntax>\unit\) ctxt t | record_type_tr _ ts = raise TERM ("record_type_tr", ts); fun record_type_scheme_tr ctxt [t, more] = record_field_types_tr more ctxt t | record_type_scheme_tr _ ts = raise TERM ("record_type_scheme_tr", ts); -fun field_tr ((Const (@{syntax_const "_field"}, _) $ Const (name, _) $ arg)) = (name, arg) +fun field_tr ((Const (\<^syntax_const>\_field\, _) $ Const (name, _) $ arg)) = (name, arg) | field_tr t = raise TERM ("field_tr", [t]); -fun fields_tr (Const (@{syntax_const "_fields"}, _) $ t $ u) = field_tr t :: fields_tr u +fun fields_tr (Const (\<^syntax_const>\_fields\, _) $ t $ u) = field_tr t :: fields_tr u | fields_tr t = [field_tr t]; fun record_fields_tr more ctxt t = @@ -706,18 +706,18 @@ | mk_ext [] = more; in mk_ext (fields_tr t) end; -fun record_tr ctxt [t] = record_fields_tr (Syntax.const @{const_syntax Unity}) ctxt t +fun record_tr ctxt [t] = record_fields_tr (Syntax.const \<^const_syntax>\Unity\) ctxt t | record_tr _ ts = raise TERM ("record_tr", ts); fun record_scheme_tr ctxt [t, more] = record_fields_tr more ctxt t | record_scheme_tr _ ts = raise TERM ("record_scheme_tr", ts); -fun field_update_tr (Const (@{syntax_const "_field_update"}, _) $ Const (name, _) $ arg) = +fun field_update_tr (Const (\<^syntax_const>\_field_update\, _) $ Const (name, _) $ arg) = Syntax.const (suffix updateN name) $ Abs (Name.uu_, dummyT, arg) | field_update_tr t = raise TERM ("field_update_tr", [t]); -fun field_updates_tr (Const (@{syntax_const "_field_updates"}, _) $ t $ u) = +fun field_updates_tr (Const (\<^syntax_const>\_field_updates\, _) $ t $ u) = field_update_tr t :: field_updates_tr u | field_updates_tr t = [field_update_tr t]; @@ -729,28 +729,28 @@ val _ = Theory.setup (Sign.parse_translation - [(@{syntax_const "_record_update"}, K record_update_tr), - (@{syntax_const "_record"}, record_tr), - (@{syntax_const "_record_scheme"}, record_scheme_tr), - (@{syntax_const "_record_type"}, record_type_tr), - (@{syntax_const "_record_type_scheme"}, record_type_scheme_tr)]); + [(\<^syntax_const>\_record_update\, K record_update_tr), + (\<^syntax_const>\_record\, record_tr), + (\<^syntax_const>\_record_scheme\, record_scheme_tr), + (\<^syntax_const>\_record_type\, record_type_tr), + (\<^syntax_const>\_record_type_scheme\, record_type_scheme_tr)]); end; (* print translations *) -val type_abbr = Attrib.setup_config_bool @{binding record_type_abbr} (K true); -val type_as_fields = Attrib.setup_config_bool @{binding record_type_as_fields} (K true); +val type_abbr = Attrib.setup_config_bool \<^binding>\record_type_abbr\ (K true); +val type_as_fields = Attrib.setup_config_bool \<^binding>\record_type_as_fields\ (K true); local (* FIXME early extern (!??) *) (* FIXME Syntax.free (??) *) -fun field_type_tr' (c, t) = Syntax.const @{syntax_const "_field_type"} $ Syntax.const c $ t; - -fun field_types_tr' (t, u) = Syntax.const @{syntax_const "_field_types"} $ t $ u; +fun field_type_tr' (c, t) = Syntax.const \<^syntax_const>\_field_type\ $ Syntax.const c $ t; + +fun field_types_tr' (t, u) = Syntax.const \<^syntax_const>\_field_types\ $ t $ u; fun record_type_tr' ctxt t = let @@ -791,9 +791,9 @@ (map (field_type_tr' o apsnd (Syntax_Phases.term_of_typ ctxt)) fields); in if not (Config.get ctxt type_as_fields) orelse null fields then raise Match - else if moreT = HOLogic.unitT then Syntax.const @{syntax_const "_record_type"} $ u + else if moreT = HOLogic.unitT then Syntax.const \<^syntax_const>\_record_type\ $ u else - Syntax.const @{syntax_const "_record_type_scheme"} $ u $ + Syntax.const \<^syntax_const>\_record_type_scheme\ $ u $ Syntax_Phases.term_of_typ ctxt moreT end; @@ -847,8 +847,8 @@ local (* FIXME Syntax.free (??) *) -fun field_tr' (c, t) = Syntax.const @{syntax_const "_field"} $ Syntax.const c $ t; -fun fields_tr' (t, u) = Syntax.const @{syntax_const "_fields"} $ t $ u; +fun field_tr' (c, t) = Syntax.const \<^syntax_const>\_field\ $ Syntax.const c $ t; +fun fields_tr' (t, u) = Syntax.const \<^syntax_const>\_fields\ $ t $ u; fun record_tr' ctxt t = let @@ -876,8 +876,8 @@ val u = foldr1 fields_tr' (map field_tr' fields); in (case more of - Const (@{const_syntax Unity}, _) => Syntax.const @{syntax_const "_record"} $ u - | _ => Syntax.const @{syntax_const "_record_scheme"} $ u $ more) + Const (\<^const_syntax>\Unity\, _) => Syntax.const \<^syntax_const>\_record\ $ u + | _ => Syntax.const \<^syntax_const>\_record_scheme\ $ u $ more) end; in @@ -903,7 +903,7 @@ SOME name => (case try Syntax_Trans.const_abs_tr' k of SOME t => - apfst (cons (Syntax.const @{syntax_const "_field_update"} $ Syntax.free name $ t)) + apfst (cons (Syntax.const \<^syntax_const>\_field_update\ $ Syntax.free name $ t)) (field_updates_tr' ctxt u) | NONE => ([], tm)) | NONE => ([], tm)) @@ -913,8 +913,8 @@ (case field_updates_tr' ctxt tm of ([], _) => raise Match | (ts, u) => - Syntax.const @{syntax_const "_record_update"} $ u $ - foldr1 (fn (v, w) => Syntax.const @{syntax_const "_field_updates"} $ v $ w) (rev ts)); + Syntax.const \<^syntax_const>\_record_update\ $ u $ + foldr1 (fn (v, w) => Syntax.const \<^syntax_const>\_field_updates\ $ v $ w) (rev ts)); in @@ -938,7 +938,7 @@ fun mk_comp_id f = let val T = range_type (fastype_of f) - in HOLogic.mk_comp (Const (@{const_name Fun.id}, T --> T), f) end; + in HOLogic.mk_comp (Const (\<^const_name>\Fun.id\, T --> T), f) end; fun get_upd_funs (upd $ _ $ t) = upd :: get_upd_funs t | get_upd_funs _ = []; @@ -1063,7 +1063,7 @@ *) val simproc = Simplifier.make_simproc @{context} "record" - {lhss = [@{term "x::'a::{}"}], + {lhss = [\<^term>\x::'a::{}\], proc = fn _ => fn ctxt => fn ct => let val thy = Proof_Context.theory_of ctxt; @@ -1147,7 +1147,7 @@ both a more update and an update to a field within it.*) val upd_simproc = Simplifier.make_simproc @{context} "record_upd" - {lhss = [@{term "x::'a::{}"}], + {lhss = [\<^term>\x::'a::{}\], proc = fn _ => fn ctxt => fn ct => let val thy = Proof_Context.theory_of ctxt; @@ -1218,7 +1218,7 @@ val (isnoop, skelf') = is_upd_noop s f term; val funT = domain_type T; fun mk_comp_local (f, f') = - Const (@{const_name Fun.comp}, funT --> funT --> funT) $ f $ f'; + Const (\<^const_name>\Fun.comp\, funT --> funT --> funT) $ f $ f'; in if isnoop then (upd $ skelf' $ lhs, rhs, vars, @@ -1270,10 +1270,10 @@ *) val eq_simproc = Simplifier.make_simproc @{context} "record_eq" - {lhss = [@{term "r = s"}], + {lhss = [\<^term>\r = s\], proc = fn _ => fn ctxt => fn ct => (case Thm.term_of ct of - Const (@{const_name HOL.eq}, Type (_, [T, _])) $ _ $ _ => + Const (\<^const_name>\HOL.eq\, Type (_, [T, _])) $ _ $ _ => (case rec_id ~1 T of "" => NONE | name => @@ -1292,13 +1292,13 @@ P t > 0: split up to given bound of record extensions.*) fun split_simproc P = Simplifier.make_simproc @{context} "record_split" - {lhss = [@{term "x::'a::{}"}], + {lhss = [\<^term>\x::'a::{}\], proc = fn _ => fn ctxt => fn ct => (case Thm.term_of ct of Const (quantifier, Type (_, [Type (_, [T, _]), _])) $ _ => - if quantifier = @{const_name Pure.all} orelse - quantifier = @{const_name All} orelse - quantifier = @{const_name Ex} + if quantifier = \<^const_name>\Pure.all\ orelse + quantifier = \<^const_name>\All\ orelse + quantifier = \<^const_name>\Ex\ then (case rec_id ~1 T of "" => NONE @@ -1310,9 +1310,9 @@ | SOME (all_thm, All_thm, Ex_thm, _) => SOME (case quantifier of - @{const_name Pure.all} => all_thm - | @{const_name All} => All_thm RS @{thm eq_reflection} - | @{const_name Ex} => Ex_thm RS @{thm eq_reflection} + \<^const_name>\Pure.all\ => all_thm + | \<^const_name>\All\ => All_thm RS @{thm eq_reflection} + | \<^const_name>\Ex\ => Ex_thm RS @{thm eq_reflection} | _ => raise Fail "split_simproc")) else NONE end) @@ -1321,7 +1321,7 @@ val ex_sel_eq_simproc = Simplifier.make_simproc @{context} "ex_sel_eq" - {lhss = [@{term "Ex t"}], + {lhss = [\<^term>\Ex t\], proc = fn _ => fn ctxt => fn ct => let val thy = Proof_Context.theory_of ctxt; @@ -1335,23 +1335,23 @@ else raise TERM ("", [x]); val sel' = Const (sel, Tsel) $ Bound 0; val (l, r) = if lr then (sel', x') else (x', sel'); - in Const (@{const_name HOL.eq}, Teq) $ l $ r end + in Const (\<^const_name>\HOL.eq\, Teq) $ l $ r end else raise TERM ("", [Const (sel, Tsel)]); - fun dest_sel_eq (Const (@{const_name HOL.eq}, Teq) $ (Const (sel, Tsel) $ Bound 0) $ X) = + fun dest_sel_eq (Const (\<^const_name>\HOL.eq\, Teq) $ (Const (sel, Tsel) $ Bound 0) $ X) = (true, Teq, (sel, Tsel), X) - | dest_sel_eq (Const (@{const_name HOL.eq}, Teq) $ X $ (Const (sel, Tsel) $ Bound 0)) = + | dest_sel_eq (Const (\<^const_name>\HOL.eq\, Teq) $ X $ (Const (sel, Tsel) $ Bound 0)) = (false, Teq, (sel, Tsel), X) | dest_sel_eq _ = raise TERM ("", []); in (case t of - Const (@{const_name Ex}, Tex) $ Abs (s, T, t) => + Const (\<^const_name>\Ex\, Tex) $ Abs (s, T, t) => (let val eq = mkeq (dest_sel_eq t) 0; val prop = Logic.list_all ([("r", T)], Logic.mk_equals - (Const (@{const_name Ex}, Tex) $ Abs (s, T, eq), @{term True})); + (Const (\<^const_name>\Ex\, Tex) $ Abs (s, T, eq), \<^term>\True\)); in SOME (Goal.prove_sorry_global thy [] [] prop (fn _ => simp_tac (put_simpset (get_simpset thy) ctxt @@ -1379,7 +1379,7 @@ val has_rec = exists_Const (fn (s, Type (_, [Type (_, [T, _]), _])) => - (s = @{const_name Pure.all} orelse s = @{const_name All} orelse s = @{const_name Ex}) + (s = \<^const_name>\Pure.all\ orelse s = \<^const_name>\All\ orelse s = \<^const_name>\Ex\) andalso is_recT T | _ => false); @@ -1427,11 +1427,11 @@ val has_rec = exists_Const (fn (s, Type (_, [Type (_, [T, _]), _])) => - (s = @{const_name Pure.all} orelse s = @{const_name All}) andalso is_recT T + (s = \<^const_name>\Pure.all\ orelse s = \<^const_name>\All\) andalso is_recT T | _ => false); - fun is_all (Const (@{const_name Pure.all}, _) $ _) = ~1 - | is_all (Const (@{const_name All}, _) $ _) = ~1 + fun is_all (Const (\<^const_name>\Pure.all\, _) $ _) = ~1 + | is_all (Const (\<^const_name>\All\, _) $ _) = ~1 | is_all _ = 0; in if has_rec goal then @@ -1586,10 +1586,10 @@ val inject_prop = (* FIXME local x x' *) let val vars_more' = map (fn (Free (x, T)) => Free (x ^ "'", T)) vars_more in HOLogic.mk_conj (HOLogic.eq_const extT $ - mk_ext vars_more $ mk_ext vars_more', @{term True}) + mk_ext vars_more $ mk_ext vars_more', \<^term>\True\) === foldr1 HOLogic.mk_conj - (map HOLogic.mk_eq (vars_more ~~ vars_more') @ [@{term True}]) + (map HOLogic.mk_eq (vars_more ~~ vars_more') @ [\<^term>\True\]) end; val induct_prop = (fold_rev Logic.all vars_more (Trueprop (P $ ext)), Trueprop (P $ s)); @@ -1688,19 +1688,19 @@ fun mk_random_eq tyco vs extN Ts = let (* FIXME local i etc. *) - val size = @{term "i::natural"}; - fun termifyT T = HOLogic.mk_prodT (T, @{typ "unit => term"}); + val size = \<^term>\i::natural\; + fun termifyT T = HOLogic.mk_prodT (T, \<^typ>\unit \ term\); val T = Type (tyco, map TFree vs); val Tm = termifyT T; val params = Name.invent_names Name.context "x" Ts; val lhs = HOLogic.mk_random T size; - val tc = HOLogic.mk_return Tm @{typ Random.seed} + val tc = HOLogic.mk_return Tm \<^typ>\Random.seed\ (HOLogic.mk_valtermify_app extN params T); val rhs = HOLogic.mk_ST (map (fn (v, T') => - ((HOLogic.mk_random T' size, @{typ Random.seed}), SOME (v, termifyT T'))) params) - tc @{typ Random.seed} (SOME Tm, @{typ Random.seed}); + ((HOLogic.mk_random T' size, \<^typ>\Random.seed\), SOME (v, termifyT T'))) params) + tc \<^typ>\Random.seed\ (SOME Tm, \<^typ>\Random.seed\); in (lhs, rhs) end @@ -1708,16 +1708,16 @@ fun mk_full_exhaustive_eq tyco vs extN Ts = let (* FIXME local i etc. *) - val size = @{term "i::natural"}; - fun termifyT T = HOLogic.mk_prodT (T, @{typ "unit => term"}); + val size = \<^term>\i::natural\; + fun termifyT T = HOLogic.mk_prodT (T, \<^typ>\unit \ term\); val T = Type (tyco, map TFree vs); - val test_function = Free ("f", termifyT T --> @{typ "(bool * term list) option"}); + val test_function = Free ("f", termifyT T --> \<^typ>\(bool \ term list) option\); val params = Name.invent_names Name.context "x" Ts; fun full_exhaustiveT T = - (termifyT T --> @{typ "(bool * Code_Evaluation.term list) option"}) --> - @{typ natural} --> @{typ "(bool * Code_Evaluation.term list) option"}; + (termifyT T --> \<^typ>\(bool \ Code_Evaluation.term list) option\) --> + \<^typ>\natural\ --> \<^typ>\(bool \ Code_Evaluation.term list) option\; fun mk_full_exhaustive T = - Const (@{const_name "Quickcheck_Exhaustive.full_exhaustive_class.full_exhaustive"}, + Const (\<^const_name>\Quickcheck_Exhaustive.full_exhaustive_class.full_exhaustive\, full_exhaustiveT T); val lhs = mk_full_exhaustive T $ test_function $ size; val tc = test_function $ (HOLogic.mk_valtermify_app extN params T); @@ -1758,8 +1758,8 @@ let val eq = HOLogic.mk_Trueprop (HOLogic.mk_eq - (Const (@{const_name HOL.equal}, extT --> extT --> HOLogic.boolT), - Const (@{const_name HOL.eq}, extT --> extT --> HOLogic.boolT))); + (Const (\<^const_name>\HOL.equal\, extT --> extT --> HOLogic.boolT), + Const (\<^const_name>\HOL.eq\, extT --> extT --> HOLogic.boolT))); fun tac ctxt eq_def = Class.intro_classes_tac ctxt [] THEN rewrite_goals_tac ctxt [Simpdata.mk_eq eq_def] @@ -1770,11 +1770,11 @@ fun mk_eq_refl ctxt = @{thm equal_refl} |> Thm.instantiate - ([((("'a", 0), @{sort equal}), Thm.ctyp_of ctxt (Logic.varifyT_global extT))], []) + ([((("'a", 0), \<^sort>\equal\), Thm.ctyp_of ctxt (Logic.varifyT_global extT))], []) |> Axclass.unoverload ctxt; - val ensure_random_record = ensure_sort_record (@{sort random}, mk_random_eq); + val ensure_random_record = ensure_sort_record (\<^sort>\random\, mk_random_eq); val ensure_exhaustive_record = - ensure_sort_record (@{sort full_exhaustive}, mk_full_exhaustive_eq); + ensure_sort_record (\<^sort>\full_exhaustive\, mk_full_exhaustive_eq); fun add_code eq_def thy = let val ctxt = Proof_Context.init_global thy; @@ -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 (\<^const_name>\Pure.eq\, _) $ t $ _) = t + | lhs_of_equation (\<^const>\Trueprop\ $ (Const (\<^const_name>\HOL.eq\, _) $ t $ _)) = t; fun add_spec_rule rule = let val head = head_of (lhs_of_equation (Thm.prop_of rule)) in @@ -1858,7 +1858,7 @@ val all_vars = parent_vars @ vars; val all_named_vars = (parent_names ~~ parent_vars) @ named_vars; - val zeta = (singleton (Name.variant_list (map #1 alphas)) "'z", @{sort type}); + val zeta = (singleton (Name.variant_list (map #1 alphas)) "'z", \<^sort>\type\); val moreT = TFree zeta; val more = Free (moreN, moreT); val full_moreN = full (Binding.name moreN); @@ -2408,18 +2408,18 @@ (* outer syntax *) val _ = - Outer_Syntax.command @{command_keyword record} "define extensible record" + Outer_Syntax.command \<^command_keyword>\record\ "define extensible record" (Parse_Spec.overloaded -- (Parse.type_args_constrained -- Parse.binding) -- - (@{keyword "="} |-- Scan.option (Parse.typ --| @{keyword "+"}) -- + (\<^keyword>\=\ |-- Scan.option (Parse.typ --| \<^keyword>\+\) -- Scan.repeat1 Parse.const_binding) >> (fn ((overloaded, x), (y, z)) => Toplevel.theory (add_record_cmd {overloaded = overloaded} x y z))); val opt_modes = - Scan.optional (@{keyword "("} |-- Parse.!!! (Scan.repeat1 Parse.name --| @{keyword ")"})) [] + Scan.optional (\<^keyword>\(\ |-- Parse.!!! (Scan.repeat1 Parse.name --| \<^keyword>\)\)) [] val _ = - Outer_Syntax.command @{command_keyword print_record} "print record definiton" + Outer_Syntax.command \<^command_keyword>\print_record\ "print record definiton" (opt_modes -- Parse.typ >> print_record); end diff -r d24dcac5eb4c -r e61557884799 src/HOL/Tools/sat.ML --- a/src/HOL/Tools/sat.ML Wed Dec 06 19:34:59 2017 +0100 +++ b/src/HOL/Tools/sat.ML Wed Dec 06 20:43:09 2017 +0100 @@ -60,12 +60,12 @@ structure SAT : SAT = struct -val trace = Attrib.setup_config_bool @{binding sat_trace} (K false); +val trace = Attrib.setup_config_bool \<^binding>\sat_trace\ (K false); fun cond_tracing ctxt msg = if Config.get ctxt trace then tracing (msg ()) else (); -val solver = Attrib.setup_config_string @{binding sat_solver} (K "cdclite"); +val solver = Attrib.setup_config_string \<^binding>\sat_solver\ (K "cdclite"); (*see HOL/Tools/sat_solver.ML for possible values*) val counter = Unsynchronized.ref 0; @@ -212,7 +212,7 @@ let fun index_of_literal chyp = (case (HOLogic.dest_Trueprop o Thm.term_of) chyp of - (Const (@{const_name Not}, _) $ atom) => + (Const (\<^const_name>\Not\, _) $ atom) => SOME (~ (the (Termtab.lookup atom_table atom))) | atom => SOME (the (Termtab.lookup atom_table atom))) handle TERM _ => NONE; (* 'chyp' is not a literal *) @@ -281,7 +281,7 @@ let (* remove premises that equal "True" *) val clauses' = filter (fn clause => - (not_equal @{term True} o HOLogic.dest_Trueprop o Thm.term_of) clause + (not_equal \<^term>\True\ o HOLogic.dest_Trueprop o Thm.term_of) clause handle TERM ("dest_Trueprop", _) => true) clauses (* remove non-clausal premises -- of course this shouldn't actually *) (* remove anything as long as 'rawsat_tac' is only called after the *) @@ -321,7 +321,7 @@ val _ = cond_tracing ctxt (fn () => "quick_and_dirty is set: proof reconstruction skipped, using oracle instead") - val False_thm = Skip_Proof.make_thm_cterm @{cprop False} + val False_thm = Skip_Proof.make_thm_cterm \<^cprop>\False\ in (* 'fold Thm.weaken (rev sorted_clauses)' is linear, while 'fold *) (* Thm.weaken sorted_clauses' would be quadratic, since we sorted *) diff -r d24dcac5eb4c -r e61557884799 src/HOL/Tools/split_rule.ML --- a/src/HOL/Tools/split_rule.ML Wed Dec 06 19:34:59 2017 +0100 +++ b/src/HOL/Tools/split_rule.ML Wed Dec 06 20:43:09 2017 +0100 @@ -17,7 +17,7 @@ (** split rules **) fun internal_case_prod_const (Ta, Tb, Tc) = - Const (@{const_name Product_Type.internal_case_prod}, + Const (\<^const_name>\Product_Type.internal_case_prod\, [[Ta, Tb] ---> Tc, HOLogic.mk_prodT (Ta, Tb)] ---> Tc); fun eval_internal_split ctxt = @@ -30,7 +30,7 @@ (*In ap_split S T u, term u expects separate arguments for the factors of S, with result type T. The call creates a new term expecting one argument of type S.*) -fun ap_split (Type (@{type_name Product_Type.prod}, [T1, T2])) T3 u = +fun ap_split (Type (\<^type_name>\Product_Type.prod\, [T1, T2])) T3 u = internal_case_prod_const (T1, T2, T3) $ Abs ("v", T1, ap_split T2 T3 @@ -111,11 +111,11 @@ val _ = Theory.setup - (Attrib.setup @{binding split_format} + (Attrib.setup \<^binding>\split_format\ (Scan.lift (Args.parens (Args.$$$ "complete") >> K (Thm.rule_attribute [] (complete_split_rule o Context.proof_of)))) "split pair-typed subterms in premises, or function arguments" #> - Attrib.setup @{binding split_rule} + Attrib.setup \<^binding>\split_rule\ (Scan.succeed (Thm.rule_attribute [] (split_rule o Context.proof_of))) "curries ALL function variables occurring in a rule's conclusion"); diff -r d24dcac5eb4c -r e61557884799 src/HOL/Tools/try0.ML --- a/src/HOL/Tools/try0.ML Wed Dec 06 19:34:59 2017 +0100 +++ b/src/HOL/Tools/try0.ML Wed Dec 06 20:43:09 2017 +0100 @@ -182,11 +182,11 @@ || Scan.repeat parse_attr >> (fn quad => fold merge_attrs quad ([], [], [], []))) x; val _ = - Outer_Syntax.command @{command_keyword try0} "try a combination of proof methods" + Outer_Syntax.command \<^command_keyword>\try0\ "try a combination of proof methods" (Scan.optional parse_attrs ([], [], [], []) #>> try0_trans); fun try_try0 auto = generic_try0 (if auto then Auto_Try else Try) NONE ([], [], [], []); -val _ = Try.tool_setup (try0N, (30, @{system_option auto_methods}, try_try0)); +val _ = Try.tool_setup (try0N, (30, \<^system_option>\auto_methods\, try_try0)); end; diff -r d24dcac5eb4c -r e61557884799 src/HOL/Tools/typedef.ML --- a/src/HOL/Tools/typedef.ML Wed Dec 06 19:34:59 2017 +0100 +++ b/src/HOL/Tools/typedef.ML Wed Dec 06 20:43:09 2017 +0100 @@ -85,7 +85,7 @@ structure Typedef_Plugin = Plugin(type T = string); -val typedef_plugin = Plugin_Name.declare_setup @{binding typedef}; +val typedef_plugin = Plugin_Name.declare_setup \<^binding>\typedef\; fun interpretation f = Typedef_Plugin.interpretation typedef_plugin @@ -99,7 +99,7 @@ (* primitive typedef axiomatization -- for fresh typedecl *) -val typedef_overloaded = Attrib.setup_config_bool @{binding typedef_overloaded} (K false); +val typedef_overloaded = Attrib.setup_config_bool \<^binding>\typedef_overloaded\ (K false); fun mk_inhabited A = let val T = HOLogic.dest_setT (Term.fastype_of A) @@ -108,7 +108,7 @@ fun mk_typedef newT oldT RepC AbsC A = let val typedefC = - Const (@{const_name type_definition}, + Const (\<^const_name>\type_definition\, (newT --> oldT) --> (oldT --> newT) --> HOLogic.mk_setT oldT --> HOLogic.boolT); in Logic.mk_implies (mk_inhabited A, HOLogic.mk_Trueprop (typedefC $ RepC $ AbsC $ A)) end; @@ -336,11 +336,11 @@ (** outer syntax **) val _ = - Outer_Syntax.local_theory_to_proof @{command_keyword typedef} + Outer_Syntax.local_theory_to_proof \<^command_keyword>\typedef\ "HOL type definition (requires non-emptiness proof)" (Parse_Spec.overloaded -- Parse.type_args_constrained -- Parse.binding -- Parse.opt_mixfix -- - (@{keyword "="} |-- Parse.term) -- - Scan.option (@{keyword "morphisms"} |-- Parse.!!! (Parse.binding -- Parse.binding)) + (\<^keyword>\=\ |-- Parse.term) -- + Scan.option (\<^keyword>\morphisms\ |-- Parse.!!! (Parse.binding -- Parse.binding)) >> (fn (((((overloaded, vs), t), mx), A), opt_morphs) => fn lthy => typedef_cmd {overloaded = overloaded} (t, vs, mx) A (SOME (make_morphisms t opt_morphs)) lthy)); diff -r d24dcac5eb4c -r e61557884799 src/HOL/Tools/value_command.ML --- a/src/HOL/Tools/value_command.ML Wed Dec 06 19:34:59 2017 +0100 +++ b/src/HOL/Tools/value_command.ML Wed Dec 06 20:43:09 2017 +0100 @@ -62,18 +62,18 @@ in Pretty.writeln p end; val opt_modes = - Scan.optional (@{keyword "("} |-- Parse.!!! (Scan.repeat1 Parse.name --| @{keyword ")"})) []; + Scan.optional (\<^keyword>\(\ |-- Parse.!!! (Scan.repeat1 Parse.name --| \<^keyword>\)\)) []; val opt_evaluator = - Scan.option (@{keyword "["} |-- Parse.name --| @{keyword "]"}) + Scan.option (\<^keyword>\[\ |-- Parse.name --| \<^keyword>\]\) val _ = - Outer_Syntax.command @{command_keyword value} "evaluate and print term" + Outer_Syntax.command \<^command_keyword>\value\ "evaluate and print term" (opt_evaluator -- opt_modes -- Parse.term >> (fn ((some_name, modes), t) => Toplevel.keep (value_cmd some_name modes t))); val _ = Theory.setup - (Thy_Output.antiquotation @{binding value} + (Thy_Output.antiquotation \<^binding>\value\ (Scan.lift opt_evaluator -- Term_Style.parse -- Args.term) (fn {source, context, ...} => fn ((some_name, style), t) => Thy_Output.output context (Thy_Output.maybe_pretty_source Thy_Output.pretty_term context source diff -r d24dcac5eb4c -r e61557884799 src/HOL/Topological_Spaces.thy --- a/src/HOL/Topological_Spaces.thy Wed Dec 06 19:34:59 2017 +0100 +++ b/src/HOL/Topological_Spaces.thy Wed Dec 06 20:43:09 2017 +0100 @@ -688,7 +688,7 @@ named_theorems tendsto_intros "introduction rules for tendsto" setup \ - Global_Theory.add_thms_dynamic (@{binding tendsto_eq_intros}, + Global_Theory.add_thms_dynamic (\<^binding>\tendsto_eq_intros\, fn context => Named_Theorems.get (Context.proof_of context) @{named_theorems tendsto_intros} |> map_filter (try (fn thm => @{thm tendsto_eq_rhs} OF [thm]))) diff -r d24dcac5eb4c -r e61557884799 src/Tools/atomize_elim.ML --- a/src/Tools/atomize_elim.ML Wed Dec 06 19:34:59 2017 +0100 +++ b/src/Tools/atomize_elim.ML Wed Dec 06 20:43:09 2017 +0100 @@ -19,7 +19,7 @@ val named_theorems = Context.>>> (Context.map_theory_result (Named_Target.theory_init #> - Named_Theorems.declare @{binding atomize_elim} "atomize_elim rewrite rule" ##> + Named_Theorems.declare \<^binding>\atomize_elim\ "atomize_elim rewrite rule" ##> Local_Theory.exit_global)); @@ -129,7 +129,7 @@ val _ = Theory.setup - (Method.setup @{binding atomize_elim} (Scan.succeed (SIMPLE_METHOD' o atomize_elim_tac)) + (Method.setup \<^binding>\atomize_elim\ (Scan.succeed (SIMPLE_METHOD' o atomize_elim_tac)) "convert obtains statement to atomic object logic goal") end diff -r d24dcac5eb4c -r e61557884799 src/Tools/case_product.ML --- a/src/Tools/case_product.ML Wed Dec 06 19:34:59 2017 +0100 +++ b/src/Tools/case_product.ML Wed Dec 06 20:43:09 2017 +0100 @@ -108,7 +108,7 @@ val _ = Theory.setup - (Attrib.setup @{binding case_product} + (Attrib.setup \<^binding>\case_product\ let fun combine_list ctxt = fold (fn x => fn y => combine_annotated ctxt y x) in diff -r d24dcac5eb4c -r e61557884799 src/Tools/coherent.ML --- a/src/Tools/coherent.ML Wed Dec 06 19:34:59 2017 +0100 +++ b/src/Tools/coherent.ML Wed Dec 06 20:43:09 2017 +0100 @@ -29,7 +29,7 @@ (** misc tools **) -val (trace, trace_setup) = Attrib.config_bool @{binding coherent_trace} (K false); +val (trace, trace_setup) = Attrib.config_bool \<^binding>\coherent_trace\ (K false); fun cond_trace ctxt msg = if Config.get ctxt trace then tracing (msg ()) else (); datatype cl_prf = @@ -227,7 +227,7 @@ val _ = Theory.setup (trace_setup #> - Method.setup @{binding coherent} + Method.setup \<^binding>\coherent\ (Attrib.thms >> (fn rules => fn ctxt => METHOD (fn facts => HEADGOAL (coherent_tac ctxt (facts @ rules))))) "prove coherent formula"); diff -r d24dcac5eb4c -r e61557884799 src/Tools/eqsubst.ML --- a/src/Tools/eqsubst.ML Wed Dec 06 19:34:59 2017 +0100 +++ b/src/Tools/eqsubst.ML Wed Dec 06 20:43:09 2017 +0100 @@ -413,7 +413,7 @@ the goal) as well as the theorems to use *) val _ = Theory.setup - (Method.setup @{binding subst} + (Method.setup \<^binding>\subst\ (Scan.lift (Args.mode "asm" -- Scan.optional (Args.parens (Scan.repeat Parse.nat)) [0]) -- Attrib.thms >> (fn ((asm, occs), inthms) => fn ctxt => SIMPLE_METHOD' ((if asm then eqsubst_asm_tac else eqsubst_tac) ctxt occs inthms))) diff -r d24dcac5eb4c -r e61557884799 src/Tools/induct.ML --- a/src/Tools/induct.ML Wed Dec 06 19:34:59 2017 +0100 +++ b/src/Tools/induct.ML Wed Dec 06 20:43:09 2017 +0100 @@ -266,7 +266,7 @@ end; val _ = - Outer_Syntax.command @{command_keyword print_induct_rules} + Outer_Syntax.command \<^command_keyword>\print_induct_rules\ "print induction and cases rules" (Scan.succeed (Toplevel.keep (print_rules o Toplevel.context_of))); @@ -384,13 +384,13 @@ val _ = Theory.local_setup - (Attrib.local_setup @{binding cases} (attrib cases_type cases_pred cases_del) + (Attrib.local_setup \<^binding>\cases\ (attrib cases_type cases_pred cases_del) "declaration of cases rule" #> - Attrib.local_setup @{binding induct} (attrib induct_type induct_pred induct_del) + Attrib.local_setup \<^binding>\induct\ (attrib induct_type induct_pred induct_del) "declaration of induction rule" #> - Attrib.local_setup @{binding coinduct} (attrib coinduct_type coinduct_pred coinduct_del) + Attrib.local_setup \<^binding>\coinduct\ (attrib coinduct_type coinduct_pred coinduct_del) "declaration of coinduction rule" #> - Attrib.local_setup @{binding induct_simp} (Attrib.add_del induct_simp_add induct_simp_del) + Attrib.local_setup \<^binding>\induct_simp\ (Attrib.add_del induct_simp_add induct_simp_del) "declaration of rules for simplifying induction or cases rules"); end; @@ -953,15 +953,15 @@ val _ = Theory.local_setup - (Method.local_setup @{binding cases} + (Method.local_setup \<^binding>\cases\ (Scan.lift (Args.mode no_simpN) -- (Parse.and_list' (Scan.repeat (unless_more_args inst)) -- Scan.option cases_rule) >> (fn (no_simp, (insts, opt_rule)) => fn _ => CONTEXT_METHOD (fn facts => Seq.DETERM (cases_context_tactic (not no_simp) insts opt_rule facts 1)))) "case analysis on types or predicates/sets" #> - gen_induct_setup @{binding induct} induct_context_tactic #> - Method.local_setup @{binding coinduct} + gen_induct_setup \<^binding>\induct\ induct_context_tactic #> + Method.local_setup \<^binding>\coinduct\ (Scan.repeat (unless_more_args inst) -- taking -- Scan.option coinduct_rule >> (fn ((insts, taking), opt_rule) => fn _ => fn facts => Seq.DETERM (coinduct_context_tactic insts taking opt_rule facts 1))) diff -r d24dcac5eb4c -r e61557884799 src/Tools/induct_tacs.ML --- a/src/Tools/induct_tacs.ML Wed Dec 06 19:34:59 2017 +0100 +++ b/src/Tools/induct_tacs.ML Wed Dec 06 20:43:09 2017 +0100 @@ -128,11 +128,11 @@ val _ = Theory.setup - (Method.setup @{binding case_tac} + (Method.setup \<^binding>\case_tac\ (Args.goal_spec -- Scan.lift (Args.embedded_inner_syntax -- Parse.for_fixes) -- opt_rule >> (fn ((quant, (s, xs)), r) => fn ctxt => SIMPLE_METHOD'' quant (case_tac ctxt s xs r))) "unstructured case analysis on types" #> - Method.setup @{binding induct_tac} + Method.setup \<^binding>\induct_tac\ (Args.goal_spec -- varss -- opt_rules >> (fn ((quant, vs), rs) => fn ctxt => SIMPLE_METHOD'' quant (induct_tac ctxt vs rs))) "unstructured induction on types"); diff -r d24dcac5eb4c -r e61557884799 src/Tools/induction.ML --- a/src/Tools/induction.ML Wed Dec 06 19:34:59 2017 +0100 +++ b/src/Tools/induction.ML Wed Dec 06 20:43:09 2017 +0100 @@ -50,6 +50,6 @@ Method.NO_CONTEXT_TACTIC ctxt (induction_context_tactic x1 x2 x3 x4 x5 x6 x7); val _ = - Theory.local_setup (Induct.gen_induct_setup @{binding induction} induction_context_tactic); + Theory.local_setup (Induct.gen_induct_setup \<^binding>\induction\ induction_context_tactic); end diff -r d24dcac5eb4c -r e61557884799 src/Tools/nbe.ML --- a/src/Tools/nbe.ML Wed Dec 06 19:34:59 2017 +0100 +++ b/src/Tools/nbe.ML Wed Dec 06 20:43:09 2017 +0100 @@ -35,7 +35,7 @@ (* generic non-sense *) -val trace = Attrib.setup_config_bool @{binding "nbe_trace"} (K false); +val trace = Attrib.setup_config_bool \<^binding>\nbe_trace\ (K false); fun traced ctxt f x = if Config.get ctxt trace then (tracing (f x); x) else x; @@ -80,7 +80,7 @@ val get_triv_classes = map fst o Triv_Class_Data.get; val (_, triv_of_class) = Context.>>> (Context.map_theory_result - (Thm.add_oracle (@{binding triv_of_class}, fn (thy, T, class) => + (Thm.add_oracle (\<^binding>\triv_of_class\, fn (thy, T, class) => Thm.global_cterm_of thy (Logic.mk_of_class (T, class))))); in @@ -594,7 +594,7 @@ in Thm.mk_binop eq lhs rhs end; val (_, raw_oracle) = Context.>>> (Context.map_theory_result - (Thm.add_oracle (@{binding normalization_by_evaluation}, + (Thm.add_oracle (\<^binding>\normalization_by_evaluation\, fn (nbe_program_idx_tab, ctxt, vs_ty_t, deps, ct) => mk_equals ctxt ct (normalize_term nbe_program_idx_tab ctxt (Thm.term_of ct) vs_ty_t deps)))); diff -r d24dcac5eb4c -r e61557884799 src/Tools/quickcheck.ML --- a/src/Tools/quickcheck.ML Wed Dec 06 19:34:59 2017 +0100 +++ b/src/Tools/quickcheck.ML Wed Dec 06 20:43:09 2017 +0100 @@ -160,30 +160,30 @@ if expect1 = expect2 then expect1 else No_Expectation; (*quickcheck configuration -- default parameters, test generators*) -val batch_tester = Attrib.setup_config_string @{binding quickcheck_batch_tester} (K ""); -val size = Attrib.setup_config_int @{binding quickcheck_size} (K 10); -val iterations = Attrib.setup_config_int @{binding quickcheck_iterations} (K 100); -val depth = Attrib.setup_config_int @{binding quickcheck_depth} (K 10); +val batch_tester = Attrib.setup_config_string \<^binding>\quickcheck_batch_tester\ (K ""); +val size = Attrib.setup_config_int \<^binding>\quickcheck_size\ (K 10); +val iterations = Attrib.setup_config_int \<^binding>\quickcheck_iterations\ (K 100); +val depth = Attrib.setup_config_int \<^binding>\quickcheck_depth\ (K 10); -val no_assms = Attrib.setup_config_bool @{binding quickcheck_no_assms} (K false); -val locale = Attrib.setup_config_string @{binding quickcheck_locale} (K "interpret expand"); -val report = Attrib.setup_config_bool @{binding quickcheck_report} (K true); -val timing = Attrib.setup_config_bool @{binding quickcheck_timing} (K false); -val timeout = Attrib.setup_config_real @{binding quickcheck_timeout} (K 30.0); +val no_assms = Attrib.setup_config_bool \<^binding>\quickcheck_no_assms\ (K false); +val locale = Attrib.setup_config_string \<^binding>\quickcheck_locale\ (K "interpret expand"); +val report = Attrib.setup_config_bool \<^binding>\quickcheck_report\ (K true); +val timing = Attrib.setup_config_bool \<^binding>\quickcheck_timing\ (K false); +val timeout = Attrib.setup_config_real \<^binding>\quickcheck_timeout\ (K 30.0); -val genuine_only = Attrib.setup_config_bool @{binding quickcheck_genuine_only} (K false); -val abort_potential = Attrib.setup_config_bool @{binding quickcheck_abort_potential} (K false); +val genuine_only = Attrib.setup_config_bool \<^binding>\quickcheck_genuine_only\ (K false); +val abort_potential = Attrib.setup_config_bool \<^binding>\quickcheck_abort_potential\ (K false); -val quiet = Attrib.setup_config_bool @{binding quickcheck_quiet} (K false); -val verbose = Attrib.setup_config_bool @{binding quickcheck_verbose} (K false); -val tag = Attrib.setup_config_string @{binding quickcheck_tag} (K ""); +val quiet = Attrib.setup_config_bool \<^binding>\quickcheck_quiet\ (K false); +val verbose = Attrib.setup_config_bool \<^binding>\quickcheck_verbose\ (K false); +val tag = Attrib.setup_config_string \<^binding>\quickcheck_tag\ (K ""); -val use_subtype = Attrib.setup_config_bool @{binding quickcheck_use_subtype} (K false); +val use_subtype = Attrib.setup_config_bool \<^binding>\quickcheck_use_subtype\ (K false); val allow_function_inversion = - Attrib.setup_config_bool @{binding quickcheck_allow_function_inversion} (K false); -val finite_types = Attrib.setup_config_bool @{binding quickcheck_finite_types} (K true); -val finite_type_size = Attrib.setup_config_int @{binding quickcheck_finite_type_size} (K 3); + Attrib.setup_config_bool \<^binding>\quickcheck_allow_function_inversion\ (K false); +val finite_types = Attrib.setup_config_bool \<^binding>\quickcheck_finite_types\ (K true); +val finite_type_size = Attrib.setup_config_int \<^binding>\quickcheck_finite_type_size\ (K 3); datatype test_params = Test_Params of {default_type: typ list, expect : expectation}; @@ -343,7 +343,7 @@ val ctxt = Proof.context_of state; val thy = Proof.theory_of state; - fun strip (Const (@{const_name Pure.all}, _) $ Abs (_, _, t)) = strip t + fun strip (Const (\<^const_name>\Pure.all\, _) $ Abs (_, _, t)) = strip t | strip t = t; val {goal = st, ...} = Proof.raw_goal state; val (gi, frees) = Logic.goal_params (Thm.prop_of st) i; @@ -522,19 +522,19 @@ val parse_arg = Parse.name -- - (Scan.optional (@{keyword "="} |-- + (Scan.optional (\<^keyword>\=\ |-- (((Parse.name || Parse.float_number) >> single) || - (@{keyword "["} |-- Parse.list1 Parse.name --| @{keyword "]"}))) ["true"]); + (\<^keyword>\[\ |-- Parse.list1 Parse.name --| \<^keyword>\]\))) ["true"]); val parse_args = - @{keyword "["} |-- Parse.list1 parse_arg --| @{keyword "]"} || Scan.succeed []; + \<^keyword>\[\ |-- Parse.list1 parse_arg --| \<^keyword>\]\ || Scan.succeed []; val _ = - Outer_Syntax.command @{command_keyword quickcheck_params} "set parameters for random testing" + Outer_Syntax.command \<^command_keyword>\quickcheck_params\ "set parameters for random testing" (parse_args >> (fn args => Toplevel.theory (quickcheck_params_cmd args))); val _ = - Outer_Syntax.command @{command_keyword quickcheck} + Outer_Syntax.command \<^command_keyword>\quickcheck\ "try to find counterexample for subgoal" (parse_args -- Scan.optional Parse.nat 1 >> (fn (args, i) => Toplevel.keep_proof (quickcheck_cmd args i))); @@ -565,7 +565,7 @@ end |> `(fn (outcome_code, _) => outcome_code = genuineN); -val _ = Try.tool_setup (quickcheckN, (20, @{system_option auto_quickcheck}, try_quickcheck)); +val _ = Try.tool_setup (quickcheckN, (20, \<^system_option>\auto_quickcheck\, try_quickcheck)); end; diff -r d24dcac5eb4c -r e61557884799 src/Tools/solve_direct.ML --- a/src/Tools/solve_direct.ML Wed Dec 06 19:34:59 2017 +0100 +++ b/src/Tools/solve_direct.ML Wed Dec 06 20:43:09 2017 +0100 @@ -33,8 +33,8 @@ (* preferences *) -val max_solutions = Attrib.setup_config_int @{binding solve_direct_max_solutions} (K 5); -val strict_warnings = Attrib.setup_config_bool @{binding solve_direct_strict_warnings} (K false); +val max_solutions = Attrib.setup_config_int \<^binding>\solve_direct_max_solutions\ (K 5); +val strict_warnings = Attrib.setup_config_bool \<^binding>\solve_direct_strict_warnings\ (K false); (* solve_direct command *) @@ -90,7 +90,7 @@ val solve_direct = do_solve_direct Normal val _ = - Outer_Syntax.command @{command_keyword solve_direct} + Outer_Syntax.command \<^command_keyword>\solve_direct\ "try to solve conjectures directly with existing theorems" (Scan.succeed (Toplevel.keep_proof (ignore o solve_direct o Toplevel.proof_of))); diff -r d24dcac5eb4c -r e61557884799 src/Tools/subtyping.ML --- a/src/Tools/subtyping.ML Wed Dec 06 19:34:59 2017 +0100 +++ b/src/Tools/subtyping.ML Wed Dec 06 20:43:09 2017 +0100 @@ -882,7 +882,7 @@ (* term check *) -val coercion_enabled = Attrib.setup_config_bool @{binding coercion_enabled} (K false); +val coercion_enabled = Attrib.setup_config_bool \<^binding>\coercion_enabled\ (K false); val _ = Theory.setup @@ -1097,16 +1097,16 @@ val _ = Theory.setup - (Attrib.setup @{binding coercion} + (Attrib.setup \<^binding>\coercion\ (Args.term >> (fn t => Thm.declaration_attribute (K (add_coercion t)))) "declaration of new coercions" #> - Attrib.setup @{binding coercion_delete} + Attrib.setup \<^binding>\coercion_delete\ (Args.term >> (fn t => Thm.declaration_attribute (K (delete_coercion t)))) "deletion of coercions" #> - Attrib.setup @{binding coercion_map} + Attrib.setup \<^binding>\coercion_map\ (Args.term >> (fn t => Thm.declaration_attribute (K (add_type_map t)))) "declaration of new map functions" #> - Attrib.setup @{binding coercion_args} + Attrib.setup \<^binding>\coercion_args\ (Args.const {proper = false, strict = false} -- Scan.lift (Scan.repeat1 parse_coerce_args) >> (fn spec => Thm.declaration_attribute (K (map_coerce_args (Symtab.update spec))))) "declaration of new constants with coercion-invariant arguments"); @@ -1115,7 +1115,7 @@ (* outer syntax commands *) val _ = - Outer_Syntax.command @{command_keyword print_coercions} + Outer_Syntax.command \<^command_keyword>\print_coercions\ "print information about coercions" (Scan.succeed (Toplevel.keep (print_coercions o Toplevel.context_of))); diff -r d24dcac5eb4c -r e61557884799 src/Tools/try.ML --- a/src/Tools/try.ML Wed Dec 06 19:34:59 2017 +0100 +++ b/src/Tools/try.ML Wed Dec 06 20:43:09 2017 +0100 @@ -73,7 +73,7 @@ |> tap (fn NONE => writeln "Tried in vain" | _ => ()); val _ = - Outer_Syntax.command @{command_keyword try} + Outer_Syntax.command \<^command_keyword>\try\ "try a combination of automatic proving and disproving tools" (Scan.succeed (Toplevel.keep_proof (ignore o try_tools o Toplevel.proof_of)));