src/ZF/Inductive.ML
author kleing
Fri, 27 May 2005 01:09:44 +0200
changeset 16095 f6af6b265d20
parent 12132 1ef58b332ca9
permissions -rw-r--r--
put global isatest settings in one file, sourced by the other scripts
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1418
diff changeset
     1
(*  Title:      ZF/Inductive.ML
578
efc648d29dd0 ZF/Inductive.thy,.ML: renamed from "inductive" to allow re-building without
lcp
parents:
diff changeset
     2
    ID:         $Id$
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1418
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
578
efc648d29dd0 ZF/Inductive.thy,.ML: renamed from "inductive" to allow re-building without
lcp
parents:
diff changeset
     4
    Copyright   1993  University of Cambridge
efc648d29dd0 ZF/Inductive.thy,.ML: renamed from "inductive" to allow re-building without
lcp
parents:
diff changeset
     5
efc648d29dd0 ZF/Inductive.thy,.ML: renamed from "inductive" to allow re-building without
lcp
parents:
diff changeset
     6
(Co)Inductive Definitions for Zermelo-Fraenkel Set Theory
efc648d29dd0 ZF/Inductive.thy,.ML: renamed from "inductive" to allow re-building without
lcp
parents:
diff changeset
     7
efc648d29dd0 ZF/Inductive.thy,.ML: renamed from "inductive" to allow re-building without
lcp
parents:
diff changeset
     8
Inductive definitions use least fixedpoints with standard products and sums
efc648d29dd0 ZF/Inductive.thy,.ML: renamed from "inductive" to allow re-building without
lcp
parents:
diff changeset
     9
Coinductive definitions use greatest fixedpoints with Quine products and sums
efc648d29dd0 ZF/Inductive.thy,.ML: renamed from "inductive" to allow re-building without
lcp
parents:
diff changeset
    10
efc648d29dd0 ZF/Inductive.thy,.ML: renamed from "inductive" to allow re-building without
lcp
parents:
diff changeset
    11
Sums are used only for mutual recursion;
efc648d29dd0 ZF/Inductive.thy,.ML: renamed from "inductive" to allow re-building without
lcp
parents:
diff changeset
    12
Products are used only to derive "streamlined" induction rules for relations
efc648d29dd0 ZF/Inductive.thy,.ML: renamed from "inductive" to allow re-building without
lcp
parents:
diff changeset
    13
*)
efc648d29dd0 ZF/Inductive.thy,.ML: renamed from "inductive" to allow re-building without
lcp
parents:
diff changeset
    14
1418
f5f97ee67cbb Improving space efficiency of inductive/datatype definitions.
paulson
parents: 1093
diff changeset
    15
val iT = Ind_Syntax.iT
4352
7ac9f3e8a97d Moved some functions from ZF/ind_syntax.ML to FOL/fologic.ML
paulson
parents: 1737
diff changeset
    16
and oT = FOLogic.oT;
1418
f5f97ee67cbb Improving space efficiency of inductive/datatype definitions.
paulson
parents: 1093
diff changeset
    17
578
efc648d29dd0 ZF/Inductive.thy,.ML: renamed from "inductive" to allow re-building without
lcp
parents:
diff changeset
    18
structure Lfp =
efc648d29dd0 ZF/Inductive.thy,.ML: renamed from "inductive" to allow re-building without
lcp
parents:
diff changeset
    19
  struct
6093
87bf8c03b169 eliminated global/local names;
wenzelm
parents: 6053
diff changeset
    20
  val oper      = Const("Fixedpt.lfp", [iT,iT-->iT]--->iT)
87bf8c03b169 eliminated global/local names;
wenzelm
parents: 6053
diff changeset
    21
  val bnd_mono  = Const("Fixedpt.bnd_mono", [iT,iT-->iT]--->oT)
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1418
diff changeset
    22
  val bnd_monoI = bnd_monoI
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1418
diff changeset
    23
  val subs      = def_lfp_subset
10216
e928bdf62014 renamed fp_Tarski to fp_unfold
paulson
parents: 6093
diff changeset
    24
  val Tarski    = def_lfp_unfold
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1418
diff changeset
    25
  val induct    = def_induct
578
efc648d29dd0 ZF/Inductive.thy,.ML: renamed from "inductive" to allow re-building without
lcp
parents:
diff changeset
    26
  end;
efc648d29dd0 ZF/Inductive.thy,.ML: renamed from "inductive" to allow re-building without
lcp
parents:
diff changeset
    27
efc648d29dd0 ZF/Inductive.thy,.ML: renamed from "inductive" to allow re-building without
lcp
parents:
diff changeset
    28
structure Standard_Prod =
efc648d29dd0 ZF/Inductive.thy,.ML: renamed from "inductive" to allow re-building without
lcp
parents:
diff changeset
    29
  struct
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1418
diff changeset
    30
  val sigma     = Const("Sigma", [iT, iT-->iT]--->iT)
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1418
diff changeset
    31
  val pair      = Const("Pair", [iT,iT]--->iT)
1737
5a4f382455ce Modified to use new functor signatures
paulson
parents: 1461
diff changeset
    32
  val split_name = "split"
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1418
diff changeset
    33
  val pair_iff  = Pair_iff
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1418
diff changeset
    34
  val split_eq  = split
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1418
diff changeset
    35
  val fsplitI   = splitI
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1418
diff changeset
    36
  val fsplitD   = splitD
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1418
diff changeset
    37
  val fsplitE   = splitE
578
efc648d29dd0 ZF/Inductive.thy,.ML: renamed from "inductive" to allow re-building without
lcp
parents:
diff changeset
    38
  end;
efc648d29dd0 ZF/Inductive.thy,.ML: renamed from "inductive" to allow re-building without
lcp
parents:
diff changeset
    39
1737
5a4f382455ce Modified to use new functor signatures
paulson
parents: 1461
diff changeset
    40
structure Standard_CP = CartProd_Fun (Standard_Prod);
5a4f382455ce Modified to use new functor signatures
paulson
parents: 1461
diff changeset
    41
578
efc648d29dd0 ZF/Inductive.thy,.ML: renamed from "inductive" to allow re-building without
lcp
parents:
diff changeset
    42
structure Standard_Sum =
efc648d29dd0 ZF/Inductive.thy,.ML: renamed from "inductive" to allow re-building without
lcp
parents:
diff changeset
    43
  struct
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1418
diff changeset
    44
  val sum       = Const("op +", [iT,iT]--->iT)
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1418
diff changeset
    45
  val inl       = Const("Inl", iT-->iT)
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1418
diff changeset
    46
  val inr       = Const("Inr", iT-->iT)
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1418
diff changeset
    47
  val elim      = Const("case", [iT-->iT, iT-->iT, iT]--->iT)
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1418
diff changeset
    48
  val case_inl  = case_Inl
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1418
diff changeset
    49
  val case_inr  = case_Inr
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1418
diff changeset
    50
  val inl_iff   = Inl_iff
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1418
diff changeset
    51
  val inr_iff   = Inr_iff
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1418
diff changeset
    52
  val distinct  = Inl_Inr_iff
578
efc648d29dd0 ZF/Inductive.thy,.ML: renamed from "inductive" to allow re-building without
lcp
parents:
diff changeset
    53
  val distinct' = Inr_Inl_iff
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 10216
diff changeset
    54
  val free_SEs  = Ind_Syntax.mk_free_SEs
1418
f5f97ee67cbb Improving space efficiency of inductive/datatype definitions.
paulson
parents: 1093
diff changeset
    55
            [distinct, distinct', inl_iff, inr_iff, Standard_Prod.pair_iff]
578
efc648d29dd0 ZF/Inductive.thy,.ML: renamed from "inductive" to allow re-building without
lcp
parents:
diff changeset
    56
  end;
1418
f5f97ee67cbb Improving space efficiency of inductive/datatype definitions.
paulson
parents: 1093
diff changeset
    57
578
efc648d29dd0 ZF/Inductive.thy,.ML: renamed from "inductive" to allow re-building without
lcp
parents:
diff changeset
    58
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 10216
diff changeset
    59
structure Ind_Package =
6053
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 4352
diff changeset
    60
    Add_inductive_def_Fun
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 4352
diff changeset
    61
      (structure Fp=Lfp and Pr=Standard_Prod and CP=Standard_CP
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 10216
diff changeset
    62
       and Su=Standard_Sum val coind = false);
578
efc648d29dd0 ZF/Inductive.thy,.ML: renamed from "inductive" to allow re-building without
lcp
parents:
diff changeset
    63
efc648d29dd0 ZF/Inductive.thy,.ML: renamed from "inductive" to allow re-building without
lcp
parents:
diff changeset
    64
efc648d29dd0 ZF/Inductive.thy,.ML: renamed from "inductive" to allow re-building without
lcp
parents:
diff changeset
    65
structure Gfp =
efc648d29dd0 ZF/Inductive.thy,.ML: renamed from "inductive" to allow re-building without
lcp
parents:
diff changeset
    66
  struct
6093
87bf8c03b169 eliminated global/local names;
wenzelm
parents: 6053
diff changeset
    67
  val oper      = Const("Fixedpt.gfp", [iT,iT-->iT]--->iT)
87bf8c03b169 eliminated global/local names;
wenzelm
parents: 6053
diff changeset
    68
  val bnd_mono  = Const("Fixedpt.bnd_mono", [iT,iT-->iT]--->oT)
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1418
diff changeset
    69
  val bnd_monoI = bnd_monoI
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1418
diff changeset
    70
  val subs      = def_gfp_subset
10216
e928bdf62014 renamed fp_Tarski to fp_unfold
paulson
parents: 6093
diff changeset
    71
  val Tarski    = def_gfp_unfold
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1418
diff changeset
    72
  val induct    = def_Collect_coinduct
578
efc648d29dd0 ZF/Inductive.thy,.ML: renamed from "inductive" to allow re-building without
lcp
parents:
diff changeset
    73
  end;
efc648d29dd0 ZF/Inductive.thy,.ML: renamed from "inductive" to allow re-building without
lcp
parents:
diff changeset
    74
efc648d29dd0 ZF/Inductive.thy,.ML: renamed from "inductive" to allow re-building without
lcp
parents:
diff changeset
    75
structure Quine_Prod =
efc648d29dd0 ZF/Inductive.thy,.ML: renamed from "inductive" to allow re-building without
lcp
parents:
diff changeset
    76
  struct
6093
87bf8c03b169 eliminated global/local names;
wenzelm
parents: 6053
diff changeset
    77
  val sigma     = Const("QPair.QSigma", [iT, iT-->iT]--->iT)
87bf8c03b169 eliminated global/local names;
wenzelm
parents: 6053
diff changeset
    78
  val pair      = Const("QPair.QPair", [iT,iT]--->iT)
87bf8c03b169 eliminated global/local names;
wenzelm
parents: 6053
diff changeset
    79
  val split_name = "QPair.qsplit"
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1418
diff changeset
    80
  val pair_iff  = QPair_iff
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1418
diff changeset
    81
  val split_eq  = qsplit
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1418
diff changeset
    82
  val fsplitI   = qsplitI
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1418
diff changeset
    83
  val fsplitD   = qsplitD
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1418
diff changeset
    84
  val fsplitE   = qsplitE
578
efc648d29dd0 ZF/Inductive.thy,.ML: renamed from "inductive" to allow re-building without
lcp
parents:
diff changeset
    85
  end;
efc648d29dd0 ZF/Inductive.thy,.ML: renamed from "inductive" to allow re-building without
lcp
parents:
diff changeset
    86
1737
5a4f382455ce Modified to use new functor signatures
paulson
parents: 1461
diff changeset
    87
structure Quine_CP = CartProd_Fun (Quine_Prod);
5a4f382455ce Modified to use new functor signatures
paulson
parents: 1461
diff changeset
    88
578
efc648d29dd0 ZF/Inductive.thy,.ML: renamed from "inductive" to allow re-building without
lcp
parents:
diff changeset
    89
structure Quine_Sum =
efc648d29dd0 ZF/Inductive.thy,.ML: renamed from "inductive" to allow re-building without
lcp
parents:
diff changeset
    90
  struct
6093
87bf8c03b169 eliminated global/local names;
wenzelm
parents: 6053
diff changeset
    91
  val sum       = Const("QPair.op <+>", [iT,iT]--->iT)
87bf8c03b169 eliminated global/local names;
wenzelm
parents: 6053
diff changeset
    92
  val inl       = Const("QPair.QInl", iT-->iT)
87bf8c03b169 eliminated global/local names;
wenzelm
parents: 6053
diff changeset
    93
  val inr       = Const("QPair.QInr", iT-->iT)
87bf8c03b169 eliminated global/local names;
wenzelm
parents: 6053
diff changeset
    94
  val elim      = Const("QPair.qcase", [iT-->iT, iT-->iT, iT]--->iT)
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1418
diff changeset
    95
  val case_inl  = qcase_QInl
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1418
diff changeset
    96
  val case_inr  = qcase_QInr
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1418
diff changeset
    97
  val inl_iff   = QInl_iff
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1418
diff changeset
    98
  val inr_iff   = QInr_iff
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1418
diff changeset
    99
  val distinct  = QInl_QInr_iff
578
efc648d29dd0 ZF/Inductive.thy,.ML: renamed from "inductive" to allow re-building without
lcp
parents:
diff changeset
   100
  val distinct' = QInr_QInl_iff
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 10216
diff changeset
   101
  val free_SEs  = Ind_Syntax.mk_free_SEs
1418
f5f97ee67cbb Improving space efficiency of inductive/datatype definitions.
paulson
parents: 1093
diff changeset
   102
            [distinct, distinct', inl_iff, inr_iff, Quine_Prod.pair_iff]
578
efc648d29dd0 ZF/Inductive.thy,.ML: renamed from "inductive" to allow re-building without
lcp
parents:
diff changeset
   103
  end;
efc648d29dd0 ZF/Inductive.thy,.ML: renamed from "inductive" to allow re-building without
lcp
parents:
diff changeset
   104
efc648d29dd0 ZF/Inductive.thy,.ML: renamed from "inductive" to allow re-building without
lcp
parents:
diff changeset
   105
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 10216
diff changeset
   106
structure CoInd_Package =
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 10216
diff changeset
   107
  Add_inductive_def_Fun(structure Fp=Gfp and Pr=Quine_Prod and CP=Quine_CP
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 10216
diff changeset
   108
    and Su=Quine_Sum val coind = true);