|
| 1 | +/- |
| 2 | +Copyright (c) 2022 Yaël Dillies. All rights reserved. |
| 3 | +Released under Apache 2.0 license as described in the file LICENSE. |
| 4 | +Authors: Yaël Dillies |
| 5 | +-/ |
| 6 | +import data.option.basic |
| 7 | + |
| 8 | +/-! |
| 9 | +# Binary map of options |
| 10 | +
|
| 11 | +This file defines the binary map of `option`. This is mostly useful to define pointwise operations |
| 12 | +on intervals. |
| 13 | +
|
| 14 | +## Main declarations |
| 15 | +
|
| 16 | +* `option.map₂`: Binary map of options. |
| 17 | +
|
| 18 | +## Notes |
| 19 | +
|
| 20 | +This file is very similar to the n-ary section of `data.set.basic`, to `data.finset.n_ary` and to |
| 21 | +`order.filter.n_ary`. Please keep them in sync. |
| 22 | +
|
| 23 | +We do not define `option.map₃` as its only purpose so far would be to prove properties of |
| 24 | +`option.map₂` and casing already fulfills this task. |
| 25 | +-/ |
| 26 | + |
| 27 | +open function |
| 28 | + |
| 29 | +namespace option |
| 30 | +variables {α α' β β' γ γ' δ δ' ε ε' : Type*} {f : α → β → γ} {a : option α} {b : option β} |
| 31 | + {c : option γ} |
| 32 | + |
| 33 | +/-- The image of a binary function `f : α → β → γ` as a function `option α → option β → option γ`. |
| 34 | +Mathematically this should be thought of as the image of the corresponding function `α × β → γ`. -/ |
| 35 | +def map₂ (f : α → β → γ) (a : option α) (b : option β) : option γ := a.bind $ λ a, b.map $ f a |
| 36 | + |
| 37 | +/-- `option.map₂` in terms of monadic operations. Note that this can't be taken as the definition |
| 38 | +because of the lack of universe polymorphism. -/ |
| 39 | +lemma map₂_def {α β γ : Type*} (f : α → β → γ) (a : option α) (b : option β) : |
| 40 | + map₂ f a b = f <$> a <*> b := by cases a; refl |
| 41 | + |
| 42 | +@[simp] lemma map₂_some_some (f : α → β → γ) (a : α) (b : β) : map₂ f (some a) (some b) = f a b := |
| 43 | +rfl |
| 44 | +lemma map₂_coe_coe (f : α → β → γ) (a : α) (b : β) : map₂ f a b = f a b := rfl |
| 45 | +@[simp] lemma map₂_none_left (f : α → β → γ) (b : option β) : map₂ f none b = none := rfl |
| 46 | +@[simp] lemma map₂_none_right (f : α → β → γ) (a : option α) : map₂ f a none = none := |
| 47 | +by cases a; refl |
| 48 | +@[simp] lemma map₂_coe_left (f : α → β → γ) (a : α) (b : option β) : |
| 49 | + map₂ f a b = b.map (λ b, f a b) := rfl |
| 50 | +@[simp] lemma map₂_coe_right (f : α → β → γ) (a : option α) (b : β) : |
| 51 | + map₂ f a b = a.map (λ a, f a b) := rfl |
| 52 | + |
| 53 | +@[simp] lemma mem_map₂_iff {c : γ} : c ∈ map₂ f a b ↔ ∃ a' b', a' ∈ a ∧ b' ∈ b ∧ f a' b' = c := |
| 54 | +by simp [map₂] |
| 55 | + |
| 56 | +@[simp] lemma map₂_eq_none_iff : map₂ f a b = none ↔ a = none ∨ b = none := |
| 57 | +by cases a; cases b; simp |
| 58 | + |
| 59 | +lemma map₂_swap (f : α → β → γ) (a : option α) (b : option β) : |
| 60 | + map₂ f a b = map₂ (λ a b, f b a) b a := |
| 61 | +by cases a; cases b; refl |
| 62 | + |
| 63 | +lemma map_map₂ (f : α → β → γ) (g : γ → δ) : (map₂ f a b).map g = map₂ (λ a b, g (f a b)) a b := |
| 64 | +by cases a; cases b; refl |
| 65 | + |
| 66 | +lemma map₂_map_left (f : γ → β → δ) (g : α → γ) : |
| 67 | + map₂ f (a.map g) b = map₂ (λ a b, f (g a) b) a b := |
| 68 | +by cases a; refl |
| 69 | + |
| 70 | +lemma map₂_map_right (f : α → γ → δ) (g : β → γ) : |
| 71 | + map₂ f a (b.map g) = map₂ (λ a b, f a (g b)) a b := |
| 72 | +by cases b; refl |
| 73 | + |
| 74 | +@[simp] lemma map₂_curry (f : α × β → γ) (a : option α) (b : option β) : |
| 75 | + map₂ (curry f) a b = option.map f (map₂ prod.mk a b) := (map_map₂ _ _).symm |
| 76 | + |
| 77 | +@[simp] lemma map_uncurry (f : α → β → γ) (x : option (α × β)) : |
| 78 | + x.map (uncurry f) = map₂ f (x.map prod.fst) (x.map prod.snd) := by cases x; refl |
| 79 | + |
| 80 | +/-! |
| 81 | +### Algebraic replacement rules |
| 82 | +
|
| 83 | +A collection of lemmas to transfer associativity, commutativity, distributivity, ... of operations |
| 84 | +to the associativity, commutativity, distributivity, ... of `option.map₂` of those operations. |
| 85 | +The proof pattern is `map₂_lemma operation_lemma`. For example, `map₂_comm mul_comm` proves that |
| 86 | +`map₂ (*) a b = map₂ (*) g f` in a `comm_semigroup`. |
| 87 | +-/ |
| 88 | + |
| 89 | +lemma map₂_assoc {f : δ → γ → ε} {g : α → β → δ} {f' : α → ε' → ε} {g' : β → γ → ε'} |
| 90 | + (h_assoc : ∀ a b c, f (g a b) c = f' a (g' b c)) : |
| 91 | + map₂ f (map₂ g a b) c = map₂ f' a (map₂ g' b c) := |
| 92 | +by cases a; cases b; cases c; simp [h_assoc] |
| 93 | + |
| 94 | +lemma map₂_comm {g : β → α → γ} (h_comm : ∀ a b, f a b = g b a) : map₂ f a b = map₂ g b a := |
| 95 | +by cases a; cases b; simp [h_comm] |
| 96 | + |
| 97 | +lemma map₂_left_comm {f : α → δ → ε} {g : β → γ → δ} {f' : α → γ → δ'} {g' : β → δ' → ε} |
| 98 | + (h_left_comm : ∀ a b c, f a (g b c) = g' b (f' a c)) : |
| 99 | + map₂ f a (map₂ g b c) = map₂ g' b (map₂ f' a c) := |
| 100 | +by cases a; cases b; cases c; simp [h_left_comm] |
| 101 | + |
| 102 | +lemma map₂_right_comm {f : δ → γ → ε} {g : α → β → δ} {f' : α → γ → δ'} {g' : δ' → β → ε} |
| 103 | + (h_right_comm : ∀ a b c, f (g a b) c = g' (f' a c) b) : |
| 104 | + map₂ f (map₂ g a b) c = map₂ g' (map₂ f' a c) b := |
| 105 | +by cases a; cases b; cases c; simp [h_right_comm] |
| 106 | + |
| 107 | +lemma map_map₂_distrib {g : γ → δ} {f' : α' → β' → δ} {g₁ : α → α'} {g₂ : β → β'} |
| 108 | + (h_distrib : ∀ a b, g (f a b) = f' (g₁ a) (g₂ b)) : |
| 109 | + (map₂ f a b).map g = map₂ f' (a.map g₁) (b.map g₂) := |
| 110 | +by cases a; cases b; simp [h_distrib] |
| 111 | + |
| 112 | +/-! |
| 113 | +The following symmetric restatement are needed because unification has a hard time figuring all the |
| 114 | +functions if you symmetrize on the spot. This is also how the other n-ary APIs do it. |
| 115 | +-/ |
| 116 | + |
| 117 | +/-- Symmetric statement to `option.map₂_map_left_comm`. -/ |
| 118 | +lemma map_map₂_distrib_left {g : γ → δ} {f' : α' → β → δ} {g' : α → α'} |
| 119 | + (h_distrib : ∀ a b, g (f a b) = f' (g' a) b) : |
| 120 | + (map₂ f a b).map g = map₂ f' (a.map g') b := |
| 121 | +by cases a; cases b; simp [h_distrib] |
| 122 | + |
| 123 | +/-- Symmetric statement to `option.map_map₂_right_comm`. -/ |
| 124 | +lemma map_map₂_distrib_right {g : γ → δ} {f' : α → β' → δ} {g' : β → β'} |
| 125 | + (h_distrib : ∀ a b, g (f a b) = f' a (g' b)) : |
| 126 | + (map₂ f a b).map g = map₂ f' a (b.map g') := |
| 127 | +by cases a; cases b; simp [h_distrib] |
| 128 | + |
| 129 | +/-- Symmetric statement to `option.map_map₂_distrib_left`. -/ |
| 130 | +lemma map₂_map_left_comm {f : α' → β → γ} {g : α → α'} {f' : α → β → δ} {g' : δ → γ} |
| 131 | + (h_left_comm : ∀ a b, f (g a) b = g' (f' a b)) : |
| 132 | + map₂ f (a.map g) b = (map₂ f' a b).map g' := |
| 133 | +by cases a; cases b; simp [h_left_comm] |
| 134 | + |
| 135 | +/-- Symmetric statement to `option.map_map₂_distrib_right`. -/ |
| 136 | +lemma map_map₂_right_comm {f : α → β' → γ} {g : β → β'} {f' : α → β → δ} {g' : δ → γ} |
| 137 | + (h_right_comm : ∀ a b, f a (g b) = g' (f' a b)) : |
| 138 | + map₂ f a (b.map g) = (map₂ f' a b).map g' := |
| 139 | +by cases a; cases b; simp [h_right_comm] |
| 140 | + |
| 141 | +lemma map_map₂_antidistrib {g : γ → δ} {f' : β' → α' → δ} {g₁ : β → β'} {g₂ : α → α'} |
| 142 | + (h_antidistrib : ∀ a b, g (f a b) = f' (g₁ b) (g₂ a)) : |
| 143 | + (map₂ f a b).map g = map₂ f' (b.map g₁) (a.map g₂) := |
| 144 | +by cases a; cases b; simp [h_antidistrib] |
| 145 | + |
| 146 | +/-- Symmetric statement to `option.map₂_map_left_anticomm`. -/ |
| 147 | +lemma map_map₂_antidistrib_left {g : γ → δ} {f' : β' → α → δ} {g' : β → β'} |
| 148 | + (h_antidistrib : ∀ a b, g (f a b) = f' (g' b) a) : |
| 149 | + (map₂ f a b).map g = map₂ f' (b.map g') a := |
| 150 | +by cases a; cases b; simp [h_antidistrib] |
| 151 | + |
| 152 | +/-- Symmetric statement to `option.map_map₂_right_anticomm`. -/ |
| 153 | +lemma map_map₂_antidistrib_right {g : γ → δ} {f' : β → α' → δ} {g' : α → α'} |
| 154 | + (h_antidistrib : ∀ a b, g (f a b) = f' b (g' a)) : |
| 155 | + (map₂ f a b).map g = map₂ f' b (a.map g') := |
| 156 | +by cases a; cases b; simp [h_antidistrib] |
| 157 | + |
| 158 | +/-- Symmetric statement to `option.map_map₂_antidistrib_left`. -/ |
| 159 | +lemma map₂_map_left_anticomm {f : α' → β → γ} {g : α → α'} {f' : β → α → δ} {g' : δ → γ} |
| 160 | + (h_left_anticomm : ∀ a b, f (g a) b = g' (f' b a)) : |
| 161 | + map₂ f (a.map g) b = (map₂ f' b a).map g' := |
| 162 | +by cases a; cases b; simp [h_left_anticomm] |
| 163 | + |
| 164 | +/-- Symmetric statement to `option.map_map₂_antidistrib_right`. -/ |
| 165 | +lemma map_map₂_right_anticomm {f : α → β' → γ} {g : β → β'} {f' : β → α → δ} {g' : δ → γ} |
| 166 | + (h_right_anticomm : ∀ a b, f a (g b) = g' (f' b a)) : |
| 167 | + map₂ f a (b.map g) = (map₂ f' b a).map g' := |
| 168 | +by cases a; cases b; simp [h_right_anticomm] |
| 169 | + |
| 170 | +end option |
0 commit comments