|
9 | 9 | public import Init.Control.Lawful.Basic |
10 | 10 | public import Init.Control.Except |
11 | 11 | import all Init.Control.Except |
| 12 | +public import Init.Control.Option |
| 13 | +import all Init.Control.Option |
12 | 14 | public import Init.Control.State |
13 | 15 | import all Init.Control.State |
14 | 16 | public import Init.Control.StateRef |
@@ -38,14 +40,8 @@ namespace ExceptT |
38 | 40 | @[simp, grind =] theorem bind_throw [Monad m] [LawfulMonad m] (f : α → ExceptT ε m β) : (throw e >>= f) = throw e := by |
39 | 41 | simp [throw, throwThe, MonadExceptOf.throw, bind, ExceptT.bind, ExceptT.bindCont, ExceptT.mk] |
40 | 42 |
|
41 | | -@[grind =] |
42 | | -theorem run_bind [Monad m] (x : ExceptT ε m α) (f : α → ExceptT ε m β) |
43 | | - : run (x >>= f : ExceptT ε m β) |
44 | | - = |
45 | | - run x >>= fun |
46 | | - | Except.ok x => run (f x) |
47 | | - | Except.error e => pure (Except.error e) := |
48 | | - rfl |
| 43 | +@[simp, grind =] theorem run_bind (f : α → OptionT m β) [Monad m] : |
| 44 | + (x >>= f).run = Option.elimM x.run (pure none) (fun x => (f x).run) := rfl |
49 | 45 |
|
50 | 46 | @[simp, grind =] theorem lift_pure [Monad m] [LawfulMonad m] (a : α) : ExceptT.lift (pure a) = (pure a : ExceptT ε m α) := by |
51 | 47 | simp [ExceptT.lift, pure, ExceptT.pure] |
@@ -110,6 +106,101 @@ instance : LawfulMonad (Except ε) := LawfulMonad.mk' |
110 | 106 | instance : LawfulApplicative (Except ε) := inferInstance |
111 | 107 | instance : LawfulFunctor (Except ε) := inferInstance |
112 | 108 |
|
| 109 | +/-! # OptionT -/ |
| 110 | + |
| 111 | +namespace OptionT |
| 112 | + |
| 113 | +@[ext] theorem ext {x y : OptionT m α} (h : x.run = y.run) : x = y := by |
| 114 | + simp [run] at h |
| 115 | + assumption |
| 116 | + |
| 117 | +@[simp] theorem run_pure [Monad m] (x : α) : run (pure x : OptionT m α) = pure (some x) := by |
| 118 | + simp [run, pure, OptionT.pure, OptionT.mk] |
| 119 | + |
| 120 | +@[simp] theorem run_lift [Monad.{u, v} m] (x : m α) : run (OptionT.lift x : OptionT m α) = (return some (← x) : m (Option α)) := by |
| 121 | + simp [run, OptionT.lift, OptionT.mk] |
| 122 | + |
| 123 | +@[simp] theorem run_throw [Monad m] : run (throw e : OptionT m β) = pure none := by |
| 124 | + simp [run, throw, throwThe, MonadExceptOf.throw, OptionT.fail, OptionT.mk] |
| 125 | + |
| 126 | +@[simp] theorem run_bind_lift [Monad m] [LawfulMonad m] (x : m α) (f : α → OptionT m β) : run (OptionT.lift x >>= f : OptionT m β) = x >>= fun a => run (f a) := by |
| 127 | + simp [OptionT.run, OptionT.lift, bind, OptionT.bind, OptionT.mk] |
| 128 | + |
| 129 | +@[simp] theorem bind_throw [Monad m] [LawfulMonad m] (f : α → OptionT m β) : (throw e >>= f) = throw e := by |
| 130 | + simp [throw, throwThe, MonadExceptOf.throw, bind, OptionT.bind, OptionT.mk, OptionT.fail] |
| 131 | + |
| 132 | +theorem run_bind [Monad m] (x : OptionT m α) |
| 133 | + : run (x >>= f : OptionT m β) |
| 134 | + = |
| 135 | + run x >>= fun |
| 136 | + | some x => run (f x) |
| 137 | + | none => pure none := by |
| 138 | + simp only [run, bind, OptionT.bind, OptionT.mk] |
| 139 | + congr |
| 140 | + |
| 141 | +@[simp] theorem lift_pure [Monad m] [LawfulMonad m] {α : Type u} (a : α) : OptionT.lift (pure a : m α) = pure a := by |
| 142 | + simp only [OptionT.lift, OptionT.mk, bind_pure_comp, map_pure, pure, OptionT.pure] |
| 143 | + |
| 144 | +@[simp] theorem run_map [Monad m] [LawfulMonad m] (f : α → β) (x : OptionT m α) |
| 145 | + : (f <$> x).run = Option.map f <$> x.run := by |
| 146 | + simp [Functor.map, Option.map, ←bind_pure_comp] |
| 147 | + apply bind_congr |
| 148 | + intro a; cases a <;> simp [OptionT.pure, OptionT.mk] |
| 149 | + |
| 150 | +protected theorem seq_eq {α β : Type u} [Monad m] (mf : OptionT m (α → β)) (x : OptionT m α) : mf <*> x = mf >>= fun f => f <$> x := |
| 151 | + rfl |
| 152 | + |
| 153 | +protected theorem bind_pure_comp [Monad m] (f : α → β) (x : OptionT m α) : x >>= pure ∘ f = f <$> x := by |
| 154 | + intros; rfl |
| 155 | + |
| 156 | +protected theorem seqLeft_eq {α β : Type u} {m : Type u → Type v} [Monad m] [LawfulMonad m] (x : OptionT m α) (y : OptionT m β) : x <* y = const β <$> x <*> y := by |
| 157 | + change (x >>= fun a => y >>= fun _ => pure a) = (const (α := α) β <$> x) >>= fun f => f <$> y |
| 158 | + rw [← OptionT.bind_pure_comp] |
| 159 | + apply ext |
| 160 | + simp [run_bind] |
| 161 | + apply bind_congr |
| 162 | + intro |
| 163 | + | none => simp |
| 164 | + | some _ => |
| 165 | + simp [←bind_pure_comp]; apply bind_congr; intro b; |
| 166 | + cases b <;> simp [const] |
| 167 | + |
| 168 | +protected theorem seqRight_eq [Monad m] [LawfulMonad m] (x : OptionT m α) (y : OptionT m β) : x *> y = const α id <$> x <*> y := by |
| 169 | + change (x >>= fun _ => y) = (const α id <$> x) >>= fun f => f <$> y |
| 170 | + rw [← OptionT.bind_pure_comp] |
| 171 | + apply ext |
| 172 | + simp [run_bind] |
| 173 | + apply bind_congr |
| 174 | + intro a; cases a <;> simp |
| 175 | + |
| 176 | +instance [Monad m] [LawfulMonad m] : LawfulMonad (OptionT m) where |
| 177 | + id_map := by intros; apply ext; simp |
| 178 | + map_const := by intros; rfl |
| 179 | + seqLeft_eq := OptionT.seqLeft_eq |
| 180 | + seqRight_eq := OptionT.seqRight_eq |
| 181 | + pure_seq := by intros; apply ext; simp [OptionT.seq_eq, run_bind] |
| 182 | + bind_pure_comp := OptionT.bind_pure_comp |
| 183 | + bind_map := by intros; rfl |
| 184 | + pure_bind := by intros; apply ext; simp [run_bind] |
| 185 | + bind_assoc := by intros; apply ext; simp [run_bind]; apply bind_congr; intro a; cases a <;> simp |
| 186 | + |
| 187 | +@[simp] theorem map_failure [Monad m] [LawfulMonad m] {α β : Type _} (f : α → β) : |
| 188 | + f <$> (failure : OptionT m α) = (failure : OptionT m β) := by |
| 189 | + simp [OptionT.mk, Functor.map, Alternative.failure, OptionT.fail, OptionT.bind] |
| 190 | + |
| 191 | +end OptionT |
| 192 | + |
| 193 | +/-! # Option -/ |
| 194 | + |
| 195 | +instance : LawfulMonad Option := LawfulMonad.mk' |
| 196 | + (id_map := fun x => by cases x <;> rfl) |
| 197 | + (pure_bind := fun _ _ => by rfl) |
| 198 | + (bind_assoc := fun a _ _ => by cases a <;> rfl) |
| 199 | + (bind_pure_comp := bind_pure_comp) |
| 200 | + |
| 201 | +instance : LawfulApplicative Option := inferInstance |
| 202 | +instance : LawfulFunctor Option := inferInstance |
| 203 | + |
113 | 204 | /-! # ReaderT -/ |
114 | 205 |
|
115 | 206 | namespace ReaderT |
|
0 commit comments