diff --git a/src/HOL/Library/Lexord.thy b/src/HOL/Library/Lexord.thy new file mode 100644 --- /dev/null +++ b/src/HOL/Library/Lexord.thy @@ -0,0 +1,208 @@ +section \Lexicographic orderings\ + +theory Lexord + imports Main +begin + +subsection \The preorder case\ + +locale lex_preordering = preordering +begin + +inductive lex_less :: \'a list \ 'a list \ bool\ (infix \[\<^bold><]\ 50) +where + Nil: \[] [\<^bold><] y # ys\ +| Cons: \x \<^bold>< y \ x # xs [\<^bold><] y # ys\ +| Cons_eq: \x \<^bold>\ y \ y \<^bold>\ x \ xs [\<^bold><] ys \ x # xs [\<^bold><] y # ys\ + +inductive lex_less_eq :: \'a list \ 'a list \ bool\ (infix \[\<^bold>\]\ 50) +where + Nil: \[] [\<^bold>\] ys\ +| Cons: \x \<^bold>< y \ x # xs [\<^bold>\] y # ys\ +| Cons_eq: \x \<^bold>\ y \ y \<^bold>\ x \ xs [\<^bold>\] ys \ x # xs [\<^bold>\] y # ys\ + +lemma lex_less_simps [simp]: + \[] [\<^bold><] y # ys\ + \\ xs [\<^bold><] []\ + \x # xs [\<^bold><] y # ys \ x \<^bold>< y \ x \<^bold>\ y \ y \<^bold>\ x \ xs [\<^bold><] ys\ + by (auto intro: lex_less.intros elim: lex_less.cases) + +lemma lex_less_eq_simps [simp]: + \[] [\<^bold>\] ys\ + \\ x # xs [\<^bold>\] []\ + \x # xs [\<^bold>\] y # ys \ x \<^bold>< y \ x \<^bold>\ y \ y \<^bold>\ x \ xs [\<^bold>\] ys\ + by (auto intro: lex_less_eq.intros elim: lex_less_eq.cases) + +lemma lex_less_code [code]: + \[] [\<^bold><] y # ys \ True\ + \xs [\<^bold><] [] \ False\ + \x # xs [\<^bold><] y # ys \ x \<^bold>< y \ x \<^bold>\ y \ y \<^bold>\ x \ xs [\<^bold><] ys\ + by simp_all + +lemma lex_less_eq_code [code]: + \[] [\<^bold>\] ys \ True\ + \x # xs [\<^bold>\] [] \ False\ + \x # xs [\<^bold>\] y # ys \ x \<^bold>< y \ x \<^bold>\ y \ y \<^bold>\ x \ xs [\<^bold>\] ys\ + by simp_all + +lemma preordering: + \preordering ([\<^bold>\]) ([\<^bold><])\ +proof + fix xs ys zs + show \xs [\<^bold>\] xs\ + by (induction xs) (simp_all add: refl) + show \xs [\<^bold>\] zs\ if \xs [\<^bold>\] ys\ \ys [\<^bold>\] zs\ + using that proof (induction arbitrary: zs) + case (Nil ys) + then show ?case by simp + next + case (Cons x y xs ys) + then show ?case + by (cases zs) (auto dest: strict_trans strict_trans2) + next + case (Cons_eq x y xs ys) + then show ?case + by (cases zs) (auto dest: strict_trans1 intro: trans) + qed + show \xs [\<^bold><] ys \ xs [\<^bold>\] ys \ \ ys [\<^bold>\] xs\ (is \?P \ ?Q\) + proof + assume ?P + then have \xs [\<^bold>\] ys\ + by induction simp_all + moreover have \\ ys [\<^bold>\] xs\ + using \?P\ + by induction (simp_all, simp_all add: strict_iff_not asym) + ultimately show ?Q .. + next + assume ?Q + then have \xs [\<^bold>\] ys\ \\ ys [\<^bold>\] xs\ + by auto + then show ?P + proof induction + case (Nil ys) + then show ?case + by (cases ys) simp_all + next + case (Cons x y xs ys) + then show ?case + by simp + next + case (Cons_eq x y xs ys) + then show ?case + by simp + qed + qed +qed + +interpretation lex: preordering \([\<^bold>\])\ \([\<^bold><])\ + by (fact preordering) + +end + + +subsection \The order case\ + +locale lex_ordering = lex_preordering + ordering +begin + +interpretation lex: preordering \([\<^bold>\])\ \([\<^bold><])\ + by (fact preordering) + +lemma less_lex_Cons_iff [simp]: + \x # xs [\<^bold><] y # ys \ x \<^bold>< y \ x = y \ xs [\<^bold><] ys\ + by (auto intro: refl antisym) + +lemma less_eq_lex_Cons_iff [simp]: + \x # xs [\<^bold>\] y # ys \ x \<^bold>< y \ x = y \ xs [\<^bold>\] ys\ + by (auto intro: refl antisym) + +lemma ordering: + \ordering ([\<^bold>\]) ([\<^bold><])\ +proof + fix xs ys + show *: \xs = ys\ if \xs [\<^bold>\] ys\ \ys [\<^bold>\] xs\ + using that proof induction + case (Nil ys) + then show ?case by (cases ys) simp + next + case (Cons x y xs ys) + then show ?case by (auto dest: asym intro: antisym) + (simp add: strict_iff_not) + next + case (Cons_eq x y xs ys) + then show ?case by (auto intro: antisym) + (simp add: strict_iff_not) + qed + show \xs [\<^bold><] ys \ xs [\<^bold>\] ys \ xs \ ys\ + by (auto simp add: lex.strict_iff_not dest: *) +qed + +interpretation lex: ordering \([\<^bold>\])\ \([\<^bold><])\ + by (fact ordering) + +end + + +subsection \Canonical instance\ + +instantiation list :: (preorder) preorder +begin + +global_interpretation lex: lex_preordering \(\) :: 'a::preorder \ 'a \ bool\ \(<) :: 'a \ 'a \ bool\ + defines less_eq_list = lex.lex_less_eq + and less_list = lex.lex_less .. + +instance + by (rule class.preorder.of_class.intro, rule preordering_preorderI, fact lex.preordering) + +end + +global_interpretation lex: lex_ordering \(\) :: 'a::order \ 'a \ bool\ \(<) :: 'a \ 'a \ bool\ + rewrites \lex_preordering.lex_less_eq (\) (<) = ((\) :: 'a list \ 'a list \ bool)\ + and \lex_preordering.lex_less (\) (<) = ((<) :: 'a list \ 'a list \ bool)\ +proof - + interpret lex_ordering \(\) :: 'a \ 'a \ bool\ \(<) :: 'a \ 'a \ bool\ .. + show \lex_ordering ((\) :: 'a \ 'a \ bool) (<)\ + by (fact lex_ordering_axioms) + show \lex_preordering.lex_less_eq (\) (<) = (\)\ + by (simp add: less_eq_list_def) + show \lex_preordering.lex_less (\) (<) = (<)\ + by (simp add: less_list_def) +qed + +instance list :: (order) order + by (rule class.order.of_class.intro, rule ordering_orderI, fact lex.ordering) + +export_code \(\) :: _ list \ _ list \ bool\ \(<) :: _ list \ _ list \ bool\ in Haskell + + +subsection \Non-canonical instance\ + +context comm_monoid_mult +begin + +definition dvd_strict :: \'a \ 'a \ bool\ + where \dvd_strict a b \ a dvd b \ \ b dvd a\ + +end + +global_interpretation dvd: lex_preordering \(dvd) :: 'a::comm_monoid_mult \ 'a \ bool\ dvd_strict + defines lex_dvd = dvd.lex_less_eq + and lex_dvd_strict = dvd.lex_less + apply (rule lex_preordering.intro) + apply standard + apply (auto simp add: dvd_strict_def) + done + +print_theorems + +global_interpretation lex_dvd: preordering lex_dvd lex_dvd_strict + by (fact dvd.preordering) + +definition \example = lex_dvd [(4::int), - 7, 8] [- 8, 13, 5]\ + +export_code example in Haskell + +value example + +end diff --git a/src/HOL/Orderings.thy b/src/HOL/Orderings.thy --- a/src/HOL/Orderings.thy +++ b/src/HOL/Orderings.thy @@ -1,1702 +1,1712 @@ (* 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_procedure.ML\ ML_file \~~/src/Provers/order_tac.ML\ subsection \Abstract ordering\ locale partial_preordering = fixes less_eq :: \'a \ 'a \ bool\ (infix \\<^bold>\\ 50) assumes refl: \a \<^bold>\ a\ \ \not \iff\: makes problems due to multiple (dual) interpretations\ and trans: \a \<^bold>\ b \ b \<^bold>\ c \ a \<^bold>\ c\ locale preordering = partial_preordering + fixes less :: \'a \ 'a \ bool\ (infix \\<^bold><\ 50) assumes strict_iff_not: \a \<^bold>< b \ a \<^bold>\ b \ \ b \<^bold>\ a\ begin lemma strict_implies_order: \a \<^bold>< b \ a \<^bold>\ b\ by (simp add: strict_iff_not) lemma irrefl: \ \not \iff\: makes problems due to multiple (dual) interpretations\ \\ a \<^bold>< a\ by (simp add: strict_iff_not) lemma asym: \a \<^bold>< b \ b \<^bold>< a \ False\ by (auto simp add: strict_iff_not) lemma strict_trans1: \a \<^bold>\ b \ b \<^bold>< c \ a \<^bold>< c\ by (auto simp add: strict_iff_not intro: trans) lemma strict_trans2: \a \<^bold>< b \ b \<^bold>\ c \ a \<^bold>< c\ by (auto simp add: strict_iff_not intro: trans) lemma strict_trans: \a \<^bold>< b \ b \<^bold>< c \ a \<^bold>< c\ by (auto intro: strict_trans1 strict_implies_order) end lemma preordering_strictI: \ \Alternative introduction rule with bias towards strict order\ 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 \preordering (\<^bold>\) (\<^bold><)\ proof fix a b show \a \<^bold>< b \ a \<^bold>\ b \ \ b \<^bold>\ a\ 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) qed lemma preordering_dualI: fixes less_eq (infix \\<^bold>\\ 50) and less (infix \\<^bold><\ 50) assumes \preordering (\a b. b \<^bold>\ a) (\a b. b \<^bold>< a)\ shows \preordering (\<^bold>\) (\<^bold><)\ proof - from assms interpret preordering \\a b. b \<^bold>\ a\ \\a b. b \<^bold>< a\ . show ?thesis by standard (auto simp: strict_iff_not refl intro: trans) qed locale ordering = partial_preordering + fixes less :: \'a \ 'a \ bool\ (infix \\<^bold><\ 50) assumes strict_iff_order: \a \<^bold>< b \ a \<^bold>\ b \ a \ b\ assumes antisym: \a \<^bold>\ b \ b \<^bold>\ a \ a = b\ begin sublocale preordering \(\<^bold>\)\ \(\<^bold><)\ proof show \a \<^bold>< b \ a \<^bold>\ b \ \ b \<^bold>\ a\ for a b by (auto simp add: strict_iff_order intro: antisym) qed 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 eq_iff: \a = b \ a \<^bold>\ b \ b \<^bold>\ a\ by (auto simp add: refl intro: antisym) end lemma ordering_strictI: \ \Alternative introduction rule with bias towards strict order\ 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 (\<^bold>\) (\<^bold><)\ 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 (\<^bold>\) (\<^bold><)\ 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 sublocale order: preordering less_eq less + dual_order: preordering greater_eq greater proof - interpret preordering less_eq less by standard (auto intro: order_trans simp add: less_le_not_le) show \preordering less_eq less\ by (fact preordering_axioms) then show \preordering greater_eq greater\ by (rule preordering_dualI) qed 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 +lemma preordering_preorderI: + \class.preorder (\<^bold>\) (\<^bold><)\ if \preordering (\<^bold>\) (\<^bold><)\ + for less_eq (infix \\<^bold>\\ 50) and less (infix \\<^bold><\ 50) +proof - + from that interpret preordering \(\<^bold>\)\ \(\<^bold><)\ . + show ?thesis + by standard (auto simp add: strict_iff_not refl intro: trans) +qed + + subsection \Partial orders\ class order = preorder + assumes order_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: order_antisym) sublocale order: ordering less_eq less + dual_order: ordering greater_eq greater proof - interpret ordering less_eq less by standard (auto intro: order_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 print_theorems 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 order_eq_iff: "x = y \ x \ y \ y \ x" by (fact order.eq_iff) lemma antisym_conv: "y \ x \ x \ y \ x = y" by (simp add: order.eq_iff) lemma less_imp_neq: "x < y \ x \ y" 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 order.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 order.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 order.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: order.antisym)+ lemma Greatest_equality: "\ P x; \y. P y \ x \ y \ \ Greatest P = x" unfolding Greatest_def by (rule the_equality) (blast intro: order.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" + fixes less (infix "\<^bold><" 50) + and less_eq (infix "\<^bold>\" 50) + assumes "\a b. a \<^bold>\ b \ a \<^bold>< b \ a = b" + assumes "\a b. a \<^bold>< b \ \ b \<^bold>< a" + assumes "\a. \ a \<^bold>< a" + assumes "\a b c. a \<^bold>< b \ b \<^bold>< c \ a \<^bold>< 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: order.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: order.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: order.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 \ structure Logic_Signature : LOGIC_SIGNATURE = struct val mk_Trueprop = HOLogic.mk_Trueprop val dest_Trueprop = HOLogic.dest_Trueprop val Trueprop_conv = HOLogic.Trueprop_conv val Not = HOLogic.Not val conj = HOLogic.conj val disj = HOLogic.disj val notI = @{thm notI} val ccontr = @{thm ccontr} val conjI = @{thm conjI} val conjE = @{thm conjE} val disjE = @{thm disjE} val not_not_conv = Conv.rewr_conv @{thm eq_reflection[OF not_not]} val de_Morgan_conj_conv = Conv.rewr_conv @{thm eq_reflection[OF de_Morgan_conj]} val de_Morgan_disj_conv = Conv.rewr_conv @{thm eq_reflection[OF de_Morgan_disj]} val conj_disj_distribL_conv = Conv.rewr_conv @{thm eq_reflection[OF conj_disj_distribL]} val conj_disj_distribR_conv = Conv.rewr_conv @{thm eq_reflection[OF conj_disj_distribR]} end structure HOL_Base_Order_Tac = Base_Order_Tac( structure Logic_Sig = Logic_Signature; (* Exclude types with specialised solvers. *) val excluded_types = [HOLogic.natT, HOLogic.intT, HOLogic.realT] ) structure HOL_Order_Tac = Order_Tac(structure Base_Tac = HOL_Base_Order_Tac) fun print_orders ctxt0 = let val ctxt = Config.put show_sorts true ctxt0 val orders = HOL_Order_Tac.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)), Pretty.brk 1] fun pretty_order ({kind = kind, ops = ops, ...}, _) = Pretty.block ([Pretty.str (@{make_string} kind), Pretty.str ":", Pretty.brk 1] @ map pretty_term ops) in Pretty.writeln (Pretty.big_list "order structures:" (map pretty_order orders)) end val _ = Outer_Syntax.command \<^command_keyword>\print_orders\ "print order structures available to transitivity reasoner" (Scan.succeed (Toplevel.keep (print_orders o Toplevel.context_of))) \ method_setup order = \ Scan.succeed (fn ctxt => SIMPLE_METHOD' (HOL_Order_Tac.tac [] ctxt)) \ "transitivity reasoner" text \Declarations to set up transitivity reasoner of partial and linear orders.\ context order begin lemma nless_le: "(\ a < b) \ (\ a \ b) \ a = b" using local.dual_order.order_iff_strict by blast local_setup \ HOL_Order_Tac.declare_order { ops = {eq = @{term \(=) :: 'a \ 'a \ bool\}, le = @{term \(\)\}, lt = @{term \(<)\}}, thms = {trans = @{thm order_trans}, refl = @{thm order_refl}, eqD1 = @{thm eq_refl}, eqD2 = @{thm eq_refl[OF sym]}, antisym = @{thm order_antisym}, contr = @{thm notE}}, conv_thms = {less_le = @{thm eq_reflection[OF less_le]}, nless_le = @{thm eq_reflection[OF nless_le]}} } \ end context linorder begin lemma nle_le: "(\ a \ b) \ b \ a \ b \ a" using not_le less_le by simp local_setup \ HOL_Order_Tac.declare_linorder { ops = {eq = @{term \(=) :: 'a \ 'a \ bool\}, le = @{term \(\)\}, lt = @{term \(<)\}}, thms = {trans = @{thm order_trans}, refl = @{thm order_refl}, eqD1 = @{thm eq_refl}, eqD2 = @{thm eq_refl[OF sym]}, antisym = @{thm order_antisym}, contr = @{thm notE}}, conv_thms = {less_le = @{thm eq_reflection[OF less_le]}, nless_le = @{thm eq_reflection[OF not_less]}, nle_le = @{thm eq_reflection[OF nle_le]}} } \ end setup \ map_theory_simpset (fn ctxt0 => ctxt0 addSolver mk_solver "Transitivity" (fn ctxt => HOL_Order_Tac.tac (Simplifier.prems_of ctxt) ctxt)) \ 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] = order.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 order.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 order.antisym) instance "fun" :: (type, order) order proof qed (auto simp add: le_fun_def intro: order.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 antisym = order.antisym lemmas eq_iff = order.eq_iff 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_eq_iff = order_class.order.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