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