| author | paulson | 
| Thu, 22 May 1997 15:10:37 +0200 | |
| changeset 3298 | 5f0ed3caa991 | 
| parent 3245 | 241838c01caf | 
| child 3302 | 404fe31fd8d2 | 
| permissions | -rw-r--r-- | 
| 2112 | 1  | 
structure Thry : Thry_sig (* LThry_sig *) =  | 
2  | 
struct  | 
|
3  | 
||
4  | 
structure USyntax = USyntax;  | 
|
5  | 
||
6  | 
open Mask;  | 
|
7  | 
structure S = USyntax;  | 
|
8  | 
||
9  | 
||
10  | 
fun THRY_ERR{func,mesg} = Utils.ERR{module = "Thry",func=func,mesg=mesg};
 | 
|
11  | 
||
12  | 
(*---------------------------------------------------------------------------  | 
|
13  | 
* Matching  | 
|
14  | 
*---------------------------------------------------------------------------*)  | 
|
15  | 
||
16  | 
local open Utils  | 
|
17  | 
infix 3 |->  | 
|
18  | 
fun tybind (x,y) = TVar (x,["term"]) |-> y  | 
|
19  | 
fun tmbind (x,y) = Var (x,type_of y) |-> y  | 
|
20  | 
in  | 
|
21  | 
fun match_term thry pat ob =  | 
|
22  | 
let val tsig = #tsig(Sign.rep_sg(sign_of thry))  | 
|
23  | 
val (ty_theta,tm_theta) = Pattern.match tsig (pat,ob)  | 
|
24  | 
in (map tmbind tm_theta, map tybind ty_theta)  | 
|
25  | 
end  | 
|
26  | 
||
27  | 
fun match_type thry pat ob =  | 
|
28  | 
map tybind(Type.typ_match (#tsig(Sign.rep_sg(sign_of thry))) ([],(pat,ob)))  | 
|
29  | 
end;  | 
|
30  | 
||
31  | 
||
32  | 
(*---------------------------------------------------------------------------  | 
|
33  | 
* Typing  | 
|
34  | 
*---------------------------------------------------------------------------*)  | 
|
35  | 
||
36  | 
fun typecheck thry = cterm_of (sign_of thry);  | 
|
37  | 
||
38  | 
||
39  | 
||
40  | 
(*----------------------------------------------------------------------------  | 
|
41  | 
* Making a definition. The argument "tm" looks like "f = WFREC R M". This  | 
|
42  | 
* entrypoint is specialized for interactive use, since it closes the theory  | 
|
43  | 
* after making the definition. This allows later interactive definitions to  | 
|
44  | 
* refer to previous ones. The name for the new theory is automatically  | 
|
45  | 
* generated from the name of the argument theory.  | 
|
46  | 
*---------------------------------------------------------------------------*)  | 
|
| 3191 | 47  | 
|
48  | 
||
49  | 
(*---------------------------------------------------------------------------  | 
|
50  | 
* TFL attempts to make definitions where the lhs is a variable. Isabelle  | 
|
51  | 
* wants it to be a constant, so here we map it to a constant. Moreover, the  | 
|
52  | 
* theory should already have the constant, so we refrain from adding the  | 
|
53  | 
* constant to the theory. We just add the axiom and return the theory.  | 
|
54  | 
*---------------------------------------------------------------------------*)  | 
|
| 2112 | 55  | 
local val (imp $ tprop $ (eeq $ _ $ _ )) = #prop(rep_thm(eq_reflection))  | 
56  | 
val Const(eeq_name, ty) = eeq  | 
|
57  | 
val prop = #2 (S.strip_type ty)  | 
|
58  | 
in  | 
|
59  | 
fun make_definition parent s tm =  | 
|
60  | 
   let val {lhs,rhs} = S.dest_eq tm
 | 
|
61  | 
       val {Name,Ty} = S.dest_var lhs
 | 
|
62  | 
       val lhs1 = S.mk_const{Name = Name, Ty = Ty}
 | 
|
63  | 
       val eeq1 = S.mk_const{Name = eeq_name, Ty = Ty --> Ty --> prop}
 | 
|
| 
3245
 
241838c01caf
Removal of redundant code (unused or already present in Isabelle.
 
paulson 
parents: 
3191 
diff
changeset
 | 
64  | 
val dtm = list_comb(eeq1,[lhs1,rhs]) (* Rename "=" to "==" *)  | 
| 3191 | 65  | 
val (_, tm', _) = Sign.infer_types (sign_of parent)  | 
| 
3245
 
241838c01caf
Removal of redundant code (unused or already present in Isabelle.
 
paulson 
parents: 
3191 
diff
changeset
 | 
66  | 
(K None) (K None) [] true ([dtm],propT)  | 
| 3191 | 67  | 
val new_thy = add_defs_i [(s,tm')] parent  | 
| 2112 | 68  | 
in  | 
| 3191 | 69  | 
(freezeT((get_axiom new_thy s) RS meta_eq_to_obj_eq), new_thy)  | 
70  | 
end;  | 
|
| 2112 | 71  | 
end;  | 
72  | 
||
73  | 
(*---------------------------------------------------------------------------  | 
|
74  | 
* Utility routine. Insert into list ordered by the key (a string). If two  | 
|
75  | 
* keys are equal, the new element replaces the old. A more efficient option  | 
|
76  | 
* for the future is needed. In fact, having the list of datatype facts be  | 
|
77  | 
* ordered is useless, since the lookup should never fail!  | 
|
78  | 
*---------------------------------------------------------------------------*)  | 
|
79  | 
fun insert (el as (x:string, _)) =  | 
|
80  | 
let fun canfind[] = [el]  | 
|
81  | 
| canfind(alist as ((y as (k,_))::rst)) =  | 
|
82  | 
if (x<k) then el::alist  | 
|
83  | 
else if (x=k) then el::rst  | 
|
84  | 
else y::canfind rst  | 
|
85  | 
in canfind  | 
|
86  | 
end;  | 
|
87  | 
||
88  | 
||
89  | 
(*---------------------------------------------------------------------------  | 
|
90  | 
* A collection of facts about datatypes  | 
|
91  | 
*---------------------------------------------------------------------------*)  | 
|
92  | 
val nat_record = Dtype.build_record (Nat.thy, ("nat",["0","Suc"]), nat_ind_tac)
 | 
|
93  | 
val prod_record =  | 
|
94  | 
let val prod_case_thms = Dtype.case_thms (sign_of Prod.thy) [split]  | 
|
95  | 
                                 (fn s => res_inst_tac [("p",s)] PairE_lemma)
 | 
|
96  | 
fun const s = Const(s, the(Sign.const_type (sign_of Prod.thy) s))  | 
|
97  | 
     in ("*", 
 | 
|
98  | 
         {constructors = [const "Pair"],
 | 
|
99  | 
case_const = const "split",  | 
|
100  | 
case_rewrites = [split RS eq_reflection],  | 
|
101  | 
case_cong = #case_cong prod_case_thms,  | 
|
102  | 
nchotomy = #nchotomy prod_case_thms})  | 
|
103  | 
end;  | 
|
104  | 
||
105  | 
(*---------------------------------------------------------------------------  | 
|
106  | 
* Hacks to make interactive mode work. Referring to "datatypes" directly  | 
|
107  | 
* is temporary, I hope!  | 
|
108  | 
*---------------------------------------------------------------------------*)  | 
|
109  | 
val match_info = fn thy =>  | 
|
| 
3245
 
241838c01caf
Removal of redundant code (unused or already present in Isabelle.
 
paulson 
parents: 
3191 
diff
changeset
 | 
110  | 
    fn "*" => Some({case_const = #case_const (#2 prod_record),
 | 
| 2112 | 111  | 
constructors = #constructors (#2 prod_record)})  | 
| 
3245
 
241838c01caf
Removal of redundant code (unused or already present in Isabelle.
 
paulson 
parents: 
3191 
diff
changeset
 | 
112  | 
     | "nat" => Some({case_const = #case_const (#2 nat_record),
 | 
| 2112 | 113  | 
constructors = #constructors (#2 nat_record)})  | 
114  | 
| ty => case assoc(!datatypes,ty)  | 
|
| 
3245
 
241838c01caf
Removal of redundant code (unused or already present in Isabelle.
 
paulson 
parents: 
3191 
diff
changeset
 | 
115  | 
of None => None  | 
| 2112 | 116  | 
                | Some{case_const,constructors, ...} =>
 | 
| 
3245
 
241838c01caf
Removal of redundant code (unused or already present in Isabelle.
 
paulson 
parents: 
3191 
diff
changeset
 | 
117  | 
                   Some{case_const=case_const, constructors=constructors}
 | 
| 2112 | 118  | 
|
119  | 
val induct_info = fn thy =>  | 
|
| 
3245
 
241838c01caf
Removal of redundant code (unused or already present in Isabelle.
 
paulson 
parents: 
3191 
diff
changeset
 | 
120  | 
    fn "*" => Some({nchotomy = #nchotomy (#2 prod_record),
 | 
| 2112 | 121  | 
constructors = #constructors (#2 prod_record)})  | 
| 
3245
 
241838c01caf
Removal of redundant code (unused or already present in Isabelle.
 
paulson 
parents: 
3191 
diff
changeset
 | 
122  | 
     | "nat" => Some({nchotomy = #nchotomy (#2 nat_record),
 | 
| 2112 | 123  | 
constructors = #constructors (#2 nat_record)})  | 
124  | 
| ty => case assoc(!datatypes,ty)  | 
|
| 
3245
 
241838c01caf
Removal of redundant code (unused or already present in Isabelle.
 
paulson 
parents: 
3191 
diff
changeset
 | 
125  | 
of None => None  | 
| 2112 | 126  | 
                | Some{nchotomy,constructors, ...} =>
 | 
| 
3245
 
241838c01caf
Removal of redundant code (unused or already present in Isabelle.
 
paulson 
parents: 
3191 
diff
changeset
 | 
127  | 
                  Some{nchotomy=nchotomy, constructors=constructors}
 | 
| 2112 | 128  | 
|
129  | 
val extract_info = fn thy =>  | 
|
130  | 
let val case_congs = map (#case_cong o #2) (!datatypes)  | 
|
131  | 
val case_rewrites = flat(map (#case_rewrites o #2) (!datatypes))  | 
|
132  | 
 in {case_congs = #case_cong (#2 prod_record)::
 | 
|
133  | 
#case_cong (#2 nat_record)::case_congs,  | 
|
134  | 
case_rewrites = #case_rewrites(#2 prod_record)@  | 
|
135  | 
#case_rewrites(#2 nat_record)@case_rewrites}  | 
|
136  | 
end;  | 
|
137  | 
||
138  | 
end; (* Thry *)  |