src/HOLCF/Tools/domain/domain_library.ML
author huffman
Fri, 10 Apr 2009 11:35:21 -0700
changeset 30910 a7cc0ef93269
parent 30595 c87a3350f5a9
child 31006 644d18da3c77
permissions -rw-r--r--
set up domain package in Domain.thy
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
     1
(*  Title:      HOLCF/Tools/domain/domain_library.ML
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
     2
    Author:     David von Oheimb
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
     3
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
     4
Library for domain command.
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
     5
*)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
     6
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
     7
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
     8
(* ----- general support ---------------------------------------------------- *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
     9
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    10
fun mapn f n []      = []
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    11
|   mapn f n (x::xs) = (f n x) :: mapn f (n+1) xs;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    12
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    13
fun foldr'' f (l,f2) = let fun itr []  = raise Fail "foldr''"
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    14
			     | itr [a] = f2 a
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    15
			     | itr (a::l) = f(a, itr l)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    16
in  itr l  end;
30190
479806475f3c use long names for old-style fold combinators;
wenzelm
parents: 27239
diff changeset
    17
fun map_cumulr f start xs = List.foldr (fn (x,(ys,res))=>case f(x,res) of (y,res2) =>
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    18
						  (y::ys,res2)) ([],start) xs;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    19
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    20
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    21
fun first  (x,_,_) = x; fun second (_,x,_) = x; fun third  (_,_,x) = x;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    22
fun upd_first  f (x,y,z) = (f x,   y,   z);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    23
fun upd_second f (x,y,z) = (  x, f y,   z);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    24
fun upd_third  f (x,y,z) = (  x,   y, f z);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    25
27239
f2f42f9fa09d pervasive RuleInsts;
wenzelm
parents: 27231
diff changeset
    26
fun atomize ctxt thm = let val r_inst = read_instantiate ctxt;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    27
    fun at  thm = case concl_of thm of
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    28
      _$(Const("op &",_)$_$_)       => at(thm RS conjunct1)@at(thm RS conjunct2)
27231
ef7cc91a0d7f atomize: proper context;
wenzelm
parents: 27153
diff changeset
    29
    | _$(Const("All" ,_)$Abs(s,_,_))=> at(thm RS (r_inst [(("x", 0), "?" ^ s)] spec))
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    30
    | _				    => [thm];
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    31
in map zero_var_indexes (at thm) end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    32
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    33
(* ----- specific support for domain ---------------------------------------- *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    34
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    35
structure Domain_Library = struct
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    36
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    37
exception Impossible of string;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    38
fun Imposs msg = raise Impossible ("Domain:"^msg);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    39
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    40
(* ----- name handling ----- *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    41
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    42
val strip_esc = let fun strip ("'" :: c :: cs) = c :: strip cs
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    43
		    |   strip ["'"] = []
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    44
		    |   strip (c :: cs) = c :: strip cs
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    45
		    |   strip [] = [];
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    46
in implode o strip o Symbol.explode end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    47
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    48
fun extern_name con = case Symbol.explode con of 
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    49
		   ("o"::"p"::" "::rest) => implode rest
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    50
		   | _ => con;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    51
fun dis_name  con = "is_"^ (extern_name con);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    52
fun dis_name_ con = "is_"^ (strip_esc   con);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    53
fun mat_name  con = "match_"^ (extern_name con);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    54
fun mat_name_ con = "match_"^ (strip_esc   con);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    55
fun pat_name  con = (extern_name con) ^ "_pat";
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    56
fun pat_name_ con = (strip_esc   con) ^ "_pat";
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    57
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    58
(* make distinct names out of the type list, 
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    59
   forbidding "o","n..","x..","f..","P.." as names *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    60
(* a number string is added if necessary *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    61
fun mk_var_names ids : string list = let
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    62
    fun nonreserved s = if s mem ["n","x","f","P"] then s^"'" else s;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    63
    fun index_vnames(vn::vns,occupied) =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    64
          (case AList.lookup (op =) occupied vn of
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    65
             NONE => if vn mem vns
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    66
                     then (vn^"1") :: index_vnames(vns,(vn,1)  ::occupied)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    67
                     else  vn      :: index_vnames(vns,          occupied)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    68
           | SOME(i) => (vn^(string_of_int (i+1)))
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    69
				   :: index_vnames(vns,(vn,i+1)::occupied))
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    70
      | index_vnames([],occupied) = [];
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    71
in index_vnames(map nonreserved ids, [("O",0),("o",0)]) end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    72
30910
a7cc0ef93269 set up domain package in Domain.thy
huffman
parents: 30595
diff changeset
    73
fun pcpo_type sg t = Sign.of_sort sg (Sign.certify_typ sg t, @{sort pcpo});
26939
1035c89b4c02 moved global pretty/string_of functions from Sign to Syntax;
wenzelm
parents: 26012
diff changeset
    74
fun string_of_typ sg = Syntax.string_of_typ_global sg o Sign.certify_typ sg;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    75
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    76
(* ----- constructor list handling ----- *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    77
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    78
type cons = (string *				(* operator name of constr *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    79
	    ((bool*int*DatatypeAux.dtyp)*	(*  (lazy,recursive element or ~1) *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    80
	      string option*			(*   selector name    *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    81
	      string)				(*   argument name    *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    82
	    list);				(* argument list      *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    83
type eq = (string *		(* name      of abstracted type *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    84
	   typ list) *		(* arguments of abstracted type *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    85
	  cons list;		(* represented type, as a constructor list *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    86
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    87
fun rec_of arg  = second (first arg);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    88
fun is_lazy arg = first (first arg);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    89
val sel_of    =       second;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    90
val     vname =       third;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    91
val upd_vname =   upd_third;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    92
fun is_rec         arg = rec_of arg >=0;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    93
fun is_nonlazy_rec arg = is_rec arg andalso not (is_lazy arg);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    94
fun nonlazy     args   = map vname (filter_out is_lazy    args);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    95
fun nonlazy_rec args   = map vname (List.filter is_nonlazy_rec args);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    96
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    97
(* ----- support for type and mixfix expressions ----- *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    98
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    99
infixr 5 -->;
30910
a7cc0ef93269 set up domain package in Domain.thy
huffman
parents: 30595
diff changeset
   100
fun mk_uT T = Type(@{type_name "u"}, [T]);
a7cc0ef93269 set up domain package in Domain.thy
huffman
parents: 30595
diff changeset
   101
fun mk_cfunT (T, U) = Type(@{type_name "->"}, [T, U]);
a7cc0ef93269 set up domain package in Domain.thy
huffman
parents: 30595
diff changeset
   102
fun mk_sprodT (T, U) = Type(@{type_name "**"}, [T, U]);
a7cc0ef93269 set up domain package in Domain.thy
huffman
parents: 30595
diff changeset
   103
fun mk_ssumT (T, U) = Type(@{type_name "++"}, [T, U]);
a7cc0ef93269 set up domain package in Domain.thy
huffman
parents: 30595
diff changeset
   104
val oneT = @{typ one};
a7cc0ef93269 set up domain package in Domain.thy
huffman
parents: 30595
diff changeset
   105
val trT = @{typ tr};
a7cc0ef93269 set up domain package in Domain.thy
huffman
parents: 30595
diff changeset
   106
a7cc0ef93269 set up domain package in Domain.thy
huffman
parents: 30595
diff changeset
   107
infixr 6 ->>;
a7cc0ef93269 set up domain package in Domain.thy
huffman
parents: 30595
diff changeset
   108
val op ->> = mk_cfunT;
a7cc0ef93269 set up domain package in Domain.thy
huffman
parents: 30595
diff changeset
   109
a7cc0ef93269 set up domain package in Domain.thy
huffman
parents: 30595
diff changeset
   110
fun mk_TFree s = TFree ("'" ^ s, @{sort pcpo});
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   111
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   112
(* ----- support for term expressions ----- *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   113
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   114
fun %: s = Free(s,dummyT);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   115
fun %# arg = %:(vname arg);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   116
fun %%: s = Const(s,dummyT);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   117
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   118
local open HOLogic in
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   119
val mk_trp = mk_Trueprop;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   120
fun mk_conj (S,T) = conj $ S $ T;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   121
fun mk_disj (S,T) = disj $ S $ T;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   122
fun mk_imp  (S,T) = imp  $ S $ T;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   123
fun mk_lam  (x,T) = Abs(x,dummyT,T);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   124
fun mk_all  (x,P) = HOLogic.mk_all (x,dummyT,P);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   125
fun mk_ex   (x,P) = mk_exists (x,dummyT,P);
24680
0d355aa59e67 TypeInfer.constrain: canonical argument order;
wenzelm
parents: 23284
diff changeset
   126
val mk_constrain = uncurry TypeInfer.constrain;
0d355aa59e67 TypeInfer.constrain: canonical argument order;
wenzelm
parents: 23284
diff changeset
   127
fun mk_constrainall (x,typ,P) = %%:"All" $ (TypeInfer.constrain (typ --> boolT) (mk_lam(x,P)));
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   128
end
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   129
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   130
fun mk_All  (x,P) = %%:"all" $ mk_lam(x,P); (* meta universal quantification *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   131
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   132
infixr 0 ===>;  fun S ===> T = %%:"==>" $ S $ T;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   133
infixr 0 ==>;   fun S ==> T = mk_trp S ===> mk_trp T;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   134
infix 0 ==;     fun S ==  T = %%:"==" $ S $ T;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   135
infix 1 ===;    fun S === T = %%:"op =" $ S $ T;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   136
infix 1 ~=;     fun S ~=  T = HOLogic.mk_not (S === T);
30595
c87a3350f5a9 proper spacing before ML antiquotations -- note that @ may be part of symbolic ML identifiers;
wenzelm
parents: 30190
diff changeset
   137
infix 1 <<;     fun S <<  T = %%: @{const_name Porder.sq_le} $ S $ T;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   138
infix 1 ~<<;    fun S ~<< T = HOLogic.mk_not (S << T);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   139
30595
c87a3350f5a9 proper spacing before ML antiquotations -- note that @ may be part of symbolic ML identifiers;
wenzelm
parents: 30190
diff changeset
   140
infix 9 `  ; fun f ` x = %%: @{const_name Rep_CFun} $ f $ x;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   141
infix 9 `% ; fun f`% s = f` %: s;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   142
infix 9 `%%; fun f`%%s = f` %%:s;
26012
f6917792f8a4 new term-building combinators
huffman
parents: 25723
diff changeset
   143
30595
c87a3350f5a9 proper spacing before ML antiquotations -- note that @ may be part of symbolic ML identifiers;
wenzelm
parents: 30190
diff changeset
   144
fun mk_adm t = %%: @{const_name adm} $ t;
c87a3350f5a9 proper spacing before ML antiquotations -- note that @ may be part of symbolic ML identifiers;
wenzelm
parents: 30190
diff changeset
   145
fun mk_compact t = %%: @{const_name compact} $ t;
c87a3350f5a9 proper spacing before ML antiquotations -- note that @ may be part of symbolic ML identifiers;
wenzelm
parents: 30190
diff changeset
   146
val ID = %%: @{const_name ID};
c87a3350f5a9 proper spacing before ML antiquotations -- note that @ may be part of symbolic ML identifiers;
wenzelm
parents: 30190
diff changeset
   147
fun mk_strictify t = %%: @{const_name strictify}`t;
c87a3350f5a9 proper spacing before ML antiquotations -- note that @ may be part of symbolic ML identifiers;
wenzelm
parents: 30190
diff changeset
   148
fun mk_cfst t = %%: @{const_name cfst}`t;
c87a3350f5a9 proper spacing before ML antiquotations -- note that @ may be part of symbolic ML identifiers;
wenzelm
parents: 30190
diff changeset
   149
fun mk_csnd t = %%: @{const_name csnd}`t;
26012
f6917792f8a4 new term-building combinators
huffman
parents: 25723
diff changeset
   150
(*val csplitN    = "Cprod.csplit";*)
f6917792f8a4 new term-building combinators
huffman
parents: 25723
diff changeset
   151
(*val sfstN      = "Sprod.sfst";*)
f6917792f8a4 new term-building combinators
huffman
parents: 25723
diff changeset
   152
(*val ssndN      = "Sprod.ssnd";*)
30595
c87a3350f5a9 proper spacing before ML antiquotations -- note that @ may be part of symbolic ML identifiers;
wenzelm
parents: 30190
diff changeset
   153
fun mk_ssplit t = %%: @{const_name ssplit}`t;
c87a3350f5a9 proper spacing before ML antiquotations -- note that @ may be part of symbolic ML identifiers;
wenzelm
parents: 30190
diff changeset
   154
fun mk_sinl t = %%: @{const_name sinl}`t;
c87a3350f5a9 proper spacing before ML antiquotations -- note that @ may be part of symbolic ML identifiers;
wenzelm
parents: 30190
diff changeset
   155
fun mk_sinr t = %%: @{const_name sinr}`t;
c87a3350f5a9 proper spacing before ML antiquotations -- note that @ may be part of symbolic ML identifiers;
wenzelm
parents: 30190
diff changeset
   156
fun mk_sscase (x, y) = %%: @{const_name sscase}`x`y;
c87a3350f5a9 proper spacing before ML antiquotations -- note that @ may be part of symbolic ML identifiers;
wenzelm
parents: 30190
diff changeset
   157
fun mk_up t = %%: @{const_name up}`t;
c87a3350f5a9 proper spacing before ML antiquotations -- note that @ may be part of symbolic ML identifiers;
wenzelm
parents: 30190
diff changeset
   158
fun mk_fup (t,u) = %%: @{const_name fup} ` t ` u;
26012
f6917792f8a4 new term-building combinators
huffman
parents: 25723
diff changeset
   159
val ONE = @{term ONE};
f6917792f8a4 new term-building combinators
huffman
parents: 25723
diff changeset
   160
val TT = @{term TT};
f6917792f8a4 new term-building combinators
huffman
parents: 25723
diff changeset
   161
val FF = @{term FF};
30595
c87a3350f5a9 proper spacing before ML antiquotations -- note that @ may be part of symbolic ML identifiers;
wenzelm
parents: 30190
diff changeset
   162
fun mk_iterate (n,f,z) = %%: @{const_name iterate} $ n ` f ` z;
c87a3350f5a9 proper spacing before ML antiquotations -- note that @ may be part of symbolic ML identifiers;
wenzelm
parents: 30190
diff changeset
   163
fun mk_fix t = %%: @{const_name fix}`t;
c87a3350f5a9 proper spacing before ML antiquotations -- note that @ may be part of symbolic ML identifiers;
wenzelm
parents: 30190
diff changeset
   164
fun mk_return t = %%: @{const_name Fixrec.return}`t;
c87a3350f5a9 proper spacing before ML antiquotations -- note that @ may be part of symbolic ML identifiers;
wenzelm
parents: 30190
diff changeset
   165
val mk_fail = %%: @{const_name Fixrec.fail};
26012
f6917792f8a4 new term-building combinators
huffman
parents: 25723
diff changeset
   166
30595
c87a3350f5a9 proper spacing before ML antiquotations -- note that @ may be part of symbolic ML identifiers;
wenzelm
parents: 30190
diff changeset
   167
fun mk_branch t = %%: @{const_name Fixrec.branch} $ t;
26012
f6917792f8a4 new term-building combinators
huffman
parents: 25723
diff changeset
   168
f6917792f8a4 new term-building combinators
huffman
parents: 25723
diff changeset
   169
val pcpoS = @{sort pcpo};
f6917792f8a4 new term-building combinators
huffman
parents: 25723
diff changeset
   170
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   171
val list_ccomb = Library.foldl (op `); (* continuous version of list_comb *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   172
fun con_app2 con f args = list_ccomb(%%:con,map f args);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   173
fun con_app con = con_app2 con %#;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   174
fun if_rec  arg f y   = if is_rec arg then f (rec_of arg) else y;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   175
fun app_rec_arg p arg = if_rec arg (fn n => fn x => (p n)`x) I (%# arg);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   176
fun prj _  _  x (   _::[]) _ = x
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   177
|   prj f1 _  x (_::y::ys) 0 = f1 x y
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   178
|   prj f1 f2 x (y::   ys) j = prj f1 f2 (f2 x y) ys (j-1);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   179
fun  proj x      = prj (fn S => K(%%:"fst" $S)) (fn S => K(%%:"snd" $S)) x;
26012
f6917792f8a4 new term-building combinators
huffman
parents: 25723
diff changeset
   180
fun cproj x      = prj (fn S => K(mk_cfst S)) (fn S => K(mk_csnd S)) x;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   181
fun lift tfn = Library.foldr (fn (x,t)=> (mk_trp(tfn x) ===> t));
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   182
30595
c87a3350f5a9 proper spacing before ML antiquotations -- note that @ may be part of symbolic ML identifiers;
wenzelm
parents: 30190
diff changeset
   183
fun /\ v T = %%: @{const_name Abs_CFun} $ mk_lam(v,T);
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   184
fun /\# (arg,T) = /\ (vname arg) T;
30595
c87a3350f5a9 proper spacing before ML antiquotations -- note that @ may be part of symbolic ML identifiers;
wenzelm
parents: 30190
diff changeset
   185
infixr 9 oo; fun S oo T = %%: @{const_name cfcomp}`S`T;
c87a3350f5a9 proper spacing before ML antiquotations -- note that @ may be part of symbolic ML identifiers;
wenzelm
parents: 30190
diff changeset
   186
val UU = %%: @{const_name UU};
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   187
fun strict f = f`UU === UU;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   188
fun defined t = t ~= UU;
30595
c87a3350f5a9 proper spacing before ML antiquotations -- note that @ may be part of symbolic ML identifiers;
wenzelm
parents: 30190
diff changeset
   189
fun cpair (t,u) = %%: @{const_name cpair}`t`u;
c87a3350f5a9 proper spacing before ML antiquotations -- note that @ may be part of symbolic ML identifiers;
wenzelm
parents: 30190
diff changeset
   190
fun spair (t,u) = %%: @{const_name spair}`t`u;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   191
fun mk_ctuple [] = HOLogic.unit (* used in match_defs *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   192
|   mk_ctuple ts = foldr1 cpair ts;
26012
f6917792f8a4 new term-building combinators
huffman
parents: 25723
diff changeset
   193
fun mk_stuple [] = ONE
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   194
|   mk_stuple ts = foldr1 spair ts;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   195
fun mk_ctupleT [] = HOLogic.unitT   (* used in match_defs *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   196
|   mk_ctupleT Ts = foldr1 HOLogic.mk_prodT Ts;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   197
fun mk_maybeT T = Type ("Fixrec.maybe",[T]);
30595
c87a3350f5a9 proper spacing before ML antiquotations -- note that @ may be part of symbolic ML identifiers;
wenzelm
parents: 30190
diff changeset
   198
fun cpair_pat (p1,p2) = %%: @{const_name cpair_pat} $ p1 $ p2;
26012
f6917792f8a4 new term-building combinators
huffman
parents: 25723
diff changeset
   199
val mk_ctuple_pat = foldr1 cpair_pat;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   200
fun lift_defined f = lift (fn x => defined (f x));
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   201
fun bound_arg vns v = Bound(length vns -find_index_eq v vns -1);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   202
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   203
fun cont_eta_contract (Const("Cfun.Abs_CFun",TT) $ Abs(a,T,body)) = 
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   204
      (case cont_eta_contract body  of
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   205
        body' as (Const("Cfun.Rep_CFun",Ta) $ f $ Bound 0) => 
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   206
	  if not (0 mem loose_bnos f) then incr_boundvars ~1 f 
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   207
	  else   Const("Cfun.Abs_CFun",TT) $ Abs(a,T,body')
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   208
      | body' => Const("Cfun.Abs_CFun",TT) $ Abs(a,T,body'))
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   209
|   cont_eta_contract(f$t) = cont_eta_contract f $ cont_eta_contract t
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   210
|   cont_eta_contract t    = t;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   211
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   212
fun idx_name dnames s n = s^(if length dnames = 1 then "" else string_of_int n);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   213
fun when_funs cons = if length cons = 1 then ["f"] 
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   214
                     else mapn (fn n => K("f"^(string_of_int n))) 1 cons;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   215
fun when_body cons funarg = let
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   216
	fun one_fun n (_,[]  ) = /\ "dummy" (funarg(1,n))
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   217
	|   one_fun n (_,args) = let
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   218
		val l2 = length args;
26012
f6917792f8a4 new term-building combinators
huffman
parents: 25723
diff changeset
   219
		fun idxs m arg = (if is_lazy arg then (fn t => mk_fup (ID, t))
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   220
					         else I) (Bound(l2-m));
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   221
		in cont_eta_contract (foldr'' 
26012
f6917792f8a4 new term-building combinators
huffman
parents: 25723
diff changeset
   222
			(fn (a,t) => mk_ssplit (/\# (a,t)))
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   223
			(args,
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   224
			fn a=> /\#(a,(list_ccomb(funarg(l2,n),mapn idxs 1 args))))
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   225
			) end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   226
in (if length cons = 1 andalso length(snd(hd cons)) <= 1
26012
f6917792f8a4 new term-building combinators
huffman
parents: 25723
diff changeset
   227
    then mk_strictify else I)
f6917792f8a4 new term-building combinators
huffman
parents: 25723
diff changeset
   228
     (foldr1 mk_sscase (mapn one_fun 1 cons)) end;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   229
end; (* struct *)