| author | Fabian Huch <huch@in.tum.de> | 
| Tue, 04 Jun 2024 09:02:36 +0200 | |
| changeset 80246 | 245dd5f82462 | 
| parent 80074 | 951c371c1cd9 | 
| child 80897 | 5328d67ec647 | 
| permissions | -rw-r--r-- | 
| 
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 | 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: 
62765 
diff
changeset
 | 
11  | 
val syntax_of: T -> Syntax.syntax  | 
| 18997 | 12  | 
val init: theory -> T  | 
13  | 
val rebuild: theory -> 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 | 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: 
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 | 22  | 
end;  | 
23  | 
||
| 33387 | 24  | 
structure Local_Syntax: LOCAL_SYNTAX =  | 
| 18997 | 25  | 
struct  | 
26  | 
||
27  | 
(* datatype T *)  | 
|
28  | 
||
| 19386 | 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 | 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: 
52144 
diff
changeset
 | 
37  | 
mixfixes: local_mixfix list};  | 
| 18997 | 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 | 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 | 44  | 
|
| 
19369
 
a4374b41c9bf
simplified handling of authentic syntax (cf. early externing in consts.ML);
 
wenzelm 
parents: 
19016 
diff
changeset
 | 
45  | 
fun is_consistent thy (Syntax {thy_syntax, ...}) =
 | 
| 
80074
 
951c371c1cd9
clarified names: discontinue odd convention from 3 decades ago;
 
wenzelm 
parents: 
62765 
diff
changeset
 | 
46  | 
Syntax.eq_syntax (Sign.syntax_of thy, thy_syntax);  | 
| 18997 | 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 | 49  | 
|
50  | 
||
51  | 
(* build syntax *)  | 
|
52  | 
||
| 
59152
 
66e6c539a36d
more frugal Local_Syntax.init -- maintain idents within context;
 
wenzelm 
parents: 
52144 
diff
changeset
 | 
53  | 
fun build_syntax thy mode mixfixes =  | 
| 18997 | 54  | 
let  | 
| 
80074
 
951c371c1cd9
clarified names: discontinue odd convention from 3 decades ago;
 
wenzelm 
parents: 
62765 
diff
changeset
 | 
55  | 
val thy_syntax = Sign.syntax_of thy;  | 
| 
35412
 
b8dead547d9e
more uniform treatment of syntax for types vs. consts;
 
wenzelm 
parents: 
33387 
diff
changeset
 | 
56  | 
fun update_gram ((true, add, m), decls) = Syntax.update_type_gram add m decls  | 
| 
 
b8dead547d9e
more uniform treatment of syntax for types vs. consts;
 
wenzelm 
parents: 
33387 
diff
changeset
 | 
57  | 
| update_gram ((false, add, m), decls) =  | 
| 59841 | 58  | 
Syntax.update_const_gram add (Sign.logical_types thy) m decls;  | 
| 
35412
 
b8dead547d9e
more uniform treatment of syntax for types vs. consts;
 
wenzelm 
parents: 
33387 
diff
changeset
 | 
59  | 
|
| 18997 | 60  | 
val local_syntax = thy_syntax  | 
| 
35412
 
b8dead547d9e
more uniform treatment of syntax for types vs. consts;
 
wenzelm 
parents: 
33387 
diff
changeset
 | 
61  | 
|> 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
 | 
62  | 
in make_syntax (thy_syntax, local_syntax, mode, mixfixes) end;  | 
| 19541 | 63  | 
|
| 
59152
 
66e6c539a36d
more frugal Local_Syntax.init -- maintain idents within context;
 
wenzelm 
parents: 
52144 
diff
changeset
 | 
64  | 
fun init thy =  | 
| 
80074
 
951c371c1cd9
clarified names: discontinue odd convention from 3 decades ago;
 
wenzelm 
parents: 
62765 
diff
changeset
 | 
65  | 
let val thy_syntax = Sign.syntax_of thy  | 
| 
59152
 
66e6c539a36d
more frugal Local_Syntax.init -- maintain idents within context;
 
wenzelm 
parents: 
52144 
diff
changeset
 | 
66  | 
in make_syntax (thy_syntax, thy_syntax, Syntax.mode_default, []) end;  | 
| 18997 | 67  | 
|
| 
59152
 
66e6c539a36d
more frugal Local_Syntax.init -- maintain idents within context;
 
wenzelm 
parents: 
52144 
diff
changeset
 | 
68  | 
fun rebuild thy (syntax as Syntax {mode, mixfixes, ...}) =
 | 
| 19541 | 69  | 
if is_consistent thy syntax then syntax  | 
| 
59152
 
66e6c539a36d
more frugal Local_Syntax.init -- maintain idents within context;
 
wenzelm 
parents: 
52144 
diff
changeset
 | 
70  | 
else build_syntax thy mode mixfixes;  | 
| 18997 | 71  | 
|
| 19541 | 72  | 
|
| 18997 | 73  | 
(* mixfix declarations *)  | 
74  | 
||
| 
35412
 
b8dead547d9e
more uniform treatment of syntax for types vs. consts;
 
wenzelm 
parents: 
33387 
diff
changeset
 | 
75  | 
datatype kind = Type | Const | Fixed;  | 
| 
 
b8dead547d9e
more uniform treatment of syntax for types vs. consts;
 
wenzelm 
parents: 
33387 
diff
changeset
 | 
76  | 
|
| 18997 | 77  | 
local  | 
78  | 
||
| 62752 | 79  | 
fun prep_mixfix _ _ (_, (_, _, Structure _)) = NONE  | 
| 
35412
 
b8dead547d9e
more uniform treatment of syntax for types vs. consts;
 
wenzelm 
parents: 
33387 
diff
changeset
 | 
80  | 
| 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
 | 
81  | 
| 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
 | 
82  | 
| prep_mixfix add mode (Fixed, (x, T, mx)) =  | 
| 
42290
 
b1f544c84040
discontinued special treatment of structure Lexicon;
 
wenzelm 
parents: 
42288 
diff
changeset
 | 
83  | 
SOME ((x, true), ((false, add, mode), (Lexicon.mark_fixed x, T, mx)));  | 
| 19016 | 84  | 
|
| 62752 | 85  | 
fun prep_struct (Fixed, (c, _, Structure _)) = SOME c  | 
| 
62765
 
5b95a12b7b19
tuned messages -- position is usually missing here;
 
wenzelm 
parents: 
62752 
diff
changeset
 | 
86  | 
  | prep_struct (_, (c, _, Structure _)) = error ("Bad structure mixfix declaration for " ^ quote c)
 | 
| 19016 | 87  | 
| prep_struct _ = NONE;  | 
| 18997 | 88  | 
|
89  | 
in  | 
|
90  | 
||
| 
59152
 
66e6c539a36d
more frugal Local_Syntax.init -- maintain idents within context;
 
wenzelm 
parents: 
52144 
diff
changeset
 | 
91  | 
fun update_syntax ctxt add raw_decls (syntax as (Syntax {mode, mixfixes, ...})) =
 | 
| 62752 | 92  | 
(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
 | 
93  | 
[] => (NONE, syntax)  | 
| 19016 | 94  | 
| decls =>  | 
95  | 
let  | 
|
| 
35412
 
b8dead547d9e
more uniform treatment of syntax for types vs. consts;
 
wenzelm 
parents: 
33387 
diff
changeset
 | 
96  | 
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
 | 
97  | 
val new_structs = map_filter prep_struct decls;  | 
| 
 
0664f10a4094
update_modesyntax: may delete 'structure' notation as well;
 
wenzelm 
parents: 
24951 
diff
changeset
 | 
98  | 
val mixfixes' = rev new_mixfixes @ mixfixes;  | 
| 18997 | 99  | 
|
| 
59152
 
66e6c539a36d
more frugal Local_Syntax.init -- maintain idents within context;
 
wenzelm 
parents: 
52144 
diff
changeset
 | 
100  | 
val idents = Syntax_Trans.get_idents ctxt;  | 
| 
 
66e6c539a36d
more frugal Local_Syntax.init -- maintain idents within context;
 
wenzelm 
parents: 
52144 
diff
changeset
 | 
101  | 
val idents' =  | 
| 
 
66e6c539a36d
more frugal Local_Syntax.init -- maintain idents within context;
 
wenzelm 
parents: 
52144 
diff
changeset
 | 
102  | 
          {structs =
 | 
| 
 
66e6c539a36d
more frugal Local_Syntax.init -- maintain idents within context;
 
wenzelm 
parents: 
52144 
diff
changeset
 | 
103  | 
if add then #structs idents @ new_structs  | 
| 
 
66e6c539a36d
more frugal Local_Syntax.init -- maintain idents within context;
 
wenzelm 
parents: 
52144 
diff
changeset
 | 
104  | 
else subtract (op =) new_structs (#structs idents),  | 
| 
 
66e6c539a36d
more frugal Local_Syntax.init -- maintain idents within context;
 
wenzelm 
parents: 
52144 
diff
changeset
 | 
105  | 
fixes = fold (fn ((x, true), _) => cons x | _ => I) mixfixes' []};  | 
| 
 
66e6c539a36d
more frugal Local_Syntax.init -- maintain idents within context;
 
wenzelm 
parents: 
52144 
diff
changeset
 | 
106  | 
|
| 
 
66e6c539a36d
more frugal Local_Syntax.init -- maintain idents within context;
 
wenzelm 
parents: 
52144 
diff
changeset
 | 
107  | 
val syntax' = build_syntax (Proof_Context.theory_of ctxt) mode mixfixes';  | 
| 
 
66e6c539a36d
more frugal Local_Syntax.init -- maintain idents within context;
 
wenzelm 
parents: 
52144 
diff
changeset
 | 
108  | 
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
 | 
109  | 
|
| 
 
66e6c539a36d
more frugal Local_Syntax.init -- maintain idents within context;
 
wenzelm 
parents: 
52144 
diff
changeset
 | 
110  | 
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
 | 
111  | 
|
| 18997 | 112  | 
end;  | 
113  | 
||
114  | 
||
| 19660 | 115  | 
(* syntax mode *)  | 
116  | 
||
| 
59152
 
66e6c539a36d
more frugal Local_Syntax.init -- maintain idents within context;
 
wenzelm 
parents: 
52144 
diff
changeset
 | 
117  | 
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
 | 
118  | 
(thy_syntax, local_syntax, mode, mixfixes));  | 
| 19660 | 119  | 
|
120  | 
fun restore_mode (Syntax {mode, ...}) = set_mode mode;
 | 
|
121  | 
||
| 
59152
 
66e6c539a36d
more frugal Local_Syntax.init -- maintain idents within context;
 
wenzelm 
parents: 
52144 
diff
changeset
 | 
122  | 
fun update_modesyntax ctxt add mode args syntax =  | 
| 
 
66e6c539a36d
more frugal Local_Syntax.init -- maintain idents within context;
 
wenzelm 
parents: 
52144 
diff
changeset
 | 
123  | 
syntax |> set_mode mode |> update_syntax ctxt add args ||> restore_mode syntax;  | 
| 19660 | 124  | 
|
| 18997 | 125  | 
end;  |