nipkow@15737: (* Title: HOL/Library/Product_ord.thy nipkow@15737: Author: Norbert Voelker nipkow@15737: *) nipkow@15737: wenzelm@17200: header {* Order on product types *} nipkow@15737: nipkow@15737: theory Product_ord 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@37765: prod_le_def: "x \ y \ fst x < fst y \ fst x \ fst y \ snd x \ snd y" haftmann@25571: haftmann@25571: definition haftmann@37765: prod_less_def: "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@28562: lemma [code]: wenzelm@47961: "(x1::'a::{ord, equal}, y1) \ (x2, y2) \ x1 < x2 \ x1 \ x2 \ y1 \ y2" wenzelm@47961: "(x1::'a::{ord, equal}, y1) < (x2, y2) \ x1 < x2 \ x1 \ x2 \ y1 < y2" haftmann@25502: unfolding prod_le_def prod_less_def by simp_all haftmann@22177: wenzelm@47961: instance prod :: (preorder, preorder) preorder wenzelm@47961: by default (auto simp: prod_le_def prod_less_def less_le_not_le intro: order_trans) nipkow@15737: wenzelm@47961: instance prod :: (order, order) order wenzelm@47961: by default (auto simp add: prod_le_def) haftmann@31040: wenzelm@47961: instance prod :: (linorder, linorder) linorder wenzelm@47961: by default (auto simp: prod_le_def) nipkow@15737: haftmann@37678: instantiation prod :: (linorder, linorder) distrib_lattice haftmann@25571: begin haftmann@25571: haftmann@25571: definition wenzelm@47961: inf_prod_def: "(inf :: 'a \ 'b \ _ \ _) = min" haftmann@25571: haftmann@25571: definition wenzelm@47961: sup_prod_def: "(sup :: 'a \ 'b \ _ \ _) = max" haftmann@25571: wenzelm@47961: instance wenzelm@47961: by default (auto simp add: inf_prod_def sup_prod_def min_max.sup_inf_distrib1) haftmann@31040: haftmann@31040: end haftmann@31040: haftmann@37678: instantiation prod :: (bot, bot) bot haftmann@31040: begin haftmann@31040: haftmann@31040: definition haftmann@31040: bot_prod_def: "bot = (bot, bot)" haftmann@31040: wenzelm@47961: instance wenzelm@47961: by default (auto simp add: bot_prod_def prod_le_def) haftmann@31040: haftmann@31040: end haftmann@31040: haftmann@37678: instantiation prod :: (top, top) top haftmann@31040: begin haftmann@31040: haftmann@31040: definition haftmann@31040: top_prod_def: "top = (top, top)" haftmann@31040: wenzelm@47961: instance wenzelm@47961: by default (auto simp add: top_prod_def prod_le_def) haftmann@22483: wenzelm@19736: end haftmann@25571: huffman@44063: text {* A stronger version of the definition holds for partial orders. *} huffman@44063: huffman@44063: lemma prod_less_eq: huffman@44063: fixes x y :: "'a::order \ 'b::ord" huffman@44063: shows "x < y \ fst x < fst y \ (fst x = fst y \ snd x < snd y)" huffman@44063: unfolding prod_less_def fst_conv snd_conv le_less by auto huffman@44063: 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@47961: case (less a\<^isub>1) note a\<^isub>1 = this wenzelm@47961: show "P (a\<^isub>1, b)" wenzelm@47961: proof (induct b rule: less_induct) wenzelm@47961: case (less b\<^isub>1) note b\<^isub>1 = this wenzelm@47961: show "P (a\<^isub>1, b\<^isub>1)" wenzelm@47961: proof (rule P) wenzelm@47961: fix p assume p: "p < (a\<^isub>1, b\<^isub>1)" wenzelm@47961: show "P p" wenzelm@47961: proof (cases "fst p < a\<^isub>1") wenzelm@47961: case True wenzelm@47961: then have "P (fst p, snd p)" by (rule a\<^isub>1) wenzelm@47961: then show ?thesis by simp wenzelm@47961: next wenzelm@47961: case False wenzelm@47961: with p have 1: "a\<^isub>1 = fst p" and 2: "snd p < b\<^isub>1" wenzelm@47961: by (simp_all add: prod_less_eq) wenzelm@47961: from 2 have "P (a\<^isub>1, snd p)" by (rule b\<^isub>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@25571: end