src/HOLCF/domain/library.ML
author wenzelm
Sat, 03 Nov 2001 01:41:26 +0100
changeset 12030 46d57d0290a2
parent 11531 d038246a62f2
child 14643 130076a81b84
permissions -rw-r--r--
GPLed;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
2276
3eb9a113029e modified file headers
oheimb
parents: 2238
diff changeset
     1
(*  Title:      HOLCF/domain/library.ML
2445
51993fea433f removed Holcfb.thy and Holcfb.ML, moving classical3 to HOL.ML as classical2
oheimb
parents: 2276
diff changeset
     2
    ID:         $Id$
12030
wenzelm
parents: 11531
diff changeset
     3
    Author:     David von Oheimb
wenzelm
parents: 11531
diff changeset
     4
    License:    GPL (GNU GENERAL PUBLIC LICENSE)
2445
51993fea433f removed Holcfb.thy and Holcfb.ML, moving classical3 to HOL.ML as classical2
oheimb
parents: 2276
diff changeset
     5
12030
wenzelm
parents: 11531
diff changeset
     6
Library for domain section.
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
     7
*)
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
     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
(* ----- general support ---------------------------------------------------- *)
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    11
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    12
fun Id x = x;
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    13
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    14
fun mapn f n []      = []
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    15
|   mapn f n (x::xs) = (f n x) :: mapn f (n+1) xs;
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    16
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1461
diff changeset
    17
fun foldr'' f (l,f2) = let fun itr []  = raise LIST "foldr''"
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1461
diff changeset
    18
			     | itr [a] = f2 a
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1461
diff changeset
    19
			     | itr (a::l) = f(a, itr l)
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1461
diff changeset
    20
in  itr l  end;
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    21
fun foldr' f l = foldr'' f (l,Id);
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1461
diff changeset
    22
fun map_cumulr f start xs = foldr (fn (x,(ys,res))=>case f(x,res) of (y,res2) =>
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1461
diff changeset
    23
						  (y::ys,res2)) (xs,([],start));
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    24
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    25
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    26
fun first  (x,_,_) = x; fun second (_,x,_) = x; fun third  (_,_,x) = x;
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    27
fun upd_first  f (x,y,z) = (f x,   y,   z);
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    28
fun upd_second f (x,y,z) = (  x, f y,   z);
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    29
fun upd_third  f (x,y,z) = (  x,   y, f z);
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    30
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1461
diff changeset
    31
fun atomize thm = let val r_inst = read_instantiate;
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1461
diff changeset
    32
    fun at  thm = case concl_of thm of
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1461
diff changeset
    33
      _$(Const("op &",_)$_$_)       => at(thm RS conjunct1)@at(thm RS conjunct2)
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1461
diff changeset
    34
    | _$(Const("All" ,_)$Abs(s,_,_))=> at(thm RS (r_inst [("x","?"^s)] spec))
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1461
diff changeset
    35
    | _				    => [thm];
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1461
diff changeset
    36
in map zero_var_indexes (at thm) end;
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    37
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1461
diff changeset
    38
(* ----- specific support for domain ---------------------------------------- *)
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    39
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    40
structure Domain_Library = struct
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    41
4122
f63c283cefaf * removed "axioms" and "generated by" section
oheimb
parents: 4043
diff changeset
    42
open HOLCFLogic;
f63c283cefaf * removed "axioms" and "generated by" section
oheimb
parents: 4043
diff changeset
    43
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    44
exception Impossible of string;
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    45
fun Imposs msg = raise Impossible ("Domain:"^msg);
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    46
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    47
(* ----- name handling ----- *)
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    48
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1461
diff changeset
    49
val strip_esc = let fun strip ("'" :: c :: cs) = c :: strip cs
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1461
diff changeset
    50
		    |   strip ["'"] = []
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1461
diff changeset
    51
		    |   strip (c :: cs) = c :: strip cs
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1461
diff changeset
    52
		    |   strip [] = [];
4709
d24168720303 Symbol.explode;
wenzelm
parents: 4188
diff changeset
    53
in implode o strip o Symbol.explode end;
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    54
4709
d24168720303 Symbol.explode;
wenzelm
parents: 4188
diff changeset
    55
fun extern_name con = case Symbol.explode con of 
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1461
diff changeset
    56
		   ("o"::"p"::" "::rest) => implode rest
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1461
diff changeset
    57
		   | _ => con;
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    58
fun dis_name  con = "is_"^ (extern_name con);
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    59
fun dis_name_ con = "is_"^ (strip_esc   con);
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    60
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1461
diff changeset
    61
(* make distinct names out of the type list, 
5292
92ea5ff34c79 minor correction: n must not be used as free variable
oheimb
parents: 5291
diff changeset
    62
   forbidding "o","n..","x..","f..","P.." as names *)
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1461
diff changeset
    63
(* a number string is added if necessary *)
4122
f63c283cefaf * removed "axioms" and "generated by" section
oheimb
parents: 4043
diff changeset
    64
fun mk_var_names ids : string list = let
5292
92ea5ff34c79 minor correction: n must not be used as free variable
oheimb
parents: 5291
diff changeset
    65
    fun nonreserved s = if s mem ["n","x","f","P"] then s^"'" else s;
1819
245721624c8d minor fix of index_vnames: name "O" is occupied (used in HOL/Relation.thy)
oheimb
parents: 1637
diff changeset
    66
    fun index_vnames(vn::vns,occupied) =
245721624c8d minor fix of index_vnames: name "O" is occupied (used in HOL/Relation.thy)
oheimb
parents: 1637
diff changeset
    67
          (case assoc(occupied,vn) of
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    68
             None => if vn mem vns
1819
245721624c8d minor fix of index_vnames: name "O" is occupied (used in HOL/Relation.thy)
oheimb
parents: 1637
diff changeset
    69
                     then (vn^"1") :: index_vnames(vns,(vn,1)  ::occupied)
245721624c8d minor fix of index_vnames: name "O" is occupied (used in HOL/Relation.thy)
oheimb
parents: 1637
diff changeset
    70
                     else  vn      :: index_vnames(vns,          occupied)
245721624c8d minor fix of index_vnames: name "O" is occupied (used in HOL/Relation.thy)
oheimb
parents: 1637
diff changeset
    71
           | Some(i) => (vn^(string_of_int (i+1)))
245721624c8d minor fix of index_vnames: name "O" is occupied (used in HOL/Relation.thy)
oheimb
parents: 1637
diff changeset
    72
				   :: index_vnames(vns,(vn,i+1)::occupied))
245721624c8d minor fix of index_vnames: name "O" is occupied (used in HOL/Relation.thy)
oheimb
parents: 1637
diff changeset
    73
      | index_vnames([],occupied) = [];
4122
f63c283cefaf * removed "axioms" and "generated by" section
oheimb
parents: 4043
diff changeset
    74
in index_vnames(map nonreserved ids, [("O",0),("o",0)]) end;
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    75
2445
51993fea433f removed Holcfb.thy and Holcfb.ML, moving classical3 to HOL.ML as classical2
oheimb
parents: 2276
diff changeset
    76
fun rep_Type  (Type  x) = x | rep_Type  _ = Imposs "library:rep_Type";
51993fea433f removed Holcfb.thy and Holcfb.ML, moving classical3 to HOL.ML as classical2
oheimb
parents: 2276
diff changeset
    77
fun rep_TFree (TFree x) = x | rep_TFree _ = Imposs "library:rep_TFree";
4008
2444085532c6 adapted domain and ax_ops package for name spaces
oheimb
parents: 3428
diff changeset
    78
val tsig_of = #tsig o Sign.rep_sg;
4122
f63c283cefaf * removed "axioms" and "generated by" section
oheimb
parents: 4043
diff changeset
    79
7652
2db14b7298c6 Sign.of_sort;
wenzelm
parents: 5439
diff changeset
    80
fun pcpo_type sg t = Sign.of_sort sg (Sign.certify_typ sg t, pcpoS);
4008
2444085532c6 adapted domain and ax_ops package for name spaces
oheimb
parents: 3428
diff changeset
    81
fun string_of_typ sg = Sign.string_of_typ sg o Sign.certify_typ sg;
2444085532c6 adapted domain and ax_ops package for name spaces
oheimb
parents: 3428
diff changeset
    82
fun str2typ sg = typ_of o read_ctyp sg;
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    83
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    84
(* ----- constructor list handling ----- *)
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    85
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1461
diff changeset
    86
type cons = (string *			(* operator name of constr *)
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1461
diff changeset
    87
	    ((bool*int)*		(*  (lazy,recursive element or ~1) *)
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1461
diff changeset
    88
	      string*			(*   selector name    *)
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1461
diff changeset
    89
	      string)			(*   argument name    *)
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1461
diff changeset
    90
	    list);			(* argument list      *)
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1461
diff changeset
    91
type eq = (string *		(* name      of abstracted type *)
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1461
diff changeset
    92
	   typ list) *		(* arguments of abstracted type *)
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1461
diff changeset
    93
	  cons list;		(* represented type, as a constructor list *)
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    94
2238
c72a23bbe762 Eta-expanded some declarations that are illegal under value polymorphism
paulson
parents: 1834
diff changeset
    95
fun rec_of arg  = snd (first arg);
c72a23bbe762 Eta-expanded some declarations that are illegal under value polymorphism
paulson
parents: 1834
diff changeset
    96
fun is_lazy arg = fst (first arg);
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    97
val sel_of    =       second;
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    98
val     vname =       third;
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    99
val upd_vname =   upd_third;
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   100
fun is_rec         arg = rec_of arg >=0;
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   101
fun is_nonlazy_rec arg = is_rec arg andalso not (is_lazy arg);
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1461
diff changeset
   102
fun nonlazy     args   = map vname (filter_out is_lazy    args);
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1461
diff changeset
   103
fun nonlazy_rec args   = map vname (filter is_nonlazy_rec args);
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1461
diff changeset
   104
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1461
diff changeset
   105
(* ----- support for type and mixfix expressions ----- *)
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1461
diff changeset
   106
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1461
diff changeset
   107
infixr 5 -->;
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   108
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   109
(* ----- support for term expressions ----- *)
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   110
11531
d038246a62f2 Renamed functions % and %% to avoid clash with syntax for proof terms.
berghofe
parents: 7652
diff changeset
   111
fun %: s = Free(s,dummyT);
d038246a62f2 Renamed functions % and %% to avoid clash with syntax for proof terms.
berghofe
parents: 7652
diff changeset
   112
fun %# arg = %:(vname arg);
d038246a62f2 Renamed functions % and %% to avoid clash with syntax for proof terms.
berghofe
parents: 7652
diff changeset
   113
fun %%: s = Const(s,dummyT);
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   114
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   115
local open HOLogic in
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   116
val mk_trp = mk_Trueprop;
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   117
fun mk_conj (S,T) = conj $ S $ T;
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   118
fun mk_disj (S,T) = disj $ S $ T;
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   119
fun mk_imp  (S,T) = imp  $ S $ T;
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   120
fun mk_lam  (x,T) = Abs(x,dummyT,T);
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   121
fun mk_all  (x,P) = HOLogic.mk_all (x,dummyT,P);
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   122
local 
11531
d038246a62f2 Renamed functions % and %% to avoid clash with syntax for proof terms.
berghofe
parents: 7652
diff changeset
   123
		    fun sg [s]     = %:s
d038246a62f2 Renamed functions % and %% to avoid clash with syntax for proof terms.
berghofe
parents: 7652
diff changeset
   124
		    |   sg (s::ss) = %%:"_classes" $ %:s $ sg ss
2445
51993fea433f removed Holcfb.thy and Holcfb.ML, moving classical3 to HOL.ML as classical2
oheimb
parents: 2276
diff changeset
   125
	 	    |   sg _       = Imposs "library:sg";
11531
d038246a62f2 Renamed functions % and %% to avoid clash with syntax for proof terms.
berghofe
parents: 7652
diff changeset
   126
		    fun sf [] = %%:"_emptysort"
d038246a62f2 Renamed functions % and %% to avoid clash with syntax for proof terms.
berghofe
parents: 7652
diff changeset
   127
		    |   sf s  = %%:"_sort" $ sg s
d038246a62f2 Renamed functions % and %% to avoid clash with syntax for proof terms.
berghofe
parents: 7652
diff changeset
   128
		    fun tf (Type (s,args)) = foldl (op $) (%:s,map tf args)
d038246a62f2 Renamed functions % and %% to avoid clash with syntax for proof terms.
berghofe
parents: 7652
diff changeset
   129
		    |   tf (TFree(s,sort)) = %%:"_ofsort" $ %:s $ sf sort
2445
51993fea433f removed Holcfb.thy and Holcfb.ML, moving classical3 to HOL.ML as classical2
oheimb
parents: 2276
diff changeset
   130
		    |   tf _               = Imposs "library:tf";
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   131
in
11531
d038246a62f2 Renamed functions % and %% to avoid clash with syntax for proof terms.
berghofe
parents: 7652
diff changeset
   132
fun mk_constrain      (typ,T) = %%:"_constrain" $ T $ tf typ;
d038246a62f2 Renamed functions % and %% to avoid clash with syntax for proof terms.
berghofe
parents: 7652
diff changeset
   133
fun mk_constrainall (x,typ,P) = %%:"All" $ (%%:"_constrainAbs" $ mk_lam(x,P) $ tf typ);
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   134
end;
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   135
fun mk_ex   (x,P) = mk_exists (x,dummyT,P);
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   136
end;
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   137
11531
d038246a62f2 Renamed functions % and %% to avoid clash with syntax for proof terms.
berghofe
parents: 7652
diff changeset
   138
fun mk_All  (x,P) = %%:"all" $ mk_lam(x,P); (* meta universal quantification *)
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   139
11531
d038246a62f2 Renamed functions % and %% to avoid clash with syntax for proof terms.
berghofe
parents: 7652
diff changeset
   140
infixr 0 ===>;fun S ===> T = %%:"==>" $ S $ T;
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   141
infixr 0 ==>;fun S ==> T = mk_trp S ===> mk_trp T;
11531
d038246a62f2 Renamed functions % and %% to avoid clash with syntax for proof terms.
berghofe
parents: 7652
diff changeset
   142
infix 0 ==;  fun S ==  T = %%:"==" $ S $ T;
d038246a62f2 Renamed functions % and %% to avoid clash with syntax for proof terms.
berghofe
parents: 7652
diff changeset
   143
infix 1 ===; fun S === T = %%:"op =" $ S $ T;
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   144
infix 1 ~=;  fun S ~=  T = mk_not (S === T);
11531
d038246a62f2 Renamed functions % and %% to avoid clash with syntax for proof terms.
berghofe
parents: 7652
diff changeset
   145
infix 1 <<;  fun S <<  T = %%:"op <<" $ S $ T;
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   146
infix 1 ~<<; fun S ~<< T = mk_not (S << T);
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   147
11531
d038246a62f2 Renamed functions % and %% to avoid clash with syntax for proof terms.
berghofe
parents: 7652
diff changeset
   148
infix 9 `  ; fun f`  x = %%:"Rep_CFun" $ f $ x;
d038246a62f2 Renamed functions % and %% to avoid clash with syntax for proof terms.
berghofe
parents: 7652
diff changeset
   149
infix 9 `% ; fun f`% s = f` %: s;
d038246a62f2 Renamed functions % and %% to avoid clash with syntax for proof terms.
berghofe
parents: 7652
diff changeset
   150
infix 9 `%%; fun f`%%s = f` %%:s;
5291
5706f0ef1d43 eliminated fabs,fapp.
slotosch
parents: 4778
diff changeset
   151
fun mk_cRep_CFun (F,As) = foldl (op `) (F,As);
11531
d038246a62f2 Renamed functions % and %% to avoid clash with syntax for proof terms.
berghofe
parents: 7652
diff changeset
   152
fun con_app2 con f args = mk_cRep_CFun(%%:con,map f args);
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   153
fun con_app con = con_app2 con %#;
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   154
fun if_rec  arg f y   = if is_rec arg then f (rec_of arg) else y;
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   155
fun app_rec_arg p arg = if_rec arg (fn n => fn x => (p n)`x) Id (%# arg);
4754
2c090aef2ca2 added cproj', and therefore extended prj
oheimb
parents: 4709
diff changeset
   156
fun prj _  _  x (   _::[]) _ = x
2c090aef2ca2 added cproj', and therefore extended prj
oheimb
parents: 4709
diff changeset
   157
|   prj f1 _  x (_::y::ys) 0 = f1 x y
2c090aef2ca2 added cproj', and therefore extended prj
oheimb
parents: 4709
diff changeset
   158
|   prj f1 f2 x (y::   ys) j = prj f1 f2 (f2 x y) ys (j-1);
2c090aef2ca2 added cproj', and therefore extended prj
oheimb
parents: 4709
diff changeset
   159
fun fix_tp (tn, args) = (tn, map (K oneT) args); (* instantiate type to
2c090aef2ca2 added cproj', and therefore extended prj
oheimb
parents: 4709
diff changeset
   160
						    avoid type varaibles *)
11531
d038246a62f2 Renamed functions % and %% to avoid clash with syntax for proof terms.
berghofe
parents: 7652
diff changeset
   161
fun  proj x      = prj (fn S => K(%%:"fst" $S)) (fn S => K(%%:"snd" $S)) x;
d038246a62f2 Renamed functions % and %% to avoid clash with syntax for proof terms.
berghofe
parents: 7652
diff changeset
   162
fun cproj x      = prj (fn S => K(%%:"cfst"`S)) (fn S => K(%%:"csnd"`S)) x;
4754
2c090aef2ca2 added cproj', and therefore extended prj
oheimb
parents: 4709
diff changeset
   163
fun cproj' T eqs = prj 
2c090aef2ca2 added cproj', and therefore extended prj
oheimb
parents: 4709
diff changeset
   164
	(fn S => fn t => Const("cfst",mk_prodT(dummyT,t)->>dummyT)`S) 
2c090aef2ca2 added cproj', and therefore extended prj
oheimb
parents: 4709
diff changeset
   165
	(fn S => fn t => Const("csnd",mk_prodT(t,dummyT)->>dummyT)`S) 
2c090aef2ca2 added cproj', and therefore extended prj
oheimb
parents: 4709
diff changeset
   166
		       T (map ((fn tp => tp ->> tp) o Type o fix_tp o fst) eqs);
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   167
fun lift tfn = foldr (fn (x,t)=> (mk_trp(tfn x) ===> t));
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   168
11531
d038246a62f2 Renamed functions % and %% to avoid clash with syntax for proof terms.
berghofe
parents: 7652
diff changeset
   169
fun /\ v T = %%:"Abs_CFun" $ mk_lam(v,T);
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   170
fun /\# (arg,T) = /\ (vname arg) T;
11531
d038246a62f2 Renamed functions % and %% to avoid clash with syntax for proof terms.
berghofe
parents: 7652
diff changeset
   171
infixr 9 oo; fun S oo T = %%:"cfcomp"`S`T;
d038246a62f2 Renamed functions % and %% to avoid clash with syntax for proof terms.
berghofe
parents: 7652
diff changeset
   172
val UU = %%:"UU";
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   173
fun strict f = f`UU === UU;
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   174
fun defined t = t ~= UU;
11531
d038246a62f2 Renamed functions % and %% to avoid clash with syntax for proof terms.
berghofe
parents: 7652
diff changeset
   175
fun cpair (S,T) = %%:"cpair"`S`T;
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   176
fun lift_defined f = lift (fn x => defined (f x));
4188
1025a27b08f9 changed libraray function find to find_index_eq, currying it
oheimb
parents: 4122
diff changeset
   177
fun bound_arg vns v = Bound(length vns -find_index_eq v vns -1);
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   178
5291
5706f0ef1d43 eliminated fabs,fapp.
slotosch
parents: 4778
diff changeset
   179
fun cont_eta_contract (Const("Abs_CFun",TT) $ Abs(a,T,body)) = 
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   180
      (case cont_eta_contract body  of
5291
5706f0ef1d43 eliminated fabs,fapp.
slotosch
parents: 4778
diff changeset
   181
        body' as (Const("Rep_CFun",Ta) $ f $ Bound 0) => 
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1461
diff changeset
   182
	  if not (0 mem loose_bnos f) then incr_boundvars ~1 f 
5291
5706f0ef1d43 eliminated fabs,fapp.
slotosch
parents: 4778
diff changeset
   183
	  else   Const("Abs_CFun",TT) $ Abs(a,T,body')
5706f0ef1d43 eliminated fabs,fapp.
slotosch
parents: 4778
diff changeset
   184
      | body' => Const("Abs_CFun",TT) $ Abs(a,T,body'))
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   185
|   cont_eta_contract(f$t) = cont_eta_contract f $ cont_eta_contract t
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   186
|   cont_eta_contract t    = t;
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   187
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1461
diff changeset
   188
fun idx_name dnames s n = s^(if length dnames = 1 then "" else string_of_int n);
1834
c780a4f39454 re-added when_funs to library.ML
oheimb
parents: 1828
diff changeset
   189
fun when_funs cons = if length cons = 1 then ["f"] 
c780a4f39454 re-added when_funs to library.ML
oheimb
parents: 1828
diff changeset
   190
                     else mapn (fn n => K("f"^(string_of_int n))) 1 cons;
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   191
fun when_body cons funarg = let
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1461
diff changeset
   192
	fun one_fun n (_,[]  ) = /\ "dummy" (funarg(1,n))
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1461
diff changeset
   193
	|   one_fun n (_,args) = let
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1461
diff changeset
   194
		val l2 = length args;
11531
d038246a62f2 Renamed functions % and %% to avoid clash with syntax for proof terms.
berghofe
parents: 7652
diff changeset
   195
		fun idxs m arg = (if is_lazy arg then fn x=> %%:"fup"`%%"ID"`x
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1461
diff changeset
   196
					         else Id) (Bound(l2-m));
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1461
diff changeset
   197
		in cont_eta_contract (foldr'' 
11531
d038246a62f2 Renamed functions % and %% to avoid clash with syntax for proof terms.
berghofe
parents: 7652
diff changeset
   198
			(fn (a,t) => %%:"ssplit"`(/\# (a,t)))
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1461
diff changeset
   199
			(args,
5291
5706f0ef1d43 eliminated fabs,fapp.
slotosch
parents: 4778
diff changeset
   200
			fn a=> /\#(a,(mk_cRep_CFun(funarg(l2,n),mapn idxs 1 args))))
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1461
diff changeset
   201
			) end;
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1461
diff changeset
   202
in (if length cons = 1 andalso length(snd(hd cons)) <= 1
11531
d038246a62f2 Renamed functions % and %% to avoid clash with syntax for proof terms.
berghofe
parents: 7652
diff changeset
   203
    then fn t => %%:"strictify"`t else Id)
d038246a62f2 Renamed functions % and %% to avoid clash with syntax for proof terms.
berghofe
parents: 7652
diff changeset
   204
     (foldr' (fn (x,y)=> %%:"sscase"`x`y) (mapn one_fun 1 cons)) end;
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   205
end; (* struct *)