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