src/HOL/Inductive.ML
author paulson
Wed, 21 Aug 1996 13:22:23 +0200
changeset 1933 8b24773de6db
parent 1465 5d7a7e439cec
child 2855 36f75c4a0047
permissions -rw-r--r--
Addition of message NS5
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1430
diff changeset
     1
(*  Title:      HOL/inductive.ML
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     2
    ID:         $Id$
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1430
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     4
    Copyright   1993  University of Cambridge
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
(Co)Inductive Definitions for HOL
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
Inductive definitions use least fixedpoints with standard products and sums
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     9
Coinductive definitions use greatest fixedpoints with Quine products and sums
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    10
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    11
Sums are used only for mutual recursion;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    12
Products are used only to derive "streamlined" induction rules for relations
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    13
*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    14
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    15
fun gen_fp_oper a (X,T,t) = 
1430
439e1476a7f8 Improving space efficiency of inductive/datatype definitions.
paulson
parents: 923
diff changeset
    16
    let val setT = Ind_Syntax.mk_setT T
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    17
    in Const(a, (setT-->setT)-->setT) $ absfree(X, setT, t)  end;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    18
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    19
structure Lfp_items =
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    20
  struct
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1430
diff changeset
    21
  val oper      = gen_fp_oper "lfp"
5d7a7e439cec expanded tabs
clasohm
parents: 1430
diff changeset
    22
  val Tarski    = def_lfp_Tarski
5d7a7e439cec expanded tabs
clasohm
parents: 1430
diff changeset
    23
  val induct    = def_induct
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    24
  end;
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
structure Gfp_items =
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    27
  struct
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1430
diff changeset
    28
  val oper      = gen_fp_oper "gfp"
5d7a7e439cec expanded tabs
clasohm
parents: 1430
diff changeset
    29
  val Tarski    = def_gfp_Tarski
5d7a7e439cec expanded tabs
clasohm
parents: 1430
diff changeset
    30
  val induct    = def_Collect_coinduct
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    31
  end;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    32
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    33
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    34
functor Ind_section_Fun (Inductive: sig include INDUCTIVE_ARG INDUCTIVE_I end) 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    35
  : sig include INTR_ELIM INDRULE end =
1430
439e1476a7f8 Improving space efficiency of inductive/datatype definitions.
paulson
parents: 923
diff changeset
    36
let
439e1476a7f8 Improving space efficiency of inductive/datatype definitions.
paulson
parents: 923
diff changeset
    37
  structure Intr_elim = Intr_elim_Fun(structure Inductive=Inductive and 
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1430
diff changeset
    38
                                          Fp=Lfp_items);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    39
1430
439e1476a7f8 Improving space efficiency of inductive/datatype definitions.
paulson
parents: 923
diff changeset
    40
  structure Indrule = Indrule_Fun
439e1476a7f8 Improving space efficiency of inductive/datatype definitions.
paulson
parents: 923
diff changeset
    41
      (structure Inductive=Inductive and Intr_elim=Intr_elim);
439e1476a7f8 Improving space efficiency of inductive/datatype definitions.
paulson
parents: 923
diff changeset
    42
in 
439e1476a7f8 Improving space efficiency of inductive/datatype definitions.
paulson
parents: 923
diff changeset
    43
   struct 
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1430
diff changeset
    44
   val thy      = Intr_elim.thy
5d7a7e439cec expanded tabs
clasohm
parents: 1430
diff changeset
    45
   val defs     = Intr_elim.defs
5d7a7e439cec expanded tabs
clasohm
parents: 1430
diff changeset
    46
   val mono     = Intr_elim.mono
5d7a7e439cec expanded tabs
clasohm
parents: 1430
diff changeset
    47
   val intrs    = Intr_elim.intrs
5d7a7e439cec expanded tabs
clasohm
parents: 1430
diff changeset
    48
   val elim     = Intr_elim.elim
5d7a7e439cec expanded tabs
clasohm
parents: 1430
diff changeset
    49
   val mk_cases = Intr_elim.mk_cases
1430
439e1476a7f8 Improving space efficiency of inductive/datatype definitions.
paulson
parents: 923
diff changeset
    50
   open Indrule 
439e1476a7f8 Improving space efficiency of inductive/datatype definitions.
paulson
parents: 923
diff changeset
    51
   end
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    52
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
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    55
structure Ind = Add_inductive_def_Fun (Lfp_items);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    56
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    57
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    58
signature INDUCTIVE_STRING =
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    59
  sig
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1430
diff changeset
    60
  val thy_name   : string               (*name of the new theory*)
5d7a7e439cec expanded tabs
clasohm
parents: 1430
diff changeset
    61
  val srec_tms   : string list          (*recursion terms*)
5d7a7e439cec expanded tabs
clasohm
parents: 1430
diff changeset
    62
  val sintrs     : string list          (*desired introduction rules*)
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    63
  end;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    64
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    65
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    66
(*For upwards compatibility: can be called directly from ML*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    67
functor Inductive_Fun
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    68
 (Inductive: sig include INDUCTIVE_STRING INDUCTIVE_ARG end)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    69
   : sig include INTR_ELIM INDRULE end =
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    70
Ind_section_Fun
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    71
   (open Inductive Ind_Syntax
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    72
    val sign = sign_of thy;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    73
    val rec_tms = map (readtm sign termTVar) srec_tms
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    74
    and intr_tms = map (readtm sign propT) sintrs;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    75
    val thy = thy |> Ind.add_fp_def_i(rec_tms, intr_tms) 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    76
                  |> add_thyname thy_name);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    77
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    78
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    79
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    80
signature COINDRULE =
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    81
  sig
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    82
  val coinduct : thm
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    83
  end;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    84
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    85
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    86
functor CoInd_section_Fun
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    87
 (Inductive: sig include INDUCTIVE_ARG INDUCTIVE_I end) 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    88
    : sig include INTR_ELIM COINDRULE end =
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    89
struct
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    90
structure Intr_elim = Intr_elim_Fun(structure Inductive=Inductive and Fp=Gfp_items);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    91
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    92
open Intr_elim 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    93
val coinduct = raw_induct
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    94
end;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    95
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
structure CoInd = Add_inductive_def_Fun(Gfp_items);