src/HOL/Library/Coinductive_List.thy
author wenzelm
Mon, 26 Oct 2009 20:02:37 +0100
changeset 33209 d36ca3960e33
parent 33197 de6285ebcc05
child 34941 156925dd67af
permissions -rw-r--r--
tuned white space;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
     1
(*  Title:      HOL/Library/Coinductive_Lists.thy
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
     2
    Author:     Lawrence C Paulson and Makarius
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
     3
*)
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
     4
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
     5
header {* Potentially infinite lists as greatest fixed-point *}
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
     6
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
     7
theory Coinductive_List
30663
0b6aff7451b2 Main is (Complex_Main) base entry point in library theories
haftmann
parents: 30198
diff changeset
     8
imports List Main
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
     9
begin
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
    10
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
    11
subsection {* List constructors over the datatype universe *}
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
    12
24860
ceb634874e0c coinduct: instantiation refers to suffix of main prop (major premise or conclusion);
wenzelm
parents: 23755
diff changeset
    13
definition "NIL = Datatype.In0 (Datatype.Numb 0)"
ceb634874e0c coinduct: instantiation refers to suffix of main prop (major premise or conclusion);
wenzelm
parents: 23755
diff changeset
    14
definition "CONS M N = Datatype.In1 (Datatype.Scons M N)"
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
    15
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
    16
lemma CONS_not_NIL [iff]: "CONS M N \<noteq> NIL"
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
    17
  and NIL_not_CONS [iff]: "NIL \<noteq> CONS M N"
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
    18
  and CONS_inject [iff]: "(CONS K M) = (CONS L N) = (K = L \<and> M = N)"
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
    19
  by (simp_all add: NIL_def CONS_def)
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
    20
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
    21
lemma CONS_mono: "M \<subseteq> M' \<Longrightarrow> N \<subseteq> N' \<Longrightarrow> CONS M N \<subseteq> CONS M' N'"
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
    22
  by (simp add: CONS_def In1_mono Scons_mono)
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
    23
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
    24
lemma CONS_UN1: "CONS M (\<Union>x. f x) = (\<Union>x. CONS M (f x))"
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
    25
    -- {* A continuity result? *}
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
    26
  by (simp add: CONS_def In1_UN1 Scons_UN1_y)
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
    27
24860
ceb634874e0c coinduct: instantiation refers to suffix of main prop (major premise or conclusion);
wenzelm
parents: 23755
diff changeset
    28
definition "List_case c h = Datatype.Case (\<lambda>_. c) (Datatype.Split h)"
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
    29
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
    30
lemma List_case_NIL [simp]: "List_case c h NIL = c"
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
    31
  and List_case_CONS [simp]: "List_case c h (CONS M N) = h M N"
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
    32
  by (simp_all add: List_case_def NIL_def CONS_def)
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
    33
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
    34
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
    35
subsection {* Corecursive lists *}
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
    36
24860
ceb634874e0c coinduct: instantiation refers to suffix of main prop (major premise or conclusion);
wenzelm
parents: 23755
diff changeset
    37
coinductive_set LList for A
ceb634874e0c coinduct: instantiation refers to suffix of main prop (major premise or conclusion);
wenzelm
parents: 23755
diff changeset
    38
where NIL [intro]:  "NIL \<in> LList A"
23755
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 22780
diff changeset
    39
  | CONS [intro]: "a \<in> A \<Longrightarrow> M \<in> LList A \<Longrightarrow> CONS a M \<in> LList A"
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
    40
23755
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 22780
diff changeset
    41
lemma LList_mono:
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 22780
diff changeset
    42
  assumes subset: "A \<subseteq> B"
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 22780
diff changeset
    43
  shows "LList A \<subseteq> LList B"
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
    44
    -- {* This justifies using @{text LList} in other recursive type definitions. *}
23755
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 22780
diff changeset
    45
proof
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 22780
diff changeset
    46
  fix x
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 22780
diff changeset
    47
  assume "x \<in> LList A"
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 22780
diff changeset
    48
  then show "x \<in> LList B"
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 22780
diff changeset
    49
  proof coinduct
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 22780
diff changeset
    50
    case LList
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 22780
diff changeset
    51
    then show ?case using subset
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 22780
diff changeset
    52
      by cases blast+
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 22780
diff changeset
    53
  qed
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 22780
diff changeset
    54
qed
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
    55
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
    56
consts
20820
58693343905f removed obsolete Datatype_Universe.thy (cf. Datatype.thy);
wenzelm
parents: 20802
diff changeset
    57
  LList_corec_aux :: "nat \<Rightarrow> ('a \<Rightarrow> ('b Datatype.item \<times> 'a) option) \<Rightarrow>
58693343905f removed obsolete Datatype_Universe.thy (cf. Datatype.thy);
wenzelm
parents: 20802
diff changeset
    58
    'a \<Rightarrow> 'b Datatype.item"
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
    59
primrec
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
    60
  "LList_corec_aux 0 f x = {}"
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
    61
  "LList_corec_aux (Suc k) f x =
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
    62
    (case f x of
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
    63
      None \<Rightarrow> NIL
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
    64
    | Some (z, w) \<Rightarrow> CONS z (LList_corec_aux k f w))"
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
    65
24860
ceb634874e0c coinduct: instantiation refers to suffix of main prop (major premise or conclusion);
wenzelm
parents: 23755
diff changeset
    66
definition "LList_corec a f = (\<Union>k. LList_corec_aux k f a)"
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
    67
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
    68
text {*
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
    69
  Note: the subsequent recursion equation for @{text LList_corec} may
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
    70
  be used with the Simplifier, provided it operates in a non-strict
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
    71
  fashion for case expressions (i.e.\ the usual @{text case}
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
    72
  congruence rule needs to be present).
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
    73
*}
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
    74
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
    75
lemma LList_corec:
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
    76
  "LList_corec a f =
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
    77
    (case f a of None \<Rightarrow> NIL | Some (z, w) \<Rightarrow> CONS z (LList_corec w f))"
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
    78
  (is "?lhs = ?rhs")
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
    79
proof
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
    80
  show "?lhs \<subseteq> ?rhs"
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
    81
    apply (unfold LList_corec_def)
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
    82
    apply (rule UN_least)
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
    83
    apply (case_tac k)
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
    84
     apply (simp_all (no_asm_simp) split: option.splits)
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
    85
    apply (rule allI impI subset_refl [THEN CONS_mono] UNIV_I [THEN UN_upper])+
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
    86
    done
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
    87
  show "?rhs \<subseteq> ?lhs"
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
    88
    apply (simp add: LList_corec_def split: option.splits)
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
    89
    apply (simp add: CONS_UN1)
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
    90
    apply safe
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
    91
     apply (rule_tac a = "Suc ?k" in UN_I, simp, simp)+
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
    92
    done
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
    93
qed
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
    94
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
    95
lemma LList_corec_type: "LList_corec a f \<in> LList UNIV"
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
    96
proof -
23755
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 22780
diff changeset
    97
  have "\<exists>x. LList_corec a f = LList_corec x f" by blast
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
    98
  then show ?thesis
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
    99
  proof coinduct
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   100
    case (LList L)
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   101
    then obtain x where L: "L = LList_corec x f" by blast
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   102
    show ?case
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   103
    proof (cases "f x")
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   104
      case None
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   105
      then have "LList_corec x f = NIL"
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   106
        by (simp add: LList_corec)
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   107
      with L have ?NIL by simp
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   108
      then show ?thesis ..
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   109
    next
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   110
      case (Some p)
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   111
      then have "LList_corec x f = CONS (fst p) (LList_corec (snd p) f)"
22780
41162a270151 Adapted to new parse translation for case expressions.
berghofe
parents: 22367
diff changeset
   112
        by (simp add: LList_corec split: prod.split)
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   113
      with L have ?CONS by auto
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   114
      then show ?thesis ..
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   115
    qed
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   116
  qed
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   117
qed
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   118
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   119
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   120
subsection {* Abstract type definition *}
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   121
24860
ceb634874e0c coinduct: instantiation refers to suffix of main prop (major premise or conclusion);
wenzelm
parents: 23755
diff changeset
   122
typedef 'a llist = "LList (range Datatype.Leaf) :: 'a Datatype.item set"
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   123
proof
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   124
  show "NIL \<in> ?llist" ..
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   125
qed
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   126
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   127
lemma NIL_type: "NIL \<in> llist"
18730
843da46f89ac tuned proofs;
wenzelm
parents: 18400
diff changeset
   128
  unfolding llist_def by (rule LList.NIL)
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   129
20820
58693343905f removed obsolete Datatype_Universe.thy (cf. Datatype.thy);
wenzelm
parents: 20802
diff changeset
   130
lemma CONS_type: "a \<in> range Datatype.Leaf \<Longrightarrow>
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   131
    M \<in> llist \<Longrightarrow> CONS a M \<in> llist"
18730
843da46f89ac tuned proofs;
wenzelm
parents: 18400
diff changeset
   132
  unfolding llist_def by (rule LList.CONS)
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   133
20820
58693343905f removed obsolete Datatype_Universe.thy (cf. Datatype.thy);
wenzelm
parents: 20802
diff changeset
   134
lemma llistI: "x \<in> LList (range Datatype.Leaf) \<Longrightarrow> x \<in> llist"
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   135
  by (simp add: llist_def)
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   136
20820
58693343905f removed obsolete Datatype_Universe.thy (cf. Datatype.thy);
wenzelm
parents: 20802
diff changeset
   137
lemma llistD: "x \<in> llist \<Longrightarrow> x \<in> LList (range Datatype.Leaf)"
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   138
  by (simp add: llist_def)
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   139
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   140
lemma Rep_llist_UNIV: "Rep_llist x \<in> LList UNIV"
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   141
proof -
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   142
  have "Rep_llist x \<in> llist" by (rule Rep_llist)
20820
58693343905f removed obsolete Datatype_Universe.thy (cf. Datatype.thy);
wenzelm
parents: 20802
diff changeset
   143
  then have "Rep_llist x \<in> LList (range Datatype.Leaf)"
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   144
    by (simp add: llist_def)
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   145
  also have "\<dots> \<subseteq> LList UNIV" by (rule LList_mono) simp
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   146
  finally show ?thesis .
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   147
qed
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   148
24860
ceb634874e0c coinduct: instantiation refers to suffix of main prop (major premise or conclusion);
wenzelm
parents: 23755
diff changeset
   149
definition "LNil = Abs_llist NIL"
ceb634874e0c coinduct: instantiation refers to suffix of main prop (major premise or conclusion);
wenzelm
parents: 23755
diff changeset
   150
definition "LCons x xs = Abs_llist (CONS (Datatype.Leaf x) (Rep_llist xs))"
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   151
28693
a1294a197d12 added rudimentary code generation
haftmann
parents: 27487
diff changeset
   152
code_datatype LNil LCons
a1294a197d12 added rudimentary code generation
haftmann
parents: 27487
diff changeset
   153
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   154
lemma LCons_not_LNil [iff]: "LCons x xs \<noteq> LNil"
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   155
  apply (simp add: LNil_def LCons_def)
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   156
  apply (subst Abs_llist_inject)
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   157
    apply (auto intro: NIL_type CONS_type Rep_llist)
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   158
  done
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   159
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   160
lemma LNil_not_LCons [iff]: "LNil \<noteq> LCons x xs"
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   161
  by (rule LCons_not_LNil [symmetric])
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   162
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   163
lemma LCons_inject [iff]: "(LCons x xs = LCons y ys) = (x = y \<and> xs = ys)"
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   164
  apply (simp add: LCons_def)
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   165
  apply (subst Abs_llist_inject)
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   166
    apply (auto simp add: Rep_llist_inject intro: CONS_type Rep_llist)
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   167
  done
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   168
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   169
lemma Rep_llist_LNil: "Rep_llist LNil = NIL"
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   170
  by (simp add: LNil_def add: Abs_llist_inverse NIL_type)
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   171
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   172
lemma Rep_llist_LCons: "Rep_llist (LCons x l) =
20820
58693343905f removed obsolete Datatype_Universe.thy (cf. Datatype.thy);
wenzelm
parents: 20802
diff changeset
   173
    CONS (Datatype.Leaf x) (Rep_llist l)"
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   174
  by (simp add: LCons_def Abs_llist_inverse CONS_type Rep_llist)
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   175
20802
1b43d9184bf5 tuned proofs;
wenzelm
parents: 20770
diff changeset
   176
lemma llist_cases [cases type: llist]:
1b43d9184bf5 tuned proofs;
wenzelm
parents: 20770
diff changeset
   177
  obtains
1b43d9184bf5 tuned proofs;
wenzelm
parents: 20770
diff changeset
   178
    (LNil) "l = LNil"
1b43d9184bf5 tuned proofs;
wenzelm
parents: 20770
diff changeset
   179
  | (LCons) x l' where "l = LCons x l'"
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   180
proof (cases l)
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   181
  case (Abs_llist L)
20820
58693343905f removed obsolete Datatype_Universe.thy (cf. Datatype.thy);
wenzelm
parents: 20802
diff changeset
   182
  from `L \<in> llist` have "L \<in> LList (range Datatype.Leaf)" by (rule llistD)
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   183
  then show ?thesis
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   184
  proof cases
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   185
    case NIL
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   186
    with Abs_llist have "l = LNil" by (simp add: LNil_def)
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   187
    with LNil show ?thesis .
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   188
  next
23755
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 22780
diff changeset
   189
    case (CONS a K)
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   190
    then have "K \<in> llist" by (blast intro: llistI)
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   191
    then obtain l' where "K = Rep_llist l'" by cases
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   192
    with CONS and Abs_llist obtain x where "l = LCons x l'"
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   193
      by (auto simp add: LCons_def Abs_llist_inject)
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   194
    with LCons show ?thesis .
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   195
  qed
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   196
qed
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   197
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   198
19086
1b3780be6cc2 new-style definitions/abbreviations;
wenzelm
parents: 18730
diff changeset
   199
definition
28693
a1294a197d12 added rudimentary code generation
haftmann
parents: 27487
diff changeset
   200
  [code del]: "llist_case c d l =
20820
58693343905f removed obsolete Datatype_Universe.thy (cf. Datatype.thy);
wenzelm
parents: 20802
diff changeset
   201
    List_case c (\<lambda>x y. d (inv Datatype.Leaf x) (Abs_llist y)) (Rep_llist l)"
20770
2c583720436e fixed translations: CONST;
wenzelm
parents: 20503
diff changeset
   202
33197
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents: 33056
diff changeset
   203
20770
2c583720436e fixed translations: CONST;
wenzelm
parents: 20503
diff changeset
   204
syntax  (* FIXME? *)
2c583720436e fixed translations: CONST;
wenzelm
parents: 20503
diff changeset
   205
  LNil :: logic
2c583720436e fixed translations: CONST;
wenzelm
parents: 20503
diff changeset
   206
  LCons :: logic
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   207
translations
20770
2c583720436e fixed translations: CONST;
wenzelm
parents: 20503
diff changeset
   208
  "case p of LNil \<Rightarrow> a | LCons x l \<Rightarrow> b" \<rightleftharpoons> "CONST llist_case a (\<lambda>x l. b) p"
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   209
28693
a1294a197d12 added rudimentary code generation
haftmann
parents: 27487
diff changeset
   210
lemma llist_case_LNil [simp, code]: "llist_case c d LNil = c"
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   211
  by (simp add: llist_case_def LNil_def
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   212
    NIL_type Abs_llist_inverse)
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   213
28693
a1294a197d12 added rudimentary code generation
haftmann
parents: 27487
diff changeset
   214
lemma llist_case_LCons [simp, code]: "llist_case c d (LCons M N) = d M N"
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   215
  by (simp add: llist_case_def LCons_def
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   216
    CONS_type Abs_llist_inverse Rep_llist Rep_llist_inverse inj_Leaf)
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   217
28693
a1294a197d12 added rudimentary code generation
haftmann
parents: 27487
diff changeset
   218
lemma llist_case_cert:
a1294a197d12 added rudimentary code generation
haftmann
parents: 27487
diff changeset
   219
  assumes "CASE \<equiv> llist_case c d"
28856
5e009a80fe6d Pure syntax: more coherent treatment of aprop, permanent TERM and &&&;
wenzelm
parents: 28702
diff changeset
   220
  shows "(CASE LNil \<equiv> c) &&& (CASE (LCons M N) \<equiv> d M N)"
28693
a1294a197d12 added rudimentary code generation
haftmann
parents: 27487
diff changeset
   221
  using assms by simp_all
a1294a197d12 added rudimentary code generation
haftmann
parents: 27487
diff changeset
   222
a1294a197d12 added rudimentary code generation
haftmann
parents: 27487
diff changeset
   223
setup {*
a1294a197d12 added rudimentary code generation
haftmann
parents: 27487
diff changeset
   224
  Code.add_case @{thm llist_case_cert}
a1294a197d12 added rudimentary code generation
haftmann
parents: 27487
diff changeset
   225
*}
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   226
19086
1b3780be6cc2 new-style definitions/abbreviations;
wenzelm
parents: 18730
diff changeset
   227
definition
28693
a1294a197d12 added rudimentary code generation
haftmann
parents: 27487
diff changeset
   228
  [code del]: "llist_corec a f =
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   229
    Abs_llist (LList_corec a
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   230
      (\<lambda>z.
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   231
        case f z of None \<Rightarrow> None
20820
58693343905f removed obsolete Datatype_Universe.thy (cf. Datatype.thy);
wenzelm
parents: 20802
diff changeset
   232
        | Some (v, w) \<Rightarrow> Some (Datatype.Leaf v, w)))"
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   233
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   234
lemma LList_corec_type2:
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   235
  "LList_corec a
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   236
    (\<lambda>z. case f z of None \<Rightarrow> None
20820
58693343905f removed obsolete Datatype_Universe.thy (cf. Datatype.thy);
wenzelm
parents: 20802
diff changeset
   237
      | Some (v, w) \<Rightarrow> Some (Datatype.Leaf v, w)) \<in> llist"
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   238
  (is "?corec a \<in> _")
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   239
proof (unfold llist_def)
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   240
  let "LList_corec a ?g" = "?corec a"
23755
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 22780
diff changeset
   241
  have "\<exists>x. ?corec a = ?corec x" by blast
20820
58693343905f removed obsolete Datatype_Universe.thy (cf. Datatype.thy);
wenzelm
parents: 20802
diff changeset
   242
  then show "?corec a \<in> LList (range Datatype.Leaf)"
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   243
  proof coinduct
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   244
    case (LList L)
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   245
    then obtain x where L: "L = ?corec x" by blast
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   246
    show ?case
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   247
    proof (cases "f x")
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   248
      case None
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   249
      then have "?corec x = NIL"
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   250
        by (simp add: LList_corec)
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   251
      with L have ?NIL by simp
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   252
      then show ?thesis ..
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   253
    next
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   254
      case (Some p)
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   255
      then have "?corec x =
20820
58693343905f removed obsolete Datatype_Universe.thy (cf. Datatype.thy);
wenzelm
parents: 20802
diff changeset
   256
          CONS (Datatype.Leaf (fst p)) (?corec (snd p))"
22780
41162a270151 Adapted to new parse translation for case expressions.
berghofe
parents: 22367
diff changeset
   257
        by (simp add: LList_corec split: prod.split)
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   258
      with L have ?CONS by auto
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   259
      then show ?thesis ..
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   260
    qed
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   261
  qed
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   262
qed
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   263
33056
791a4655cae3 renamed "nitpick_const_xxx" attributes to "nitpick_xxx" and "nitpick_ind_intros" to "nitpick_intros"
blanchet
parents: 32960
diff changeset
   264
lemma llist_corec [code, nitpick_simp]:
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   265
  "llist_corec a f =
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   266
    (case f a of None \<Rightarrow> LNil | Some (z, w) \<Rightarrow> LCons z (llist_corec w f))"
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   267
proof (cases "f a")
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   268
  case None
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   269
  then show ?thesis
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   270
    by (simp add: llist_corec_def LList_corec LNil_def)
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   271
next
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   272
  case (Some p)
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   273
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   274
  let "?corec a" = "llist_corec a f"
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   275
  let "?rep_corec a" =
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   276
    "LList_corec a
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   277
      (\<lambda>z. case f z of None \<Rightarrow> None
20820
58693343905f removed obsolete Datatype_Universe.thy (cf. Datatype.thy);
wenzelm
parents: 20802
diff changeset
   278
        | Some (v, w) \<Rightarrow> Some (Datatype.Leaf v, w))"
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   279
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   280
  have "?corec a = Abs_llist (?rep_corec a)"
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   281
    by (simp only: llist_corec_def)
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   282
  also from Some have "?rep_corec a =
20820
58693343905f removed obsolete Datatype_Universe.thy (cf. Datatype.thy);
wenzelm
parents: 20802
diff changeset
   283
      CONS (Datatype.Leaf (fst p)) (?rep_corec (snd p))"
22780
41162a270151 Adapted to new parse translation for case expressions.
berghofe
parents: 22367
diff changeset
   284
    by (simp add: LList_corec split: prod.split)
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   285
  also have "?rep_corec (snd p) = Rep_llist (?corec (snd p))"
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   286
    by (simp only: llist_corec_def Abs_llist_inverse LList_corec_type2)
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   287
  finally have "?corec a = LCons (fst p) (?corec (snd p))"
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   288
    by (simp only: LCons_def)
22780
41162a270151 Adapted to new parse translation for case expressions.
berghofe
parents: 22367
diff changeset
   289
  with Some show ?thesis by (simp split: prod.split)
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   290
qed
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   291
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   292
22367
6860f09242bf tuned document;
wenzelm
parents: 21404
diff changeset
   293
subsection {* Equality as greatest fixed-point -- the bisimulation principle *}
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   294
24860
ceb634874e0c coinduct: instantiation refers to suffix of main prop (major premise or conclusion);
wenzelm
parents: 23755
diff changeset
   295
coinductive_set EqLList for r
ceb634874e0c coinduct: instantiation refers to suffix of main prop (major premise or conclusion);
wenzelm
parents: 23755
diff changeset
   296
where EqNIL: "(NIL, NIL) \<in> EqLList r"
23755
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 22780
diff changeset
   297
  | EqCONS: "(a, b) \<in> r \<Longrightarrow> (M, N) \<in> EqLList r \<Longrightarrow>
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   298
      (CONS a M, CONS b N) \<in> EqLList r"
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   299
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   300
lemma EqLList_unfold:
30198
922f944f03b2 name changes
nipkow
parents: 28856
diff changeset
   301
    "EqLList r = dsum (Id_on {Datatype.Numb 0}) (dprod r (EqLList r))"
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   302
  by (fast intro!: EqLList.intros [unfolded NIL_def CONS_def]
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   303
           elim: EqLList.cases [unfolded NIL_def CONS_def])
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   304
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   305
lemma EqLList_implies_ntrunc_equality:
30198
922f944f03b2 name changes
nipkow
parents: 28856
diff changeset
   306
    "(M, N) \<in> EqLList (Id_on A) \<Longrightarrow> ntrunc k M = ntrunc k N"
20503
503ac4c5ef91 induct method: renamed 'fixing' to 'arbitrary';
wenzelm
parents: 19086
diff changeset
   307
  apply (induct k arbitrary: M N rule: nat_less_induct)
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   308
  apply (erule EqLList.cases)
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   309
   apply (safe del: equalityI)
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   310
  apply (case_tac n)
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   311
   apply simp
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   312
  apply (rename_tac n')
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   313
  apply (case_tac n')
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   314
   apply (simp_all add: CONS_def less_Suc_eq)
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   315
  done
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   316
30198
922f944f03b2 name changes
nipkow
parents: 28856
diff changeset
   317
lemma Domain_EqLList: "Domain (EqLList (Id_on A)) \<subseteq> LList A"
23755
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 22780
diff changeset
   318
  apply (rule subsetI)
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 22780
diff changeset
   319
  apply (erule LList.coinduct)
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 22780
diff changeset
   320
  apply (subst (asm) EqLList_unfold)
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 22780
diff changeset
   321
  apply (auto simp add: NIL_def CONS_def)
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   322
  done
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   323
30198
922f944f03b2 name changes
nipkow
parents: 28856
diff changeset
   324
lemma EqLList_Id_on: "EqLList (Id_on A) = Id_on (LList A)"
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   325
  (is "?lhs = ?rhs")
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   326
proof
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   327
  show "?lhs \<subseteq> ?rhs"
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   328
    apply (rule subsetI)
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   329
    apply (rule_tac p = x in PairE)
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   330
    apply clarify
30198
922f944f03b2 name changes
nipkow
parents: 28856
diff changeset
   331
    apply (rule Id_on_eqI)
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   332
     apply (rule EqLList_implies_ntrunc_equality [THEN ntrunc_equality],
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   333
       assumption)
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   334
    apply (erule DomainI [THEN Domain_EqLList [THEN subsetD]])
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   335
    done
23755
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 22780
diff changeset
   336
  {
30198
922f944f03b2 name changes
nipkow
parents: 28856
diff changeset
   337
    fix M N assume "(M, N) \<in> Id_on (LList A)"
922f944f03b2 name changes
nipkow
parents: 28856
diff changeset
   338
    then have "(M, N) \<in> EqLList (Id_on A)"
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   339
    proof coinduct
23755
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 22780
diff changeset
   340
      case (EqLList M N)
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 22780
diff changeset
   341
      then obtain L where L: "L \<in> LList A" and MN: "M = L" "N = L" by blast
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   342
      from L show ?case
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   343
      proof cases
23755
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 22780
diff changeset
   344
        case NIL with MN have ?EqNIL by simp
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   345
        then show ?thesis ..
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   346
      next
30198
922f944f03b2 name changes
nipkow
parents: 28856
diff changeset
   347
        case CONS with MN have ?EqCONS by (simp add: Id_onI)
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   348
        then show ?thesis ..
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   349
      qed
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   350
    qed
23755
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 22780
diff changeset
   351
  }
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 22780
diff changeset
   352
  then show "?rhs \<subseteq> ?lhs" by auto
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   353
qed
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   354
30198
922f944f03b2 name changes
nipkow
parents: 28856
diff changeset
   355
lemma EqLList_Id_on_iff [iff]: "(p \<in> EqLList (Id_on A)) = (p \<in> Id_on (LList A))"
922f944f03b2 name changes
nipkow
parents: 28856
diff changeset
   356
  by (simp only: EqLList_Id_on)
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   357
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   358
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   359
text {*
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   360
  To show two LLists are equal, exhibit a bisimulation!  (Also admits
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   361
  true equality.)
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   362
*}
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   363
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   364
lemma LList_equalityI
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   365
  [consumes 1, case_names EqLList, case_conclusion EqLList EqNIL EqCONS]:
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   366
  assumes r: "(M, N) \<in> r"
23755
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 22780
diff changeset
   367
    and step: "\<And>M N. (M, N) \<in> r \<Longrightarrow>
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 22780
diff changeset
   368
      M = NIL \<and> N = NIL \<or>
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 22780
diff changeset
   369
        (\<exists>a b M' N'.
30198
922f944f03b2 name changes
nipkow
parents: 28856
diff changeset
   370
          M = CONS a M' \<and> N = CONS b N' \<and> (a, b) \<in> Id_on A \<and>
922f944f03b2 name changes
nipkow
parents: 28856
diff changeset
   371
            ((M', N') \<in> r \<or> (M', N') \<in> EqLList (Id_on A)))"
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   372
  shows "M = N"
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   373
proof -
30198
922f944f03b2 name changes
nipkow
parents: 28856
diff changeset
   374
  from r have "(M, N) \<in> EqLList (Id_on A)"
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   375
  proof coinduct
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   376
    case EqLList
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   377
    then show ?case by (rule step)
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   378
  qed
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   379
  then show ?thesis by auto
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   380
qed
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   381
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   382
lemma LList_fun_equalityI
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   383
  [consumes 1, case_names NIL_type NIL CONS, case_conclusion CONS EqNIL EqCONS]:
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   384
  assumes M: "M \<in> LList A"
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   385
    and fun_NIL: "g NIL \<in> LList A"  "f NIL = g NIL"
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   386
    and fun_CONS: "\<And>x l. x \<in> A \<Longrightarrow> l \<in> LList A \<Longrightarrow>
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   387
            (f (CONS x l), g (CONS x l)) = (NIL, NIL) \<or>
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   388
            (\<exists>M N a b.
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   389
              (f (CONS x l), g (CONS x l)) = (CONS a M, CONS b N) \<and>
30198
922f944f03b2 name changes
nipkow
parents: 28856
diff changeset
   390
                (a, b) \<in> Id_on A \<and>
922f944f03b2 name changes
nipkow
parents: 28856
diff changeset
   391
                (M, N) \<in> {(f u, g u) | u. u \<in> LList A} \<union> Id_on (LList A))"
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   392
      (is "\<And>x l. _ \<Longrightarrow> _ \<Longrightarrow> ?fun_CONS x l")
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   393
  shows "f M = g M"
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   394
proof -
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   395
  let ?bisim = "{(f L, g L) | L. L \<in> LList A}"
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   396
  have "(f M, g M) \<in> ?bisim" using M by blast
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   397
  then show ?thesis
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   398
  proof (coinduct taking: A rule: LList_equalityI)
23755
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 22780
diff changeset
   399
    case (EqLList M N)
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 22780
diff changeset
   400
    then obtain L where MN: "M = f L" "N = g L" and L: "L \<in> LList A" by blast
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   401
    from L show ?case
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   402
    proof (cases L)
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   403
      case NIL
30198
922f944f03b2 name changes
nipkow
parents: 28856
diff changeset
   404
      with fun_NIL and MN have "(M, N) \<in> Id_on (LList A)" by auto
922f944f03b2 name changes
nipkow
parents: 28856
diff changeset
   405
      then have "(M, N) \<in> EqLList (Id_on A)" ..
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   406
      then show ?thesis by cases simp_all
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   407
    next
23755
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 22780
diff changeset
   408
      case (CONS a K)
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   409
      from fun_CONS and `a \<in> A` `K \<in> LList A`
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   410
      have "?fun_CONS a K" (is "?NIL \<or> ?CONS") .
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   411
      then show ?thesis
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   412
      proof
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   413
        assume ?NIL
30198
922f944f03b2 name changes
nipkow
parents: 28856
diff changeset
   414
        with MN CONS have "(M, N) \<in> Id_on (LList A)" by auto
922f944f03b2 name changes
nipkow
parents: 28856
diff changeset
   415
        then have "(M, N) \<in> EqLList (Id_on A)" ..
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   416
        then show ?thesis by cases simp_all
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   417
      next
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   418
        assume ?CONS
23755
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 22780
diff changeset
   419
        with CONS obtain a b M' N' where
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 22780
diff changeset
   420
            fg: "(f L, g L) = (CONS a M', CONS b N')"
30198
922f944f03b2 name changes
nipkow
parents: 28856
diff changeset
   421
          and ab: "(a, b) \<in> Id_on A"
922f944f03b2 name changes
nipkow
parents: 28856
diff changeset
   422
          and M'N': "(M', N') \<in> ?bisim \<union> Id_on (LList A)"
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   423
          by blast
23755
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 22780
diff changeset
   424
        from M'N' show ?thesis
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   425
        proof
23755
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 22780
diff changeset
   426
          assume "(M', N') \<in> ?bisim"
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 22780
diff changeset
   427
          with MN fg ab show ?thesis by simp
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   428
        next
30198
922f944f03b2 name changes
nipkow
parents: 28856
diff changeset
   429
          assume "(M', N') \<in> Id_on (LList A)"
922f944f03b2 name changes
nipkow
parents: 28856
diff changeset
   430
          then have "(M', N') \<in> EqLList (Id_on A)" ..
23755
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 22780
diff changeset
   431
          with MN fg ab show ?thesis by simp
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   432
        qed
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   433
      qed
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   434
    qed
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   435
  qed
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   436
qed
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   437
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   438
text {*
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   439
  Finality of @{text "llist A"}: Uniqueness of functions defined by corecursion.
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   440
*}
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   441
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   442
lemma equals_LList_corec:
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   443
  assumes h: "\<And>x. h x =
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   444
    (case f x of None \<Rightarrow> NIL | Some (z, w) \<Rightarrow> CONS z (h w))"
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   445
  shows "h x = (\<lambda>x. LList_corec x f) x"
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   446
proof -
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   447
  def h' \<equiv> "\<lambda>x. LList_corec x f"
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   448
  then have h': "\<And>x. h' x =
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   449
      (case f x of None \<Rightarrow> NIL | Some (z, w) \<Rightarrow> CONS z (h' w))"
18730
843da46f89ac tuned proofs;
wenzelm
parents: 18400
diff changeset
   450
    unfolding h'_def by (simp add: LList_corec)
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   451
  have "(h x, h' x) \<in> {(h u, h' u) | u. True}" by blast
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   452
  then show "h x = h' x"
24863
307b979b1f54 tuned proofs (via polymorphic taking'');
wenzelm
parents: 24860
diff changeset
   453
  proof (coinduct taking: UNIV rule: LList_equalityI)
23755
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 22780
diff changeset
   454
    case (EqLList M N)
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 22780
diff changeset
   455
    then obtain x where MN: "M = h x" "N = h' x" by blast
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   456
    show ?case
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   457
    proof (cases "f x")
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   458
      case None
23755
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 22780
diff changeset
   459
      with h h' MN have ?EqNIL by simp
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   460
      then show ?thesis ..
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   461
    next
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   462
      case (Some p)
23755
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 22780
diff changeset
   463
      with h h' MN have "M = CONS (fst p) (h (snd p))"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32655
diff changeset
   464
        and "N = CONS (fst p) (h' (snd p))"
23755
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 22780
diff changeset
   465
        by (simp_all split: prod.split)
30198
922f944f03b2 name changes
nipkow
parents: 28856
diff changeset
   466
      then have ?EqCONS by (auto iff: Id_on_iff)
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   467
      then show ?thesis ..
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   468
    qed
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   469
  qed
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   470
qed
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   471
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   472
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   473
lemma llist_equalityI
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   474
  [consumes 1, case_names Eqllist, case_conclusion Eqllist EqLNil EqLCons]:
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   475
  assumes r: "(l1, l2) \<in> r"
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   476
    and step: "\<And>q. q \<in> r \<Longrightarrow>
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   477
      q = (LNil, LNil) \<or>
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   478
        (\<exists>l1 l2 a b.
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   479
          q = (LCons a l1, LCons b l2) \<and> a = b \<and>
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   480
            ((l1, l2) \<in> r \<or> l1 = l2))"
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   481
      (is "\<And>q. _ \<Longrightarrow> ?EqLNil q \<or> ?EqLCons q")
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   482
  shows "l1 = l2"
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   483
proof -
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   484
  def M \<equiv> "Rep_llist l1" and N \<equiv> "Rep_llist l2"
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   485
  with r have "(M, N) \<in> {(Rep_llist l1, Rep_llist l2) | l1 l2. (l1, l2) \<in> r}"
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   486
    by blast
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   487
  then have "M = N"
24863
307b979b1f54 tuned proofs (via polymorphic taking'');
wenzelm
parents: 24860
diff changeset
   488
  proof (coinduct taking: UNIV rule: LList_equalityI)
23755
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 22780
diff changeset
   489
    case (EqLList M N)
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   490
    then obtain l1 l2 where
23755
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 22780
diff changeset
   491
        MN: "M = Rep_llist l1" "N = Rep_llist l2" and r: "(l1, l2) \<in> r"
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   492
      by auto
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   493
    from step [OF r] show ?case
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   494
    proof
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   495
      assume "?EqLNil (l1, l2)"
23755
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 22780
diff changeset
   496
      with MN have ?EqNIL by (simp add: Rep_llist_LNil)
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   497
      then show ?thesis ..
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   498
    next
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   499
      assume "?EqLCons (l1, l2)"
23755
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 22780
diff changeset
   500
      with MN have ?EqCONS
30198
922f944f03b2 name changes
nipkow
parents: 28856
diff changeset
   501
        by (force simp add: Rep_llist_LCons EqLList_Id_on intro: Rep_llist_UNIV)
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   502
      then show ?thesis ..
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   503
    qed
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   504
  qed
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   505
  then show ?thesis by (simp add: M_def N_def Rep_llist_inject)
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   506
qed
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   507
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   508
lemma llist_fun_equalityI
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   509
  [case_names LNil LCons, case_conclusion LCons EqLNil EqLCons]:
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   510
  assumes fun_LNil: "f LNil = g LNil"
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   511
    and fun_LCons: "\<And>x l.
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   512
      (f (LCons x l), g (LCons x l)) = (LNil, LNil) \<or>
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   513
        (\<exists>l1 l2 a b.
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   514
          (f (LCons x l), g (LCons x l)) = (LCons a l1, LCons b l2) \<and>
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   515
            a = b \<and> ((l1, l2) \<in> {(f u, g u) | u. True} \<or> l1 = l2))"
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   516
      (is "\<And>x l. ?fun_LCons x l")
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   517
  shows "f l = g l"
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   518
proof -
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   519
  have "(f l, g l) \<in> {(f l, g l) | l. True}" by blast
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   520
  then show ?thesis
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   521
  proof (coinduct rule: llist_equalityI)
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   522
    case (Eqllist q)
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   523
    then obtain l where q: "q = (f l, g l)" by blast
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   524
    show ?case
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   525
    proof (cases l)
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   526
      case LNil
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   527
      with fun_LNil and q have "q = (g LNil, g LNil)" by simp
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   528
      then show ?thesis by (cases "g LNil") simp_all
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   529
    next
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   530
      case (LCons x l')
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   531
      with `?fun_LCons x l'` q LCons show ?thesis by blast
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   532
    qed
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   533
  qed
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   534
qed
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   535
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   536
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   537
subsection {* Derived operations -- both on the set and abstract type *}
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   538
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   539
subsubsection {* @{text Lconst} *}
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   540
24860
ceb634874e0c coinduct: instantiation refers to suffix of main prop (major premise or conclusion);
wenzelm
parents: 23755
diff changeset
   541
definition "Lconst M \<equiv> lfp (\<lambda>N. CONS M N)"
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   542
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   543
lemma Lconst_fun_mono: "mono (CONS M)"
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   544
  by (simp add: monoI CONS_mono)
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   545
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   546
lemma Lconst: "Lconst M = CONS M (Lconst M)"
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   547
  by (rule Lconst_def [THEN def_lfp_unfold]) (rule Lconst_fun_mono)
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   548
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   549
lemma Lconst_type:
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   550
  assumes "M \<in> A"
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   551
  shows "Lconst M \<in> LList A"
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   552
proof -
23755
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 22780
diff changeset
   553
  have "Lconst M \<in> {Lconst (id M)}" by simp
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   554
  then show ?thesis
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   555
  proof coinduct
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   556
    case (LList N)
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   557
    then have "N = Lconst M" by simp
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   558
    also have "\<dots> = CONS M (Lconst M)" by (rule Lconst)
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   559
    finally have ?CONS using `M \<in> A` by simp
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   560
    then show ?case ..
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   561
  qed
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   562
qed
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   563
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   564
lemma Lconst_eq_LList_corec: "Lconst M = LList_corec M (\<lambda>x. Some (x, x))"
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   565
  apply (rule equals_LList_corec)
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   566
  apply simp
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   567
  apply (rule Lconst)
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   568
  done
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   569
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   570
lemma gfp_Lconst_eq_LList_corec:
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   571
    "gfp (\<lambda>N. CONS M N) = LList_corec M (\<lambda>x. Some(x, x))"
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   572
  apply (rule equals_LList_corec)
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   573
  apply simp
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   574
  apply (rule Lconst_fun_mono [THEN gfp_unfold])
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   575
  done
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   576
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   577
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   578
subsubsection {* @{text Lmap} and @{text lmap} *}
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   579
19086
1b3780be6cc2 new-style definitions/abbreviations;
wenzelm
parents: 18730
diff changeset
   580
definition
1b3780be6cc2 new-style definitions/abbreviations;
wenzelm
parents: 18730
diff changeset
   581
  "Lmap f M = LList_corec M (List_case None (\<lambda>x M'. Some (f x, M')))"
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 20820
diff changeset
   582
definition
19086
1b3780be6cc2 new-style definitions/abbreviations;
wenzelm
parents: 18730
diff changeset
   583
  "lmap f l = llist_corec l
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   584
    (\<lambda>z.
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   585
      case z of LNil \<Rightarrow> None
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   586
      | LCons y z \<Rightarrow> Some (f y, z))"
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   587
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   588
lemma Lmap_NIL [simp]: "Lmap f NIL = NIL"
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   589
  and Lmap_CONS [simp]: "Lmap f (CONS M N) = CONS (f M) (Lmap f N)"
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   590
  by (simp_all add: Lmap_def LList_corec)
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   591
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   592
lemma Lmap_type:
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   593
  assumes M: "M \<in> LList A"
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   594
    and f: "\<And>x. x \<in> A \<Longrightarrow> f x \<in> B"
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   595
  shows "Lmap f M \<in> LList B"
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   596
proof -
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   597
  from M have "Lmap f M \<in> {Lmap f N | N. N \<in> LList A}" by blast
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   598
  then show ?thesis
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   599
  proof coinduct
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   600
    case (LList L)
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   601
    then obtain N where L: "L = Lmap f N" and N: "N \<in> LList A" by blast
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   602
    from N show ?case
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   603
    proof cases
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   604
      case NIL
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   605
      with L have ?NIL by simp
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   606
      then show ?thesis ..
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   607
    next
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   608
      case (CONS K a)
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   609
      with f L have ?CONS by auto
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   610
      then show ?thesis ..
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   611
    qed
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   612
  qed
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   613
qed
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   614
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   615
lemma Lmap_compose:
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   616
  assumes M: "M \<in> LList A"
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   617
  shows "Lmap (f o g) M = Lmap f (Lmap g M)"  (is "?lhs M = ?rhs M")
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   618
proof -
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   619
  have "(?lhs M, ?rhs M) \<in> {(?lhs N, ?rhs N) | N. N \<in> LList A}"
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   620
    using M by blast
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   621
  then show ?thesis
24863
307b979b1f54 tuned proofs (via polymorphic taking'');
wenzelm
parents: 24860
diff changeset
   622
  proof (coinduct taking: "range (\<lambda>N. N)" rule: LList_equalityI)
23755
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 22780
diff changeset
   623
    case (EqLList L M)
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 22780
diff changeset
   624
    then obtain N where LM: "L = ?lhs N" "M = ?rhs N" and N: "N \<in> LList A" by blast
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   625
    from N show ?case
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   626
    proof cases
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   627
      case NIL
23755
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 22780
diff changeset
   628
      with LM have ?EqNIL by simp
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   629
      then show ?thesis ..
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   630
    next
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   631
      case CONS
23755
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 22780
diff changeset
   632
      with LM have ?EqCONS by auto
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   633
      then show ?thesis ..
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   634
    qed
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   635
  qed
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   636
qed
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   637
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   638
lemma Lmap_ident:
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   639
  assumes M: "M \<in> LList A"
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   640
  shows "Lmap (\<lambda>x. x) M = M"  (is "?lmap M = _")
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   641
proof -
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   642
  have "(?lmap M, M) \<in> {(?lmap N, N) | N. N \<in> LList A}" using M by blast
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   643
  then show ?thesis
24863
307b979b1f54 tuned proofs (via polymorphic taking'');
wenzelm
parents: 24860
diff changeset
   644
  proof (coinduct taking: "range (\<lambda>N. N)" rule: LList_equalityI)
23755
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 22780
diff changeset
   645
    case (EqLList L M)
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 22780
diff changeset
   646
    then obtain N where LM: "L = ?lmap N" "M = N" and N: "N \<in> LList A" by blast
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   647
    from N show ?case
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   648
    proof cases
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   649
      case NIL
23755
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 22780
diff changeset
   650
      with LM have ?EqNIL by simp
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   651
      then show ?thesis ..
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   652
    next
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   653
      case CONS
23755
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 22780
diff changeset
   654
      with LM have ?EqCONS by auto
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   655
      then show ?thesis ..
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   656
    qed
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   657
  qed
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   658
qed
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   659
33056
791a4655cae3 renamed "nitpick_const_xxx" attributes to "nitpick_xxx" and "nitpick_ind_intros" to "nitpick_intros"
blanchet
parents: 32960
diff changeset
   660
lemma lmap_LNil [simp, nitpick_simp]: "lmap f LNil = LNil"
791a4655cae3 renamed "nitpick_const_xxx" attributes to "nitpick_xxx" and "nitpick_ind_intros" to "nitpick_intros"
blanchet
parents: 32960
diff changeset
   661
  and lmap_LCons [simp, nitpick_simp]:
32655
dd84779cd191 Added "nitpick_const_simp" tags to lazy list theories.
blanchet
parents: 30971
diff changeset
   662
  "lmap f (LCons M N) = LCons (f M) (lmap f N)"
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   663
  by (simp_all add: lmap_def llist_corec)
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   664
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   665
lemma lmap_compose [simp]: "lmap (f o g) l = lmap f (lmap g l)"
24860
ceb634874e0c coinduct: instantiation refers to suffix of main prop (major premise or conclusion);
wenzelm
parents: 23755
diff changeset
   666
  by (coinduct l rule: llist_fun_equalityI) auto
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   667
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   668
lemma lmap_ident [simp]: "lmap (\<lambda>x. x) l = l"
24860
ceb634874e0c coinduct: instantiation refers to suffix of main prop (major premise or conclusion);
wenzelm
parents: 23755
diff changeset
   669
  by (coinduct l rule: llist_fun_equalityI) auto
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   670
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   671
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   672
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   673
subsubsection {* @{text Lappend} *}
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   674
19086
1b3780be6cc2 new-style definitions/abbreviations;
wenzelm
parents: 18730
diff changeset
   675
definition
1b3780be6cc2 new-style definitions/abbreviations;
wenzelm
parents: 18730
diff changeset
   676
  "Lappend M N = LList_corec (M, N)
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   677
    (split (List_case
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   678
        (List_case None (\<lambda>N1 N2. Some (N1, (NIL, N2))))
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   679
        (\<lambda>M1 M2 N. Some (M1, (M2, N)))))"
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 20820
diff changeset
   680
definition
19086
1b3780be6cc2 new-style definitions/abbreviations;
wenzelm
parents: 18730
diff changeset
   681
  "lappend l n = llist_corec (l, n)
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   682
    (split (llist_case
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   683
        (llist_case None (\<lambda>n1 n2. Some (n1, (LNil, n2))))
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   684
        (\<lambda>l1 l2 n. Some (l1, (l2, n)))))"
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   685
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   686
lemma Lappend_NIL_NIL [simp]:
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   687
    "Lappend NIL NIL = NIL"
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   688
  and Lappend_NIL_CONS [simp]:
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   689
    "Lappend NIL (CONS N N') = CONS N (Lappend NIL N')"
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   690
  and Lappend_CONS [simp]:
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   691
    "Lappend (CONS M M') N = CONS M (Lappend M' N)"
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   692
  by (simp_all add: Lappend_def LList_corec)
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   693
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   694
lemma Lappend_NIL [simp]: "M \<in> LList A \<Longrightarrow> Lappend NIL M = M"
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   695
  by (erule LList_fun_equalityI) auto
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   696
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   697
lemma Lappend_NIL2: "M \<in> LList A \<Longrightarrow> Lappend M NIL = M"
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   698
  by (erule LList_fun_equalityI) auto
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   699
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   700
lemma Lappend_type:
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   701
  assumes M: "M \<in> LList A" and N: "N \<in> LList A"
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   702
  shows "Lappend M N \<in> LList A"
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   703
proof -
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   704
  have "Lappend M N \<in> {Lappend u v | u v. u \<in> LList A \<and> v \<in> LList A}"
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   705
    using M N by blast
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   706
  then show ?thesis
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   707
  proof coinduct
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   708
    case (LList L)
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   709
    then obtain M N where L: "L = Lappend M N"
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   710
        and M: "M \<in> LList A" and N: "N \<in> LList A"
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   711
      by blast
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   712
    from M show ?case
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   713
    proof cases
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   714
      case NIL
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   715
      from N show ?thesis
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   716
      proof cases
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   717
        case NIL
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   718
        with L and `M = NIL` have ?NIL by simp
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   719
        then show ?thesis ..
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   720
      next
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   721
        case CONS
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   722
        with L and `M = NIL` have ?CONS by simp
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   723
        then show ?thesis ..
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   724
      qed
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   725
    next
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   726
      case CONS
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   727
      with L N have ?CONS by auto
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   728
      then show ?thesis ..
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   729
    qed
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   730
  qed
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   731
qed
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   732
33056
791a4655cae3 renamed "nitpick_const_xxx" attributes to "nitpick_xxx" and "nitpick_ind_intros" to "nitpick_intros"
blanchet
parents: 32960
diff changeset
   733
lemma lappend_LNil_LNil [simp, nitpick_simp]: "lappend LNil LNil = LNil"
791a4655cae3 renamed "nitpick_const_xxx" attributes to "nitpick_xxx" and "nitpick_ind_intros" to "nitpick_intros"
blanchet
parents: 32960
diff changeset
   734
  and lappend_LNil_LCons [simp, nitpick_simp]: "lappend LNil (LCons l l') = LCons l (lappend LNil l')"
791a4655cae3 renamed "nitpick_const_xxx" attributes to "nitpick_xxx" and "nitpick_ind_intros" to "nitpick_intros"
blanchet
parents: 32960
diff changeset
   735
  and lappend_LCons [simp, nitpick_simp]: "lappend (LCons l l') m = LCons l (lappend l' m)"
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   736
  by (simp_all add: lappend_def llist_corec)
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   737
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   738
lemma lappend_LNil1 [simp]: "lappend LNil l = l"
24860
ceb634874e0c coinduct: instantiation refers to suffix of main prop (major premise or conclusion);
wenzelm
parents: 23755
diff changeset
   739
  by (coinduct l rule: llist_fun_equalityI) auto
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   740
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   741
lemma lappend_LNil2 [simp]: "lappend l LNil = l"
24860
ceb634874e0c coinduct: instantiation refers to suffix of main prop (major premise or conclusion);
wenzelm
parents: 23755
diff changeset
   742
  by (coinduct l rule: llist_fun_equalityI) auto
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   743
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   744
lemma lappend_assoc: "lappend (lappend l1 l2) l3 = lappend l1 (lappend l2 l3)"
24860
ceb634874e0c coinduct: instantiation refers to suffix of main prop (major premise or conclusion);
wenzelm
parents: 23755
diff changeset
   745
  by (coinduct l1 rule: llist_fun_equalityI) auto
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   746
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   747
lemma lmap_lappend_distrib: "lmap f (lappend l n) = lappend (lmap f l) (lmap f n)"
24860
ceb634874e0c coinduct: instantiation refers to suffix of main prop (major premise or conclusion);
wenzelm
parents: 23755
diff changeset
   748
  by (coinduct l rule: llist_fun_equalityI) auto
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   749
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   750
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   751
subsection{* iterates *}
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   752
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   753
text {* @{text llist_fun_equalityI} cannot be used here! *}
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   754
19086
1b3780be6cc2 new-style definitions/abbreviations;
wenzelm
parents: 18730
diff changeset
   755
definition
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 20820
diff changeset
   756
  iterates :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a llist" where
19086
1b3780be6cc2 new-style definitions/abbreviations;
wenzelm
parents: 18730
diff changeset
   757
  "iterates f a = llist_corec a (\<lambda>x. Some (x, f x))"
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   758
33056
791a4655cae3 renamed "nitpick_const_xxx" attributes to "nitpick_xxx" and "nitpick_ind_intros" to "nitpick_intros"
blanchet
parents: 32960
diff changeset
   759
lemma iterates [nitpick_simp]: "iterates f x = LCons x (iterates f (f x))"
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   760
  apply (unfold iterates_def)
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   761
  apply (subst llist_corec)
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   762
  apply simp
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   763
  done
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   764
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   765
lemma lmap_iterates: "lmap f (iterates f x) = iterates f (f x)"
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   766
proof -
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   767
  have "(lmap f (iterates f x), iterates f (f x)) \<in>
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   768
    {(lmap f (iterates f u), iterates f (f u)) | u. True}" by blast
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   769
  then show ?thesis
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   770
  proof (coinduct rule: llist_equalityI)
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   771
    case (Eqllist q)
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   772
    then obtain x where q: "q = (lmap f (iterates f x), iterates f (f x))"
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   773
      by blast
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   774
    also have "iterates f (f x) = LCons (f x) (iterates f (f (f x)))"
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   775
      by (subst iterates) rule
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   776
    also have "iterates f x = LCons x (iterates f (f x))"
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   777
      by (subst iterates) rule
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   778
    finally have ?EqLCons by auto
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   779
    then show ?case ..
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   780
  qed
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   781
qed
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   782
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   783
lemma iterates_lmap: "iterates f x = LCons x (lmap f (iterates f x))"
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   784
  by (subst lmap_iterates) (rule iterates)
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   785
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   786
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   787
subsection{* A rather complex proof about iterates -- cf.\ Andy Pitts *}
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   788
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   789
lemma funpow_lmap:
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   790
  fixes f :: "'a \<Rightarrow> 'a"
30971
7fbebf75b3ef funpow and relpow with shared "^^" syntax
haftmann
parents: 30952
diff changeset
   791
  shows "(lmap f ^^ n) (LCons b l) = LCons ((f ^^ n) b) ((lmap f ^^ n) l)"
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   792
  by (induct n) simp_all
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   793
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   794
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   795
lemma iterates_equality:
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   796
  assumes h: "\<And>x. h x = LCons x (lmap f (h x))"
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   797
  shows "h = iterates f"
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   798
proof
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   799
  fix x
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   800
  have "(h x, iterates f x) \<in>
30971
7fbebf75b3ef funpow and relpow with shared "^^" syntax
haftmann
parents: 30952
diff changeset
   801
      {((lmap f ^^ n) (h u), (lmap f ^^ n) (iterates f u)) | u n. True}"
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   802
  proof -
30971
7fbebf75b3ef funpow and relpow with shared "^^" syntax
haftmann
parents: 30952
diff changeset
   803
    have "(h x, iterates f x) = ((lmap f ^^ 0) (h x), (lmap f ^^ 0) (iterates f x))"
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   804
      by simp
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   805
    then show ?thesis by blast
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   806
  qed
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   807
  then show "h x = iterates f x"
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   808
  proof (coinduct rule: llist_equalityI)
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   809
    case (Eqllist q)
30971
7fbebf75b3ef funpow and relpow with shared "^^" syntax
haftmann
parents: 30952
diff changeset
   810
    then obtain u n where "q = ((lmap f ^^ n) (h u), (lmap f ^^ n) (iterates f u))"
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   811
        (is "_ = (?q1, ?q2)")
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   812
      by auto
30971
7fbebf75b3ef funpow and relpow with shared "^^" syntax
haftmann
parents: 30952
diff changeset
   813
    also have "?q1 = LCons ((f ^^ n) u) ((lmap f ^^ Suc n) (h u))"
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   814
    proof -
30971
7fbebf75b3ef funpow and relpow with shared "^^" syntax
haftmann
parents: 30952
diff changeset
   815
      have "?q1 = (lmap f ^^ n) (LCons u (lmap f (h u)))"
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   816
        by (subst h) rule
30971
7fbebf75b3ef funpow and relpow with shared "^^" syntax
haftmann
parents: 30952
diff changeset
   817
      also have "\<dots> = LCons ((f ^^ n) u) ((lmap f ^^ n) (lmap f (h u)))"
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   818
        by (rule funpow_lmap)
30971
7fbebf75b3ef funpow and relpow with shared "^^" syntax
haftmann
parents: 30952
diff changeset
   819
      also have "(lmap f ^^ n) (lmap f (h u)) = (lmap f ^^ Suc n) (h u)"
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   820
        by (simp add: funpow_swap1)
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   821
      finally show ?thesis .
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   822
    qed
30971
7fbebf75b3ef funpow and relpow with shared "^^" syntax
haftmann
parents: 30952
diff changeset
   823
    also have "?q2 = LCons ((f ^^ n) u) ((lmap f ^^ Suc n) (iterates f u))"
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   824
    proof -
30971
7fbebf75b3ef funpow and relpow with shared "^^" syntax
haftmann
parents: 30952
diff changeset
   825
      have "?q2 = (lmap f ^^ n) (LCons u (iterates f (f u)))"
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   826
        by (subst iterates) rule
30971
7fbebf75b3ef funpow and relpow with shared "^^" syntax
haftmann
parents: 30952
diff changeset
   827
      also have "\<dots> = LCons ((f ^^ n) u) ((lmap f ^^ n) (iterates f (f u)))"
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   828
        by (rule funpow_lmap)
30971
7fbebf75b3ef funpow and relpow with shared "^^" syntax
haftmann
parents: 30952
diff changeset
   829
      also have "(lmap f ^^ n) (iterates f (f u)) = (lmap f ^^ Suc n) (iterates f u)"
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   830
        by (simp add: lmap_iterates funpow_swap1)
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   831
      finally show ?thesis .
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   832
    qed
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   833
    finally have ?EqLCons by (auto simp del: funpow.simps)
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   834
    then show ?case ..
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   835
  qed
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   836
qed
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   837
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   838
lemma lappend_iterates: "lappend (iterates f x) l = iterates f x"
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   839
proof -
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   840
  have "(lappend (iterates f x) l, iterates f x) \<in>
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   841
    {(lappend (iterates f u) l, iterates f u) | u. True}" by blast
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   842
  then show ?thesis
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   843
  proof (coinduct rule: llist_equalityI)
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   844
    case (Eqllist q)
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   845
    then obtain x where "q = (lappend (iterates f x) l, iterates f x)" by blast
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   846
    also have "iterates f x = LCons x (iterates f (f x))" by (rule iterates)
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   847
    finally have ?EqLCons by auto
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   848
    then show ?case ..
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   849
  qed
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   850
qed
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   851
33197
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents: 33056
diff changeset
   852
setup {*
33209
d36ca3960e33 tuned white space;
wenzelm
parents: 33197
diff changeset
   853
  Nitpick.register_codatatype @{typ "'a llist"} @{const_name llist_case}
d36ca3960e33 tuned white space;
wenzelm
parents: 33197
diff changeset
   854
    (map dest_Const [@{term LNil}, @{term LCons}])
33197
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents: 33056
diff changeset
   855
*}
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents: 33056
diff changeset
   856
18400
6cc32c77d402 Potentially infinite lists as greatest fixed-point.
wenzelm
parents:
diff changeset
   857
end