| author | haftmann | 
| Fri, 11 Sep 2009 09:04:51 +0200 | |
| changeset 32560 | c83dab2c5988 | 
| parent 29606 | fedb8be05f24 | 
| child 32738 | 15bb09ca0378 | 
| permissions | -rw-r--r-- | 
| 18997 | 1 | (* Title: Pure/Isar/local_syntax.ML | 
| 2 | Author: Makarius | |
| 3 | ||
| 4 | Local syntax depending on theory syntax. | |
| 5 | *) | |
| 6 | ||
| 7 | val show_structs = ref false; | |
| 8 | ||
| 9 | signature LOCAL_SYNTAX = | |
| 10 | sig | |
| 11 | type T | |
| 12 | val syn_of: T -> Syntax.syntax | |
| 13 | val structs_of: T -> string list | |
| 14 | val init: theory -> T | |
| 15 | val rebuild: theory -> T -> T | |
| 19660 | 16 | val add_syntax: theory -> (bool * (string * typ * mixfix)) list -> T -> T | 
| 20785 | 17 | val set_mode: Syntax.mode -> T -> T | 
| 19541 | 18 | val restore_mode: T -> T -> T | 
| 24951 
834b8c2b9553
replaced add_modesyntax by general update_modesyntax (add or del);
 wenzelm parents: 
20785diff
changeset | 19 | val update_modesyntax: theory -> bool -> Syntax.mode -> | 
| 
834b8c2b9553
replaced add_modesyntax by general update_modesyntax (add or del);
 wenzelm parents: 
20785diff
changeset | 20 | (bool * (string * typ * mixfix)) list -> T -> T | 
| 19369 
a4374b41c9bf
simplified handling of authentic syntax (cf. early externing in consts.ML);
 wenzelm parents: 
19016diff
changeset | 21 | val extern_term: T -> term -> term | 
| 18997 | 22 | end; | 
| 23 | ||
| 19016 | 24 | structure LocalSyntax: LOCAL_SYNTAX = | 
| 18997 | 25 | struct | 
| 26 | ||
| 27 | (* datatype T *) | |
| 28 | ||
| 19386 | 29 | type local_mixfix = | 
| 24951 
834b8c2b9553
replaced add_modesyntax by general update_modesyntax (add or del);
 wenzelm parents: 
20785diff
changeset | 30 | (string * bool) * (*name, fixed?*) | 
| 
834b8c2b9553
replaced add_modesyntax by general update_modesyntax (add or del);
 wenzelm parents: 
20785diff
changeset | 31 | ((bool * Syntax.mode) * (string * typ * mixfix)); (*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, | 
| 19386 | 37 | mixfixes: local_mixfix list, | 
| 19369 
a4374b41c9bf
simplified handling of authentic syntax (cf. early externing in consts.ML);
 wenzelm parents: 
19016diff
changeset | 38 | idents: string list * string list}; | 
| 18997 | 39 | |
| 19541 | 40 | fun make_syntax (thy_syntax, local_syntax, mode, mixfixes, idents) = | 
| 18997 | 41 |   Syntax {thy_syntax = thy_syntax, local_syntax = local_syntax,
 | 
| 19541 | 42 | mode = mode, mixfixes = mixfixes, idents = idents}; | 
| 18997 | 43 | |
| 19541 | 44 | fun map_syntax f (Syntax {thy_syntax, local_syntax, mode, mixfixes, idents}) =
 | 
| 45 | make_syntax (f (thy_syntax, local_syntax, mode, mixfixes, idents)); | |
| 18997 | 46 | |
| 19369 
a4374b41c9bf
simplified handling of authentic syntax (cf. early externing in consts.ML);
 wenzelm parents: 
19016diff
changeset | 47 | fun is_consistent thy (Syntax {thy_syntax, ...}) =
 | 
| 18997 | 48 | Syntax.eq_syntax (Sign.syn_of thy, thy_syntax); | 
| 49 | ||
| 50 | fun syn_of (Syntax {local_syntax, ...}) = local_syntax;
 | |
| 51 | fun idents_of (Syntax {idents, ...}) = idents;
 | |
| 52 | val structs_of = #1 o idents_of; | |
| 53 | ||
| 54 | ||
| 55 | (* build syntax *) | |
| 56 | ||
| 19541 | 57 | fun build_syntax thy mode mixfixes (idents as (structs, _)) = | 
| 18997 | 58 | let | 
| 59 | val thy_syntax = Sign.syn_of thy; | |
| 60 | val is_logtype = Sign.is_logtype thy; | |
| 25385 
42df506673ed
update_modesyntax: based on Syntax.update_const_gram (avoids duplicates);
 wenzelm parents: 
24960diff
changeset | 61 | fun const_gram ((true, m), decls) = Syntax.update_const_gram is_logtype m decls | 
| 24951 
834b8c2b9553
replaced add_modesyntax by general update_modesyntax (add or del);
 wenzelm parents: 
20785diff
changeset | 62 | | const_gram ((false, m), decls) = Syntax.remove_const_gram is_logtype m decls; | 
| 18997 | 63 | val (atrs, trs, trs', atrs') = Syntax.struct_trfuns structs; | 
| 64 | val local_syntax = thy_syntax | |
| 25390 
8bfa6566ac6b
syntax operations: turned extend'' into update'' (absorb duplicates);
 wenzelm parents: 
25385diff
changeset | 65 | |> Syntax.update_trfuns | 
| 19386 | 66 | (map Syntax.mk_trfun atrs, map Syntax.mk_trfun trs, | 
| 67 | map Syntax.mk_trfun trs', map Syntax.mk_trfun atrs') | |
| 24951 
834b8c2b9553
replaced add_modesyntax by general update_modesyntax (add or del);
 wenzelm parents: 
20785diff
changeset | 68 | |> fold const_gram (AList.coalesce (op =) (rev (map snd mixfixes))); | 
| 19541 | 69 | in make_syntax (thy_syntax, local_syntax, mode, mixfixes, idents) end; | 
| 70 | ||
| 24960 | 71 | fun init thy = build_syntax thy Syntax.mode_default [] ([], []); | 
| 18997 | 72 | |
| 19541 | 73 | fun rebuild thy (syntax as Syntax {mode, mixfixes, idents, ...}) =
 | 
| 74 | if is_consistent thy syntax then syntax | |
| 75 | else build_syntax thy mode mixfixes idents; | |
| 18997 | 76 | |
| 19541 | 77 | |
| 18997 | 78 | (* mixfix declarations *) | 
| 79 | ||
| 80 | local | |
| 81 | ||
| 24951 
834b8c2b9553
replaced add_modesyntax by general update_modesyntax (add or del);
 wenzelm parents: 
20785diff
changeset | 82 | fun prep_mixfix _ (_, (_, _, Structure)) = NONE | 
| 
834b8c2b9553
replaced add_modesyntax by general update_modesyntax (add or del);
 wenzelm parents: 
20785diff
changeset | 83 | | prep_mixfix mode (fixed, (x, T, mx)) = | 
| 
834b8c2b9553
replaced add_modesyntax by general update_modesyntax (add or del);
 wenzelm parents: 
20785diff
changeset | 84 | let val c = if fixed then Syntax.fixedN ^ x else x | 
| 
834b8c2b9553
replaced add_modesyntax by general update_modesyntax (add or del);
 wenzelm parents: 
20785diff
changeset | 85 | in SOME ((x, fixed), (mode, (c, T, mx))) end; | 
| 19016 | 86 | |
| 87 | fun prep_struct (fixed, (c, _, Structure)) = | |
| 88 | if fixed then SOME c | |
| 24951 
834b8c2b9553
replaced add_modesyntax by general update_modesyntax (add or del);
 wenzelm parents: 
20785diff
changeset | 89 |       else error ("Bad mixfix declaration for constant: " ^ 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: 
20785diff
changeset | 94 | fun update_syntax add thy raw_decls | 
| 
834b8c2b9553
replaced add_modesyntax by general update_modesyntax (add or del);
 wenzelm parents: 
20785diff
changeset | 95 |     (syntax as (Syntax {mode, mixfixes, idents = (structs, _), ...})) =
 | 
| 
834b8c2b9553
replaced add_modesyntax by general update_modesyntax (add or del);
 wenzelm parents: 
20785diff
changeset | 96 | (case filter_out (fn (_, (_, _, mx)) => mx = NoSyn) raw_decls of | 
| 19016 | 97 | [] => syntax | 
| 98 | | decls => | |
| 99 | let | |
| 24954 
0664f10a4094
update_modesyntax: may delete 'structure' notation as well;
 wenzelm parents: 
24951diff
changeset | 100 | val new_mixfixes = map_filter (prep_mixfix (add, mode)) decls; | 
| 
0664f10a4094
update_modesyntax: may delete 'structure' notation as well;
 wenzelm parents: 
24951diff
changeset | 101 | val new_structs = map_filter prep_struct decls; | 
| 
0664f10a4094
update_modesyntax: may delete 'structure' notation as well;
 wenzelm parents: 
24951diff
changeset | 102 | val mixfixes' = rev new_mixfixes @ mixfixes; | 
| 
0664f10a4094
update_modesyntax: may delete 'structure' notation as well;
 wenzelm parents: 
24951diff
changeset | 103 | val structs' = | 
| 
0664f10a4094
update_modesyntax: may delete 'structure' notation as well;
 wenzelm parents: 
24951diff
changeset | 104 | if add then structs @ new_structs | 
| 
0664f10a4094
update_modesyntax: may delete 'structure' notation as well;
 wenzelm parents: 
24951diff
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: 
20785diff
changeset | 109 | val add_syntax = update_syntax true; | 
| 
834b8c2b9553
replaced add_modesyntax by general update_modesyntax (add or del);
 wenzelm parents: 
20785diff
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: 
20785diff
changeset | 121 | fun update_modesyntax thy add mode args syntax = | 
| 
834b8c2b9553
replaced add_modesyntax by general update_modesyntax (add or del);
 wenzelm parents: 
20785diff
changeset | 122 | syntax |> set_mode mode |> update_syntax add thy args |> restore_mode syntax; | 
| 19660 | 123 | |
| 124 | ||
| 18997 | 125 | (* extern_term *) | 
| 126 | ||
| 19369 
a4374b41c9bf
simplified handling of authentic syntax (cf. early externing in consts.ML);
 wenzelm parents: 
19016diff
changeset | 127 | fun extern_term syntax = | 
| 18997 | 128 | let | 
| 19369 
a4374b41c9bf
simplified handling of authentic syntax (cf. early externing in consts.ML);
 wenzelm parents: 
19016diff
changeset | 129 | val (structs, fixes) = idents_of syntax; | 
| 18997 | 130 | fun map_free (t as Free (x, T)) = | 
| 19618 | 131 | let val i = find_index (fn s => s = x) structs + 1 in | 
| 18997 | 132 | if i = 0 andalso member (op =) fixes x then | 
| 133 | Const (Syntax.fixedN ^ x, T) | |
| 134 | else if i = 1 andalso not (! show_structs) then | |
| 135 | Syntax.const "_struct" $ Syntax.const "_indexdefault" | |
| 136 | else t | |
| 137 | end | |
| 138 | | map_free t = t; | |
| 19369 
a4374b41c9bf
simplified handling of authentic syntax (cf. early externing in consts.ML);
 wenzelm parents: 
19016diff
changeset | 139 | in Term.map_aterms map_free end; | 
| 18997 | 140 | |
| 141 | end; |