diff --git a/src/HOL/Lattices.thy b/src/HOL/Lattices.thy --- a/src/HOL/Lattices.thy +++ b/src/HOL/Lattices.thy @@ -1,977 +1,986 @@ (* Title: HOL/Lattices.thy Author: Tobias Nipkow *) section \Abstract lattices\ theory Lattices imports Groups begin subsection \Abstract semilattice\ text \ These locales provide a basic structure for interpretation into bigger structures; extensions require careful thinking, otherwise undesired effects may occur due to interpretation. \ locale semilattice = abel_semigroup + assumes idem [simp]: "a \<^bold>* a = a" begin lemma left_idem [simp]: "a \<^bold>* (a \<^bold>* b) = a \<^bold>* b" by (simp add: assoc [symmetric]) lemma right_idem [simp]: "(a \<^bold>* b) \<^bold>* b = a \<^bold>* b" by (simp add: assoc) end locale semilattice_neutr = semilattice + comm_monoid locale semilattice_order = semilattice + fixes less_eq :: "'a \ 'a \ bool" (infix "\<^bold>\" 50) and less :: "'a \ 'a \ bool" (infix "\<^bold><" 50) assumes order_iff: "a \<^bold>\ b \ a = a \<^bold>* b" and strict_order_iff: "a \<^bold>< b \ a = a \<^bold>* b \ a \ b" begin lemma orderI: "a = a \<^bold>* b \ a \<^bold>\ b" by (simp add: order_iff) lemma orderE: assumes "a \<^bold>\ b" obtains "a = a \<^bold>* b" using assms by (unfold order_iff) sublocale ordering less_eq less proof show "a \<^bold>< b \ a \<^bold>\ b \ a \ b" for a b by (simp add: order_iff strict_order_iff) next show "a \<^bold>\ a" for a by (simp add: order_iff) next fix a b assume "a \<^bold>\ b" "b \<^bold>\ a" then have "a = a \<^bold>* b" "a \<^bold>* b = b" by (simp_all add: order_iff commute) then show "a = b" by simp next fix a b c assume "a \<^bold>\ b" "b \<^bold>\ c" then have "a = a \<^bold>* b" "b = b \<^bold>* c" by (simp_all add: order_iff commute) then have "a = a \<^bold>* (b \<^bold>* c)" by simp then have "a = (a \<^bold>* b) \<^bold>* c" by (simp add: assoc) with \a = a \<^bold>* b\ [symmetric] have "a = a \<^bold>* c" by simp then show "a \<^bold>\ c" by (rule orderI) qed lemma cobounded1 [simp]: "a \<^bold>* b \<^bold>\ a" by (simp add: order_iff commute) lemma cobounded2 [simp]: "a \<^bold>* b \<^bold>\ b" by (simp add: order_iff) lemma boundedI: assumes "a \<^bold>\ b" and "a \<^bold>\ c" shows "a \<^bold>\ b \<^bold>* c" proof (rule orderI) from assms obtain "a \<^bold>* b = a" and "a \<^bold>* c = a" by (auto elim!: orderE) then show "a = a \<^bold>* (b \<^bold>* c)" by (simp add: assoc [symmetric]) qed lemma boundedE: assumes "a \<^bold>\ b \<^bold>* c" obtains "a \<^bold>\ b" and "a \<^bold>\ c" using assms by (blast intro: trans cobounded1 cobounded2) lemma bounded_iff [simp]: "a \<^bold>\ b \<^bold>* c \ a \<^bold>\ b \ a \<^bold>\ c" by (blast intro: boundedI elim: boundedE) lemma strict_boundedE: assumes "a \<^bold>< b \<^bold>* c" obtains "a \<^bold>< b" and "a \<^bold>< c" using assms by (auto simp add: commute strict_iff_order elim: orderE intro!: that)+ lemma coboundedI1: "a \<^bold>\ c \ a \<^bold>* b \<^bold>\ c" by (rule trans) auto lemma coboundedI2: "b \<^bold>\ c \ a \<^bold>* b \<^bold>\ c" by (rule trans) auto lemma strict_coboundedI1: "a \<^bold>< c \ a \<^bold>* b \<^bold>< c" using irrefl by (auto intro: not_eq_order_implies_strict coboundedI1 strict_implies_order elim: strict_boundedE) lemma strict_coboundedI2: "b \<^bold>< c \ a \<^bold>* b \<^bold>< c" using strict_coboundedI1 [of b c a] by (simp add: commute) lemma mono: "a \<^bold>\ c \ b \<^bold>\ d \ a \<^bold>* b \<^bold>\ c \<^bold>* d" by (blast intro: boundedI coboundedI1 coboundedI2) lemma absorb1: "a \<^bold>\ b \ a \<^bold>* b = a" by (rule antisym) (auto simp: refl) lemma absorb2: "b \<^bold>\ a \ a \<^bold>* b = b" by (rule antisym) (auto simp: refl) lemma absorb_iff1: "a \<^bold>\ b \ a \<^bold>* b = a" using order_iff by auto lemma absorb_iff2: "b \<^bold>\ a \ a \<^bold>* b = b" using order_iff by (auto simp add: commute) end locale semilattice_neutr_order = semilattice_neutr + semilattice_order begin sublocale ordering_top less_eq less "\<^bold>1" by standard (simp add: order_iff) +lemma eq_neutr_iff [simp]: \a \<^bold>* b = \<^bold>1 \ a = \<^bold>1 \ b = \<^bold>1\ + by (simp add: eq_iff) + +lemma neutr_eq_iff [simp]: \\<^bold>1 = a \<^bold>* b \ a = \<^bold>1 \ b = \<^bold>1\ + by (simp add: eq_iff) + end text \Passive interpretations for boolean operators\ lemma semilattice_neutr_and: "semilattice_neutr HOL.conj True" by standard auto lemma semilattice_neutr_or: "semilattice_neutr HOL.disj False" by standard auto subsection \Syntactic infimum and supremum operations\ class inf = fixes inf :: "'a \ 'a \ 'a" (infixl "\" 70) class sup = fixes sup :: "'a \ 'a \ 'a" (infixl "\" 65) subsection \Concrete lattices\ class semilattice_inf = order + inf + assumes inf_le1 [simp]: "x \ y \ x" and inf_le2 [simp]: "x \ y \ y" and inf_greatest: "x \ y \ x \ z \ x \ y \ z" class semilattice_sup = order + sup + assumes sup_ge1 [simp]: "x \ x \ y" and sup_ge2 [simp]: "y \ x \ y" and sup_least: "y \ x \ z \ x \ y \ z \ x" begin text \Dual lattice.\ lemma dual_semilattice: "class.semilattice_inf sup greater_eq greater" by (rule class.semilattice_inf.intro, rule dual_order) (unfold_locales, simp_all add: sup_least) end class lattice = semilattice_inf + semilattice_sup subsubsection \Intro and elim rules\ context semilattice_inf begin lemma le_infI1: "a \ x \ a \ b \ x" by (rule order_trans) auto lemma le_infI2: "b \ x \ a \ b \ x" by (rule order_trans) auto lemma le_infI: "x \ a \ x \ b \ x \ a \ b" by (fact inf_greatest) (* FIXME: duplicate lemma *) lemma le_infE: "x \ a \ b \ (x \ a \ x \ b \ P) \ P" by (blast intro: order_trans inf_le1 inf_le2) lemma le_inf_iff: "x \ y \ z \ x \ y \ x \ z" by (blast intro: le_infI elim: le_infE) lemma le_iff_inf: "x \ y \ x \ y = x" by (auto intro: le_infI1 antisym dest: eq_iff [THEN iffD1] simp add: le_inf_iff) lemma inf_mono: "a \ c \ b \ d \ a \ b \ c \ d" by (fast intro: inf_greatest le_infI1 le_infI2) lemma mono_inf: "mono f \ f (A \ B) \ f A \ f B" for f :: "'a \ 'b::semilattice_inf" by (auto simp add: mono_def intro: Lattices.inf_greatest) end context semilattice_sup begin lemma le_supI1: "x \ a \ x \ a \ b" by (rule order_trans) auto lemma le_supI2: "x \ b \ x \ a \ b" by (rule order_trans) auto lemma le_supI: "a \ x \ b \ x \ a \ b \ x" by (fact sup_least) (* FIXME: duplicate lemma *) lemma le_supE: "a \ b \ x \ (a \ x \ b \ x \ P) \ P" by (blast intro: order_trans sup_ge1 sup_ge2) lemma le_sup_iff: "x \ y \ z \ x \ z \ y \ z" by (blast intro: le_supI elim: le_supE) lemma le_iff_sup: "x \ y \ x \ y = y" by (auto intro: le_supI2 antisym dest: eq_iff [THEN iffD1] simp add: le_sup_iff) lemma sup_mono: "a \ c \ b \ d \ a \ b \ c \ d" by (fast intro: sup_least le_supI1 le_supI2) lemma mono_sup: "mono f \ f A \ f B \ f (A \ B)" for f :: "'a \ 'b::semilattice_sup" by (auto simp add: mono_def intro: Lattices.sup_least) end subsubsection \Equational laws\ context semilattice_inf begin sublocale inf: semilattice inf proof fix a b c show "(a \ b) \ c = a \ (b \ c)" by (rule antisym) (auto intro: le_infI1 le_infI2 simp add: le_inf_iff) show "a \ b = b \ a" by (rule antisym) (auto simp add: le_inf_iff) show "a \ a = a" by (rule antisym) (auto simp add: le_inf_iff) qed sublocale inf: semilattice_order inf less_eq less by standard (auto simp add: le_iff_inf less_le) lemma inf_assoc: "(x \ y) \ z = x \ (y \ z)" by (fact inf.assoc) lemma inf_commute: "(x \ y) = (y \ x)" by (fact inf.commute) lemma inf_left_commute: "x \ (y \ z) = y \ (x \ z)" by (fact inf.left_commute) lemma inf_idem: "x \ x = x" by (fact inf.idem) (* already simp *) lemma inf_left_idem: "x \ (x \ y) = x \ y" by (fact inf.left_idem) (* already simp *) lemma inf_right_idem: "(x \ y) \ y = x \ y" by (fact inf.right_idem) (* already simp *) lemma inf_absorb1: "x \ y \ x \ y = x" by (rule antisym) auto lemma inf_absorb2: "y \ x \ x \ y = y" by (rule antisym) auto lemmas inf_aci = inf_commute inf_assoc inf_left_commute inf_left_idem end context semilattice_sup begin sublocale sup: semilattice sup proof fix a b c show "(a \ b) \ c = a \ (b \ c)" by (rule antisym) (auto intro: le_supI1 le_supI2 simp add: le_sup_iff) show "a \ b = b \ a" by (rule antisym) (auto simp add: le_sup_iff) show "a \ a = a" by (rule antisym) (auto simp add: le_sup_iff) qed sublocale sup: semilattice_order sup greater_eq greater by standard (auto simp add: le_iff_sup sup.commute less_le) lemma sup_assoc: "(x \ y) \ z = x \ (y \ z)" by (fact sup.assoc) lemma sup_commute: "(x \ y) = (y \ x)" by (fact sup.commute) lemma sup_left_commute: "x \ (y \ z) = y \ (x \ z)" by (fact sup.left_commute) lemma sup_idem: "x \ x = x" by (fact sup.idem) (* already simp *) lemma sup_left_idem [simp]: "x \ (x \ y) = x \ y" by (fact sup.left_idem) lemma sup_absorb1: "y \ x \ x \ y = x" by (rule antisym) auto lemma sup_absorb2: "x \ y \ x \ y = y" by (rule antisym) auto lemmas sup_aci = sup_commute sup_assoc sup_left_commute sup_left_idem end context lattice begin lemma dual_lattice: "class.lattice sup (\) (>) inf" by (rule class.lattice.intro, rule dual_semilattice, rule class.semilattice_sup.intro, rule dual_order) (unfold_locales, auto) lemma inf_sup_absorb [simp]: "x \ (x \ y) = x" by (blast intro: antisym inf_le1 inf_greatest sup_ge1) lemma sup_inf_absorb [simp]: "x \ (x \ y) = x" by (blast intro: antisym sup_ge1 sup_least inf_le1) lemmas inf_sup_aci = inf_aci sup_aci lemmas inf_sup_ord = inf_le1 inf_le2 sup_ge1 sup_ge2 text \Towards distributivity.\ lemma distrib_sup_le: "x \ (y \ z) \ (x \ y) \ (x \ z)" by (auto intro: le_infI1 le_infI2 le_supI1 le_supI2) lemma distrib_inf_le: "(x \ y) \ (x \ z) \ x \ (y \ z)" by (auto intro: le_infI1 le_infI2 le_supI1 le_supI2) text \If you have one of them, you have them all.\ lemma distrib_imp1: assumes distrib: "\x y z. x \ (y \ z) = (x \ y) \ (x \ z)" shows "x \ (y \ z) = (x \ y) \ (x \ z)" proof- have "x \ (y \ z) = (x \ (x \ z)) \ (y \ z)" by simp also have "\ = x \ (z \ (x \ y))" by (simp add: distrib inf_commute sup_assoc del: sup_inf_absorb) also have "\ = ((x \ y) \ x) \ ((x \ y) \ z)" by (simp add: inf_commute) also have "\ = (x \ y) \ (x \ z)" by(simp add:distrib) finally show ?thesis . qed lemma distrib_imp2: assumes distrib: "\x y z. x \ (y \ z) = (x \ y) \ (x \ z)" shows "x \ (y \ z) = (x \ y) \ (x \ z)" proof- have "x \ (y \ z) = (x \ (x \ z)) \ (y \ z)" by simp also have "\ = x \ (z \ (x \ y))" by (simp add: distrib sup_commute inf_assoc del: inf_sup_absorb) also have "\ = ((x \ y) \ x) \ ((x \ y) \ z)" by (simp add: sup_commute) also have "\ = (x \ y) \ (x \ z)" by (simp add:distrib) finally show ?thesis . qed end subsubsection \Strict order\ context semilattice_inf begin lemma less_infI1: "a < x \ a \ b < x" by (auto simp add: less_le inf_absorb1 intro: le_infI1) lemma less_infI2: "b < x \ a \ b < x" by (auto simp add: less_le inf_absorb2 intro: le_infI2) end context semilattice_sup begin lemma less_supI1: "x < a \ x < a \ b" using dual_semilattice by (rule semilattice_inf.less_infI1) lemma less_supI2: "x < b \ x < a \ b" using dual_semilattice by (rule semilattice_inf.less_infI2) end subsection \Distributive lattices\ class distrib_lattice = lattice + assumes sup_inf_distrib1: "x \ (y \ z) = (x \ y) \ (x \ z)" context distrib_lattice begin lemma sup_inf_distrib2: "(y \ z) \ x = (y \ x) \ (z \ x)" by (simp add: sup_commute sup_inf_distrib1) lemma inf_sup_distrib1: "x \ (y \ z) = (x \ y) \ (x \ z)" by (rule distrib_imp2 [OF sup_inf_distrib1]) lemma inf_sup_distrib2: "(y \ z) \ x = (y \ x) \ (z \ x)" by (simp add: inf_commute inf_sup_distrib1) lemma dual_distrib_lattice: "class.distrib_lattice sup (\) (>) inf" by (rule class.distrib_lattice.intro, rule dual_lattice) (unfold_locales, fact inf_sup_distrib1) lemmas sup_inf_distrib = sup_inf_distrib1 sup_inf_distrib2 lemmas inf_sup_distrib = inf_sup_distrib1 inf_sup_distrib2 lemmas distrib = sup_inf_distrib1 sup_inf_distrib2 inf_sup_distrib1 inf_sup_distrib2 end subsection \Bounded lattices and boolean algebras\ class bounded_semilattice_inf_top = semilattice_inf + order_top begin sublocale inf_top: semilattice_neutr inf top + inf_top: semilattice_neutr_order inf top less_eq less proof show "x \ \ = x" for x by (rule inf_absorb1) simp qed +lemma inf_top_left: "\ \ x = x" + by (fact inf_top.left_neutral) + +lemma inf_top_right: "x \ \ = x" + by (fact inf_top.right_neutral) + +lemma inf_eq_top_iff: "x \ y = \ \ x = \ \ y = \" + by (fact inf_top.eq_neutr_iff) + +lemma top_eq_inf_iff: "\ = x \ y \ x = \ \ y = \" + by (fact inf_top.neutr_eq_iff) + end class bounded_semilattice_sup_bot = semilattice_sup + order_bot begin sublocale sup_bot: semilattice_neutr sup bot + sup_bot: semilattice_neutr_order sup bot greater_eq greater proof show "x \ \ = x" for x by (rule sup_absorb1) simp qed +lemma sup_bot_left: "\ \ x = x" + by (fact sup_bot.left_neutral) + +lemma sup_bot_right: "x \ \ = x" + by (fact sup_bot.right_neutral) + +lemma sup_eq_bot_iff: "x \ y = \ \ x = \ \ y = \" + by (fact sup_bot.eq_neutr_iff) + +lemma bot_eq_sup_iff: "\ = x \ y \ x = \ \ y = \" + by (fact sup_bot.neutr_eq_iff) + end class bounded_lattice_bot = lattice + order_bot begin subclass bounded_semilattice_sup_bot .. lemma inf_bot_left [simp]: "\ \ x = \" by (rule inf_absorb1) simp lemma inf_bot_right [simp]: "x \ \ = \" by (rule inf_absorb2) simp -lemma sup_bot_left: "\ \ x = x" - by (fact sup_bot.left_neutral) - -lemma sup_bot_right: "x \ \ = x" - by (fact sup_bot.right_neutral) - -lemma sup_eq_bot_iff [simp]: "x \ y = \ \ x = \ \ y = \" - by (simp add: eq_iff) - -lemma bot_eq_sup_iff [simp]: "\ = x \ y \ x = \ \ y = \" - by (simp add: eq_iff) - end class bounded_lattice_top = lattice + order_top begin subclass bounded_semilattice_inf_top .. lemma sup_top_left [simp]: "\ \ x = \" by (rule sup_absorb1) simp lemma sup_top_right [simp]: "x \ \ = \" by (rule sup_absorb2) simp -lemma inf_top_left: "\ \ x = x" - by (fact inf_top.left_neutral) - -lemma inf_top_right: "x \ \ = x" - by (fact inf_top.right_neutral) - -lemma inf_eq_top_iff [simp]: "x \ y = \ \ x = \ \ y = \" - by (simp add: eq_iff) - end class bounded_lattice = lattice + order_bot + order_top begin subclass bounded_lattice_bot .. subclass bounded_lattice_top .. lemma dual_bounded_lattice: "class.bounded_lattice sup greater_eq greater inf \ \" by unfold_locales (auto simp add: less_le_not_le) end class boolean_algebra = distrib_lattice + bounded_lattice + minus + uminus + assumes inf_compl_bot: "x \ - x = \" and sup_compl_top: "x \ - x = \" assumes diff_eq: "x - y = x \ - y" begin lemma dual_boolean_algebra: "class.boolean_algebra (\x y. x \ - y) uminus sup greater_eq greater inf \ \" by (rule class.boolean_algebra.intro, rule dual_bounded_lattice, rule dual_distrib_lattice) (unfold_locales, auto simp add: inf_compl_bot sup_compl_top diff_eq) lemma compl_inf_bot [simp]: "- x \ x = \" by (simp add: inf_commute inf_compl_bot) lemma compl_sup_top [simp]: "- x \ x = \" by (simp add: sup_commute sup_compl_top) lemma compl_unique: assumes "x \ y = \" and "x \ y = \" shows "- x = y" proof - have "(x \ - x) \ (- x \ y) = (x \ y) \ (- x \ y)" using inf_compl_bot assms(1) by simp then have "(- x \ x) \ (- x \ y) = (y \ x) \ (y \ - x)" by (simp add: inf_commute) then have "- x \ (x \ y) = y \ (x \ - x)" by (simp add: inf_sup_distrib1) then have "- x \ \ = y \ \" using sup_compl_top assms(2) by simp then show "- x = y" by simp qed lemma double_compl [simp]: "- (- x) = x" using compl_inf_bot compl_sup_top by (rule compl_unique) lemma compl_eq_compl_iff [simp]: "- x = - y \ x = y" proof assume "- x = - y" then have "- (- x) = - (- y)" by (rule arg_cong) then show "x = y" by simp next assume "x = y" then show "- x = - y" by simp qed lemma compl_bot_eq [simp]: "- \ = \" proof - from sup_compl_top have "\ \ - \ = \" . then show ?thesis by simp qed lemma compl_top_eq [simp]: "- \ = \" proof - from inf_compl_bot have "\ \ - \ = \" . then show ?thesis by simp qed lemma compl_inf [simp]: "- (x \ y) = - x \ - y" proof (rule compl_unique) have "(x \ y) \ (- x \ - y) = (y \ (x \ - x)) \ (x \ (y \ - y))" by (simp only: inf_sup_distrib inf_aci) then show "(x \ y) \ (- x \ - y) = \" by (simp add: inf_compl_bot) next have "(x \ y) \ (- x \ - y) = (- y \ (x \ - x)) \ (- x \ (y \ - y))" by (simp only: sup_inf_distrib sup_aci) then show "(x \ y) \ (- x \ - y) = \" by (simp add: sup_compl_top) qed lemma compl_sup [simp]: "- (x \ y) = - x \ - y" using dual_boolean_algebra by (rule boolean_algebra.compl_inf) lemma compl_mono: assumes "x \ y" shows "- y \ - x" proof - from assms have "x \ y = y" by (simp only: le_iff_sup) then have "- (x \ y) = - y" by simp then have "- x \ - y = - y" by simp then have "- y \ - x = - y" by (simp only: inf_commute) then show ?thesis by (simp only: le_iff_inf) qed lemma compl_le_compl_iff [simp]: "- x \ - y \ y \ x" by (auto dest: compl_mono) lemma compl_le_swap1: assumes "y \ - x" shows "x \ -y" proof - from assms have "- (- x) \ - y" by (simp only: compl_le_compl_iff) then show ?thesis by simp qed lemma compl_le_swap2: assumes "- y \ x" shows "- x \ y" proof - from assms have "- x \ - (- y)" by (simp only: compl_le_compl_iff) then show ?thesis by simp qed lemma compl_less_compl_iff: "- x < - y \ y < x" (* TODO: declare [simp] ? *) by (auto simp add: less_le) lemma compl_less_swap1: assumes "y < - x" shows "x < - y" proof - from assms have "- (- x) < - y" by (simp only: compl_less_compl_iff) then show ?thesis by simp qed lemma compl_less_swap2: assumes "- y < x" shows "- x < y" proof - from assms have "- x < - (- y)" by (simp only: compl_less_compl_iff) then show ?thesis by simp qed lemma sup_cancel_left1: "sup (sup x a) (sup (- x) b) = top" by (simp add: inf_sup_aci sup_compl_top) lemma sup_cancel_left2: "sup (sup (- x) a) (sup x b) = top" by (simp add: inf_sup_aci sup_compl_top) lemma inf_cancel_left1: "inf (inf x a) (inf (- x) b) = bot" by (simp add: inf_sup_aci inf_compl_bot) lemma inf_cancel_left2: "inf (inf (- x) a) (inf x b) = bot" by (simp add: inf_sup_aci inf_compl_bot) declare inf_compl_bot [simp] and sup_compl_top [simp] lemma sup_compl_top_left1 [simp]: "sup (- x) (sup x y) = top" by (simp add: sup_assoc[symmetric]) lemma sup_compl_top_left2 [simp]: "sup x (sup (- x) y) = top" using sup_compl_top_left1[of "- x" y] by simp lemma inf_compl_bot_left1 [simp]: "inf (- x) (inf x y) = bot" by (simp add: inf_assoc[symmetric]) lemma inf_compl_bot_left2 [simp]: "inf x (inf (- x) y) = bot" using inf_compl_bot_left1[of "- x" y] by simp lemma inf_compl_bot_right [simp]: "inf x (inf y (- x)) = bot" by (subst inf_left_commute) simp end locale boolean_algebra_cancel begin lemma sup1: "(A::'a::semilattice_sup) \ sup k a \ sup A b \ sup k (sup a b)" by (simp only: ac_simps) lemma sup2: "(B::'a::semilattice_sup) \ sup k b \ sup a B \ sup k (sup a b)" by (simp only: ac_simps) lemma sup0: "(a::'a::bounded_semilattice_sup_bot) \ sup a bot" by simp lemma inf1: "(A::'a::semilattice_inf) \ inf k a \ inf A b \ inf k (inf a b)" by (simp only: ac_simps) lemma inf2: "(B::'a::semilattice_inf) \ inf k b \ inf a B \ inf k (inf a b)" by (simp only: ac_simps) lemma inf0: "(a::'a::bounded_semilattice_inf_top) \ inf a top" by simp end ML_file \Tools/boolean_algebra_cancel.ML\ simproc_setup boolean_algebra_cancel_sup ("sup a b::'a::boolean_algebra") = \fn phi => fn ss => try Boolean_Algebra_Cancel.cancel_sup_conv\ simproc_setup boolean_algebra_cancel_inf ("inf a b::'a::boolean_algebra") = \fn phi => fn ss => try Boolean_Algebra_Cancel.cancel_inf_conv\ subsection \\min/max\ as special case of lattice\ context linorder begin sublocale min: semilattice_order min less_eq less + max: semilattice_order max greater_eq greater by standard (auto simp add: min_def max_def) lemma min_le_iff_disj: "min x y \ z \ x \ z \ y \ z" unfolding min_def using linear by (auto intro: order_trans) lemma le_max_iff_disj: "z \ max x y \ z \ x \ z \ y" unfolding max_def using linear by (auto intro: order_trans) lemma min_less_iff_disj: "min x y < z \ x < z \ y < z" unfolding min_def le_less using less_linear by (auto intro: less_trans) lemma less_max_iff_disj: "z < max x y \ z < x \ z < y" unfolding max_def le_less using less_linear by (auto intro: less_trans) lemma min_less_iff_conj [simp]: "z < min x y \ z < x \ z < y" unfolding min_def le_less using less_linear by (auto intro: less_trans) lemma max_less_iff_conj [simp]: "max x y < z \ x < z \ y < z" unfolding max_def le_less using less_linear by (auto intro: less_trans) lemma min_max_distrib1: "min (max b c) a = max (min b a) (min c a)" by (auto simp add: min_def max_def not_le dest: le_less_trans less_trans intro: antisym) lemma min_max_distrib2: "min a (max b c) = max (min a b) (min a c)" by (auto simp add: min_def max_def not_le dest: le_less_trans less_trans intro: antisym) lemma max_min_distrib1: "max (min b c) a = min (max b a) (max c a)" by (auto simp add: min_def max_def not_le dest: le_less_trans less_trans intro: antisym) lemma max_min_distrib2: "max a (min b c) = min (max a b) (max a c)" by (auto simp add: min_def max_def not_le dest: le_less_trans less_trans intro: antisym) lemmas min_max_distribs = min_max_distrib1 min_max_distrib2 max_min_distrib1 max_min_distrib2 lemma split_min [no_atp]: "P (min i j) \ (i \ j \ P i) \ (\ i \ j \ P j)" by (simp add: min_def) lemma split_max [no_atp]: "P (max i j) \ (i \ j \ P j) \ (\ i \ j \ P i)" by (simp add: max_def) lemma split_min_lin [no_atp]: \P (min a b) \ (b = a \ P a) \ (a < b \ P a) \ (b < a \ P b)\ by (cases a b rule: linorder_cases) (auto simp add: min.absorb1 min.absorb2) lemma split_max_lin [no_atp]: \P (max a b) \ (b = a \ P a) \ (a < b \ P b) \ (b < a \ P a)\ by (cases a b rule: linorder_cases) (auto simp add: max.absorb1 max.absorb2) lemma min_of_mono: "mono f \ min (f m) (f n) = f (min m n)" for f :: "'a \ 'b::linorder" by (auto simp: mono_def Orderings.min_def min_def intro: Orderings.antisym) lemma max_of_mono: "mono f \ max (f m) (f n) = f (max m n)" for f :: "'a \ 'b::linorder" by (auto simp: mono_def Orderings.max_def max_def intro: Orderings.antisym) end lemma max_of_antimono: "antimono f \ max (f x) (f y) = f (min x y)" and min_of_antimono: "antimono f \ min (f x) (f y) = f (max x y)" for f::"'a::linorder \ 'b::linorder" by (auto simp: antimono_def Orderings.max_def min_def intro!: antisym) lemma inf_min: "inf = (min :: 'a::{semilattice_inf,linorder} \ 'a \ 'a)" by (auto intro: antisym simp add: min_def fun_eq_iff) lemma sup_max: "sup = (max :: 'a::{semilattice_sup,linorder} \ 'a \ 'a)" by (auto intro: antisym simp add: max_def fun_eq_iff) subsection \Uniqueness of inf and sup\ lemma (in semilattice_inf) inf_unique: fixes f (infixl "\" 70) 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 \ x \ y" by (rule le_infI) (rule le1, rule le2) 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 semilattice_sup) sup_unique: fixes f (infixl "\" 70) 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 \ x \ y" by (rule le_supI) (rule ge1, rule ge2) 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 subsection \Lattice on \<^typ>\bool\\ instantiation bool :: boolean_algebra begin definition bool_Compl_def [simp]: "uminus = Not" definition bool_diff_def [simp]: "A - B \ A \ \ B" definition [simp]: "P \ Q \ P \ Q" definition [simp]: "P \ Q \ P \ Q" instance by standard auto end lemma sup_boolI1: "P \ P \ Q" by simp lemma sup_boolI2: "Q \ P \ Q" by simp lemma sup_boolE: "P \ Q \ (P \ R) \ (Q \ R) \ R" by auto subsection \Lattice on \<^typ>\_ \ _\\ instantiation "fun" :: (type, semilattice_sup) semilattice_sup begin definition "f \ g = (\x. f x \ g x)" lemma sup_apply [simp, code]: "(f \ g) x = f x \ g x" by (simp add: sup_fun_def) instance by standard (simp_all add: le_fun_def) end instantiation "fun" :: (type, semilattice_inf) semilattice_inf begin definition "f \ g = (\x. f x \ g x)" lemma inf_apply [simp, code]: "(f \ g) x = f x \ g x" by (simp add: inf_fun_def) instance by standard (simp_all add: le_fun_def) end instance "fun" :: (type, lattice) lattice .. instance "fun" :: (type, distrib_lattice) distrib_lattice by standard (rule ext, simp add: sup_inf_distrib1) instance "fun" :: (type, bounded_lattice) bounded_lattice .. instantiation "fun" :: (type, uminus) uminus begin definition fun_Compl_def: "- A = (\x. - A x)" lemma uminus_apply [simp, code]: "(- A) x = - (A x)" by (simp add: fun_Compl_def) instance .. end instantiation "fun" :: (type, minus) minus begin definition fun_diff_def: "A - B = (\x. A x - B x)" lemma minus_apply [simp, code]: "(A - B) x = A x - B x" by (simp add: fun_diff_def) instance .. end instance "fun" :: (type, boolean_algebra) boolean_algebra by standard (rule ext, simp_all add: inf_compl_bot sup_compl_top diff_eq)+ subsection \Lattice on unary and binary predicates\ lemma inf1I: "A x \ B x \ (A \ B) x" by (simp add: inf_fun_def) lemma inf2I: "A x y \ B x y \ (A \ B) x y" by (simp add: inf_fun_def) lemma inf1E: "(A \ B) x \ (A x \ B x \ P) \ P" by (simp add: inf_fun_def) lemma inf2E: "(A \ B) x y \ (A x y \ B x y \ P) \ P" by (simp add: inf_fun_def) lemma inf1D1: "(A \ B) x \ A x" by (rule inf1E) lemma inf2D1: "(A \ B) x y \ A x y" by (rule inf2E) lemma inf1D2: "(A \ B) x \ B x" by (rule inf1E) lemma inf2D2: "(A \ B) x y \ B x y" by (rule inf2E) lemma sup1I1: "A x \ (A \ B) x" by (simp add: sup_fun_def) lemma sup2I1: "A x y \ (A \ B) x y" by (simp add: sup_fun_def) lemma sup1I2: "B x \ (A \ B) x" by (simp add: sup_fun_def) lemma sup2I2: "B x y \ (A \ B) x y" by (simp add: sup_fun_def) lemma sup1E: "(A \ B) x \ (A x \ P) \ (B x \ P) \ P" by (simp add: sup_fun_def) iprover lemma sup2E: "(A \ B) x y \ (A x y \ P) \ (B x y \ P) \ P" by (simp add: sup_fun_def) iprover text \ \<^medskip> Classical introduction rule: no commitment to \A\ vs \B\.\ lemma sup1CI: "(\ B x \ A x) \ (A \ B) x" by (auto simp add: sup_fun_def) lemma sup2CI: "(\ B x y \ A x y) \ (A \ B) x y" by (auto simp add: sup_fun_def) end diff --git a/src/HOL/Nat.thy b/src/HOL/Nat.thy --- a/src/HOL/Nat.thy +++ b/src/HOL/Nat.thy @@ -1,2545 +1,2551 @@ (* Title: HOL/Nat.thy Author: Tobias Nipkow Author: Lawrence C Paulson Author: Markus Wenzel *) section \Natural numbers\ theory Nat imports Inductive Typedef Fun Rings begin subsection \Type \ind\\ typedecl ind axiomatization Zero_Rep :: ind and Suc_Rep :: "ind \ ind" \ \The axiom of infinity in 2 parts:\ where Suc_Rep_inject: "Suc_Rep x = Suc_Rep y \ x = y" and Suc_Rep_not_Zero_Rep: "Suc_Rep x \ Zero_Rep" subsection \Type nat\ text \Type definition\ inductive Nat :: "ind \ bool" where Zero_RepI: "Nat Zero_Rep" | Suc_RepI: "Nat i \ Nat (Suc_Rep i)" typedef nat = "{n. Nat n}" morphisms Rep_Nat Abs_Nat using Nat.Zero_RepI by auto lemma Nat_Rep_Nat: "Nat (Rep_Nat n)" using Rep_Nat by simp lemma Nat_Abs_Nat_inverse: "Nat n \ Rep_Nat (Abs_Nat n) = n" using Abs_Nat_inverse by simp lemma Nat_Abs_Nat_inject: "Nat n \ Nat m \ Abs_Nat n = Abs_Nat m \ n = m" using Abs_Nat_inject by simp instantiation nat :: zero begin definition Zero_nat_def: "0 = Abs_Nat Zero_Rep" instance .. end definition Suc :: "nat \ nat" where "Suc n = Abs_Nat (Suc_Rep (Rep_Nat n))" lemma Suc_not_Zero: "Suc m \ 0" by (simp add: Zero_nat_def Suc_def Suc_RepI Zero_RepI Nat_Abs_Nat_inject Suc_Rep_not_Zero_Rep Nat_Rep_Nat) lemma Zero_not_Suc: "0 \ Suc m" by (rule not_sym) (rule Suc_not_Zero) lemma Suc_Rep_inject': "Suc_Rep x = Suc_Rep y \ x = y" by (rule iffI, rule Suc_Rep_inject) simp_all lemma nat_induct0: assumes "P 0" and "\n. P n \ P (Suc n)" shows "P n" proof - have "P (Abs_Nat (Rep_Nat n))" using assms unfolding Zero_nat_def Suc_def by (iprover intro: Nat_Rep_Nat [THEN Nat.induct] elim: Nat_Abs_Nat_inverse [THEN subst]) then show ?thesis by (simp add: Rep_Nat_inverse) qed free_constructors case_nat for "0 :: nat" | Suc pred where "pred (0 :: nat) = (0 :: nat)" apply atomize_elim apply (rename_tac n, induct_tac n rule: nat_induct0, auto) apply (simp add: Suc_def Nat_Abs_Nat_inject Nat_Rep_Nat Suc_RepI Suc_Rep_inject' Rep_Nat_inject) apply (simp only: Suc_not_Zero) done \ \Avoid name clashes by prefixing the output of \old_rep_datatype\ with \old\.\ setup \Sign.mandatory_path "old"\ old_rep_datatype "0 :: nat" Suc by (erule nat_induct0) auto setup \Sign.parent_path\ \ \But erase the prefix for properties that are not generated by \free_constructors\.\ setup \Sign.mandatory_path "nat"\ declare old.nat.inject[iff del] and old.nat.distinct(1)[simp del, induct_simp del] lemmas induct = old.nat.induct lemmas inducts = old.nat.inducts lemmas rec = old.nat.rec lemmas simps = nat.inject nat.distinct nat.case nat.rec setup \Sign.parent_path\ abbreviation rec_nat :: "'a \ (nat \ 'a \ 'a) \ nat \ 'a" where "rec_nat \ old.rec_nat" declare nat.sel[code del] hide_const (open) Nat.pred \ \hide everything related to the selector\ hide_fact nat.case_eq_if nat.collapse nat.expand nat.sel nat.exhaust_sel nat.split_sel nat.split_sel_asm lemma nat_exhaust [case_names 0 Suc, cases type: nat]: "(y = 0 \ P) \ (\nat. y = Suc nat \ P) \ P" \ \for backward compatibility -- names of variables differ\ by (rule old.nat.exhaust) lemma nat_induct [case_names 0 Suc, induct type: nat]: fixes n assumes "P 0" and "\n. P n \ P (Suc n)" shows "P n" \ \for backward compatibility -- names of variables differ\ using assms by (rule nat.induct) hide_fact nat_exhaust nat_induct0 ML \ val nat_basic_lfp_sugar = let val ctr_sugar = the (Ctr_Sugar.ctr_sugar_of_global \<^theory> \<^type_name>\nat\); val recx = Logic.varify_types_global \<^term>\rec_nat\; val C = body_type (fastype_of recx); in {T = HOLogic.natT, fp_res_index = 0, C = C, fun_arg_Tsss = [[], [[HOLogic.natT, C]]], ctr_sugar = ctr_sugar, recx = recx, rec_thms = @{thms nat.rec}} end; \ setup \ let fun basic_lfp_sugars_of _ [\<^typ>\nat\] _ _ ctxt = ([], [0], [nat_basic_lfp_sugar], [], [], [], TrueI (*dummy*), [], false, ctxt) | basic_lfp_sugars_of bs arg_Ts callers callssss ctxt = BNF_LFP_Rec_Sugar.default_basic_lfp_sugars_of bs arg_Ts callers callssss ctxt; in BNF_LFP_Rec_Sugar.register_lfp_rec_extension {nested_simps = [], special_endgame_tac = K (K (K (K no_tac))), is_new_datatype = K (K true), basic_lfp_sugars_of = basic_lfp_sugars_of, rewrite_nested_rec_call = NONE} end \ text \Injectiveness and distinctness lemmas\ lemma inj_Suc [simp]: "inj_on Suc N" by (simp add: inj_on_def) lemma bij_betw_Suc [simp]: "bij_betw Suc M N \ Suc ` M = N" by (simp add: bij_betw_def) lemma Suc_neq_Zero: "Suc m = 0 \ R" by (rule notE) (rule Suc_not_Zero) lemma Zero_neq_Suc: "0 = Suc m \ R" by (rule Suc_neq_Zero) (erule sym) lemma Suc_inject: "Suc x = Suc y \ x = y" by (rule inj_Suc [THEN injD]) lemma n_not_Suc_n: "n \ Suc n" by (induct n) simp_all lemma Suc_n_not_n: "Suc n \ n" by (rule not_sym) (rule n_not_Suc_n) text \A special form of induction for reasoning about \<^term>\m < n\ and \<^term>\m - n\.\ lemma diff_induct: assumes "\x. P x 0" and "\y. P 0 (Suc y)" and "\x y. P x y \ P (Suc x) (Suc y)" shows "P m n" proof (induct n arbitrary: m) case 0 show ?case by (rule assms(1)) next case (Suc n) show ?case proof (induct m) case 0 show ?case by (rule assms(2)) next case (Suc m) from \P m n\ show ?case by (rule assms(3)) qed qed subsection \Arithmetic operators\ instantiation nat :: comm_monoid_diff begin primrec plus_nat where add_0: "0 + n = (n::nat)" | add_Suc: "Suc m + n = Suc (m + n)" lemma add_0_right [simp]: "m + 0 = m" for m :: nat by (induct m) simp_all lemma add_Suc_right [simp]: "m + Suc n = Suc (m + n)" by (induct m) simp_all declare add_0 [code] lemma add_Suc_shift [code]: "Suc m + n = m + Suc n" by simp primrec minus_nat where diff_0 [code]: "m - 0 = (m::nat)" | diff_Suc: "m - Suc n = (case m - n of 0 \ 0 | Suc k \ k)" declare diff_Suc [simp del] lemma diff_0_eq_0 [simp, code]: "0 - n = 0" for n :: nat by (induct n) (simp_all add: diff_Suc) lemma diff_Suc_Suc [simp, code]: "Suc m - Suc n = m - n" by (induct n) (simp_all add: diff_Suc) instance proof fix n m q :: nat show "(n + m) + q = n + (m + q)" by (induct n) simp_all show "n + m = m + n" by (induct n) simp_all show "m + n - m = n" by (induct m) simp_all show "n - m - q = n - (m + q)" by (induct q) (simp_all add: diff_Suc) show "0 + n = n" by simp show "0 - n = 0" by simp qed end hide_fact (open) add_0 add_0_right diff_0 instantiation nat :: comm_semiring_1_cancel begin definition One_nat_def [simp]: "1 = Suc 0" primrec times_nat where mult_0: "0 * n = (0::nat)" | mult_Suc: "Suc m * n = n + (m * n)" lemma mult_0_right [simp]: "m * 0 = 0" for m :: nat by (induct m) simp_all lemma mult_Suc_right [simp]: "m * Suc n = m + (m * n)" by (induct m) (simp_all add: add.left_commute) lemma add_mult_distrib: "(m + n) * k = (m * k) + (n * k)" for m n k :: nat by (induct m) (simp_all add: add.assoc) instance proof fix k n m q :: nat show "0 \ (1::nat)" by simp show "1 * n = n" by simp show "n * m = m * n" by (induct n) simp_all show "(n * m) * q = n * (m * q)" by (induct n) (simp_all add: add_mult_distrib) show "(n + m) * q = n * q + m * q" by (rule add_mult_distrib) show "k * (m - n) = (k * m) - (k * n)" by (induct m n rule: diff_induct) simp_all qed end subsubsection \Addition\ text \Reasoning about \m + 0 = 0\, etc.\ lemma add_is_0 [iff]: "m + n = 0 \ m = 0 \ n = 0" for m n :: nat by (cases m) simp_all lemma add_is_1: "m + n = Suc 0 \ m = Suc 0 \ n = 0 \ m = 0 \ n = Suc 0" by (cases m) simp_all lemma one_is_add: "Suc 0 = m + n \ m = Suc 0 \ n = 0 \ m = 0 \ n = Suc 0" by (rule trans, rule eq_commute, rule add_is_1) lemma add_eq_self_zero: "m + n = m \ n = 0" for m n :: nat by (induct m) simp_all lemma plus_1_eq_Suc: "plus 1 = Suc" by (simp add: fun_eq_iff) lemma Suc_eq_plus1: "Suc n = n + 1" by simp lemma Suc_eq_plus1_left: "Suc n = 1 + n" by simp subsubsection \Difference\ lemma Suc_diff_diff [simp]: "(Suc m - n) - Suc k = m - n - k" by (simp add: diff_diff_add) lemma diff_Suc_1 [simp]: "Suc n - 1 = n" by simp subsubsection \Multiplication\ lemma mult_is_0 [simp]: "m * n = 0 \ m = 0 \ n = 0" for m n :: nat by (induct m) auto lemma mult_eq_1_iff [simp]: "m * n = Suc 0 \ m = Suc 0 \ n = Suc 0" proof (induct m) case 0 then show ?case by simp next case (Suc m) then show ?case by (induct n) auto qed lemma one_eq_mult_iff [simp]: "Suc 0 = m * n \ m = Suc 0 \ n = Suc 0" by (simp add: eq_commute flip: mult_eq_1_iff) lemma nat_mult_eq_1_iff [simp]: "m * n = 1 \ m = 1 \ n = 1" and nat_1_eq_mult_iff [simp]: "1 = m * n \ m = 1 \ n = 1" for m n :: nat by auto lemma mult_cancel1 [simp]: "k * m = k * n \ m = n \ k = 0" for k m n :: nat proof - have "k \ 0 \ k * m = k * n \ m = n" proof (induct n arbitrary: m) case 0 then show "m = 0" by simp next case (Suc n) then show "m = Suc n" by (cases m) (simp_all add: eq_commute [of 0]) qed then show ?thesis by auto qed lemma mult_cancel2 [simp]: "m * k = n * k \ m = n \ k = 0" for k m n :: nat by (simp add: mult.commute) lemma Suc_mult_cancel1: "Suc k * m = Suc k * n \ m = n" by (subst mult_cancel1) simp subsection \Orders on \<^typ>\nat\\ subsubsection \Operation definition\ instantiation nat :: linorder begin primrec less_eq_nat where "(0::nat) \ n \ True" | "Suc m \ n \ (case n of 0 \ False | Suc n \ m \ n)" declare less_eq_nat.simps [simp del] lemma le0 [iff]: "0 \ n" for n :: nat by (simp add: less_eq_nat.simps) lemma [code]: "0 \ n \ True" for n :: nat by simp definition less_nat where less_eq_Suc_le: "n < m \ Suc n \ m" lemma Suc_le_mono [iff]: "Suc n \ Suc m \ n \ m" by (simp add: less_eq_nat.simps(2)) lemma Suc_le_eq [code]: "Suc m \ n \ m < n" unfolding less_eq_Suc_le .. lemma le_0_eq [iff]: "n \ 0 \ n = 0" for n :: nat by (induct n) (simp_all add: less_eq_nat.simps(2)) lemma not_less0 [iff]: "\ n < 0" for n :: nat by (simp add: less_eq_Suc_le) lemma less_nat_zero_code [code]: "n < 0 \ False" for n :: nat by simp lemma Suc_less_eq [iff]: "Suc m < Suc n \ m < n" by (simp add: less_eq_Suc_le) lemma less_Suc_eq_le [code]: "m < Suc n \ m \ n" by (simp add: less_eq_Suc_le) lemma Suc_less_eq2: "Suc n < m \ (\m'. m = Suc m' \ n < m')" by (cases m) auto lemma le_SucI: "m \ n \ m \ Suc n" by (induct m arbitrary: n) (simp_all add: less_eq_nat.simps(2) split: nat.splits) lemma Suc_leD: "Suc m \ n \ m \ n" by (cases n) (auto intro: le_SucI) lemma less_SucI: "m < n \ m < Suc n" by (simp add: less_eq_Suc_le) (erule Suc_leD) lemma Suc_lessD: "Suc m < n \ m < n" by (simp add: less_eq_Suc_le) (erule Suc_leD) instance proof fix n m q :: nat show "n < m \ n \ m \ \ m \ n" proof (induct n arbitrary: m) case 0 then show ?case by (cases m) (simp_all add: less_eq_Suc_le) next case (Suc n) then show ?case by (cases m) (simp_all add: less_eq_Suc_le) qed show "n \ n" by (induct n) simp_all then show "n = m" if "n \ m" and "m \ n" using that by (induct n arbitrary: m) (simp_all add: less_eq_nat.simps(2) split: nat.splits) show "n \ q" if "n \ m" and "m \ q" using that proof (induct n arbitrary: m q) case 0 show ?case by simp next case (Suc n) then show ?case by (simp_all (no_asm_use) add: less_eq_nat.simps(2) split: nat.splits, clarify, simp_all (no_asm_use) add: less_eq_nat.simps(2) split: nat.splits, clarify, simp_all (no_asm_use) add: less_eq_nat.simps(2) split: nat.splits) qed show "n \ m \ m \ n" by (induct n arbitrary: m) (simp_all add: less_eq_nat.simps(2) split: nat.splits) qed end instantiation nat :: order_bot begin definition bot_nat :: nat where "bot_nat = 0" instance by standard (simp add: bot_nat_def) end instance nat :: no_top by standard (auto intro: less_Suc_eq_le [THEN iffD2]) subsubsection \Introduction properties\ lemma lessI [iff]: "n < Suc n" by (simp add: less_Suc_eq_le) lemma zero_less_Suc [iff]: "0 < Suc n" by (simp add: less_Suc_eq_le) subsubsection \Elimination properties\ lemma less_not_refl: "\ n < n" for n :: nat by (rule order_less_irrefl) lemma less_not_refl2: "n < m \ m \ n" for m n :: nat by (rule not_sym) (rule less_imp_neq) lemma less_not_refl3: "s < t \ s \ t" for s t :: nat by (rule less_imp_neq) lemma less_irrefl_nat: "n < n \ R" for n :: nat by (rule notE, rule less_not_refl) lemma less_zeroE: "n < 0 \ R" for n :: nat by (rule notE) (rule not_less0) lemma less_Suc_eq: "m < Suc n \ m < n \ m = n" unfolding less_Suc_eq_le le_less .. lemma less_Suc0 [iff]: "(n < Suc 0) = (n = 0)" by (simp add: less_Suc_eq) lemma less_one [iff]: "n < 1 \ n = 0" for n :: nat unfolding One_nat_def by (rule less_Suc0) lemma Suc_mono: "m < n \ Suc m < Suc n" by simp text \"Less than" is antisymmetric, sort of.\ lemma less_antisym: "\ n < m \ n < Suc m \ m = n" unfolding not_less less_Suc_eq_le by (rule antisym) lemma nat_neq_iff: "m \ n \ m < n \ n < m" for m n :: nat by (rule linorder_neq_iff) subsubsection \Inductive (?) properties\ lemma Suc_lessI: "m < n \ Suc m \ n \ Suc m < n" unfolding less_eq_Suc_le [of m] le_less by simp lemma lessE: assumes major: "i < k" and 1: "k = Suc i \ P" and 2: "\j. i < j \ k = Suc j \ P" shows P proof - from major have "\j. i \ j \ k = Suc j" unfolding less_eq_Suc_le by (induct k) simp_all then have "(\j. i < j \ k = Suc j) \ k = Suc i" by (auto simp add: less_le) with 1 2 show P by auto qed lemma less_SucE: assumes major: "m < Suc n" and less: "m < n \ P" and eq: "m = n \ P" shows P proof (rule major [THEN lessE]) show "Suc n = Suc m \ P" using eq by blast show "\j. \m < j; Suc n = Suc j\ \ P" by (blast intro: less) qed lemma Suc_lessE: assumes major: "Suc i < k" and minor: "\j. i < j \ k = Suc j \ P" shows P proof (rule major [THEN lessE]) show "k = Suc (Suc i) \ P" using lessI minor by iprover show "\j. \Suc i < j; k = Suc j\ \ P" using Suc_lessD minor by iprover qed lemma Suc_less_SucD: "Suc m < Suc n \ m < n" by simp lemma less_trans_Suc: assumes le: "i < j" shows "j < k \ Suc i < k" proof (induct k) case 0 then show ?case by simp next case (Suc k) with le show ?case by simp (auto simp add: less_Suc_eq dest: Suc_lessD) qed text \Can be used with \less_Suc_eq\ to get \<^prop>\n = m \ n < m\.\ lemma not_less_eq: "\ m < n \ n < Suc m" by (simp only: not_less less_Suc_eq_le) lemma not_less_eq_eq: "\ m \ n \ Suc n \ m" by (simp only: not_le Suc_le_eq) text \Properties of "less than or equal".\ lemma le_imp_less_Suc: "m \ n \ m < Suc n" by (simp only: less_Suc_eq_le) lemma Suc_n_not_le_n: "\ Suc n \ n" by (simp add: not_le less_Suc_eq_le) lemma le_Suc_eq: "m \ Suc n \ m \ n \ m = Suc n" by (simp add: less_Suc_eq_le [symmetric] less_Suc_eq) lemma le_SucE: "m \ Suc n \ (m \ n \ R) \ (m = Suc n \ R) \ R" by (drule le_Suc_eq [THEN iffD1], iprover+) lemma Suc_leI: "m < n \ Suc m \ n" by (simp only: Suc_le_eq) text \Stronger version of \Suc_leD\.\ lemma Suc_le_lessD: "Suc m \ n \ m < n" by (simp only: Suc_le_eq) lemma less_imp_le_nat: "m < n \ m \ n" for m n :: nat unfolding less_eq_Suc_le by (rule Suc_leD) text \For instance, \(Suc m < Suc n) = (Suc m \ n) = (m < n)\\ lemmas le_simps = less_imp_le_nat less_Suc_eq_le Suc_le_eq text \Equivalence of \m \ n\ and \m < n \ m = n\\ lemma less_or_eq_imp_le: "m < n \ m = n \ m \ n" for m n :: nat unfolding le_less . lemma le_eq_less_or_eq: "m \ n \ m < n \ m = n" for m n :: nat by (rule le_less) text \Useful with \blast\.\ lemma eq_imp_le: "m = n \ m \ n" for m n :: nat by auto lemma le_refl: "n \ n" for n :: nat by simp lemma le_trans: "i \ j \ j \ k \ i \ k" for i j k :: nat by (rule order_trans) lemma le_antisym: "m \ n \ n \ m \ m = n" for m n :: nat by (rule antisym) lemma nat_less_le: "m < n \ m \ n \ m \ n" for m n :: nat by (rule less_le) lemma le_neq_implies_less: "m \ n \ m \ n \ m < n" for m n :: nat unfolding less_le .. lemma nat_le_linear: "m \ n \ n \ m" for m n :: nat by (rule linear) lemmas linorder_neqE_nat = linorder_neqE [where 'a = nat] lemma le_less_Suc_eq: "m \ n \ n < Suc m \ n = m" unfolding less_Suc_eq_le by auto lemma not_less_less_Suc_eq: "\ n < m \ n < Suc m \ n = m" unfolding not_less by (rule le_less_Suc_eq) lemmas not_less_simps = not_less_less_Suc_eq le_less_Suc_eq lemma not0_implies_Suc: "n \ 0 \ \m. n = Suc m" by (cases n) simp_all lemma gr0_implies_Suc: "n > 0 \ \m. n = Suc m" by (cases n) simp_all lemma gr_implies_not0: "m < n \ n \ 0" for m n :: nat by (cases n) simp_all lemma neq0_conv[iff]: "n \ 0 \ 0 < n" for n :: nat by (cases n) simp_all text \This theorem is useful with \blast\\ lemma gr0I: "(n = 0 \ False) \ 0 < n" for n :: nat by (rule neq0_conv[THEN iffD1]) iprover lemma gr0_conv_Suc: "0 < n \ (\m. n = Suc m)" by (fast intro: not0_implies_Suc) lemma not_gr0 [iff]: "\ 0 < n \ n = 0" for n :: nat using neq0_conv by blast lemma Suc_le_D: "Suc n \ m' \ \m. m' = Suc m" by (induct m') simp_all text \Useful in certain inductive arguments\ lemma less_Suc_eq_0_disj: "m < Suc n \ m = 0 \ (\j. m = Suc j \ j < n)" by (cases m) simp_all lemma All_less_Suc: "(\i < Suc n. P i) = (P n \ (\i < n. P i))" by (auto simp: less_Suc_eq) lemma All_less_Suc2: "(\i < Suc n. P i) = (P 0 \ (\i < n. P(Suc i)))" by (auto simp: less_Suc_eq_0_disj) lemma Ex_less_Suc: "(\i < Suc n. P i) = (P n \ (\i < n. P i))" by (auto simp: less_Suc_eq) lemma Ex_less_Suc2: "(\i < Suc n. P i) = (P 0 \ (\i < n. P(Suc i)))" by (auto simp: less_Suc_eq_0_disj) text \@{term mono} (non-strict) doesn't imply increasing, as the function could be constant\ lemma strict_mono_imp_increasing: fixes n::nat assumes "strict_mono f" shows "f n \ n" proof (induction n) case 0 then show ?case by auto next case (Suc n) then show ?case unfolding not_less_eq_eq [symmetric] using Suc_n_not_le_n assms order_trans strict_mono_less_eq by blast qed subsubsection \Monotonicity of Addition\ lemma Suc_pred [simp]: "n > 0 \ Suc (n - Suc 0) = n" by (simp add: diff_Suc split: nat.split) lemma Suc_diff_1 [simp]: "0 < n \ Suc (n - 1) = n" unfolding One_nat_def by (rule Suc_pred) lemma nat_add_left_cancel_le [simp]: "k + m \ k + n \ m \ n" for k m n :: nat by (induct k) simp_all lemma nat_add_left_cancel_less [simp]: "k + m < k + n \ m < n" for k m n :: nat by (induct k) simp_all lemma add_gr_0 [iff]: "m + n > 0 \ m > 0 \ n > 0" for m n :: nat by (auto dest: gr0_implies_Suc) text \strict, in 1st argument\ lemma add_less_mono1: "i < j \ i + k < j + k" for i j k :: nat by (induct k) simp_all text \strict, in both arguments\ lemma add_less_mono: fixes i j k l :: nat assumes "i < j" "k < l" shows "i + k < j + l" proof - have "i + k < j + k" by (simp add: add_less_mono1 assms) also have "... < j + l" using \i < j\ by (induction j) (auto simp: assms) finally show ?thesis . qed lemma less_imp_Suc_add: "m < n \ \k. n = Suc (m + k)" proof (induct n) case 0 then show ?case by simp next case Suc then show ?case by (simp add: order_le_less) (blast elim!: less_SucE intro!: Nat.add_0_right [symmetric] add_Suc_right [symmetric]) qed lemma le_Suc_ex: "k \ l \ (\n. l = k + n)" for k l :: nat by (auto simp: less_Suc_eq_le[symmetric] dest: less_imp_Suc_add) lemma less_natE: assumes \m < n\ obtains q where \n = Suc (m + q)\ using assms by (auto dest: less_imp_Suc_add intro: that) text \strict, in 1st argument; proof is by induction on \k > 0\\ lemma mult_less_mono2: fixes i j :: nat assumes "i < j" and "0 < k" shows "k * i < k * j" using \0 < k\ proof (induct k) case 0 then show ?case by simp next case (Suc k) with \i < j\ show ?case by (cases k) (simp_all add: add_less_mono) qed text \Addition is the inverse of subtraction: if \<^term>\n \ m\ then \<^term>\n + (m - n) = m\.\ lemma add_diff_inverse_nat: "\ m < n \ n + (m - n) = m" for m n :: nat by (induct m n rule: diff_induct) simp_all lemma nat_le_iff_add: "m \ n \ (\k. n = m + k)" for m n :: nat using nat_add_left_cancel_le[of m 0] by (auto dest: le_Suc_ex) text \The naturals form an ordered \semidom\ and a \dioid\.\ instance nat :: linordered_semidom proof fix m n q :: nat show "0 < (1::nat)" by simp show "m \ n \ q + m \ q + n" by simp show "m < n \ 0 < q \ q * m < q * n" by (simp add: mult_less_mono2) show "m \ 0 \ n \ 0 \ m * n \ 0" by simp show "n \ m \ (m - n) + n = m" by (simp add: add_diff_inverse_nat add.commute linorder_not_less) qed instance nat :: dioid by standard (rule nat_le_iff_add) declare le0[simp del] \ \This is now @{thm zero_le}\ declare le_0_eq[simp del] \ \This is now @{thm le_zero_eq}\ declare not_less0[simp del] \ \This is now @{thm not_less_zero}\ declare not_gr0[simp del] \ \This is now @{thm not_gr_zero}\ instance nat :: ordered_cancel_comm_monoid_add .. instance nat :: ordered_cancel_comm_monoid_diff .. subsubsection \\<^term>\min\ and \<^term>\max\\ +global_interpretation bot_nat_0: ordering_top \(\)\ \(>)\ \0::nat\ + by standard simp + +global_interpretation max_nat: semilattice_neutr_order max \0::nat\ \(\)\ \(>)\ + by standard (simp add: max_def) + lemma mono_Suc: "mono Suc" by (rule monoI) simp lemma min_0L [simp]: "min 0 n = 0" for n :: nat by (rule min_absorb1) simp lemma min_0R [simp]: "min n 0 = 0" for n :: nat by (rule min_absorb2) simp lemma min_Suc_Suc [simp]: "min (Suc m) (Suc n) = Suc (min m n)" by (simp add: mono_Suc min_of_mono) lemma min_Suc1: "min (Suc n) m = (case m of 0 \ 0 | Suc m' \ Suc(min n m'))" by (simp split: nat.split) lemma min_Suc2: "min m (Suc n) = (case m of 0 \ 0 | Suc m' \ Suc(min m' n))" by (simp split: nat.split) lemma max_0L [simp]: "max 0 n = n" for n :: nat - by (rule max_absorb2) simp + by (fact max_nat.left_neutral) lemma max_0R [simp]: "max n 0 = n" for n :: nat - by (rule max_absorb1) simp + by (fact max_nat.right_neutral) lemma max_Suc_Suc [simp]: "max (Suc m) (Suc n) = Suc (max m n)" by (simp add: mono_Suc max_of_mono) lemma max_Suc1: "max (Suc n) m = (case m of 0 \ Suc n | Suc m' \ Suc (max n m'))" by (simp split: nat.split) lemma max_Suc2: "max m (Suc n) = (case m of 0 \ Suc n | Suc m' \ Suc (max m' n))" by (simp split: nat.split) -lemma max_0_iff[simp]: "max m n = (0::nat) \ m = 0 \ n = 0" -by(cases m, auto simp: max_Suc1 split: nat.split) +lemma max_0_iff: "max m n = (0::nat) \ m = 0 \ n = 0" + by (fact max_nat.eq_neutr_iff) lemma nat_mult_min_left: "min m n * q = min (m * q) (n * q)" for m n q :: nat by (simp add: min_def not_le) (auto dest: mult_right_le_imp_le mult_right_less_imp_less le_less_trans) lemma nat_mult_min_right: "m * min n q = min (m * n) (m * q)" for m n q :: nat by (simp add: min_def not_le) (auto dest: mult_left_le_imp_le mult_left_less_imp_less le_less_trans) lemma nat_add_max_left: "max m n + q = max (m + q) (n + q)" for m n q :: nat by (simp add: max_def) lemma nat_add_max_right: "m + max n q = max (m + n) (m + q)" for m n q :: nat by (simp add: max_def) lemma nat_mult_max_left: "max m n * q = max (m * q) (n * q)" for m n q :: nat by (simp add: max_def not_le) (auto dest: mult_right_le_imp_le mult_right_less_imp_less le_less_trans) lemma nat_mult_max_right: "m * max n q = max (m * n) (m * q)" for m n q :: nat by (simp add: max_def not_le) (auto dest: mult_left_le_imp_le mult_left_less_imp_less le_less_trans) subsubsection \Additional theorems about \<^term>\(\)\\ text \Complete induction, aka course-of-values induction\ instance nat :: wellorder proof fix P and n :: nat assume step: "(\m. m < n \ P m) \ P n" for n :: nat have "\q. q \ n \ P q" proof (induct n) case (0 n) have "P 0" by (rule step) auto with 0 show ?case by auto next case (Suc m n) then have "n \ m \ n = Suc m" by (simp add: le_Suc_eq) then show ?case proof assume "n \ m" then show "P n" by (rule Suc(1)) next assume n: "n = Suc m" show "P n" by (rule step) (rule Suc(1), simp add: n le_simps) qed qed then show "P n" by auto qed lemma Least_eq_0[simp]: "P 0 \ Least P = 0" for P :: "nat \ bool" by (rule Least_equality[OF _ le0]) lemma Least_Suc: assumes "P n" "\ P 0" shows "(LEAST n. P n) = Suc (LEAST m. P (Suc m))" proof (cases n) case (Suc m) show ?thesis proof (rule antisym) show "(LEAST x. P x) \ Suc (LEAST x. P (Suc x))" using assms Suc by (force intro: LeastI Least_le) have \
: "P (LEAST x. P x)" by (blast intro: LeastI assms) show "Suc (LEAST m. P (Suc m)) \ (LEAST n. P n)" proof (cases "(LEAST n. P n)") case 0 then show ?thesis using \
by (simp add: assms) next case Suc with \
show ?thesis by (auto simp: Least_le) qed qed qed (use assms in auto) lemma Least_Suc2: "P n \ Q m \ \ P 0 \ \k. P (Suc k) = Q k \ Least P = Suc (Least Q)" by (erule (1) Least_Suc [THEN ssubst]) simp lemma ex_least_nat_le: fixes P :: "nat \ bool" assumes "P n" "\ P 0" shows "\k\n. (\i P i) \ P k" proof (cases n) case (Suc m) with assms show ?thesis by (blast intro: Least_le LeastI_ex dest: not_less_Least) qed (use assms in auto) lemma ex_least_nat_less: fixes P :: "nat \ bool" assumes "P n" "\ P 0" shows "\ki\k. \ P i) \ P (Suc k)" proof (cases n) case (Suc m) then obtain k where k: "k \ n" "\i P i" "P k" using ex_least_nat_le [OF assms] by blast show ?thesis by (cases k) (use assms k less_eq_Suc_le in auto) qed (use assms in auto) lemma nat_less_induct: fixes P :: "nat \ bool" assumes "\n. \m. m < n \ P m \ P n" shows "P n" using assms less_induct by blast lemma measure_induct_rule [case_names less]: fixes f :: "'a \ 'b::wellorder" assumes step: "\x. (\y. f y < f x \ P y) \ P x" shows "P a" by (induct m \ "f a" arbitrary: a rule: less_induct) (auto intro: step) text \old style induction rules:\ lemma measure_induct: fixes f :: "'a \ 'b::wellorder" shows "(\x. \y. f y < f x \ P y \ P x) \ P a" by (rule measure_induct_rule [of f P a]) iprover lemma full_nat_induct: assumes step: "\n. (\m. Suc m \ n \ P m) \ P n" shows "P n" by (rule less_induct) (auto intro: step simp:le_simps) text\An induction rule for establishing binary relations\ lemma less_Suc_induct [consumes 1]: assumes less: "i < j" and step: "\i. P i (Suc i)" and trans: "\i j k. i < j \ j < k \ P i j \ P j k \ P i k" shows "P i j" proof - from less obtain k where j: "j = Suc (i + k)" by (auto dest: less_imp_Suc_add) have "P i (Suc (i + k))" proof (induct k) case 0 show ?case by (simp add: step) next case (Suc k) have "0 + i < Suc k + i" by (rule add_less_mono1) simp then have "i < Suc (i + k)" by (simp add: add.commute) from trans[OF this lessI Suc step] show ?case by simp qed then show "P i j" by (simp add: j) qed text \ The method of infinite descent, frequently used in number theory. Provided by Roelof Oosterhuis. \P n\ is true for all natural numbers if \<^item> case ``0'': given \n = 0\ prove \P n\ \<^item> case ``smaller'': given \n > 0\ and \\ P n\ prove there exists a smaller natural number \m\ such that \\ P m\. \ lemma infinite_descent: "(\n. \ P n \ \m P m) \ P n" for P :: "nat \ bool" \ \compact version without explicit base case\ by (induct n rule: less_induct) auto lemma infinite_descent0 [case_names 0 smaller]: fixes P :: "nat \ bool" assumes "P 0" and "\n. n > 0 \ \ P n \ \m. m < n \ \ P m" shows "P n" proof (rule infinite_descent) show "\n. \ P n \ \m P m" using assms by (case_tac "n > 0") auto qed text \ Infinite descent using a mapping to \nat\: \P x\ is true for all \x \ D\ if there exists a \V \ D \ nat\ and \<^item> case ``0'': given \V x = 0\ prove \P x\ \<^item> ``smaller'': given \V x > 0\ and \\ P x\ prove there exists a \y \ D\ such that \V y < V x\ and \\ P y\. \ corollary infinite_descent0_measure [case_names 0 smaller]: fixes V :: "'a \ nat" assumes 1: "\x. V x = 0 \ P x" and 2: "\x. V x > 0 \ \ P x \ \y. V y < V x \ \ P y" shows "P x" proof - obtain n where "n = V x" by auto moreover have "\x. V x = n \ P x" proof (induct n rule: infinite_descent0) case 0 with 1 show "P x" by auto next case (smaller n) then obtain x where *: "V x = n " and "V x > 0 \ \ P x" by auto with 2 obtain y where "V y < V x \ \ P y" by auto with * obtain m where "m = V y \ m < n \ \ P y" by auto then show ?case by auto qed ultimately show "P x" by auto qed text \Again, without explicit base case:\ lemma infinite_descent_measure: fixes V :: "'a \ nat" assumes "\x. \ P x \ \y. V y < V x \ \ P y" shows "P x" proof - from assms obtain n where "n = V x" by auto moreover have "\x. V x = n \ P x" proof (induct n rule: infinite_descent, auto) show "\m < V x. \y. V y = m \ \ P y" if "\ P x" for x using assms and that by auto qed ultimately show "P x" by auto qed text \A (clumsy) way of lifting \<\ monotonicity to \\\ monotonicity\ lemma less_mono_imp_le_mono: fixes f :: "nat \ nat" and i j :: nat assumes "\i j::nat. i < j \ f i < f j" and "i \ j" shows "f i \ f j" using assms by (auto simp add: order_le_less) text \non-strict, in 1st argument\ lemma add_le_mono1: "i \ j \ i + k \ j + k" for i j k :: nat by (rule add_right_mono) text \non-strict, in both arguments\ lemma add_le_mono: "i \ j \ k \ l \ i + k \ j + l" for i j k l :: nat by (rule add_mono) lemma le_add2: "n \ m + n" for m n :: nat by simp lemma le_add1: "n \ n + m" for m n :: nat by simp lemma less_add_Suc1: "i < Suc (i + m)" by (rule le_less_trans, rule le_add1, rule lessI) lemma less_add_Suc2: "i < Suc (m + i)" by (rule le_less_trans, rule le_add2, rule lessI) lemma less_iff_Suc_add: "m < n \ (\k. n = Suc (m + k))" by (iprover intro!: less_add_Suc1 less_imp_Suc_add) lemma trans_le_add1: "i \ j \ i \ j + m" for i j m :: nat by (rule le_trans, assumption, rule le_add1) lemma trans_le_add2: "i \ j \ i \ m + j" for i j m :: nat by (rule le_trans, assumption, rule le_add2) lemma trans_less_add1: "i < j \ i < j + m" for i j m :: nat by (rule less_le_trans, assumption, rule le_add1) lemma trans_less_add2: "i < j \ i < m + j" for i j m :: nat by (rule less_le_trans, assumption, rule le_add2) lemma add_lessD1: "i + j < k \ i < k" for i j k :: nat by (rule le_less_trans [of _ "i+j"]) (simp_all add: le_add1) lemma not_add_less1 [iff]: "\ i + j < i" for i j :: nat by simp lemma not_add_less2 [iff]: "\ j + i < i" for i j :: nat by simp lemma add_leD1: "m + k \ n \ m \ n" for k m n :: nat by (rule order_trans [of _ "m + k"]) (simp_all add: le_add1) lemma add_leD2: "m + k \ n \ k \ n" for k m n :: nat by (force simp add: add.commute dest: add_leD1) lemma add_leE: "m + k \ n \ (m \ n \ k \ n \ R) \ R" for k m n :: nat by (blast dest: add_leD1 add_leD2) text \needs \\k\ for \ac_simps\ to work\ lemma less_add_eq_less: "\k. k < l \ m + l = k + n \ m < n" for l m n :: nat by (force simp del: add_Suc_right simp add: less_iff_Suc_add add_Suc_right [symmetric] ac_simps) subsubsection \More results about difference\ lemma Suc_diff_le: "n \ m \ Suc m - n = Suc (m - n)" by (induct m n rule: diff_induct) simp_all lemma diff_less_Suc: "m - n < Suc m" by (induct m n rule: diff_induct) (auto simp: less_Suc_eq) lemma diff_le_self [simp]: "m - n \ m" for m n :: nat by (induct m n rule: diff_induct) (simp_all add: le_SucI) lemma less_imp_diff_less: "j < k \ j - n < k" for j k n :: nat by (rule le_less_trans, rule diff_le_self) lemma diff_Suc_less [simp]: "0 < n \ n - Suc i < n" by (cases n) (auto simp add: le_simps) lemma diff_add_assoc: "k \ j \ (i + j) - k = i + (j - k)" for i j k :: nat by (fact ordered_cancel_comm_monoid_diff_class.diff_add_assoc) lemma add_diff_assoc [simp]: "k \ j \ i + (j - k) = i + j - k" for i j k :: nat by (fact ordered_cancel_comm_monoid_diff_class.add_diff_assoc) lemma diff_add_assoc2: "k \ j \ (j + i) - k = (j - k) + i" for i j k :: nat by (fact ordered_cancel_comm_monoid_diff_class.diff_add_assoc2) lemma add_diff_assoc2 [simp]: "k \ j \ j - k + i = j + i - k" for i j k :: nat by (fact ordered_cancel_comm_monoid_diff_class.add_diff_assoc2) lemma le_imp_diff_is_add: "i \ j \ (j - i = k) = (j = k + i)" for i j k :: nat by auto lemma diff_is_0_eq [simp]: "m - n = 0 \ m \ n" for m n :: nat by (induct m n rule: diff_induct) simp_all lemma diff_is_0_eq' [simp]: "m \ n \ m - n = 0" for m n :: nat by (rule iffD2, rule diff_is_0_eq) lemma zero_less_diff [simp]: "0 < n - m \ m < n" for m n :: nat by (induct m n rule: diff_induct) simp_all lemma less_imp_add_positive: assumes "i < j" shows "\k::nat. 0 < k \ i + k = j" proof from assms show "0 < j - i \ i + (j - i) = j" by (simp add: order_less_imp_le) qed text \a nice rewrite for bounded subtraction\ lemma nat_minus_add_max: "n - m + m = max n m" for m n :: nat by (simp add: max_def not_le order_less_imp_le) lemma nat_diff_split: "P (a - b) \ (a < b \ P 0) \ (\d. a = b + d \ P d)" for a b :: nat \ \elimination of \-\ on \nat\\ by (cases "a < b") (auto simp add: not_less le_less dest!: add_eq_self_zero [OF sym]) lemma nat_diff_split_asm: "P (a - b) \ \ (a < b \ \ P 0 \ (\d. a = b + d \ \ P d))" for a b :: nat \ \elimination of \-\ on \nat\ in assumptions\ by (auto split: nat_diff_split) lemma Suc_pred': "0 < n \ n = Suc(n - 1)" by simp lemma add_eq_if: "m + n = (if m = 0 then n else Suc ((m - 1) + n))" unfolding One_nat_def by (cases m) simp_all lemma mult_eq_if: "m * n = (if m = 0 then 0 else n + ((m - 1) * n))" for m n :: nat by (cases m) simp_all lemma Suc_diff_eq_diff_pred: "0 < n \ Suc m - n = m - (n - 1)" by (cases n) simp_all lemma diff_Suc_eq_diff_pred: "m - Suc n = (m - 1) - n" by (cases m) simp_all lemma Let_Suc [simp]: "Let (Suc n) f \ f (Suc n)" by (fact Let_def) subsubsection \Monotonicity of multiplication\ lemma mult_le_mono1: "i \ j \ i * k \ j * k" for i j k :: nat by (simp add: mult_right_mono) lemma mult_le_mono2: "i \ j \ k * i \ k * j" for i j k :: nat by (simp add: mult_left_mono) text \\\\ monotonicity, BOTH arguments\ lemma mult_le_mono: "i \ j \ k \ l \ i * k \ j * l" for i j k l :: nat by (simp add: mult_mono) lemma mult_less_mono1: "i < j \ 0 < k \ i * k < j * k" for i j k :: nat by (simp add: mult_strict_right_mono) text \Differs from the standard \zero_less_mult_iff\ in that there are no negative numbers.\ lemma nat_0_less_mult_iff [simp]: "0 < m * n \ 0 < m \ 0 < n" for m n :: nat proof (induct m) case 0 then show ?case by simp next case (Suc m) then show ?case by (cases n) simp_all qed lemma one_le_mult_iff [simp]: "Suc 0 \ m * n \ Suc 0 \ m \ Suc 0 \ n" proof (induct m) case 0 then show ?case by simp next case (Suc m) then show ?case by (cases n) simp_all qed lemma mult_less_cancel2 [simp]: "m * k < n * k \ 0 < k \ m < n" for k m n :: nat proof (intro iffI conjI) assume m: "m * k < n * k" then show "0 < k" by (cases k) auto show "m < n" proof (cases k) case 0 then show ?thesis using m by auto next case (Suc k') then show ?thesis using m by (simp flip: linorder_not_le) (blast intro: add_mono mult_le_mono1) qed next assume "0 < k \ m < n" then show "m * k < n * k" by (blast intro: mult_less_mono1) qed lemma mult_less_cancel1 [simp]: "k * m < k * n \ 0 < k \ m < n" for k m n :: nat by (simp add: mult.commute [of k]) lemma mult_le_cancel1 [simp]: "k * m \ k * n \ (0 < k \ m \ n)" for k m n :: nat by (simp add: linorder_not_less [symmetric], auto) lemma mult_le_cancel2 [simp]: "m * k \ n * k \ (0 < k \ m \ n)" for k m n :: nat by (simp add: linorder_not_less [symmetric], auto) lemma Suc_mult_less_cancel1: "Suc k * m < Suc k * n \ m < n" by (subst mult_less_cancel1) simp lemma Suc_mult_le_cancel1: "Suc k * m \ Suc k * n \ m \ n" by (subst mult_le_cancel1) simp lemma le_square: "m \ m * m" for m :: nat by (cases m) (auto intro: le_add1) lemma le_cube: "m \ m * (m * m)" for m :: nat by (cases m) (auto intro: le_add1) text \Lemma for \gcd\\ lemma mult_eq_self_implies_10: fixes m n :: nat assumes "m = m * n" shows "n = 1 \ m = 0" proof (rule disjCI) assume "m \ 0" show "n = 1" proof (cases n "1::nat" rule: linorder_cases) case greater show ?thesis using assms mult_less_mono2 [OF greater, of m] \m \ 0\ by auto qed (use assms \m \ 0\ in auto) qed lemma mono_times_nat: fixes n :: nat assumes "n > 0" shows "mono (times n)" proof fix m q :: nat assume "m \ q" with assms show "n * m \ n * q" by simp qed text \The lattice order on \<^typ>\nat\.\ instantiation nat :: distrib_lattice begin definition "(inf :: nat \ nat \ nat) = min" definition "(sup :: nat \ nat \ nat) = max" instance by intro_classes (auto simp add: inf_nat_def sup_nat_def max_def not_le min_def intro: order_less_imp_le antisym elim!: order_trans order_less_trans) end subsection \Natural operation of natural numbers on functions\ text \ We use the same logical constant for the power operations on functions and relations, in order to share the same syntax. \ consts compow :: "nat \ 'a \ 'a" abbreviation compower :: "'a \ nat \ 'a" (infixr "^^" 80) where "f ^^ n \ compow n f" notation (latex output) compower ("(_\<^bsup>_\<^esup>)" [1000] 1000) text \\f ^^ n = f \ \ \ f\, the \n\-fold composition of \f\\ overloading funpow \ "compow :: nat \ ('a \ 'a) \ ('a \ 'a)" begin primrec funpow :: "nat \ ('a \ 'a) \ 'a \ 'a" where "funpow 0 f = id" | "funpow (Suc n) f = f \ funpow n f" end lemma funpow_0 [simp]: "(f ^^ 0) x = x" by simp lemma funpow_Suc_right: "f ^^ Suc n = f ^^ n \ f" proof (induct n) case 0 then show ?case by simp next fix n assume "f ^^ Suc n = f ^^ n \ f" then show "f ^^ Suc (Suc n) = f ^^ Suc n \ f" by (simp add: o_assoc) qed lemmas funpow_simps_right = funpow.simps(1) funpow_Suc_right text \For code generation.\ definition funpow :: "nat \ ('a \ 'a) \ 'a \ 'a" where funpow_code_def [code_abbrev]: "funpow = compow" lemma [code]: "funpow (Suc n) f = f \ funpow n f" "funpow 0 f = id" by (simp_all add: funpow_code_def) hide_const (open) funpow lemma funpow_add: "f ^^ (m + n) = f ^^ m \ f ^^ n" by (induct m) simp_all lemma funpow_mult: "(f ^^ m) ^^ n = f ^^ (m * n)" for f :: "'a \ 'a" by (induct n) (simp_all add: funpow_add) lemma funpow_swap1: "f ((f ^^ n) x) = (f ^^ n) (f x)" proof - have "f ((f ^^ n) x) = (f ^^ (n + 1)) x" by simp also have "\ = (f ^^ n \ f ^^ 1) x" by (simp only: funpow_add) also have "\ = (f ^^ n) (f x)" by simp finally show ?thesis . qed lemma comp_funpow: "comp f ^^ n = comp (f ^^ n)" for f :: "'a \ 'a" by (induct n) simp_all lemma Suc_funpow[simp]: "Suc ^^ n = ((+) n)" by (induct n) simp_all lemma id_funpow[simp]: "id ^^ n = id" by (induct n) simp_all lemma funpow_mono: "mono f \ A \ B \ (f ^^ n) A \ (f ^^ n) B" for f :: "'a \ ('a::order)" by (induct n arbitrary: A B) (auto simp del: funpow.simps(2) simp add: funpow_Suc_right mono_def) lemma funpow_mono2: assumes "mono f" and "i \ j" and "x \ y" and "x \ f x" shows "(f ^^ i) x \ (f ^^ j) y" using assms(2,3) proof (induct j arbitrary: y) case 0 then show ?case by simp next case (Suc j) show ?case proof(cases "i = Suc j") case True with assms(1) Suc show ?thesis by (simp del: funpow.simps add: funpow_simps_right monoD funpow_mono) next case False with assms(1,4) Suc show ?thesis by (simp del: funpow.simps add: funpow_simps_right le_eq_less_or_eq less_Suc_eq_le) (simp add: Suc.hyps monoD order_subst1) qed qed lemma inj_fn[simp]: fixes f::"'a \ 'a" assumes "inj f" shows "inj (f^^n)" proof (induction n) case Suc thus ?case using inj_compose[OF assms Suc.IH] by (simp del: comp_apply) qed simp lemma surj_fn[simp]: fixes f::"'a \ 'a" assumes "surj f" shows "surj (f^^n)" proof (induction n) case Suc thus ?case by (simp add: comp_surj[OF Suc.IH assms] del: comp_apply) qed simp lemma bij_fn[simp]: fixes f::"'a \ 'a" assumes "bij f" shows "bij (f^^n)" by (rule bijI[OF inj_fn[OF bij_is_inj[OF assms]] surj_fn[OF bij_is_surj[OF assms]]]) subsection \Kleene iteration\ lemma Kleene_iter_lpfp: fixes f :: "'a::order_bot \ 'a" assumes "mono f" and "f p \ p" shows "(f ^^ k) bot \ p" proof (induct k) case 0 show ?case by simp next case Suc show ?case using monoD[OF assms(1) Suc] assms(2) by simp qed lemma lfp_Kleene_iter: assumes "mono f" and "(f ^^ Suc k) bot = (f ^^ k) bot" shows "lfp f = (f ^^ k) bot" proof (rule antisym) show "lfp f \ (f ^^ k) bot" proof (rule lfp_lowerbound) show "f ((f ^^ k) bot) \ (f ^^ k) bot" using assms(2) by simp qed show "(f ^^ k) bot \ lfp f" using Kleene_iter_lpfp[OF assms(1)] lfp_unfold[OF assms(1)] by simp qed lemma mono_pow: "mono f \ mono (f ^^ n)" for f :: "'a \ 'a::complete_lattice" by (induct n) (auto simp: mono_def) lemma lfp_funpow: assumes f: "mono f" shows "lfp (f ^^ Suc n) = lfp f" proof (rule antisym) show "lfp f \ lfp (f ^^ Suc n)" proof (rule lfp_lowerbound) have "f (lfp (f ^^ Suc n)) = lfp (\x. f ((f ^^ n) x))" unfolding funpow_Suc_right by (simp add: lfp_rolling f mono_pow comp_def) then show "f (lfp (f ^^ Suc n)) \ lfp (f ^^ Suc n)" by (simp add: comp_def) qed have "(f ^^ n) (lfp f) = lfp f" for n by (induct n) (auto intro: f lfp_fixpoint) then show "lfp (f ^^ Suc n) \ lfp f" by (intro lfp_lowerbound) (simp del: funpow.simps) qed lemma gfp_funpow: assumes f: "mono f" shows "gfp (f ^^ Suc n) = gfp f" proof (rule antisym) show "gfp f \ gfp (f ^^ Suc n)" proof (rule gfp_upperbound) have "f (gfp (f ^^ Suc n)) = gfp (\x. f ((f ^^ n) x))" unfolding funpow_Suc_right by (simp add: gfp_rolling f mono_pow comp_def) then show "f (gfp (f ^^ Suc n)) \ gfp (f ^^ Suc n)" by (simp add: comp_def) qed have "(f ^^ n) (gfp f) = gfp f" for n by (induct n) (auto intro: f gfp_fixpoint) then show "gfp (f ^^ Suc n) \ gfp f" by (intro gfp_upperbound) (simp del: funpow.simps) qed lemma Kleene_iter_gpfp: fixes f :: "'a::order_top \ 'a" assumes "mono f" and "p \ f p" shows "p \ (f ^^ k) top" proof (induct k) case 0 show ?case by simp next case Suc show ?case using monoD[OF assms(1) Suc] assms(2) by simp qed lemma gfp_Kleene_iter: assumes "mono f" and "(f ^^ Suc k) top = (f ^^ k) top" shows "gfp f = (f ^^ k) top" (is "?lhs = ?rhs") proof (rule antisym) have "?rhs \ f ?rhs" using assms(2) by simp then show "?rhs \ ?lhs" by (rule gfp_upperbound) show "?lhs \ ?rhs" using Kleene_iter_gpfp[OF assms(1)] gfp_unfold[OF assms(1)] by simp qed subsection \Embedding of the naturals into any \semiring_1\: \<^term>\of_nat\\ context semiring_1 begin definition of_nat :: "nat \ 'a" where "of_nat n = (plus 1 ^^ n) 0" lemma of_nat_simps [simp]: shows of_nat_0: "of_nat 0 = 0" and of_nat_Suc: "of_nat (Suc m) = 1 + of_nat m" by (simp_all add: of_nat_def) lemma of_nat_1 [simp]: "of_nat 1 = 1" by (simp add: of_nat_def) lemma of_nat_add [simp]: "of_nat (m + n) = of_nat m + of_nat n" by (induct m) (simp_all add: ac_simps) lemma of_nat_mult [simp]: "of_nat (m * n) = of_nat m * of_nat n" by (induct m) (simp_all add: ac_simps distrib_right) lemma mult_of_nat_commute: "of_nat x * y = y * of_nat x" by (induct x) (simp_all add: algebra_simps) primrec of_nat_aux :: "('a \ 'a) \ nat \ 'a \ 'a" where "of_nat_aux inc 0 i = i" | "of_nat_aux inc (Suc n) i = of_nat_aux inc n (inc i)" \ \tail recursive\ lemma of_nat_code: "of_nat n = of_nat_aux (\i. i + 1) n 0" proof (induct n) case 0 then show ?case by simp next case (Suc n) have "\i. of_nat_aux (\i. i + 1) n (i + 1) = of_nat_aux (\i. i + 1) n i + 1" by (induct n) simp_all from this [of 0] have "of_nat_aux (\i. i + 1) n 1 = of_nat_aux (\i. i + 1) n 0 + 1" by simp with Suc show ?case by (simp add: add.commute) qed lemma of_nat_of_bool [simp]: "of_nat (of_bool P) = of_bool P" by auto end declare of_nat_code [code] context semiring_1_cancel begin lemma of_nat_diff: \of_nat (m - n) = of_nat m - of_nat n\ if \n \ m\ proof - from that obtain q where \m = n + q\ by (blast dest: le_Suc_ex) then show ?thesis by simp qed end text \Class for unital semirings with characteristic zero. Includes non-ordered rings like the complex numbers.\ class semiring_char_0 = semiring_1 + assumes inj_of_nat: "inj of_nat" begin lemma of_nat_eq_iff [simp]: "of_nat m = of_nat n \ m = n" by (auto intro: inj_of_nat injD) text \Special cases where either operand is zero\ lemma of_nat_0_eq_iff [simp]: "0 = of_nat n \ 0 = n" by (fact of_nat_eq_iff [of 0 n, unfolded of_nat_0]) lemma of_nat_eq_0_iff [simp]: "of_nat m = 0 \ m = 0" by (fact of_nat_eq_iff [of m 0, unfolded of_nat_0]) lemma of_nat_1_eq_iff [simp]: "1 = of_nat n \ n=1" using of_nat_eq_iff by fastforce lemma of_nat_eq_1_iff [simp]: "of_nat n = 1 \ n=1" using of_nat_eq_iff by fastforce lemma of_nat_neq_0 [simp]: "of_nat (Suc n) \ 0" unfolding of_nat_eq_0_iff by simp lemma of_nat_0_neq [simp]: "0 \ of_nat (Suc n)" unfolding of_nat_0_eq_iff by simp end class ring_char_0 = ring_1 + semiring_char_0 context linordered_nonzero_semiring begin lemma of_nat_0_le_iff [simp]: "0 \ of_nat n" by (induct n) simp_all lemma of_nat_less_0_iff [simp]: "\ of_nat m < 0" by (simp add: not_less) lemma of_nat_mono[simp]: "i \ j \ of_nat i \ of_nat j" by (auto simp: le_iff_add intro!: add_increasing2) lemma of_nat_less_iff [simp]: "of_nat m < of_nat n \ m < n" proof(induct m n rule: diff_induct) case (1 m) then show ?case by auto next case (2 n) then show ?case by (simp add: add_pos_nonneg) next case (3 m n) then show ?case by (auto simp: add_commute [of 1] add_mono1 not_less add_right_mono leD) qed lemma of_nat_le_iff [simp]: "of_nat m \ of_nat n \ m \ n" by (simp add: not_less [symmetric] linorder_not_less [symmetric]) lemma less_imp_of_nat_less: "m < n \ of_nat m < of_nat n" by simp lemma of_nat_less_imp_less: "of_nat m < of_nat n \ m < n" by simp text \Every \linordered_nonzero_semiring\ has characteristic zero.\ subclass semiring_char_0 by standard (auto intro!: injI simp add: eq_iff) text \Special cases where either operand is zero\ lemma of_nat_le_0_iff [simp]: "of_nat m \ 0 \ m = 0" by (rule of_nat_le_iff [of _ 0, simplified]) lemma of_nat_0_less_iff [simp]: "0 < of_nat n \ 0 < n" by (rule of_nat_less_iff [of 0, simplified]) end context linordered_nonzero_semiring begin lemma of_nat_max: "of_nat (max x y) = max (of_nat x) (of_nat y)" by (auto simp: max_def ord_class.max_def) lemma of_nat_min: "of_nat (min x y) = min (of_nat x) (of_nat y)" by (auto simp: min_def ord_class.min_def) end context linordered_semidom begin subclass linordered_nonzero_semiring .. subclass semiring_char_0 .. end context linordered_idom begin lemma abs_of_nat [simp]: "\of_nat n\ = of_nat n" by (simp add: abs_if) lemma sgn_of_nat [simp]: "sgn (of_nat n) = of_bool (n > 0)" by simp end lemma of_nat_id [simp]: "of_nat n = n" by (induct n) simp_all lemma of_nat_eq_id [simp]: "of_nat = id" by (auto simp add: fun_eq_iff) subsection \The set of natural numbers\ context semiring_1 begin definition Nats :: "'a set" ("\") where "\ = range of_nat" lemma of_nat_in_Nats [simp]: "of_nat n \ \" by (simp add: Nats_def) lemma Nats_0 [simp]: "0 \ \" using of_nat_0 [symmetric] unfolding Nats_def by (rule range_eqI) lemma Nats_1 [simp]: "1 \ \" using of_nat_1 [symmetric] unfolding Nats_def by (rule range_eqI) lemma Nats_add [simp]: "a \ \ \ b \ \ \ a + b \ \" unfolding Nats_def using of_nat_add [symmetric] by (blast intro: range_eqI) lemma Nats_mult [simp]: "a \ \ \ b \ \ \ a * b \ \" unfolding Nats_def using of_nat_mult [symmetric] by (blast intro: range_eqI) lemma Nats_cases [cases set: Nats]: assumes "x \ \" obtains (of_nat) n where "x = of_nat n" unfolding Nats_def proof - from \x \ \\ have "x \ range of_nat" unfolding Nats_def . then obtain n where "x = of_nat n" .. then show thesis .. qed lemma Nats_induct [case_names of_nat, induct set: Nats]: "x \ \ \ (\n. P (of_nat n)) \ P x" by (rule Nats_cases) auto end lemma Nats_diff [simp]: fixes a:: "'a::linordered_idom" assumes "a \ \" "b \ \" "b \ a" shows "a - b \ \" proof - obtain i where i: "a = of_nat i" using Nats_cases assms by blast obtain j where j: "b = of_nat j" using Nats_cases assms by blast have "j \ i" using \b \ a\ i j of_nat_le_iff by blast then have *: "of_nat i - of_nat j = (of_nat (i-j) :: 'a)" by (simp add: of_nat_diff) then show ?thesis by (simp add: * i j) qed subsection \Further arithmetic facts concerning the natural numbers\ lemma subst_equals: assumes "t = s" and "u = t" shows "u = s" using assms(2,1) by (rule trans) locale nat_arith begin lemma add1: "(A::'a::comm_monoid_add) \ k + a \ A + b \ k + (a + b)" by (simp only: ac_simps) lemma add2: "(B::'a::comm_monoid_add) \ k + b \ a + B \ k + (a + b)" by (simp only: ac_simps) lemma suc1: "A == k + a \ Suc A \ k + Suc a" by (simp only: add_Suc_right) lemma rule0: "(a::'a::comm_monoid_add) \ a + 0" by (simp only: add_0_right) end ML_file \Tools/nat_arith.ML\ simproc_setup nateq_cancel_sums ("(l::nat) + m = n" | "(l::nat) = m + n" | "Suc m = n" | "m = Suc n") = \fn phi => try o Nat_Arith.cancel_eq_conv\ simproc_setup natless_cancel_sums ("(l::nat) + m < n" | "(l::nat) < m + n" | "Suc m < n" | "m < Suc n") = \fn phi => try o Nat_Arith.cancel_less_conv\ simproc_setup natle_cancel_sums ("(l::nat) + m \ n" | "(l::nat) \ m + n" | "Suc m \ n" | "m \ Suc n") = \fn phi => try o Nat_Arith.cancel_le_conv\ simproc_setup natdiff_cancel_sums ("(l::nat) + m - n" | "(l::nat) - (m + n)" | "Suc m - n" | "m - Suc n") = \fn phi => try o Nat_Arith.cancel_diff_conv\ context order begin lemma lift_Suc_mono_le: assumes mono: "\n. f n \ f (Suc n)" and "n \ n'" shows "f n \ f n'" proof (cases "n < n'") case True then show ?thesis by (induct n n' rule: less_Suc_induct) (auto intro: mono) next case False with \n \ n'\ show ?thesis by auto qed lemma lift_Suc_antimono_le: assumes mono: "\n. f n \ f (Suc n)" and "n \ n'" shows "f n \ f n'" proof (cases "n < n'") case True then show ?thesis by (induct n n' rule: less_Suc_induct) (auto intro: mono) next case False with \n \ n'\ show ?thesis by auto qed lemma lift_Suc_mono_less: assumes mono: "\n. f n < f (Suc n)" and "n < n'" shows "f n < f n'" using \n < n'\ by (induct n n' rule: less_Suc_induct) (auto intro: mono) lemma lift_Suc_mono_less_iff: "(\n. f n < f (Suc n)) \ f n < f m \ n < m" by (blast intro: less_asym' lift_Suc_mono_less [of f] dest: linorder_not_less[THEN iffD1] le_eq_less_or_eq [THEN iffD1]) end lemma mono_iff_le_Suc: "mono f \ (\n. f n \ f (Suc n))" unfolding mono_def by (auto intro: lift_Suc_mono_le [of f]) lemma antimono_iff_le_Suc: "antimono f \ (\n. f (Suc n) \ f n)" unfolding antimono_def by (auto intro: lift_Suc_antimono_le [of f]) lemma mono_nat_linear_lb: fixes f :: "nat \ nat" assumes "\m n. m < n \ f m < f n" shows "f m + k \ f (m + k)" proof (induct k) case 0 then show ?case by simp next case (Suc k) then have "Suc (f m + k) \ Suc (f (m + k))" by simp also from assms [of "m + k" "Suc (m + k)"] have "Suc (f (m + k)) \ f (Suc (m + k))" by (simp add: Suc_le_eq) finally show ?case by simp qed text \Subtraction laws, mostly by Clemens Ballarin\ lemma diff_less_mono: fixes a b c :: nat assumes "a < b" and "c \ a" shows "a - c < b - c" proof - from assms obtain d e where "b = c + (d + e)" and "a = c + e" and "d > 0" by (auto dest!: le_Suc_ex less_imp_Suc_add simp add: ac_simps) then show ?thesis by simp qed lemma less_diff_conv: "i < j - k \ i + k < j" for i j k :: nat by (cases "k \ j") (auto simp add: not_le dest: less_imp_Suc_add le_Suc_ex) lemma less_diff_conv2: "k \ j \ j - k < i \ j < i + k" for j k i :: nat by (auto dest: le_Suc_ex) lemma le_diff_conv: "j - k \ i \ j \ i + k" for j k i :: nat by (cases "k \ j") (auto simp add: not_le dest!: less_imp_Suc_add le_Suc_ex) lemma diff_diff_cancel [simp]: "i \ n \ n - (n - i) = i" for i n :: nat by (auto dest: le_Suc_ex) lemma diff_less [simp]: "0 < n \ 0 < m \ m - n < m" for i n :: nat by (auto dest: less_imp_Suc_add) text \Simplification of relational expressions involving subtraction\ lemma diff_diff_eq: "k \ m \ k \ n \ m - k - (n - k) = m - n" for m n k :: nat by (auto dest!: le_Suc_ex) hide_fact (open) diff_diff_eq lemma eq_diff_iff: "k \ m \ k \ n \ m - k = n - k \ m = n" for m n k :: nat by (auto dest: le_Suc_ex) lemma less_diff_iff: "k \ m \ k \ n \ m - k < n - k \ m < n" for m n k :: nat by (auto dest!: le_Suc_ex) lemma le_diff_iff: "k \ m \ k \ n \ m - k \ n - k \ m \ n" for m n k :: nat by (auto dest!: le_Suc_ex) lemma le_diff_iff': "a \ c \ b \ c \ c - a \ c - b \ b \ a" for a b c :: nat by (force dest: le_Suc_ex) text \(Anti)Monotonicity of subtraction -- by Stephan Merz\ lemma diff_le_mono: "m \ n \ m - l \ n - l" for m n l :: nat by (auto dest: less_imp_le less_imp_Suc_add split: nat_diff_split) lemma diff_le_mono2: "m \ n \ l - n \ l - m" for m n l :: nat by (auto dest: less_imp_le le_Suc_ex less_imp_Suc_add less_le_trans split: nat_diff_split) lemma diff_less_mono2: "m < n \ m < l \ l - n < l - m" for m n l :: nat by (auto dest: less_imp_Suc_add split: nat_diff_split) lemma diffs0_imp_equal: "m - n = 0 \ n - m = 0 \ m = n" for m n :: nat by (simp split: nat_diff_split) lemma min_diff: "min (m - i) (n - i) = min m n - i" for m n i :: nat by (cases m n rule: le_cases) (auto simp add: not_le min.absorb1 min.absorb2 min.absorb_iff1 [symmetric] diff_le_mono) lemma inj_on_diff_nat: fixes k :: nat assumes "\n. n \ N \ k \ n" shows "inj_on (\n. n - k) N" proof (rule inj_onI) fix x y assume a: "x \ N" "y \ N" "x - k = y - k" with assms have "x - k + k = y - k + k" by auto with a assms show "x = y" by (auto simp add: eq_diff_iff) qed text \Rewriting to pull differences out\ lemma diff_diff_right [simp]: "k \ j \ i - (j - k) = i + k - j" for i j k :: nat by (fact diff_diff_right) lemma diff_Suc_diff_eq1 [simp]: assumes "k \ j" shows "i - Suc (j - k) = i + k - Suc j" proof - from assms have *: "Suc (j - k) = Suc j - k" by (simp add: Suc_diff_le) from assms have "k \ Suc j" by (rule order_trans) simp with diff_diff_right [of k "Suc j" i] * show ?thesis by simp qed lemma diff_Suc_diff_eq2 [simp]: assumes "k \ j" shows "Suc (j - k) - i = Suc j - (k + i)" proof - from assms obtain n where "j = k + n" by (auto dest: le_Suc_ex) moreover have "Suc n - i = (k + Suc n) - (k + i)" using add_diff_cancel_left [of k "Suc n" i] by simp ultimately show ?thesis by simp qed lemma Suc_diff_Suc: assumes "n < m" shows "Suc (m - Suc n) = m - n" proof - from assms obtain q where "m = n + Suc q" by (auto dest: less_imp_Suc_add) moreover define r where "r = Suc q" ultimately have "Suc (m - Suc n) = r" and "m = n + r" by simp_all then show ?thesis by simp qed lemma one_less_mult: "Suc 0 < n \ Suc 0 < m \ Suc 0 < m * n" using less_1_mult [of n m] by (simp add: ac_simps) lemma n_less_m_mult_n: "0 < n \ Suc 0 < m \ n < m * n" using mult_strict_right_mono [of 1 m n] by simp lemma n_less_n_mult_m: "0 < n \ Suc 0 < m \ n < n * m" using mult_strict_left_mono [of 1 m n] by simp text \Induction starting beyond zero\ lemma nat_induct_at_least [consumes 1, case_names base Suc]: "P n" if "n \ m" "P m" "\n. n \ m \ P n \ P (Suc n)" proof - define q where "q = n - m" with \n \ m\ have "n = m + q" by simp moreover have "P (m + q)" by (induction q) (use that in simp_all) ultimately show "P n" by simp qed lemma nat_induct_non_zero [consumes 1, case_names 1 Suc]: "P n" if "n > 0" "P 1" "\n. n > 0 \ P n \ P (Suc n)" proof - from \n > 0\ have "n \ 1" by (cases n) simp_all moreover note \P 1\ moreover have "\n. n \ 1 \ P n \ P (Suc n)" using \\n. n > 0 \ P n \ P (Suc n)\ by (simp add: Suc_le_eq) ultimately show "P n" by (rule nat_induct_at_least) qed text \Specialized induction principles that work "backwards":\ lemma inc_induct [consumes 1, case_names base step]: assumes less: "i \ j" and base: "P j" and step: "\n. i \ n \ n < j \ P (Suc n) \ P n" shows "P i" using less step proof (induct "j - i" arbitrary: i) case (0 i) then have "i = j" by simp with base show ?case by simp next case (Suc d n) from Suc.hyps have "n \ j" by auto with Suc have "n < j" by (simp add: less_le) from \Suc d = j - n\ have "d + 1 = j - n" by simp then have "d + 1 - 1 = j - n - 1" by simp then have "d = j - n - 1" by simp then have "d = j - (n + 1)" by (simp add: diff_diff_eq) then have "d = j - Suc n" by simp moreover from \n < j\ have "Suc n \ j" by (simp add: Suc_le_eq) ultimately have "P (Suc n)" proof (rule Suc.hyps) fix q assume "Suc n \ q" then have "n \ q" by (simp add: Suc_le_eq less_imp_le) moreover assume "q < j" moreover assume "P (Suc q)" ultimately show "P q" by (rule Suc.prems) qed with order_refl \n < j\ show "P n" by (rule Suc.prems) qed lemma strict_inc_induct [consumes 1, case_names base step]: assumes less: "i < j" and base: "\i. j = Suc i \ P i" and step: "\i. i < j \ P (Suc i) \ P i" shows "P i" using less proof (induct "j - i - 1" arbitrary: i) case (0 i) from \i < j\ obtain n where "j = i + n" and "n > 0" by (auto dest!: less_imp_Suc_add) with 0 have "j = Suc i" by (auto intro: order_antisym simp add: Suc_le_eq) with base show ?case by simp next case (Suc d i) from \Suc d = j - i - 1\ have *: "Suc d = j - Suc i" by (simp add: diff_diff_add) then have "Suc d - 1 = j - Suc i - 1" by simp then have "d = j - Suc i - 1" by simp moreover from * have "j - Suc i \ 0" by auto then have "Suc i < j" by (simp add: not_le) ultimately have "P (Suc i)" by (rule Suc.hyps) with \i < j\ show "P i" by (rule step) qed lemma zero_induct_lemma: "P k \ (\n. P (Suc n) \ P n) \ P (k - i)" using inc_induct[of "k - i" k P, simplified] by blast lemma zero_induct: "P k \ (\n. P (Suc n) \ P n) \ P 0" using inc_induct[of 0 k P] by blast text \Further induction rule similar to @{thm inc_induct}.\ lemma dec_induct [consumes 1, case_names base step]: "i \ j \ P i \ (\n. i \ n \ n < j \ P n \ P (Suc n)) \ P j" proof (induct j arbitrary: i) case 0 then show ?case by simp next case (Suc j) from Suc.prems consider "i \ j" | "i = Suc j" by (auto simp add: le_Suc_eq) then show ?case proof cases case 1 moreover have "j < Suc j" by simp moreover have "P j" using \i \ j\ \P i\ proof (rule Suc.hyps) fix q assume "i \ q" moreover assume "q < j" then have "q < Suc j" by (simp add: less_Suc_eq) moreover assume "P q" ultimately show "P (Suc q)" by (rule Suc.prems) qed ultimately show "P (Suc j)" by (rule Suc.prems) next case 2 with \P i\ show "P (Suc j)" by simp qed qed lemma transitive_stepwise_le: assumes "m \ n" "\x. R x x" "\x y z. R x y \ R y z \ R x z" and "\n. R n (Suc n)" shows "R m n" using \m \ n\ by (induction rule: dec_induct) (use assms in blast)+ subsubsection \Greatest operator\ lemma ex_has_greatest_nat: "P (k::nat) \ \y. P y \ y \ b \ \x. P x \ (\y. P y \ y \ x)" proof (induction "b-k" arbitrary: b k rule: less_induct) case less show ?case proof cases assume "\n>k. P n" then obtain n where "n>k" "P n" by blast have "n \ b" using \P n\ less.prems(2) by auto hence "b-n < b-k" by(rule diff_less_mono2[OF \k less_le_trans[OF \k]]) from less.hyps[OF this \P n\ less.prems(2)] show ?thesis . next assume "\ (\n>k. P n)" hence "\y. P y \ y \ k" by (auto simp: not_less) thus ?thesis using less.prems(1) by auto qed qed lemma fixes k::nat assumes "P k" and minor: "\y. P y \ y \ b" shows GreatestI_nat: "P (Greatest P)" and Greatest_le_nat: "k \ Greatest P" proof - obtain x where "P x" "\y. P y \ y \ x" using assms ex_has_greatest_nat by blast with \P k\ show "P (Greatest P)" "k \ Greatest P" using GreatestI2_order by blast+ qed lemma GreatestI_ex_nat: "\ \k::nat. P k; \y. P y \ y \ b \ \ P (Greatest P)" by (blast intro: GreatestI_nat) subsection \Monotonicity of \funpow\\ lemma funpow_increasing: "m \ n \ mono f \ (f ^^ n) \ \ (f ^^ m) \" for f :: "'a::{lattice,order_top} \ 'a" by (induct rule: inc_induct) (auto simp del: funpow.simps(2) simp add: funpow_Suc_right intro: order_trans[OF _ funpow_mono]) lemma funpow_decreasing: "m \ n \ mono f \ (f ^^ m) \ \ (f ^^ n) \" for f :: "'a::{lattice,order_bot} \ 'a" by (induct rule: dec_induct) (auto simp del: funpow.simps(2) simp add: funpow_Suc_right intro: order_trans[OF _ funpow_mono]) lemma mono_funpow: "mono Q \ mono (\i. (Q ^^ i) \)" for Q :: "'a::{lattice,order_bot} \ 'a" by (auto intro!: funpow_decreasing simp: mono_def) lemma antimono_funpow: "mono Q \ antimono (\i. (Q ^^ i) \)" for Q :: "'a::{lattice,order_top} \ 'a" by (auto intro!: funpow_increasing simp: antimono_def) subsection \The divides relation on \<^typ>\nat\\ lemma dvd_1_left [iff]: "Suc 0 dvd k" by (simp add: dvd_def) lemma dvd_1_iff_1 [simp]: "m dvd Suc 0 \ m = Suc 0" by (simp add: dvd_def) lemma nat_dvd_1_iff_1 [simp]: "m dvd 1 \ m = 1" for m :: nat by (simp add: dvd_def) lemma dvd_antisym: "m dvd n \ n dvd m \ m = n" for m n :: nat unfolding dvd_def by (force dest: mult_eq_self_implies_10 simp add: mult.assoc) lemma dvd_diff_nat [simp]: "k dvd m \ k dvd n \ k dvd (m - n)" for k m n :: nat unfolding dvd_def by (blast intro: right_diff_distrib' [symmetric]) lemma dvd_diffD: fixes k m n :: nat assumes "k dvd m - n" "k dvd n" "n \ m" shows "k dvd m" proof - have "k dvd n + (m - n)" using assms by (blast intro: dvd_add) with assms show ?thesis by simp qed lemma dvd_diffD1: "k dvd m - n \ k dvd m \ n \ m \ k dvd n" for k m n :: nat by (drule_tac m = m in dvd_diff_nat) auto lemma dvd_mult_cancel: fixes m n k :: nat assumes "k * m dvd k * n" and "0 < k" shows "m dvd n" proof - from assms(1) obtain q where "k * n = (k * m) * q" .. then have "k * n = k * (m * q)" by (simp add: ac_simps) with \0 < k\ have "n = m * q" by (auto simp add: mult_left_cancel) then show ?thesis .. qed lemma dvd_mult_cancel1: fixes m n :: nat assumes "0 < m" shows "m * n dvd m \ n = 1" proof assume "m * n dvd m" then have "m * n dvd m * 1" by simp then have "n dvd 1" by (iprover intro: assms dvd_mult_cancel) then show "n = 1" by auto qed auto lemma dvd_mult_cancel2: "0 < m \ n * m dvd m \ n = 1" for m n :: nat using dvd_mult_cancel1 [of m n] by (simp add: ac_simps) lemma dvd_imp_le: "k dvd n \ 0 < n \ k \ n" for k n :: nat by (auto elim!: dvdE) (auto simp add: gr0_conv_Suc) lemma nat_dvd_not_less: "0 < m \ m < n \ \ n dvd m" for m n :: nat by (auto elim!: dvdE) (auto simp add: gr0_conv_Suc) lemma less_eq_dvd_minus: fixes m n :: nat assumes "m \ n" shows "m dvd n \ m dvd n - m" proof - from assms have "n = m + (n - m)" by simp then obtain q where "n = m + q" .. then show ?thesis by (simp add: add.commute [of m]) qed lemma dvd_minus_self: "m dvd n - m \ n < m \ m dvd n" for m n :: nat by (cases "n < m") (auto elim!: dvdE simp add: not_less le_imp_diff_is_add dest: less_imp_le) lemma dvd_minus_add: fixes m n q r :: nat assumes "q \ n" "q \ r * m" shows "m dvd n - q \ m dvd n + (r * m - q)" proof - have "m dvd n - q \ m dvd r * m + (n - q)" using dvd_add_times_triv_left_iff [of m r] by simp also from assms have "\ \ m dvd r * m + n - q" by simp also from assms have "\ \ m dvd (r * m - q) + n" by simp also have "\ \ m dvd n + (r * m - q)" by (simp add: add.commute) finally show ?thesis . qed subsection \Aliasses\ lemma nat_mult_1: "1 * n = n" for n :: nat by (fact mult_1_left) lemma nat_mult_1_right: "n * 1 = n" for n :: nat by (fact mult_1_right) lemma diff_mult_distrib: "(m - n) * k = (m * k) - (n * k)" for k m n :: nat by (fact left_diff_distrib') lemma diff_mult_distrib2: "k * (m - n) = (k * m) - (k * n)" for k m n :: nat by (fact right_diff_distrib') (*Used in AUTO2 and Groups.le_diff_conv2 (with variables renamed) doesn't work for some reason*) lemma le_diff_conv2: "k \ j \ (i \ j - k) = (i + k \ j)" for i j k :: nat by (fact le_diff_conv2) lemma diff_self_eq_0 [simp]: "m - m = 0" for m :: nat by (fact diff_cancel) lemma diff_diff_left [simp]: "i - j - k = i - (j + k)" for i j k :: nat by (fact diff_diff_add) lemma diff_commute: "i - j - k = i - k - j" for i j k :: nat by (fact diff_right_commute) lemma diff_add_inverse: "(n + m) - n = m" for m n :: nat by (fact add_diff_cancel_left') lemma diff_add_inverse2: "(m + n) - n = m" for m n :: nat by (fact add_diff_cancel_right') lemma diff_cancel: "(k + m) - (k + n) = m - n" for k m n :: nat by (fact add_diff_cancel_left) lemma diff_cancel2: "(m + k) - (n + k) = m - n" for k m n :: nat by (fact add_diff_cancel_right) lemma diff_add_0: "n - (n + m) = 0" for m n :: nat by (fact diff_add_zero) lemma add_mult_distrib2: "k * (m + n) = (k * m) + (k * n)" for k m n :: nat by (fact distrib_left) lemmas nat_distrib = add_mult_distrib distrib_left diff_mult_distrib diff_mult_distrib2 subsection \Size of a datatype value\ class size = fixes size :: "'a \ nat" \ \see further theory \Wellfounded\\ instantiation nat :: size begin definition size_nat where [simp, code]: "size (n::nat) = n" instance .. end lemmas size_nat = size_nat_def lemma size_neq_size_imp_neq: "size x \ size y \ x \ y" by (erule contrapos_nn) (rule arg_cong) subsection \Code module namespace\ code_identifier code_module Nat \ (SML) Arith and (OCaml) Arith and (Haskell) Arith hide_const (open) of_nat_aux end diff --git a/src/HOL/Orderings.thy b/src/HOL/Orderings.thy --- a/src/HOL/Orderings.thy +++ b/src/HOL/Orderings.thy @@ -1,1742 +1,1745 @@ (* Title: HOL/Orderings.thy Author: Tobias Nipkow, Markus Wenzel, and Larry Paulson *) section \Abstract orderings\ theory Orderings imports HOL keywords "print_orders" :: diag begin ML_file \~~/src/Provers/order.ML\ subsection \Abstract ordering\ locale ordering = fixes less_eq :: "'a \ 'a \ bool" (infix "\<^bold>\" 50) and less :: "'a \ 'a \ bool" (infix "\<^bold><" 50) assumes strict_iff_order: "a \<^bold>< b \ a \<^bold>\ b \ a \ b" assumes refl: "a \<^bold>\ a" \ \not \iff\: makes problems due to multiple (dual) interpretations\ and antisym: "a \<^bold>\ b \ b \<^bold>\ a \ a = b" and trans: "a \<^bold>\ b \ b \<^bold>\ c \ a \<^bold>\ c" begin lemma strict_implies_order: "a \<^bold>< b \ a \<^bold>\ b" by (simp add: strict_iff_order) lemma strict_implies_not_eq: "a \<^bold>< b \ a \ b" by (simp add: strict_iff_order) lemma not_eq_order_implies_strict: "a \ b \ a \<^bold>\ b \ a \<^bold>< b" by (simp add: strict_iff_order) lemma order_iff_strict: "a \<^bold>\ b \ a \<^bold>< b \ a = b" by (auto simp add: strict_iff_order refl) lemma irrefl: \ \not \iff\: makes problems due to multiple (dual) interpretations\ "\ a \<^bold>< a" by (simp add: strict_iff_order) lemma asym: "a \<^bold>< b \ b \<^bold>< a \ False" by (auto simp add: strict_iff_order intro: antisym) lemma strict_trans1: "a \<^bold>\ b \ b \<^bold>< c \ a \<^bold>< c" by (auto simp add: strict_iff_order intro: trans antisym) lemma strict_trans2: "a \<^bold>< b \ b \<^bold>\ c \ a \<^bold>< c" by (auto simp add: strict_iff_order intro: trans antisym) lemma strict_trans: "a \<^bold>< b \ b \<^bold>< c \ a \<^bold>< c" by (auto intro: strict_trans1 strict_implies_order) +lemma eq_iff: "a = b \ a \<^bold>\ b \ b \<^bold>\ a" + by (auto simp add: refl intro: antisym) + end text \Alternative introduction rule with bias towards strict order\ lemma ordering_strictI: fixes less_eq (infix "\<^bold>\" 50) and less (infix "\<^bold><" 50) assumes less_eq_less: "\a b. a \<^bold>\ b \ a \<^bold>< b \ a = b" assumes asym: "\a b. a \<^bold>< b \ \ b \<^bold>< a" assumes irrefl: "\a. \ a \<^bold>< a" assumes trans: "\a b c. a \<^bold>< b \ b \<^bold>< c \ a \<^bold>< c" shows "ordering less_eq less" proof fix a b show "a \<^bold>< b \ a \<^bold>\ b \ a \ b" by (auto simp add: less_eq_less asym irrefl) next fix a show "a \<^bold>\ a" by (auto simp add: less_eq_less) next fix a b c assume "a \<^bold>\ b" and "b \<^bold>\ c" then show "a \<^bold>\ c" by (auto simp add: less_eq_less intro: trans) next fix a b assume "a \<^bold>\ b" and "b \<^bold>\ a" then show "a = b" by (auto simp add: less_eq_less asym) qed lemma ordering_dualI: fixes less_eq (infix "\<^bold>\" 50) and less (infix "\<^bold><" 50) assumes "ordering (\a b. b \<^bold>\ a) (\a b. b \<^bold>< a)" shows "ordering less_eq less" proof - from assms interpret ordering "\a b. b \<^bold>\ a" "\a b. b \<^bold>< a" . show ?thesis by standard (auto simp: strict_iff_order refl intro: antisym trans) qed locale ordering_top = ordering + fixes top :: "'a" ("\<^bold>\") assumes extremum [simp]: "a \<^bold>\ \<^bold>\" begin lemma extremum_uniqueI: "\<^bold>\ \<^bold>\ a \ a = \<^bold>\" by (rule antisym) auto lemma extremum_unique: "\<^bold>\ \<^bold>\ a \ a = \<^bold>\" by (auto intro: antisym) lemma extremum_strict [simp]: "\ (\<^bold>\ \<^bold>< a)" using extremum [of a] by (auto simp add: order_iff_strict intro: asym irrefl) lemma not_eq_extremum: "a \ \<^bold>\ \ a \<^bold>< \<^bold>\" by (auto simp add: order_iff_strict intro: not_eq_order_implies_strict extremum) end subsection \Syntactic orders\ class ord = fixes less_eq :: "'a \ 'a \ bool" and less :: "'a \ 'a \ bool" begin notation less_eq ("'(\')") and less_eq ("(_/ \ _)" [51, 51] 50) and less ("'(<')") and less ("(_/ < _)" [51, 51] 50) abbreviation (input) greater_eq (infix "\" 50) where "x \ y \ y \ x" abbreviation (input) greater (infix ">" 50) where "x > y \ y < x" notation (ASCII) less_eq ("'(<=')") and less_eq ("(_/ <= _)" [51, 51] 50) notation (input) greater_eq (infix ">=" 50) end subsection \Quasi orders\ class preorder = ord + assumes less_le_not_le: "x < y \ x \ y \ \ (y \ x)" and order_refl [iff]: "x \ x" and order_trans: "x \ y \ y \ z \ x \ z" begin text \Reflexivity.\ lemma eq_refl: "x = y \ x \ y" \ \This form is useful with the classical reasoner.\ by (erule ssubst) (rule order_refl) lemma less_irrefl [iff]: "\ x < x" by (simp add: less_le_not_le) lemma less_imp_le: "x < y \ x \ y" by (simp add: less_le_not_le) text \Asymmetry.\ lemma less_not_sym: "x < y \ \ (y < x)" by (simp add: less_le_not_le) lemma less_asym: "x < y \ (\ P \ y < x) \ P" by (drule less_not_sym, erule contrapos_np) simp text \Transitivity.\ lemma less_trans: "x < y \ y < z \ x < z" by (auto simp add: less_le_not_le intro: order_trans) lemma le_less_trans: "x \ y \ y < z \ x < z" by (auto simp add: less_le_not_le intro: order_trans) lemma less_le_trans: "x < y \ y \ z \ x < z" by (auto simp add: less_le_not_le intro: order_trans) text \Useful for simplification, but too risky to include by default.\ lemma less_imp_not_less: "x < y \ (\ y < x) \ True" by (blast elim: less_asym) lemma less_imp_triv: "x < y \ (y < x \ P) \ True" by (blast elim: less_asym) text \Transitivity rules for calculational reasoning\ lemma less_asym': "a < b \ b < a \ P" by (rule less_asym) text \Dual order\ lemma dual_preorder: "class.preorder (\) (>)" by standard (auto simp add: less_le_not_le intro: order_trans) end subsection \Partial orders\ class order = preorder + assumes antisym: "x \ y \ y \ x \ x = y" begin lemma less_le: "x < y \ x \ y \ x \ y" by (auto simp add: less_le_not_le intro: antisym) sublocale order: ordering less_eq less + dual_order: ordering greater_eq greater proof - interpret ordering less_eq less by standard (auto intro: antisym order_trans simp add: less_le) show "ordering less_eq less" by (fact ordering_axioms) then show "ordering greater_eq greater" by (rule ordering_dualI) qed text \Reflexivity.\ lemma le_less: "x \ y \ x < y \ x = y" \ \NOT suitable for iff, since it can cause PROOF FAILED.\ by (fact order.order_iff_strict) lemma le_imp_less_or_eq: "x \ y \ x < y \ x = y" by (simp add: less_le) text \Useful for simplification, but too risky to include by default.\ lemma less_imp_not_eq: "x < y \ (x = y) \ False" by auto lemma less_imp_not_eq2: "x < y \ (y = x) \ False" by auto text \Transitivity rules for calculational reasoning\ lemma neq_le_trans: "a \ b \ a \ b \ a < b" by (fact order.not_eq_order_implies_strict) lemma le_neq_trans: "a \ b \ a \ b \ a < b" by (rule order.not_eq_order_implies_strict) text \Asymmetry.\ lemma eq_iff: "x = y \ x \ y \ y \ x" -by (blast intro: antisym) + by (fact order.eq_iff) lemma antisym_conv: "y \ x \ x \ y \ x = y" -by (blast intro: antisym) + by (simp add: eq_iff) lemma less_imp_neq: "x < y \ x \ y" -by (fact order.strict_implies_not_eq) + by (fact order.strict_implies_not_eq) lemma antisym_conv1: "\ x < y \ x \ y \ x = y" by (simp add: local.le_less) lemma antisym_conv2: "x \ y \ \ x < y \ x = y" by (simp add: local.less_le) lemma leD: "y \ x \ \ x < y" by (auto simp: less_le antisym) text \Least value operator\ definition (in ord) Least :: "('a \ bool) \ 'a" (binder "LEAST " 10) where "Least P = (THE x. P x \ (\y. P y \ x \ y))" lemma Least_equality: assumes "P x" and "\y. P y \ x \ y" shows "Least P = x" unfolding Least_def by (rule the_equality) (blast intro: assms antisym)+ lemma LeastI2_order: assumes "P x" and "\y. P y \ x \ y" and "\x. P x \ \y. P y \ x \ y \ Q x" shows "Q (Least P)" unfolding Least_def by (rule theI2) (blast intro: assms antisym)+ lemma Least_ex1: assumes "\!x. P x \ (\y. P y \ x \ y)" shows Least1I: "P (Least P)" and Least1_le: "P z \ Least P \ z" using theI'[OF assms] unfolding Least_def by auto text \Greatest value operator\ definition Greatest :: "('a \ bool) \ 'a" (binder "GREATEST " 10) where "Greatest P = (THE x. P x \ (\y. P y \ x \ y))" lemma GreatestI2_order: "\ P x; \y. P y \ x \ y; \x. \ P x; \y. P y \ x \ y \ \ Q x \ \ Q (Greatest P)" unfolding Greatest_def by (rule theI2) (blast intro: antisym)+ lemma Greatest_equality: "\ P x; \y. P y \ x \ y \ \ Greatest P = x" unfolding Greatest_def by (rule the_equality) (blast intro: antisym)+ end lemma ordering_orderI: fixes less_eq (infix "\<^bold>\" 50) and less (infix "\<^bold><" 50) assumes "ordering less_eq less" shows "class.order less_eq less" proof - from assms interpret ordering less_eq less . show ?thesis by standard (auto intro: antisym trans simp add: refl strict_iff_order) qed lemma order_strictI: fixes less (infix "\" 50) and less_eq (infix "\" 50) assumes "\a b. a \ b \ a \ b \ a = b" assumes "\a b. a \ b \ \ b \ a" assumes "\a. \ a \ a" assumes "\a b c. a \ b \ b \ c \ a \ c" shows "class.order less_eq less" by (rule ordering_orderI) (rule ordering_strictI, (fact assms)+) context order begin text \Dual order\ lemma dual_order: "class.order (\) (>)" using dual_order.ordering_axioms by (rule ordering_orderI) end subsection \Linear (total) orders\ class linorder = order + assumes linear: "x \ y \ y \ x" begin lemma less_linear: "x < y \ x = y \ y < x" unfolding less_le using less_le linear by blast lemma le_less_linear: "x \ y \ y < x" by (simp add: le_less less_linear) lemma le_cases [case_names le ge]: "(x \ y \ P) \ (y \ x \ P) \ P" using linear by blast lemma (in linorder) le_cases3: "\\x \ y; y \ z\ \ P; \y \ x; x \ z\ \ P; \x \ z; z \ y\ \ P; \z \ y; y \ x\ \ P; \y \ z; z \ x\ \ P; \z \ x; x \ y\ \ P\ \ P" by (blast intro: le_cases) lemma linorder_cases [case_names less equal greater]: "(x < y \ P) \ (x = y \ P) \ (y < x \ P) \ P" using less_linear by blast lemma linorder_wlog[case_names le sym]: "(\a b. a \ b \ P a b) \ (\a b. P b a \ P a b) \ P a b" by (cases rule: le_cases[of a b]) blast+ lemma not_less: "\ x < y \ y \ x" unfolding less_le using linear by (blast intro: antisym) lemma not_less_iff_gr_or_eq: "\(x < y) \ (x > y \ x = y)" by (auto simp add:not_less le_less) lemma not_le: "\ x \ y \ y < x" unfolding less_le using linear by (blast intro: antisym) lemma neq_iff: "x \ y \ x < y \ y < x" by (cut_tac x = x and y = y in less_linear, auto) lemma neqE: "x \ y \ (x < y \ R) \ (y < x \ R) \ R" by (simp add: neq_iff) blast lemma antisym_conv3: "\ y < x \ \ x < y \ x = y" by (blast intro: antisym dest: not_less [THEN iffD1]) lemma leI: "\ x < y \ y \ x" unfolding not_less . lemma not_le_imp_less: "\ y \ x \ x < y" unfolding not_le . lemma linorder_less_wlog[case_names less refl sym]: "\\a b. a < b \ P a b; \a. P a a; \a b. P b a \ P a b\ \ P a b" using antisym_conv3 by blast text \Dual order\ lemma dual_linorder: "class.linorder (\) (>)" by (rule class.linorder.intro, rule dual_order) (unfold_locales, rule linear) end text \Alternative introduction rule with bias towards strict order\ lemma linorder_strictI: fixes less_eq (infix "\<^bold>\" 50) and less (infix "\<^bold><" 50) assumes "class.order less_eq less" assumes trichotomy: "\a b. a \<^bold>< b \ a = b \ b \<^bold>< a" shows "class.linorder less_eq less" proof - interpret order less_eq less by (fact \class.order less_eq less\) show ?thesis proof fix a b show "a \<^bold>\ b \ b \<^bold>\ a" using trichotomy by (auto simp add: le_less) qed qed subsection \Reasoning tools setup\ ML \ signature ORDERS = sig val print_structures: Proof.context -> unit val order_tac: Proof.context -> thm list -> int -> tactic val add_struct: string * term list -> string -> attribute val del_struct: string * term list -> attribute end; structure Orders: ORDERS = struct (* context data *) fun struct_eq ((s1: string, ts1), (s2, ts2)) = s1 = s2 andalso eq_list (op aconv) (ts1, ts2); structure Data = Generic_Data ( type T = ((string * term list) * Order_Tac.less_arith) list; (* Order structures: identifier of the structure, list of operations and record of theorems needed to set up the transitivity reasoner, identifier and operations identify the structure uniquely. *) val empty = []; val extend = I; fun merge data = AList.join struct_eq (K fst) data; ); fun print_structures ctxt = let val structs = Data.get (Context.Proof ctxt); fun pretty_term t = Pretty.block [Pretty.quote (Syntax.pretty_term ctxt t), Pretty.brk 1, Pretty.str "::", Pretty.brk 1, Pretty.quote (Syntax.pretty_typ ctxt (type_of t))]; fun pretty_struct ((s, ts), _) = Pretty.block [Pretty.str s, Pretty.str ":", Pretty.brk 1, Pretty.enclose "(" ")" (Pretty.breaks (map pretty_term ts))]; in Pretty.writeln (Pretty.big_list "order structures:" (map pretty_struct structs)) end; val _ = Outer_Syntax.command \<^command_keyword>\print_orders\ "print order structures available to transitivity reasoner" (Scan.succeed (Toplevel.keep (print_structures o Toplevel.context_of))); (* tactics *) fun struct_tac ((s, ops), thms) ctxt facts = let val [eq, le, less] = ops; fun decomp thy (\<^const>\Trueprop\ $ t) = let fun excluded t = (* exclude numeric types: linear arithmetic subsumes transitivity *) let val T = type_of t in T = HOLogic.natT orelse T = HOLogic.intT orelse T = HOLogic.realT end; fun rel (bin_op $ t1 $ t2) = if excluded t1 then NONE else if Pattern.matches thy (eq, bin_op) then SOME (t1, "=", t2) else if Pattern.matches thy (le, bin_op) then SOME (t1, "<=", t2) else if Pattern.matches thy (less, bin_op) then SOME (t1, "<", t2) else NONE | rel _ = NONE; fun dec (Const (\<^const_name>\Not\, _) $ t) = (case rel t of NONE => NONE | SOME (t1, rel, t2) => SOME (t1, "~" ^ rel, t2)) | dec x = rel x; in dec t end | decomp _ _ = NONE; in (case s of "order" => Order_Tac.partial_tac decomp thms ctxt facts | "linorder" => Order_Tac.linear_tac decomp thms ctxt facts | _ => error ("Unknown order kind " ^ quote s ^ " encountered in transitivity reasoner")) end fun order_tac ctxt facts = FIRST' (map (fn s => CHANGED o struct_tac s ctxt facts) (Data.get (Context.Proof ctxt))); (* attributes *) fun add_struct s tag = Thm.declaration_attribute (fn thm => Data.map (AList.map_default struct_eq (s, Order_Tac.empty TrueI) (Order_Tac.update tag thm))); fun del_struct s = Thm.declaration_attribute (fn _ => Data.map (AList.delete struct_eq s)); end; \ attribute_setup order = \ Scan.lift ((Args.add -- Args.name >> (fn (_, s) => SOME s) || Args.del >> K NONE) --| Args.colon (* FIXME || Scan.succeed true *) ) -- Scan.lift Args.name -- Scan.repeat Args.term >> (fn ((SOME tag, n), ts) => Orders.add_struct (n, ts) tag | ((NONE, n), ts) => Orders.del_struct (n, ts)) \ "theorems controlling transitivity reasoner" method_setup order = \ Scan.succeed (fn ctxt => SIMPLE_METHOD' (Orders.order_tac ctxt [])) \ "transitivity reasoner" text \Declarations to set up transitivity reasoner of partial and linear orders.\ context order begin (* The type constraint on @{term (=}) below is necessary since the operation is not a parameter of the locale. *) declare less_irrefl [THEN notE, order add less_reflE: order "(=) :: 'a \ 'a \ bool" "(<=)" "(<)"] declare order_refl [order add le_refl: order "(=) :: 'a => 'a => bool" "(<=)" "(<)"] declare less_imp_le [order add less_imp_le: order "(=) :: 'a => 'a => bool" "(<=)" "(<)"] declare antisym [order add eqI: order "(=) :: 'a => 'a => bool" "(<=)" "(<)"] declare eq_refl [order add eqD1: order "(=) :: 'a => 'a => bool" "(<=)" "(<)"] declare sym [THEN eq_refl, order add eqD2: order "(=) :: 'a => 'a => bool" "(<=)" "(<)"] declare less_trans [order add less_trans: order "(=) :: 'a => 'a => bool" "(<=)" "(<)"] declare less_le_trans [order add less_le_trans: order "(=) :: 'a => 'a => bool" "(<=)" "(<)"] declare le_less_trans [order add le_less_trans: order "(=) :: 'a => 'a => bool" "(<=)" "(<)"] declare order_trans [order add le_trans: order "(=) :: 'a => 'a => bool" "(<=)" "(<)"] declare le_neq_trans [order add le_neq_trans: order "(=) :: 'a => 'a => bool" "(<=)" "(<)"] declare neq_le_trans [order add neq_le_trans: order "(=) :: 'a => 'a => bool" "(<=)" "(<)"] declare less_imp_neq [order add less_imp_neq: order "(=) :: 'a => 'a => bool" "(<=)" "(<)"] declare eq_neq_eq_imp_neq [order add eq_neq_eq_imp_neq: order "(=) :: 'a => 'a => bool" "(<=)" "(<)"] declare not_sym [order add not_sym: order "(=) :: 'a => 'a => bool" "(<=)" "(<)"] end context linorder begin declare [[order del: order "(=) :: 'a => 'a => bool" "(<=)" "(<)"]] declare less_irrefl [THEN notE, order add less_reflE: linorder "(=) :: 'a => 'a => bool" "(<=)" "(<)"] declare order_refl [order add le_refl: linorder "(=) :: 'a => 'a => bool" "(<=)" "(<)"] declare less_imp_le [order add less_imp_le: linorder "(=) :: 'a => 'a => bool" "(<=)" "(<)"] declare not_less [THEN iffD2, order add not_lessI: linorder "(=) :: 'a => 'a => bool" "(<=)" "(<)"] declare not_le [THEN iffD2, order add not_leI: linorder "(=) :: 'a => 'a => bool" "(<=)" "(<)"] declare not_less [THEN iffD1, order add not_lessD: linorder "(=) :: 'a => 'a => bool" "(<=)" "(<)"] declare not_le [THEN iffD1, order add not_leD: linorder "(=) :: 'a => 'a => bool" "(<=)" "(<)"] declare antisym [order add eqI: linorder "(=) :: 'a => 'a => bool" "(<=)" "(<)"] declare eq_refl [order add eqD1: linorder "(=) :: 'a => 'a => bool" "(<=)" "(<)"] declare sym [THEN eq_refl, order add eqD2: linorder "(=) :: 'a => 'a => bool" "(<=)" "(<)"] declare less_trans [order add less_trans: linorder "(=) :: 'a => 'a => bool" "(<=)" "(<)"] declare less_le_trans [order add less_le_trans: linorder "(=) :: 'a => 'a => bool" "(<=)" "(<)"] declare le_less_trans [order add le_less_trans: linorder "(=) :: 'a => 'a => bool" "(<=)" "(<)"] declare order_trans [order add le_trans: linorder "(=) :: 'a => 'a => bool" "(<=)" "(<)"] declare le_neq_trans [order add le_neq_trans: linorder "(=) :: 'a => 'a => bool" "(<=)" "(<)"] declare neq_le_trans [order add neq_le_trans: linorder "(=) :: 'a => 'a => bool" "(<=)" "(<)"] declare less_imp_neq [order add less_imp_neq: linorder "(=) :: 'a => 'a => bool" "(<=)" "(<)"] declare eq_neq_eq_imp_neq [order add eq_neq_eq_imp_neq: linorder "(=) :: 'a => 'a => bool" "(<=)" "(<)"] declare not_sym [order add not_sym: linorder "(=) :: 'a => 'a => bool" "(<=)" "(<)"] end setup \ map_theory_simpset (fn ctxt0 => ctxt0 addSolver mk_solver "Transitivity" (fn ctxt => Orders.order_tac ctxt (Simplifier.prems_of ctxt))) (*Adding the transitivity reasoners also as safe solvers showed a slight speed up, but the reasoning strength appears to be not higher (at least no breaking of additional proofs in the entire HOL distribution, as of 5 March 2004, was observed).*) \ ML \ local fun prp t thm = Thm.prop_of thm = t; (* FIXME proper aconv!? *) in fun antisym_le_simproc ctxt ct = (case Thm.term_of ct of (le as Const (_, T)) $ r $ s => (let val prems = Simplifier.prems_of ctxt; val less = Const (\<^const_name>\less\, T); val t = HOLogic.mk_Trueprop(le $ s $ r); in (case find_first (prp t) prems of NONE => let val t = HOLogic.mk_Trueprop(HOLogic.Not $ (less $ r $ s)) in (case find_first (prp t) prems of NONE => NONE | SOME thm => SOME(mk_meta_eq(thm RS @{thm antisym_conv1}))) end | SOME thm => SOME (mk_meta_eq (thm RS @{thm order_class.antisym_conv}))) end handle THM _ => NONE) | _ => NONE); fun antisym_less_simproc ctxt ct = (case Thm.term_of ct of NotC $ ((less as Const(_,T)) $ r $ s) => (let val prems = Simplifier.prems_of ctxt; val le = Const (\<^const_name>\less_eq\, T); val t = HOLogic.mk_Trueprop(le $ r $ s); in (case find_first (prp t) prems of NONE => let val t = HOLogic.mk_Trueprop (NotC $ (less $ s $ r)) in (case find_first (prp t) prems of NONE => NONE | SOME thm => SOME (mk_meta_eq(thm RS @{thm linorder_class.antisym_conv3}))) end | SOME thm => SOME (mk_meta_eq (thm RS @{thm antisym_conv2}))) end handle THM _ => NONE) | _ => NONE); end; \ simproc_setup antisym_le ("(x::'a::order) \ y") = "K antisym_le_simproc" simproc_setup antisym_less ("\ (x::'a::linorder) < y") = "K antisym_less_simproc" subsection \Bounded quantifiers\ syntax (ASCII) "_All_less" :: "[idt, 'a, bool] => bool" ("(3ALL _<_./ _)" [0, 0, 10] 10) "_Ex_less" :: "[idt, 'a, bool] => bool" ("(3EX _<_./ _)" [0, 0, 10] 10) "_All_less_eq" :: "[idt, 'a, bool] => bool" ("(3ALL _<=_./ _)" [0, 0, 10] 10) "_Ex_less_eq" :: "[idt, 'a, bool] => bool" ("(3EX _<=_./ _)" [0, 0, 10] 10) "_All_greater" :: "[idt, 'a, bool] => bool" ("(3ALL _>_./ _)" [0, 0, 10] 10) "_Ex_greater" :: "[idt, 'a, bool] => bool" ("(3EX _>_./ _)" [0, 0, 10] 10) "_All_greater_eq" :: "[idt, 'a, bool] => bool" ("(3ALL _>=_./ _)" [0, 0, 10] 10) "_Ex_greater_eq" :: "[idt, 'a, bool] => bool" ("(3EX _>=_./ _)" [0, 0, 10] 10) "_All_neq" :: "[idt, 'a, bool] => bool" ("(3ALL _~=_./ _)" [0, 0, 10] 10) "_Ex_neq" :: "[idt, 'a, bool] => bool" ("(3EX _~=_./ _)" [0, 0, 10] 10) syntax "_All_less" :: "[idt, 'a, bool] => bool" ("(3\_<_./ _)" [0, 0, 10] 10) "_Ex_less" :: "[idt, 'a, bool] => bool" ("(3\_<_./ _)" [0, 0, 10] 10) "_All_less_eq" :: "[idt, 'a, bool] => bool" ("(3\_\_./ _)" [0, 0, 10] 10) "_Ex_less_eq" :: "[idt, 'a, bool] => bool" ("(3\_\_./ _)" [0, 0, 10] 10) "_All_greater" :: "[idt, 'a, bool] => bool" ("(3\_>_./ _)" [0, 0, 10] 10) "_Ex_greater" :: "[idt, 'a, bool] => bool" ("(3\_>_./ _)" [0, 0, 10] 10) "_All_greater_eq" :: "[idt, 'a, bool] => bool" ("(3\_\_./ _)" [0, 0, 10] 10) "_Ex_greater_eq" :: "[idt, 'a, bool] => bool" ("(3\_\_./ _)" [0, 0, 10] 10) "_All_neq" :: "[idt, 'a, bool] => bool" ("(3\_\_./ _)" [0, 0, 10] 10) "_Ex_neq" :: "[idt, 'a, bool] => bool" ("(3\_\_./ _)" [0, 0, 10] 10) syntax (input) "_All_less" :: "[idt, 'a, bool] => bool" ("(3! _<_./ _)" [0, 0, 10] 10) "_Ex_less" :: "[idt, 'a, bool] => bool" ("(3? _<_./ _)" [0, 0, 10] 10) "_All_less_eq" :: "[idt, 'a, bool] => bool" ("(3! _<=_./ _)" [0, 0, 10] 10) "_Ex_less_eq" :: "[idt, 'a, bool] => bool" ("(3? _<=_./ _)" [0, 0, 10] 10) "_All_neq" :: "[idt, 'a, bool] => bool" ("(3! _~=_./ _)" [0, 0, 10] 10) "_Ex_neq" :: "[idt, 'a, bool] => bool" ("(3? _~=_./ _)" [0, 0, 10] 10) translations "\x "\x. x < y \ P" "\x "\x. x < y \ P" "\x\y. P" \ "\x. x \ y \ P" "\x\y. P" \ "\x. x \ y \ P" "\x>y. P" \ "\x. x > y \ P" "\x>y. P" \ "\x. x > y \ P" "\x\y. P" \ "\x. x \ y \ P" "\x\y. P" \ "\x. x \ y \ P" "\x\y. P" \ "\x. x \ y \ P" "\x\y. P" \ "\x. x \ y \ P" print_translation \ let val All_binder = Mixfix.binder_name \<^const_syntax>\All\; val Ex_binder = Mixfix.binder_name \<^const_syntax>\Ex\; val impl = \<^const_syntax>\HOL.implies\; val conj = \<^const_syntax>\HOL.conj\; val less = \<^const_syntax>\less\; val less_eq = \<^const_syntax>\less_eq\; val trans = [((All_binder, impl, less), (\<^syntax_const>\_All_less\, \<^syntax_const>\_All_greater\)), ((All_binder, impl, less_eq), (\<^syntax_const>\_All_less_eq\, \<^syntax_const>\_All_greater_eq\)), ((Ex_binder, conj, less), (\<^syntax_const>\_Ex_less\, \<^syntax_const>\_Ex_greater\)), ((Ex_binder, conj, less_eq), (\<^syntax_const>\_Ex_less_eq\, \<^syntax_const>\_Ex_greater_eq\))]; fun matches_bound v t = (case t of Const (\<^syntax_const>\_bound\, _) $ Free (v', _) => v = v' | _ => false); fun contains_var v = Term.exists_subterm (fn Free (x, _) => x = v | _ => false); fun mk x c n P = Syntax.const c $ Syntax_Trans.mark_bound_body x $ n $ P; fun tr' q = (q, fn _ => (fn [Const (\<^syntax_const>\_bound\, _) $ Free (v, T), Const (c, _) $ (Const (d, _) $ t $ u) $ P] => (case AList.lookup (=) trans (q, c, d) of NONE => raise Match | SOME (l, g) => if matches_bound v t andalso not (contains_var v u) then mk (v, T) l u P else if matches_bound v u andalso not (contains_var v t) then mk (v, T) g t P else raise Match) | _ => raise Match)); in [tr' All_binder, tr' Ex_binder] end \ subsection \Transitivity reasoning\ context ord begin lemma ord_le_eq_trans: "a \ b \ b = c \ a \ c" by (rule subst) lemma ord_eq_le_trans: "a = b \ b \ c \ a \ c" by (rule ssubst) lemma ord_less_eq_trans: "a < b \ b = c \ a < c" by (rule subst) lemma ord_eq_less_trans: "a = b \ b < c \ a < c" by (rule ssubst) end lemma order_less_subst2: "(a::'a::order) < b ==> f b < (c::'c::order) ==> (!!x y. x < y ==> f x < f y) ==> f a < c" proof - assume r: "!!x y. x < y ==> f x < f y" assume "a < b" hence "f a < f b" by (rule r) also assume "f b < c" finally (less_trans) show ?thesis . qed lemma order_less_subst1: "(a::'a::order) < f b ==> (b::'b::order) < c ==> (!!x y. x < y ==> f x < f y) ==> a < f c" proof - assume r: "!!x y. x < y ==> f x < f y" assume "a < f b" also assume "b < c" hence "f b < f c" by (rule r) finally (less_trans) show ?thesis . qed lemma order_le_less_subst2: "(a::'a::order) <= b ==> f b < (c::'c::order) ==> (!!x y. x <= y ==> f x <= f y) ==> f a < c" proof - assume r: "!!x y. x <= y ==> f x <= f y" assume "a <= b" hence "f a <= f b" by (rule r) also assume "f b < c" finally (le_less_trans) show ?thesis . qed lemma order_le_less_subst1: "(a::'a::order) <= f b ==> (b::'b::order) < c ==> (!!x y. x < y ==> f x < f y) ==> a < f c" proof - assume r: "!!x y. x < y ==> f x < f y" assume "a <= f b" also assume "b < c" hence "f b < f c" by (rule r) finally (le_less_trans) show ?thesis . qed lemma order_less_le_subst2: "(a::'a::order) < b ==> f b <= (c::'c::order) ==> (!!x y. x < y ==> f x < f y) ==> f a < c" proof - assume r: "!!x y. x < y ==> f x < f y" assume "a < b" hence "f a < f b" by (rule r) also assume "f b <= c" finally (less_le_trans) show ?thesis . qed lemma order_less_le_subst1: "(a::'a::order) < f b ==> (b::'b::order) <= c ==> (!!x y. x <= y ==> f x <= f y) ==> a < f c" proof - assume r: "!!x y. x <= y ==> f x <= f y" assume "a < f b" also assume "b <= c" hence "f b <= f c" by (rule r) finally (less_le_trans) show ?thesis . qed lemma order_subst1: "(a::'a::order) <= f b ==> (b::'b::order) <= c ==> (!!x y. x <= y ==> f x <= f y) ==> a <= f c" proof - assume r: "!!x y. x <= y ==> f x <= f y" assume "a <= f b" also assume "b <= c" hence "f b <= f c" by (rule r) finally (order_trans) show ?thesis . qed lemma order_subst2: "(a::'a::order) <= b ==> f b <= (c::'c::order) ==> (!!x y. x <= y ==> f x <= f y) ==> f a <= c" proof - assume r: "!!x y. x <= y ==> f x <= f y" assume "a <= b" hence "f a <= f b" by (rule r) also assume "f b <= c" finally (order_trans) show ?thesis . qed lemma ord_le_eq_subst: "a <= b ==> f b = c ==> (!!x y. x <= y ==> f x <= f y) ==> f a <= c" proof - assume r: "!!x y. x <= y ==> f x <= f y" assume "a <= b" hence "f a <= f b" by (rule r) also assume "f b = c" finally (ord_le_eq_trans) show ?thesis . qed lemma ord_eq_le_subst: "a = f b ==> b <= c ==> (!!x y. x <= y ==> f x <= f y) ==> a <= f c" proof - assume r: "!!x y. x <= y ==> f x <= f y" assume "a = f b" also assume "b <= c" hence "f b <= f c" by (rule r) finally (ord_eq_le_trans) show ?thesis . qed lemma ord_less_eq_subst: "a < b ==> f b = c ==> (!!x y. x < y ==> f x < f y) ==> f a < c" proof - assume r: "!!x y. x < y ==> f x < f y" assume "a < b" hence "f a < f b" by (rule r) also assume "f b = c" finally (ord_less_eq_trans) show ?thesis . qed lemma ord_eq_less_subst: "a = f b ==> b < c ==> (!!x y. x < y ==> f x < f y) ==> a < f c" proof - assume r: "!!x y. x < y ==> f x < f y" assume "a = f b" also assume "b < c" hence "f b < f c" by (rule r) finally (ord_eq_less_trans) show ?thesis . qed text \ Note that this list of rules is in reverse order of priorities. \ lemmas [trans] = order_less_subst2 order_less_subst1 order_le_less_subst2 order_le_less_subst1 order_less_le_subst2 order_less_le_subst1 order_subst2 order_subst1 ord_le_eq_subst ord_eq_le_subst ord_less_eq_subst ord_eq_less_subst forw_subst back_subst rev_mp mp lemmas (in order) [trans] = neq_le_trans le_neq_trans lemmas (in preorder) [trans] = less_trans less_asym' le_less_trans less_le_trans order_trans lemmas (in order) [trans] = antisym lemmas (in ord) [trans] = ord_le_eq_trans ord_eq_le_trans ord_less_eq_trans ord_eq_less_trans lemmas [trans] = trans lemmas order_trans_rules = order_less_subst2 order_less_subst1 order_le_less_subst2 order_le_less_subst1 order_less_le_subst2 order_less_le_subst1 order_subst2 order_subst1 ord_le_eq_subst ord_eq_le_subst ord_less_eq_subst ord_eq_less_subst forw_subst back_subst rev_mp mp neq_le_trans le_neq_trans less_trans less_asym' le_less_trans less_le_trans order_trans antisym ord_le_eq_trans ord_eq_le_trans ord_less_eq_trans ord_eq_less_trans trans text \These support proving chains of decreasing inequalities a >= b >= c ... in Isar proofs.\ lemma xt1 [no_atp]: "a = b \ b > c \ a > c" "a > b \ b = c \ a > c" "a = b \ b \ c \ a \ c" "a \ b \ b = c \ a \ c" "(x::'a::order) \ y \ y \ x \ x = y" "(x::'a::order) \ y \ y \ z \ x \ z" "(x::'a::order) > y \ y \ z \ x > z" "(x::'a::order) \ y \ y > z \ x > z" "(a::'a::order) > b \ b > a \ P" "(x::'a::order) > y \ y > z \ x > z" "(a::'a::order) \ b \ a \ b \ a > b" "(a::'a::order) \ b \ a \ b \ a > b" "a = f b \ b > c \ (\x y. x > y \ f x > f y) \ a > f c" "a > b \ f b = c \ (\x y. x > y \ f x > f y) \ f a > c" "a = f b \ b \ c \ (\x y. x \ y \ f x \ f y) \ a \ f c" "a \ b \ f b = c \ (\x y. x \ y \ f x \ f y) \ f a \ c" by auto lemma xt2 [no_atp]: "(a::'a::order) >= f b ==> b >= c ==> (!!x y. x >= y ==> f x >= f y) ==> a >= f c" by (subgoal_tac "f b >= f c", force, force) lemma xt3 [no_atp]: "(a::'a::order) >= b ==> (f b::'b::order) >= c ==> (!!x y. x >= y ==> f x >= f y) ==> f a >= c" by (subgoal_tac "f a >= f b", force, force) lemma xt4 [no_atp]: "(a::'a::order) > f b ==> (b::'b::order) >= c ==> (!!x y. x >= y ==> f x >= f y) ==> a > f c" by (subgoal_tac "f b >= f c", force, force) lemma xt5 [no_atp]: "(a::'a::order) > b ==> (f b::'b::order) >= c==> (!!x y. x > y ==> f x > f y) ==> f a > c" by (subgoal_tac "f a > f b", force, force) lemma xt6 [no_atp]: "(a::'a::order) >= f b ==> b > c ==> (!!x y. x > y ==> f x > f y) ==> a > f c" by (subgoal_tac "f b > f c", force, force) lemma xt7 [no_atp]: "(a::'a::order) >= b ==> (f b::'b::order) > c ==> (!!x y. x >= y ==> f x >= f y) ==> f a > c" by (subgoal_tac "f a >= f b", force, force) lemma xt8 [no_atp]: "(a::'a::order) > f b ==> (b::'b::order) > c ==> (!!x y. x > y ==> f x > f y) ==> a > f c" by (subgoal_tac "f b > f c", force, force) lemma xt9 [no_atp]: "(a::'a::order) > b ==> (f b::'b::order) > c ==> (!!x y. x > y ==> f x > f y) ==> f a > c" by (subgoal_tac "f a > f b", force, force) lemmas xtrans = xt1 xt2 xt3 xt4 xt5 xt6 xt7 xt8 xt9 (* Since "a >= b" abbreviates "b <= a", the abbreviation "..." stands for the wrong thing in an Isar proof. The extra transitivity rules can be used as follows: lemma "(a::'a::order) > z" proof - have "a >= b" (is "_ >= ?rhs") sorry also have "?rhs >= c" (is "_ >= ?rhs") sorry also (xtrans) have "?rhs = d" (is "_ = ?rhs") sorry also (xtrans) have "?rhs >= e" (is "_ >= ?rhs") sorry also (xtrans) have "?rhs > f" (is "_ > ?rhs") sorry also (xtrans) have "?rhs > z" sorry finally (xtrans) show ?thesis . qed Alternatively, one can use "declare xtrans [trans]" and then leave out the "(xtrans)" above. *) subsection \Monotonicity\ context order begin definition mono :: "('a \ 'b::order) \ bool" where "mono f \ (\x y. x \ y \ f x \ f y)" lemma monoI [intro?]: fixes f :: "'a \ 'b::order" shows "(\x y. x \ y \ f x \ f y) \ mono f" unfolding mono_def by iprover lemma monoD [dest?]: fixes f :: "'a \ 'b::order" shows "mono f \ x \ y \ f x \ f y" unfolding mono_def by iprover lemma monoE: fixes f :: "'a \ 'b::order" assumes "mono f" assumes "x \ y" obtains "f x \ f y" proof from assms show "f x \ f y" by (simp add: mono_def) qed definition antimono :: "('a \ 'b::order) \ bool" where "antimono f \ (\x y. x \ y \ f x \ f y)" lemma antimonoI [intro?]: fixes f :: "'a \ 'b::order" shows "(\x y. x \ y \ f x \ f y) \ antimono f" unfolding antimono_def by iprover lemma antimonoD [dest?]: fixes f :: "'a \ 'b::order" shows "antimono f \ x \ y \ f x \ f y" unfolding antimono_def by iprover lemma antimonoE: fixes f :: "'a \ 'b::order" assumes "antimono f" assumes "x \ y" obtains "f x \ f y" proof from assms show "f x \ f y" by (simp add: antimono_def) qed definition strict_mono :: "('a \ 'b::order) \ bool" where "strict_mono f \ (\x y. x < y \ f x < f y)" lemma strict_monoI [intro?]: assumes "\x y. x < y \ f x < f y" shows "strict_mono f" using assms unfolding strict_mono_def by auto lemma strict_monoD [dest?]: "strict_mono f \ x < y \ f x < f y" unfolding strict_mono_def by auto lemma strict_mono_mono [dest?]: assumes "strict_mono f" shows "mono f" proof (rule monoI) fix x y assume "x \ y" show "f x \ f y" proof (cases "x = y") case True then show ?thesis by simp next case False with \x \ y\ have "x < y" by simp with assms strict_monoD have "f x < f y" by auto then show ?thesis by simp qed qed end context linorder begin lemma mono_invE: fixes f :: "'a \ 'b::order" assumes "mono f" assumes "f x < f y" obtains "x \ y" proof show "x \ y" proof (rule ccontr) assume "\ x \ y" then have "y \ x" by simp with \mono f\ obtain "f y \ f x" by (rule monoE) with \f x < f y\ show False by simp qed qed lemma mono_strict_invE: fixes f :: "'a \ 'b::order" assumes "mono f" assumes "f x < f y" obtains "x < y" proof show "x < y" proof (rule ccontr) assume "\ x < y" then have "y \ x" by simp with \mono f\ obtain "f y \ f x" by (rule monoE) with \f x < f y\ show False by simp qed qed lemma strict_mono_eq: assumes "strict_mono f" shows "f x = f y \ x = y" proof assume "f x = f y" show "x = y" proof (cases x y rule: linorder_cases) case less with assms strict_monoD have "f x < f y" by auto with \f x = f y\ show ?thesis by simp next case equal then show ?thesis . next case greater with assms strict_monoD have "f y < f x" by auto with \f x = f y\ show ?thesis by simp qed qed simp lemma strict_mono_less_eq: assumes "strict_mono f" shows "f x \ f y \ x \ y" proof assume "x \ y" with assms strict_mono_mono monoD show "f x \ f y" by auto next assume "f x \ f y" show "x \ y" proof (rule ccontr) assume "\ x \ y" then have "y < x" by simp with assms strict_monoD have "f y < f x" by auto with \f x \ f y\ show False by simp qed qed lemma strict_mono_less: assumes "strict_mono f" shows "f x < f y \ x < y" using assms by (auto simp add: less_le Orderings.less_le strict_mono_eq strict_mono_less_eq) end subsection \min and max -- fundamental\ definition (in ord) min :: "'a \ 'a \ 'a" where "min a b = (if a \ b then a else b)" definition (in ord) max :: "'a \ 'a \ 'a" where "max a b = (if a \ b then b else a)" lemma min_absorb1: "x \ y \ min x y = x" by (simp add: min_def) lemma max_absorb2: "x \ y \ max x y = y" by (simp add: max_def) lemma min_absorb2: "(y::'a::order) \ x \ min x y = y" by (simp add:min_def) lemma max_absorb1: "(y::'a::order) \ x \ max x y = x" by (simp add: max_def) lemma max_min_same [simp]: fixes x y :: "'a :: linorder" shows "max x (min x y) = x" "max (min x y) x = x" "max (min x y) y = y" "max y (min x y) = y" by(auto simp add: max_def min_def) subsection \(Unique) top and bottom elements\ class bot = fixes bot :: 'a ("\") class order_bot = order + bot + assumes bot_least: "\ \ a" begin sublocale bot: ordering_top greater_eq greater bot by standard (fact bot_least) lemma le_bot: "a \ \ \ a = \" by (fact bot.extremum_uniqueI) lemma bot_unique: "a \ \ \ a = \" by (fact bot.extremum_unique) lemma not_less_bot: "\ a < \" by (fact bot.extremum_strict) lemma bot_less: "a \ \ \ \ < a" by (fact bot.not_eq_extremum) lemma max_bot[simp]: "max bot x = x" by(simp add: max_def bot_unique) lemma max_bot2[simp]: "max x bot = x" by(simp add: max_def bot_unique) lemma min_bot[simp]: "min bot x = bot" by(simp add: min_def bot_unique) lemma min_bot2[simp]: "min x bot = bot" by(simp add: min_def bot_unique) end class top = fixes top :: 'a ("\") class order_top = order + top + assumes top_greatest: "a \ \" begin sublocale top: ordering_top less_eq less top by standard (fact top_greatest) lemma top_le: "\ \ a \ a = \" by (fact top.extremum_uniqueI) lemma top_unique: "\ \ a \ a = \" by (fact top.extremum_unique) lemma not_top_less: "\ \ < a" by (fact top.extremum_strict) lemma less_top: "a \ \ \ a < \" by (fact top.not_eq_extremum) lemma max_top[simp]: "max top x = top" by(simp add: max_def top_unique) lemma max_top2[simp]: "max x top = top" by(simp add: max_def top_unique) lemma min_top[simp]: "min top x = x" by(simp add: min_def top_unique) lemma min_top2[simp]: "min x top = x" by(simp add: min_def top_unique) end subsection \Dense orders\ class dense_order = order + assumes dense: "x < y \ (\z. x < z \ z < y)" class dense_linorder = linorder + dense_order begin lemma dense_le: fixes y z :: 'a assumes "\x. x < y \ x \ z" shows "y \ z" proof (rule ccontr) assume "\ ?thesis" hence "z < y" by simp from dense[OF this] obtain x where "x < y" and "z < x" by safe moreover have "x \ z" using assms[OF \x < y\] . ultimately show False by auto qed lemma dense_le_bounded: fixes x y z :: 'a assumes "x < y" assumes *: "\w. \ x < w ; w < y \ \ w \ z" shows "y \ z" proof (rule dense_le) fix w assume "w < y" from dense[OF \x < y\] obtain u where "x < u" "u < y" by safe from linear[of u w] show "w \ z" proof (rule disjE) assume "u \ w" from less_le_trans[OF \x < u\ \u \ w\] \w < y\ show "w \ z" by (rule *) next assume "w \ u" from \w \ u\ *[OF \x < u\ \u < y\] show "w \ z" by (rule order_trans) qed qed lemma dense_ge: fixes y z :: 'a assumes "\x. z < x \ y \ x" shows "y \ z" proof (rule ccontr) assume "\ ?thesis" hence "z < y" by simp from dense[OF this] obtain x where "x < y" and "z < x" by safe moreover have "y \ x" using assms[OF \z < x\] . ultimately show False by auto qed lemma dense_ge_bounded: fixes x y z :: 'a assumes "z < x" assumes *: "\w. \ z < w ; w < x \ \ y \ w" shows "y \ z" proof (rule dense_ge) fix w assume "z < w" from dense[OF \z < x\] obtain u where "z < u" "u < x" by safe from linear[of u w] show "y \ w" proof (rule disjE) assume "w \ u" from \z < w\ le_less_trans[OF \w \ u\ \u < x\] show "y \ w" by (rule *) next assume "u \ w" from *[OF \z < u\ \u < x\] \u \ w\ show "y \ w" by (rule order_trans) qed qed end class no_top = order + assumes gt_ex: "\y. x < y" class no_bot = order + assumes lt_ex: "\y. y < x" class unbounded_dense_linorder = dense_linorder + no_top + no_bot subsection \Wellorders\ class wellorder = linorder + assumes less_induct [case_names less]: "(\x. (\y. y < x \ P y) \ P x) \ P a" begin lemma wellorder_Least_lemma: fixes k :: 'a assumes "P k" shows LeastI: "P (LEAST x. P x)" and Least_le: "(LEAST x. P x) \ k" proof - have "P (LEAST x. P x) \ (LEAST x. P x) \ k" using assms proof (induct k rule: less_induct) case (less x) then have "P x" by simp show ?case proof (rule classical) assume assm: "\ (P (LEAST a. P a) \ (LEAST a. P a) \ x)" have "\y. P y \ x \ y" proof (rule classical) fix y assume "P y" and "\ x \ y" with less have "P (LEAST a. P a)" and "(LEAST a. P a) \ y" by (auto simp add: not_le) with assm have "x < (LEAST a. P a)" and "(LEAST a. P a) \ y" by auto then show "x \ y" by auto qed with \P x\ have Least: "(LEAST a. P a) = x" by (rule Least_equality) with \P x\ show ?thesis by simp qed qed then show "P (LEAST x. P x)" and "(LEAST x. P x) \ k" by auto qed \ \The following 3 lemmas are due to Brian Huffman\ lemma LeastI_ex: "\x. P x \ P (Least P)" by (erule exE) (erule LeastI) lemma LeastI2: "P a \ (\x. P x \ Q x) \ Q (Least P)" by (blast intro: LeastI) lemma LeastI2_ex: "\a. P a \ (\x. P x \ Q x) \ Q (Least P)" by (blast intro: LeastI_ex) lemma LeastI2_wellorder: assumes "P a" and "\a. \ P a; \b. P b \ a \ b \ \ Q a" shows "Q (Least P)" proof (rule LeastI2_order) show "P (Least P)" using \P a\ by (rule LeastI) next fix y assume "P y" thus "Least P \ y" by (rule Least_le) next fix x assume "P x" "\y. P y \ x \ y" thus "Q x" by (rule assms(2)) qed lemma LeastI2_wellorder_ex: assumes "\x. P x" and "\a. \ P a; \b. P b \ a \ b \ \ Q a" shows "Q (Least P)" using assms by clarify (blast intro!: LeastI2_wellorder) lemma not_less_Least: "k < (LEAST x. P x) \ \ P k" apply (simp add: not_le [symmetric]) apply (erule contrapos_nn) apply (erule Least_le) done lemma exists_least_iff: "(\n. P n) \ (\n. P n \ (\m < n. \ P m))" (is "?lhs \ ?rhs") proof assume ?rhs thus ?lhs by blast next assume H: ?lhs then obtain n where n: "P n" by blast let ?x = "Least P" { fix m assume m: "m < ?x" from not_less_Least[OF m] have "\ P m" . } with LeastI_ex[OF H] show ?rhs by blast qed end subsection \Order on \<^typ>\bool\\ instantiation bool :: "{order_bot, order_top, linorder}" begin definition le_bool_def [simp]: "P \ Q \ P \ Q" definition [simp]: "(P::bool) < Q \ \ P \ Q" definition [simp]: "\ \ False" definition [simp]: "\ \ True" instance proof qed auto end lemma le_boolI: "(P \ Q) \ P \ Q" by simp lemma le_boolI': "P \ Q \ P \ Q" by simp lemma le_boolE: "P \ Q \ P \ (Q \ R) \ R" by simp lemma le_boolD: "P \ Q \ P \ Q" by simp lemma bot_boolE: "\ \ P" by simp lemma top_boolI: \ by simp lemma [code]: "False \ b \ True" "True \ b \ b" "False < b \ b" "True < b \ False" by simp_all subsection \Order on \<^typ>\_ \ _\\ instantiation "fun" :: (type, ord) ord begin definition le_fun_def: "f \ g \ (\x. f x \ g x)" definition "(f::'a \ 'b) < g \ f \ g \ \ (g \ f)" instance .. end instance "fun" :: (type, preorder) preorder proof qed (auto simp add: le_fun_def less_fun_def intro: order_trans antisym) instance "fun" :: (type, order) order proof qed (auto simp add: le_fun_def intro: antisym) instantiation "fun" :: (type, bot) bot begin definition "\ = (\x. \)" instance .. end instantiation "fun" :: (type, order_bot) order_bot begin lemma bot_apply [simp, code]: "\ x = \" by (simp add: bot_fun_def) instance proof qed (simp add: le_fun_def) end instantiation "fun" :: (type, top) top begin definition [no_atp]: "\ = (\x. \)" instance .. end instantiation "fun" :: (type, order_top) order_top begin lemma top_apply [simp, code]: "\ x = \" by (simp add: top_fun_def) instance proof qed (simp add: le_fun_def) end lemma le_funI: "(\x. f x \ g x) \ f \ g" unfolding le_fun_def by simp lemma le_funE: "f \ g \ (f x \ g x \ P) \ P" unfolding le_fun_def by simp lemma le_funD: "f \ g \ f x \ g x" by (rule le_funE) lemma mono_compose: "mono Q \ mono (\i x. Q i (f x))" unfolding mono_def le_fun_def by auto subsection \Order on unary and binary predicates\ lemma predicate1I: assumes PQ: "\x. P x \ Q x" shows "P \ Q" apply (rule le_funI) apply (rule le_boolI) apply (rule PQ) apply assumption done lemma predicate1D: "P \ Q \ P x \ Q x" apply (erule le_funE) apply (erule le_boolE) apply assumption+ done lemma rev_predicate1D: "P x \ P \ Q \ Q x" by (rule predicate1D) lemma predicate2I: assumes PQ: "\x y. P x y \ Q x y" shows "P \ Q" apply (rule le_funI)+ apply (rule le_boolI) apply (rule PQ) apply assumption done lemma predicate2D: "P \ Q \ P x y \ Q x y" apply (erule le_funE)+ apply (erule le_boolE) apply assumption+ done lemma rev_predicate2D: "P x y \ P \ Q \ Q x y" by (rule predicate2D) lemma bot1E [no_atp]: "\ x \ P" by (simp add: bot_fun_def) lemma bot2E: "\ x y \ P" by (simp add: bot_fun_def) lemma top1I: "\ x" by (simp add: top_fun_def) lemma top2I: "\ x y" by (simp add: top_fun_def) subsection \Name duplicates\ lemmas order_eq_refl = preorder_class.eq_refl lemmas order_less_irrefl = preorder_class.less_irrefl lemmas order_less_imp_le = preorder_class.less_imp_le lemmas order_less_not_sym = preorder_class.less_not_sym lemmas order_less_asym = preorder_class.less_asym lemmas order_less_trans = preorder_class.less_trans lemmas order_le_less_trans = preorder_class.le_less_trans lemmas order_less_le_trans = preorder_class.less_le_trans lemmas order_less_imp_not_less = preorder_class.less_imp_not_less lemmas order_less_imp_triv = preorder_class.less_imp_triv lemmas order_less_asym' = preorder_class.less_asym' lemmas order_less_le = order_class.less_le lemmas order_le_less = order_class.le_less lemmas order_le_imp_less_or_eq = order_class.le_imp_less_or_eq lemmas order_less_imp_not_eq = order_class.less_imp_not_eq lemmas order_less_imp_not_eq2 = order_class.less_imp_not_eq2 lemmas order_neq_le_trans = order_class.neq_le_trans lemmas order_le_neq_trans = order_class.le_neq_trans lemmas order_antisym = order_class.antisym lemmas order_eq_iff = order_class.eq_iff lemmas order_antisym_conv = order_class.antisym_conv lemmas linorder_linear = linorder_class.linear lemmas linorder_less_linear = linorder_class.less_linear lemmas linorder_le_less_linear = linorder_class.le_less_linear lemmas linorder_le_cases = linorder_class.le_cases lemmas linorder_not_less = linorder_class.not_less lemmas linorder_not_le = linorder_class.not_le lemmas linorder_neq_iff = linorder_class.neq_iff lemmas linorder_neqE = linorder_class.neqE end