src/CCL/ex/Stream.thy
changeset 20140 98acc6d0fab6
parent 17456 bcf7544875b2
child 23894 1a4167d761ac
--- a/src/CCL/ex/Stream.thy	Mon Jul 17 18:42:38 2006 +0200
+++ b/src/CCL/ex/Stream.thy	Tue Jul 18 02:22:38 2006 +0200
@@ -19,6 +19,127 @@
   iter1_def:   "iter1(f,a) == letrec iter x be x$iter(f(x)) in iter(a)"
   iter2_def:   "iter2(f,a) == letrec iter x be x$map(f,iter(x)) in iter(a)"
 
-ML {* use_legacy_bindings (the_context ()) *}
+
+(*
+Proving properties about infinite lists using coinduction:
+    Lists(A)  is the set of all finite and infinite lists of elements of A.
+    ILists(A) is the set of infinite lists of elements of A.
+*)
+
+
+subsection {* Map of composition is composition of maps *}
+
+lemma map_comp:
+  assumes 1: "l:Lists(A)"
+  shows "map(f o g,l) = map(f,map(g,l))"
+  apply (tactic {* eq_coinduct3_tac
+    "{p. EX x y. p=<x,y> & (EX l:Lists (A) .x=map (f o g,l) & y=map (f,map (g,l)))}" 1 *})
+   apply (blast intro: 1)
+  apply safe
+  apply (drule ListsXH [THEN iffD1])
+  apply (tactic "EQgen_tac (simpset ()) [] 1")
+   apply fastsimp
+  done
+
+(*** Mapping the identity function leaves a list unchanged ***)
+
+lemma map_id:
+  assumes 1: "l:Lists(A)"
+  shows "map(%x. x,l) = l"
+  apply (tactic {* eq_coinduct3_tac
+    "{p. EX x y. p=<x,y> & (EX l:Lists (A) .x=map (%x. x,l) & y=l) }" 1 *})
+  apply (blast intro: 1)
+  apply safe
+  apply (drule ListsXH [THEN iffD1])
+  apply (tactic "EQgen_tac (simpset ()) [] 1")
+  apply blast
+  done
+
+
+subsection {* Mapping distributes over append *}
+
+lemma map_append:
+  assumes "l:Lists(A)"
+    and "m:Lists(A)"
+  shows "map(f,l@m) = map(f,l) @ map(f,m)"
+  apply (tactic {* eq_coinduct3_tac
+    "{p. EX x y. p=<x,y> & (EX l:Lists (A). EX m:Lists (A). x=map (f,l@m) & y=map (f,l) @ map (f,m))}" 1 *})
+  apply (blast intro: prems)
+  apply safe
+  apply (drule ListsXH [THEN iffD1])
+  apply (tactic "EQgen_tac (simpset ()) [] 1")
+  apply (drule ListsXH [THEN iffD1])
+  apply (tactic "EQgen_tac (simpset ()) [] 1")
+  apply blast
+  done
+
+
+subsection {* Append is associative *}
+
+lemma append_assoc:
+  assumes "k:Lists(A)"
+    and "l:Lists(A)"
+    and "m:Lists(A)"
+  shows "k @ l @ m = (k @ l) @ m"
+  apply (tactic {* eq_coinduct3_tac
+    "{p. EX x y. p=<x,y> & (EX k:Lists (A). EX l:Lists (A). EX m:Lists (A). x=k @ l @ m & y= (k @ l) @ m) }" 1*})
+  apply (blast intro: prems)
+  apply safe
+  apply (drule ListsXH [THEN iffD1])
+  apply (tactic "EQgen_tac (simpset ()) [] 1")
+   prefer 2
+   apply blast
+  apply (tactic {* DEPTH_SOLVE (etac (XH_to_E (thm "ListsXH")) 1
+    THEN EQgen_tac (simpset ()) [] 1) *})
+  done
+
+
+subsection {* Appending anything to an infinite list doesn't alter it *}
+
+lemma ilist_append:
+  assumes "l:ILists(A)"
+  shows "l @ m = l"
+  apply (tactic {* eq_coinduct3_tac
+    "{p. EX x y. p=<x,y> & (EX l:ILists (A) .EX m. x=l@m & y=l)}" 1 *})
+  apply (blast intro: prems)
+  apply safe
+  apply (drule IListsXH [THEN iffD1])
+  apply (tactic "EQgen_tac (simpset ()) [] 1")
+  apply blast
+  done
+
+(*** The equivalance of two versions of an iteration function       ***)
+(*                                                                    *)
+(*        fun iter1(f,a) = a$iter1(f,f(a))                            *)
+(*        fun iter2(f,a) = a$map(f,iter2(f,a))                        *)
+
+lemma iter1B: "iter1(f,a) = a$iter1(f,f(a))"
+  apply (unfold iter1_def)
+  apply (rule letrecB [THEN trans])
+  apply simp
+  done
+
+lemma iter2B: "iter2(f,a) = a $ map(f,iter2(f,a))"
+  apply (unfold iter2_def)
+  apply (rule letrecB [THEN trans])
+  apply (rule refl)
+  done
+
+lemma iter2Blemma:
+  "n:Nat ==>  
+    map(f) ^ n ` iter2(f,a) = (f ^ n ` a) $ (map(f) ^ n ` map(f,iter2(f,a)))"
+  apply (rule_tac P = "%x. ?lhs (x) = ?rhs" in iter2B [THEN ssubst])
+  apply (simp add: nmapBcons)
+  done
+
+lemma iter1_iter2_eq: "iter1(f,a) = iter2(f,a)"
+  apply (tactic {* eq_coinduct3_tac
+    "{p. EX x y. p=<x,y> & (EX n:Nat. x=iter1 (f,f^n`a) & y=map (f) ^n`iter2 (f,a))}" 1*})
+  apply (fast intro!: napplyBzero [symmetric] napplyBzero [symmetric, THEN arg_cong])
+  apply (tactic {* EQgen_tac (simpset ()) [thm "iter1B", thm "iter2Blemma"] 1 *})
+  apply (subst napply_f, assumption)
+  apply (rule_tac f1 = f in napplyBsucc [THEN subst])
+  apply blast
+  done
 
 end