| author | wenzelm | 
| Tue, 10 Dec 2024 21:06:04 +0100 | |
| changeset 81572 | 693a95492008 | 
| parent 81149 | 0e506128c14a | 
| child 81591 | d570d215e380 | 
| permissions | -rw-r--r-- | 
| 42240 
5a4d30cd47a7
moved Isar/local_syntax.ML to Syntax/local_syntax.ML;
 wenzelm parents: 
35429diff
changeset | 1 | (* Title: Pure/Syntax/local_syntax.ML | 
| 18997 | 2 | Author: Makarius | 
| 3 | ||
| 51655 | 4 | Local syntax depending on theory syntax, with special support for | 
| 5 | implicit structure references. | |
| 18997 | 6 | *) | 
| 7 | ||
| 8 | signature LOCAL_SYNTAX = | |
| 9 | sig | |
| 10 | type T | |
| 80074 
951c371c1cd9
clarified names: discontinue odd convention from 3 decades ago;
 wenzelm parents: 
62765diff
changeset | 11 | val syntax_of: T -> Syntax.syntax | 
| 18997 | 12 | val init: theory -> T | 
| 80897 
5328d67ec647
more explicit context for syn_ext/mixfix operations, but it often degenerates to background theory;
 wenzelm parents: 
80074diff
changeset | 13 | val rebuild: Proof.context -> T -> T | 
| 35412 
b8dead547d9e
more uniform treatment of syntax for types vs. consts;
 wenzelm parents: 
33387diff
changeset | 14 | datatype kind = Type | Const | Fixed | 
| 59152 
66e6c539a36d
more frugal Local_Syntax.init -- maintain idents within context;
 wenzelm parents: 
52144diff
changeset | 15 | val add_syntax: Proof.context -> (kind * (string * typ * mixfix)) list -> | 
| 
66e6c539a36d
more frugal Local_Syntax.init -- maintain idents within context;
 wenzelm parents: 
52144diff
changeset | 16 |     T -> {structs: string list, fixes: string list} option * T
 | 
| 20785 | 17 | val set_mode: Syntax.mode -> T -> T | 
| 19541 | 18 | val restore_mode: T -> T -> T | 
| 59152 
66e6c539a36d
more frugal Local_Syntax.init -- maintain idents within context;
 wenzelm parents: 
52144diff
changeset | 19 | val update_modesyntax: Proof.context -> bool -> Syntax.mode -> | 
| 
66e6c539a36d
more frugal Local_Syntax.init -- maintain idents within context;
 wenzelm parents: 
52144diff
changeset | 20 | (kind * (string * typ * mixfix)) list -> | 
| 
66e6c539a36d
more frugal Local_Syntax.init -- maintain idents within context;
 wenzelm parents: 
52144diff
changeset | 21 |     T -> {structs: string list, fixes: string list} option * T
 | 
| 18997 | 22 | end; | 
| 23 | ||
| 33387 | 24 | structure Local_Syntax: LOCAL_SYNTAX = | 
| 18997 | 25 | struct | 
| 26 | ||
| 27 | (* datatype T *) | |
| 28 | ||
| 19386 | 29 | type local_mixfix = | 
| 81149 | 30 |   {name: string, is_fixed: bool} *
 | 
| 31 |   ({is_type: bool, add: bool, mode: Syntax.mode} * (string * typ * mixfix));
 | |
| 19386 | 32 | |
| 18997 | 33 | datatype T = Syntax of | 
| 34 |  {thy_syntax: Syntax.syntax,
 | |
| 35 | local_syntax: Syntax.syntax, | |
| 20785 | 36 | mode: Syntax.mode, | 
| 59152 
66e6c539a36d
more frugal Local_Syntax.init -- maintain idents within context;
 wenzelm parents: 
52144diff
changeset | 37 | mixfixes: local_mixfix list}; | 
| 18997 | 38 | |
| 59152 
66e6c539a36d
more frugal Local_Syntax.init -- maintain idents within context;
 wenzelm parents: 
52144diff
changeset | 39 | fun make_syntax (thy_syntax, local_syntax, mode, mixfixes) = | 
| 
66e6c539a36d
more frugal Local_Syntax.init -- maintain idents within context;
 wenzelm parents: 
52144diff
changeset | 40 |   Syntax {thy_syntax = thy_syntax, local_syntax = local_syntax, mode = mode, mixfixes = mixfixes};
 | 
| 18997 | 41 | |
| 59152 
66e6c539a36d
more frugal Local_Syntax.init -- maintain idents within context;
 wenzelm parents: 
52144diff
changeset | 42 | fun map_syntax f (Syntax {thy_syntax, local_syntax, mode, mixfixes}) =
 | 
| 
66e6c539a36d
more frugal Local_Syntax.init -- maintain idents within context;
 wenzelm parents: 
52144diff
changeset | 43 | make_syntax (f (thy_syntax, local_syntax, mode, mixfixes)); | 
| 18997 | 44 | |
| 80897 
5328d67ec647
more explicit context for syn_ext/mixfix operations, but it often degenerates to background theory;
 wenzelm parents: 
80074diff
changeset | 45 | fun is_consistent ctxt (Syntax {thy_syntax, ...}) =
 | 
| 
5328d67ec647
more explicit context for syn_ext/mixfix operations, but it often degenerates to background theory;
 wenzelm parents: 
80074diff
changeset | 46 | Syntax.eq_syntax (Sign.syntax_of (Proof_Context.theory_of ctxt), thy_syntax); | 
| 18997 | 47 | |
| 80074 
951c371c1cd9
clarified names: discontinue odd convention from 3 decades ago;
 wenzelm parents: 
62765diff
changeset | 48 | fun syntax_of (Syntax {local_syntax, ...}) = local_syntax;
 | 
| 18997 | 49 | |
| 50 | ||
| 51 | (* build syntax *) | |
| 52 | ||
| 80897 
5328d67ec647
more explicit context for syn_ext/mixfix operations, but it often degenerates to background theory;
 wenzelm parents: 
80074diff
changeset | 53 | fun build_syntax ctxt mode mixfixes = | 
| 18997 | 54 | let | 
| 80897 
5328d67ec647
more explicit context for syn_ext/mixfix operations, but it often degenerates to background theory;
 wenzelm parents: 
80074diff
changeset | 55 | val thy = Proof_Context.theory_of ctxt; | 
| 80074 
951c371c1cd9
clarified names: discontinue odd convention from 3 decades ago;
 wenzelm parents: 
62765diff
changeset | 56 | val thy_syntax = Sign.syntax_of thy; | 
| 81149 | 57 |     fun update_gram ({is_type = true, add, mode}, decls) =
 | 
| 58 | Syntax.update_type_gram ctxt add mode decls | |
| 59 |       | update_gram ({is_type = false, add, mode}, decls) =
 | |
| 60 | Syntax.update_const_gram ctxt add (Sign.logical_types thy) mode decls; | |
| 35412 
b8dead547d9e
more uniform treatment of syntax for types vs. consts;
 wenzelm parents: 
33387diff
changeset | 61 | |
| 18997 | 62 | val local_syntax = thy_syntax | 
| 35412 
b8dead547d9e
more uniform treatment of syntax for types vs. consts;
 wenzelm parents: 
33387diff
changeset | 63 | |> fold update_gram (AList.coalesce (op =) (rev (map snd mixfixes))); | 
| 59152 
66e6c539a36d
more frugal Local_Syntax.init -- maintain idents within context;
 wenzelm parents: 
52144diff
changeset | 64 | in make_syntax (thy_syntax, local_syntax, mode, mixfixes) end; | 
| 19541 | 65 | |
| 59152 
66e6c539a36d
more frugal Local_Syntax.init -- maintain idents within context;
 wenzelm parents: 
52144diff
changeset | 66 | fun init thy = | 
| 80074 
951c371c1cd9
clarified names: discontinue odd convention from 3 decades ago;
 wenzelm parents: 
62765diff
changeset | 67 | let val thy_syntax = Sign.syntax_of thy | 
| 59152 
66e6c539a36d
more frugal Local_Syntax.init -- maintain idents within context;
 wenzelm parents: 
52144diff
changeset | 68 | in make_syntax (thy_syntax, thy_syntax, Syntax.mode_default, []) end; | 
| 18997 | 69 | |
| 80897 
5328d67ec647
more explicit context for syn_ext/mixfix operations, but it often degenerates to background theory;
 wenzelm parents: 
80074diff
changeset | 70 | fun rebuild ctxt (syntax as Syntax {mode, mixfixes, ...}) =
 | 
| 
5328d67ec647
more explicit context for syn_ext/mixfix operations, but it often degenerates to background theory;
 wenzelm parents: 
80074diff
changeset | 71 | if is_consistent ctxt syntax then syntax | 
| 
5328d67ec647
more explicit context for syn_ext/mixfix operations, but it often degenerates to background theory;
 wenzelm parents: 
80074diff
changeset | 72 | else build_syntax ctxt mode mixfixes; | 
| 18997 | 73 | |
| 19541 | 74 | |
| 18997 | 75 | (* mixfix declarations *) | 
| 76 | ||
| 35412 
b8dead547d9e
more uniform treatment of syntax for types vs. consts;
 wenzelm parents: 
33387diff
changeset | 77 | datatype kind = Type | Const | Fixed; | 
| 
b8dead547d9e
more uniform treatment of syntax for types vs. consts;
 wenzelm parents: 
33387diff
changeset | 78 | |
| 18997 | 79 | local | 
| 80 | ||
| 81149 | 81 | fun prep_mixfix _ _ (_, (_, _, Structure _)) = (NONE: local_mixfix option) | 
| 82 | | prep_mixfix add mode (Type, decl as (x, _, _)) = | |
| 83 |       SOME ({name = x, is_fixed = false}, ({is_type = true, add = add, mode = mode}, decl))
 | |
| 84 | | prep_mixfix add mode (Const, decl as (x, _, _)) = | |
| 85 |       SOME ({name = x, is_fixed = false}, ({is_type = false, add = add, mode = mode}, decl))
 | |
| 35412 
b8dead547d9e
more uniform treatment of syntax for types vs. consts;
 wenzelm parents: 
33387diff
changeset | 86 | | prep_mixfix add mode (Fixed, (x, T, mx)) = | 
| 81149 | 87 |       SOME (({name = x, is_fixed = true},
 | 
| 88 |         ({is_type = false, add = add, mode = mode}, (Lexicon.mark_fixed x, T, mx))));
 | |
| 19016 | 89 | |
| 62752 | 90 | fun prep_struct (Fixed, (c, _, Structure _)) = SOME c | 
| 62765 
5b95a12b7b19
tuned messages -- position is usually missing here;
 wenzelm parents: 
62752diff
changeset | 91 |   | prep_struct (_, (c, _, Structure _)) = error ("Bad structure mixfix declaration for " ^ quote c)
 | 
| 19016 | 92 | | prep_struct _ = NONE; | 
| 18997 | 93 | |
| 94 | in | |
| 95 | ||
| 59152 
66e6c539a36d
more frugal Local_Syntax.init -- maintain idents within context;
 wenzelm parents: 
52144diff
changeset | 96 | fun update_syntax ctxt add raw_decls (syntax as (Syntax {mode, mixfixes, ...})) =
 | 
| 62752 | 97 | (case filter_out (Mixfix.is_empty o #3 o #2) raw_decls of | 
| 59152 
66e6c539a36d
more frugal Local_Syntax.init -- maintain idents within context;
 wenzelm parents: 
52144diff
changeset | 98 | [] => (NONE, syntax) | 
| 19016 | 99 | | decls => | 
| 100 | let | |
| 35412 
b8dead547d9e
more uniform treatment of syntax for types vs. consts;
 wenzelm parents: 
33387diff
changeset | 101 | val new_mixfixes = map_filter (prep_mixfix add mode) decls; | 
| 24954 
0664f10a4094
update_modesyntax: may delete 'structure' notation as well;
 wenzelm parents: 
24951diff
changeset | 102 | val new_structs = map_filter prep_struct decls; | 
| 
0664f10a4094
update_modesyntax: may delete 'structure' notation as well;
 wenzelm parents: 
24951diff
changeset | 103 | val mixfixes' = rev new_mixfixes @ mixfixes; | 
| 18997 | 104 | |
| 59152 
66e6c539a36d
more frugal Local_Syntax.init -- maintain idents within context;
 wenzelm parents: 
52144diff
changeset | 105 | val idents = Syntax_Trans.get_idents ctxt; | 
| 
66e6c539a36d
more frugal Local_Syntax.init -- maintain idents within context;
 wenzelm parents: 
52144diff
changeset | 106 | val idents' = | 
| 
66e6c539a36d
more frugal Local_Syntax.init -- maintain idents within context;
 wenzelm parents: 
52144diff
changeset | 107 |           {structs =
 | 
| 
66e6c539a36d
more frugal Local_Syntax.init -- maintain idents within context;
 wenzelm parents: 
52144diff
changeset | 108 | if add then #structs idents @ new_structs | 
| 
66e6c539a36d
more frugal Local_Syntax.init -- maintain idents within context;
 wenzelm parents: 
52144diff
changeset | 109 | else subtract (op =) new_structs (#structs idents), | 
| 81149 | 110 |            fixes = fold (fn ({name = x, is_fixed = true}, _) => cons x | _ => I) mixfixes' []};
 | 
| 59152 
66e6c539a36d
more frugal Local_Syntax.init -- maintain idents within context;
 wenzelm parents: 
52144diff
changeset | 111 | |
| 80897 
5328d67ec647
more explicit context for syn_ext/mixfix operations, but it often degenerates to background theory;
 wenzelm parents: 
80074diff
changeset | 112 | val syntax' = build_syntax ctxt mode mixfixes'; | 
| 59152 
66e6c539a36d
more frugal Local_Syntax.init -- maintain idents within context;
 wenzelm parents: 
52144diff
changeset | 113 | in (if idents = idents' then NONE else SOME idents', syntax') end); | 
| 
66e6c539a36d
more frugal Local_Syntax.init -- maintain idents within context;
 wenzelm parents: 
52144diff
changeset | 114 | |
| 
66e6c539a36d
more frugal Local_Syntax.init -- maintain idents within context;
 wenzelm parents: 
52144diff
changeset | 115 | fun add_syntax ctxt = update_syntax ctxt true; | 
| 24951 
834b8c2b9553
replaced add_modesyntax by general update_modesyntax (add or del);
 wenzelm parents: 
20785diff
changeset | 116 | |
| 18997 | 117 | end; | 
| 118 | ||
| 119 | ||
| 19660 | 120 | (* syntax mode *) | 
| 121 | ||
| 59152 
66e6c539a36d
more frugal Local_Syntax.init -- maintain idents within context;
 wenzelm parents: 
52144diff
changeset | 122 | fun set_mode mode = map_syntax (fn (thy_syntax, local_syntax, _, mixfixes) => | 
| 
66e6c539a36d
more frugal Local_Syntax.init -- maintain idents within context;
 wenzelm parents: 
52144diff
changeset | 123 | (thy_syntax, local_syntax, mode, mixfixes)); | 
| 19660 | 124 | |
| 125 | fun restore_mode (Syntax {mode, ...}) = set_mode mode;
 | |
| 126 | ||
| 59152 
66e6c539a36d
more frugal Local_Syntax.init -- maintain idents within context;
 wenzelm parents: 
52144diff
changeset | 127 | fun update_modesyntax ctxt add mode args syntax = | 
| 
66e6c539a36d
more frugal Local_Syntax.init -- maintain idents within context;
 wenzelm parents: 
52144diff
changeset | 128 | syntax |> set_mode mode |> update_syntax ctxt add args ||> restore_mode syntax; | 
| 19660 | 129 | |
| 18997 | 130 | end; |