| author | nipkow | 
| Fri, 02 Nov 2007 08:59:15 +0100 | |
| changeset 25261 | 3dc292be0b54 | 
| parent 24487 | e92b74b3a254 | 
| child 25999 | f8bcd311d501 | 
| permissions | -rw-r--r-- | 
| 
24236
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
1  | 
(* Title: Pure/Syntax/simple_syntax.ML  | 
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
2  | 
ID: $Id$  | 
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
3  | 
Author: Makarius  | 
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
4  | 
|
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
5  | 
Simple syntax for types and terms --- for bootstrapping Pure.  | 
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
6  | 
*)  | 
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
7  | 
|
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
8  | 
signature SIMPLE_SYNTAX =  | 
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
9  | 
sig  | 
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
10  | 
val read_typ: string -> typ  | 
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
11  | 
val read_term: string -> term  | 
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
12  | 
val read_prop: string -> term  | 
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
13  | 
end;  | 
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
14  | 
|
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
15  | 
structure SimpleSyntax: SIMPLE_SYNTAX =  | 
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
16  | 
struct  | 
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
17  | 
|
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
18  | 
(* scanning tokens *)  | 
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
19  | 
|
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
20  | 
local  | 
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
21  | 
|
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
22  | 
val eof = Lexicon.EndToken;  | 
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
23  | 
fun is_eof tk = tk = Lexicon.EndToken;  | 
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
24  | 
|
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
25  | 
val stopper = (eof, is_eof);  | 
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
26  | 
val not_eof = not o is_eof;  | 
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
27  | 
|
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
28  | 
val lexicon = Scan.make_lexicon  | 
| 24246 | 29  | 
  (map Symbol.explode ["!!", "%", "(", ")", ".", "::", "==", "==>", "=?=", "=>", "&&", "CONST"]);
 | 
| 
24236
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
30  | 
|
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
31  | 
in  | 
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
32  | 
|
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
33  | 
fun read scan s =  | 
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
34  | 
(case  | 
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
35  | 
Symbol.explode s |>  | 
| 24246 | 36  | 
Lexicon.tokenize lexicon false |>  | 
| 
24236
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
37  | 
Scan.read stopper scan of  | 
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
38  | 
SOME x => x  | 
| 24246 | 39  | 
  | NONE => error ("Malformed input: " ^ quote s));
 | 
| 
24236
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
40  | 
|
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
41  | 
end;  | 
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
42  | 
|
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
43  | 
|
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
44  | 
(* basic scanners *)  | 
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
45  | 
|
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
46  | 
fun $$ s = Scan.some (fn Lexicon.Token s' => if s = s' then SOME s else NONE | _ => NONE);  | 
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
47  | 
fun enum1 s scan = scan -- Scan.repeat ($$ s |-- scan) >> op ::;  | 
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
48  | 
fun enum2 s scan = scan -- Scan.repeat1 ($$ s |-- scan) >> op ::;  | 
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
49  | 
|
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
50  | 
val tfree = Scan.some (fn Lexicon.TFreeSy s => SOME s | _ => NONE);  | 
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
51  | 
val ident = Scan.some (fn Lexicon.IdentSy s => SOME s | _ => NONE);  | 
| 24246 | 52  | 
val var =  | 
53  | 
Scan.some (fn Lexicon.VarSy s => SOME (Lexicon.read_indexname (unprefix "?" s)) | _ => NONE);  | 
|
| 
24236
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
54  | 
val long_ident = Scan.some (fn Lexicon.LongIdentSy s => SOME s | _ => NONE);  | 
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
55  | 
val const = long_ident || ident;  | 
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
56  | 
|
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
57  | 
|
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
58  | 
(* types *)  | 
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
59  | 
|
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
60  | 
(*  | 
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
61  | 
typ = typ1 => ... => typ1  | 
| 24237 | 62  | 
| typ1  | 
| 
24236
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
63  | 
typ1 = typ2 const ... const  | 
| 24237 | 64  | 
| typ2  | 
65  | 
typ2 = tfree  | 
|
66  | 
| const  | 
|
67  | 
| ( typ )  | 
|
| 
24236
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
68  | 
*)  | 
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
69  | 
|
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
70  | 
fun typ x =  | 
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
71  | 
(enum1 "=>" typ1 >> (op ---> o split_last)) x  | 
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
72  | 
and typ1 x =  | 
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
73  | 
(typ2 -- Scan.repeat const >> (fn (T, cs) => fold (fn c => fn U => Type (c, [U])) cs T)) x  | 
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
74  | 
and typ2 x =  | 
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
75  | 
(tfree >> (fn a => TFree (a, [])) ||  | 
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
76  | 
const >> (fn c => Type (c, [])) ||  | 
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
77  | 
  $$ "(" |-- typ --| $$ ")") x;
 | 
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
78  | 
|
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
79  | 
val read_typ = read typ;  | 
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
80  | 
|
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
81  | 
|
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
82  | 
(* terms *)  | 
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
83  | 
|
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
84  | 
(*  | 
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
85  | 
term = !!ident :: typ. term  | 
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
86  | 
| term1  | 
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
87  | 
term1 = term2 ==> ... ==> term2  | 
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
88  | 
| term2  | 
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
89  | 
term2 = term3 == term2  | 
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
90  | 
| term3 =?= term2  | 
| 24246 | 91  | 
| term3 && term2  | 
| 
24236
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
92  | 
| term3  | 
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
93  | 
term3 = ident :: typ  | 
| 24246 | 94  | 
| var :: typ  | 
| 
24236
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
95  | 
| CONST const :: typ  | 
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
96  | 
| %ident :: typ. term3  | 
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
97  | 
| term4  | 
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
98  | 
term4 = term5 ... term5  | 
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
99  | 
| term5  | 
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
100  | 
term5 = ident  | 
| 24246 | 101  | 
| var  | 
| 
24236
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
102  | 
| CONST const  | 
| 24237 | 103  | 
| ( term )  | 
| 
24236
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
104  | 
*)  | 
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
105  | 
|
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
106  | 
local  | 
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
107  | 
|
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
108  | 
val constraint = $$ "::" |-- typ;  | 
| 24246 | 109  | 
val idt = ident -- constraint;  | 
110  | 
val bind = idt --| $$ ".";  | 
|
| 
24236
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
111  | 
|
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
112  | 
fun term env T x =  | 
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
113  | 
($$ "!!" |-- bind :|-- (fn v =>  | 
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
114  | 
term (v :: env) propT >> (fn B => Term.all (#2 v) $ lambda (Free v) B)) ||  | 
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
115  | 
term1 env T) x  | 
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
116  | 
and term1 env T x =  | 
| 24262 | 117  | 
(enum2 "==>" (term2 env propT) >> foldr1 Logic.mk_implies ||  | 
| 
24236
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
118  | 
term2 env T) x  | 
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
119  | 
and term2 env T x =  | 
| 24246 | 120  | 
(equal env "==" || equal env "=?=" ||  | 
| 24262 | 121  | 
term3 env propT -- ($$ "&&" |-- term2 env propT) >> Logic.mk_conjunction ||  | 
| 24246 | 122  | 
term3 env T) x  | 
| 
24236
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
123  | 
and equal env eq x =  | 
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
124  | 
(term3 env dummyT -- ($$ eq |-- term2 env dummyT) >> (fn (t, u) =>  | 
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
125  | 
Const (eq, Term.fastype_of t --> Term.fastype_of u --> propT) $ t $ u)) x  | 
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
126  | 
and term3 env T x =  | 
| 24246 | 127  | 
(idt >> Free ||  | 
128  | 
var -- constraint >> Var ||  | 
|
| 
24236
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
129  | 
$$ "CONST" |-- const -- constraint >> Const ||  | 
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
130  | 
$$ "%" |-- bind :|-- (fn v => term3 (v :: env) dummyT >> lambda (Free v)) ||  | 
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
131  | 
term4 env T) x  | 
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
132  | 
and term4 env T x =  | 
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
133  | 
(term5 env dummyT -- Scan.repeat1 (term5 env dummyT) >> Term.list_comb ||  | 
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
134  | 
term5 env T) x  | 
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
135  | 
and term5 env T x =  | 
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
136  | 
(ident >> (fn a => Free (a, the_default T (AList.lookup (op =) env a))) ||  | 
| 24246 | 137  | 
var >> (fn xi => Var (xi, T)) ||  | 
| 
24236
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
138  | 
$$ "CONST" |-- const >> (fn c => Const (c, T)) ||  | 
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
139  | 
  $$ "(" |-- term env T --| $$ ")") x;
 | 
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
140  | 
|
| 24246 | 141  | 
fun read_tm T s =  | 
142  | 
let val t = read (term [] T) s in  | 
|
143  | 
if can (Term.map_types Term.no_dummyT) t then t  | 
|
144  | 
    else error ("Unspecified types in input: " ^ quote s)
 | 
|
145  | 
end;  | 
|
146  | 
||
| 24487 | 147  | 
in  | 
148  | 
||
| 24246 | 149  | 
val read_term = read_tm dummyT;  | 
150  | 
val read_prop = read_tm propT;  | 
|
| 
24236
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
151  | 
|
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
152  | 
end;  | 
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
153  | 
|
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
154  | 
end;  | 
| 
 
d8b05073edc7
Simple syntax for types and terms --- for bootstrapping Pure.
 
wenzelm 
parents:  
diff
changeset
 | 
155  |