| author | wenzelm | 
| Thu, 11 Oct 2012 15:06:27 +0200 | |
| changeset 49817 | 85b37aece3b3 | 
| parent 42290 | b1f544c84040 | 
| child 51655 | 28d6eb23522c | 
| 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  | 
||
4  | 
Local syntax depending on theory syntax.  | 
|
5  | 
*)  | 
|
6  | 
||
7  | 
signature LOCAL_SYNTAX =  | 
|
8  | 
sig  | 
|
9  | 
type T  | 
|
10  | 
val syn_of: T -> Syntax.syntax  | 
|
| 
35429
 
afa8cf9e63d8
authentic syntax for classes and type constructors;
 
wenzelm 
parents: 
35412 
diff
changeset
 | 
11  | 
  val idents_of: T -> {structs: string list, fixes: string list}
 | 
| 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  | 
| 
 
b8dead547d9e
more uniform treatment of syntax for types vs. consts;
 
wenzelm 
parents: 
33387 
diff
changeset
 | 
15  | 
val add_syntax: theory -> (kind * (string * typ * mixfix)) list -> T -> T  | 
| 20785 | 16  | 
val set_mode: Syntax.mode -> T -> T  | 
| 19541 | 17  | 
val restore_mode: T -> T -> T  | 
| 
24951
 
834b8c2b9553
replaced add_modesyntax by general update_modesyntax (add or del);
 
wenzelm 
parents: 
20785 
diff
changeset
 | 
18  | 
val update_modesyntax: theory -> bool -> Syntax.mode ->  | 
| 
35412
 
b8dead547d9e
more uniform treatment of syntax for types vs. consts;
 
wenzelm 
parents: 
33387 
diff
changeset
 | 
19  | 
(kind * (string * typ * mixfix)) list -> T -> T  | 
| 18997 | 20  | 
end;  | 
21  | 
||
| 33387 | 22  | 
structure Local_Syntax: LOCAL_SYNTAX =  | 
| 18997 | 23  | 
struct  | 
24  | 
||
25  | 
(* datatype T *)  | 
|
26  | 
||
| 19386 | 27  | 
type local_mixfix =  | 
| 
35412
 
b8dead547d9e
more uniform treatment of syntax for types vs. consts;
 
wenzelm 
parents: 
33387 
diff
changeset
 | 
28  | 
(string * bool) * (*name, fixed?*)  | 
| 
 
b8dead547d9e
more uniform treatment of syntax for types vs. consts;
 
wenzelm 
parents: 
33387 
diff
changeset
 | 
29  | 
((bool * bool * Syntax.mode) * (string * typ * mixfix)); (*type?, add?, mode, declaration*)  | 
| 19386 | 30  | 
|
| 18997 | 31  | 
datatype T = Syntax of  | 
32  | 
 {thy_syntax: Syntax.syntax,
 | 
|
33  | 
local_syntax: Syntax.syntax,  | 
|
| 20785 | 34  | 
mode: Syntax.mode,  | 
| 19386 | 35  | 
mixfixes: local_mixfix list,  | 
| 
19369
 
a4374b41c9bf
simplified handling of authentic syntax (cf. early externing in consts.ML);
 
wenzelm 
parents: 
19016 
diff
changeset
 | 
36  | 
idents: string list * string list};  | 
| 18997 | 37  | 
|
| 19541 | 38  | 
fun make_syntax (thy_syntax, local_syntax, mode, mixfixes, idents) =  | 
| 18997 | 39  | 
  Syntax {thy_syntax = thy_syntax, local_syntax = local_syntax,
 | 
| 19541 | 40  | 
mode = mode, mixfixes = mixfixes, idents = idents};  | 
| 18997 | 41  | 
|
| 19541 | 42  | 
fun map_syntax f (Syntax {thy_syntax, local_syntax, mode, mixfixes, idents}) =
 | 
43  | 
make_syntax (f (thy_syntax, local_syntax, mode, mixfixes, idents));  | 
|
| 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, ...}) =
 | 
| 18997 | 46  | 
Syntax.eq_syntax (Sign.syn_of thy, thy_syntax);  | 
47  | 
||
48  | 
fun syn_of (Syntax {local_syntax, ...}) = local_syntax;
 | 
|
| 
35429
 
afa8cf9e63d8
authentic syntax for classes and type constructors;
 
wenzelm 
parents: 
35412 
diff
changeset
 | 
49  | 
fun idents_of (Syntax {idents = (structs, fixes), ...}) = {structs = structs, fixes = fixes};
 | 
| 18997 | 50  | 
|
51  | 
||
52  | 
(* build syntax *)  | 
|
53  | 
||
| 19541 | 54  | 
fun build_syntax thy mode mixfixes (idents as (structs, _)) =  | 
| 18997 | 55  | 
let  | 
56  | 
val thy_syntax = Sign.syn_of thy;  | 
|
| 
35412
 
b8dead547d9e
more uniform treatment of syntax for types vs. consts;
 
wenzelm 
parents: 
33387 
diff
changeset
 | 
57  | 
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
 | 
58  | 
| update_gram ((false, add, m), decls) =  | 
| 
 
b8dead547d9e
more uniform treatment of syntax for types vs. consts;
 
wenzelm 
parents: 
33387 
diff
changeset
 | 
59  | 
Syntax.update_const_gram add (Sign.is_logtype thy) m decls;  | 
| 
 
b8dead547d9e
more uniform treatment of syntax for types vs. consts;
 
wenzelm 
parents: 
33387 
diff
changeset
 | 
60  | 
|
| 42284 | 61  | 
val (atrs, trs, trs', atrs') = Syntax_Trans.struct_trfuns structs;  | 
| 18997 | 62  | 
val local_syntax = thy_syntax  | 
| 
25390
 
8bfa6566ac6b
syntax operations: turned extend'' into update'' (absorb duplicates);
 
wenzelm 
parents: 
25385 
diff
changeset
 | 
63  | 
|> Syntax.update_trfuns  | 
| 
42288
 
2074b31650e6
discontinued special treatment of structure Syntax_Ext (formerly Syn_Ext);
 
wenzelm 
parents: 
42284 
diff
changeset
 | 
64  | 
(map Syntax_Ext.mk_trfun atrs, map Syntax_Ext.mk_trfun trs,  | 
| 
 
2074b31650e6
discontinued special treatment of structure Syntax_Ext (formerly Syn_Ext);
 
wenzelm 
parents: 
42284 
diff
changeset
 | 
65  | 
map Syntax_Ext.mk_trfun trs', map Syntax_Ext.mk_trfun atrs')  | 
| 
35412
 
b8dead547d9e
more uniform treatment of syntax for types vs. consts;
 
wenzelm 
parents: 
33387 
diff
changeset
 | 
66  | 
|> fold update_gram (AList.coalesce (op =) (rev (map snd mixfixes)));  | 
| 19541 | 67  | 
in make_syntax (thy_syntax, local_syntax, mode, mixfixes, idents) end;  | 
68  | 
||
| 24960 | 69  | 
fun init thy = build_syntax thy Syntax.mode_default [] ([], []);  | 
| 18997 | 70  | 
|
| 19541 | 71  | 
fun rebuild thy (syntax as Syntax {mode, mixfixes, idents, ...}) =
 | 
72  | 
if is_consistent thy syntax then syntax  | 
|
73  | 
else build_syntax thy mode mixfixes idents;  | 
|
| 18997 | 74  | 
|
| 19541 | 75  | 
|
| 18997 | 76  | 
(* mixfix declarations *)  | 
77  | 
||
| 
35412
 
b8dead547d9e
more uniform treatment of syntax for types vs. consts;
 
wenzelm 
parents: 
33387 
diff
changeset
 | 
78  | 
datatype kind = Type | Const | Fixed;  | 
| 
 
b8dead547d9e
more uniform treatment of syntax for types vs. consts;
 
wenzelm 
parents: 
33387 
diff
changeset
 | 
79  | 
|
| 18997 | 80  | 
local  | 
81  | 
||
| 
35412
 
b8dead547d9e
more uniform treatment of syntax for types vs. consts;
 
wenzelm 
parents: 
33387 
diff
changeset
 | 
82  | 
fun prep_mixfix _ _ (_, (_, _, Structure)) = NONE  | 
| 
 
b8dead547d9e
more uniform treatment of syntax for types vs. consts;
 
wenzelm 
parents: 
33387 
diff
changeset
 | 
83  | 
| 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
 | 
84  | 
| 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
 | 
85  | 
| prep_mixfix add mode (Fixed, (x, T, mx)) =  | 
| 
42290
 
b1f544c84040
discontinued special treatment of structure Lexicon;
 
wenzelm 
parents: 
42288 
diff
changeset
 | 
86  | 
SOME ((x, true), ((false, add, mode), (Lexicon.mark_fixed x, T, mx)));  | 
| 19016 | 87  | 
|
| 
35412
 
b8dead547d9e
more uniform treatment of syntax for types vs. consts;
 
wenzelm 
parents: 
33387 
diff
changeset
 | 
88  | 
fun prep_struct (Fixed, (c, _, Structure)) = SOME c  | 
| 
 
b8dead547d9e
more uniform treatment of syntax for types vs. consts;
 
wenzelm 
parents: 
33387 
diff
changeset
 | 
89  | 
  | prep_struct (_, (c, _, Structure)) = error ("Bad mixfix declaration for " ^ quote c)
 | 
| 19016 | 90  | 
| prep_struct _ = NONE;  | 
| 18997 | 91  | 
|
92  | 
in  | 
|
93  | 
||
| 
24951
 
834b8c2b9553
replaced add_modesyntax by general update_modesyntax (add or del);
 
wenzelm 
parents: 
20785 
diff
changeset
 | 
94  | 
fun update_syntax add thy raw_decls  | 
| 
 
834b8c2b9553
replaced add_modesyntax by general update_modesyntax (add or del);
 
wenzelm 
parents: 
20785 
diff
changeset
 | 
95  | 
    (syntax as (Syntax {mode, mixfixes, idents = (structs, _), ...})) =
 | 
| 
 
834b8c2b9553
replaced add_modesyntax by general update_modesyntax (add or del);
 
wenzelm 
parents: 
20785 
diff
changeset
 | 
96  | 
(case filter_out (fn (_, (_, _, mx)) => mx = NoSyn) raw_decls of  | 
| 19016 | 97  | 
[] => syntax  | 
98  | 
| decls =>  | 
|
99  | 
let  | 
|
| 
35412
 
b8dead547d9e
more uniform treatment of syntax for types vs. consts;
 
wenzelm 
parents: 
33387 
diff
changeset
 | 
100  | 
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
 | 
101  | 
val new_structs = map_filter prep_struct decls;  | 
| 
 
0664f10a4094
update_modesyntax: may delete 'structure' notation as well;
 
wenzelm 
parents: 
24951 
diff
changeset
 | 
102  | 
val mixfixes' = rev new_mixfixes @ mixfixes;  | 
| 
 
0664f10a4094
update_modesyntax: may delete 'structure' notation as well;
 
wenzelm 
parents: 
24951 
diff
changeset
 | 
103  | 
val structs' =  | 
| 
 
0664f10a4094
update_modesyntax: may delete 'structure' notation as well;
 
wenzelm 
parents: 
24951 
diff
changeset
 | 
104  | 
if add then structs @ new_structs  | 
| 
 
0664f10a4094
update_modesyntax: may delete 'structure' notation as well;
 
wenzelm 
parents: 
24951 
diff
changeset
 | 
105  | 
else subtract (op =) new_structs structs;  | 
| 20785 | 106  | 
val fixes' = fold (fn ((x, true), _) => cons x | _ => I) mixfixes' [];  | 
| 19541 | 107  | 
in build_syntax thy mode mixfixes' (structs', fixes') end);  | 
| 18997 | 108  | 
|
| 
24951
 
834b8c2b9553
replaced add_modesyntax by general update_modesyntax (add or del);
 
wenzelm 
parents: 
20785 
diff
changeset
 | 
109  | 
val add_syntax = update_syntax true;  | 
| 
 
834b8c2b9553
replaced add_modesyntax by general update_modesyntax (add or del);
 
wenzelm 
parents: 
20785 
diff
changeset
 | 
110  | 
|
| 18997 | 111  | 
end;  | 
112  | 
||
113  | 
||
| 19660 | 114  | 
(* syntax mode *)  | 
115  | 
||
116  | 
fun set_mode mode = map_syntax (fn (thy_syntax, local_syntax, _, mixfixes, idents) =>  | 
|
117  | 
(thy_syntax, local_syntax, mode, mixfixes, idents));  | 
|
118  | 
||
119  | 
fun restore_mode (Syntax {mode, ...}) = set_mode mode;
 | 
|
120  | 
||
| 
24951
 
834b8c2b9553
replaced add_modesyntax by general update_modesyntax (add or del);
 
wenzelm 
parents: 
20785 
diff
changeset
 | 
121  | 
fun update_modesyntax thy add mode args syntax =  | 
| 
 
834b8c2b9553
replaced add_modesyntax by general update_modesyntax (add or del);
 
wenzelm 
parents: 
20785 
diff
changeset
 | 
122  | 
syntax |> set_mode mode |> update_syntax add thy args |> restore_mode syntax;  | 
| 19660 | 123  | 
|
| 18997 | 124  | 
end;  |