Datatype.ML
author nipkow
Sat, 13 Aug 1994 16:33:53 +0200
changeset 101 5f99df1e26c4
parent 96 d94d0b324b4b
child 103 c57ab3ce997e
permissions -rw-r--r--
Added primitive recursive functions (Norbert Voelker's code) to the datatype package.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
80
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
     1
(*  Title:       HOL/Datatype
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
     2
    ID:          $Id$
101
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
     3
    Author:      Max Breitling / Carsten Clasohm /
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
     4
                 Norbert Voelker / Tobias Nipkow
53
5e0570ea8b70 ML-like datatype decalaration facility. Axiomatic.
nipkow
parents:
diff changeset
     5
    Copyright    1994 TU Muenchen
5e0570ea8b70 ML-like datatype decalaration facility. Axiomatic.
nipkow
parents:
diff changeset
     6
*)
5e0570ea8b70 ML-like datatype decalaration facility. Axiomatic.
nipkow
parents:
diff changeset
     7
5e0570ea8b70 ML-like datatype decalaration facility. Axiomatic.
nipkow
parents:
diff changeset
     8
80
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
     9
(*choice between Ci_neg1 and Ci_neg2 axioms depends on number of constructors*)
96
d94d0b324b4b Lots of simplifications
nipkow
parents: 92
diff changeset
    10
local
d94d0b324b4b Lots of simplifications
nipkow
parents: 92
diff changeset
    11
d94d0b324b4b Lots of simplifications
nipkow
parents: 92
diff changeset
    12
val dtK = 5
101
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
    13
val pars = enclose "(" ")";
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
    14
val brackets = enclose "[" "]";
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
    15
val mk_list = brackets o commas;
96
d94d0b324b4b Lots of simplifications
nipkow
parents: 92
diff changeset
    16
92
bcd0ee8d71aa Hidden dtK and Impossible with a "local" clause
nipkow
parents: 91
diff changeset
    17
in
80
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
    18
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
    19
local open ThyParse in
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
    20
val datatype_decls =
101
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
    21
  let val tvar = type_var >> (fn s => "dtVar" ^ s);
53
5e0570ea8b70 ML-like datatype decalaration facility. Axiomatic.
nipkow
parents:
diff changeset
    22
80
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
    23
      val type_var_list = 
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
    24
        tvar >> (fn s => [s]) || "(" $$-- list1 tvar --$$ ")";
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
    25
    
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
    26
      val typ =
96
d94d0b324b4b Lots of simplifications
nipkow
parents: 92
diff changeset
    27
         ident                  >> (fn s => "dtTyp([]," ^ quote s ^")")
80
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
    28
        ||
96
d94d0b324b4b Lots of simplifications
nipkow
parents: 92
diff changeset
    29
         type_var_list -- ident >> (fn (ts, id) => "dtTyp(" ^ mk_list ts ^
d94d0b324b4b Lots of simplifications
nipkow
parents: 92
diff changeset
    30
  				  "," ^ quote id ^ ")")
80
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
    31
        ||
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
    32
         tvar;
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
    33
    
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
    34
      val typ_list = "(" $$-- list1 typ --$$ ")" || empty;
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
    35
  
91
a94029edb01f added mixfix annotations to constructor declarations
clasohm
parents: 87
diff changeset
    36
      val cons = name -- typ_list -- opt_mixfix;
80
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
    37
  
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
    38
      fun constructs ts =
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
    39
        ( cons --$$ "|" -- constructs >> op::
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
    40
         ||
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
    41
          cons                        >> (fn c => [c])) ts;  
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
    42
  
91
a94029edb01f added mixfix annotations to constructor declarations
clasohm
parents: 87
diff changeset
    43
      val mk_cons = map (fn ((s, ts), syn) => 
a94029edb01f added mixfix annotations to constructor declarations
clasohm
parents: 87
diff changeset
    44
                           pars (commas [s, mk_list ts, syn]));
101
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
    45
80
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
    46
      (*remove all quotes from a string*)
96
d94d0b324b4b Lots of simplifications
nipkow
parents: 92
diff changeset
    47
      val rem_quotes = implode o filter (fn c => c <> "\"") o explode;
101
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
    48
96
d94d0b324b4b Lots of simplifications
nipkow
parents: 92
diff changeset
    49
      (*generate names of distinct axioms*)
d94d0b324b4b Lots of simplifications
nipkow
parents: 92
diff changeset
    50
      fun rules_distinct cs tname = 
d94d0b324b4b Lots of simplifications
nipkow
parents: 92
diff changeset
    51
        let val uqcs = map (fn ((s,_),_) => rem_quotes s) cs;
d94d0b324b4b Lots of simplifications
nipkow
parents: 92
diff changeset
    52
            (*combine all constructor names with all others w/o duplicates*)
d94d0b324b4b Lots of simplifications
nipkow
parents: 92
diff changeset
    53
            fun negOne c = map (fn c2 => quote (c ^ "_not_" ^ c2));
80
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
    54
            fun neg1 [] = []
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
    55
              | neg1 (c1 :: cs) = (negOne c1 cs) @ (neg1 cs)
96
d94d0b324b4b Lots of simplifications
nipkow
parents: 92
diff changeset
    56
        in if length uqcs < dtK then neg1 uqcs
d94d0b324b4b Lots of simplifications
nipkow
parents: 92
diff changeset
    57
           else quote (tname ^ "_ord_distinct") ::
d94d0b324b4b Lots of simplifications
nipkow
parents: 92
diff changeset
    58
                map (fn c => quote (tname ^ "_ord_" ^ c)) uqcs
80
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
    59
        end;
101
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
    60
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
    61
       fun rule_names tname cons pre =
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
    62
         mk_list (map (fn ((s,_),_) => quote(tname ^ pre ^ rem_quotes s)) cons)
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
    63
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
    64
       fun rules tname cons pre =
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
    65
         " map (get_axiom thy) " ^ rule_names tname cons pre       
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
    66
91
a94029edb01f added mixfix annotations to constructor declarations
clasohm
parents: 87
diff changeset
    67
      (*generate string for calling 'add_datatype'*)
80
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
    68
      fun mk_params ((ts, tname), cons) =
101
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
    69
       ("val (thy," ^ tname ^ "_add_primrec) =  add_datatype\n" ^
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
    70
       pars (commas [mk_list ts, quote tname, mk_list (mk_cons cons)]) ^
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
    71
       " thy\n\
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
    72
       \val thy=thy",
80
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
    73
       "structure " ^ tname ^ " =\n\
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
    74
       \struct\n\
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
    75
       \  val inject = map (get_axiom thy) " ^
91
a94029edb01f added mixfix annotations to constructor declarations
clasohm
parents: 87
diff changeset
    76
         mk_list (map (fn ((s,_), _) => quote ("inject_" ^ rem_quotes s)) 
96
d94d0b324b4b Lots of simplifications
nipkow
parents: 92
diff changeset
    77
                      (filter_out (null o snd o fst) cons)) ^ ";\n\
d94d0b324b4b Lots of simplifications
nipkow
parents: 92
diff changeset
    78
       \  val distinct = " ^ (if length cons < dtK then "let val distinct' = " else "")
d94d0b324b4b Lots of simplifications
nipkow
parents: 92
diff changeset
    79
         ^ "map (get_axiom thy) " ^ mk_list (rules_distinct cons tname) ^ 
80
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
    80
         (if length cons < dtK then 
96
d94d0b324b4b Lots of simplifications
nipkow
parents: 92
diff changeset
    81
           "  in distinct' @ (map (fn t => sym COMP (t RS contrapos)) distinct') end"
80
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
    82
          else "") ^ ";\n\
87
b0ea0e55dfe8 added prefix to name of induct axiom
clasohm
parents: 80
diff changeset
    83
       \  val induct = get_axiom thy \"" ^ tname ^ "_induct\";\n\
101
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
    84
       \  val cases =" ^ rules tname cons "_case_" ^ ";\n\
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
    85
       \  val recs =" ^ rules tname cons "_rec_" ^ ";\n\
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
    86
       \  val simps = inject @ distinct @ cases @ recs;\n\
96
d94d0b324b4b Lots of simplifications
nipkow
parents: 92
diff changeset
    87
       \  fun induct_tac a = res_inst_tac[(" ^ quote tname ^ ", a)]induct;\n\
101
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
    88
       \end;\n")
80
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
    89
  in (type_var_list || empty) -- ident --$$ "=" -- constructs >> mk_params end
101
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
    90
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
    91
val primrec_decl =
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
    92
  let fun mkstrings((fname,tname),axms) =
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
    93
        let fun prove (name,eqn) =
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
    94
             "val "^name^"= prove_goalw thy [get_axiom thy \""^fname^"_def\"] "
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
    95
                 ^ eqn ^"\n\
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
    96
             \(fn _ => [resolve_tac " ^ tname^".recs 1])"
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
    97
        in ("|> " ^ tname^"_add_primrec " ^ mk_list (map snd axms),
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
    98
            cat_lines(map prove axms))
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
    99
        end
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   100
  in ident -- ident -- repeat1 (ident -- string)  >> mkstrings end
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   101
53
5e0570ea8b70 ML-like datatype decalaration facility. Axiomatic.
nipkow
parents:
diff changeset
   102
end;
5e0570ea8b70 ML-like datatype decalaration facility. Axiomatic.
nipkow
parents:
diff changeset
   103
80
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   104
(*used for constructor parameters*)
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   105
datatype dt_type = dtVar of string |
96
d94d0b324b4b Lots of simplifications
nipkow
parents: 92
diff changeset
   106
                   dtTyp of dt_type list * string |
92
bcd0ee8d71aa Hidden dtK and Impossible with a "local" clause
nipkow
parents: 91
diff changeset
   107
                   dtRek of dt_type list * string;
53
5e0570ea8b70 ML-like datatype decalaration facility. Axiomatic.
nipkow
parents:
diff changeset
   108
92
bcd0ee8d71aa Hidden dtK and Impossible with a "local" clause
nipkow
parents: 91
diff changeset
   109
local open Syntax.Mixfix
bcd0ee8d71aa Hidden dtK and Impossible with a "local" clause
nipkow
parents: 91
diff changeset
   110
      exception Impossible
101
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   111
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   112
val is_rek = (fn dtRek _ => true  |  _  => false);
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   113
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   114
(* ------------------------------------------------------------------------- *)
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   115
(* Die Funktionen fuer das Umsetzen von Gleichungen in eine Definition mit   *)
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   116
(* dem prim-Rek. Kombinator                                                  *)
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   117
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   118
(*** Part 1: handling a single equation   ***)
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   119
 
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   120
(* filter REK type args by correspondence with targs. Reverses order *) 
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   121
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   122
fun rek_args (args, targs) = 
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   123
let fun h (x :: xs, tx :: txs, res) 
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   124
           = h(xs,txs,if is_rek tx then x :: res else res )
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   125
     |  h ([],[],res) = res
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   126
in h (args,targs,[])
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   127
end;
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   128
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   129
(* abstract over all recursive calls of f in t with param v in vs.
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   130
   Name in abstraction is variant of v w.r.t. free names in t. 
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   131
   Also returns reversed list of new variables names with types. 
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   132
   Checks that there are no free occurences of f left. 
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   133
*) 
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   134
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   135
fun abstract_recs f vs t  = 
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   136
let val tfrees = add_term_names(t,[]); 
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   137
    fun h [] vns t = if fst(dest_Const f) mem add_term_names(t,[]) 
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   138
		     then raise Impossible 
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   139
		     else (t,vns)
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   140
     |  h (v::vs) vns t
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   141
        = let val vn = variant tfrees (fst(dest_Free v))
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   142
          in  h vs (vn::vns) (Abs(vn, dummyT, abstract_over(f $ v,t)))
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   143
          end;
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   144
in h vs [] t
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   145
end;
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   146
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   147
(* For every defining equation, I need to abstract over arguments and
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   148
   over the recursive calls. Cant do it simply minded in this order, because 
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   149
   abstracting over v turns (Free v) into a bound variable, so that
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   150
   abstract_recs does not apply anymore.  
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   151
   abstract_arecs_funct performs the following steps 
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   152
    * abstract over (f xi) (reverse order) 
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   153
    * remove outermost length(rargs) abstractions
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   154
    * increase loose bound variables index by #cargs
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   155
    * apply the carg abstraction (reverse order) 
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   156
    * add length(rargs) lambdas. 
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   157
    Using lower level operations on term and arithmetic, this could probably
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   158
    be made more efficient. 
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   159
*) 
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   160
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   161
(* remove n outermost abstractions from a term *)
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   162
fun rem_Abs 0 t = t
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   163
 |  rem_Abs n (Abs(s,T,t)) = rem_Abs (n-1) t
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   164
;
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   165
(* add one abstraction for for every variable in vs *)  
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   166
fun add_Abs []      t = t
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   167
 |  add_Abs (vname::vs) t = Abs(vname, dummyT, add_Abs vs t)
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   168
; 
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   169
fun abstract_arecs funct rargs args t = 
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   170
let val (arecs,vns) = abstract_recs funct rargs t;
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   171
in  add_Abs vns 
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   172
    ( list_abs_free
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   173
        ( map dest_Free args
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   174
        , incr_boundvars (length args) (rem_Abs (length rargs) arecs)))
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   175
end;
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   176
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   177
(*** part 2. Processing of list of equations ***) 
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   178
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   179
(* Take list of constructors cs and equations eqns. 
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   180
   Find for ever element c of cs a corresponding eq in eqns. 
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   181
   Check that the function name is unique and there are no double params.  
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   182
   Derive term from equation using abstract_arecs and instantiate types. 
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   183
   Assume: equation list eqns nonempty
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   184
           length(eqns) = length(cs) 
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   185
           every constant name identifies a constant and its type. 
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   186
   In h: first parameter reqs reflects the remaining equations. 
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   187
*)
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   188
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   189
fun funs_from_eqns cs eqns =
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   190
let fun dest_eq ( Const("Trueprop",_) $ (Const ("op =",_)
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   191
                 $ (f $ capp) $ right))
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   192
	         = (f, strip_comb(capp), right);
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   193
    val fname = (fn (Const(f,_),_,_) => f) (dest_eq(hd eqns));
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   194
    fun h []   []        []       res = res
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   195
     |  h _    (_ :: _)  []       _   = raise Impossible
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   196
     |  h _    []        (_ :: _) _   = raise Impossible
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   197
     |  h reqs (eq::eqs) (c::cs)  res =
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   198
	let
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   199
          val (f,(Const(cname_eq,_),args),rhs) = dest_eq eq;
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   200
          val (cname,targs,syn) = c;
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   201
        in
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   202
	  if (cname_eq <> const_name cname syn) then h reqs eqs (c::cs) res
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   203
          else
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   204
          if fst(dest_Const(f)) = fname
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   205
             andalso (duplicates (map (fst o dest_Free) args) = [])
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   206
          then let val reqs' = reqs \ eq
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   207
               in h reqs' reqs' cs
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   208
                    (abstract_arecs f (rek_args(args,targs)) args rhs :: res)
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   209
	       end
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   210
          else raise Impossible
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   211
        end
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   212
in (fname, h eqns eqns cs []) end;
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   213
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   214
(* take datatype and eqns and return a properly type-instantiated 
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   215
   application of the prim-rec-combinator which solves eqns.
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   216
*)
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   217
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   218
fun instant_types thy t =
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   219
let val rs = Sign.rep_sg(sign_of thy);  
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   220
in  fst(Type.infer_types( #tsig rs,#const_tab rs, K None, K None
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   221
		        , TVar(("",0),[]), t))
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   222
end;
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   223
92
bcd0ee8d71aa Hidden dtK and Impossible with a "local" clause
nipkow
parents: 91
diff changeset
   224
in
101
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   225
80
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   226
fun add_datatype (typevars, tname, cons_list') thy = 
96
d94d0b324b4b Lots of simplifications
nipkow
parents: 92
diff changeset
   227
  let (*check if constructor names are unique*)
91
a94029edb01f added mixfix annotations to constructor declarations
clasohm
parents: 87
diff changeset
   228
      fun check_cons (cs : (string * 'b * 'c) list) =
a94029edb01f added mixfix annotations to constructor declarations
clasohm
parents: 87
diff changeset
   229
        (case findrep (map #1 cs) of
80
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   230
           [] => true
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   231
         | c::_ => error("Constructor \"" ^ c ^ "\" occurs twice"));
53
5e0570ea8b70 ML-like datatype decalaration facility. Axiomatic.
nipkow
parents:
diff changeset
   232
80
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   233
      (*search for free type variables and convert recursive *)
91
a94029edb01f added mixfix annotations to constructor declarations
clasohm
parents: 87
diff changeset
   234
      fun analyse_types (cons, typlist, syn) =
96
d94d0b324b4b Lots of simplifications
nipkow
parents: 92
diff changeset
   235
            let fun analyse(t as dtVar v) =
d94d0b324b4b Lots of simplifications
nipkow
parents: 92
diff changeset
   236
                     if t mem typevars then t
80
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   237
                     else error ("Variable " ^ v ^ " is free.")
96
d94d0b324b4b Lots of simplifications
nipkow
parents: 92
diff changeset
   238
                  | analyse(dtTyp(typl,s)) =
d94d0b324b4b Lots of simplifications
nipkow
parents: 92
diff changeset
   239
                     if tname <> s then dtTyp(analyses typl, s)
d94d0b324b4b Lots of simplifications
nipkow
parents: 92
diff changeset
   240
                     else if typevars = typl then dtRek(typl, s)
80
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   241
                     else error (s ^ " used in different ways")
96
d94d0b324b4b Lots of simplifications
nipkow
parents: 92
diff changeset
   242
                  | analyse(dtRek _) = raise Impossible
d94d0b324b4b Lots of simplifications
nipkow
parents: 92
diff changeset
   243
                 and analyses ts = map analyse ts;
d94d0b324b4b Lots of simplifications
nipkow
parents: 92
diff changeset
   244
            in (cons, analyses typlist, syn) end;
53
5e0570ea8b70 ML-like datatype decalaration facility. Axiomatic.
nipkow
parents:
diff changeset
   245
80
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   246
      (*test if there are elements that are not recursive, i.e. if the type is
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   247
        not empty*)
91
a94029edb01f added mixfix annotations to constructor declarations
clasohm
parents: 87
diff changeset
   248
      fun one_not_rek (cs : ('a * dt_type list * 'c) list) = 
92
bcd0ee8d71aa Hidden dtK and Impossible with a "local" clause
nipkow
parents: 91
diff changeset
   249
        let val contains_no_rek = forall (fn dtRek _ => false | _ => true);
91
a94029edb01f added mixfix annotations to constructor declarations
clasohm
parents: 87
diff changeset
   250
        in exists (contains_no_rek o #2) cs orelse
80
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   251
           error("Empty type not allowed!") end;
53
5e0570ea8b70 ML-like datatype decalaration facility. Axiomatic.
nipkow
parents:
diff changeset
   252
91
a94029edb01f added mixfix annotations to constructor declarations
clasohm
parents: 87
diff changeset
   253
      val dummy = check_cons cons_list';
a94029edb01f added mixfix annotations to constructor declarations
clasohm
parents: 87
diff changeset
   254
      val cons_list = map analyse_types cons_list';
a94029edb01f added mixfix annotations to constructor declarations
clasohm
parents: 87
diff changeset
   255
      val dummy = one_not_rek cons_list;
53
5e0570ea8b70 ML-like datatype decalaration facility. Axiomatic.
nipkow
parents:
diff changeset
   256
80
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   257
      (*Pretty printers for type lists;
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   258
        pp_typlist1: parentheses, pp_typlist2: brackets*)
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   259
      fun pp_typ (dtVar s) = s
96
d94d0b324b4b Lots of simplifications
nipkow
parents: 92
diff changeset
   260
        | pp_typ (dtTyp (typvars, id)) =
d94d0b324b4b Lots of simplifications
nipkow
parents: 92
diff changeset
   261
            if null typvars then id else (pp_typlist1 typvars) ^ id
92
bcd0ee8d71aa Hidden dtK and Impossible with a "local" clause
nipkow
parents: 91
diff changeset
   262
        | pp_typ (dtRek (typvars, id)) = (pp_typlist1 typvars) ^ id
80
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   263
      and
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   264
          pp_typlist' ts = commas (map pp_typ ts)
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   265
      and
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   266
          pp_typlist1 ts = if null ts then "" else pars (pp_typlist' ts);
53
5e0570ea8b70 ML-like datatype decalaration facility. Axiomatic.
nipkow
parents:
diff changeset
   267
80
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   268
      fun pp_typlist2 ts = if null ts then "" else brackets (pp_typlist' ts);
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   269
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   270
      fun Args(var, delim, n, m) = if n = m then var ^ string_of_int(n) 
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   271
                                   else var ^ string_of_int(n) ^ delim ^ 
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   272
			    	        Args(var, delim, n+1, m);
53
5e0570ea8b70 ML-like datatype decalaration facility. Axiomatic.
nipkow
parents:
diff changeset
   273
80
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   274
      (* Generate syntax translation for case rules *)
91
a94029edb01f added mixfix annotations to constructor declarations
clasohm
parents: 87
diff changeset
   275
      fun calc_xrules c_nr y_nr ((id, typlist, syn) :: cs) = 
a94029edb01f added mixfix annotations to constructor declarations
clasohm
parents: 87
diff changeset
   276
            let val name = const_name id syn;
a94029edb01f added mixfix annotations to constructor declarations
clasohm
parents: 87
diff changeset
   277
                val arity = length typlist;
80
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   278
                val body  = "z" ^ string_of_int(c_nr);
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   279
                val args1 = if arity=0 then ""
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   280
                            else pars (Args ("y", ",", y_nr, y_nr+arity-1));
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   281
                val args2 = if arity=0 then ""
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   282
                            else "% " ^ Args ("y", " ", y_nr, y_nr+arity-1) 
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   283
                            ^ ". ";
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   284
                val (rest1,rest2) = 
91
a94029edb01f added mixfix annotations to constructor declarations
clasohm
parents: 87
diff changeset
   285
		  if null cs then ("","")
80
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   286
                  else let val (h1, h2) = calc_xrules (c_nr+1) (y_nr+arity) cs
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   287
                       in (" | " ^ h1, ", " ^ h2) end;
91
a94029edb01f added mixfix annotations to constructor declarations
clasohm
parents: 87
diff changeset
   288
            in (name ^ args1 ^ " => " ^ body ^ rest1, args2 ^ body ^ rest2) end
80
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   289
        | calc_xrules _ _ [] = raise Impossible;
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   290
      
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   291
      val xrules =
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   292
         let val (first_part, scnd_part) = calc_xrules 1 1 cons_list
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   293
         in  [("logic", "case x of " ^ first_part) <->
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   294
              ("logic", tname ^ "_case(x, " ^ scnd_part ^ ")" )]
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   295
         end;
53
5e0570ea8b70 ML-like datatype decalaration facility. Axiomatic.
nipkow
parents:
diff changeset
   296
80
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   297
      (*type declarations for constructors*)
96
d94d0b324b4b Lots of simplifications
nipkow
parents: 92
diff changeset
   298
      fun const_type (id, typlist, syn) =
91
a94029edb01f added mixfix annotations to constructor declarations
clasohm
parents: 87
diff changeset
   299
           (id,  
80
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   300
            (if null typlist then "" else pp_typlist2 typlist ^ " => ") ^
96
d94d0b324b4b Lots of simplifications
nipkow
parents: 92
diff changeset
   301
             pp_typlist1 typevars ^ tname, syn);
53
5e0570ea8b70 ML-like datatype decalaration facility. Axiomatic.
nipkow
parents:
diff changeset
   302
80
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   303
      fun create_typevar (dtVar s) typlist =
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   304
            if (dtVar s) mem typlist then 
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   305
	      create_typevar (dtVar (s ^ "'")) typlist 
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   306
            else s
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   307
        | create_typevar _ _ = raise Impossible;
53
5e0570ea8b70 ML-like datatype decalaration facility. Axiomatic.
nipkow
parents:
diff changeset
   308
92
bcd0ee8d71aa Hidden dtK and Impossible with a "local" clause
nipkow
parents: 91
diff changeset
   309
      fun assumpt (dtRek _ :: ts, v :: vs ,found) =
96
d94d0b324b4b Lots of simplifications
nipkow
parents: 92
diff changeset
   310
            let val h = if found then ";P(" ^ v ^ ")" else "[| P(" ^ v ^ ")"
80
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   311
            in h ^ (assumpt (ts, vs, true)) end
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   312
        | assumpt (t :: ts, v :: vs, found) = assumpt (ts, vs, found)
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   313
        | assumpt ([], [], found) = if found then "|] ==>" else ""
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   314
        | assumpt _ = raise Impossible;
53
5e0570ea8b70 ML-like datatype decalaration facility. Axiomatic.
nipkow
parents:
diff changeset
   315
80
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   316
      (*insert type with suggested name 'varname' into table*)
96
d94d0b324b4b Lots of simplifications
nipkow
parents: 92
diff changeset
   317
      fun insert typ varname ((tri as (t, s, n)) :: xs) = 
80
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   318
            if typ = t then (t, s, n+1) :: xs
96
d94d0b324b4b Lots of simplifications
nipkow
parents: 92
diff changeset
   319
            else tri :: (if varname = s then insert typ (varname ^ "'") xs
d94d0b324b4b Lots of simplifications
nipkow
parents: 92
diff changeset
   320
                         else insert typ varname xs)
80
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   321
        | insert typ varname [] = [(typ, varname, 1)];
53
5e0570ea8b70 ML-like datatype decalaration facility. Axiomatic.
nipkow
parents:
diff changeset
   322
96
d94d0b324b4b Lots of simplifications
nipkow
parents: 92
diff changeset
   323
      fun typid(dtRek(_,id)) = id
d94d0b324b4b Lots of simplifications
nipkow
parents: 92
diff changeset
   324
        | typid(dtVar s) = implode (tl (explode s))
d94d0b324b4b Lots of simplifications
nipkow
parents: 92
diff changeset
   325
        | typid(dtTyp(_,id)) = id;
d94d0b324b4b Lots of simplifications
nipkow
parents: 92
diff changeset
   326
d94d0b324b4b Lots of simplifications
nipkow
parents: 92
diff changeset
   327
      val insert_types = foldl (fn (tab,typ) => insert typ (typid typ) tab);
53
5e0570ea8b70 ML-like datatype decalaration facility. Axiomatic.
nipkow
parents:
diff changeset
   328
92
bcd0ee8d71aa Hidden dtK and Impossible with a "local" clause
nipkow
parents: 91
diff changeset
   329
      fun update(dtRek _, s, v :: vs, (dtRek _) :: ts) = s :: vs
80
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   330
        | update(t, s, v :: vs, t1 :: ts) = 
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   331
            if t=t1 then s :: vs
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   332
                    else v :: (update (t, s, vs, ts))
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   333
        | update _ = raise Impossible;
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   334
      
92
bcd0ee8d71aa Hidden dtK and Impossible with a "local" clause
nipkow
parents: 91
diff changeset
   335
      fun update_n (dtRek r1, s, v :: vs, (dtRek r2) :: ts, n) =
80
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   336
            if r1 = r2 then (s ^ string_of_int n) :: 
92
bcd0ee8d71aa Hidden dtK and Impossible with a "local" clause
nipkow
parents: 91
diff changeset
   337
                            (update_n (dtRek r1, s, vs, ts, n+1))
bcd0ee8d71aa Hidden dtK and Impossible with a "local" clause
nipkow
parents: 91
diff changeset
   338
                       else v :: (update_n (dtRek r1, s, vs, ts, n))
80
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   339
        | update_n (t, s, v :: vs, t1 :: ts, n) = 
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   340
            if t = t1 then (s ^ string_of_int n) :: 
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   341
                           (update_n (t, s, vs, ts, n+1))
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   342
                      else v :: (update_n (t, s, vs, ts, n))
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   343
        | update_n (_,_,[],[],_) = []
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   344
        | update_n _ = raise Impossible;
53
5e0570ea8b70 ML-like datatype decalaration facility. Axiomatic.
nipkow
parents:
diff changeset
   345
80
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   346
      (*insert type variables into table*)
96
d94d0b324b4b Lots of simplifications
nipkow
parents: 92
diff changeset
   347
      fun convert typs =
d94d0b324b4b Lots of simplifications
nipkow
parents: 92
diff changeset
   348
        let fun conv(vars, (t, s, n)) =
d94d0b324b4b Lots of simplifications
nipkow
parents: 92
diff changeset
   349
              if n=1 then update (t, s, vars, typs)
d94d0b324b4b Lots of simplifications
nipkow
parents: 92
diff changeset
   350
                     else update_n (t, s, vars, typs, 1)
d94d0b324b4b Lots of simplifications
nipkow
parents: 92
diff changeset
   351
        in foldl conv end;
53
5e0570ea8b70 ML-like datatype decalaration facility. Axiomatic.
nipkow
parents:
diff changeset
   352
80
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   353
      fun empty_list n = replicate n "";
53
5e0570ea8b70 ML-like datatype decalaration facility. Axiomatic.
nipkow
parents:
diff changeset
   354
91
a94029edb01f added mixfix annotations to constructor declarations
clasohm
parents: 87
diff changeset
   355
      fun t_inducting ((id, typl, syn) :: cs) =
a94029edb01f added mixfix annotations to constructor declarations
clasohm
parents: 87
diff changeset
   356
            let val name = const_name id syn;
96
d94d0b324b4b Lots of simplifications
nipkow
parents: 92
diff changeset
   357
                val tab = insert_types([],typl);
80
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   358
                val arity = length typl;
96
d94d0b324b4b Lots of simplifications
nipkow
parents: 92
diff changeset
   359
                val var_list = convert typl (empty_list arity,tab); 
80
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   360
                val h = if arity = 0 then " P(" ^ name ^ ")"
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   361
                        else " !!" ^ (space_implode " " var_list) ^ "." ^
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   362
                             (assumpt (typl, var_list, false)) ^ "P(" ^ 
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   363
                             name ^ "(" ^ (commas var_list) ^ "))";
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   364
                val rest = t_inducting cs;
91
a94029edb01f added mixfix annotations to constructor declarations
clasohm
parents: 87
diff changeset
   365
            in if rest = "" then h else h ^ "; " ^ rest end
80
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   366
        | t_inducting [] = "";
53
5e0570ea8b70 ML-like datatype decalaration facility. Axiomatic.
nipkow
parents:
diff changeset
   367
80
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   368
      fun t_induct cl typ_name=
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   369
        "[|" ^ t_inducting cl ^ "|] ==> P(" ^ typ_name ^ ")";
53
5e0570ea8b70 ML-like datatype decalaration facility. Axiomatic.
nipkow
parents:
diff changeset
   370
91
a94029edb01f added mixfix annotations to constructor declarations
clasohm
parents: 87
diff changeset
   371
      fun case_typlist typevar ((_, typlist, _) :: cs) =
80
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   372
           let val h = if (length typlist) > 0 then 
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   373
		         (pp_typlist2 typlist) ^ "=>"
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   374
                       else ""
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   375
           in "," ^ h ^ typevar ^ (case_typlist typevar cs) end
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   376
        | case_typlist _ [] = "";
53
5e0570ea8b70 ML-like datatype decalaration facility. Axiomatic.
nipkow
parents:
diff changeset
   377
101
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   378
      val t_case = tname ^ "_case";
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   379
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   380
      fun case_rules arity n ((id, typlist, syn) :: cs) =
91
a94029edb01f added mixfix annotations to constructor declarations
clasohm
parents: 87
diff changeset
   381
            let val name = const_name id syn;
a94029edb01f added mixfix annotations to constructor declarations
clasohm
parents: 87
diff changeset
   382
                val args = if null typlist then ""
96
d94d0b324b4b Lots of simplifications
nipkow
parents: 92
diff changeset
   383
  			   else pars(Args("x", ",", 1, length typlist))
80
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   384
            in (t_case ^ "_" ^ id,
91
a94029edb01f added mixfix annotations to constructor declarations
clasohm
parents: 87
diff changeset
   385
                t_case ^ "(" ^ name ^ args ^ "," ^ Args ("f", ",", 1, arity) 
80
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   386
                ^ ") = f" ^ string_of_int(n) ^ args)
101
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   387
               :: (case_rules arity (n+1) cs)
80
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   388
            end
101
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   389
        | case_rules _ _ [] = [];
53
5e0570ea8b70 ML-like datatype decalaration facility. Axiomatic.
nipkow
parents:
diff changeset
   390
80
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   391
      val datatype_arity = length typevars;
53
5e0570ea8b70 ML-like datatype decalaration facility. Axiomatic.
nipkow
parents:
diff changeset
   392
91
a94029edb01f added mixfix annotations to constructor declarations
clasohm
parents: 87
diff changeset
   393
      val types = [(tname, datatype_arity, NoSyn)];
53
5e0570ea8b70 ML-like datatype decalaration facility. Axiomatic.
nipkow
parents:
diff changeset
   394
80
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   395
      val arities = 
91
a94029edb01f added mixfix annotations to constructor declarations
clasohm
parents: 87
diff changeset
   396
        let val term_list = replicate datatype_arity ["term"];
a94029edb01f added mixfix annotations to constructor declarations
clasohm
parents: 87
diff changeset
   397
        in [(tname, term_list, ["term"])] end;
53
5e0570ea8b70 ML-like datatype decalaration facility. Axiomatic.
nipkow
parents:
diff changeset
   398
80
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   399
      val datatype_name = pp_typlist1 typevars ^ tname;
53
5e0570ea8b70 ML-like datatype decalaration facility. Axiomatic.
nipkow
parents:
diff changeset
   400
80
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   401
      val (case_const, rules_case) =
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   402
         let val typevar = create_typevar (dtVar "'beta") typevars;
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   403
             val arity = length cons_list;
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   404
             val dekl = (t_case, "[" ^ pp_typlist1 typevars ^ tname ^
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   405
                       case_typlist typevar cons_list ^ "]=>" ^ typevar, NoSyn)
101
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   406
             val rules = case_rules arity 1 cons_list;
80
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   407
         in (dekl, rules) end;
53
5e0570ea8b70 ML-like datatype decalaration facility. Axiomatic.
nipkow
parents:
diff changeset
   408
101
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   409
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   410
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   411
(* -------------------------------------------------------------------- *)
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   412
(* Die Funktionen fuer die t_rec - Funktion                             *)
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   413
(* Analog zu t_case bis auf Hinzufuegen rek. Aufrufe pro Konstruktor 	*)
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   414
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   415
      val t_rec = tname ^ "_rec"
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   416
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   417
fun add_reks typevar ts = 
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   418
  let val tv = dtVar typevar; 
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   419
      fun h (t::ts) res = h ts (if is_rek(t) then tv::res else res)
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   420
	| h [] res  = res        
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   421
  in  h ts ts
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   422
  end;
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   423
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   424
fun rec_typlist typevar ((c,ts,_)::cs) = 
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   425
    let val h = if (length ts) > 0 
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   426
	        then (pp_typlist2 (add_reks typevar ts)) ^ "=>"
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   427
                else ""
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   428
    in "," ^ h ^ typevar ^ (rec_typlist typevar cs)
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   429
    end
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   430
  | rec_typlist _ [] = "";
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   431
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   432
fun arg_reks arity ts = 
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   433
  let fun arg_rek (t::ts) n res  = 
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   434
        let val h = t_rec ^"(" ^ "x" ^string_of_int(n) 
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   435
			       ^"," ^Args("f",",",1,arity) ^")," 
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   436
        in arg_rek ts (n+1) (if is_rek(t) then res ^ h else res)
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   437
        end 
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   438
      | arg_rek [] _ res = res        
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   439
  in  arg_rek ts 1 ""
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   440
  end;
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   441
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   442
fun rec_rules arity n ((id,ts,syn)::cs) =
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   443
  let val name = const_name id syn;
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   444
      val lts = length ts 
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   445
      val args = if (lts = 0) then ""
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   446
	         else "(" ^ Args("x",",",1,lts) ^ ")" 
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   447
      val rargs = if (lts = 0) then ""
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   448
	          else "("^ arg_reks arity ts ^ Args("x",",",1,lts) ^")"
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   449
  in     
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   450
    ( t_rec ^ "_" ^ id
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   451
    , t_rec ^ "(" ^ name ^ args ^ "," ^ Args("f",",",1,arity) ^ ") = f"
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   452
      ^ string_of_int(n) ^ rargs) 
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   453
     :: (rec_rules arity (n+1) cs)
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   454
  end
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   455
  | rec_rules _ _ [] = [];
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   456
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   457
val (rec_const,rules_rec) =
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   458
   let val typevar = create_typevar (dtVar "'beta") typevars
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   459
       val arity = length cons_list
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   460
       val dekl = (t_rec,
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   461
                    "[" ^ (pp_typlist1 typevars) ^ tname ^
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   462
                      (rec_typlist typevar cons_list) ^ "]=>" ^ typevar,
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   463
                    NoSyn)
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   464
       val rules = rec_rules arity 1 cons_list
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   465
       in (dekl,rules)
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   466
   end;
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   467
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   468
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   469
80
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   470
      val consts = 
96
d94d0b324b4b Lots of simplifications
nipkow
parents: 92
diff changeset
   471
        map const_type cons_list
91
a94029edb01f added mixfix annotations to constructor declarations
clasohm
parents: 87
diff changeset
   472
	@ (if length cons_list < dtK then []
80
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   473
	   else [(tname ^ "_ord", datatype_name ^ "=>nat", NoSyn)])
101
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   474
	@ [case_const,rec_const];
53
5e0570ea8b70 ML-like datatype decalaration facility. Axiomatic.
nipkow
parents:
diff changeset
   475
80
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   476
      (*generate 'var_n, ..., var_m'*)
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   477
      fun Args(var, delim, n, m) = 
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   478
        space_implode delim (map (fn n => var^string_of_int(n)) (n upto m));
53
5e0570ea8b70 ML-like datatype decalaration facility. Axiomatic.
nipkow
parents:
diff changeset
   479
80
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   480
      (*generate 'name_1', ..., 'name_n'*)
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   481
      fun C_exp(name, n, var) =
96
d94d0b324b4b Lots of simplifications
nipkow
parents: 92
diff changeset
   482
        if n > 0 then name ^ pars(Args(var, ",", 1, n)) else name;
53
5e0570ea8b70 ML-like datatype decalaration facility. Axiomatic.
nipkow
parents:
diff changeset
   483
80
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   484
      (*generate 'x_n = y_n, ..., x_m = y_m'*)
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   485
      fun Arg_eql(n,m) = 
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   486
        if n=m then "x" ^ string_of_int(n) ^ "=y" ^ string_of_int(n) 
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   487
        else "x" ^ string_of_int(n) ^ "=y" ^ string_of_int(n) ^ " & " ^ 
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   488
             Arg_eql(n+1, m);
53
5e0570ea8b70 ML-like datatype decalaration facility. Axiomatic.
nipkow
parents:
diff changeset
   489
91
a94029edb01f added mixfix annotations to constructor declarations
clasohm
parents: 87
diff changeset
   490
      fun Ci_ing ((id, typlist, syn) :: cs) =
a94029edb01f added mixfix annotations to constructor declarations
clasohm
parents: 87
diff changeset
   491
            let val name = const_name id syn;
a94029edb01f added mixfix annotations to constructor declarations
clasohm
parents: 87
diff changeset
   492
                val arity = length typlist;
96
d94d0b324b4b Lots of simplifications
nipkow
parents: 92
diff changeset
   493
            in if arity = 0 then Ci_ing cs
d94d0b324b4b Lots of simplifications
nipkow
parents: 92
diff changeset
   494
               else ("inject_" ^ id,
80
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   495
                     "(" ^ C_exp(name,arity,"x") ^ "=" ^ C_exp(name,arity,"y") 
91
a94029edb01f added mixfix annotations to constructor declarations
clasohm
parents: 87
diff changeset
   496
                     ^ ") = (" ^ Arg_eql (1, arity) ^ ")") :: (Ci_ing cs)
80
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   497
            end
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   498
        | Ci_ing [] = [];
53
5e0570ea8b70 ML-like datatype decalaration facility. Axiomatic.
nipkow
parents:
diff changeset
   499
96
d94d0b324b4b Lots of simplifications
nipkow
parents: 92
diff changeset
   500
      fun Ci_negOne (id1, tl1, syn1) (id2, tl2, syn2) =
d94d0b324b4b Lots of simplifications
nipkow
parents: 92
diff changeset
   501
           let val name1 = const_name id1 syn1;
91
a94029edb01f added mixfix annotations to constructor declarations
clasohm
parents: 87
diff changeset
   502
               val name2 = const_name id2 syn2;
96
d94d0b324b4b Lots of simplifications
nipkow
parents: 92
diff changeset
   503
               val ax = C_exp(name1, length tl1, "x") ^ "~=" ^
d94d0b324b4b Lots of simplifications
nipkow
parents: 92
diff changeset
   504
                        C_exp(name2, length tl2, "y")
d94d0b324b4b Lots of simplifications
nipkow
parents: 92
diff changeset
   505
           in (id1 ^ "_not_" ^ id2, ax) end;
53
5e0570ea8b70 ML-like datatype decalaration facility. Axiomatic.
nipkow
parents:
diff changeset
   506
80
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   507
      fun Ci_neg1 [] = []
96
d94d0b324b4b Lots of simplifications
nipkow
parents: 92
diff changeset
   508
        | Ci_neg1 (c1::cs) = (map (Ci_negOne c1) cs) @ Ci_neg1 cs;
53
5e0570ea8b70 ML-like datatype decalaration facility. Axiomatic.
nipkow
parents:
diff changeset
   509
80
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   510
      fun suc_expr n = 
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   511
        if n=0 then "0" else "Suc(" ^ suc_expr(n-1) ^ ")";
53
5e0570ea8b70 ML-like datatype decalaration facility. Axiomatic.
nipkow
parents:
diff changeset
   512
96
d94d0b324b4b Lots of simplifications
nipkow
parents: 92
diff changeset
   513
      fun Ci_neg2() =
80
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   514
        let val ord_t = tname ^ "_ord";
96
d94d0b324b4b Lots of simplifications
nipkow
parents: 92
diff changeset
   515
            val cis = cons_list ~~ (0 upto (length cons_list - 1))
d94d0b324b4b Lots of simplifications
nipkow
parents: 92
diff changeset
   516
            fun Ci_neg2equals ((id, typlist, syn), n) =
d94d0b324b4b Lots of simplifications
nipkow
parents: 92
diff changeset
   517
              let val name = const_name id syn;
d94d0b324b4b Lots of simplifications
nipkow
parents: 92
diff changeset
   518
                  val ax = ord_t ^ "(" ^ (C_exp(name, length typlist, "x")) 
d94d0b324b4b Lots of simplifications
nipkow
parents: 92
diff changeset
   519
                                 ^ ") = " ^ (suc_expr n)
d94d0b324b4b Lots of simplifications
nipkow
parents: 92
diff changeset
   520
              in (ord_t ^ "_" ^ id, ax) end
d94d0b324b4b Lots of simplifications
nipkow
parents: 92
diff changeset
   521
        in (ord_t ^ "_distinct", ord_t^"(x) ~= "^ord_t^"(y) ==> x ~= y") ::
d94d0b324b4b Lots of simplifications
nipkow
parents: 92
diff changeset
   522
           (map Ci_neg2equals cis)
80
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   523
        end;
53
5e0570ea8b70 ML-like datatype decalaration facility. Axiomatic.
nipkow
parents:
diff changeset
   524
96
d94d0b324b4b Lots of simplifications
nipkow
parents: 92
diff changeset
   525
      val rules_distinct = if length cons_list < dtK then Ci_neg1 cons_list
d94d0b324b4b Lots of simplifications
nipkow
parents: 92
diff changeset
   526
                           else Ci_neg2();
53
5e0570ea8b70 ML-like datatype decalaration facility. Axiomatic.
nipkow
parents:
diff changeset
   527
80
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   528
      val rules_inject = Ci_ing cons_list;
53
5e0570ea8b70 ML-like datatype decalaration facility. Axiomatic.
nipkow
parents:
diff changeset
   529
87
b0ea0e55dfe8 added prefix to name of induct axiom
clasohm
parents: 80
diff changeset
   530
      val rule_induct = (tname ^ "_induct", t_induct cons_list tname);
60
43e36c15a831 new field "simps" added
nipkow
parents: 53
diff changeset
   531
101
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   532
      val rules = rule_induct ::
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   533
                  (rules_inject @ rules_distinct @ rules_case @ rules_rec);
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   534
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   535
      fun add_primrec eqns thy =
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   536
      let val rec_comb = Const(t_rec,dummyT)
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   537
          val teqns = map (fn eq => snd(read_axm (sign_of thy) ("",eq))) eqns
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   538
          val (fname,rfuns) = funs_from_eqns cons_list teqns
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   539
          val rhs = Abs(tname, dummyT,
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   540
                        list_comb(rec_comb, Bound 0 :: rev rfuns))
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   541
          val def = Const("==",dummyT) $ Const(fname,dummyT) $ rhs
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   542
          val tdef = instant_types thy def
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   543
      in add_defns_i [(fname ^ "_def", tdef)] thy end;
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   544
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   545
  in (thy
80
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   546
     |> add_types types
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   547
     |> add_arities arities
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   548
     |> add_consts consts
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   549
     |> add_trrules xrules
101
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   550
     |> add_axioms rules,
5f99df1e26c4 Added primitive recursive functions (Norbert Voelker's code) to the datatype
nipkow
parents: 96
diff changeset
   551
     add_primrec)
80
d3d727449d7b datatypes must now be defined using a thy file section
clasohm
parents: 60
diff changeset
   552
  end
92
bcd0ee8d71aa Hidden dtK and Impossible with a "local" clause
nipkow
parents: 91
diff changeset
   553
end
bcd0ee8d71aa Hidden dtK and Impossible with a "local" clause
nipkow
parents: 91
diff changeset
   554
end;