# HG changeset patch # User huffman # Date 1285944057 25200 # Node ID 20c74cf9c112826b3bd2ba180492ce9d8fcb8c1f # Parent da88519e6a0bb655ad970d0950a5c25eb92e0275 added lemmas to List_Cpo.thy diff -r da88519e6a0b -r 20c74cf9c112 src/HOLCF/Library/List_Cpo.thy --- a/src/HOLCF/Library/List_Cpo.thy Thu Sep 30 19:42:12 2010 -0700 +++ b/src/HOLCF/Library/List_Cpo.thy Fri Oct 01 07:40:57 2010 -0700 @@ -50,6 +50,24 @@ lemma below_Nil_iff [simp]: "xs \ [] \ xs = []" by (cases xs, simp_all) +lemma list_below_induct [consumes 1, case_names Nil Cons]: + assumes "xs \ ys" + assumes 1: "P [] []" + assumes 2: "\x y xs ys. \x \ y; xs \ ys; P xs ys\ \ P (x # xs) (y # ys)" + shows "P xs ys" +using `xs \ ys` +proof (induct xs arbitrary: ys) + case Nil thus ?case by (simp add: 1) +next + case (Cons x xs) thus ?case by (cases ys, simp_all add: 2) +qed + +lemma list_below_cases: + assumes "xs \ ys" + obtains "xs = []" and "ys = []" | + x y xs' ys' where "xs = x # xs'" and "ys = y # ys'" +using assms by (cases xs, simp, cases ys, auto) + text "Thanks to Joachim Breitner" lemma list_Cons_below: @@ -77,25 +95,36 @@ lemma below_same_length: "xs \ ys \ length xs = length ys" unfolding below_list_def by (rule list_all2_lengthD) +lemma list_chain_induct [consumes 1, case_names Nil Cons]: + assumes "chain S" + assumes 1: "P (\i. [])" + assumes 2: "\A B. chain A \ chain B \ P B \ P (\i. A i # B i)" + shows "P S" +using `chain S` +proof (induct "S 0" arbitrary: S) + case Nil + have "\i. S 0 \ S i" by (simp add: chain_mono [OF `chain S`]) + with Nil have "\i. S i = []" by simp + thus ?case by (simp add: 1) +next + case (Cons x xs) + have "\i. S 0 \ S i" by (simp add: chain_mono [OF `chain S`]) + hence *: "\i. S i \ []" by (rule all_forward, insert Cons) auto + have "chain (\i. hd (S i))" and "chain (\i. tl (S i))" + using `chain S` by simp_all + moreover have "P (\i. tl (S i))" + using `chain S` and `x # xs = S 0` [symmetric] + by (simp add: Cons(1)) + ultimately have "P (\i. hd (S i) # tl (S i))" + by (rule 2) + thus "P S" by (simp add: *) +qed + lemma list_chain_cases: assumes S: "chain S" - obtains "\i. S i = []" | - A B where "chain A" and "chain B" and "\i. S i = A i # B i" -proof (cases "S 0") - case Nil - have "\i. S 0 \ S i" by (simp add: chain_mono [OF S]) - with Nil have "\i. S i = []" by simp - thus ?thesis .. -next - case (Cons x xs) - have "\i. S 0 \ S i" by (simp add: chain_mono [OF S]) - hence *: "\i. S i \ []" by (rule all_forward) (auto simp add: Cons) - let ?A = "\i. hd (S i)" - let ?B = "\i. tl (S i)" - from S have "chain ?A" and "chain ?B" by simp_all - moreover have "\i. S i = ?A i # ?B i" by (simp add: *) - ultimately show ?thesis .. -qed + obtains "S = (\i. [])" | + A B where "chain A" and "chain B" and "S = (\i. A i # B i)" +using S by (induct rule: list_chain_induct) simp_all subsection {* Lists are a complete partial order *} @@ -107,38 +136,15 @@ unfolding is_lub_def is_ub_def Ball_def [symmetric] by (clarsimp, case_tac u, simp_all) -lemma list_cpo_lemma: - fixes S :: "nat \ 'a::cpo list" - assumes "chain S" and "\i. length (S i) = n" - shows "\x. range S <<| x" -using assms - apply (induct n arbitrary: S) - apply (subgoal_tac "S = (\i. [])") - apply (fast intro: lub_const) - apply (simp add: fun_eq_iff) - apply (drule_tac x="\i. tl (S i)" in meta_spec, clarsimp) - apply (rule_tac x="(\i. hd (S i)) # x" in exI) - apply (subgoal_tac "range (\i. hd (S i) # tl (S i)) = range S") - apply (erule subst) - apply (rule is_lub_Cons) - apply (rule thelubE [OF _ refl]) - apply (erule ch2ch_hd) - apply assumption - apply (rule_tac f="\S. range S" in arg_cong) - apply (rule ext) - apply (rule hd_Cons_tl) - apply (drule_tac x=i in spec, auto) -done - instance list :: (cpo) cpo proof fix S :: "nat \ 'a list" - assume "chain S" - hence "\i. S 0 \ S i" by (simp add: chain_mono) - hence "\i. length (S i) = length (S 0)" - by (fast intro: below_same_length [symmetric]) - with `chain S` show "\x. range S <<| x" - by (rule list_cpo_lemma) + assume "chain S" thus "\x. range S <<| x" + proof (induct rule: list_chain_induct) + case Nil thus ?case by (auto intro: lub_const) + next + case (Cons A B) thus ?case by (auto intro: is_lub_Cons cpo_lubI) + qed qed subsection {* Continuity of list operations *} @@ -196,6 +202,37 @@ shows "cont (\x. case l of [] \ f1 x | y # ys \ f2 x y ys)" using assms by (cases l) auto +text {* Lemma for proving continuity of recursive list functions: *} + +lemma list_contI: + fixes f :: "'a::cpo list \ 'b::cpo" + assumes f: "\x xs. f (x # xs) = g x xs (f xs)" + assumes g1: "\xs y. cont (\x. g x xs y)" + assumes g2: "\x y. cont (\xs. g x xs y)" + assumes g3: "\x xs. cont (\y. g x xs y)" + shows "cont f" +proof (rule contI2) + obtain h where h: "\x xs y. g x xs y = h\x\xs\y" + proof + fix x xs y show "g x xs y = (\ x xs y. g x xs y)\x\xs\y" + by (simp add: cont2cont_LAM g1 g2 g3) + qed + show mono: "monofun f" + apply (rule monofunI) + apply (erule list_below_induct) + apply simp + apply (simp add: f h monofun_cfun) + done + fix Y :: "nat \ 'a list" + assume "chain Y" thus "f (\i. Y i) \ (\i. f (Y i))" + apply (induct rule: list_chain_induct) + apply simp + apply (simp add: lub_Cons f h) + apply (simp add: contlub_cfun [symmetric] ch2ch_monofun [OF mono]) + apply (simp add: monofun_cfun) + done +qed + text {* There are probably lots of other list operations that also deserve to have continuity lemmas. I'll add more as they are needed. *}