| author | paulson | 
| Wed, 13 Oct 1999 12:08:05 +0200 | |
| changeset 7845 | 3561e0da35b8 | 
| parent 7653 | 0408848fa7d4 | 
| child 11728 | b5f6963b193c | 
| permissions | -rw-r--r-- | 
| 
2445
 
51993fea433f
removed Holcfb.thy and Holcfb.ML, moving classical3 to HOL.ML as classical2
 
oheimb 
parents: 
1637 
diff
changeset
 | 
1  | 
(* Title: HOLCF/domain/extender.ML  | 
| 
 
51993fea433f
removed Holcfb.thy and Holcfb.ML, moving classical3 to HOL.ML as classical2
 
oheimb 
parents: 
1637 
diff
changeset
 | 
2  | 
ID: $Id$  | 
| 
 
51993fea433f
removed Holcfb.thy and Holcfb.ML, moving classical3 to HOL.ML as classical2
 
oheimb 
parents: 
1637 
diff
changeset
 | 
3  | 
Author : David von Oheimb  | 
| 
 
51993fea433f
removed Holcfb.thy and Holcfb.ML, moving classical3 to HOL.ML as classical2
 
oheimb 
parents: 
1637 
diff
changeset
 | 
4  | 
Copyright 1995, 1996 TU Muenchen  | 
| 
 
51993fea433f
removed Holcfb.thy and Holcfb.ML, moving classical3 to HOL.ML as classical2
 
oheimb 
parents: 
1637 
diff
changeset
 | 
5  | 
|
| 
 
51993fea433f
removed Holcfb.thy and Holcfb.ML, moving classical3 to HOL.ML as classical2
 
oheimb 
parents: 
1637 
diff
changeset
 | 
6  | 
theory extender for domain section  | 
| 1274 | 7  | 
*)  | 
8  | 
||
| 
2446
 
c2a9bf6c0948
The previous log message was wrong. The correct one is:
 
oheimb 
parents: 
2445 
diff
changeset
 | 
9  | 
|
| 
1637
 
b8a8ae2e5de1
Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
 
oheimb 
parents: 
1461 
diff
changeset
 | 
10  | 
structure Domain_Extender =  | 
| 1274 | 11  | 
struct  | 
12  | 
||
13  | 
local  | 
|
14  | 
||
15  | 
open Domain_Library;  | 
|
16  | 
||
| 
2445
 
51993fea433f
removed Holcfb.thy and Holcfb.ML, moving classical3 to HOL.ML as classical2
 
oheimb 
parents: 
1637 
diff
changeset
 | 
17  | 
(* ----- general testing and preprocessing of constructor list -------------- *)  | 
| 1274 | 18  | 
|
| 4008 | 19  | 
fun check_and_sort_domain (dtnvs: (string * typ list) list, cons'' :  | 
| 4122 | 20  | 
((string * mixfix * (bool*string*typ) list) list) list) sg =  | 
| 4008 | 21  | 
let  | 
| 7653 | 22  | 
val defaultS = Sign.defaultS sg;  | 
| 1274 | 23  | 
val test_dupl_typs = (case duplicates (map fst dtnvs) of  | 
| 
1637
 
b8a8ae2e5de1
Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
 
oheimb 
parents: 
1461 
diff
changeset
 | 
24  | 
	[] => false | dups => error ("Duplicate types: " ^ commas_quote dups));
 | 
| 4008 | 25  | 
val test_dupl_cons = (case duplicates (map first (flat cons'')) of  | 
26  | 
	[] => false | dups => error ("Duplicate constructors: " 
 | 
|
27  | 
^ commas_quote dups));  | 
|
28  | 
val test_dupl_sels = (case duplicates  | 
|
29  | 
(map second (flat (map third (flat cons'')))) of  | 
|
| 
2445
 
51993fea433f
removed Holcfb.thy and Holcfb.ML, moving classical3 to HOL.ML as classical2
 
oheimb 
parents: 
1637 
diff
changeset
 | 
30  | 
        [] => false | dups => error("Duplicate selectors: "^commas_quote dups));
 | 
| 
 
51993fea433f
removed Holcfb.thy and Holcfb.ML, moving classical3 to HOL.ML as classical2
 
oheimb 
parents: 
1637 
diff
changeset
 | 
31  | 
val test_dupl_tvars = exists(fn s=>case duplicates(map(fst o rep_TFree)s)of  | 
| 
 
51993fea433f
removed Holcfb.thy and Holcfb.ML, moving classical3 to HOL.ML as classical2
 
oheimb 
parents: 
1637 
diff
changeset
 | 
32  | 
	[] => false | dups => error("Duplicate type arguments: " 
 | 
| 4008 | 33  | 
^commas_quote dups)) (map snd dtnvs);  | 
34  | 
(* test for free type variables, illegal sort constraints on rhs,  | 
|
35  | 
non-pcpo-types and invalid use of recursive type;  | 
|
36  | 
replace sorts in type variables on rhs *)  | 
|
37  | 
fun analyse_equation ((dname,typevars),cons') =  | 
|
38  | 
let  | 
|
39  | 
val tvars = map rep_TFree typevars;  | 
|
| 4030 | 40  | 
fun distinct_name s = "'"^Sign.base_name dname^"_"^s;  | 
41  | 
val distinct_typevars = map (fn (n,sort) =>  | 
|
42  | 
TFree (distinct_name n,sort)) tvars;  | 
|
| 4008 | 43  | 
fun rm_sorts (TFree(s,_)) = TFree(s,[])  | 
44  | 
| rm_sorts (Type(s,ts)) = Type(s,remove_sorts ts)  | 
|
45  | 
| rm_sorts (TVar(s,_)) = TVar(s,[])  | 
|
46  | 
and remove_sorts l = map rm_sorts l;  | 
|
47  | 
fun analyse(TFree(v,s)) = (case assoc_string(tvars,v) of  | 
|
48  | 
		    None      => error ("Free type variable " ^ v ^ " on rhs.")
 | 
|
49  | 
| Some sort => if eq_set_string (s,defaultS) orelse  | 
|
| 4030 | 50  | 
eq_set_string (s,sort )  | 
51  | 
then TFree(distinct_name v,sort)  | 
|
| 4008 | 52  | 
				 else error ("Additional constraint on rhs "^
 | 
53  | 
"for type variable "^quote v))  | 
|
| 4753 | 54  | 
(** BUG OR FEATURE?: mutual recursion may use different arguments **)  | 
55  | 
| analyse(Type(s,typl)) = (case assoc_string((*dtnvs*)  | 
|
56  | 
[(dname,typevars)],s) of  | 
|
57  | 
None => Type(s,map analyse typl)  | 
|
58  | 
| Some typevars => if remove_sorts typevars = remove_sorts typl  | 
|
59  | 
then Type(s,map analyse typl)  | 
|
| 4008 | 60  | 
				else error ("Recursion of type " ^ s ^ 
 | 
| 4753 | 61  | 
" with different arguments"))  | 
62  | 
| analyse(TVar _) = Imposs "extender:analyse";  | 
|
63  | 
fun check_pcpo t = (pcpo_type sg t orelse error(  | 
|
64  | 
"Type not of sort pcpo: "^string_of_typ sg t); t);  | 
|
| 4008 | 65  | 
val analyse_con = upd_third (map (upd_third (check_pcpo o analyse)));  | 
| 4030 | 66  | 
in ((dname,distinct_typevars), map analyse_con cons') end;  | 
| 4008 | 67  | 
in ListPair.map analyse_equation (dtnvs,cons'')  | 
68  | 
end; (* let *)  | 
|
69  | 
||
| 4030 | 70  | 
(* ----- calls for building new thy and thms -------------------------------- *)  | 
| 1274 | 71  | 
|
72  | 
in  | 
|
73  | 
||
| 4008 | 74  | 
fun add_domain (comp_dnam,eqs''') thy''' = let  | 
75  | 
val sg''' = sign_of thy''';  | 
|
76  | 
val dtnvs = map ((fn (dname,vs) =>  | 
|
77  | 
(Sign.full_name sg''' dname,map (str2typ sg''') vs))  | 
|
78  | 
o fst) eqs''';  | 
|
79  | 
val cons''' = map snd eqs''';  | 
|
80  | 
fun thy_type (dname,tvars) = (Sign.base_name dname, length tvars, NoSyn);  | 
|
81  | 
fun thy_arity (dname,tvars) = (dname, map (snd o rep_TFree) tvars, pcpoS);  | 
|
82  | 
val thy'' = thy''' |> Theory.add_types (map thy_type dtnvs)  | 
|
83  | 
|> Theory.add_arities_i (map thy_arity dtnvs);  | 
|
84  | 
val sg'' = sign_of thy'';  | 
|
85  | 
val cons''=map (map (upd_third (map (upd_third (str2typ sg''))))) cons''';  | 
|
86  | 
val eqs' = check_and_sort_domain (dtnvs,cons'') sg'';  | 
|
87  | 
val thy' = thy'' |> Domain_Syntax.add_syntax (comp_dnam,eqs');  | 
|
| 1274 | 88  | 
val dts = map (Type o fst) eqs';  | 
| 
4188
 
1025a27b08f9
changed libraray function find to find_index_eq, currying it
 
oheimb 
parents: 
4122 
diff
changeset
 | 
89  | 
fun strip ss = drop (find_index_eq "'" ss +1, ss);  | 
| 4709 | 90  | 
fun typid (Type (id,_) ) = hd (Symbol.explode (Sign.base_name id))  | 
91  | 
| typid (TFree (id,_) ) = hd (strip (tl (Symbol.explode (Sign.base_name id))))  | 
|
92  | 
| typid (TVar ((id,_),_)) = hd (tl (Symbol.explode (Sign.base_name id)));  | 
|
| 1274 | 93  | 
fun cons cons' = (map (fn (con,syn,args) =>  | 
| 4122 | 94  | 
((Syntax.const_name con syn),  | 
| 4008 | 95  | 
ListPair.map (fn ((lazy,sel,tp),vn) => ((lazy,  | 
| 
4188
 
1025a27b08f9
changed libraray function find to find_index_eq, currying it
 
oheimb 
parents: 
4122 
diff
changeset
 | 
96  | 
find_index_eq tp dts),  | 
| 
1637
 
b8a8ae2e5de1
Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
 
oheimb 
parents: 
1461 
diff
changeset
 | 
97  | 
sel,vn))  | 
| 4122 | 98  | 
(args,(mk_var_names(map (typid o third) args)))  | 
| 
1637
 
b8a8ae2e5de1
Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
 
oheimb 
parents: 
1461 
diff
changeset
 | 
99  | 
)) cons') : cons list;  | 
| 1274 | 100  | 
val eqs = map (fn (dtnvs,cons') => (dtnvs,cons cons')) eqs' : eq list;  | 
| 4122 | 101  | 
val thy = thy' |> Domain_Axioms.add_axioms (comp_dnam,eqs);  | 
| 4043 | 102  | 
in (foldl (fn (thy0,eq) => Domain_Theorems.theorems (eq,eqs) thy0) (thy,eqs))  | 
103  | 
|> Domain_Theorems.comp_theorems (comp_dnam, eqs) end;  | 
|
| 1274 | 104  | 
|
105  | 
end (* local *)  | 
|
106  | 
end (* struct *)  |