diff -r 787b69fffaae -r 46a94aa3ec8e src/HOL/Metis_Examples/Abstraction.thy --- a/src/HOL/Metis_Examples/Abstraction.thy Mon Jan 31 16:09:23 2022 +0100 +++ b/src/HOL/Metis_Examples/Abstraction.thy Mon Jan 31 16:09:23 2022 +0100 @@ -8,14 +8,14 @@ section \Example Featuring Metis's Support for Lambda-Abstractions\ theory Abstraction -imports "HOL-Library.FuncSet" + imports "HOL-Library.FuncSet" begin (* For Christoph Benzmüller *) lemma "x < 1 \ ((=) = (=)) \ ((=) = (=)) \ x < (2::nat)" by (metis nat_1_add_1 trans_less_add2) -lemma "((=) ) = (\x y. y = x)" +lemma "(=) = (\x y. y = x)" by metis consts