diff -r 250e1da3204b -r af5ef0d4d655 src/HOL/Lattices.thy --- a/src/HOL/Lattices.thy Tue Oct 16 23:12:38 2007 +0200 +++ b/src/HOL/Lattices.thy Tue Oct 16 23:12:45 2007 +0200 @@ -33,18 +33,20 @@ lemmas antisym_intro [intro!] = antisym lemmas (in -) [rule del] = antisym_intro -lemma le_infI1[intro]: "a \ x \ a \ b \ x" -apply(subgoal_tac "a \ b \ a") - apply(blast intro: order_trans) -apply simp -done +lemma le_infI1[intro]: + assumes "a \ x" + shows "a \ b \ x" +proof (rule order_trans) + show "a \ b \ a" and "a \ x" using assms by simp +qed lemmas (in -) [rule del] = le_infI1 -lemma le_infI2[intro]: "b \ x \ a \ b \ x" -apply(subgoal_tac "a \ b \ b") - apply(blast intro: order_trans) -apply simp -done +lemma le_infI2[intro]: + assumes "b \ x" + shows "a \ b \ x" +proof (rule order_trans) + show "a \ b \ b" and "b \ x" using assms by simp +qed lemmas (in -) [rule del] = le_infI2 lemma le_infI[intro!]: "x \ a \ x \ b \ x \ a \ b" @@ -75,17 +77,11 @@ lemmas (in -) [rule del] = antisym_intro lemma le_supI1[intro]: "x \ a \ x \ a \ b" -apply(subgoal_tac "a \ a \ b") - apply(blast intro: order_trans) -apply simp -done + by (rule order_trans) auto lemmas (in -) [rule del] = le_supI1 lemma le_supI2[intro]: "x \ b \ x \ a \ b" -apply(subgoal_tac "b \ a \ b") - apply(blast intro: order_trans) -apply simp -done + by (rule order_trans) auto lemmas (in -) [rule del] = le_supI2 lemma le_supI[intro!]: "a \ x \ b \ x \ a \ b \ x" @@ -255,26 +251,26 @@ lemma (in lower_semilattice) inf_unique: fixes f (infixl "\" 70) - assumes le1: "\x y. x \ y \<^loc>\ x" and le2: "\x y. x \ y \<^loc>\ y" - and greatest: "\x y z. x \<^loc>\ y \ x \<^loc>\ z \ x \<^loc>\ y \ z" + assumes le1: "\x y. x \ y \ x" and le2: "\x y. x \ y \ y" + and greatest: "\x y z. x \ y \ x \ z \ x \ y \ z" shows "x \ y = x \ y" proof (rule antisym) - show "x \ y \<^loc>\ x \ y" by (rule le_infI) (rule le1, rule le2) + show "x \ y \ x \ y" by (rule le_infI) (rule le1, rule le2) next - have leI: "\x y z. x \<^loc>\ y \ x \<^loc>\ z \ x \<^loc>\ y \ z" by (blast intro: greatest) - show "x \ y \<^loc>\ x \ y" by (rule leI) simp_all + have leI: "\x y z. x \ y \ x \ z \ x \ y \ z" by (blast intro: greatest) + show "x \ y \ x \ y" by (rule leI) simp_all qed lemma (in upper_semilattice) sup_unique: fixes f (infixl "\" 70) - assumes ge1 [simp]: "\x y. x \<^loc>\ x \ y" and ge2: "\x y. y \<^loc>\ x \ y" - and least: "\x y z. y \<^loc>\ x \ z \<^loc>\ x \ y \ z \<^loc>\ x" + assumes ge1 [simp]: "\x y. x \ x \ y" and ge2: "\x y. y \ x \ y" + and least: "\x y z. y \ x \ z \ x \ y \ z \ x" shows "x \ y = x \ y" proof (rule antisym) - show "x \ y \<^loc>\ x \ y" by (rule le_supI) (rule ge1, rule ge2) + show "x \ y \ x \ y" by (rule le_supI) (rule ge1, rule ge2) next - have leI: "\x y z. x \<^loc>\ z \ y \<^loc>\ z \ x \ y \<^loc>\ z" by (blast intro: least) - show "x \ y \<^loc>\ x \ y" by (rule leI) simp_all + have leI: "\x y z. x \ z \ y \ z \ x \ y \ z" by (blast intro: least) + show "x \ y \ x \ y" by (rule leI) simp_all qed @@ -282,9 +278,9 @@ special case of @{const inf}/@{const sup} *} lemma (in linorder) distrib_lattice_min_max: - "distrib_lattice (op \<^loc>\) (op \<^loc><) min max" + "distrib_lattice (op \) (op <) min max" proof unfold_locales - have aux: "\x y \ 'a. x \<^loc>< y \ y \<^loc>\ x \ x = y" + have aux: "\x y \ 'a. x < y \ y \ x \ x = y" by (auto simp add: less_le antisym) fix x y z show "max x (min y z) = min (max x y) (max x z)" @@ -333,10 +329,10 @@ and Sup_least: "(\x. x \ A \ x \ z) \ \A \ z" begin -lemma Inf_Sup: "\A = \{b. \a \ A. b \<^loc>\ a}" +lemma Inf_Sup: "\A = \{b. \a \ A. b \ a}" by (auto intro: Inf_lower Inf_greatest Sup_upper Sup_least) -lemma Sup_Inf: "\A = \{b. \a \ A. a \<^loc>\ b}" +lemma Sup_Inf: "\A = \{b. \a \ A. a \ b}" by (auto intro: Inf_lower Inf_greatest Sup_upper Sup_least) lemma Inf_Univ: "\UNIV = \{}" @@ -411,10 +407,10 @@ where "bot = Sup {}" -lemma top_greatest [simp]: "x \<^loc>\ top" +lemma top_greatest [simp]: "x \ top" by (unfold top_def, rule Inf_greatest, simp) -lemma bot_least [simp]: "bot \<^loc>\ x" +lemma bot_least [simp]: "bot \ x" by (unfold bot_def, rule Sup_least, simp) definition @@ -584,4 +580,16 @@ lemmas inf_aci = inf_ACI lemmas sup_aci = sup_ACI +no_notation + inf (infixl "\" 70) + +no_notation + sup (infixl "\" 65) + +no_notation + Inf ("\_" [900] 900) + +no_notation + Sup ("\_" [900] 900) + end