| author | wenzelm | 
| Sun, 22 Jul 2007 13:53:46 +0200 | |
| changeset 23896 | 26f92c405337 | 
| parent 20785 | d60f81c56fd4 | 
| child 24951 | 834b8c2b9553 | 
| permissions | -rw-r--r-- | 
| 18997 | 1  | 
(* Title: Pure/Isar/local_syntax.ML  | 
2  | 
ID: $Id$  | 
|
3  | 
Author: Makarius  | 
|
4  | 
||
5  | 
Local syntax depending on theory syntax.  | 
|
6  | 
*)  | 
|
7  | 
||
8  | 
val show_structs = ref false;  | 
|
9  | 
||
10  | 
signature LOCAL_SYNTAX =  | 
|
11  | 
sig  | 
|
12  | 
type T  | 
|
13  | 
val syn_of: T -> Syntax.syntax  | 
|
14  | 
val structs_of: T -> string list  | 
|
15  | 
val init: theory -> T  | 
|
16  | 
val rebuild: theory -> T -> T  | 
|
| 19660 | 17  | 
val add_syntax: theory -> (bool * (string * typ * mixfix)) list -> T -> T  | 
| 20785 | 18  | 
val set_mode: Syntax.mode -> T -> T  | 
| 19541 | 19  | 
val restore_mode: T -> T -> T  | 
| 20785 | 20  | 
val add_modesyntax: theory -> Syntax.mode -> (bool * (string * typ * mixfix)) list -> T -> T  | 
| 
19369
 
a4374b41c9bf
simplified handling of authentic syntax (cf. early externing in consts.ML);
 
wenzelm 
parents: 
19016 
diff
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 =  | 
| 20785 | 30  | 
(string * bool) * (*name, fixed?*)  | 
31  | 
(Syntax.mode * (string * typ * mixfix)); (*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: 
19016 
diff
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: 
19016 
diff
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;  | 
|
61  | 
val (atrs, trs, trs', atrs') = Syntax.struct_trfuns structs;  | 
|
62  | 
val local_syntax = thy_syntax  | 
|
63  | 
|> Syntax.extend_trfuns  | 
|
| 19386 | 64  | 
(map Syntax.mk_trfun atrs, map Syntax.mk_trfun trs,  | 
65  | 
map Syntax.mk_trfun trs', map Syntax.mk_trfun atrs')  | 
|
66  | 
|> fold (uncurry (Syntax.extend_const_gram is_logtype))  | 
|
| 
19454
 
46a7e133f802
moved coalesce to AList, added equality predicates to library
 
haftmann 
parents: 
19386 
diff
changeset
 | 
67  | 
(AList.coalesce (op =) (rev (map snd mixfixes)));  | 
| 19541 | 68  | 
in make_syntax (thy_syntax, local_syntax, mode, mixfixes, idents) end;  | 
69  | 
||
70  | 
fun init thy = build_syntax thy Syntax.default_mode [] ([], []);  | 
|
| 18997 | 71  | 
|
| 19541 | 72  | 
fun rebuild thy (syntax as Syntax {mode, mixfixes, idents, ...}) =
 | 
73  | 
if is_consistent thy syntax then syntax  | 
|
74  | 
else build_syntax thy mode mixfixes idents;  | 
|
| 18997 | 75  | 
|
| 19541 | 76  | 
|
| 18997 | 77  | 
(* mixfix declarations *)  | 
78  | 
||
79  | 
local  | 
|
80  | 
||
| 19016 | 81  | 
fun mixfix_nosyn (_, (_, _, mx)) = mx = NoSyn;  | 
82  | 
fun mixfix_struct (_, (_, _, mx)) = mx = Structure;  | 
|
83  | 
||
| 20785 | 84  | 
fun add_mixfix mode (fixed, (x, T, mx)) =  | 
85  | 
let val c = if fixed then Syntax.fixedN ^ x else x  | 
|
86  | 
in cons ((x, fixed), (mode, (c, T, mx))) end;  | 
|
| 19016 | 87  | 
|
88  | 
fun prep_struct (fixed, (c, _, Structure)) =  | 
|
89  | 
if fixed then SOME c  | 
|
90  | 
      else error ("Bad mixfix declaration for const: " ^ quote c)
 | 
|
91  | 
| prep_struct _ = NONE;  | 
|
| 18997 | 92  | 
|
93  | 
in  | 
|
94  | 
||
| 19541 | 95  | 
fun add_syntax thy raw_decls (syntax as (Syntax {mode, mixfixes, idents = (structs, _), ...})) =
 | 
| 19016 | 96  | 
(case filter_out mixfix_nosyn raw_decls of  | 
97  | 
[] => syntax  | 
|
98  | 
| decls =>  | 
|
99  | 
let  | 
|
| 19541 | 100  | 
val mixfixes' = mixfixes |> fold (add_mixfix mode) (filter_out mixfix_struct decls);  | 
| 20785 | 101  | 
val fixes' = fold (fn ((x, true), _) => cons x | _ => I) mixfixes' [];  | 
| 
19482
 
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
 
wenzelm 
parents: 
19454 
diff
changeset
 | 
102  | 
val structs' = structs @ map_filter prep_struct decls;  | 
| 19541 | 103  | 
in build_syntax thy mode mixfixes' (structs', fixes') end);  | 
| 18997 | 104  | 
|
105  | 
end;  | 
|
106  | 
||
107  | 
||
| 19660 | 108  | 
(* syntax mode *)  | 
109  | 
||
110  | 
fun set_mode mode = map_syntax (fn (thy_syntax, local_syntax, _, mixfixes, idents) =>  | 
|
111  | 
(thy_syntax, local_syntax, mode, mixfixes, idents));  | 
|
112  | 
||
113  | 
fun restore_mode (Syntax {mode, ...}) = set_mode mode;
 | 
|
114  | 
||
115  | 
fun add_modesyntax thy mode args syntax =  | 
|
116  | 
syntax |> set_mode mode |> add_syntax thy args |> restore_mode syntax;  | 
|
117  | 
||
118  | 
||
| 18997 | 119  | 
(* extern_term *)  | 
120  | 
||
| 
19369
 
a4374b41c9bf
simplified handling of authentic syntax (cf. early externing in consts.ML);
 
wenzelm 
parents: 
19016 
diff
changeset
 | 
121  | 
fun extern_term syntax =  | 
| 18997 | 122  | 
let  | 
| 
19369
 
a4374b41c9bf
simplified handling of authentic syntax (cf. early externing in consts.ML);
 
wenzelm 
parents: 
19016 
diff
changeset
 | 
123  | 
val (structs, fixes) = idents_of syntax;  | 
| 18997 | 124  | 
fun map_free (t as Free (x, T)) =  | 
| 19618 | 125  | 
let val i = find_index (fn s => s = x) structs + 1 in  | 
| 18997 | 126  | 
if i = 0 andalso member (op =) fixes x then  | 
127  | 
Const (Syntax.fixedN ^ x, T)  | 
|
128  | 
else if i = 1 andalso not (! show_structs) then  | 
|
129  | 
Syntax.const "_struct" $ Syntax.const "_indexdefault"  | 
|
130  | 
else t  | 
|
131  | 
end  | 
|
132  | 
| map_free t = t;  | 
|
| 
19369
 
a4374b41c9bf
simplified handling of authentic syntax (cf. early externing in consts.ML);
 
wenzelm 
parents: 
19016 
diff
changeset
 | 
133  | 
in Term.map_aterms map_free end;  | 
| 18997 | 134  | 
|
135  | 
end;  |