src/HOL/datatype.ML
author clasohm
Thu, 30 Mar 1995 13:07:59 +0200
changeset 980 33e3054b2871
parent 964 5f690b184f91
child 1279 f59b4f9f2cdc
permissions -rw-r--r--
removed unnecessary parentheses from the generated rules
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     1
(* Title:       HOL/datatype.ML
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     2
   ID:          $Id$
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     3
   Author:      Max Breitling, Carsten Clasohm, Tobias Nipkow, Norbert Voelker
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     4
   Copyright 1995 TU Muenchen
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     5
*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     6
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     7
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     8
(*used for constructor parameters*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     9
datatype dt_type = dtVar of string |
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    10
  dtTyp of dt_type list * string |
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    11
  dtRek of dt_type list * string;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    12
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    13
structure Datatype =
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    14
struct
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    15
local 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    16
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    17
val mysort = sort;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    18
open ThyParse HOLogic;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    19
exception Impossible;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    20
exception RecError of string;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    21
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    22
val is_dtRek = (fn dtRek _ => true  |  _  => false);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    23
fun opt_parens s = if s = "" then "" else enclose "(" ")" s; 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    24
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    25
(* ----------------------------------------------------------------------- *)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    26
(* Derivation of the primrec combinator application from the equations     *)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    27
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    28
(* substitute fname(ls,xk,rs) by yk(ls,rs) in t for (xk,yk) in pairs  *) 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    29
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    30
fun subst_apps (_,_) [] t = t
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    31
  | subst_apps (fname,rpos) pairs t =
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    32
    let 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    33
    fun subst (Abs(a,T,t)) = Abs(a,T,subst t)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    34
      | subst (funct $ body) = 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    35
	let val (f,b) = strip_comb (funct$body)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    36
	in 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    37
	  if is_Const f andalso fst(dest_Const f) = fname 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    38
	    then 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    39
	      let val (ls,rest) = (take(rpos,b), drop(rpos,b));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    40
		val (xk,rs) = (hd rest,tl rest)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    41
		  handle LIST _ => raise RecError "not enough arguments \
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    42
		   \ in recursive application on rhs"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    43
              in 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    44
		(case assoc (pairs,xk) of 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    45
		   None => raise RecError 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    46
		     ("illegal occurence of " ^ fname ^ " on rhs")
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    47
		 | Some(U) => list_comb(U,map subst (ls @ rs)))
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    48
	      end
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    49
	  else list_comb(f, map subst b)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    50
	end
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    51
      | subst(t) = t
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    52
    in subst t end;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    53
  
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    54
(* abstract rhs *)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    55
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    56
fun abst_rec (fname,rpos,tc,ls,cargs,rs,rhs) =       
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    57
  let val rargs = (map fst o 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    58
		   (filter (fn (a,T) => is_dtRek T))) (cargs ~~ tc);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    59
      val subs = map (fn (s,T) => (s,dummyT))
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    60
	           (rev(rename_wrt_term rhs rargs));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    61
      val subst_rhs = subst_apps (fname,rpos)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    62
	                (map Free rargs ~~ map Free subs) rhs;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    63
  in 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    64
      list_abs_free (cargs @ subs @ ls @ rs, subst_rhs) 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    65
  end;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    66
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    67
(* parsing the prim rec equations *)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    68
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    69
fun dest_eq ( Const("Trueprop",_) $ (Const ("op =",_) $ lhs $ rhs))
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    70
                 = (lhs, rhs)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    71
   | dest_eq _ = raise RecError "not a proper equation"; 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    72
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    73
fun dest_rec eq = 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    74
  let val (lhs,rhs) = dest_eq eq; 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    75
    val (name,args) = strip_comb lhs; 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    76
    val (ls',rest)  = take_prefix is_Free args; 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    77
    val (middle,rs') = take_suffix is_Free rest;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    78
    val rpos = length ls';
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    79
    val (c,cargs') = strip_comb (hd middle)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    80
      handle LIST "hd" => raise RecError "constructor missing";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    81
    val (ls,cargs,rs) = (map dest_Free ls', map dest_Free cargs'
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    82
			 , map dest_Free rs')
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    83
      handle TERM ("dest_Free",_) => 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    84
	  raise RecError "constructor has illegal argument in pattern";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    85
  in 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    86
    if length middle > 1 then 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    87
      raise RecError "more than one non-variable in pattern"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    88
    else if not(null(findrep (map fst (ls @ rs @ cargs)))) then 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    89
      raise RecError "repeated variable name in pattern" 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    90
	 else (fst(dest_Const name) handle TERM _ => 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    91
	       raise RecError "function is not declared as constant in theory"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    92
		 ,rpos,ls,fst( dest_Const c),cargs,rs,rhs)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    93
  end; 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    94
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    95
(* check function specified for all constructors and sort function terms *)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    96
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    97
fun check_and_sort (n,its) = 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    98
  if length its = n 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    99
    then map snd (mysort (fn ((i : int,_),(j,_)) => i<j) its)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   100
  else raise error "Primrec definition error:\n\
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   101
   \Please give an equation for every constructor";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   102
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   103
(* translate rec equations into function arguments suitable for rec comb *)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   104
(* theory parameter needed for printing error messages                   *) 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   105
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   106
fun trans_recs _ _ [] = error("No primrec equations.")
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   107
  | trans_recs thy cs' (eq1::eqs) = 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   108
    let val (name1,rpos1,ls1,_,_,_,_) = dest_rec eq1
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   109
      handle RecError s =>
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   110
	error("Primrec definition error: " ^ s ^ ":\n" 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   111
	      ^ "   " ^ Sign.string_of_term (sign_of thy) eq1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   112
      val tcs = map (fn (_,c,T,_,_) => (c,T)) cs';  
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   113
      val cs = map fst tcs;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   114
      fun trans_recs' _ [] = []
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   115
        | trans_recs' cis (eq::eqs) = 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   116
	  let val (name,rpos,ls,c,cargs,rs,rhs) = dest_rec eq; 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   117
	    val tc = assoc(tcs,c);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   118
	    val i = (1 + find (c,cs))  handle LIST "find" => 0; 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   119
	  in
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   120
	  if name <> name1 then 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   121
	    raise RecError "function names inconsistent"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   122
	  else if rpos <> rpos1 then 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   123
	    raise RecError "position of rec. argument inconsistent"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   124
	  else if i = 0 then 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   125
	    raise RecError "illegal argument in pattern" 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   126
	  else if i mem cis then
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   127
	    raise RecError "constructor already occured as pattern "
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   128
	       else (i,abst_rec (name,rpos,the tc,ls,cargs,rs,rhs))
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   129
		     :: trans_recs' (i::cis) eqs 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   130
	  end
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   131
	  handle RecError s =>
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   132
	        error("Primrec definition error\n" ^ s ^ "\n" 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   133
		      ^ "   " ^ Sign.string_of_term (sign_of thy) eq);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   134
    in (  name1, ls1
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   135
	, check_and_sort (length cs, trans_recs' [] (eq1::eqs)))
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   136
    end ;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   137
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   138
in
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   139
  fun add_datatype (typevars, tname, cons_list') thy = 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   140
    let
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   141
      fun typid(dtRek(_,id)) = id
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   142
        | typid(dtVar s) = implode (tl (explode s))
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   143
        | typid(dtTyp(_,id)) = id;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   144
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   145
      fun index_vnames(vn::vns,tab) =
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   146
            (case assoc(tab,vn) of
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   147
               None => if vn mem vns
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   148
                       then (vn^"1") :: index_vnames(vns,(vn,2)::tab)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   149
                       else vn :: index_vnames(vns,tab)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   150
             | Some(i) => (vn^(string_of_int i)) ::
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   151
                          index_vnames(vns,(vn,i+1)::tab))
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   152
        | index_vnames([],tab) = [];
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   153
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   154
      fun mk_var_names types = index_vnames(map typid types,[]);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   155
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   156
      (*search for free type variables and convert recursive *)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   157
      fun analyse_types (cons, types, syn) =
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   158
	let fun analyse(t as dtVar v) =
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   159
                  if t mem typevars then t
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   160
                  else error ("Free type variable " ^ v ^ " on rhs.")
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   161
	      | analyse(dtTyp(typl,s)) =
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   162
		  if tname <> s then dtTyp(analyses typl, s)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   163
                  else if typevars = typl then dtRek(typl, s)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   164
                       else error (s ^ " used in different ways")
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   165
	      | analyse(dtRek _) = raise Impossible
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   166
	    and analyses ts = map analyse ts;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   167
	in (cons, Syntax.const_name cons syn, analyses types,
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   168
            mk_var_names types, syn)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   169
        end;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   170
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   171
     (*test if all elements are recursive, i.e. if the type is empty*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   172
      
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   173
      fun non_empty (cs : ('a * 'b * dt_type list * 'c *'d) list) = 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   174
	not(forall (exists is_dtRek o #3) cs) orelse
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   175
	error("Empty datatype not allowed!");
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   176
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   177
      val cons_list = map analyse_types cons_list';
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   178
      val dummy = non_empty cons_list;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   179
      val num_of_cons = length cons_list;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   180
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   181
     (* Auxiliary functions to construct argument and equation lists *)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   182
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   183
     (*generate 'var_n, ..., var_m'*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   184
      fun Args(var, delim, n, m) = 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   185
	space_implode delim (map (fn n => var^string_of_int(n)) (n upto m));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   186
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   187
      fun C_exp name vns = name ^ opt_parens(space_implode ") (" vns);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   188
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   189
     (*Arg_eqs([x1,...,xn],[y1,...,yn]) = "x1 = y1 & ... & xn = yn" *)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   190
      fun arg_eqs vns vns' =
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   191
        let fun mkeq(x,x') = x ^ "=" ^ x'
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   192
        in space_implode " & " (map mkeq (vns~~vns')) end;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   193
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   194
     (*Pretty printers for type lists;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   195
       pp_typlist1: parentheses, pp_typlist2: brackets*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   196
      fun pp_typ (dtVar s) = s
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   197
        | pp_typ (dtTyp (typvars, id)) =
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   198
	  if null typvars then id else (pp_typlist1 typvars) ^ id
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   199
        | pp_typ (dtRek (typvars, id)) = (pp_typlist1 typvars) ^ id
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   200
      and
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   201
	pp_typlist' ts = commas (map pp_typ ts)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   202
      and
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   203
	pp_typlist1 ts = if null ts then "" else parens (pp_typlist' ts);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   204
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   205
      fun pp_typlist2 ts = if null ts then "" else brackets (pp_typlist' ts);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   206
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   207
     (* Generate syntax translation for case rules *)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   208
      fun calc_xrules c_nr y_nr ((_, name, _, vns, _) :: cs) = 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   209
	let val arity = length vns;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   210
	  val body  = "z" ^ string_of_int(c_nr);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   211
	  val args1 = if arity=0 then ""
980
33e3054b2871 removed unnecessary parentheses from the generated rules
clasohm
parents: 964
diff changeset
   212
		      else " " ^ Args ("y", " ", y_nr, y_nr+arity-1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   213
	  val args2 = if arity=0 then ""
964
5f690b184f91 fixed two severe bugs in calc_xrules and case_rule
clasohm
parents: 923
diff changeset
   214
		      else "(% " ^ Args ("y", " ", y_nr, y_nr+arity-1) 
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   215
			^ ". ";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   216
	  val (rest1,rest2) = 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   217
	    if null cs then ("","")
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   218
	    else let val (h1, h2) = calc_xrules (c_nr+1) (y_nr+arity) cs
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   219
	    in (" | " ^ h1, " " ^ h2) end;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   220
	in (name ^ args1 ^ " => " ^ body ^ rest1,
964
5f690b184f91 fixed two severe bugs in calc_xrules and case_rule
clasohm
parents: 923
diff changeset
   221
            args2 ^ body ^ (if args2 = "" then "" else ")") ^ rest2)
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   222
        end
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   223
        | calc_xrules _ _ [] = raise Impossible;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   224
      
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   225
      val xrules =
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   226
	let val (first_part, scnd_part) = calc_xrules 1 1 cons_list
980
33e3054b2871 removed unnecessary parentheses from the generated rules
clasohm
parents: 964
diff changeset
   227
	in [("logic", "case x of " ^ first_part) <->
33e3054b2871 removed unnecessary parentheses from the generated rules
clasohm
parents: 964
diff changeset
   228
	     ("logic", tname ^ "_case " ^ scnd_part ^ " x")]
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   229
	end;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   230
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   231
     (*type declarations for constructors*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   232
      fun const_type (id, _, typlist, _, syn) =
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   233
	(id,  
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   234
	 (if null typlist then "" else pp_typlist2 typlist ^ " => ") ^
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   235
	    pp_typlist1 typevars ^ tname, syn);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   236
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   237
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   238
      fun assumpt (dtRek _ :: ts, v :: vs ,found) =
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   239
	let val h = if found then ";P(" ^ v ^ ")" else "[| P(" ^ v ^ ")"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   240
	in h ^ (assumpt (ts, vs, true)) end
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   241
        | assumpt (t :: ts, v :: vs, found) = assumpt (ts, vs, found)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   242
      | assumpt ([], [], found) = if found then "|] ==>" else ""
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   243
        | assumpt _ = raise Impossible;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   244
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   245
      fun t_inducting ((_, name, types, vns, _) :: cs) =
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   246
	let
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   247
	  val h = if null types then " P(" ^ name ^ ")"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   248
		  else " !!" ^ (space_implode " " vns) ^ "." ^
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   249
		    (assumpt (types, vns, false)) ^
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   250
                    "P(" ^ C_exp name vns ^ ")";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   251
	  val rest = t_inducting cs;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   252
	in if rest = "" then h else h ^ "; " ^ rest end
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   253
        | t_inducting [] = "";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   254
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   255
      fun t_induct cl typ_name =
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   256
        "[|" ^ t_inducting cl ^ "|] ==> P(" ^ typ_name ^ ")";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   257
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   258
      fun gen_typlist typevar f ((_, _, ts, _, _) :: cs) =
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   259
	let val h = if (length ts) > 0
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   260
		      then pp_typlist2(f ts) ^ "=>"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   261
		    else ""
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   262
	in h ^ typevar ^  "," ^ (gen_typlist typevar f cs) end
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   263
        | gen_typlist _ _ [] = "";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   264
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   265
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   266
(* -------------------------------------------------------------------- *)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   267
(* The case constant and rules 	        				*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   268
 		
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   269
      val t_case = tname ^ "_case";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   270
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   271
      fun case_rule n (id, name, _, vns, _) =
980
33e3054b2871 removed unnecessary parentheses from the generated rules
clasohm
parents: 964
diff changeset
   272
	let val args = if vns = [] then "" else " " ^ space_implode " " vns
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   273
	in (t_case ^ "_" ^ id,
980
33e3054b2871 removed unnecessary parentheses from the generated rules
clasohm
parents: 964
diff changeset
   274
	    t_case ^ " " ^ Args("f", " ", 1, num_of_cons)
33e3054b2871 removed unnecessary parentheses from the generated rules
clasohm
parents: 964
diff changeset
   275
	    ^ " (" ^ name ^ args ^ ") = f"^string_of_int(n) ^ args)
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   276
	end
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   277
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   278
      fun case_rules n (c :: cs) = case_rule n c :: case_rules(n+1) cs
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   279
        | case_rules _ [] = [];
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   280
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   281
      val datatype_arity = length typevars;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   282
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   283
      val types = [(tname, datatype_arity, NoSyn)];
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   284
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   285
      val arities = 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   286
        let val term_list = replicate datatype_arity termS;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   287
        in [(tname, term_list, termS)] 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   288
	end;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   289
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   290
      val datatype_name = pp_typlist1 typevars ^ tname;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   291
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   292
      val new_tvar_name = variant (map (fn dtVar s => s) typevars) "'z";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   293
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   294
      val case_const =
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   295
	(t_case,
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   296
	 "[" ^ gen_typlist new_tvar_name I cons_list 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   297
	 ^  pp_typlist1 typevars ^ tname ^ "] =>" ^ new_tvar_name,
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   298
	 NoSyn);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   299
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   300
      val rules_case = case_rules 1 cons_list;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   301
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   302
(* -------------------------------------------------------------------- *)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   303
(* The prim-rec combinator						*) 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   304
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   305
      val t_rec = tname ^ "_rec"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   306
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   307
(* adding type variables for dtRek types to end of list of dt_types      *)   
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   308
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   309
      fun add_reks ts = 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   310
	ts @ map (fn _ => dtVar new_tvar_name) (filter is_dtRek ts); 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   311
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   312
(* positions of the dtRek types in a list of dt_types, starting from 1  *)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   313
      fun rek_vars ts vns = map snd (filter (is_dtRek o fst) (ts ~~ vns))
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   314
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   315
      fun rec_rule n (id,name,ts,vns,_) = 
964
5f690b184f91 fixed two severe bugs in calc_xrules and case_rule
clasohm
parents: 923
diff changeset
   316
	let val args = opt_parens(space_implode ") (" vns)
5f690b184f91 fixed two severe bugs in calc_xrules and case_rule
clasohm
parents: 923
diff changeset
   317
	  val fargs = opt_parens(Args("f", ") (", 1, num_of_cons))
5f690b184f91 fixed two severe bugs in calc_xrules and case_rule
clasohm
parents: 923
diff changeset
   318
	  fun rarg vn = t_rec ^ fargs ^ " (" ^ vn ^ ")"
5f690b184f91 fixed two severe bugs in calc_xrules and case_rule
clasohm
parents: 923
diff changeset
   319
	  val rargs = opt_parens(space_implode ") ("
5f690b184f91 fixed two severe bugs in calc_xrules and case_rule
clasohm
parents: 923
diff changeset
   320
                                 (map rarg (rek_vars ts vns)))
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   321
	in
964
5f690b184f91 fixed two severe bugs in calc_xrules and case_rule
clasohm
parents: 923
diff changeset
   322
	  (t_rec ^ "_" ^ id,
5f690b184f91 fixed two severe bugs in calc_xrules and case_rule
clasohm
parents: 923
diff changeset
   323
	   t_rec ^ fargs ^ " (" ^ name ^ args ^ ") = f"
5f690b184f91 fixed two severe bugs in calc_xrules and case_rule
clasohm
parents: 923
diff changeset
   324
	   ^ string_of_int(n) ^ args ^ rargs)
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   325
	end
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   326
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   327
      fun rec_rules n (c::cs) = rec_rule n c :: rec_rules (n+1) cs 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   328
	| rec_rules _ [] = [];
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   329
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   330
      val rec_const =
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   331
	(t_rec,
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   332
	 "[" ^ (gen_typlist new_tvar_name add_reks cons_list) 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   333
	 ^ (pp_typlist1 typevars) ^ tname ^ "] =>" ^ new_tvar_name,
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   334
	 NoSyn);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   335
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   336
      val rules_rec = rec_rules 1 cons_list
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   337
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   338
(* -------------------------------------------------------------------- *)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   339
      val consts = 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   340
	map const_type cons_list
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   341
	@ (if num_of_cons < dtK then []
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   342
	   else [(tname ^ "_ord", datatype_name ^ "=>nat", NoSyn)])
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   343
	@ [case_const,rec_const];
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   344
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   345
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   346
      fun Ci_ing ((id, name, _, vns, _) :: cs) =
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   347
	   if null vns then Ci_ing cs
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   348
	   else let val vns' = variantlist(vns,vns)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   349
                in ("inject_" ^ id,
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   350
		    "(" ^ (C_exp name vns) ^ "=" ^ (C_exp name vns')
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   351
		    ^ ") = (" ^ (arg_eqs vns vns') ^ ")") :: (Ci_ing cs)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   352
                end
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   353
	| Ci_ing [] = [];
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   354
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   355
      fun Ci_negOne (id1,name1,_,vns1,_) (id2,name2,_,vns2,_) =
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   356
            let val vns2' = variantlist(vns2,vns1)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   357
                val ax = C_exp name1 vns1 ^ "~=" ^ C_exp name2 vns2'
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   358
	in (id1 ^ "_not_" ^ id2, ax) end;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   359
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   360
      fun Ci_neg1 [] = []
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   361
	| Ci_neg1 (c1::cs) = (map (Ci_negOne c1) cs) @ Ci_neg1 cs;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   362
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   363
      fun suc_expr n = 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   364
	if n=0 then "0" else "Suc(" ^ suc_expr(n-1) ^ ")";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   365
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   366
      fun Ci_neg2() =
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   367
	let val ord_t = tname ^ "_ord";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   368
	  val cis = cons_list ~~ (0 upto (num_of_cons - 1))
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   369
	  fun Ci_neg2equals ((id, name, _, vns, _), n) =
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   370
	    let val ax = ord_t ^ "(" ^ (C_exp name vns) ^ ") = " ^ (suc_expr n)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   371
	    in (ord_t ^ "_" ^ id, ax) end
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   372
	in (ord_t ^ "_distinct", ord_t^"(x) ~= "^ord_t^"(y) ==> x ~= y") ::
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   373
	  (map Ci_neg2equals cis)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   374
	end;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   375
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   376
      val rules_distinct = if num_of_cons < dtK then Ci_neg1 cons_list
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   377
			   else Ci_neg2();
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   378
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   379
      val rules_inject = Ci_ing cons_list;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   380
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   381
      val rule_induct = (tname ^ "_induct", t_induct cons_list tname);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   382
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   383
      val rules = rule_induct ::
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   384
	(rules_inject @ rules_distinct @ rules_case @ rules_rec);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   385
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   386
      fun add_primrec eqns thy =
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   387
	let val rec_comb = Const(t_rec,dummyT)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   388
	  val teqns = map (fn neq => snd(read_axm (sign_of thy) neq)) eqns
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   389
	  val (fname,ls,fns) = trans_recs thy cons_list teqns
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   390
	  val rhs = 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   391
	    list_abs_free
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   392
	    (ls @ [(tname,dummyT)]
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   393
	     ,list_comb(rec_comb
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   394
			, fns @ map Bound (0 ::(length ls downto 1))));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   395
          val sg = sign_of thy;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   396
          val defpair =  mk_defpair (Const(fname,dummyT),rhs)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   397
	  val defpairT as (_, _ $ Const(_,T) $ _ ) = inferT_axm sg defpair;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   398
	  val varT = Type.varifyT T;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   399
          val ftyp = the (Sign.const_type sg fname);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   400
	in
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   401
	  if Type.typ_instance (#tsig(Sign.rep_sg sg), ftyp, varT)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   402
	  then add_defs_i [defpairT] thy
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   403
	  else error("Primrec definition error: \ntype of " ^ fname 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   404
		     ^ " is not instance of type deduced from equations")
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   405
	end;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   406
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   407
    in 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   408
      (thy
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   409
      |> add_types types
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   410
      |> add_arities arities
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   411
      |> add_consts consts
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   412
      |> add_trrules xrules
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   413
      |> add_axioms rules,add_primrec)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   414
    end
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   415
end
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   416
end
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   417
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   418
(*
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   419
Informal description of functions used in datatype.ML for the Isabelle/HOL
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   420
implementation of prim. rec. function definitions. (N. Voelker, Feb. 1995) 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   421
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   422
* subst_apps (fname,rpos) pairs t:
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   423
   substitute the term 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   424
       fname(ls,xk,rs) 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   425
   by 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   426
      yk(ls,rs) 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   427
   in t for (xk,yk) in pairs, where rpos = length ls. 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   428
   Applied with : 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   429
     fname = function name 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   430
     rpos = position of recursive argument 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   431
     pairs = list of pairs (xk,yk), where 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   432
          xk are the rec. arguments of the constructor in the pattern,
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   433
          yk is a variable with name derived from xk 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   434
     t = rhs of equation 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   435
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   436
* abst_rec (fname,rpos,tc,ls,cargs,rs,rhs)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   437
  - filter recursive arguments from constructor arguments cargs,
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   438
  - perform substitutions on rhs, 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   439
  - derive list subs of new variable names yk for use in subst_apps, 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   440
  - abstract rhs with respect to cargs, subs, ls and rs. 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   441
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   442
* dest_eq t 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   443
  destruct a term denoting an equation into lhs and rhs. 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   444
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   445
* dest_req eq 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   446
  destruct an equation of the form 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   447
      name (vl1..vlrpos, Ci(vi1..vin), vr1..vrn) = rhs
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   448
  into 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   449
  - function name  (name) 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   450
  - position of the first non-variable parameter  (rpos)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   451
  - the list of first rpos parameters (ls = [vl1..vlrpos]) 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   452
  - the constructor (fst( dest_Const c) = Ci)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   453
  - the arguments of the constructor (cargs = [vi1..vin])
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   454
  - the rest of the variables in the pattern (rs = [vr1..vrn])
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   455
  - the right hand side of the equation (rhs).  
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   456
 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   457
* check_and_sort (n,its)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   458
  check that  n = length its holds, and sort elements of its by 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   459
  first component. 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   460
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   461
* trans_recs thy cs' (eq1::eqs)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   462
  destruct eq1 into name1, rpos1, ls1, etc.. 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   463
  get constructor list with and without type (tcs resp. cs) from cs',  
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   464
  for every equation:  
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   465
    destruct it into (name,rpos,ls,c,cargs,rs,rhs)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   466
    get typed constructor tc from c and tcs 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   467
    determine the index i of the constructor 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   468
    check function name and position of rec. argument by comparison
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   469
    with first equation 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   470
    check for repeated variable names in pattern
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   471
    derive function term f_i which is used as argument of the rec. combinator
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   472
    sort the terms f_i according to i and return them together
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   473
      with the function name and the parameter of the definition (ls). 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   474
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   475
* Application:
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   476
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   477
  The rec. combinator is applied to the function terms resulting from
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   478
  trans_rec. This results in a function which takes the recursive arg. 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   479
  as first parameter and then the arguments corresponding to ls. The
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   480
  order of parameters is corrected by setting the rhs equal to 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   481
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   482
  list_abs_free
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   483
	    (ls @ [(tname,dummyT)]
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   484
	     ,list_comb(rec_comb
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   485
			, fns @ map Bound (0 ::(length ls downto 1))));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   486
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   487
  Note the de-Bruijn indices counting the number of lambdas between the
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   488
  variable and its binding. 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   489
*)