@@ -261,6 +261,22 @@ protected theorem well_founded : ∀ (f : r ↪r s) (h : well_founded s), well_f
261
261
protected theorem is_well_order : ∀ (f : r ↪r s) [is_well_order β s], is_well_order α r
262
262
| f H := by exactI {wf := f.well_founded H.wf, ..f.is_strict_total_order'}
263
263
264
+ /--
265
+ To define an relation embedding from an antisymmetric relation `r` to a reflexive relation `s` it
266
+ suffices to give a function together with a proof that it satisfies `s (f a) (f b) ↔ r a b`.
267
+ -/
268
+ def of_map_rel_iff (f : α → β) [is_antisymm α r] [is_refl β s]
269
+ (hf : ∀ a b, s (f a) (f b) ↔ r a b) : r ↪r s :=
270
+ { to_fun := f,
271
+ inj' := λ x y h, antisymm ((hf _ _).1 (h ▸ refl _)) ((hf _ _).1 (h ▸ refl _)),
272
+ map_rel_iff' := hf }
273
+
274
+ @[simp]
275
+ lemma of_map_rel_iff_coe (f : α → β) [is_antisymm α r] [is_refl β s]
276
+ (hf : ∀ a b, s (f a) (f b) ↔ r a b) :
277
+ ⇑(of_map_rel_iff f hf : r ↪r s) = f :=
278
+ rfl
279
+
264
280
/-- It suffices to prove `f` is monotone between strict relations
265
281
to show it is a relation embedding. -/
266
282
def of_monotone [is_trichotomous α r] [is_asymm β s] (f : α → β)
@@ -326,12 +342,21 @@ f.lt_embedding.is_well_order
326
342
protected def dual : order_dual α ↪o order_dual β :=
327
343
⟨f.to_embedding, λ a b, f.map_rel_iff⟩
328
344
329
- /-- A sctrictly monotone map from a linear order is an order embedding. --/
345
+ /--
346
+ To define an order embedding from a partial order to a preorder it suffices to give a function
347
+ together with a proof that it satisfies `f a ≤ f b ↔ a ≤ b`.
348
+ -/
349
+ def of_map_rel_iff {α β} [partial_order α] [preorder β] (f : α → β)
350
+ (hf : ∀ a b, f a ≤ f b ↔ a ≤ b) : α ↪o β :=
351
+ rel_embedding.of_map_rel_iff f hf
352
+
353
+ @[simp] lemma coe_of_map_rel_iff {α β} [partial_order α] [preorder β] {f : α → β} (h) :
354
+ ⇑(of_map_rel_iff f h) = f := rfl
355
+
356
+ /-- A strictly monotone map from a linear order is an order embedding. --/
330
357
def of_strict_mono {α β} [linear_order α] [preorder β] (f : α → β)
331
358
(h : strict_mono f) : α ↪o β :=
332
- { to_fun := f,
333
- inj' := strict_mono.injective h,
334
- map_rel_iff' := λ a b, h.le_iff_le }
359
+ of_map_rel_iff f (λ _ _, h.le_iff_le)
335
360
336
361
@[simp] lemma coe_of_strict_mono {α β} [linear_order α] [preorder β] {f : α → β}
337
362
(h : strict_mono f) : ⇑(of_strict_mono f h) = f := rfl
0 commit comments