1274
|
1 |
(* library.ML
|
|
2 |
ID: $Id$
|
|
3 |
Author: David von Oheimb
|
|
4 |
Created: 18-Jul-95 extracted from syntax.ML, axioms.ML, extender.ML, interface.ML
|
|
5 |
Updated: 30-Aug-95
|
|
6 |
Copyright 1995 TU Muenchen
|
|
7 |
*)
|
|
8 |
|
|
9 |
(* ----- general support ---------------------------------------------------------- *)
|
|
10 |
|
|
11 |
fun Id x = x;
|
|
12 |
|
|
13 |
fun mapn f n [] = []
|
|
14 |
| mapn f n (x::xs) = (f n x) :: mapn f (n+1) xs;
|
|
15 |
|
|
16 |
fun foldr'' f (l,f2) =
|
|
17 |
let fun itr [] = raise LIST "foldr''"
|
|
18 |
| itr [a] = f2 a
|
|
19 |
| itr (a::l) = f(a, itr l)
|
|
20 |
in itr l end;
|
|
21 |
fun foldr' f l = foldr'' f (l,Id);
|
|
22 |
fun map_cumulr f start xs = foldr (fn (x,(ys,res)) => case f(x,res) of (y,res2) =>
|
1461
|
23 |
(y::ys,res2)) (xs,([],start));
|
1274
|
24 |
|
|
25 |
|
|
26 |
fun first (x,_,_) = x; fun second (_,x,_) = x; fun third (_,_,x) = x;
|
|
27 |
fun upd_first f (x,y,z) = (f x, y, z);
|
|
28 |
fun upd_second f (x,y,z) = ( x, f y, z);
|
|
29 |
fun upd_third f (x,y,z) = ( x, y, f z);
|
|
30 |
|
|
31 |
(* fn : ('a -> 'a) -> ('a -> 'a) -> 'a -> 'b list -> int -> 'a *)
|
|
32 |
fun bin_branchr f1 f2 y is j = let
|
|
33 |
fun bb y 1 _ = y
|
|
34 |
| bb y _ 0 = f1 y
|
|
35 |
| bb y i j = if i=2 then (f2 y) else bb (f2 y) (i-1) (j-1)
|
|
36 |
in bb y (length is) j end;
|
|
37 |
|
|
38 |
fun atomize thm = case concl_of thm of
|
|
39 |
_ $ (Const("op &",_) $ _ $ _) => atomize (thm RS conjunct1) @
|
1461
|
40 |
atomize (thm RS conjunct2)
|
1274
|
41 |
| _ $ (Const("All" ,_) $ Abs (s,_,_)) => atomize (thm RS
|
1461
|
42 |
(read_instantiate [("x","?"^s)] spec))
|
|
43 |
| _ => [thm];
|
1274
|
44 |
|
|
45 |
(* ----- specific support for domain ---------------------------------------------- *)
|
|
46 |
|
|
47 |
structure Domain_Library = struct
|
|
48 |
|
|
49 |
exception Impossible of string;
|
|
50 |
fun Imposs msg = raise Impossible ("Domain:"^msg);
|
|
51 |
|
|
52 |
(* ----- name handling ----- *)
|
|
53 |
|
|
54 |
val strip_esc = let
|
|
55 |
fun strip ("'" :: c :: cs) = c :: strip cs
|
|
56 |
| strip ["'"] = []
|
|
57 |
| strip (c :: cs) = c :: strip cs
|
|
58 |
| strip [] = [];
|
|
59 |
in implode o strip o explode end;
|
|
60 |
|
|
61 |
fun extern_name con = case explode con of
|
1461
|
62 |
("o"::"p"::" "::rest) => implode rest
|
|
63 |
| _ => con;
|
1274
|
64 |
fun dis_name con = "is_"^ (extern_name con);
|
|
65 |
fun dis_name_ con = "is_"^ (strip_esc con);
|
|
66 |
|
|
67 |
(*make distinct names out of the type list,
|
|
68 |
forbidding "o", "x..","f..","P.." as names *)
|
|
69 |
(*a number string is added if necessary *)
|
|
70 |
fun mk_var_names types : string list = let
|
|
71 |
fun typid (Type (id,_) ) = hd (explode id)
|
|
72 |
| typid (TFree (id,_) ) = hd (tl (explode id))
|
|
73 |
| typid (TVar ((id,_),_)) = hd (tl (explode id));
|
|
74 |
fun nonreserved id = let val cs = explode id in
|
1461
|
75 |
if not(hd cs mem ["x","f","P"]) then id
|
|
76 |
else implode(chr(1+ord (hd cs))::tl cs) end;
|
1274
|
77 |
fun index_vnames(vn::vns,tab) =
|
|
78 |
(case assoc(tab,vn) of
|
|
79 |
None => if vn mem vns
|
|
80 |
then (vn^"1") :: index_vnames(vns,(vn,2)::tab)
|
|
81 |
else vn :: index_vnames(vns, tab)
|
|
82 |
| Some(i) => (vn^(string_of_int i)) :: index_vnames(vns,(vn,i+1)::tab))
|
|
83 |
| index_vnames([],tab) = [];
|
|
84 |
in index_vnames(map (nonreserved o typid) types,[("o",1)]) end;
|
|
85 |
|
|
86 |
fun type_name_vars (Type(name,typevars)) = (name,typevars)
|
|
87 |
| type_name_vars _ = Imposs "library:type_name_vars";
|
|
88 |
|
|
89 |
(* ----- support for type and mixfix expressions ----- *)
|
|
90 |
|
|
91 |
fun mk_tvar s = TFree("'"^s,["pcpo"]);
|
|
92 |
fun mk_typ t (S,T) = Type(t,[S,T]);
|
|
93 |
infixr 5 -->;
|
|
94 |
infixr 6 ~>; val op ~> = mk_typ "->";
|
|
95 |
val NoSyn' = ThyOps.Mixfix NoSyn;
|
|
96 |
|
|
97 |
(* ----- constructor list handling ----- *)
|
|
98 |
|
1461
|
99 |
type cons = (string * (* operator name of constr *)
|
|
100 |
((bool*int)* (* (lazy,recursive element or ~1) *)
|
|
101 |
string* (* selector name *)
|
|
102 |
string) (* argument name *)
|
|
103 |
list); (* argument list *)
|
|
104 |
type eq = (string * (* name of abstracted type *)
|
|
105 |
typ list) * (* arguments of abstracted type *)
|
|
106 |
cons list; (* represented type, as a constructor list *)
|
1274
|
107 |
|
|
108 |
val rec_of = snd o first;
|
|
109 |
val is_lazy = fst o first;
|
|
110 |
val sel_of = second;
|
|
111 |
val vname = third;
|
|
112 |
val upd_vname = upd_third;
|
|
113 |
fun is_rec arg = rec_of arg >=0;
|
|
114 |
fun is_nonlazy_rec arg = is_rec arg andalso not (is_lazy arg);
|
|
115 |
fun nonlazy args = map vname (filter_out is_lazy args);
|
|
116 |
fun is_one_con_one_arg p cons = length cons = 1 andalso let val args = snd(hd cons) in
|
1461
|
117 |
length args = 1 andalso p (hd args) end;
|
1274
|
118 |
|
|
119 |
(* ----- support for term expressions ----- *)
|
|
120 |
|
|
121 |
fun % s = Free(s,dummyT);
|
|
122 |
fun %# arg = %(vname arg);
|
|
123 |
fun %% s = Const(s,dummyT);
|
|
124 |
|
|
125 |
local open HOLogic in
|
|
126 |
val mk_trp = mk_Trueprop;
|
|
127 |
fun mk_conj (S,T) = conj $ S $ T;
|
|
128 |
fun mk_disj (S,T) = disj $ S $ T;
|
|
129 |
fun mk_imp (S,T) = imp $ S $ T;
|
|
130 |
fun mk_lam (x,T) = Abs(x,dummyT,T);
|
|
131 |
fun mk_all (x,P) = HOLogic.mk_all (x,dummyT,P);
|
|
132 |
local
|
1461
|
133 |
fun tf (Type (s,args)) = foldl (op $) (%s,map tf args)
|
|
134 |
| tf (TFree(s,_ )) = %s
|
|
135 |
| tf _ = Imposs "mk_constrainall";
|
1274
|
136 |
in
|
|
137 |
fun mk_constrain (typ,T) = %%"_constrain" $ T $ tf typ;
|
|
138 |
fun mk_constrainall (x,typ,P) = %%"All" $ (%%"_constrainAbs"$Abs(x,dummyT,P)$tf typ);
|
|
139 |
end;
|
1461
|
140 |
|
1274
|
141 |
fun mk_ex (x,P) = mk_exists (x,dummyT,P);
|
|
142 |
fun mk_not P = Const("not" ,boolT --> boolT) $ P;
|
|
143 |
end;
|
|
144 |
|
|
145 |
fun mk_All (x,P) = %%"all" $ mk_lam(x,P); (* meta universal quantification *)
|
|
146 |
|
|
147 |
infixr 0 ===>;fun S ===> T = Const("==>", dummyT) $ S $ T;
|
|
148 |
infixr 0 ==>;fun S ==> T = mk_trp S ===> mk_trp T;
|
|
149 |
infix 0 ==; fun S == T = Const("==", dummyT) $ S $ T;
|
|
150 |
infix 1 ===; fun S === T = Const("op =", dummyT) $ S $ T;
|
|
151 |
infix 1 ~=; fun S ~= T = mk_not (S === T);
|
|
152 |
infix 1 <<; fun S << T = Const("op <<", dummyT) $ S $ T;
|
|
153 |
infix 1 ~<<; fun S ~<< T = mk_not (S << T);
|
|
154 |
|
|
155 |
infix 9 ` ; fun f` x = %%"fapp" $ f $ x;
|
|
156 |
infix 9 `% ; fun f`% s = f` % s;
|
|
157 |
infix 9 `%%; fun f`%%s = f` %%s;
|
|
158 |
fun mk_cfapp (F,As) = foldl (op `) (F,As);
|
|
159 |
fun con_app2 con f args = mk_cfapp(%%con,map f args);
|
|
160 |
fun con_app con = con_app2 con %#;
|
|
161 |
fun if_rec arg f y = if is_rec arg then f (rec_of arg) else y;
|
|
162 |
fun app_rec_arg p arg = if_rec arg (fn n => fn x => (p n)`x) Id (%# arg);
|
|
163 |
val cproj = bin_branchr (fn S => %%"cfst"`S) (fn S => %%"csnd"`S);
|
|
164 |
val proj = bin_branchr (fn S => %%"fst" $S) (fn S => %%"snd" $S);
|
|
165 |
fun lift tfn = foldr (fn (x,t)=> (mk_trp(tfn x) ===> t));
|
|
166 |
|
|
167 |
fun /\ v T = %%"fabs" $ mk_lam(v,T);
|
|
168 |
fun /\# (arg,T) = /\ (vname arg) T;
|
|
169 |
infixr 9 oo; fun S oo T = %%"cfcomp"`S`T;
|
|
170 |
val UU = %%"UU";
|
|
171 |
fun strict f = f`UU === UU;
|
|
172 |
fun defined t = t ~= UU;
|
|
173 |
fun cpair (S,T) = %%"cpair"`S`T;
|
|
174 |
fun lift_defined f = lift (fn x => defined (f x));
|
|
175 |
fun bound_arg vns v = Bound(length vns-find(v,vns)-1);
|
|
176 |
|
|
177 |
fun cont_eta_contract (Const("fabs",TT) $ Abs(a,T,body)) =
|
|
178 |
(case cont_eta_contract body of
|
|
179 |
body' as (Const("fapp",Ta) $ f $ Bound 0) =>
|
1461
|
180 |
if not (0 mem loose_bnos f) then incr_boundvars ~1 f
|
|
181 |
else Const("fabs",TT) $ Abs(a,T,body')
|
1274
|
182 |
| body' => Const("fabs",TT) $ Abs(a,T,body'))
|
|
183 |
| cont_eta_contract(f$t) = cont_eta_contract f $ cont_eta_contract t
|
|
184 |
| cont_eta_contract t = t;
|
|
185 |
|
|
186 |
fun idx_name dnames s n = s ^ (if length dnames = 1 then "" else string_of_int n);
|
|
187 |
fun when_funs cons = if length cons = 1 then ["f"]
|
1461
|
188 |
else mapn (fn n => K("f"^(string_of_int n))) 1 cons;
|
1274
|
189 |
fun when_body cons funarg = let
|
1461
|
190 |
fun one_fun n (_,[] ) = /\ "dummy" (funarg(1,n))
|
|
191 |
| one_fun n (_,args) = let
|
|
192 |
val l2 = length args;
|
|
193 |
fun idxs m arg = (if is_lazy arg then fn x=> %%"lift"`%%"ID"`x
|
|
194 |
else Id) (Bound(l2-m));
|
|
195 |
in cont_eta_contract (foldr''
|
|
196 |
(fn (a,t) => %%"ssplit"`(/\# (a,t)))
|
|
197 |
(args,
|
|
198 |
fn a => /\# (a,(mk_cfapp(funarg (l2,n),mapn idxs 1 args))))
|
|
199 |
) end;
|
1274
|
200 |
in foldr' (fn (x,y)=> %%"sswhen"`x`y) (mapn one_fun 1 cons) end;
|
|
201 |
|
|
202 |
end; (* struct *)
|