# HG changeset patch # User haftmann # Date 1162904586 -3600 # Node ID 1c8580913738fb2e0a69dc51551dceee18f76480 # Parent 7c9337a0e30a4a616f7849f4f407b6889708471e made locale partial_order compatible with axclass order; changed import order; consecutive changes diff -r 7c9337a0e30a -r 1c8580913738 src/HOL/Lattice_Locales.thy --- a/src/HOL/Lattice_Locales.thy Tue Nov 07 14:03:04 2006 +0100 +++ b/src/HOL/Lattice_Locales.thy Tue Nov 07 14:03:06 2006 +0100 @@ -6,7 +6,7 @@ header {* Lattices via Locales *} theory Lattice_Locales -imports HOL +imports Orderings begin subsection{* Lattices *} @@ -16,12 +16,6 @@ Finite_Set}. In the longer term it may be better to define arbitrary sups and infs via @{text THE}. *} -locale partial_order = - fixes below :: "'a \ 'a \ bool" (infixl "\" 50) - assumes refl[iff]: "x \ x" - and trans: "x \ y \ y \ z \ x \ z" - and antisym: "x \ y \ y \ x \ x = y" - locale lower_semilattice = partial_order + fixes inf :: "'a \ 'a \ 'a" (infixl "\" 70) assumes inf_le1: "x \ y \ x" and inf_le2: "x \ y \ y" @@ -71,10 +65,13 @@ by(blast intro: antisym sup_ge2 sup_greatest refl) -lemma (in lower_semilattice) below_inf_conv[simp]: +lemma (in lower_semilattice) less_eq_inf_conv [simp]: "x \ y \ z = (x \ y \ x \ z)" by(blast intro: antisym inf_le1 inf_le2 inf_least refl trans) +lemmas (in lower_semilattice) below_inf_conv = less_eq_inf_conv + -- {* a duplicate for backward compatibility *} + lemma (in upper_semilattice) above_sup_conv[simp]: "x \ y \ z = (x \ z \ y \ z)" by(blast intro: antisym sup_ge1 sup_ge2 sup_greatest refl trans) @@ -169,4 +166,171 @@ sup_inf_distrib1 sup_inf_distrib2 inf_sup_distrib1 inf_sup_distrib2 +subsection {* Least value operator and min/max -- properties *} + +(*FIXME: derive more of the min/max laws generically via semilattices*) + +lemma LeastI2_order: + "[| P (x::'a::order); + !!y. P y ==> x <= y; + !!x. [| P x; ALL y. P y --> x \ y |] ==> Q x |] + ==> Q (Least P)" + apply (unfold Least_def) + apply (rule theI2) + apply (blast intro: order_antisym)+ + done + +lemma Least_equality: + "[| P (k::'a::order); !!x. P x ==> k <= x |] ==> (LEAST x. P x) = k" + apply (simp add: Least_def) + apply (rule the_equality) + apply (auto intro!: order_antisym) + done + +lemma min_leastL: "(!!x. least <= x) ==> min least x = least" + by (simp add: min_def) + +lemma max_leastL: "(!!x. least <= x) ==> max least x = x" + by (simp add: max_def) + +lemma min_leastR: "(\x\'a\order. least \ x) \ min x least = least" + apply (simp add: min_def) + apply (blast intro: order_antisym) + done + +lemma max_leastR: "(\x\'a\order. least \ x) \ max x least = x" + apply (simp add: max_def) + apply (blast intro: order_antisym) + done + +lemma min_of_mono: + "(!!x y. (f x <= f y) = (x <= y)) ==> min (f m) (f n) = f (min m n)" + by (simp add: min_def) + +lemma max_of_mono: + "(!!x y. (f x <= f y) = (x <= y)) ==> max (f m) (f n) = f (max m n)" + by (simp add: max_def) + +text{* Instantiate locales: *} + +interpretation min_max: + lower_semilattice["op \" "op <" "min :: 'a::linorder \ 'a \ 'a"] +apply unfold_locales +apply(simp add:min_def linorder_not_le order_less_imp_le) +apply(simp add:min_def linorder_not_le order_less_imp_le) +apply(simp add:min_def linorder_not_le order_less_imp_le) +done + +interpretation min_max: + upper_semilattice["op \" "op <" "max :: 'a::linorder \ 'a \ 'a"] +apply unfold_locales +apply(simp add: max_def linorder_not_le order_less_imp_le) +apply(simp add: max_def linorder_not_le order_less_imp_le) +apply(simp add: max_def linorder_not_le order_less_imp_le) +done + +interpretation min_max: + lattice["op \" "op <" "min :: 'a::linorder \ 'a \ 'a" "max"] + by unfold_locales + +interpretation min_max: + distrib_lattice["op \" "op <" "min :: 'a::linorder \ 'a \ 'a" "max"] +apply unfold_locales +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 + +lemma le_max_iff_disj: "!!z::'a::linorder. (z <= max x y) = (z <= x | z <= y)" + apply(simp add:max_def) + apply (insert linorder_linear) + apply (blast intro: order_trans) + done + +lemmas le_maxI1 = min_max.sup_ge1 +lemmas le_maxI2 = min_max.sup_ge2 + +lemma less_max_iff_disj: "!!z::'a::linorder. (z < max x y) = (z < x | z < y)" + apply (simp add: max_def order_le_less) + apply (insert linorder_less_linear) + apply (blast intro: order_less_trans) + done + +lemma max_less_iff_conj [simp]: + "!!z::'a::linorder. (max x y < z) = (x < z & y < z)" + apply (simp add: order_le_less max_def) + apply (insert linorder_less_linear) + apply (blast intro: order_less_trans) + done + +lemma min_less_iff_conj [simp]: + "!!z::'a::linorder. (z < min x y) = (z < x & z < y)" + apply (simp add: order_le_less min_def) + apply (insert linorder_less_linear) + apply (blast intro: order_less_trans) + done + +lemma min_le_iff_disj: "!!z::'a::linorder. (min x y <= z) = (x <= z | y <= z)" + apply (simp add: min_def) + apply (insert linorder_linear) + apply (blast intro: order_trans) + done + +lemma min_less_iff_disj: "!!z::'a::linorder. (min x y < z) = (x < z | y < z)" + apply (simp add: min_def order_le_less) + apply (insert linorder_less_linear) + apply (blast intro: order_less_trans) + done + +lemmas max_ac = min_max.sup_assoc min_max.sup_commute + mk_left_commute[of max,OF min_max.sup_assoc min_max.sup_commute] + +lemmas min_ac = min_max.inf_assoc min_max.inf_commute + mk_left_commute[of min,OF min_max.inf_assoc min_max.inf_commute] + +lemma split_min: + "P (min (i::'a::linorder) j) = ((i <= j --> P(i)) & (~ i <= j --> P(j)))" + by (simp add: min_def) + +lemma split_max: + "P (max (i::'a::linorder) j) = ((i <= j --> P(j)) & (~ i <= j --> P(i)))" + by (simp add: max_def) + +text {* ML legacy bindings *} + +ML {* +val Least_def = thm "Least_def"; +val Least_equality = thm "Least_equality"; +val min_def = thm "min_def"; +val min_of_mono = thm "min_of_mono"; +val max_def = thm "max_def"; +val max_of_mono = thm "max_of_mono"; +val min_leastL = thm "min_leastL"; +val max_leastL = thm "max_leastL"; +val min_leastR = thm "min_leastR"; +val max_leastR = thm "max_leastR"; +val le_max_iff_disj = thm "le_max_iff_disj"; +val le_maxI1 = thm "le_maxI1"; +val le_maxI2 = thm "le_maxI2"; +val less_max_iff_disj = thm "less_max_iff_disj"; +val max_less_iff_conj = thm "max_less_iff_conj"; +val min_less_iff_conj = thm "min_less_iff_conj"; +val min_le_iff_disj = thm "min_le_iff_disj"; +val min_less_iff_disj = thm "min_less_iff_disj"; +val split_min = thm "split_min"; +val split_max = thm "split_max"; +*} + end diff -r 7c9337a0e30a -r 1c8580913738 src/HOL/Orderings.ML --- a/src/HOL/Orderings.ML Tue Nov 07 14:03:04 2006 +0100 +++ b/src/HOL/Orderings.ML Tue Nov 07 14:03:06 2006 +0100 @@ -1,18 +1,5 @@ (* legacy ML bindings *) -val Least_def = thm "Least_def"; -val Least_equality = thm "Least_equality"; -val mono_def = thm "mono_def"; -val monoI = thm "monoI"; -val monoD = thm "monoD"; -val min_def = thm "min_def"; -val min_of_mono = thm "min_of_mono"; -val max_def = thm "max_def"; -val max_of_mono = thm "max_of_mono"; -val min_leastL = thm "min_leastL"; -val max_leastL = thm "max_leastL"; -val min_leastR = thm "min_leastR"; -val max_leastR = thm "max_leastR"; val order_eq_refl = thm "order_eq_refl"; val order_less_irrefl = thm "order_less_irrefl"; val order_le_less = thm "order_le_less"; @@ -33,22 +20,11 @@ val linorder_not_le = thm "linorder_not_le"; val linorder_neq_iff = thm "linorder_neq_iff"; val linorder_neqE = thm "linorder_neqE"; -(* -val min_same = thm "min_same"; -val max_same = thm "max_same"; -*) -val le_max_iff_disj = thm "le_max_iff_disj"; -val le_maxI1 = thm "le_maxI1"; -val le_maxI2 = thm "le_maxI2"; -val less_max_iff_disj = thm "less_max_iff_disj"; -val max_less_iff_conj = thm "max_less_iff_conj"; -val min_less_iff_conj = thm "min_less_iff_conj"; -val min_le_iff_disj = thm "min_le_iff_disj"; -val min_less_iff_disj = thm "min_less_iff_disj"; -val split_min = thm "split_min"; -val split_max = thm "split_max"; val order_refl = thm "order_refl"; val order_trans = thm "order_trans"; val order_antisym = thm "order_antisym"; val order_less_le = thm "order_less_le"; val linorder_linear = thm "linorder_linear"; +val mono_def = thm "mono_def"; +val monoI = thm "monoI"; +val monoD = thm "monoD"; diff -r 7c9337a0e30a -r 1c8580913738 src/HOL/Orderings.thy --- a/src/HOL/Orderings.thy Tue Nov 07 14:03:04 2006 +0100 +++ b/src/HOL/Orderings.thy Tue Nov 07 14:03:06 2006 +0100 @@ -6,7 +6,7 @@ header {* Abstract orderings *} theory Orderings -imports Code_Generator Lattice_Locales +imports Code_Generator begin section {* Abstract orderings *} @@ -69,23 +69,29 @@ subsection {* Partial orderings *} +locale partial_order = + fixes below :: "'a \ 'a \ bool" (infixl "\" 50) + fixes less :: "'a \ 'a \ bool" (infixl "\" 50) + assumes refl [iff]: "x \ x" + and trans: "x \ y \ y \ z \ x \ z" + and antisym: "x \ y \ y \ x \ x = y" + and less_le: "(x \ y) = (x \ y \ x \ y)" + axclass order < ord order_refl [iff]: "x <= x" order_trans: "x <= y ==> y <= z ==> x <= z" order_antisym: "x <= y ==> y <= x ==> x = y" order_less_le: "(x < y) = (x <= y & x ~= y)" -text {* Connection to locale: *} - interpretation order: - partial_order["op \ :: 'a::order \ 'a \ bool"] + partial_order ["op \ \ 'a\order \ 'a \ bool" "op < \ 'a\order \ 'a \ bool"] apply(rule partial_order.intro) -apply(rule order_refl, erule (1) order_trans, erule (1) order_antisym) +apply(rule order_refl, erule (1) order_trans, erule (1) order_antisym, rule order_less_le) done text {* Reflexivity. *} -lemma order_eq_refl: "!!x::'a::order. x = y ==> x <= y" +lemma order_eq_refl: "(x \ 'a\order) = y \ x \ y" -- {* This form is useful with the classical reasoner. *} apply (erule ssubst) apply (rule order_refl) @@ -169,11 +175,18 @@ by (rule order_less_asym) -subsection {* Total orderings *} +subsection {* Linear (total) orderings *} + +locale linear_order = partial_order + + assumes linear: "x \ y \ y \ x" axclass linorder < order linorder_linear: "x <= y | y <= x" +interpretation linorder: + linear_order ["op \ \ 'a\linorder \ 'a \ bool" "op < \ 'a\linorder \ 'a \ bool"] + by unfold_locales (rule linorder_linear) + lemma linorder_less_linear: "!!x::'a::linorder. x B \ f A \ f B" + +lemmas monoI [intro?] = mono.intro + and monoD [dest?] = mono.mono constdefs Least :: "('a::ord => bool) => 'a" (binder "LEAST " 10) "Least P == THE x. P x & (ALL y. P y --> x <= y)" -- {* We can no longer use LeastM because the latter requires Hilbert-AC. *} -lemma LeastI2_order: - "[| P (x::'a::order); - !!y. P y ==> x <= y; - !!x. [| P x; ALL y. P y --> x \ y |] ==> Q x |] - ==> Q (Least P)" - apply (unfold Least_def) - apply (rule theI2) - apply (blast intro: order_antisym)+ - done - -lemma Least_equality: - "[| P (k::'a::order); !!x. P x ==> k <= x |] ==> (LEAST x. P x) = k" - apply (simp add: Least_def) - apply (rule the_equality) - apply (auto intro!: order_antisym) - done - -locale mono = - fixes f - assumes mono: "A <= B ==> f A <= f B" - -lemmas monoI [intro?] = mono.intro - and monoD [dest?] = mono.mono - constdefs min :: "['a::ord, 'a] => 'a" "min a b == (if a <= b then a else b)" max :: "['a::ord, 'a] => 'a" "max a b == (if a <= b then b else a)" -lemma min_leastR: "(\x\'a\order. least \ x) \ min x least = least" - apply (simp add: min_def) - apply (blast intro: order_antisym) - done - -lemma max_leastR: "(\x\'a\order. least \ x) \ max x least = x" - apply (simp add: max_def) - apply (blast intro: order_antisym) - done - -lemma min_leastL: "(!!x. least <= x) ==> min least x = least" - by (simp add: min_def) - -lemma max_leastL: "(!!x. least <= x) ==> max least x = x" - by (simp add: max_def) - -lemma min_of_mono: - "(!!x y. (f x <= f y) = (x <= y)) ==> min (f m) (f n) = f (min m n)" - by (simp add: min_def) - -lemma max_of_mono: - "(!!x y. (f x <= f y) = (x <= y)) ==> max (f m) (f n) = f (max m n)" - by (simp add: max_def) - -text{* Instantiate locales: *} - -interpretation min_max: - lower_semilattice["op \" "min :: 'a::linorder \ 'a \ 'a"] -apply unfold_locales -apply(simp add:min_def linorder_not_le order_less_imp_le) -apply(simp add:min_def linorder_not_le order_less_imp_le) -apply(simp add:min_def linorder_not_le order_less_imp_le) -done - -interpretation min_max: - upper_semilattice["op \" "max :: 'a::linorder \ 'a \ 'a"] -apply unfold_locales -apply(simp add: max_def linorder_not_le order_less_imp_le) -apply(simp add: max_def linorder_not_le order_less_imp_le) -apply(simp add: max_def linorder_not_le order_less_imp_le) -done - -interpretation min_max: - lattice["op \" "min :: 'a::linorder \ 'a \ 'a" "max"] - by unfold_locales - -interpretation min_max: - distrib_lattice["op \" "min :: 'a::linorder \ 'a \ 'a" "max"] -apply unfold_locales -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 - -lemma le_max_iff_disj: "!!z::'a::linorder. (z <= max x y) = (z <= x | z <= y)" - apply(simp add:max_def) - apply (insert linorder_linear) - apply (blast intro: order_trans) - done - -lemmas le_maxI1 = min_max.sup_ge1 -lemmas le_maxI2 = min_max.sup_ge2 - -lemma less_max_iff_disj: "!!z::'a::linorder. (z < max x y) = (z < x | z < y)" - apply (simp add: max_def order_le_less) - apply (insert linorder_less_linear) - apply (blast intro: order_less_trans) - done - -lemma max_less_iff_conj [simp]: - "!!z::'a::linorder. (max x y < z) = (x < z & y < z)" - apply (simp add: order_le_less max_def) - apply (insert linorder_less_linear) - apply (blast intro: order_less_trans) - done - -lemma min_less_iff_conj [simp]: - "!!z::'a::linorder. (z < min x y) = (z < x & z < y)" - apply (simp add: order_le_less min_def) - apply (insert linorder_less_linear) - apply (blast intro: order_less_trans) - done - -lemma min_le_iff_disj: "!!z::'a::linorder. (min x y <= z) = (x <= z | y <= z)" - apply (simp add: min_def) - apply (insert linorder_linear) - apply (blast intro: order_trans) - done - -lemma min_less_iff_disj: "!!z::'a::linorder. (min x y < z) = (x < z | y < z)" - apply (simp add: min_def order_le_less) - apply (insert linorder_less_linear) - apply (blast intro: order_less_trans) - done - -lemmas max_ac = min_max.sup_assoc min_max.sup_commute - mk_left_commute[of max,OF min_max.sup_assoc min_max.sup_commute] - -lemmas min_ac = min_max.inf_assoc min_max.inf_commute - mk_left_commute[of min,OF min_max.inf_assoc min_max.inf_commute] - -lemma split_min: - "P (min (i::'a::linorder) j) = ((i <= j --> P(i)) & (~ i <= j --> P(j)))" - by (simp add: min_def) - -lemma split_max: - "P (max (i::'a::linorder) j) = ((i <= j --> P(j)) & (~ i <= j --> P(i)))" - by (simp add: max_def) - end