# HG changeset patch # User huffman # Date 1120177859 -7200 # Node ID d28314d2dce3b8234f73acd44d1f556f38fc4ac1 # Parent 53d4e0f2839baf8d598a7f9913b5e9b65ec3ac42 cleaned up diff -r 53d4e0f2839b -r d28314d2dce3 src/HOLCF/Pcpo.thy --- a/src/HOLCF/Pcpo.thy Fri Jul 01 01:50:46 2005 +0200 +++ b/src/HOLCF/Pcpo.thy Fri Jul 01 02:30:59 2005 +0200 @@ -17,31 +17,35 @@ axclass cpo < po -- {* class axiom: *} - cpo: "chain S ==> ? x. range S <<| x" + cpo: "chain S \ \x. range S <<| x" text {* in cpo's everthing equal to THE lub has lub properties for every chain *} -lemma thelubE: "[| chain(S); lub(range(S)) = (l::'a::cpo) |] ==> range(S) <<| l" +lemma thelubE: "\chain S; (\i. S i) = (l::'a::cpo)\ \ range S <<| l" by (blast dest: cpo intro: lubI) text {* Properties of the lub *} -lemma is_ub_thelub: "chain (S::nat => 'a::cpo) ==> S(x) << lub(range(S))" +lemma is_ub_thelub: "chain (S::nat \ 'a::cpo) \ S x \ (\i. S i)" by (blast dest: cpo intro: lubI [THEN is_ub_lub]) -lemma is_lub_thelub: "[| chain (S::nat => 'a::cpo); range(S) <| x |] ==> lub(range S) << x" +lemma is_lub_thelub: + "\chain (S::nat \ 'a::cpo); range S <| x\ \ (\i. S i) \ x" by (blast dest: cpo intro: lubI [THEN is_lub_lub]) -lemma lub_range_mono: "[| range X <= range Y; chain Y; chain (X::nat=>'a::cpo) |] ==> lub(range X) << lub(range Y)" +lemma lub_range_mono: + "\range X \ range Y; chain Y; chain (X::nat \ 'a::cpo)\ + \ (\i. X i) \ (\i. Y i)" apply (erule is_lub_thelub) apply (rule ub_rangeI) -apply (subgoal_tac "? j. X i = Y j") +apply (subgoal_tac "\j. X i = Y j") apply clarsimp apply (erule is_ub_thelub) apply auto done -lemma lub_range_shift: "chain (Y::nat=>'a::cpo) ==> lub(range (%i. Y(i + j))) = lub(range Y)" +lemma lub_range_shift: + "chain (Y::nat \ 'a::cpo) \ (\i. Y (i + j)) = (\i. Y i)" apply (rule antisym_less) apply (rule lub_range_mono) apply fast @@ -57,7 +61,8 @@ apply (rule le_add1) done -lemma maxinch_is_thelub: "chain Y ==> max_in_chain i Y = (lub(range(Y)) = ((Y i)::'a::cpo))" +lemma maxinch_is_thelub: + "chain Y \ max_in_chain i Y = ((\i. Y i) = ((Y i)::'a::cpo))" apply (rule iffI) apply (fast intro!: thelubI lub_finch1) apply (unfold max_in_chain_def) @@ -67,10 +72,11 @@ apply (force elim!: is_ub_thelub) done -text {* the @{text "<<"} relation between two chains is preserved by their lubs *} +text {* the @{text "\"} relation between two chains is preserved by their lubs *} -lemma lub_mono: "[|chain(C1::(nat=>'a::cpo));chain(C2); ALL k. C1(k) << C2(k)|] - ==> lub(range(C1)) << lub(range(C2))" +lemma lub_mono: + "\chain (X::nat \ 'a::cpo); chain Y; \k. X k \ Y k\ + \ (\i. X i) \ (\i. Y i)" apply (erule is_lub_thelub) apply (rule ub_rangeI) apply (rule trans_less) @@ -80,49 +86,47 @@ text {* the = relation between two chains is preserved by their lubs *} -lemma lub_equal: "[| chain(C1::(nat=>'a::cpo));chain(C2);ALL k. C1(k)=C2(k)|] - ==> lub(range(C1))=lub(range(C2))" +lemma lub_equal: + "\chain (X::nat \ 'a::cpo); chain Y; \k. X k = Y k\ + \ (\i. X i) = (\i. Y i)" by (simp only: expand_fun_eq [symmetric]) text {* more results about mono and = of lubs of chains *} -lemma lub_mono2: "[|EX j. ALL i. j X(i::nat)=Y(i);chain(X::nat=>'a::cpo);chain(Y)|] - ==> lub(range(X))<\j::nat. \i>j. X i = Y i; chain (X::nat=>'a::cpo); chain Y\ + \ (\i. X i) \ (\i. Y i)" apply (erule exE) apply (rule is_lub_thelub) apply assumption apply (rule ub_rangeI) -apply (case_tac "j X(i)=Y(i); chain(X::nat=>'a::cpo); chain(Y)|] - ==> lub(range(X))=lub(range(Y))" +lemma lub_equal2: + "\\j. \i>j. X i = Y i; chain (X::nat \ 'a::cpo); chain Y\ + \ (\i. X i) = (\i. Y i)" by (blast intro: antisym_less lub_mono2 sym) -lemma lub_mono3: "[|chain(Y::nat=>'a::cpo);chain(X); - ALL i. EX j. Y(i)<< X(j)|]==> lub(range(Y))<chain (Y::nat \ 'a::cpo); chain X; \i. \j. Y i \ X j\ + \ (\i. Y i) \ (\i. X i)" apply (rule is_lub_thelub) apply assumption apply (rule ub_rangeI) apply (erule allE) apply (erule exE) -apply (rule trans_less) -apply (rule_tac [2] is_ub_thelub) -prefer 2 apply (assumption) -apply assumption +apply (erule trans_less) +apply (erule is_ub_thelub) done lemma ch2ch_lub: @@ -177,20 +181,18 @@ text {* The class pcpo of pointed cpos *} axclass pcpo < cpo - least: "? x.!y. x<x. \y. x \ y" -consts - UU :: "'a::pcpo" +constdefs + UU :: "'a::pcpo" + "UU \ THE x. ALL y. x \ y" syntax (xsymbols) - UU :: "'a::pcpo" ("\") - -defs - UU_def: "UU == THE x. ALL y. x<") text {* derive the old rule minimal *} -lemma UU_least: "ALL z. UU << z" +lemma UU_least: "\z. \ \ z" apply (unfold UU_def) apply (rule theI') apply (rule ex_ex1I) @@ -198,13 +200,12 @@ apply (blast intro: antisym_less) done -lemmas minimal = UU_least [THEN spec, standard] - -declare minimal [iff] +lemma minimal [iff]: "\ \ x" +by (rule UU_least [THEN spec]) -text {* useful lemmas about @{term UU} *} +text {* useful lemmas about @{term \} *} -lemma eq_UU_iff: "(x=UU)=(x<) = (x \ \)" apply (rule iffI) apply (erule ssubst) apply (rule refl_less) @@ -213,34 +214,33 @@ apply (rule minimal) done -lemma UU_I: "x << UU ==> x = UU" +lemma UU_I: "x \ \ \ x = \" by (subst eq_UU_iff) -lemma not_less2not_eq: "~(x::'a::po)< ~x=y" +lemma not_less2not_eq: "\ (x::'a::po) \ y \ x \ y" by auto -lemma chain_UU_I: "[|chain(Y);lub(range(Y))=UU|] ==> ALL i. Y(i)=UU" +lemma chain_UU_I: "\chain Y; (\i. Y i) = \\ \ \i. Y i = \" apply (rule allI) -apply (rule antisym_less) -apply (rule_tac [2] minimal) +apply (rule UU_I) apply (erule subst) apply (erule is_ub_thelub) done -lemma chain_UU_I_inverse: "ALL i. Y(i::nat)=UU ==> lub(range(Y::(nat=>'a::pcpo)))=UU" +lemma chain_UU_I_inverse: "\i::nat. Y i = \ \ (\i. Y i) = \" apply (rule lub_chain_maxelem) apply (erule spec) apply simp done -lemma chain_UU_I_inverse2: "~lub(range(Y::(nat=>'a::pcpo)))=UU ==> EX i.~ Y(i)=UU" +lemma chain_UU_I_inverse2: "(\i. Y i) \ \ \ \i::nat. Y i \ \" by (blast intro: chain_UU_I_inverse) -lemma notUU_I: "[| x< ~y=UU" +lemma notUU_I: "\x \ y; x \ \\ \ y \ \" by (blast intro: UU_I) lemma chain_mono2: - "[|EX j. ~Y(j)=UU;chain(Y::nat=>'a::pcpo)|] ==> EX j. ALL i. j~Y(i)=UU" + "\\j::nat. Y j \ \; chain Y\ \ \j. \i>j. Y i \ \" by (blast dest: notUU_I chain_mono) subsection {* Chain-finite and flat cpos *} @@ -248,17 +248,17 @@ text {* further useful classes for HOLCF domains *} axclass chfin < po - chfin: "!Y. chain Y-->(? n. max_in_chain n Y)" + chfin: "\Y. chain Y \ (\n. max_in_chain n Y)" axclass flat < pcpo - ax_flat: "! x y. x << y --> (x = UU) | (x=y)" + ax_flat: "\x y. x \ y \ (x = \) \ (x = y)" text {* some properties for chfin and flat *} text {* chfin types are cpo *} lemma chfin_imp_cpo: - "chain (S::nat=>'a::chfin) ==> EX x. range S <<| x" + "chain (S::nat \ 'a::chfin) \ \x. range S <<| x" apply (frule chfin [rule_format]) apply (blast intro: lub_finch1) done @@ -269,14 +269,14 @@ text {* flat types are chfin *} lemma flat_imp_chfin: - "ALL Y::nat=>'a::flat. chain Y --> (EX n. max_in_chain n Y)" + "\Y::nat \ 'a::flat. chain Y \ (\n. max_in_chain n Y)" apply (unfold max_in_chain_def) apply clarify -apply (case_tac "ALL i. Y (i) =UU") +apply (case_tac "\i. Y i = \") apply simp apply simp apply (erule exE) -apply (rule_tac x = "i" in exI) +apply (rule_tac x="i" in exI) apply clarify apply (erule le_imp_less_or_eq [THEN disjE]) apply safe @@ -288,18 +288,18 @@ text {* flat subclass of chfin @{text "-->"} @{text adm_flat} not needed *} -lemma flat_eq: "(a::'a::flat) ~= UU ==> a << b = (a = b)" +lemma flat_eq: "(a::'a::flat) \ \ \ a \ b = (a = b)" by (safe dest!: ax_flat [rule_format]) -lemma chfin2finch: "chain (Y::nat=>'a::chfin) ==> finite_chain Y" +lemma chfin2finch: "chain (Y::nat \ 'a::chfin) \ finite_chain Y" by (simp add: chfin finite_chain_def) text {* lemmata for improved admissibility introdution rule *} lemma infinite_chain_adm_lemma: -"[|chain Y; ALL i. P (Y i); - (!!Y. [| chain Y; ALL i. P (Y i); ~ finite_chain Y |] ==> P (lub(range Y))) - |] ==> P (lub (range Y))" + "\chain Y; \i. P (Y i); + \Y. \chain Y; \i. P (Y i); \ finite_chain Y\ \ P (\i. Y i)\ + \ P (\i. Y i)" apply (case_tac "finite_chain Y") prefer 2 apply fast apply (unfold finite_chain_def) @@ -310,10 +310,9 @@ done lemma increasing_chain_adm_lemma: -"[|chain Y; ALL i. P (Y i); - (!!Y. [| chain Y; ALL i. P (Y i); - ALL i. EX j. i < j & Y i ~= Y j & Y i << Y j|] - ==> P (lub (range Y))) |] ==> P (lub (range Y))" + "\chain Y; \i. P (Y i); \Y. \chain Y; \i. P (Y i); + \i. \j>i. Y i \ Y j \ Y i \ Y j\ \ P (\i. Y i)\ + \ P (\i. Y i)" apply (erule infinite_chain_adm_lemma) apply assumption apply (erule thin_rl) @@ -322,4 +321,4 @@ apply (fast dest: le_imp_less_or_eq elim: chain_mono) done -end +end