diff -r cdf84288d93c -r 970b9ab6c439 src/FOL/ex/Intuitionistic.thy --- a/src/FOL/ex/Intuitionistic.thy Sat Apr 02 17:03:35 2022 +0000 +++ b/src/FOL/ex/Intuitionistic.thy Sun Apr 03 14:48:55 2022 +0100 @@ -65,6 +65,14 @@ \ (((F \ A) \ B) \ I) \ E\ by (tactic \IntPr.fast_tac \<^context> 1\) +text \Admissibility of the excluded middle for negated formulae\ +lemma \(P \ \P \ \Q) \ \Q\ + by (tactic \IntPr.fast_tac \<^context> 1\) + +text \The same in a more general form, no ex falso quodlibet\ +lemma \(P \ (P\R) \ Q \ R) \ Q \ R\ + by (tactic \IntPr.fast_tac \<^context> 1\) + subsection \Lemmas for the propositional double-negation translation\