# HG changeset patch # User nipkow # Date 1107627851 -3600 # Node ID dd4ab096f082ec0aff19619843d1428740d2a7c7 # Parent 419dc5ffe8bc797dcd4b4f48b17caa1b467b3a3f Added Lattice locale diff -r 419dc5ffe8bc -r dd4ab096f082 src/HOL/Finite_Set.thy --- a/src/HOL/Finite_Set.thy Fri Feb 04 18:35:46 2005 +0100 +++ b/src/HOL/Finite_Set.thy Sat Feb 05 19:24:11 2005 +0100 @@ -1932,18 +1932,18 @@ subsubsection{* Semi-Lattices *} locale ACIfSL = ACIf + - fixes below :: "'a \ 'a \ bool" (infixl "\" 50) - assumes below_def: "x \ y = (x\y = x)" + fixes below :: "'a \ 'a \ bool" (infixl "\" 50) + assumes below_def: "(x \ y) = (x\y = x)" locale ACIfSLlin = ACIfSL + assumes lin: "x\y \ {x,y}" -lemma (in ACIfSL) below_refl[simp]: "x \ x" +lemma (in ACIfSL) below_refl[simp]: "x \ x" by(simp add: below_def idem) -lemma (in ACIfSL) below_f_conv[simp]: "x \ y \ z = (x \ y \ x \ z)" +lemma (in ACIfSL) below_f_conv[simp]: "x \ y \ z = (x \ y \ x \ z)" proof - assume "x \ y \ z" + assume "x \ y \ z" hence xyzx: "x \ (y \ z) = x" by(simp add: below_def) have "x \ y = x" proof - @@ -1959,40 +1959,40 @@ also have "\ = x" by(rule xyzx) finally show ?thesis . qed - ultimately show "x \ y \ x \ z" by(simp add: below_def) + ultimately show "x \ y \ x \ z" by(simp add: below_def) next - assume a: "x \ y \ x \ z" + assume a: "x \ y \ x \ z" hence y: "x \ y = x" and z: "x \ z = x" by(simp_all add: below_def) have "x \ (y \ z) = (x \ y) \ z" by(simp add:assoc) also have "x \ y = x" using a by(simp_all add: below_def) also have "x \ z = x" using a by(simp_all add: below_def) - finally show "x \ y \ z" by(simp_all add: below_def) + finally show "x \ y \ z" by(simp_all add: below_def) qed lemma (in ACIfSLlin) above_f_conv: - "x \ y \ z = (x \ z \ y \ z)" + "x \ y \ z = (x \ z \ y \ z)" proof - assume a: "x \ y \ z" + assume a: "x \ y \ z" have "x \ y = x \ x \ y = y" using lin[of x y] by simp - thus "x \ z \ y \ z" + thus "x \ z \ y \ z" proof - assume "x \ y = x" hence "x \ z" by(rule subst)(rule a) thus ?thesis .. + assume "x \ y = x" hence "x \ z" by(rule subst)(rule a) thus ?thesis .. next - assume "x \ y = y" hence "y \ z" by(rule subst)(rule a) thus ?thesis .. + assume "x \ y = y" hence "y \ z" by(rule subst)(rule a) thus ?thesis .. qed next - assume "x \ z \ y \ z" - thus "x \ y \ z" + assume "x \ z \ y \ z" + thus "x \ y \ z" proof - assume a: "x \ z" + assume a: "x \ z" have "(x \ y) \ z = (x \ z) \ y" by(simp add:ACI) also have "x \ z = x" using a by(simp add:below_def) - finally show "x \ y \ z" by(simp add:below_def) + finally show "x \ y \ z" by(simp add:below_def) next - assume a: "y \ z" + assume a: "y \ z" have "(x \ y) \ z = x \ (y \ z)" by(simp add:ACI) also have "y \ z = y" using a by(simp add:below_def) - finally show "x \ y \ z" by(simp add:below_def) + finally show "x \ y \ z" by(simp add:below_def) qed qed @@ -2033,13 +2033,13 @@ lemma (in ACIfSL) below_fold1_iff: assumes A: "finite A" "A \ {}" -shows "x \ fold1 f A = (\a\A. x \ a)" +shows "x \ fold1 f A = (\a\A. x \ a)" using A by(induct rule:finite_ne_induct) simp_all lemma (in ACIfSL) fold1_belowI: assumes A: "finite A" "A \ {}" -shows "a \ A \ fold1 f A \ a" +shows "a \ A \ fold1 f A \ a" using A proof (induct rule:finite_ne_induct) case singleton thus ?case by simp @@ -2051,7 +2051,7 @@ assume "a = x" thus ?thesis using insert by(simp add:below_def ACI) next assume "a \ F" - hence bel: "fold1 op \ F \ a" by(rule insert) + hence bel: "fold1 op \ F \ a" by(rule insert) have "fold1 op \ (insert x F) \ a = x \ (fold1 op \ F \ a)" using insert by(simp add:below_def ACI) also have "fold1 op \ F \ a = fold1 op \ F" @@ -2064,10 +2064,160 @@ lemma (in ACIfSLlin) fold1_below_iff: assumes A: "finite A" "A \ {}" -shows "fold1 f A \ x = (\a\A. a \ x)" +shows "fold1 f A \ x = (\a\A. a \ x)" using A by(induct rule:finite_ne_induct)(simp_all add:above_f_conv) +subsubsection{* Lattices *} + +locale Lattice = + fixes below :: "'a \ 'a \ bool" (infixl "\" 50) + and inf :: "'a \ 'a \ 'a" (infixl "\" 70) + and sup :: "'a \ 'a \ 'a" (infixl "\" 65) + and Inf :: "'a set \ 'a" ("\_" [900] 900) + and Sup :: "'a set \ 'a" ("\_" [900] 900) + assumes refl: "x \ x" + and trans: "x \ y \ y \ z \ x \ z" + and antisym: "x \ y \ y \ x \ x = y" + and inf_le1: "x \ y \ x" and inf_le2: "x \ y \ y" + and inf_least: "x \ y \ x \ z \ x \ y \ z" + and sup_ge1: "x \ x \ y" and sup_ge2: "y \ x \ y" + and sup_greatest: "y \ x \ z \ x \ y \ z \ x" +(* FIXME *) + and sup_inf_distrib1: "x \ y \ z = (x \ y) \ (x \ z)" + and sup_inf_distrib2: "x \ y \ z = (x \ z) \ (y \ z)" + defines "Inf == fold1 inf" and "Sup == fold1 sup" + + +lemma (in Lattice) inf_comm: "(x \ y) = (y \ x)" +by(blast intro: antisym inf_le1 inf_le2 inf_least) + +lemma (in Lattice) sup_comm: "(x \ y) = (y \ x)" +by(blast intro: antisym sup_ge1 sup_ge2 sup_greatest) + +lemma (in Lattice) inf_assoc: "(x \ y) \ z = x \ (y \ z)" +by(blast intro: antisym inf_le1 inf_le2 inf_least trans) + +lemma (in Lattice) sup_assoc: "(x \ y) \ z = x \ (y \ z)" +by(blast intro!: antisym sup_ge1 sup_ge2 intro: sup_greatest trans) + +lemma (in Lattice) inf_idem: "x \ x = x" +by(blast intro: antisym inf_le1 inf_le2 inf_least refl) + +lemma (in Lattice) sup_idem: "x \ x = x" +by(blast intro: antisym sup_ge1 sup_ge2 sup_greatest refl) + +text{* Lattices are semilattices *} + +lemma (in Lattice) ACf_inf: "ACf inf" +by(blast intro: ACf.intro inf_comm inf_assoc) + +lemma (in Lattice) ACf_sup: "ACf sup" +by(blast intro: ACf.intro sup_comm sup_assoc) + +lemma (in Lattice) ACIf_inf: "ACIf inf" +apply(rule ACIf.intro) +apply(rule ACf_inf) +apply(rule ACIf_axioms.intro) +apply(rule inf_idem) +done + +lemma (in Lattice) ACIf_sup: "ACIf sup" +apply(rule ACIf.intro) +apply(rule ACf_sup) +apply(rule ACIf_axioms.intro) +apply(rule sup_idem) +done + +lemma (in Lattice) ACIfSL_inf: "ACIfSL inf (op \)" +apply(rule ACIfSL.intro) +apply(rule ACf_inf) +apply(rule ACIf.axioms[OF ACIf_inf]) +apply(rule ACIfSL_axioms.intro) +apply(rule iffI) + apply(blast intro: antisym inf_le1 inf_le2 inf_least refl) +apply(erule subst) +apply(rule inf_le2) +done + +lemma (in Lattice) ACIfSL_sup: "ACIfSL sup (%x y. y \ x)" +apply(rule ACIfSL.intro) +apply(rule ACf_sup) +apply(rule ACIf.axioms[OF ACIf_sup]) +apply(rule ACIfSL_axioms.intro) +apply(rule iffI) + apply(blast intro: antisym sup_ge1 sup_ge2 sup_greatest refl) +apply(erule subst) +apply(rule sup_ge2) +done + +text{* Fold laws in lattices *} + +lemma (in Lattice) Inf_le_Sup: "\ finite A; A \ {} \ \ \A \ \A" +apply(unfold Sup_def Inf_def) +apply(subgoal_tac "EX a. a:A") +prefer 2 apply blast +apply(erule exE) +apply(rule trans) +apply(erule (2) ACIfSL.fold1_belowI[OF ACIfSL_inf]) +apply(erule (2) ACIfSL.fold1_belowI[OF ACIfSL_sup]) +done + +lemma (in Lattice) sup_Inf1_distrib: +assumes A: "finite A" "A \ {}" +shows "(x \ \A) = \{x \ a|a. a \ A}" +using A +proof (induct rule: finite_ne_induct) + case singleton thus ?case by(simp add:Inf_def) +next + case (insert y A) + have fin: "finite {x \ a |a. a \ A}" + by(fast intro: finite_surj[where f = "%a. x \ a", OF insert(0)]) + have "x \ \ (insert y A) = x \ (y \ \ A)" + using insert by(simp add:ACf.fold1_insert_def[OF ACf_inf Inf_def]) + also have "\ = (x \ y) \ (x \ \ A)" by(rule sup_inf_distrib1) + also have "x \ \ A = \{x \ a|a. a \ A}" using insert by simp + also have "(x \ y) \ \ = \ (insert (x \ y) {x \ a |a. a \ A})" + using insert by(simp add:ACIf.fold1_insert2_def[OF ACIf_inf Inf_def fin]) + also have "insert (x\y) {x\a |a. a \ A} = {x\a |a. a \ insert y A}" + by blast + finally show ?case . +qed + + +lemma (in Lattice) sup_Inf2_distrib: +assumes A: "finite A" "A \ {}" and B: "finite B" "B \ {}" +shows "(\A \ \B) = \{a \ b|a b. a \ A \ b \ B}" +using A +proof (induct rule: finite_ne_induct) + case singleton thus ?case + by(simp add: sup_Inf1_distrib[OF B] fold1_singleton_def[OF Inf_def]) +next + case (insert x A) + have finB: "finite {x \ b |b. b \ B}" + by(fast intro: finite_surj[where f = "%b. x \ b", OF B(0)]) + have finAB: "finite {a \ b |a b. a \ A \ b \ B}" + proof - + have "{a \ b |a b. a \ A \ b \ B} = (UN a:A. UN b:B. {a \ b})" + by blast + thus ?thesis by(simp add: insert(0) B(0)) + qed + have ne: "{a \ b |a b. a \ A \ b \ B} \ {}" using insert B by blast + have "\(insert x A) \ \B = (x \ \A) \ \B" + using insert by(simp add:ACIf.fold1_insert2_def[OF ACIf_inf Inf_def]) + also have "\ = (x \ \B) \ (\A \ \B)" by(rule sup_inf_distrib2) + also have "\ = \{x \ b|b. b \ B} \ \{a \ b|a b. a \ A \ b \ B}" + using insert by(simp add:sup_Inf1_distrib[OF B]) + also have "\ = \({x\b |b. b \ B} \ {a\b |a b. a \ A \ b \ B})" + (is "_ = \?M") + using B insert + by(simp add:Inf_def ACIf.fold1_Un2[OF ACIf_inf finB _ finAB ne]) + also have "?M = {a \ b |a b. a \ insert x A \ b \ B}" + by blast + finally show ?case . +qed + + subsection{*Min and Max*} @@ -2141,6 +2291,50 @@ apply(auto simp:max_def) done +lemma Lattice_min_max: "Lattice (op \) (min :: 'a::linorder \ 'a \ 'a) max" +apply(rule Lattice.intro) +apply simp +apply(erule (1) order_trans) +apply(erule (1) order_antisym) +apply(simp add:min_def max_def linorder_not_le order_less_imp_le) +apply(simp add:min_def max_def linorder_not_le order_less_imp_le) +apply(simp add:min_def max_def linorder_not_le order_less_imp_le) +apply(simp add:min_def max_def linorder_not_le order_less_imp_le) +apply(simp add:min_def max_def linorder_not_le order_less_imp_le) +apply(simp add:min_def max_def linorder_not_le order_less_imp_le) +apply(rule_tac x=x and y=y in linorder_le_cases) +apply(rule_tac x=x and y=z in linorder_le_cases) +apply(rule_tac x=y and y=z in linorder_le_cases) +apply(simp add:min_def max_def) +apply(simp add:min_def max_def) +apply(rule_tac x=y and y=z in linorder_le_cases) +apply(simp add:min_def max_def) +apply(simp add:min_def max_def) +apply(rule_tac x=x and y=z in linorder_le_cases) +apply(rule_tac x=y and y=z in linorder_le_cases) +apply(simp add:min_def max_def) +apply(simp add:min_def max_def) +apply(rule_tac x=y and y=z in linorder_le_cases) +apply(simp add:min_def max_def) +apply(simp add:min_def max_def) + +apply(rule_tac x=x and y=y in linorder_le_cases) +apply(rule_tac x=x and y=z in linorder_le_cases) +apply(rule_tac x=y and y=z in linorder_le_cases) +apply(simp add:min_def max_def) +apply(simp add:min_def max_def) +apply(rule_tac x=y and y=z in linorder_le_cases) +apply(simp add:min_def max_def) +apply(simp add:min_def max_def) +apply(rule_tac x=x and y=z in linorder_le_cases) +apply(rule_tac x=y and y=z in linorder_le_cases) +apply(simp add:min_def max_def) +apply(simp add:min_def max_def) +apply(rule_tac x=y and y=z in linorder_le_cases) +apply(simp add:min_def max_def) +apply(simp add:min_def max_def) +done + text{* Now we instantiate the recursion equations and declare them simplification rules: *} @@ -2184,4 +2378,13 @@ "\ finite A; A \ {} \ \ (x \ Max A) = (\a\A. x \ a)" by(simp add: Max_def ACIfSLlin.fold1_below_iff[OF ACIfSLlin_max]) +lemma Min_le_Max: + "\ finite A; A \ {} \ \ Min A \ Max A" +by(simp add: Min_def Max_def Lattice.Inf_le_Sup[OF Lattice_min_max]) + +lemma max_Min2_distrib: + "\ finite A; A \ {}; finite B; B \ {} \ \ + max (Min A) (Min B) = Min{ max a b |a b. a \ A \ b \ B}" +by(simp add: Min_def Max_def Lattice.sup_Inf2_distrib[OF Lattice_min_max]) + end