haftmann@51115: (* Title: HOL/Library/Product_Lexorder.thy nipkow@15737: Author: Norbert Voelker nipkow@15737: *) nipkow@15737: haftmann@51115: header {* Lexicographic order on product types *} nipkow@15737: haftmann@51115: theory Product_Lexorder haftmann@30738: imports Main nipkow@15737: begin nipkow@15737: haftmann@37678: instantiation prod :: (ord, ord) ord haftmann@25571: begin haftmann@25571: haftmann@25571: definition haftmann@51115: "x \ y \ fst x < fst y \ fst x \ fst y \ snd x \ snd y" haftmann@25571: haftmann@25571: definition haftmann@51115: "x < y \ fst x < fst y \ fst x \ fst y \ snd x < snd y" haftmann@25571: haftmann@25571: instance .. haftmann@25571: haftmann@25571: end nipkow@15737: haftmann@51115: lemma less_eq_prod_simp [simp, code]: haftmann@51115: "(x1, y1) \ (x2, y2) \ x1 < x2 \ x1 \ x2 \ y1 \ y2" haftmann@51115: by (simp add: less_eq_prod_def) haftmann@51115: haftmann@51115: lemma less_prod_simp [simp, code]: haftmann@51115: "(x1, y1) < (x2, y2) \ x1 < x2 \ x1 \ x2 \ y1 < y2" haftmann@51115: by (simp add: less_prod_def) haftmann@51115: haftmann@51115: text {* A stronger version for partial orders. *} haftmann@51115: haftmann@51115: lemma less_prod_def': haftmann@51115: fixes x y :: "'a::order \ 'b::ord" haftmann@51115: shows "x < y \ fst x < fst y \ fst x = fst y \ snd x < snd y" haftmann@51115: by (auto simp add: less_prod_def le_less) haftmann@22177: wenzelm@47961: instance prod :: (preorder, preorder) preorder haftmann@51115: by default (auto simp: less_eq_prod_def less_prod_def less_le_not_le intro: order_trans) nipkow@15737: wenzelm@47961: instance prod :: (order, order) order haftmann@51115: by default (auto simp add: less_eq_prod_def) haftmann@31040: wenzelm@47961: instance prod :: (linorder, linorder) linorder haftmann@51115: by default (auto simp: less_eq_prod_def) nipkow@15737: haftmann@37678: instantiation prod :: (linorder, linorder) distrib_lattice haftmann@25571: begin haftmann@25571: haftmann@25571: definition haftmann@51115: "(inf :: 'a \ 'b \ _ \ _) = min" haftmann@25571: haftmann@25571: definition haftmann@51115: "(sup :: 'a \ 'b \ _ \ _) = max" haftmann@25571: wenzelm@47961: instance haftmann@54863: by default (auto simp add: inf_prod_def sup_prod_def max_min_distrib2) haftmann@31040: haftmann@31040: end haftmann@31040: haftmann@37678: instantiation prod :: (bot, bot) bot haftmann@31040: begin haftmann@31040: haftmann@31040: definition haftmann@51115: "bot = (bot, bot)" haftmann@31040: haftmann@52729: instance .. haftmann@31040: haftmann@31040: end haftmann@31040: haftmann@52729: instance prod :: (order_bot, order_bot) order_bot haftmann@52729: by default (auto simp add: bot_prod_def) haftmann@52729: haftmann@37678: instantiation prod :: (top, top) top haftmann@31040: begin haftmann@31040: haftmann@31040: definition haftmann@51115: "top = (top, top)" haftmann@31040: haftmann@52729: instance .. haftmann@22483: wenzelm@19736: end haftmann@25571: haftmann@52729: instance prod :: (order_top, order_top) order_top haftmann@52729: by default (auto simp add: top_prod_def) haftmann@52729: huffman@44063: instance prod :: (wellorder, wellorder) wellorder huffman@44063: proof huffman@44063: fix P :: "'a \ 'b \ bool" and z :: "'a \ 'b" huffman@44063: assume P: "\x. (\y. y < x \ P y) \ P x" huffman@44063: show "P z" huffman@44063: proof (induct z) huffman@44063: case (Pair a b) huffman@44063: show "P (a, b)" wenzelm@47961: proof (induct a arbitrary: b rule: less_induct) wenzelm@53015: case (less a\<^sub>1) note a\<^sub>1 = this wenzelm@53015: show "P (a\<^sub>1, b)" wenzelm@47961: proof (induct b rule: less_induct) wenzelm@53015: case (less b\<^sub>1) note b\<^sub>1 = this wenzelm@53015: show "P (a\<^sub>1, b\<^sub>1)" wenzelm@47961: proof (rule P) wenzelm@53015: fix p assume p: "p < (a\<^sub>1, b\<^sub>1)" wenzelm@47961: show "P p" wenzelm@53015: proof (cases "fst p < a\<^sub>1") wenzelm@47961: case True wenzelm@53015: then have "P (fst p, snd p)" by (rule a\<^sub>1) wenzelm@47961: then show ?thesis by simp wenzelm@47961: next wenzelm@47961: case False wenzelm@53015: with p have 1: "a\<^sub>1 = fst p" and 2: "snd p < b\<^sub>1" haftmann@51115: by (simp_all add: less_prod_def') wenzelm@53015: from 2 have "P (a\<^sub>1, snd p)" by (rule b\<^sub>1) wenzelm@47961: with 1 show ?thesis by simp wenzelm@47961: qed wenzelm@47961: qed wenzelm@47961: qed wenzelm@47961: qed huffman@44063: qed huffman@44063: qed huffman@44063: haftmann@51115: text {* Legacy lemma bindings *} haftmann@51115: haftmann@51115: lemmas prod_le_def = less_eq_prod_def haftmann@51115: lemmas prod_less_def = less_prod_def haftmann@51115: lemmas prod_less_eq = less_prod_def' haftmann@51115: haftmann@25571: end haftmann@51115: