| author | nipkow | 
| Fri, 23 Nov 2012 23:07:38 +0100 | |
| changeset 50180 | c6626861c31a | 
| parent 48743 | a72f8ffecf31 | 
| child 55104 | 8284c0d5bf52 | 
| permissions | -rw-r--r-- | 
| 11523 | 1  | 
(* Title: Pure/General/scan.ML  | 
2  | 
Author: Markus Wenzel and Tobias Nipkow, TU Muenchen  | 
|
| 
6116
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
3  | 
|
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
4  | 
Generic scanners (for potentially infinite input).  | 
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
5  | 
*)  | 
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
6  | 
|
| 24025 | 7  | 
infix 5 -- :-- :|-- |-- --| ^^;  | 
| 25999 | 8  | 
infixr 5 ::: @@@;  | 
| 
6116
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
9  | 
infix 3 >>;  | 
| 23699 | 10  | 
infixr 0 ||;  | 
| 
6116
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
11  | 
|
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
12  | 
signature BASIC_SCAN =  | 
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
13  | 
sig  | 
| 
43947
 
9b00f09f7721
defer evaluation of Scan.message, for improved performance in the frequent situation where failure is handled later (e.g. via ||);
 
wenzelm 
parents: 
40627 
diff
changeset
 | 
14  | 
type message = unit -> string  | 
| 14677 | 15  | 
(*error msg handler*)  | 
| 
43947
 
9b00f09f7721
defer evaluation of Scan.message, for improved performance in the frequent situation where failure is handled later (e.g. via ||);
 
wenzelm 
parents: 
40627 
diff
changeset
 | 
16  | 
  val !! : ('a * message option -> message) -> ('a -> 'b) -> 'a -> 'b
 | 
| 14677 | 17  | 
(*apply function*)  | 
| 
6116
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
18  | 
  val >> : ('a -> 'b * 'c) * ('b -> 'd) -> 'a -> 'd * 'c
 | 
| 14677 | 19  | 
(*alternative*)  | 
| 
6116
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
20  | 
  val || : ('a -> 'b) * ('a -> 'b) -> 'a -> 'b
 | 
| 14677 | 21  | 
(*sequential pairing*)  | 
| 
6116
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
22  | 
  val -- : ('a -> 'b * 'c) * ('c -> 'd * 'e) -> 'a -> ('b * 'd) * 'e
 | 
| 14677 | 23  | 
(*dependent pairing*)  | 
| 
6116
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
24  | 
  val :-- : ('a -> 'b * 'c) * ('b -> 'c -> 'd * 'e) -> 'a -> ('b * 'd) * 'e
 | 
| 24025 | 25  | 
(*projections*)  | 
26  | 
  val :|-- : ('a -> 'b * 'c) * ('b -> 'c -> 'd * 'e) -> 'a -> 'd * 'e
 | 
|
| 
6116
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
27  | 
  val |-- : ('a -> 'b * 'c) * ('c -> 'd * 'e) -> 'a -> 'd * 'e
 | 
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
28  | 
  val --| : ('a -> 'b * 'c) * ('c -> 'd * 'e) -> 'a -> 'b * 'e
 | 
| 14677 | 29  | 
(*concatenation*)  | 
| 
6116
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
30  | 
  val ^^ : ('a -> string * 'b) * ('b -> string * 'c) -> 'a -> string * 'c
 | 
| 25999 | 31  | 
  val ::: : ('a -> 'b * 'c) * ('c -> 'b list * 'd) -> 'a -> 'b list * 'd
 | 
32  | 
  val @@@ : ('a -> 'b list * 'c) * ('c -> 'b list * 'd) -> 'a -> 'b list * 'd
 | 
|
| 14677 | 33  | 
(*one element literal*)  | 
| 19291 | 34  | 
val $$ : string -> string list -> string * string list  | 
| 19306 | 35  | 
val ~$$ : string -> string list -> string * string list  | 
| 
6116
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
36  | 
end;  | 
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
37  | 
|
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
38  | 
signature SCAN =  | 
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
39  | 
sig  | 
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
40  | 
include BASIC_SCAN  | 
| 23699 | 41  | 
  val prompt: string -> ('a -> 'b) -> 'a -> 'b
 | 
| 
48743
 
a72f8ffecf31
refined recovery of scan errors: longest prefix of delimited token after failure, otherwise just one symbol;
 
wenzelm 
parents: 
43947 
diff
changeset
 | 
42  | 
  val permissive: ('a -> 'b) -> 'a -> 'b
 | 
| 23699 | 43  | 
  val error: ('a -> 'b) -> 'a -> 'b
 | 
44  | 
  val catch: ('a -> 'b) -> 'a -> 'b    (*exception Fail*)
 | 
|
| 
6116
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
45  | 
val fail: 'a -> 'b  | 
| 
43947
 
9b00f09f7721
defer evaluation of Scan.message, for improved performance in the frequent situation where failure is handled later (e.g. via ||);
 
wenzelm 
parents: 
40627 
diff
changeset
 | 
46  | 
  val fail_with: ('a -> message) -> 'a -> 'b
 | 
| 
6116
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
47  | 
val succeed: 'a -> 'b -> 'a * 'b  | 
| 15664 | 48  | 
  val some: ('a -> 'b option) -> 'a list -> 'b * 'a list
 | 
49  | 
  val one: ('a -> bool) -> 'a list -> 'a * 'a list
 | 
|
| 19291 | 50  | 
val this: string list -> string list -> string list * string list  | 
| 14927 | 51  | 
val this_string: string -> string list -> string * string list  | 
| 
21858
 
05f57309170c
avoid conflict with Alice keywords: renamed pack -> implode, unpack -> explode, any -> many, avoided assert;
 
wenzelm 
parents: 
19473 
diff
changeset
 | 
52  | 
  val many: ('a -> bool) -> 'a list -> 'a list * 'a list
 | 
| 
 
05f57309170c
avoid conflict with Alice keywords: renamed pack -> implode, unpack -> explode, any -> many, avoided assert;
 
wenzelm 
parents: 
19473 
diff
changeset
 | 
53  | 
  val many1: ('a -> bool) -> 'a list -> 'a list * 'a list
 | 
| 
6116
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
54  | 
  val optional: ('a -> 'b * 'a) -> 'b -> 'a -> 'b * 'a
 | 
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
55  | 
  val option: ('a -> 'b * 'a) -> 'a -> 'b option * 'a
 | 
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
56  | 
  val repeat: ('a -> 'b * 'a) -> 'a -> 'b list * 'a
 | 
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
57  | 
  val repeat1: ('a -> 'b * 'a) -> 'a -> 'b list * 'a
 | 
| 23699 | 58  | 
  val single: ('a -> 'b * 'a) -> 'a -> 'b list * 'a
 | 
59  | 
  val bulk: ('a -> 'b * 'a) -> 'a -> 'b list * 'a
 | 
|
| 
6116
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
60  | 
  val max: ('a * 'a -> bool) -> ('b -> 'a * 'b) -> ('b -> 'a * 'b) -> 'b -> 'a * 'b
 | 
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
61  | 
  val ahead: ('a -> 'b * 'c) -> 'a -> 'b * 'a
 | 
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
62  | 
  val unless: ('a -> 'b * 'a) -> ('a -> 'c * 'd) -> 'a -> 'c * 'd
 | 
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
63  | 
  val first: ('a -> 'b) list -> 'a -> 'b
 | 
| 14677 | 64  | 
  val state: 'a * 'b -> 'a * ('a * 'b)
 | 
| 
6116
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
65  | 
  val depend: ('a -> 'b -> ('c * 'd) * 'e) -> 'a * 'b -> 'd * ('c * 'e)
 | 
| 15664 | 66  | 
  val peek: ('a -> 'b -> 'c * 'd) -> 'a * 'b -> 'c * ('a * 'd)
 | 
67  | 
  val pass: 'a -> ('a * 'b -> 'c * ('d * 'e)) -> 'b -> 'c * 'e
 | 
|
| 
6116
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
68  | 
  val lift: ('a -> 'b * 'c) -> 'd * 'a -> 'b * ('d * 'c)
 | 
| 23699 | 69  | 
  val unlift: (unit * 'a -> 'b * ('c * 'd)) -> 'a -> 'b * 'd
 | 
| 15664 | 70  | 
  val trace: ('a list -> 'b * 'c list) -> 'a list -> ('b * 'a list) * 'c list
 | 
| 
27731
 
a7444ded92cf
abstract type stopper, may depend on final input;
 
wenzelm 
parents: 
25999 
diff
changeset
 | 
71  | 
type 'a stopper  | 
| 
 
a7444ded92cf
abstract type stopper, may depend on final input;
 
wenzelm 
parents: 
25999 
diff
changeset
 | 
72  | 
  val stopper: ('a list -> 'a) -> ('a -> bool) -> 'a stopper
 | 
| 
 
a7444ded92cf
abstract type stopper, may depend on final input;
 
wenzelm 
parents: 
25999 
diff
changeset
 | 
73  | 
val is_stopper: 'a stopper -> 'a -> bool  | 
| 
 
a7444ded92cf
abstract type stopper, may depend on final input;
 
wenzelm 
parents: 
25999 
diff
changeset
 | 
74  | 
  val finite': 'a stopper -> ('b * 'a list -> 'c * ('d * 'a list))
 | 
| 
6116
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
75  | 
    -> 'b * 'a list -> 'c * ('d * 'a list)
 | 
| 
27731
 
a7444ded92cf
abstract type stopper, may depend on final input;
 
wenzelm 
parents: 
25999 
diff
changeset
 | 
76  | 
  val finite: 'a stopper -> ('a list -> 'b * 'a list) -> 'a list -> 'b * 'a list
 | 
| 
 
a7444ded92cf
abstract type stopper, may depend on final input;
 
wenzelm 
parents: 
25999 
diff
changeset
 | 
77  | 
  val read: 'a stopper -> ('a list -> 'b * 'a list) -> 'a list -> 'b option
 | 
| 
 
a7444ded92cf
abstract type stopper, may depend on final input;
 
wenzelm 
parents: 
25999 
diff
changeset
 | 
78  | 
val drain: string -> (string -> 'a -> 'b list * 'a) -> 'b stopper ->  | 
| 23699 | 79  | 
    ('c * 'b list -> 'd * ('e * 'b list)) -> ('c * 'b list) * 'a -> ('d * ('e * 'b list)) * 'a
 | 
| 
6116
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
80  | 
type lexicon  | 
| 
27765
 
5df443dd9deb
reorganized lexicon: allow scanning of annotated symbols, tuned representation and interfaces;
 
wenzelm 
parents: 
27731 
diff
changeset
 | 
81  | 
val is_literal: lexicon -> string list -> bool  | 
| 
 
5df443dd9deb
reorganized lexicon: allow scanning of annotated symbols, tuned representation and interfaces;
 
wenzelm 
parents: 
27731 
diff
changeset
 | 
82  | 
val literal: lexicon -> (string * 'a) list -> (string * 'a) list * (string * 'a) list  | 
| 
6116
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
83  | 
val empty_lexicon: lexicon  | 
| 
27765
 
5df443dd9deb
reorganized lexicon: allow scanning of annotated symbols, tuned representation and interfaces;
 
wenzelm 
parents: 
27731 
diff
changeset
 | 
84  | 
val extend_lexicon: string list -> lexicon -> lexicon  | 
| 
 
5df443dd9deb
reorganized lexicon: allow scanning of annotated symbols, tuned representation and interfaces;
 
wenzelm 
parents: 
27731 
diff
changeset
 | 
85  | 
val make_lexicon: string list list -> lexicon  | 
| 
 
5df443dd9deb
reorganized lexicon: allow scanning of annotated symbols, tuned representation and interfaces;
 
wenzelm 
parents: 
27731 
diff
changeset
 | 
86  | 
val dest_lexicon: lexicon -> string list  | 
| 
 
5df443dd9deb
reorganized lexicon: allow scanning of annotated symbols, tuned representation and interfaces;
 
wenzelm 
parents: 
27731 
diff
changeset
 | 
87  | 
val merge_lexicons: lexicon * lexicon -> lexicon  | 
| 
6116
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
88  | 
end;  | 
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
89  | 
|
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
90  | 
structure Scan: SCAN =  | 
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
91  | 
struct  | 
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
92  | 
|
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
93  | 
|
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
94  | 
(** scanners **)  | 
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
95  | 
|
| 23699 | 96  | 
(* exceptions *)  | 
97  | 
||
| 
43947
 
9b00f09f7721
defer evaluation of Scan.message, for improved performance in the frequent situation where failure is handled later (e.g. via ||);
 
wenzelm 
parents: 
40627 
diff
changeset
 | 
98  | 
type message = unit -> string;  | 
| 
 
9b00f09f7721
defer evaluation of Scan.message, for improved performance in the frequent situation where failure is handled later (e.g. via ||);
 
wenzelm 
parents: 
40627 
diff
changeset
 | 
99  | 
|
| 11523 | 100  | 
exception MORE of string option; (*need more input (prompt)*)  | 
| 
43947
 
9b00f09f7721
defer evaluation of Scan.message, for improved performance in the frequent situation where failure is handled later (e.g. via ||);
 
wenzelm 
parents: 
40627 
diff
changeset
 | 
101  | 
exception FAIL of message option; (*try alternatives (reason of failure)*)  | 
| 
 
9b00f09f7721
defer evaluation of Scan.message, for improved performance in the frequent situation where failure is handled later (e.g. via ||);
 
wenzelm 
parents: 
40627 
diff
changeset
 | 
102  | 
exception ABORT of message; (*dead end*)  | 
| 
6116
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
103  | 
|
| 23699 | 104  | 
fun !! err scan xs = scan xs handle FAIL msg => raise ABORT (err (xs, msg));  | 
105  | 
fun permissive scan xs = scan xs handle MORE _ => raise FAIL NONE | ABORT _ => raise FAIL NONE;  | 
|
106  | 
fun strict scan xs = scan xs handle MORE _ => raise FAIL NONE;  | 
|
107  | 
fun prompt str scan xs = scan xs handle MORE NONE => raise MORE (SOME str);  | 
|
| 
43947
 
9b00f09f7721
defer evaluation of Scan.message, for improved performance in the frequent situation where failure is handled later (e.g. via ||);
 
wenzelm 
parents: 
40627 
diff
changeset
 | 
108  | 
fun error scan xs = scan xs handle ABORT msg => Library.error (msg ());  | 
| 23699 | 109  | 
|
110  | 
fun catch scan xs = scan xs  | 
|
| 
43947
 
9b00f09f7721
defer evaluation of Scan.message, for improved performance in the frequent situation where failure is handled later (e.g. via ||);
 
wenzelm 
parents: 
40627 
diff
changeset
 | 
111  | 
handle ABORT msg => raise Fail (msg ())  | 
| 
 
9b00f09f7721
defer evaluation of Scan.message, for improved performance in the frequent situation where failure is handled later (e.g. via ||);
 
wenzelm 
parents: 
40627 
diff
changeset
 | 
112  | 
| FAIL msg => raise Fail (case msg of NONE => "Syntax error" | SOME m => m ());  | 
| 23699 | 113  | 
|
| 
6116
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
114  | 
|
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
115  | 
(* scanner combinators *)  | 
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
116  | 
|
| 19306 | 117  | 
fun (scan >> f) xs = scan xs |>> f;  | 
| 14078 | 118  | 
|
| 19306 | 119  | 
fun (scan1 || scan2) xs = scan1 xs handle FAIL _ => scan2 xs;  | 
| 14078 | 120  | 
|
| 19306 | 121  | 
fun (scan1 :-- scan2) xs =  | 
| 
14108
 
eaf3c75f2c8e
Restored old (tail recursive!) version of repeat.
 
berghofe 
parents: 
14078 
diff
changeset
 | 
122  | 
let  | 
| 19306 | 123  | 
val (x, ys) = scan1 xs;  | 
124  | 
val (y, zs) = scan2 x ys;  | 
|
125  | 
in ((x, y), zs) end;  | 
|
| 14078 | 126  | 
|
| 19306 | 127  | 
fun (scan1 -- scan2) = scan1 :-- (fn _ => scan2);  | 
| 24025 | 128  | 
fun (scan1 :|-- scan2) = scan1 :-- scan2 >> #2;  | 
| 19306 | 129  | 
fun (scan1 |-- scan2) = scan1 -- scan2 >> #2;  | 
130  | 
fun (scan1 --| scan2) = scan1 -- scan2 >> #1;  | 
|
131  | 
fun (scan1 ^^ scan2) = scan1 -- scan2 >> op ^;  | 
|
| 25999 | 132  | 
fun (scan1 ::: scan2) = scan1 -- scan2 >> op ::;  | 
133  | 
fun (scan1 @@@ scan2) = scan1 -- scan2 >> op @;  | 
|
| 
6116
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
134  | 
|
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
135  | 
|
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
136  | 
(* generic scanners *)  | 
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
137  | 
|
| 15531 | 138  | 
fun fail _ = raise FAIL NONE;  | 
139  | 
fun fail_with msg_of xs = raise FAIL (SOME (msg_of xs));  | 
|
| 
6116
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
140  | 
fun succeed y xs = (y, xs);  | 
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
141  | 
|
| 15664 | 142  | 
fun some _ [] = raise MORE NONE  | 
143  | 
| some f (x :: xs) =  | 
|
144  | 
(case f x of SOME y => (y, xs) | _ => raise FAIL NONE);  | 
|
145  | 
||
| 15531 | 146  | 
fun one _ [] = raise MORE NONE  | 
| 
6116
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
147  | 
| one pred (x :: xs) =  | 
| 15531 | 148  | 
if pred x then (x, xs) else raise FAIL NONE;  | 
| 
6116
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
149  | 
|
| 19306 | 150  | 
fun $$ a = one (fn s: string => s = a);  | 
151  | 
fun ~$$ a = one (fn s: string => s <> a);  | 
|
| 
6116
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
152  | 
|
| 14833 | 153  | 
fun this ys xs =  | 
| 14726 | 154  | 
let  | 
155  | 
fun drop_prefix [] xs = xs  | 
|
| 15531 | 156  | 
| drop_prefix (_ :: _) [] = raise MORE NONE  | 
| 14726 | 157  | 
| drop_prefix (y :: ys) (x :: xs) =  | 
| 19291 | 158  | 
if (y: string) = x then drop_prefix ys xs else raise FAIL NONE;  | 
| 14726 | 159  | 
in (ys, drop_prefix ys xs) end;  | 
160  | 
||
| 
40627
 
becf5d5187cc
renamed raw "explode" function to "raw_explode" to emphasize its meaning;
 
wenzelm 
parents: 
38875 
diff
changeset
 | 
161  | 
fun this_string s = this (raw_explode s) >> K s; (*primitive string -- no symbols here!*)  | 
| 14907 | 162  | 
|
| 
21858
 
05f57309170c
avoid conflict with Alice keywords: renamed pack -> implode, unpack -> explode, any -> many, avoided assert;
 
wenzelm 
parents: 
19473 
diff
changeset
 | 
163  | 
fun many _ [] = raise MORE NONE  | 
| 
 
05f57309170c
avoid conflict with Alice keywords: renamed pack -> implode, unpack -> explode, any -> many, avoided assert;
 
wenzelm 
parents: 
19473 
diff
changeset
 | 
164  | 
| many pred (lst as x :: xs) =  | 
| 
 
05f57309170c
avoid conflict with Alice keywords: renamed pack -> implode, unpack -> explode, any -> many, avoided assert;
 
wenzelm 
parents: 
19473 
diff
changeset
 | 
165  | 
if pred x then apfst (cons x) (many pred xs)  | 
| 
6116
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
166  | 
else ([], lst);  | 
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
167  | 
|
| 25999 | 168  | 
fun many1 pred = one pred ::: many pred;  | 
| 
6116
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
169  | 
|
| 15664 | 170  | 
fun optional scan def = scan || succeed def;  | 
171  | 
fun option scan = (scan >> SOME) || succeed NONE;  | 
|
| 
6116
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
172  | 
|
| 13795 | 173  | 
fun repeat scan =  | 
| 15664 | 174  | 
let  | 
175  | 
fun rep ys xs =  | 
|
176  | 
(case (SOME (scan xs) handle FAIL _ => NONE) of  | 
|
177  | 
NONE => (rev ys, xs)  | 
|
178  | 
| SOME (y, xs') => rep (y :: ys) xs');  | 
|
| 
14108
 
eaf3c75f2c8e
Restored old (tail recursive!) version of repeat.
 
berghofe 
parents: 
14078 
diff
changeset
 | 
179  | 
in rep [] end;  | 
| 13795 | 180  | 
|
| 25999 | 181  | 
fun repeat1 scan = scan ::: repeat scan;  | 
| 
6116
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
182  | 
|
| 23699 | 183  | 
fun single scan = scan >> (fn x => [x]);  | 
184  | 
fun bulk scan = scan -- repeat (permissive scan) >> (op ::);  | 
|
185  | 
||
| 
6116
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
186  | 
fun max leq scan1 scan2 xs =  | 
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
187  | 
(case (option scan1 xs, option scan2 xs) of  | 
| 15531 | 188  | 
((NONE, _), (NONE, _)) => raise FAIL NONE (*looses FAIL msg!*)  | 
189  | 
| ((SOME tok1, xs'), (NONE, _)) => (tok1, xs')  | 
|
190  | 
| ((NONE, _), (SOME tok2, xs')) => (tok2, xs')  | 
|
191  | 
| ((SOME tok1, xs1'), (SOME tok2, xs2')) =>  | 
|
| 
6116
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
192  | 
if leq (tok2, tok1) then (tok1, xs1') else (tok2, xs2'));  | 
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
193  | 
|
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
194  | 
fun ahead scan xs = (fst (scan xs), xs);  | 
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
195  | 
|
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
196  | 
fun unless test scan =  | 
| 15531 | 197  | 
ahead (option test) :-- (fn NONE => scan | _ => fail) >> #2;  | 
| 
6116
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
198  | 
|
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
199  | 
fun first [] = fail  | 
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
200  | 
| first (scan :: scans) = scan || first scans;  | 
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
201  | 
|
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
202  | 
|
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
203  | 
(* state based scanners *)  | 
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
204  | 
|
| 9122 | 205  | 
fun state (st, xs) = (st, (st, xs));  | 
206  | 
||
| 
6116
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
207  | 
fun depend scan (st, xs) =  | 
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
208  | 
let val ((st', y), xs') = scan st xs  | 
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
209  | 
in (y, (st', xs')) end;  | 
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
210  | 
|
| 15664 | 211  | 
fun peek scan = depend (fn st => scan st >> pair st);  | 
212  | 
||
213  | 
fun pass st scan xs =  | 
|
214  | 
let val (y, (_, xs')) = scan (st, xs)  | 
|
215  | 
in (y, xs') end;  | 
|
216  | 
||
| 
6116
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
217  | 
fun lift scan (st, xs) =  | 
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
218  | 
let val (y, xs') = scan xs  | 
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
219  | 
in (y, (st, xs')) end;  | 
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
220  | 
|
| 15664 | 221  | 
fun unlift scan = pass () scan;  | 
222  | 
||
223  | 
||
224  | 
(* trace input *)  | 
|
225  | 
||
| 23699 | 226  | 
fun trace scan xs =  | 
227  | 
let val (y, xs') = scan xs  | 
|
| 33957 | 228  | 
in ((y, take (length xs - length xs') xs), xs') end;  | 
| 
6116
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
229  | 
|
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
230  | 
|
| 
27731
 
a7444ded92cf
abstract type stopper, may depend on final input;
 
wenzelm 
parents: 
25999 
diff
changeset
 | 
231  | 
(* stopper *)  | 
| 
 
a7444ded92cf
abstract type stopper, may depend on final input;
 
wenzelm 
parents: 
25999 
diff
changeset
 | 
232  | 
|
| 
 
a7444ded92cf
abstract type stopper, may depend on final input;
 
wenzelm 
parents: 
25999 
diff
changeset
 | 
233  | 
datatype 'a stopper = Stopper of ('a list -> 'a) * ('a -> bool);
 | 
| 
 
a7444ded92cf
abstract type stopper, may depend on final input;
 
wenzelm 
parents: 
25999 
diff
changeset
 | 
234  | 
|
| 
 
a7444ded92cf
abstract type stopper, may depend on final input;
 
wenzelm 
parents: 
25999 
diff
changeset
 | 
235  | 
fun stopper mk_stopper is_stopper = Stopper (mk_stopper, is_stopper);  | 
| 
 
a7444ded92cf
abstract type stopper, may depend on final input;
 
wenzelm 
parents: 
25999 
diff
changeset
 | 
236  | 
fun is_stopper (Stopper (_, is_stopper)) = is_stopper;  | 
| 
 
a7444ded92cf
abstract type stopper, may depend on final input;
 
wenzelm 
parents: 
25999 
diff
changeset
 | 
237  | 
|
| 
 
a7444ded92cf
abstract type stopper, may depend on final input;
 
wenzelm 
parents: 
25999 
diff
changeset
 | 
238  | 
|
| 
6116
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
239  | 
(* finite scans *)  | 
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
240  | 
|
| 
27731
 
a7444ded92cf
abstract type stopper, may depend on final input;
 
wenzelm 
parents: 
25999 
diff
changeset
 | 
241  | 
fun finite' (Stopper (mk_stopper, is_stopper)) scan (state, input) =  | 
| 
6116
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
242  | 
let  | 
| 
43947
 
9b00f09f7721
defer evaluation of Scan.message, for improved performance in the frequent situation where failure is handled later (e.g. via ||);
 
wenzelm 
parents: 
40627 
diff
changeset
 | 
243  | 
fun lost () = raise ABORT (fn () => "Bad scanner: lost stopper of finite scan!");  | 
| 
6116
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
244  | 
|
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
245  | 
fun stop [] = lost ()  | 
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
246  | 
| stop lst =  | 
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
247  | 
let val (xs, x) = split_last lst  | 
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
248  | 
in if is_stopper x then ((), xs) else lost () end;  | 
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
249  | 
in  | 
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
250  | 
if exists is_stopper input then  | 
| 
43947
 
9b00f09f7721
defer evaluation of Scan.message, for improved performance in the frequent situation where failure is handled later (e.g. via ||);
 
wenzelm 
parents: 
40627 
diff
changeset
 | 
251  | 
raise ABORT (fn () => "Stopper may not occur in input of finite scan!")  | 
| 
27731
 
a7444ded92cf
abstract type stopper, may depend on final input;
 
wenzelm 
parents: 
25999 
diff
changeset
 | 
252  | 
else (strict scan --| lift stop) (state, input @ [mk_stopper input])  | 
| 
6116
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
253  | 
end;  | 
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
254  | 
|
| 15664 | 255  | 
fun finite stopper scan = unlift (finite' stopper (lift scan));  | 
| 
6116
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
256  | 
|
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
257  | 
fun read stopper scan xs =  | 
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
258  | 
(case error (finite stopper (option scan)) xs of  | 
| 15531 | 259  | 
(y as SOME _, []) => y  | 
260  | 
| _ => NONE);  | 
|
| 
6116
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
261  | 
|
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
262  | 
|
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
263  | 
(* infinite scans -- draining state-based source *)  | 
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
264  | 
|
| 23699 | 265  | 
fun drain def_prompt get stopper scan ((state, xs), src) =  | 
266  | 
(scan (state, xs), src) handle MORE prompt =>  | 
|
267  | 
(case get (the_default def_prompt prompt) src of  | 
|
| 
6116
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
268  | 
([], _) => (finite' stopper scan (state, xs), src)  | 
| 23699 | 269  | 
| (xs', src') => drain def_prompt get stopper scan ((state, xs @ xs'), src'));  | 
| 
6116
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
270  | 
|
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
271  | 
|
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
272  | 
|
| 
27782
 
377810fd718e
datatype lexicon: alternative representation using nested Symtab.table;
 
wenzelm 
parents: 
27765 
diff
changeset
 | 
273  | 
(** datatype lexicon -- position tree **)  | 
| 
 
377810fd718e
datatype lexicon: alternative representation using nested Symtab.table;
 
wenzelm 
parents: 
27765 
diff
changeset
 | 
274  | 
|
| 
 
377810fd718e
datatype lexicon: alternative representation using nested Symtab.table;
 
wenzelm 
parents: 
27765 
diff
changeset
 | 
275  | 
datatype lexicon = Lexicon of (bool * lexicon) Symtab.table;  | 
| 
 
377810fd718e
datatype lexicon: alternative representation using nested Symtab.table;
 
wenzelm 
parents: 
27765 
diff
changeset
 | 
276  | 
|
| 
 
377810fd718e
datatype lexicon: alternative representation using nested Symtab.table;
 
wenzelm 
parents: 
27765 
diff
changeset
 | 
277  | 
val empty_lexicon = Lexicon Symtab.empty;  | 
| 
6116
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
278  | 
|
| 
27782
 
377810fd718e
datatype lexicon: alternative representation using nested Symtab.table;
 
wenzelm 
parents: 
27765 
diff
changeset
 | 
279  | 
fun is_literal _ [] = false  | 
| 32784 | 280  | 
| is_literal (Lexicon tab) (c :: cs) =  | 
| 
27782
 
377810fd718e
datatype lexicon: alternative representation using nested Symtab.table;
 
wenzelm 
parents: 
27765 
diff
changeset
 | 
281  | 
(case Symtab.lookup tab c of  | 
| 
 
377810fd718e
datatype lexicon: alternative representation using nested Symtab.table;
 
wenzelm 
parents: 
27765 
diff
changeset
 | 
282  | 
SOME (tip, lex) => tip andalso null cs orelse is_literal lex cs  | 
| 
 
377810fd718e
datatype lexicon: alternative representation using nested Symtab.table;
 
wenzelm 
parents: 
27765 
diff
changeset
 | 
283  | 
| NONE => false);  | 
| 
6116
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
284  | 
|
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
285  | 
|
| 
27782
 
377810fd718e
datatype lexicon: alternative representation using nested Symtab.table;
 
wenzelm 
parents: 
27765 
diff
changeset
 | 
286  | 
(* scan longest match *)  | 
| 
27765
 
5df443dd9deb
reorganized lexicon: allow scanning of annotated symbols, tuned representation and interfaces;
 
wenzelm 
parents: 
27731 
diff
changeset
 | 
287  | 
|
| 
27782
 
377810fd718e
datatype lexicon: alternative representation using nested Symtab.table;
 
wenzelm 
parents: 
27765 
diff
changeset
 | 
288  | 
fun literal lexicon =  | 
| 
27765
 
5df443dd9deb
reorganized lexicon: allow scanning of annotated symbols, tuned representation and interfaces;
 
wenzelm 
parents: 
27731 
diff
changeset
 | 
289  | 
let  | 
| 
27782
 
377810fd718e
datatype lexicon: alternative representation using nested Symtab.table;
 
wenzelm 
parents: 
27765 
diff
changeset
 | 
290  | 
fun finish (SOME (res, rest)) = (rev res, rest)  | 
| 
 
377810fd718e
datatype lexicon: alternative representation using nested Symtab.table;
 
wenzelm 
parents: 
27765 
diff
changeset
 | 
291  | 
| finish NONE = raise FAIL NONE;  | 
| 
 
377810fd718e
datatype lexicon: alternative representation using nested Symtab.table;
 
wenzelm 
parents: 
27765 
diff
changeset
 | 
292  | 
fun scan _ res (Lexicon tab) [] = if Symtab.is_empty tab then finish res else raise MORE NONE  | 
| 32784 | 293  | 
| scan path res (Lexicon tab) (c :: cs) =  | 
| 
27782
 
377810fd718e
datatype lexicon: alternative representation using nested Symtab.table;
 
wenzelm 
parents: 
27765 
diff
changeset
 | 
294  | 
(case Symtab.lookup tab (fst c) of  | 
| 
 
377810fd718e
datatype lexicon: alternative representation using nested Symtab.table;
 
wenzelm 
parents: 
27765 
diff
changeset
 | 
295  | 
SOME (tip, lex) =>  | 
| 
 
377810fd718e
datatype lexicon: alternative representation using nested Symtab.table;
 
wenzelm 
parents: 
27765 
diff
changeset
 | 
296  | 
let val path' = c :: path  | 
| 
 
377810fd718e
datatype lexicon: alternative representation using nested Symtab.table;
 
wenzelm 
parents: 
27765 
diff
changeset
 | 
297  | 
in scan path' (if tip then SOME (path', cs) else res) lex cs end  | 
| 27784 | 298  | 
| NONE => finish res);  | 
| 
27782
 
377810fd718e
datatype lexicon: alternative representation using nested Symtab.table;
 
wenzelm 
parents: 
27765 
diff
changeset
 | 
299  | 
in scan [] NONE lexicon end;  | 
| 
6116
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
300  | 
|
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
301  | 
|
| 
27765
 
5df443dd9deb
reorganized lexicon: allow scanning of annotated symbols, tuned representation and interfaces;
 
wenzelm 
parents: 
27731 
diff
changeset
 | 
302  | 
(* build lexicons *)  | 
| 
6116
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
303  | 
|
| 
27765
 
5df443dd9deb
reorganized lexicon: allow scanning of annotated symbols, tuned representation and interfaces;
 
wenzelm 
parents: 
27731 
diff
changeset
 | 
304  | 
fun extend_lexicon chrs lexicon =  | 
| 
 
5df443dd9deb
reorganized lexicon: allow scanning of annotated symbols, tuned representation and interfaces;
 
wenzelm 
parents: 
27731 
diff
changeset
 | 
305  | 
let  | 
| 
27782
 
377810fd718e
datatype lexicon: alternative representation using nested Symtab.table;
 
wenzelm 
parents: 
27765 
diff
changeset
 | 
306  | 
fun ext [] lex = lex  | 
| 32784 | 307  | 
| ext (c :: cs) (Lexicon tab) =  | 
| 
27782
 
377810fd718e
datatype lexicon: alternative representation using nested Symtab.table;
 
wenzelm 
parents: 
27765 
diff
changeset
 | 
308  | 
(case Symtab.lookup tab c of  | 
| 
 
377810fd718e
datatype lexicon: alternative representation using nested Symtab.table;
 
wenzelm 
parents: 
27765 
diff
changeset
 | 
309  | 
SOME (tip, lex) => Lexicon (Symtab.update (c, (tip orelse null cs, ext cs lex)) tab)  | 
| 
 
377810fd718e
datatype lexicon: alternative representation using nested Symtab.table;
 
wenzelm 
parents: 
27765 
diff
changeset
 | 
310  | 
| NONE => Lexicon (Symtab.update (c, (null cs, ext cs empty_lexicon)) tab));  | 
| 
27765
 
5df443dd9deb
reorganized lexicon: allow scanning of annotated symbols, tuned representation and interfaces;
 
wenzelm 
parents: 
27731 
diff
changeset
 | 
311  | 
in if is_literal lexicon chrs then lexicon else ext chrs lexicon end;  | 
| 
6116
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
312  | 
|
| 
27765
 
5df443dd9deb
reorganized lexicon: allow scanning of annotated symbols, tuned representation and interfaces;
 
wenzelm 
parents: 
27731 
diff
changeset
 | 
313  | 
fun make_lexicon chrss = fold extend_lexicon chrss empty_lexicon;  | 
| 
6116
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
314  | 
|
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
315  | 
|
| 
27765
 
5df443dd9deb
reorganized lexicon: allow scanning of annotated symbols, tuned representation and interfaces;
 
wenzelm 
parents: 
27731 
diff
changeset
 | 
316  | 
(* merge lexicons *)  | 
| 
6116
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
317  | 
|
| 
27782
 
377810fd718e
datatype lexicon: alternative representation using nested Symtab.table;
 
wenzelm 
parents: 
27765 
diff
changeset
 | 
318  | 
fun dest path (Lexicon tab) = Symtab.fold (fn (d, (tip, lex)) =>  | 
| 
6116
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
319  | 
let  | 
| 
27782
 
377810fd718e
datatype lexicon: alternative representation using nested Symtab.table;
 
wenzelm 
parents: 
27765 
diff
changeset
 | 
320  | 
val path' = d :: path;  | 
| 
 
377810fd718e
datatype lexicon: alternative representation using nested Symtab.table;
 
wenzelm 
parents: 
27765 
diff
changeset
 | 
321  | 
val content = dest path' lex;  | 
| 
 
377810fd718e
datatype lexicon: alternative representation using nested Symtab.table;
 
wenzelm 
parents: 
27765 
diff
changeset
 | 
322  | 
in append (if tip then rev path' :: content else content) end) tab [];  | 
| 
27765
 
5df443dd9deb
reorganized lexicon: allow scanning of annotated symbols, tuned representation and interfaces;
 
wenzelm 
parents: 
27731 
diff
changeset
 | 
323  | 
|
| 
27782
 
377810fd718e
datatype lexicon: alternative representation using nested Symtab.table;
 
wenzelm 
parents: 
27765 
diff
changeset
 | 
324  | 
val dest_lexicon = map implode o dest [];  | 
| 
 
377810fd718e
datatype lexicon: alternative representation using nested Symtab.table;
 
wenzelm 
parents: 
27765 
diff
changeset
 | 
325  | 
fun merge_lexicons (lex1, lex2) = fold extend_lexicon (dest [] lex2) lex1;  | 
| 
6116
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
326  | 
|
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
327  | 
end;  | 
| 
 
8ba2f25610f7
files scan.ML, source.ML, symbol.ML, pretty.ML moved to Pure/General;
 
wenzelm 
parents:  
diff
changeset
 | 
328  | 
|
| 
37216
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
33957 
diff
changeset
 | 
329  | 
structure Basic_Scan: BASIC_SCAN = Scan;  | 
| 
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
33957 
diff
changeset
 | 
330  | 
open Basic_Scan;  |