# HG changeset patch # User wenzelm # Date 1027462252 -7200 # Node ID 666137b488a4728d06050f460eca06940a024eb6 # Parent 181a293aa37a9aae8a2aa4afd8f8a1727f08dc6c predicate defs via locales; diff -r 181a293aa37a -r 666137b488a4 src/HOL/HOL.thy --- a/src/HOL/HOL.thy Wed Jul 24 00:09:44 2002 +0200 +++ b/src/HOL/HOL.thy Wed Jul 24 00:10:52 2002 +0200 @@ -601,15 +601,17 @@ subsubsection {* Monotonicity *} -constdefs - mono :: "['a::ord => 'b::ord] => bool" - "mono f == ALL A B. A <= B --> f A <= f B" +locale mono = + fixes f + assumes mono: "A <= B ==> f A <= f B" -lemma monoI [intro?]: "(!!A B. A <= B ==> f A <= f B) ==> mono f" - by (unfold mono_def) rules +lemmas monoI [intro?] = mono.intro [OF mono_axioms.intro] + and monoD [dest?] = mono.mono -lemma monoD [dest?]: "mono f ==> A <= B ==> f A <= f B" - by (unfold mono_def) rules +lemma mono_def: "mono f == ALL A B. A <= B --> f A <= f B" + -- compatibility + by (simp only: atomize_eq mono_def mono_axioms_def) + constdefs min :: "['a::ord, 'a] => 'a" diff -r 181a293aa37a -r 666137b488a4 src/HOL/Record.thy --- a/src/HOL/Record.thy Wed Jul 24 00:09:44 2002 +0200 +++ b/src/HOL/Record.thy Wed Jul 24 00:10:52 2002 +0200 @@ -11,118 +11,57 @@ subsection {* Abstract product types *} -constdefs - product_type :: "('p => 'a * 'b) => ('a * 'b => 'p) => - ('a => 'b => 'p) => ('p => 'a) => ('p => 'b) => bool" - "product_type Rep Abs pair dest1 dest2 == - type_definition Rep Abs UNIV \ - pair = (\a b. Abs (a, b)) \ - dest1 = (\p. fst (Rep p)) \ - dest2 = (\p. snd (Rep p))" - -lemma product_typeI: - "type_definition Rep Abs UNIV ==> - pair == \a b. Abs (a, b) ==> - dest1 == (\p. fst (Rep p)) ==> - dest2 == (\p. snd (Rep p)) ==> - product_type Rep Abs pair dest1 dest2" - by (simp add: product_type_def) +locale product_type = + fixes Rep and Abs and pair and dest1 and dest2 + assumes "typedef": "type_definition Rep Abs UNIV" + and pair: "pair == (\a b. Abs (a, b))" + and dest1: "dest1 == (\p. fst (Rep p))" + and dest2: "dest2 == (\p. snd (Rep p))" -lemma product_type_typedef: - "product_type Rep Abs pair dest1 dest2 ==> type_definition Rep Abs UNIV" - by (unfold product_type_def) blast - -lemma product_type_pair: - "product_type Rep Abs pair dest1 dest2 ==> pair a b = Abs (a, b)" - by (unfold product_type_def) blast +lemmas product_typeI = + product_type.intro [OF product_type_axioms.intro] -lemma product_type_dest1: - "product_type Rep Abs pair dest1 dest2 ==> dest1 p = fst (Rep p)" - by (unfold product_type_def) blast - -lemma product_type_dest2: - "product_type Rep Abs pair dest1 dest2 ==> dest2 p = snd (Rep p)" - by (unfold product_type_def) blast - +lemma (in product_type) + "inject": "(pair x y = pair x' y') = (x = x' \ y = y')" + by (simp add: pair type_definition.Abs_inject [OF "typedef"]) -theorem product_type_inject: - "product_type Rep Abs pair dest1 dest2 ==> - (pair x y = pair x' y') = (x = x' \ y = y')" -proof - - case rule_context - show ?thesis - by (simp add: product_type_pair [OF rule_context] - Abs_inject [OF product_type_typedef [OF rule_context]]) -qed +lemma (in product_type) conv1: "dest1 (pair x y) = x" + by (simp add: pair dest1 type_definition.Abs_inverse [OF "typedef"]) -theorem product_type_conv1: - "product_type Rep Abs pair dest1 dest2 ==> dest1 (pair x y) = x" -proof - - case rule_context - show ?thesis - by (simp add: product_type_pair [OF rule_context] - product_type_dest1 [OF rule_context] - Abs_inverse [OF product_type_typedef [OF rule_context]]) -qed +lemma (in product_type) conv2: "dest2 (pair x y) = y" + by (simp add: pair dest2 type_definition.Abs_inverse [OF "typedef"]) -theorem product_type_conv2: - "product_type Rep Abs pair dest1 dest2 ==> dest2 (pair x y) = y" -proof - - case rule_context - show ?thesis - by (simp add: product_type_pair [OF rule_context] - product_type_dest2 [OF rule_context] - Abs_inverse [OF product_type_typedef [OF rule_context]]) -qed - -theorem product_type_induct [induct set: product_type]: - "product_type Rep Abs pair dest1 dest2 ==> - (!!x y. P (pair x y)) ==> P p" -proof - - assume hyp: "!!x y. P (pair x y)" - assume prod_type: "product_type Rep Abs pair dest1 dest2" - show "P p" - proof (rule Abs_induct [OF product_type_typedef [OF prod_type]]) - fix pair show "P (Abs pair)" - proof (rule prod_induct) - fix x y from hyp show "P (Abs (x, y))" - by (simp only: product_type_pair [OF prod_type]) - qed +lemma (in product_type) induct [induct type]: + assumes hyp: "!!x y. P (pair x y)" + shows "P p" +proof (rule type_definition.Abs_induct [OF "typedef"]) + fix q show "P (Abs q)" + proof (induct q) + fix x y have "P (pair x y)" by (rule hyp) + also have "pair x y = Abs (x, y)" by (simp only: pair) + finally show "P (Abs (x, y))" . qed qed -theorem product_type_cases [cases set: product_type]: - "product_type Rep Abs pair dest1 dest2 ==> - (!!x y. p = pair x y ==> C) ==> C" -proof - - assume prod_type: "product_type Rep Abs pair dest1 dest2" - assume "!!x y. p = pair x y ==> C" - with prod_type show C - by (induct p) (simp only: product_type_inject [OF prod_type], blast) -qed +lemma (in product_type) cases [cases type]: + "(!!x y. p = pair x y ==> C) ==> C" + by (induct p) (auto simp add: "inject") -theorem product_type_surjective_pairing: - "product_type Rep Abs pair dest1 dest2 ==> - p = pair (dest1 p) (dest2 p)" -proof - - case rule_context - thus ?thesis by (induct p) - (simp add: product_type_conv1 [OF rule_context] product_type_conv2 [OF rule_context]) -qed +lemma (in product_type) surjective_pairing: + "p = pair (dest1 p) (dest2 p)" + by (induct p) (simp only: conv1 conv2) -theorem product_type_split_paired_all: - "product_type Rep Abs pair dest1 dest2 ==> - (!!x. PROP P x) == (!!a b. PROP P (pair a b))" +lemma (in product_type) split_paired_all: + "(!!x. PROP P x) == (!!a b. PROP P (pair a b))" proof fix a b assume "!!x. PROP P x" thus "PROP P (pair a b)" . next - case rule_context fix x assume "!!a b. PROP P (pair a b)" hence "PROP P (pair (dest1 x) (dest2 x))" . - thus "PROP P x" by (simp only: product_type_surjective_pairing [OF rule_context, symmetric]) + thus "PROP P x" by (simp only: surjective_pairing [symmetric]) qed diff -r 181a293aa37a -r 666137b488a4 src/HOL/Typedef.thy --- a/src/HOL/Typedef.thy Wed Jul 24 00:09:44 2002 +0200 +++ b/src/HOL/Typedef.thy Wed Jul 24 00:10:52 2002 +0200 @@ -8,105 +8,79 @@ theory Typedef = Set files ("Tools/typedef_package.ML"): -constdefs - type_definition :: "('a => 'b) => ('b => 'a) => 'b set => bool" - "type_definition Rep Abs A == - (\x. Rep x \ A) \ - (\x. Abs (Rep x) = x) \ - (\y \ A. Rep (Abs y) = y)" - -- {* This will be stated as an axiom for each typedef! *} +locale type_definition = + fixes Rep and Abs and A + assumes Rep: "Rep x \ A" + and Rep_inverse: "Abs (Rep x) = x" + and Abs_inverse: "y \ A ==> Rep (Abs y) = y" + -- {* This will be axiomatized for each typedef! *} -lemma type_definitionI [intro]: - "(!!x. Rep x \ A) ==> - (!!x. Abs (Rep x) = x) ==> - (!!y. y \ A ==> Rep (Abs y) = y) ==> - type_definition Rep Abs A" - by (unfold type_definition_def) blast - -theorem Rep: "type_definition Rep Abs A ==> Rep x \ A" - by (unfold type_definition_def) blast +lemmas type_definitionI [intro] = + type_definition.intro [OF type_definition_axioms.intro] -theorem Rep_inverse: "type_definition Rep Abs A ==> Abs (Rep x) = x" - by (unfold type_definition_def) blast - -theorem Abs_inverse: "type_definition Rep Abs A ==> y \ A ==> Rep (Abs y) = y" - by (unfold type_definition_def) blast +lemma (in type_definition) Rep_inject: + "(Rep x = Rep y) = (x = y)" +proof + assume "Rep x = Rep y" + hence "Abs (Rep x) = Abs (Rep y)" by (simp only:) + also have "Abs (Rep x) = x" by (rule Rep_inverse) + also have "Abs (Rep y) = y" by (rule Rep_inverse) + finally show "x = y" . +next + assume "x = y" + thus "Rep x = Rep y" by (simp only:) +qed -theorem Rep_inject: "type_definition Rep Abs A ==> (Rep x = Rep y) = (x = y)" -proof - - assume tydef: "type_definition Rep Abs A" - show ?thesis - proof - assume "Rep x = Rep y" - hence "Abs (Rep x) = Abs (Rep y)" by (simp only:) - thus "x = y" by (simp only: Rep_inverse [OF tydef]) - next - assume "x = y" - thus "Rep x = Rep y" by simp - qed +lemma (in type_definition) Abs_inject: + assumes x: "x \ A" and y: "y \ A" + shows "(Abs x = Abs y) = (x = y)" +proof + assume "Abs x = Abs y" + hence "Rep (Abs x) = Rep (Abs y)" by (simp only:) + also from x have "Rep (Abs x) = x" by (rule Abs_inverse) + also from y have "Rep (Abs y) = y" by (rule Abs_inverse) + finally show "x = y" . +next + assume "x = y" + thus "Abs x = Abs y" by (simp only:) qed -theorem Abs_inject: - "type_definition Rep Abs A ==> x \ A ==> y \ A ==> (Abs x = Abs y) = (x = y)" -proof - - assume tydef: "type_definition Rep Abs A" - assume x: "x \ A" and y: "y \ A" - show ?thesis - proof - assume "Abs x = Abs y" - hence "Rep (Abs x) = Rep (Abs y)" by simp - moreover from x have "Rep (Abs x) = x" by (rule Abs_inverse [OF tydef]) - moreover from y have "Rep (Abs y) = y" by (rule Abs_inverse [OF tydef]) - ultimately show "x = y" by (simp only:) - next - assume "x = y" - thus "Abs x = Abs y" by simp - qed +lemma (in type_definition) Rep_cases [cases set]: + assumes y: "y \ A" + and hyp: "!!x. y = Rep x ==> P" + shows P +proof (rule hyp) + from y have "Rep (Abs y) = y" by (rule Abs_inverse) + thus "y = Rep (Abs y)" .. qed -theorem Rep_cases: - "type_definition Rep Abs A ==> y \ A ==> (!!x. y = Rep x ==> P) ==> P" -proof - - assume tydef: "type_definition Rep Abs A" - assume y: "y \ A" and r: "(!!x. y = Rep x ==> P)" - show P - proof (rule r) - from y have "Rep (Abs y) = y" by (rule Abs_inverse [OF tydef]) - thus "y = Rep (Abs y)" .. - qed +lemma (in type_definition) Abs_cases [cases type]: + assumes r: "!!y. x = Abs y ==> y \ A ==> P" + shows P +proof (rule r) + have "Abs (Rep x) = x" by (rule Rep_inverse) + thus "x = Abs (Rep x)" .. + show "Rep x \ A" by (rule Rep) qed -theorem Abs_cases: - "type_definition Rep Abs A ==> (!!y. x = Abs y ==> y \ A ==> P) ==> P" +lemma (in type_definition) Rep_induct [induct set]: + assumes y: "y \ A" + and hyp: "!!x. P (Rep x)" + shows "P y" proof - - assume tydef: "type_definition Rep Abs A" - assume r: "!!y. x = Abs y ==> y \ A ==> P" - show P - proof (rule r) - have "Abs (Rep x) = x" by (rule Rep_inverse [OF tydef]) - thus "x = Abs (Rep x)" .. - show "Rep x \ A" by (rule Rep [OF tydef]) - qed + have "P (Rep (Abs y))" by (rule hyp) + also from y have "Rep (Abs y) = y" by (rule Abs_inverse) + finally show "P y" . qed -theorem Rep_induct: - "type_definition Rep Abs A ==> y \ A ==> (!!x. P (Rep x)) ==> P y" +lemma (in type_definition) Abs_induct [induct type]: + assumes r: "!!y. y \ A ==> P (Abs y)" + shows "P x" proof - - assume tydef: "type_definition Rep Abs A" - assume "!!x. P (Rep x)" hence "P (Rep (Abs y))" . - moreover assume "y \ A" hence "Rep (Abs y) = y" by (rule Abs_inverse [OF tydef]) - ultimately show "P y" by (simp only:) -qed - -theorem Abs_induct: - "type_definition Rep Abs A ==> (!!y. y \ A ==> P (Abs y)) ==> P x" -proof - - assume tydef: "type_definition Rep Abs A" - assume r: "!!y. y \ A ==> P (Abs y)" - have "Rep x \ A" by (rule Rep [OF tydef]) + have "Rep x \ A" by (rule Rep) hence "P (Abs (Rep x))" by (rule r) - moreover have "Abs (Rep x) = x" by (rule Rep_inverse [OF tydef]) - ultimately show "P x" by (simp only:) + also have "Abs (Rep x) = x" by (rule Rep_inverse) + finally show "P x" . qed use "Tools/typedef_package.ML"