diff -r 8d2d68c8c051 -r 95c9af7483b1 src/HOL/Algebra/FiniteProduct.thy --- a/src/HOL/Algebra/FiniteProduct.thy Wed Nov 13 15:14:48 2024 +0100 +++ b/src/HOL/Algebra/FiniteProduct.thy Wed Nov 13 20:10:34 2024 +0100 @@ -115,8 +115,7 @@ then show ?thesis using \(A,y) \ foldSetD D (\) e\ by auto next - case (2 x' A' y') - note A' = this + case A': (2 x' A' y') show ?thesis using foldSetD.cases [OF \(A,y) \ foldSetD D (\) e\] proof cases @@ -124,8 +123,7 @@ then show ?thesis using \(A,x) \ foldSetD D (\) e\ by auto next - case (2 x'' A'' y'') - note A'' = this + case A'': (2 x'' A'' y'') show ?thesis proof (cases "x' = x''") case True