src/HOL/Inductive.thy
author immler@in.tum.de
Wed, 03 Jun 2009 16:56:41 +0200
changeset 31409 d8537ba165b5
parent 29281 b22ccb3998db
child 31604 eb2f9d709296
permissions -rw-r--r--
split preparing clauses and writing problemfile; included results of count_constants in return-type of prover; optionally pass counted constants to prover; removed unused external_prover from signature
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
7700
38b6d2643630 load / setup datatype package;
wenzelm
parents: 7357
diff changeset
     1
(*  Title:      HOL/Inductive.thy
10402
5e82d6cafb5f inductive_atomize, inductive_rulify;
wenzelm
parents: 10312
diff changeset
     2
    Author:     Markus Wenzel, TU Muenchen
11688
56833637db2a generic induct_method.ML;
wenzelm
parents: 11439
diff changeset
     3
*)
10727
2ccafccb81c0 added inductive_conj;
wenzelm
parents: 10434
diff changeset
     4
24915
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
     5
header {* Knaster-Tarski Fixpoint Theorem and inductive definitions *}
1187
bc94f00e47ba Includes Sum.thy as a parent for mutual recursion
lcp
parents: 923
diff changeset
     6
15131
c69542757a4d New theory header syntax.
nipkow
parents: 14981
diff changeset
     7
theory Inductive 
24915
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
     8
imports Lattices Sum_Type
16417
9bc16273c2d4 migrated theory headers to new format
haftmann
parents: 15140
diff changeset
     9
uses
10402
5e82d6cafb5f inductive_atomize, inductive_rulify;
wenzelm
parents: 10312
diff changeset
    10
  ("Tools/inductive_package.ML")
24625
0398a5e802d3 separated code for inductive sequences from inductive_codegen.ML
haftmann
parents: 24349
diff changeset
    11
  "Tools/dseq.ML"
12437
6d4e02b6dd43 Moved code generator setup from Recdef to Inductive.
berghofe
parents: 12023
diff changeset
    12
  ("Tools/inductive_codegen.ML")
10402
5e82d6cafb5f inductive_atomize, inductive_rulify;
wenzelm
parents: 10312
diff changeset
    13
  ("Tools/datatype_aux.ML")
5e82d6cafb5f inductive_atomize, inductive_rulify;
wenzelm
parents: 10312
diff changeset
    14
  ("Tools/datatype_prop.ML")
5e82d6cafb5f inductive_atomize, inductive_rulify;
wenzelm
parents: 10312
diff changeset
    15
  ("Tools/datatype_rep_proofs.ML")
5e82d6cafb5f inductive_atomize, inductive_rulify;
wenzelm
parents: 10312
diff changeset
    16
  ("Tools/datatype_abs_proofs.ML")
22783
e5f947e0ade8 Added datatype_case.
berghofe
parents: 22218
diff changeset
    17
  ("Tools/datatype_case.ML")
10402
5e82d6cafb5f inductive_atomize, inductive_rulify;
wenzelm
parents: 10312
diff changeset
    18
  ("Tools/datatype_package.ML")
25557
ea6b11021e79 added new primrec package
haftmann
parents: 25534
diff changeset
    19
  ("Tools/old_primrec_package.ML")
15131
c69542757a4d New theory header syntax.
nipkow
parents: 14981
diff changeset
    20
  ("Tools/primrec_package.ML")
25534
d0b74fdd6067 simplified infrastructure for code generator operational equality
haftmann
parents: 25510
diff changeset
    21
  ("Tools/datatype_codegen.ML")
15131
c69542757a4d New theory header syntax.
nipkow
parents: 14981
diff changeset
    22
begin
10727
2ccafccb81c0 added inductive_conj;
wenzelm
parents: 10434
diff changeset
    23
24915
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
    24
subsection {* Least and greatest fixed points *}
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
    25
26013
8764a1f1253b Theorem Inductive.lfp_ordinal_induct generalized to complete lattices
haftmann
parents: 25557
diff changeset
    26
context complete_lattice
8764a1f1253b Theorem Inductive.lfp_ordinal_induct generalized to complete lattices
haftmann
parents: 25557
diff changeset
    27
begin
8764a1f1253b Theorem Inductive.lfp_ordinal_induct generalized to complete lattices
haftmann
parents: 25557
diff changeset
    28
24915
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
    29
definition
26013
8764a1f1253b Theorem Inductive.lfp_ordinal_induct generalized to complete lattices
haftmann
parents: 25557
diff changeset
    30
  lfp :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a" where
24915
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
    31
  "lfp f = Inf {u. f u \<le> u}"    --{*least fixed point*}
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
    32
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
    33
definition
26013
8764a1f1253b Theorem Inductive.lfp_ordinal_induct generalized to complete lattices
haftmann
parents: 25557
diff changeset
    34
  gfp :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a" where
24915
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
    35
  "gfp f = Sup {u. u \<le> f u}"    --{*greatest fixed point*}
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
    36
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
    37
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
    38
subsection{* Proof of Knaster-Tarski Theorem using @{term lfp} *}
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
    39
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
    40
text{*@{term "lfp f"} is the least upper bound of 
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
    41
      the set @{term "{u. f(u) \<le> u}"} *}
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
    42
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
    43
lemma lfp_lowerbound: "f A \<le> A ==> lfp f \<le> A"
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
    44
  by (auto simp add: lfp_def intro: Inf_lower)
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
    45
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
    46
lemma lfp_greatest: "(!!u. f u \<le> u ==> A \<le> u) ==> A \<le> lfp f"
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
    47
  by (auto simp add: lfp_def intro: Inf_greatest)
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
    48
26013
8764a1f1253b Theorem Inductive.lfp_ordinal_induct generalized to complete lattices
haftmann
parents: 25557
diff changeset
    49
end
8764a1f1253b Theorem Inductive.lfp_ordinal_induct generalized to complete lattices
haftmann
parents: 25557
diff changeset
    50
24915
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
    51
lemma lfp_lemma2: "mono f ==> f (lfp f) \<le> lfp f"
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
    52
  by (iprover intro: lfp_greatest order_trans monoD lfp_lowerbound)
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
    53
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
    54
lemma lfp_lemma3: "mono f ==> lfp f \<le> f (lfp f)"
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
    55
  by (iprover intro: lfp_lemma2 monoD lfp_lowerbound)
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
    56
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
    57
lemma lfp_unfold: "mono f ==> lfp f = f (lfp f)"
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
    58
  by (iprover intro: order_antisym lfp_lemma2 lfp_lemma3)
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
    59
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
    60
lemma lfp_const: "lfp (\<lambda>x. t) = t"
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
    61
  by (rule lfp_unfold) (simp add:mono_def)
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
    62
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
    63
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
    64
subsection {* General induction rules for least fixed points *}
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
    65
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
    66
theorem lfp_induct:
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
    67
  assumes mono: "mono f" and ind: "f (inf (lfp f) P) <= P"
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
    68
  shows "lfp f <= P"
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
    69
proof -
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
    70
  have "inf (lfp f) P <= lfp f" by (rule inf_le1)
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
    71
  with mono have "f (inf (lfp f) P) <= f (lfp f)" ..
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
    72
  also from mono have "f (lfp f) = lfp f" by (rule lfp_unfold [symmetric])
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
    73
  finally have "f (inf (lfp f) P) <= lfp f" .
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
    74
  from this and ind have "f (inf (lfp f) P) <= inf (lfp f) P" by (rule le_infI)
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
    75
  hence "lfp f <= inf (lfp f) P" by (rule lfp_lowerbound)
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
    76
  also have "inf (lfp f) P <= P" by (rule inf_le2)
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
    77
  finally show ?thesis .
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
    78
qed
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
    79
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
    80
lemma lfp_induct_set:
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
    81
  assumes lfp: "a: lfp(f)"
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
    82
      and mono: "mono(f)"
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
    83
      and indhyp: "!!x. [| x: f(lfp(f) Int {x. P(x)}) |] ==> P(x)"
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
    84
  shows "P(a)"
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
    85
  by (rule lfp_induct [THEN subsetD, THEN CollectD, OF mono _ lfp])
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
    86
    (auto simp: inf_set_eq intro: indhyp)
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
    87
26013
8764a1f1253b Theorem Inductive.lfp_ordinal_induct generalized to complete lattices
haftmann
parents: 25557
diff changeset
    88
lemma lfp_ordinal_induct:
8764a1f1253b Theorem Inductive.lfp_ordinal_induct generalized to complete lattices
haftmann
parents: 25557
diff changeset
    89
  fixes f :: "'a\<Colon>complete_lattice \<Rightarrow> 'a"
8764a1f1253b Theorem Inductive.lfp_ordinal_induct generalized to complete lattices
haftmann
parents: 25557
diff changeset
    90
  assumes mono: "mono f"
8764a1f1253b Theorem Inductive.lfp_ordinal_induct generalized to complete lattices
haftmann
parents: 25557
diff changeset
    91
  and P_f: "\<And>S. P S \<Longrightarrow> P (f S)"
8764a1f1253b Theorem Inductive.lfp_ordinal_induct generalized to complete lattices
haftmann
parents: 25557
diff changeset
    92
  and P_Union: "\<And>M. \<forall>S\<in>M. P S \<Longrightarrow> P (Sup M)"
8764a1f1253b Theorem Inductive.lfp_ordinal_induct generalized to complete lattices
haftmann
parents: 25557
diff changeset
    93
  shows "P (lfp f)"
8764a1f1253b Theorem Inductive.lfp_ordinal_induct generalized to complete lattices
haftmann
parents: 25557
diff changeset
    94
proof -
8764a1f1253b Theorem Inductive.lfp_ordinal_induct generalized to complete lattices
haftmann
parents: 25557
diff changeset
    95
  let ?M = "{S. S \<le> lfp f \<and> P S}"
8764a1f1253b Theorem Inductive.lfp_ordinal_induct generalized to complete lattices
haftmann
parents: 25557
diff changeset
    96
  have "P (Sup ?M)" using P_Union by simp
8764a1f1253b Theorem Inductive.lfp_ordinal_induct generalized to complete lattices
haftmann
parents: 25557
diff changeset
    97
  also have "Sup ?M = lfp f"
8764a1f1253b Theorem Inductive.lfp_ordinal_induct generalized to complete lattices
haftmann
parents: 25557
diff changeset
    98
  proof (rule antisym)
8764a1f1253b Theorem Inductive.lfp_ordinal_induct generalized to complete lattices
haftmann
parents: 25557
diff changeset
    99
    show "Sup ?M \<le> lfp f" by (blast intro: Sup_least)
8764a1f1253b Theorem Inductive.lfp_ordinal_induct generalized to complete lattices
haftmann
parents: 25557
diff changeset
   100
    hence "f (Sup ?M) \<le> f (lfp f)" by (rule mono [THEN monoD])
8764a1f1253b Theorem Inductive.lfp_ordinal_induct generalized to complete lattices
haftmann
parents: 25557
diff changeset
   101
    hence "f (Sup ?M) \<le> lfp f" using mono [THEN lfp_unfold] by simp
8764a1f1253b Theorem Inductive.lfp_ordinal_induct generalized to complete lattices
haftmann
parents: 25557
diff changeset
   102
    hence "f (Sup ?M) \<in> ?M" using P_f P_Union by simp
8764a1f1253b Theorem Inductive.lfp_ordinal_induct generalized to complete lattices
haftmann
parents: 25557
diff changeset
   103
    hence "f (Sup ?M) \<le> Sup ?M" by (rule Sup_upper)
8764a1f1253b Theorem Inductive.lfp_ordinal_induct generalized to complete lattices
haftmann
parents: 25557
diff changeset
   104
    thus "lfp f \<le> Sup ?M" by (rule lfp_lowerbound)
8764a1f1253b Theorem Inductive.lfp_ordinal_induct generalized to complete lattices
haftmann
parents: 25557
diff changeset
   105
  qed
8764a1f1253b Theorem Inductive.lfp_ordinal_induct generalized to complete lattices
haftmann
parents: 25557
diff changeset
   106
  finally show ?thesis .
8764a1f1253b Theorem Inductive.lfp_ordinal_induct generalized to complete lattices
haftmann
parents: 25557
diff changeset
   107
qed 
8764a1f1253b Theorem Inductive.lfp_ordinal_induct generalized to complete lattices
haftmann
parents: 25557
diff changeset
   108
8764a1f1253b Theorem Inductive.lfp_ordinal_induct generalized to complete lattices
haftmann
parents: 25557
diff changeset
   109
lemma lfp_ordinal_induct_set: 
24915
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   110
  assumes mono: "mono f"
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   111
  and P_f: "!!S. P S ==> P(f S)"
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   112
  and P_Union: "!!M. !S:M. P S ==> P(Union M)"
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   113
  shows "P(lfp f)"
26793
e36a92ff543e Instantiated some rules to avoid problems with HO unification.
berghofe
parents: 26013
diff changeset
   114
  using assms unfolding Sup_set_eq [symmetric]
e36a92ff543e Instantiated some rules to avoid problems with HO unification.
berghofe
parents: 26013
diff changeset
   115
  by (rule lfp_ordinal_induct [where P=P])
24915
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   116
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   117
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   118
text{*Definition forms of @{text lfp_unfold} and @{text lfp_induct}, 
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   119
    to control unfolding*}
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   120
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   121
lemma def_lfp_unfold: "[| h==lfp(f);  mono(f) |] ==> h = f(h)"
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   122
by (auto intro!: lfp_unfold)
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   123
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   124
lemma def_lfp_induct: 
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   125
    "[| A == lfp(f); mono(f);
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   126
        f (inf A P) \<le> P
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   127
     |] ==> A \<le> P"
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   128
  by (blast intro: lfp_induct)
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   129
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   130
lemma def_lfp_induct_set: 
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   131
    "[| A == lfp(f);  mono(f);   a:A;                    
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   132
        !!x. [| x: f(A Int {x. P(x)}) |] ==> P(x)         
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   133
     |] ==> P(a)"
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   134
  by (blast intro: lfp_induct_set)
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   135
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   136
(*Monotonicity of lfp!*)
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   137
lemma lfp_mono: "(!!Z. f Z \<le> g Z) ==> lfp f \<le> lfp g"
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   138
  by (rule lfp_lowerbound [THEN lfp_greatest], blast intro: order_trans)
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   139
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   140
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   141
subsection {* Proof of Knaster-Tarski Theorem using @{term gfp} *}
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   142
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   143
text{*@{term "gfp f"} is the greatest lower bound of 
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   144
      the set @{term "{u. u \<le> f(u)}"} *}
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   145
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   146
lemma gfp_upperbound: "X \<le> f X ==> X \<le> gfp f"
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   147
  by (auto simp add: gfp_def intro: Sup_upper)
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   148
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   149
lemma gfp_least: "(!!u. u \<le> f u ==> u \<le> X) ==> gfp f \<le> X"
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   150
  by (auto simp add: gfp_def intro: Sup_least)
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   151
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   152
lemma gfp_lemma2: "mono f ==> gfp f \<le> f (gfp f)"
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   153
  by (iprover intro: gfp_least order_trans monoD gfp_upperbound)
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   154
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   155
lemma gfp_lemma3: "mono f ==> f (gfp f) \<le> gfp f"
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   156
  by (iprover intro: gfp_lemma2 monoD gfp_upperbound)
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   157
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   158
lemma gfp_unfold: "mono f ==> gfp f = f (gfp f)"
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   159
  by (iprover intro: order_antisym gfp_lemma2 gfp_lemma3)
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   160
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   161
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   162
subsection {* Coinduction rules for greatest fixed points *}
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   163
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   164
text{*weak version*}
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   165
lemma weak_coinduct: "[| a: X;  X \<subseteq> f(X) |] ==> a : gfp(f)"
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   166
by (rule gfp_upperbound [THEN subsetD], auto)
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   167
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   168
lemma weak_coinduct_image: "!!X. [| a : X; g`X \<subseteq> f (g`X) |] ==> g a : gfp f"
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   169
apply (erule gfp_upperbound [THEN subsetD])
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   170
apply (erule imageI)
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   171
done
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   172
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   173
lemma coinduct_lemma:
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   174
     "[| X \<le> f (sup X (gfp f));  mono f |] ==> sup X (gfp f) \<le> f (sup X (gfp f))"
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   175
  apply (frule gfp_lemma2)
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   176
  apply (drule mono_sup)
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   177
  apply (rule le_supI)
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   178
  apply assumption
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   179
  apply (rule order_trans)
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   180
  apply (rule order_trans)
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   181
  apply assumption
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   182
  apply (rule sup_ge2)
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   183
  apply assumption
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   184
  done
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   185
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   186
text{*strong version, thanks to Coen and Frost*}
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   187
lemma coinduct_set: "[| mono(f);  a: X;  X \<subseteq> f(X Un gfp(f)) |] ==> a : gfp(f)"
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   188
by (blast intro: weak_coinduct [OF _ coinduct_lemma, simplified sup_set_eq])
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   189
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   190
lemma coinduct: "[| mono(f); X \<le> f (sup X (gfp f)) |] ==> X \<le> gfp(f)"
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   191
  apply (rule order_trans)
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   192
  apply (rule sup_ge1)
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   193
  apply (erule gfp_upperbound [OF coinduct_lemma])
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   194
  apply assumption
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   195
  done
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   196
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   197
lemma gfp_fun_UnI2: "[| mono(f);  a: gfp(f) |] ==> a: f(X Un gfp(f))"
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   198
by (blast dest: gfp_lemma2 mono_Un)
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   199
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   200
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   201
subsection {* Even Stronger Coinduction Rule, by Martin Coen *}
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   202
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   203
text{* Weakens the condition @{term "X \<subseteq> f(X)"} to one expressed using both
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   204
  @{term lfp} and @{term gfp}*}
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   205
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   206
lemma coinduct3_mono_lemma: "mono(f) ==> mono(%x. f(x) Un X Un B)"
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   207
by (iprover intro: subset_refl monoI Un_mono monoD)
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   208
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   209
lemma coinduct3_lemma:
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   210
     "[| X \<subseteq> f(lfp(%x. f(x) Un X Un gfp(f)));  mono(f) |]
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   211
      ==> lfp(%x. f(x) Un X Un gfp(f)) \<subseteq> f(lfp(%x. f(x) Un X Un gfp(f)))"
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   212
apply (rule subset_trans)
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   213
apply (erule coinduct3_mono_lemma [THEN lfp_lemma3])
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   214
apply (rule Un_least [THEN Un_least])
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   215
apply (rule subset_refl, assumption)
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   216
apply (rule gfp_unfold [THEN equalityD1, THEN subset_trans], assumption)
26793
e36a92ff543e Instantiated some rules to avoid problems with HO unification.
berghofe
parents: 26013
diff changeset
   217
apply (rule monoD [where f=f], assumption)
24915
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   218
apply (subst coinduct3_mono_lemma [THEN lfp_unfold], auto)
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   219
done
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   220
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   221
lemma coinduct3: 
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   222
  "[| mono(f);  a:X;  X \<subseteq> f(lfp(%x. f(x) Un X Un gfp(f))) |] ==> a : gfp(f)"
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   223
apply (rule coinduct3_lemma [THEN [2] weak_coinduct])
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   224
apply (rule coinduct3_mono_lemma [THEN lfp_unfold, THEN ssubst], auto)
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   225
done
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   226
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   227
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   228
text{*Definition forms of @{text gfp_unfold} and @{text coinduct}, 
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   229
    to control unfolding*}
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   230
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   231
lemma def_gfp_unfold: "[| A==gfp(f);  mono(f) |] ==> A = f(A)"
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   232
by (auto intro!: gfp_unfold)
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   233
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   234
lemma def_coinduct:
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   235
     "[| A==gfp(f);  mono(f);  X \<le> f(sup X A) |] ==> X \<le> A"
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   236
by (iprover intro!: coinduct)
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   237
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   238
lemma def_coinduct_set:
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   239
     "[| A==gfp(f);  mono(f);  a:X;  X \<subseteq> f(X Un A) |] ==> a: A"
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   240
by (auto intro!: coinduct_set)
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   241
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   242
(*The version used in the induction/coinduction package*)
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   243
lemma def_Collect_coinduct:
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   244
    "[| A == gfp(%w. Collect(P(w)));  mono(%w. Collect(P(w)));   
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   245
        a: X;  !!z. z: X ==> P (X Un A) z |] ==>  
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   246
     a : A"
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   247
apply (erule def_coinduct_set, auto) 
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   248
done
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   249
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   250
lemma def_coinduct3:
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   251
    "[| A==gfp(f); mono(f);  a:X;  X \<subseteq> f(lfp(%x. f(x) Un X Un A)) |] ==> a: A"
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   252
by (auto intro!: coinduct3)
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   253
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   254
text{*Monotonicity of @{term gfp}!*}
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   255
lemma gfp_mono: "(!!Z. f Z \<le> g Z) ==> gfp f \<le> gfp g"
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   256
  by (rule gfp_upperbound [THEN gfp_least], blast intro: order_trans)
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   257
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   258
23734
0e11b904b3a3 Added new package for inductive sets.
berghofe
parents: 23708
diff changeset
   259
subsection {* Inductive predicates and sets *}
11688
56833637db2a generic induct_method.ML;
wenzelm
parents: 11439
diff changeset
   260
56833637db2a generic induct_method.ML;
wenzelm
parents: 11439
diff changeset
   261
text {* Inversion of injective functions. *}
11436
3f74b80d979f private "myinv" (uses "The" instead of "Eps");
wenzelm
parents: 11325
diff changeset
   262
3f74b80d979f private "myinv" (uses "The" instead of "Eps");
wenzelm
parents: 11325
diff changeset
   263
constdefs
3f74b80d979f private "myinv" (uses "The" instead of "Eps");
wenzelm
parents: 11325
diff changeset
   264
  myinv :: "('a => 'b) => ('b => 'a)"
3f74b80d979f private "myinv" (uses "The" instead of "Eps");
wenzelm
parents: 11325
diff changeset
   265
  "myinv (f :: 'a => 'b) == \<lambda>y. THE x. f x = y"
3f74b80d979f private "myinv" (uses "The" instead of "Eps");
wenzelm
parents: 11325
diff changeset
   266
3f74b80d979f private "myinv" (uses "The" instead of "Eps");
wenzelm
parents: 11325
diff changeset
   267
lemma myinv_f_f: "inj f ==> myinv f (f x) = x"
3f74b80d979f private "myinv" (uses "The" instead of "Eps");
wenzelm
parents: 11325
diff changeset
   268
proof -
3f74b80d979f private "myinv" (uses "The" instead of "Eps");
wenzelm
parents: 11325
diff changeset
   269
  assume "inj f"
3f74b80d979f private "myinv" (uses "The" instead of "Eps");
wenzelm
parents: 11325
diff changeset
   270
  hence "(THE x'. f x' = f x) = (THE x'. x' = x)"
3f74b80d979f private "myinv" (uses "The" instead of "Eps");
wenzelm
parents: 11325
diff changeset
   271
    by (simp only: inj_eq)
3f74b80d979f private "myinv" (uses "The" instead of "Eps");
wenzelm
parents: 11325
diff changeset
   272
  also have "... = x" by (rule the_eq_trivial)
11439
wenzelm
parents: 11436
diff changeset
   273
  finally show ?thesis by (unfold myinv_def)
11436
3f74b80d979f private "myinv" (uses "The" instead of "Eps");
wenzelm
parents: 11325
diff changeset
   274
qed
3f74b80d979f private "myinv" (uses "The" instead of "Eps");
wenzelm
parents: 11325
diff changeset
   275
3f74b80d979f private "myinv" (uses "The" instead of "Eps");
wenzelm
parents: 11325
diff changeset
   276
lemma f_myinv_f: "inj f ==> y \<in> range f ==> f (myinv f y) = y"
3f74b80d979f private "myinv" (uses "The" instead of "Eps");
wenzelm
parents: 11325
diff changeset
   277
proof (unfold myinv_def)
3f74b80d979f private "myinv" (uses "The" instead of "Eps");
wenzelm
parents: 11325
diff changeset
   278
  assume inj: "inj f"
3f74b80d979f private "myinv" (uses "The" instead of "Eps");
wenzelm
parents: 11325
diff changeset
   279
  assume "y \<in> range f"
3f74b80d979f private "myinv" (uses "The" instead of "Eps");
wenzelm
parents: 11325
diff changeset
   280
  then obtain x where "y = f x" ..
3f74b80d979f private "myinv" (uses "The" instead of "Eps");
wenzelm
parents: 11325
diff changeset
   281
  hence x: "f x = y" ..
3f74b80d979f private "myinv" (uses "The" instead of "Eps");
wenzelm
parents: 11325
diff changeset
   282
  thus "f (THE x. f x = y) = y"
3f74b80d979f private "myinv" (uses "The" instead of "Eps");
wenzelm
parents: 11325
diff changeset
   283
  proof (rule theI)
3f74b80d979f private "myinv" (uses "The" instead of "Eps");
wenzelm
parents: 11325
diff changeset
   284
    fix x' assume "f x' = y"
3f74b80d979f private "myinv" (uses "The" instead of "Eps");
wenzelm
parents: 11325
diff changeset
   285
    with x have "f x' = f x" by simp
3f74b80d979f private "myinv" (uses "The" instead of "Eps");
wenzelm
parents: 11325
diff changeset
   286
    with inj show "x' = x" by (rule injD)
3f74b80d979f private "myinv" (uses "The" instead of "Eps");
wenzelm
parents: 11325
diff changeset
   287
  qed
3f74b80d979f private "myinv" (uses "The" instead of "Eps");
wenzelm
parents: 11325
diff changeset
   288
qed
3f74b80d979f private "myinv" (uses "The" instead of "Eps");
wenzelm
parents: 11325
diff changeset
   289
3f74b80d979f private "myinv" (uses "The" instead of "Eps");
wenzelm
parents: 11325
diff changeset
   290
hide const myinv
3f74b80d979f private "myinv" (uses "The" instead of "Eps");
wenzelm
parents: 11325
diff changeset
   291
3f74b80d979f private "myinv" (uses "The" instead of "Eps");
wenzelm
parents: 11325
diff changeset
   292
11688
56833637db2a generic induct_method.ML;
wenzelm
parents: 11439
diff changeset
   293
text {* Package setup. *}
10402
5e82d6cafb5f inductive_atomize, inductive_rulify;
wenzelm
parents: 10312
diff changeset
   294
23734
0e11b904b3a3 Added new package for inductive sets.
berghofe
parents: 23708
diff changeset
   295
theorems basic_monos =
22218
30a8890d2967 dropped lemma duplicates in HOL.thy
haftmann
parents: 21018
diff changeset
   296
  subset_refl imp_refl disj_mono conj_mono ex_mono all_mono if_bool_eq_conj
11688
56833637db2a generic induct_method.ML;
wenzelm
parents: 11439
diff changeset
   297
  Collect_mono in_mono vimage_mono
56833637db2a generic induct_method.ML;
wenzelm
parents: 11439
diff changeset
   298
  imp_conv_disj not_not de_Morgan_disj de_Morgan_conj
56833637db2a generic induct_method.ML;
wenzelm
parents: 11439
diff changeset
   299
  not_all not_ex
56833637db2a generic induct_method.ML;
wenzelm
parents: 11439
diff changeset
   300
  Ball_def Bex_def
18456
8cc35e95450a updated auxiliary facts for induct method;
wenzelm
parents: 17009
diff changeset
   301
  induct_rulify_fallback
11688
56833637db2a generic induct_method.ML;
wenzelm
parents: 11439
diff changeset
   302
24915
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   303
ML {*
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   304
val def_lfp_unfold = @{thm def_lfp_unfold}
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   305
val def_gfp_unfold = @{thm def_gfp_unfold}
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   306
val def_lfp_induct = @{thm def_lfp_induct}
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   307
val def_coinduct = @{thm def_coinduct}
25510
38c15efe603b adjustions to due to instance target
haftmann
parents: 24915
diff changeset
   308
val inf_bool_eq = @{thm inf_bool_eq} RS @{thm eq_reflection}
38c15efe603b adjustions to due to instance target
haftmann
parents: 24915
diff changeset
   309
val inf_fun_eq = @{thm inf_fun_eq} RS @{thm eq_reflection}
38c15efe603b adjustions to due to instance target
haftmann
parents: 24915
diff changeset
   310
val sup_bool_eq = @{thm sup_bool_eq} RS @{thm eq_reflection}
38c15efe603b adjustions to due to instance target
haftmann
parents: 24915
diff changeset
   311
val sup_fun_eq = @{thm sup_fun_eq} RS @{thm eq_reflection}
24915
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   312
val le_boolI = @{thm le_boolI}
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   313
val le_boolI' = @{thm le_boolI'}
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   314
val le_funI = @{thm le_funI}
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   315
val le_boolE = @{thm le_boolE}
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   316
val le_funE = @{thm le_funE}
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   317
val le_boolD = @{thm le_boolD}
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   318
val le_funD = @{thm le_funD}
25510
38c15efe603b adjustions to due to instance target
haftmann
parents: 24915
diff changeset
   319
val le_bool_def = @{thm le_bool_def} RS @{thm eq_reflection}
38c15efe603b adjustions to due to instance target
haftmann
parents: 24915
diff changeset
   320
val le_fun_def = @{thm le_fun_def} RS @{thm eq_reflection}
24915
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   321
*}
fc90277c0dd7 integrated FixedPoint into Inductive
haftmann
parents: 24845
diff changeset
   322
21018
e6b8d6784792 Added new package for inductive definitions, moved old package
berghofe
parents: 20604
diff changeset
   323
use "Tools/inductive_package.ML"
e6b8d6784792 Added new package for inductive definitions, moved old package
berghofe
parents: 20604
diff changeset
   324
setup InductivePackage.setup
e6b8d6784792 Added new package for inductive definitions, moved old package
berghofe
parents: 20604
diff changeset
   325
23734
0e11b904b3a3 Added new package for inductive sets.
berghofe
parents: 23708
diff changeset
   326
theorems [mono] =
22218
30a8890d2967 dropped lemma duplicates in HOL.thy
haftmann
parents: 21018
diff changeset
   327
  imp_refl disj_mono conj_mono ex_mono all_mono if_bool_eq_conj
21018
e6b8d6784792 Added new package for inductive definitions, moved old package
berghofe
parents: 20604
diff changeset
   328
  imp_conv_disj not_not de_Morgan_disj de_Morgan_conj
e6b8d6784792 Added new package for inductive definitions, moved old package
berghofe
parents: 20604
diff changeset
   329
  not_all not_ex
e6b8d6784792 Added new package for inductive definitions, moved old package
berghofe
parents: 20604
diff changeset
   330
  Ball_def Bex_def
e6b8d6784792 Added new package for inductive definitions, moved old package
berghofe
parents: 20604
diff changeset
   331
  induct_rulify_fallback
e6b8d6784792 Added new package for inductive definitions, moved old package
berghofe
parents: 20604
diff changeset
   332
11688
56833637db2a generic induct_method.ML;
wenzelm
parents: 11439
diff changeset
   333
12023
wenzelm
parents: 11990
diff changeset
   334
subsection {* Inductive datatypes and primitive recursion *}
11688
56833637db2a generic induct_method.ML;
wenzelm
parents: 11439
diff changeset
   335
11825
ef7d619e2c88 moved InductMethod.setup to theory HOL;
wenzelm
parents: 11688
diff changeset
   336
text {* Package setup. *}
ef7d619e2c88 moved InductMethod.setup to theory HOL;
wenzelm
parents: 11688
diff changeset
   337
10402
5e82d6cafb5f inductive_atomize, inductive_rulify;
wenzelm
parents: 10312
diff changeset
   338
use "Tools/datatype_aux.ML"
5e82d6cafb5f inductive_atomize, inductive_rulify;
wenzelm
parents: 10312
diff changeset
   339
use "Tools/datatype_prop.ML"
5e82d6cafb5f inductive_atomize, inductive_rulify;
wenzelm
parents: 10312
diff changeset
   340
use "Tools/datatype_rep_proofs.ML"
5e82d6cafb5f inductive_atomize, inductive_rulify;
wenzelm
parents: 10312
diff changeset
   341
use "Tools/datatype_abs_proofs.ML"
22783
e5f947e0ade8 Added datatype_case.
berghofe
parents: 22218
diff changeset
   342
use "Tools/datatype_case.ML"
10402
5e82d6cafb5f inductive_atomize, inductive_rulify;
wenzelm
parents: 10312
diff changeset
   343
use "Tools/datatype_package.ML"
7700
38b6d2643630 load / setup datatype package;
wenzelm
parents: 7357
diff changeset
   344
setup DatatypePackage.setup
25557
ea6b11021e79 added new primrec package
haftmann
parents: 25534
diff changeset
   345
use "Tools/old_primrec_package.ML"
24699
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24626
diff changeset
   346
use "Tools/primrec_package.ML"
12437
6d4e02b6dd43 Moved code generator setup from Recdef to Inductive.
berghofe
parents: 12023
diff changeset
   347
25534
d0b74fdd6067 simplified infrastructure for code generator operational equality
haftmann
parents: 25510
diff changeset
   348
use "Tools/datatype_codegen.ML"
d0b74fdd6067 simplified infrastructure for code generator operational equality
haftmann
parents: 25510
diff changeset
   349
setup DatatypeCodegen.setup
d0b74fdd6067 simplified infrastructure for code generator operational equality
haftmann
parents: 25510
diff changeset
   350
12437
6d4e02b6dd43 Moved code generator setup from Recdef to Inductive.
berghofe
parents: 12023
diff changeset
   351
use "Tools/inductive_codegen.ML"
6d4e02b6dd43 Moved code generator setup from Recdef to Inductive.
berghofe
parents: 12023
diff changeset
   352
setup InductiveCodegen.setup
6d4e02b6dd43 Moved code generator setup from Recdef to Inductive.
berghofe
parents: 12023
diff changeset
   353
23526
936dc616a7fb Added pattern maatching for lambda abstraction
nipkow
parents: 23389
diff changeset
   354
text{* Lambda-abstractions with pattern matching: *}
936dc616a7fb Added pattern maatching for lambda abstraction
nipkow
parents: 23389
diff changeset
   355
936dc616a7fb Added pattern maatching for lambda abstraction
nipkow
parents: 23389
diff changeset
   356
syntax
23529
958f9d9cfb63 Fixed problem with patterns in lambdas
nipkow
parents: 23526
diff changeset
   357
  "_lam_pats_syntax" :: "cases_syn => 'a => 'b"               ("(%_)" 10)
23526
936dc616a7fb Added pattern maatching for lambda abstraction
nipkow
parents: 23389
diff changeset
   358
syntax (xsymbols)
23529
958f9d9cfb63 Fixed problem with patterns in lambdas
nipkow
parents: 23526
diff changeset
   359
  "_lam_pats_syntax" :: "cases_syn => 'a => 'b"               ("(\<lambda>_)" 10)
23526
936dc616a7fb Added pattern maatching for lambda abstraction
nipkow
parents: 23389
diff changeset
   360
23529
958f9d9cfb63 Fixed problem with patterns in lambdas
nipkow
parents: 23526
diff changeset
   361
parse_translation (advanced) {*
958f9d9cfb63 Fixed problem with patterns in lambdas
nipkow
parents: 23526
diff changeset
   362
let
958f9d9cfb63 Fixed problem with patterns in lambdas
nipkow
parents: 23526
diff changeset
   363
  fun fun_tr ctxt [cs] =
958f9d9cfb63 Fixed problem with patterns in lambdas
nipkow
parents: 23526
diff changeset
   364
    let
29281
b22ccb3998db eliminated OldTerm.add_term_free_names;
wenzelm
parents: 29270
diff changeset
   365
      val x = Free (Name.variant (Term.add_free_names cs []) "x", dummyT);
24349
0dd8782fb02d Final mods for list comprehension
nipkow
parents: 23734
diff changeset
   366
      val ft = DatatypeCase.case_tr true DatatypePackage.datatype_of_constr
0dd8782fb02d Final mods for list comprehension
nipkow
parents: 23734
diff changeset
   367
                 ctxt [x, cs]
23529
958f9d9cfb63 Fixed problem with patterns in lambdas
nipkow
parents: 23526
diff changeset
   368
    in lambda x ft end
958f9d9cfb63 Fixed problem with patterns in lambdas
nipkow
parents: 23526
diff changeset
   369
in [("_lam_pats_syntax", fun_tr)] end
23526
936dc616a7fb Added pattern maatching for lambda abstraction
nipkow
parents: 23389
diff changeset
   370
*}
936dc616a7fb Added pattern maatching for lambda abstraction
nipkow
parents: 23389
diff changeset
   371
936dc616a7fb Added pattern maatching for lambda abstraction
nipkow
parents: 23389
diff changeset
   372
end