# HG changeset patch # User wenzelm # Date 1632172502 -7200 # Node ID a9b20bc32fa6b23773c7741bc30d2d63a51aaa29 # Parent 78c1699c7672149d4192948287f5ab9d4885d135 localized command 'syntax' and 'no_syntax'; diff -r 78c1699c7672 -r a9b20bc32fa6 src/Doc/Isar_Ref/Inner_Syntax.thy --- a/src/Doc/Isar_Ref/Inner_Syntax.thy Mon Sep 20 21:56:10 2021 +0200 +++ b/src/Doc/Isar_Ref/Inner_Syntax.thy Mon Sep 20 23:15:02 2021 +0200 @@ -1077,8 +1077,8 @@ text \ \begin{tabular}{rcll} @{command_def "nonterminal"} & : & \theory \ theory\ \\ - @{command_def "syntax"} & : & \theory \ theory\ \\ - @{command_def "no_syntax"} & : & \theory \ theory\ \\ + @{command_def "syntax"} & : & \local_theory \ local_theory\ \\ + @{command_def "no_syntax"} & : & \local_theory \ local_theory\ \\ @{command_def "translations"} & : & \theory \ theory\ \\ @{command_def "no_translations"} & : & \theory \ theory\ \\ @{attribute_def syntax_ast_trace} & : & \attribute\ & default \false\ \\ diff -r 78c1699c7672 -r a9b20bc32fa6 src/Pure/Isar/local_theory.ML --- a/src/Pure/Isar/local_theory.ML Mon Sep 20 21:56:10 2021 +0200 +++ b/src/Pure/Isar/local_theory.ML Mon Sep 20 23:15:02 2021 +0200 @@ -60,6 +60,10 @@ val pretty: local_theory -> Pretty.T list val add_thms_dynamic: binding * (Context.generic -> thm list) -> local_theory -> local_theory val set_defsort: sort -> local_theory -> local_theory + val syntax: bool -> Syntax.mode -> (string * typ * mixfix) list -> + local_theory -> local_theory + val syntax_cmd: bool -> Syntax.mode -> (string * string * mixfix) list -> + local_theory -> local_theory val type_notation: bool -> Syntax.mode -> (typ * mixfix) list -> local_theory -> local_theory val notation: bool -> Syntax.mode -> (term * mixfix) list -> local_theory -> local_theory val type_alias: binding -> string -> local_theory -> local_theory @@ -297,6 +301,22 @@ (K (Context.mapping (Sign.set_defsort S) (Proof_Context.set_defsort S))); +(* syntax *) + +fun gen_syntax prep_type add mode raw_args lthy = + let + val args = map (fn (c, T, mx) => (c, prep_type lthy T, mx)) raw_args; + val args' = map (fn (c, T, mx) => (c, T, Mixfix.reset_pos mx)) args; + val _ = lthy |> Context_Position.is_visible lthy ? Proof_Context.syntax add mode args; + in + declaration {syntax = true, pervasive = false} + (fn _ => Proof_Context.generic_syntax add mode args') lthy + end; + +val syntax = gen_syntax (K I); +val syntax_cmd = gen_syntax Proof_Context.read_typ_syntax; + + (* notation *) fun type_notation add mode raw_args lthy = diff -r 78c1699c7672 -r a9b20bc32fa6 src/Pure/Isar/proof_context.ML --- a/src/Pure/Isar/proof_context.ML Mon Sep 20 21:56:10 2021 +0200 +++ b/src/Pure/Isar/proof_context.ML Mon Sep 20 23:15:02 2021 +0200 @@ -156,6 +156,10 @@ val check_case: Proof.context -> bool -> string * Position.T -> binding option list -> Rule_Cases.T val check_syntax_const: Proof.context -> string * Position.T -> string + val syntax: bool -> Syntax.mode -> (string * typ * mixfix) list -> + Proof.context -> Proof.context + val generic_syntax: bool -> Syntax.mode -> (string * typ * mixfix) list -> + Context.generic -> Context.generic val type_notation: bool -> Syntax.mode -> (typ * mixfix) list -> Proof.context -> Proof.context val notation: bool -> Syntax.mode -> (term * mixfix) list -> Proof.context -> Proof.context val generic_type_notation: bool -> Syntax.mode -> (typ * mixfix) list -> morphism -> @@ -1130,6 +1134,13 @@ if is_some (Syntax.lookup_const (syn_of ctxt) c) then c else error ("Unknown syntax const: " ^ quote c ^ Position.here pos); +fun syntax add mode args ctxt = + let val args' = map (pair Local_Syntax.Const) args + in ctxt |> map_syntax (#2 o Local_Syntax.update_modesyntax ctxt add mode args') end; + +fun generic_syntax add mode args = + Context.mapping (Sign.syntax add mode args) (syntax add mode args); + (* notation *) diff -r 78c1699c7672 -r a9b20bc32fa6 src/Pure/Pure.thy --- a/src/Pure/Pure.thy Mon Sep 20 21:56:10 2021 +0200 +++ b/src/Pure/Pure.thy Mon Sep 20 23:15:02 2021 +0200 @@ -384,14 +384,14 @@ (Parse.and_list1 Parse.binding >> (Toplevel.theory o Sign.add_nonterminals_global)); val _ = - Outer_Syntax.command \<^command_keyword>\syntax\ "add raw syntax clauses" + Outer_Syntax.local_theory \<^command_keyword>\syntax\ "add raw syntax clauses" (Parse.syntax_mode -- Scan.repeat1 Parse.const_decl - >> (Toplevel.theory o uncurry (Sign.syntax_cmd true))); + >> uncurry (Local_Theory.syntax_cmd true)); val _ = - Outer_Syntax.command \<^command_keyword>\no_syntax\ "delete raw syntax clauses" + Outer_Syntax.local_theory \<^command_keyword>\no_syntax\ "delete raw syntax clauses" (Parse.syntax_mode -- Scan.repeat1 Parse.const_decl - >> (Toplevel.theory o uncurry (Sign.syntax_cmd false))); + >> uncurry (Local_Theory.syntax_cmd false)); val trans_pat = Scan.optional