(* Title: HOLCF/Porder.thy ID: $Id: Porder.thy,v 1.49 2008/03/26 23:27:16 huffman Exp $ Author: Franz Regensburger and Brian Huffman *) header {* Partial orders *} theory Porder imports Datatype Finite_Set begin subsection {* Type class for partial orders *} class sq_ord = type + fixes sq_le :: "'a => 'a => bool" notation sq_le (infixl "<<" 55) notation (xsymbols) sq_le (infixl "\<sqsubseteq>" 55) class preorder = sq_ord + assumes refl_less [iff]: "x \<sqsubseteq> x" assumes trans_less: "[|x \<sqsubseteq> y; y \<sqsubseteq> z|] ==> x \<sqsubseteq> z" class po = preorder + assumes antisym_less: "[|x \<sqsubseteq> y; y \<sqsubseteq> x|] ==> x = y" text {* minimal fixes least element *} lemma minimal2UU[OF allI] : "∀x::'a::po. uu \<sqsubseteq> x ==> uu = (THE u. ∀y. u \<sqsubseteq> y)" by (blast intro: theI2 antisym_less) text {* the reverse law of anti-symmetry of @{term "op <<"} *} lemma antisym_less_inverse: "(x::'a::po) = y ==> x \<sqsubseteq> y ∧ y \<sqsubseteq> x" by simp lemma box_less: "[|(a::'a::po) \<sqsubseteq> b; c \<sqsubseteq> a; b \<sqsubseteq> d|] ==> c \<sqsubseteq> d" by (rule trans_less [OF trans_less]) lemma po_eq_conv: "((x::'a::po) = y) = (x \<sqsubseteq> y ∧ y \<sqsubseteq> x)" by (fast elim!: antisym_less_inverse intro!: antisym_less) lemma rev_trans_less: "[|(y::'a::po) \<sqsubseteq> z; x \<sqsubseteq> y|] ==> x \<sqsubseteq> z" by (rule trans_less) lemma sq_ord_less_eq_trans: "[|a \<sqsubseteq> b; b = c|] ==> a \<sqsubseteq> c" by (rule subst) lemma sq_ord_eq_less_trans: "[|a = b; b \<sqsubseteq> c|] ==> a \<sqsubseteq> c" by (rule ssubst) lemmas HOLCF_trans_rules [trans] = trans_less antisym_less sq_ord_less_eq_trans sq_ord_eq_less_trans subsection {* Upper bounds *} definition is_ub :: "['a set, 'a::po] => bool" (infixl "<|" 55) where "(S <| x) = (∀y. y ∈ S --> y \<sqsubseteq> x)" lemma is_ubI: "(!!x. x ∈ S ==> x \<sqsubseteq> u) ==> S <| u" by (simp add: is_ub_def) lemma is_ubD: "[|S <| u; x ∈ S|] ==> x \<sqsubseteq> u" by (simp add: is_ub_def) lemma ub_imageI: "(!!x. x ∈ S ==> f x \<sqsubseteq> u) ==> (λx. f x) ` S <| u" unfolding is_ub_def by fast lemma ub_imageD: "[|f ` S <| u; x ∈ S|] ==> f x \<sqsubseteq> u" unfolding is_ub_def by fast lemma ub_rangeI: "(!!i. S i \<sqsubseteq> x) ==> range S <| x" unfolding is_ub_def by fast lemma ub_rangeD: "range S <| x ==> S i \<sqsubseteq> x" unfolding is_ub_def by fast lemma is_ub_empty [simp]: "{} <| u" unfolding is_ub_def by fast lemma is_ub_insert [simp]: "(insert x A) <| y = (x \<sqsubseteq> y ∧ A <| y)" unfolding is_ub_def by fast lemma is_ub_upward: "[|S <| x; x \<sqsubseteq> y|] ==> S <| y" unfolding is_ub_def by (fast intro: trans_less) subsection {* Least upper bounds *} definition is_lub :: "['a set, 'a::po] => bool" (infixl "<<|" 55) where "(S <<| x) = (S <| x ∧ (∀u. S <| u --> x \<sqsubseteq> u))" definition lub :: "'a set => 'a::po" where "lub S = (THE x. S <<| x)" syntax "_BLub" :: "[pttrn, 'a set, 'b] => 'b" ("(3LUB _:_./ _)" [0,0, 10] 10) syntax (xsymbols) "_BLub" :: "[pttrn, 'a set, 'b] => 'b" ("(3\<Squnion>_∈_./ _)" [0,0, 10] 10) translations "LUB x:A. t" == "CONST lub ((%x. t) ` A)" abbreviation Lub (binder "LUB " 10) where "LUB n. t n == lub (range t)" notation (xsymbols) Lub (binder "\<Squnion> " 10) text {* access to some definition as inference rule *} lemma is_lubD1: "S <<| x ==> S <| x" unfolding is_lub_def by fast lemma is_lub_lub: "[|S <<| x; S <| u|] ==> x \<sqsubseteq> u" unfolding is_lub_def by fast lemma is_lubI: "[|S <| x; !!u. S <| u ==> x \<sqsubseteq> u|] ==> S <<| x" unfolding is_lub_def by fast text {* lubs are unique *} lemma unique_lub: "[|S <<| x; S <<| y|] ==> x = y" apply (unfold is_lub_def is_ub_def) apply (blast intro: antisym_less) done text {* technical lemmas about @{term lub} and @{term is_lub} *} lemma lubI: "M <<| x ==> M <<| lub M" apply (unfold lub_def) apply (rule theI) apply assumption apply (erule (1) unique_lub) done lemma thelubI: "M <<| l ==> lub M = l" by (rule unique_lub [OF lubI]) lemma is_lub_singleton: "{x} <<| x" by (simp add: is_lub_def) lemma lub_singleton [simp]: "lub {x} = x" by (rule thelubI [OF is_lub_singleton]) lemma is_lub_bin: "x \<sqsubseteq> y ==> {x, y} <<| y" by (simp add: is_lub_def) lemma lub_bin: "x \<sqsubseteq> y ==> lub {x, y} = y" by (rule is_lub_bin [THEN thelubI]) lemma is_lub_maximal: "[|S <| x; x ∈ S|] ==> S <<| x" by (erule is_lubI, erule (1) is_ubD) lemma lub_maximal: "[|S <| x; x ∈ S|] ==> lub S = x" by (rule is_lub_maximal [THEN thelubI]) subsection {* Countable chains *} definition -- {* Here we use countable chains and I prefer to code them as functions! *} chain :: "(nat => 'a::po) => bool" where "chain Y = (∀i. Y i \<sqsubseteq> Y (Suc i))" lemma chainI: "(!!i. Y i \<sqsubseteq> Y (Suc i)) ==> chain Y" unfolding chain_def by fast lemma chainE: "chain Y ==> Y i \<sqsubseteq> Y (Suc i)" unfolding chain_def by fast text {* chains are monotone functions *} lemma chain_mono: assumes Y: "chain Y" shows "i ≤ j ==> Y i \<sqsubseteq> Y j" apply (induct j) apply simp apply (erule le_SucE) apply (rule trans_less [OF _ chainE [OF Y]]) apply simp apply simp done lemma chain_mono_less: "[|chain Y; i < j|] ==> Y i \<sqsubseteq> Y j" by (erule chain_mono, simp) lemma chain_shift: "chain Y ==> chain (λi. Y (i + j))" apply (rule chainI) apply simp apply (erule chainE) done text {* technical lemmas about (least) upper bounds of chains *} lemma is_ub_lub: "range S <<| x ==> S i \<sqsubseteq> x" by (rule is_lubD1 [THEN ub_rangeD]) lemma is_ub_range_shift: "chain S ==> range (λi. S (i + j)) <| x = range S <| x" apply (rule iffI) apply (rule ub_rangeI) apply (rule_tac y="S (i + j)" in trans_less) apply (erule chain_mono) apply (rule le_add1) apply (erule ub_rangeD) apply (rule ub_rangeI) apply (erule ub_rangeD) done lemma is_lub_range_shift: "chain S ==> range (λi. S (i + j)) <<| x = range S <<| x" by (simp add: is_lub_def is_ub_range_shift) text {* the lub of a constant chain is the constant *} lemma chain_const [simp]: "chain (λi. c)" by (simp add: chainI) lemma lub_const: "range (λx. c) <<| c" by (blast dest: ub_rangeD intro: is_lubI ub_rangeI) lemma thelub_const [simp]: "(\<Squnion>i. c) = c" by (rule lub_const [THEN thelubI]) subsection {* Totally ordered sets *} definition -- {* Arbitrary chains are total orders *} tord :: "'a::po set => bool" where "tord S = (∀x y. x ∈ S ∧ y ∈ S --> (x \<sqsubseteq> y ∨ y \<sqsubseteq> x))" text {* The range of a chain is a totally ordered *} lemma chain_tord: "chain Y ==> tord (range Y)" unfolding tord_def apply (clarify, rename_tac i j) apply (rule_tac x=i and y=j in linorder_le_cases) apply (fast intro: chain_mono)+ done lemma finite_tord_has_max: "[|finite S; S ≠ {}; tord S|] ==> ∃y∈S. ∀x∈S. x \<sqsubseteq> y" apply (induct S rule: finite_ne_induct) apply simp apply (drule meta_mp, simp add: tord_def) apply (erule bexE, rename_tac z) apply (subgoal_tac "x \<sqsubseteq> z ∨ z \<sqsubseteq> x") apply (erule disjE) apply (rule_tac x="z" in bexI, simp, simp) apply (rule_tac x="x" in bexI) apply (clarsimp elim!: rev_trans_less) apply simp apply (simp add: tord_def) done subsection {* Finite chains *} definition -- {* finite chains, needed for monotony of continuous functions *} max_in_chain :: "[nat, nat => 'a::po] => bool" where "max_in_chain i C = (∀j. i ≤ j --> C i = C j)" definition finite_chain :: "(nat => 'a::po) => bool" where "finite_chain C = (chain C ∧ (∃i. max_in_chain i C))" text {* results about finite chains *} lemma max_in_chainI: "(!!j. i ≤ j ==> Y i = Y j) ==> max_in_chain i Y" unfolding max_in_chain_def by fast lemma max_in_chainD: "[|max_in_chain i Y; i ≤ j|] ==> Y i = Y j" unfolding max_in_chain_def by fast lemma lub_finch1: "[|chain C; max_in_chain i C|] ==> range C <<| C i" apply (rule is_lubI) apply (rule ub_rangeI, rename_tac j) apply (rule_tac x=i and y=j in linorder_le_cases) apply (drule (1) max_in_chainD, simp) apply (erule (1) chain_mono) apply (erule ub_rangeD) done lemma lub_finch2: "finite_chain C ==> range C <<| C (LEAST i. max_in_chain i C)" apply (unfold finite_chain_def) apply (erule conjE) apply (erule LeastI2_ex) apply (erule (1) lub_finch1) done lemma finch_imp_finite_range: "finite_chain Y ==> finite (range Y)" apply (unfold finite_chain_def, clarify) apply (rule_tac f="Y" and n="Suc i" in nat_seg_image_imp_finite) apply (rule equalityI) apply (rule subsetI) apply (erule rangeE, rename_tac j) apply (rule_tac x=i and y=j in linorder_le_cases) apply (subgoal_tac "Y j = Y i", simp) apply (simp add: max_in_chain_def) apply simp apply fast done lemma finite_range_imp_finch: "[|chain Y; finite (range Y)|] ==> finite_chain Y" apply (subgoal_tac "∃y∈range Y. ∀x∈range Y. x \<sqsubseteq> y") apply (clarsimp, rename_tac i) apply (subgoal_tac "max_in_chain i Y") apply (simp add: finite_chain_def exI) apply (simp add: max_in_chain_def po_eq_conv chain_mono) apply (erule finite_tord_has_max, simp) apply (erule chain_tord) done lemma bin_chain: "x \<sqsubseteq> y ==> chain (λi. if i=0 then x else y)" by (rule chainI, simp) lemma bin_chainmax: "x \<sqsubseteq> y ==> max_in_chain (Suc 0) (λi. if i=0 then x else y)" by (unfold max_in_chain_def, simp) lemma lub_bin_chain: "x \<sqsubseteq> y ==> range (λi::nat. if i=0 then x else y) <<| y" apply (frule bin_chain) apply (drule bin_chainmax) apply (drule (1) lub_finch1) apply simp done text {* the maximal element in a chain is its lub *} lemma lub_chain_maxelem: "[|Y i = c; ∀i. Y i \<sqsubseteq> c|] ==> lub (range Y) = c" by (blast dest: ub_rangeD intro: thelubI is_lubI ub_rangeI) subsection {* Directed sets *} definition directed :: "'a::po set => bool" where "directed S = ((∃x. x ∈ S) ∧ (∀x∈S. ∀y∈S. ∃z∈S. x \<sqsubseteq> z ∧ y \<sqsubseteq> z))" lemma directedI: assumes 1: "∃z. z ∈ S" assumes 2: "!!x y. [|x ∈ S; y ∈ S|] ==> ∃z∈S. x \<sqsubseteq> z ∧ y \<sqsubseteq> z" shows "directed S" unfolding directed_def using prems by fast lemma directedD1: "directed S ==> ∃z. z ∈ S" unfolding directed_def by fast lemma directedD2: "[|directed S; x ∈ S; y ∈ S|] ==> ∃z∈S. x \<sqsubseteq> z ∧ y \<sqsubseteq> z" unfolding directed_def by fast lemma directedE1: assumes S: "directed S" obtains z where "z ∈ S" by (insert directedD1 [OF S], fast) lemma directedE2: assumes S: "directed S" assumes x: "x ∈ S" and y: "y ∈ S" obtains z where "z ∈ S" "x \<sqsubseteq> z" "y \<sqsubseteq> z" by (insert directedD2 [OF S x y], fast) lemma directed_finiteI: assumes U: "!!U. [|finite U; U ⊆ S|] ==> ∃z∈S. U <| z" shows "directed S" proof (rule directedI) have "finite {}" and "{} ⊆ S" by simp_all hence "∃z∈S. {} <| z" by (rule U) thus "∃z. z ∈ S" by simp next fix x y assume "x ∈ S" and "y ∈ S" hence "finite {x, y}" and "{x, y} ⊆ S" by simp_all hence "∃z∈S. {x, y} <| z" by (rule U) thus "∃z∈S. x \<sqsubseteq> z ∧ y \<sqsubseteq> z" by simp qed lemma directed_finiteD: assumes S: "directed S" shows "[|finite U; U ⊆ S|] ==> ∃z∈S. U <| z" proof (induct U set: finite) case empty from S have "∃z. z ∈ S" by (rule directedD1) thus "∃z∈S. {} <| z" by simp next case (insert x F) from `insert x F ⊆ S` have xS: "x ∈ S" and FS: "F ⊆ S" by simp_all from FS have "∃y∈S. F <| y" by fact then obtain y where yS: "y ∈ S" and Fy: "F <| y" .. obtain z where zS: "z ∈ S" and xz: "x \<sqsubseteq> z" and yz: "y \<sqsubseteq> z" using S xS yS by (rule directedE2) from Fy yz have "F <| z" by (rule is_ub_upward) with xz have "insert x F <| z" by simp with zS show "∃z∈S. insert x F <| z" .. qed lemma not_directed_empty [simp]: "¬ directed {}" by (rule notI, drule directedD1, simp) lemma directed_singleton: "directed {x}" by (rule directedI, auto) lemma directed_bin: "x \<sqsubseteq> y ==> directed {x, y}" by (rule directedI, auto) lemma directed_chain: "chain S ==> directed (range S)" apply (rule directedI) apply (rule_tac x="S 0" in exI, simp) apply (clarify, rename_tac m n) apply (rule_tac x="S (max m n)" in bexI) apply (simp add: chain_mono) apply simp done end
lemma minimal2UU:
(!!x. uu << x) ==> uu = (THE u. ∀y. u << y)
lemma antisym_less_inverse:
x = y ==> x << y ∧ y << x
lemma box_less:
[| a << b; c << a; b << d |] ==> c << d
lemma po_eq_conv:
(x = y) = (x << y ∧ y << x)
lemma rev_trans_less:
[| y << z; x << y |] ==> x << z
lemma sq_ord_less_eq_trans:
[| a << b; b = c |] ==> a << c
lemma sq_ord_eq_less_trans:
[| a = b; b << c |] ==> a << c
lemma HOLCF_trans_rules:
[| x << y; y << z |] ==> x << z
[| x << y; y << x |] ==> x = y
[| a << b; b = c |] ==> a << c
[| a = b; b << c |] ==> a << c
lemma is_ubI:
(!!x. x ∈ S ==> x << u) ==> S <| u
lemma is_ubD:
[| S <| u; x ∈ S |] ==> x << u
lemma ub_imageI:
(!!x. x ∈ S ==> f x << u) ==> f ` S <| u
lemma ub_imageD:
[| f ` S <| u; x ∈ S |] ==> f x << u
lemma ub_rangeI:
(!!i. S i << x) ==> range S <| x
lemma ub_rangeD:
range S <| x ==> S i << x
lemma is_ub_empty:
{} <| u
lemma is_ub_insert:
insert x A <| y = (x << y ∧ A <| y)
lemma is_ub_upward:
[| S <| x; x << y |] ==> S <| y
lemma is_lubD1:
S <<| x ==> S <| x
lemma is_lub_lub:
[| S <<| x; S <| u |] ==> x << u
lemma is_lubI:
[| S <| x; !!u. S <| u ==> x << u |] ==> S <<| x
lemma unique_lub:
[| S <<| x; S <<| y |] ==> x = y
lemma lubI:
M <<| x ==> M <<| lub M
lemma thelubI:
M <<| l ==> lub M = l
lemma is_lub_singleton:
{x} <<| x
lemma lub_singleton:
lub {x} = x
lemma is_lub_bin:
x << y ==> {x, y} <<| y
lemma lub_bin:
x << y ==> lub {x, y} = y
lemma is_lub_maximal:
[| S <| x; x ∈ S |] ==> S <<| x
lemma lub_maximal:
[| S <| x; x ∈ S |] ==> lub S = x
lemma chainI:
(!!i. Y i << Y (Suc i)) ==> chain Y
lemma chainE:
chain Y ==> Y i << Y (Suc i)
lemma chain_mono:
[| chain Y; i ≤ j |] ==> Y i << Y j
lemma chain_mono_less:
[| chain Y; i < j |] ==> Y i << Y j
lemma chain_shift:
chain Y ==> chain (λi. Y (i + j))
lemma is_ub_lub:
range S <<| x ==> S i << x
lemma is_ub_range_shift:
chain S ==> range (λi. S (i + j)) <| x = range S <| x
lemma is_lub_range_shift:
chain S ==> range (λi. S (i + j)) <<| x = range S <<| x
lemma chain_const:
chain (λi. c)
lemma lub_const:
range (λx. c) <<| c
lemma thelub_const:
(LUB i. c) = c
lemma chain_tord:
chain Y ==> tord (range Y)
lemma finite_tord_has_max:
[| finite S; S ≠ {}; tord S |] ==> ∃y∈S. ∀x∈S. x << y
lemma max_in_chainI:
(!!j. i ≤ j ==> Y i = Y j) ==> max_in_chain i Y
lemma max_in_chainD:
[| max_in_chain i Y; i ≤ j |] ==> Y i = Y j
lemma lub_finch1:
[| chain C; max_in_chain i C |] ==> range C <<| C i
lemma lub_finch2:
finite_chain C ==> range C <<| C (LEAST i. max_in_chain i C)
lemma finch_imp_finite_range:
finite_chain Y ==> finite (range Y)
lemma finite_range_imp_finch:
[| chain Y; finite (range Y) |] ==> finite_chain Y
lemma bin_chain:
x << y ==> chain (λi. if i = 0 then x else y)
lemma bin_chainmax:
x << y ==> max_in_chain (Suc 0) (λi. if i = 0 then x else y)
lemma lub_bin_chain:
x << y ==> range (λi. if i = 0 then x else y) <<| y
lemma lub_chain_maxelem:
[| Y i = c; ∀i. Y i << c |] ==> Lub Y = c
lemma directedI:
[| ∃z. z ∈ S; !!x y. [| x ∈ S; y ∈ S |] ==> ∃z∈S. x << z ∧ y << z |]
==> directed S
lemma directedD1:
directed S ==> ∃z. z ∈ S
lemma directedD2:
[| directed S; x ∈ S; y ∈ S |] ==> ∃z∈S. x << z ∧ y << z
lemma directedE1:
[| directed S; !!z. z ∈ S ==> thesis |] ==> thesis
lemma directedE2:
[| directed S; x ∈ S; y ∈ S; !!z. [| z ∈ S; x << z; y << z |] ==> thesis |]
==> thesis
lemma directed_finiteI:
(!!U. [| finite U; U ⊆ S |] ==> ∃z∈S. U <| z) ==> directed S
lemma directed_finiteD:
[| directed S; finite U; U ⊆ S |] ==> ∃z∈S. U <| z
lemma not_directed_empty:
¬ directed {}
lemma directed_singleton:
directed {x}
lemma directed_bin:
x << y ==> directed {x, y}
lemma directed_chain:
chain S ==> directed (range S)