src/ZF/Inductive.thy
author wenzelm
Fri, 20 Jan 2023 19:52:52 +0100
changeset 77028 f5896dea6fce
parent 76213 e44d86131648
permissions -rw-r--r--
more direct check of bibtex entries via Isabelle/Scala;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
68490
eb53f944c8cd simplified ZF theory names (in contrast to 6a0801279f4c): session-qualification already achieves disjointness;
wenzelm
parents: 63435
diff changeset
     1
(*  Title:      ZF/Inductive.thy
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
     2
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
     3
    Copyright   1993  University of Cambridge
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
     4
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
     5
Inductive definitions use least fixedpoints with standard products and sums
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
     6
Coinductive definitions use greatest fixedpoints with Quine products and sums
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
     7
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
     8
Sums are used only for mutual recursion;
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
     9
Products are used only to derive "streamlined" induction rules for relations
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    10
*)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    11
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 58871
diff changeset
    12
section\<open>Inductive and Coinductive Definitions\<close>
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    13
68490
eb53f944c8cd simplified ZF theory names (in contrast to 6a0801279f4c): session-qualification already achieves disjointness;
wenzelm
parents: 63435
diff changeset
    14
theory Inductive
eb53f944c8cd simplified ZF theory names (in contrast to 6a0801279f4c): session-qualification already achieves disjointness;
wenzelm
parents: 63435
diff changeset
    15
imports Fixedpt QPair Nat
46947
b8c7eb0c2f89 declare minor keywords via theory header;
wenzelm
parents: 46821
diff changeset
    16
keywords
56146
8453d35e4684 discontinued somewhat pointless "thy_script" keyword kind;
wenzelm
parents: 48891
diff changeset
    17
  "inductive" "coinductive" "inductive_cases" "rep_datatype" "primrec" :: thy_decl and
46950
d0181abdbdac declare command keywords via theory header, including strict checking outside Pure;
wenzelm
parents: 46947
diff changeset
    18
  "domains" "intros" "monos" "con_defs" "type_intros" "type_elims"
63435
7743df69a6b4 clarified keywords;
wenzelm
parents: 60770
diff changeset
    19
    "elimination" "induction" "case_eqns" "recursor_eqns" :: quasi_command
26189
9808cca5c54d misc cleanup of embedded ML code;
wenzelm
parents: 26056
diff changeset
    20
begin
9808cca5c54d misc cleanup of embedded ML code;
wenzelm
parents: 26056
diff changeset
    21
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 74296
diff changeset
    22
lemma def_swap_iff: "a \<equiv> b \<Longrightarrow> a = c \<longleftrightarrow> c = b"
26189
9808cca5c54d misc cleanup of embedded ML code;
wenzelm
parents: 26056
diff changeset
    23
  by blast
9808cca5c54d misc cleanup of embedded ML code;
wenzelm
parents: 26056
diff changeset
    24
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 74296
diff changeset
    25
lemma def_trans: "f \<equiv> g \<Longrightarrow> g(a) = b \<Longrightarrow> f(a) = b"
26189
9808cca5c54d misc cleanup of embedded ML code;
wenzelm
parents: 26056
diff changeset
    26
  by simp
9808cca5c54d misc cleanup of embedded ML code;
wenzelm
parents: 26056
diff changeset
    27
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 74296
diff changeset
    28
lemma refl_thin: "\<And>P. a = a \<Longrightarrow> P \<Longrightarrow> P" .
26189
9808cca5c54d misc cleanup of embedded ML code;
wenzelm
parents: 26056
diff changeset
    29
69605
a96320074298 isabelle update -u path_cartouches;
wenzelm
parents: 69597
diff changeset
    30
ML_file \<open>ind_syntax.ML\<close>
a96320074298 isabelle update -u path_cartouches;
wenzelm
parents: 69597
diff changeset
    31
ML_file \<open>Tools/ind_cases.ML\<close>
a96320074298 isabelle update -u path_cartouches;
wenzelm
parents: 69597
diff changeset
    32
ML_file \<open>Tools/cartprod.ML\<close>
a96320074298 isabelle update -u path_cartouches;
wenzelm
parents: 69597
diff changeset
    33
ML_file \<open>Tools/inductive_package.ML\<close>
a96320074298 isabelle update -u path_cartouches;
wenzelm
parents: 69597
diff changeset
    34
ML_file \<open>Tools/induct_tacs.ML\<close>
a96320074298 isabelle update -u path_cartouches;
wenzelm
parents: 69597
diff changeset
    35
ML_file \<open>Tools/primrec_package.ML\<close>
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    36
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 58871
diff changeset
    37
ML \<open>
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    38
structure Lfp =
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    39
  struct
74296
abc878973216 clarified antiquotation;
wenzelm
parents: 69605
diff changeset
    40
  val oper      = \<^Const>\<open>lfp\<close>
abc878973216 clarified antiquotation;
wenzelm
parents: 69605
diff changeset
    41
  val bnd_mono  = \<^Const>\<open>bnd_mono\<close>
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    42
  val bnd_monoI = @{thm bnd_monoI}
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    43
  val subs      = @{thm def_lfp_subset}
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    44
  val Tarski    = @{thm def_lfp_unfold}
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    45
  val induct    = @{thm def_induct}
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    46
  end;
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    47
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    48
structure Standard_Prod =
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    49
  struct
74296
abc878973216 clarified antiquotation;
wenzelm
parents: 69605
diff changeset
    50
  val sigma     = \<^Const>\<open>Sigma\<close>
abc878973216 clarified antiquotation;
wenzelm
parents: 69605
diff changeset
    51
  val pair      = \<^Const>\<open>Pair\<close>
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 68490
diff changeset
    52
  val split_name = \<^const_name>\<open>split\<close>
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    53
  val pair_iff  = @{thm Pair_iff}
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    54
  val split_eq  = @{thm split}
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    55
  val fsplitI   = @{thm splitI}
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    56
  val fsplitD   = @{thm splitD}
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    57
  val fsplitE   = @{thm splitE}
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    58
  end;
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    59
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    60
structure Standard_CP = CartProd_Fun (Standard_Prod);
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    61
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    62
structure Standard_Sum =
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    63
  struct
74296
abc878973216 clarified antiquotation;
wenzelm
parents: 69605
diff changeset
    64
  val sum       = \<^Const>\<open>sum\<close>
abc878973216 clarified antiquotation;
wenzelm
parents: 69605
diff changeset
    65
  val inl       = \<^Const>\<open>Inl\<close>
abc878973216 clarified antiquotation;
wenzelm
parents: 69605
diff changeset
    66
  val inr       = \<^Const>\<open>Inr\<close>
abc878973216 clarified antiquotation;
wenzelm
parents: 69605
diff changeset
    67
  val elim      = \<^Const>\<open>case\<close>
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    68
  val case_inl  = @{thm case_Inl}
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    69
  val case_inr  = @{thm case_Inr}
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    70
  val inl_iff   = @{thm Inl_iff}
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    71
  val inr_iff   = @{thm Inr_iff}
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    72
  val distinct  = @{thm Inl_Inr_iff}
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    73
  val distinct' = @{thm Inr_Inl_iff}
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    74
  val free_SEs  = Ind_Syntax.mk_free_SEs
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    75
            [distinct, distinct', inl_iff, inr_iff, Standard_Prod.pair_iff]
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    76
  end;
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    77
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    78
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    79
structure Ind_Package =
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    80
    Add_inductive_def_Fun
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    81
      (structure Fp=Lfp and Pr=Standard_Prod and CP=Standard_CP
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    82
       and Su=Standard_Sum val coind = false);
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    83
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    84
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    85
structure Gfp =
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    86
  struct
74296
abc878973216 clarified antiquotation;
wenzelm
parents: 69605
diff changeset
    87
  val oper      = \<^Const>\<open>gfp\<close>
abc878973216 clarified antiquotation;
wenzelm
parents: 69605
diff changeset
    88
  val bnd_mono  = \<^Const>\<open>bnd_mono\<close>
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    89
  val bnd_monoI = @{thm bnd_monoI}
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    90
  val subs      = @{thm def_gfp_subset}
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    91
  val Tarski    = @{thm def_gfp_unfold}
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    92
  val induct    = @{thm def_Collect_coinduct}
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    93
  end;
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    94
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    95
structure Quine_Prod =
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    96
  struct
74296
abc878973216 clarified antiquotation;
wenzelm
parents: 69605
diff changeset
    97
  val sigma     = \<^Const>\<open>QSigma\<close>
abc878973216 clarified antiquotation;
wenzelm
parents: 69605
diff changeset
    98
  val pair      = \<^Const>\<open>QPair\<close>
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 68490
diff changeset
    99
  val split_name = \<^const_name>\<open>qsplit\<close>
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   100
  val pair_iff  = @{thm QPair_iff}
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   101
  val split_eq  = @{thm qsplit}
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   102
  val fsplitI   = @{thm qsplitI}
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   103
  val fsplitD   = @{thm qsplitD}
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   104
  val fsplitE   = @{thm qsplitE}
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   105
  end;
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   106
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   107
structure Quine_CP = CartProd_Fun (Quine_Prod);
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   108
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   109
structure Quine_Sum =
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   110
  struct
74296
abc878973216 clarified antiquotation;
wenzelm
parents: 69605
diff changeset
   111
  val sum       = \<^Const>\<open>qsum\<close>
abc878973216 clarified antiquotation;
wenzelm
parents: 69605
diff changeset
   112
  val inl       = \<^Const>\<open>QInl\<close>
abc878973216 clarified antiquotation;
wenzelm
parents: 69605
diff changeset
   113
  val inr       = \<^Const>\<open>QInr\<close>
abc878973216 clarified antiquotation;
wenzelm
parents: 69605
diff changeset
   114
  val elim      = \<^Const>\<open>qcase\<close>
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   115
  val case_inl  = @{thm qcase_QInl}
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   116
  val case_inr  = @{thm qcase_QInr}
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   117
  val inl_iff   = @{thm QInl_iff}
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   118
  val inr_iff   = @{thm QInr_iff}
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   119
  val distinct  = @{thm QInl_QInr_iff}
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   120
  val distinct' = @{thm QInr_QInl_iff}
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   121
  val free_SEs  = Ind_Syntax.mk_free_SEs
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   122
            [distinct, distinct', inl_iff, inr_iff, Quine_Prod.pair_iff]
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   123
  end;
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   124
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   125
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   126
structure CoInd_Package =
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   127
  Add_inductive_def_Fun(structure Fp=Gfp and Pr=Quine_Prod and CP=Quine_CP
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   128
    and Su=Quine_Sum val coind = true);
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   129
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 58871
diff changeset
   130
\<close>
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   131
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
   132
end