# HG changeset patch # User huffman # Date 1287690696 25200 # Node ID 23a1cfdb5acbf41a3cde54397c12115c7d4df9a0 # Parent 54159b52f3392238b6d59695c2e6df67d881bf12 simplify some proofs, convert to Isar style diff -r 54159b52f339 -r 23a1cfdb5acb src/HOLCF/Up.thy --- a/src/HOLCF/Up.thy Thu Oct 21 12:03:49 2010 -0700 +++ b/src/HOLCF/Up.thy Thu Oct 21 12:51:36 2010 -0700 @@ -73,97 +73,62 @@ lemma is_lub_Iup: "range S <<| x \ range (\i. Iup (S i)) <<| Iup x" -apply (rule is_lubI) -apply (rule ub_rangeI) -apply (subst Iup_below) -apply (erule is_ub_lub) -apply (case_tac u) -apply (drule ub_rangeD) -apply simp -apply simp -apply (erule is_lub_lub) -apply (rule ub_rangeI) -apply (drule_tac i=i in ub_rangeD) -apply simp -done - -text {* Now some lemmas about chains of @{typ "'a u"} elements *} - -lemma up_lemma1: "z \ Ibottom \ Iup (THE a. Iup a = z) = z" -by (case_tac z, simp_all) - -lemma up_lemma2: - "\chain Y; Y j \ Ibottom\ \ Y (i + j) \ Ibottom" -apply (erule contrapos_nn) -apply (drule_tac i="j" and j="i + j" in chain_mono) -apply (rule le_add2) -apply (case_tac "Y j") -apply assumption -apply simp -done - -lemma up_lemma3: - "\chain Y; Y j \ Ibottom\ \ Iup (THE a. Iup a = Y (i + j)) = Y (i + j)" -by (rule up_lemma1 [OF up_lemma2]) - -lemma up_lemma4: - "\chain Y; Y j \ Ibottom\ \ chain (\i. THE a. Iup a = Y (i + j))" -apply (rule chainI) -apply (rule Iup_below [THEN iffD1]) -apply (subst up_lemma3, assumption+)+ -apply (simp add: chainE) -done - -lemma up_lemma5: - "\chain Y; Y j \ Ibottom\ \ - (\i. Y (i + j)) = (\i. Iup (THE a. Iup a = Y (i + j)))" -by (rule ext, rule up_lemma3 [symmetric]) - -lemma up_lemma6: - "\chain Y; Y j \ Ibottom\ - \ range Y <<| Iup (\i. THE a. Iup a = Y(i + j))" -apply (rule_tac j1 = j in is_lub_range_shift [THEN iffD1]) -apply assumption -apply (subst up_lemma5, assumption+) -apply (rule is_lub_Iup) -apply (rule cpo_lubI) -apply (erule (1) up_lemma4) -done +unfolding is_lub_def is_ub_def ball_simps +by (auto simp add: below_up_def split: u.split) lemma up_chain_lemma: - "chain Y \ - (\A. chain A \ (\i. Y i) = Iup (\i. A i) \ - (\j. \i. Y (i + j) = Iup (A i))) \ (Y = (\i. Ibottom))" -apply (rule disjCI) -apply (simp add: fun_eq_iff) -apply (erule exE, rename_tac j) -apply (rule_tac x="\i. THE a. Iup a = Y (i + j)" in exI) -apply (simp add: up_lemma4) -apply (simp add: up_lemma6 [THEN thelubI]) -apply (rule_tac x=j in exI) -apply (simp add: up_lemma3) -done - -lemma cpo_up: "chain (Y::nat \ 'a u) \ \x. range Y <<| x" -apply (frule up_chain_lemma, safe) -apply (rule_tac x="Iup (\i. A i)" in exI) -apply (erule_tac j="j" in is_lub_range_shift [THEN iffD1, standard]) -apply (simp add: is_lub_Iup cpo_lubI) -apply (rule exI, rule lub_const) -done + assumes Y: "chain Y" obtains "\i. Y i = Ibottom" + | A k where "\i. Iup (A i) = Y (i + k)" and "chain A" and "range Y <<| Iup (\i. A i)" +proof (cases "\k. Y k \ Ibottom") + case True + then obtain k where k: "Y k \ Ibottom" .. + def A \ "\i. THE a. Iup a = Y (i + k)" + have Iup_A: "\i. Iup (A i) = Y (i + k)" + proof + fix i :: nat + from Y le_add2 have "Y k \ Y (i + k)" by (rule chain_mono) + with k have "Y (i + k) \ Ibottom" by (cases "Y k", auto) + thus "Iup (A i) = Y (i + k)" + by (cases "Y (i + k)", simp_all add: A_def) + qed + from Y have chain_A: "chain A" + unfolding chain_def Iup_below [symmetric] + by (simp add: Iup_A) + hence "range A <<| (\i. A i)" + by (rule cpo_lubI) + hence "range (\i. Iup (A i)) <<| Iup (\i. A i)" + by (rule is_lub_Iup) + hence "range (\i. Y (i + k)) <<| Iup (\i. A i)" + by (simp only: Iup_A) + hence "range (\i. Y i) <<| Iup (\i. A i)" + by (simp only: is_lub_range_shift [OF Y]) + with Iup_A chain_A show ?thesis .. +next + case False + then have "\i. Y i = Ibottom" by simp + then show ?thesis .. +qed instance u :: (cpo) cpo -by intro_classes (rule cpo_up) +proof + fix S :: "nat \ 'a u" + assume S: "chain S" + thus "\x. range (\i. S i) <<| x" + proof (rule up_chain_lemma) + assume "\i. S i = Ibottom" + hence "range (\i. S i) <<| Ibottom" + by (simp add: lub_const) + thus ?thesis .. + next + fix A assume "range S <<| Iup (\i. A i)" + thus ?thesis .. + qed +qed subsection {* Lifted cpo is pointed *} -lemma least_up: "\x::'a u. \y. x \ y" -apply (rule_tac x = "Ibottom" in exI) -apply (rule minimal_up [THEN allI]) -done - instance u :: (cpo) pcpo -by intro_classes (rule least_up) +by intro_classes fast text {* for compatibility with old HOLCF-Version *} lemma inst_up_pcpo: "\ = Ibottom" @@ -192,13 +157,22 @@ done lemma cont_Ifup2: "cont (\x. Ifup f x)" -apply (rule contI) -apply (frule up_chain_lemma, safe) -apply (rule_tac j="j" in is_lub_range_shift [THEN iffD1, standard]) -apply (erule monofun_Ifup2 [THEN ch2ch_monofun]) -apply (simp add: cont_cfun_arg) -apply (simp add: lub_const) -done +proof (rule contI2) + fix Y assume Y: "chain Y" and Y': "chain (\i. Ifup f (Y i))" + from Y show "Ifup f (\i. Y i) \ (\i. Ifup f (Y i))" + proof (rule up_chain_lemma) + fix A and k + assume A: "\i. Iup (A i) = Y (i + k)" + assume "chain A" and "range Y <<| Iup (\i. A i)" + hence "Ifup f (\i. Y i) = (\i. Ifup f (Iup (A i)))" + by (simp add: thelubI contlub_cfun_arg) + also have "\ = (\i. Ifup f (Y (i + k)))" + by (simp add: A) + also have "\ = (\i. Ifup f (Y i))" + using Y' by (rule lub_range_shift) + finally show ?thesis by simp + qed simp +qed (rule monofun_Ifup2) subsection {* Continuous versions of constants *} @@ -251,18 +225,20 @@ text {* lifting preserves chain-finiteness *} lemma up_chain_cases: - "chain Y \ - (\A. chain A \ (\i. Y i) = up\(\i. A i) \ - (\j. \i. Y (i + j) = up\(A i))) \ Y = (\i. \)" -by (simp add: inst_up_pcpo up_def cont_Iup up_chain_lemma) + assumes Y: "chain Y" obtains "\i. Y i = \" + | A k where "\i. up\(A i) = Y (i + k)" and "chain A" and "(\i. Y i) = up\(\i. A i)" +apply (rule up_chain_lemma [OF Y]) +apply (simp_all add: inst_up_pcpo up_def cont_Iup thelubI) +done lemma compact_up: "compact x \ compact (up\x)" apply (rule compactI2) -apply (drule up_chain_cases, safe) +apply (erule up_chain_cases) +apply simp apply (drule (1) compactD2, simp) -apply (erule exE, rule_tac x="i + j" in exI) -apply simp -apply simp +apply (erule exE) +apply (drule_tac f="up" and x="x" in monofun_cfun_arg) +apply (simp, erule exI) done lemma compact_upD: "compact (up\x) \ compact x"