# HG changeset patch # User nipkow # Date 1547584280 -3600 # Node ID 41ff40bf15308a6e3b5313c3a76a26b4b5e8807c # Parent a03a63b81f446eafc88b91b29cf6eeed4a080dd9 moved and renamed class diff -r a03a63b81f44 -r 41ff40bf1530 src/HOL/Analysis/Cartesian_Euclidean_Space.thy --- a/src/HOL/Analysis/Cartesian_Euclidean_Space.thy Mon Jan 14 18:35:03 2019 +0000 +++ b/src/HOL/Analysis/Cartesian_Euclidean_Space.thy Tue Jan 15 21:31:20 2019 +0100 @@ -873,16 +873,6 @@ fixes a::num1 shows "a = 1" by (rule num1_eqI) -instantiation num1 :: cart_one -begin - -instance -proof - show "CARD(1) = Suc 0" by auto -qed - -end - instantiation num1 :: linorder begin definition "a < b \ Rep_num1 a < Rep_num1 b" definition "a \ b \ Rep_num1 a \ Rep_num1 b" diff -r a03a63b81f44 -r 41ff40bf1530 src/HOL/Analysis/Finite_Cartesian_Product.thy --- a/src/HOL/Analysis/Finite_Cartesian_Product.thy Mon Jan 14 18:35:03 2019 +0000 +++ b/src/HOL/Analysis/Finite_Cartesian_Product.thy Tue Jan 15 21:31:20 2019 +0100 @@ -259,32 +259,20 @@ text\The ordering on one-dimensional vectors is linear.\ -class cart_one = - assumes UNIV_one: "card (UNIV :: 'a set) = Suc 0" -begin - -subclass finite -proof - from UNIV_one show "finite (UNIV :: 'a set)" - by (auto intro!: card_ge_0_finite) -qed - -end - instance vec:: (order, finite) order by standard (auto simp: less_eq_vec_def less_vec_def vec_eq_iff intro: order.trans order.antisym order.strict_implies_order) -instance vec :: (linorder, cart_one) linorder +instance vec :: (linorder, CARD_1) linorder proof obtain a :: 'b where all: "\P. (\i. P i) \ P a" proof - - have "card (UNIV :: 'b set) = Suc 0" by (rule UNIV_one) + have "CARD ('b) = 1" by (rule CARD_1) then obtain b :: 'b where "UNIV = {b}" by (auto iff: card_Suc_eq) then have "\P. (\i\UNIV. P i) \ P b" by auto then show thesis by (auto intro: that) qed - fix x y :: "'a^'b::cart_one" + fix x y :: "'a^'b::CARD_1" note [simp] = less_eq_vec_def less_vec_def all vec_eq_iff field_simps show "x \ y \ y \ x" by auto qed diff -r a03a63b81f44 -r 41ff40bf1530 src/HOL/Library/Cardinality.thy --- a/src/HOL/Library/Cardinality.thy Mon Jan 14 18:35:03 2019 +0000 +++ b/src/HOL/Library/Cardinality.thy Tue Jan 15 21:31:20 2019 +0100 @@ -153,6 +153,19 @@ lemma one_le_card_finite [simp]: "Suc 0 \ CARD('a::finite)" by (simp add: less_Suc_eq_le [symmetric]) + +class CARD_1 = + assumes CARD_1: "CARD ('a) = 1" +begin + +subclass finite +proof + from CARD_1 show "finite (UNIV :: 'a set)" + by (auto intro!: card_ge_0_finite) +qed + +end + text \Class for cardinality "at least 2"\ class card2 = finite + diff -r a03a63b81f44 -r 41ff40bf1530 src/HOL/Library/Numeral_Type.thy --- a/src/HOL/Library/Numeral_Type.thy Mon Jan 14 18:35:03 2019 +0000 +++ b/src/HOL/Library/Numeral_Type.thy Tue Jan 15 21:31:20 2019 +0100 @@ -52,6 +52,16 @@ using finite by (rule finite_imageI) qed +instantiation num1 :: CARD_1 +begin + +instance +proof + show "CARD(num1) = 1" by auto +qed + +end + instance bit0 :: (finite) card2 proof show "finite (UNIV::'a bit0 set)"