simple-counterexamples/SimpleCounterexamples/RunMEvil.lean
2025-04-08 21:54:27 -04:00

24 lines
705 B
Text
Raw Blame History

This file contains ambiguous Unicode characters

This file contains Unicode characters that might be confused with other characters. If you think that this is intentional, you can safely ignore this warning. Use the Escape button to reveal them.

structure EmptyM (_ : Type u) : Type v
instance : Monad EmptyM where
pure _ := ⟨⟩
bind _ _ := ⟨⟩
theorem runMEvil (runM : ∀α m, [Monad m] → m αα) : False :=
runM PEmpty EmptyM ⟨⟩ |>.elim
instance : LawfulMonad EmptyM where
map_const := by intros; trivial
id_map := by intros; trivial
seqLeft_eq := by intros; trivial
seqRight_eq := by intros; trivial
pure_seq := by intros; trivial
bind_pure_comp := by intros; trivial
bind_map := by intros; trivial
pure_bind := by intros; trivial
bind_assoc := by intros; trivial
theorem lawfulRunMEvil (runM : ∀α m, [Monad m] → [LawfulMonad m] → m αα)
: False :=
runM PEmpty EmptyM ⟨⟩ |>.elim