# HG changeset patch # User wenzelm # Date 1734019637 -3600 # Node ID 257f93d40d7c3dbb028e08a0ea2b37edd0d10b34 # Parent adbd2e1407cc8362ca1ea1d519994343290dad0a tuned proofs; diff -r adbd2e1407cc -r 257f93d40d7c src/HOL/Typedef.thy --- a/src/HOL/Typedef.thy Thu Dec 12 17:07:09 2024 +0100 +++ b/src/HOL/Typedef.thy Thu Dec 12 17:07:17 2024 +0100 @@ -23,12 +23,11 @@ proof assume "Rep x = Rep y" then have "Abs (Rep x) = Abs (Rep y)" by (simp only:) - moreover have "Abs (Rep x) = x" by (rule Rep_inverse) - moreover have "Abs (Rep y) = y" by (rule Rep_inverse) - ultimately show "x = y" by simp + 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" - then show "Rep x = Rep y" by (simp only:) + show "x = y \ Rep x = Rep y" by (simp only:) qed lemma Abs_inject: @@ -37,12 +36,11 @@ proof assume "Abs x = Abs y" then have "Rep (Abs x) = Rep (Abs y)" by (simp only:) - moreover from \x \ A\ have "Rep (Abs x) = x" by (rule Abs_inverse) - moreover from \y \ A\ have "Rep (Abs y) = y" by (rule Abs_inverse) - ultimately show "x = y" by simp + also from \x \ A\ have "Rep (Abs x) = x" by (rule Abs_inverse) + also from \y \ A\ have "Rep (Abs y) = y" by (rule Abs_inverse) + finally show "x = y" . next - assume "x = y" - then show "Abs x = Abs y" by (simp only:) + show "x = y \ Abs x = Abs y" by (simp only:) qed lemma Rep_cases [cases set]: @@ -69,8 +67,8 @@ shows "P y" proof - have "P (Rep (Abs y))" by (rule hyp) - moreover from y have "Rep (Abs y) = y" by (rule Abs_inverse) - ultimately show "P y" by simp + also from y have "Rep (Abs y) = y" by (rule Abs_inverse) + finally show "P y" . qed lemma Abs_induct [induct type]: @@ -79,8 +77,8 @@ proof - have "Rep x \ A" by (rule Rep) then have "P (Abs (Rep x))" by (rule r) - moreover have "Abs (Rep x) = x" by (rule Rep_inverse) - ultimately show "P x" by simp + also have "Abs (Rep x) = x" by (rule Rep_inverse) + finally show "P x" . qed lemma Rep_range: "range Rep = A" @@ -99,10 +97,11 @@ show "Abs ` A \ UNIV" by (rule subset_UNIV) show "UNIV \ Abs ` A" proof - fix x - have "x = Abs (Rep x)" by (rule Rep_inverse [symmetric]) - moreover have "Rep x \ A" by (rule Rep) - ultimately show "x \ Abs ` A" by (rule image_eqI) + show "x \ Abs ` A" for x + proof (rule image_eqI) + show "x = Abs (Rep x)" by (rule Rep_inverse [symmetric]) + show "Rep x \ A" by (rule Rep) + qed qed qed