diff -r 3377a830b727 -r eccbfaf2bc0e src/HOL/ex/PER.thy --- a/src/HOL/ex/PER.thy Sun Oct 01 18:29:28 2006 +0200 +++ b/src/HOL/ex/PER.thy Sun Oct 01 18:29:30 2006 +0200 @@ -49,10 +49,10 @@ "domain = {x. x \ x}" lemma domainI [intro]: "x \ x ==> x \ domain" - by (unfold domain_def) blast + unfolding domain_def by blast lemma domainD [dest]: "x \ domain ==> x \ x" - by (unfold domain_def) blast + unfolding domain_def by blast theorem domainI' [elim?]: "x \ y ==> x \ domain" proof @@ -75,18 +75,18 @@ lemma partial_equiv_funI [intro?]: "(!!x y. x \ domain ==> y \ domain ==> x \ y ==> f x \ g y) ==> f \ g" - by (unfold eqv_fun_def) blast + unfolding eqv_fun_def by blast lemma partial_equiv_funD [dest?]: "f \ g ==> x \ domain ==> y \ domain ==> x \ y ==> f x \ g y" - by (unfold eqv_fun_def) blast + unfolding eqv_fun_def by blast text {* The class of partial equivalence relations is closed under function spaces (in \emph{both} argument positions). *} -instance fun :: (partial_equiv, partial_equiv) partial_equiv +instance "fun" :: (partial_equiv, partial_equiv) partial_equiv proof fix f g h :: "'a::partial_equiv => 'b::partial_equiv" assume fg: "f \ g" @@ -94,9 +94,9 @@ proof fix x y :: 'a assume x: "x \ domain" and y: "y \ domain" - assume "x \ y" hence "y \ x" .. + assume "x \ y" then have "y \ x" .. with fg y x have "f y \ g x" .. - thus "g x \ f y" .. + then show "g x \ f y" .. qed assume gh: "g \ h" show "f \ h" @@ -154,10 +154,10 @@ by blast lemma quotI [intro]: "{x. a \ x} \ quot" - by (unfold quot_def) blast + unfolding quot_def by blast lemma quotE [elim]: "R \ quot ==> (!!a. R = {x. a \ x} ==> C) ==> C" - by (unfold quot_def) blast + unfolding quot_def by blast text {* \medskip Abstracted equivalence classes are the canonical @@ -171,14 +171,14 @@ theorem quot_rep: "\a. A = \a\" proof (cases A) fix R assume R: "A = Abs_quot R" - assume "R \ quot" hence "\a. R = {x. a \ x}" by blast + assume "R \ quot" then have "\a. R = {x. a \ x}" by blast with R have "\a. A = Abs_quot {x. a \ x}" by blast - thus ?thesis by (unfold eqv_class_def) + then show ?thesis by (unfold eqv_class_def) qed -lemma quot_cases [case_names rep, cases type: quot]: - "(!!a. A = \a\ ==> C) ==> C" - by (insert quot_rep) blast +lemma quot_cases [cases type: quot]: + obtains (rep) a where "A = \a\" + using quot_rep by blast subsection {* Equality on quotients *} @@ -204,17 +204,17 @@ finally show "a \ x" . qed qed - thus ?thesis by (simp only: eqv_class_def) + then show ?thesis by (simp only: eqv_class_def) qed theorem eqv_class_eqD' [dest?]: "\a\ = \b\ ==> a \ domain ==> a \ b" proof (unfold eqv_class_def) assume "Abs_quot {x. a \ x} = Abs_quot {x. b \ x}" - hence "{x. a \ x} = {x. b \ x}" by (simp only: Abs_quot_inject quotI) - moreover assume "a \ domain" hence "a \ a" .. + then have "{x. a \ x} = {x. b \ x}" by (simp only: Abs_quot_inject quotI) + moreover assume "a \ domain" then have "a \ a" .. ultimately have "a \ {x. b \ x}" by blast - hence "b \ a" by blast - thus "a \ b" .. + then have "b \ a" by blast + then show "a \ b" .. qed theorem eqv_class_eqD [dest?]: "\a\ = \b\ ==> a \ (b::'a::equiv)" @@ -223,10 +223,10 @@ qed lemma eqv_class_eq' [simp]: "a \ domain ==> (\a\ = \b\) = (a \ b)" - by (insert eqv_class_eqI eqv_class_eqD') blast + using eqv_class_eqI eqv_class_eqD' by blast lemma eqv_class_eq [simp]: "(\a\ = \b\) = (a \ (b::'a::equiv))" - by (insert eqv_class_eqI eqv_class_eqD) blast + using eqv_class_eqI eqv_class_eqD by blast subsection {* Picking representing elements *} @@ -242,8 +242,8 @@ proof (rule someI2) show "\a\ = \a\" .. fix x assume "\a\ = \x\" - hence "a \ x" .. - thus "x \ a" .. + then have "a \ x" .. + then show "x \ a" .. qed qed @@ -255,8 +255,8 @@ theorem pick_inverse: "\pick A\ = (A::'a::equiv quot)" proof (cases A) fix a assume a: "A = \a\" - hence "pick A \ a" by simp - hence "\pick A\ = \a\" by simp + then have "pick A \ a" by simp + then have "\pick A\ = \a\" by simp with a show ?thesis by simp qed