Skip to content
This repository was archived by the owner on Jul 24, 2024. It is now read-only.

Commit 680733c

Browse files
committed
feat(order/hom/basic): compl as a dual order isomorphism (#11630)
1 parent ff241e1 commit 680733c

File tree

1 file changed

+23
-1
lines changed

1 file changed

+23
-1
lines changed

src/order/hom/basic.lean

Lines changed: 23 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -54,11 +54,13 @@ because the more bundled version usually does not work with dot notation.
5454
`order_dual α →o order_dual β`;
5555
* `order_hom.dual_iso`: order isomorphism between `α →o β` and
5656
`order_dual (order_dual α →o order_dual β)`;
57+
* `order_iso.compl`: order isomorphism `α ≃o order_dual α` given by taking complements in a
58+
boolean algebra;
5759
5860
We also define two functions to convert other bundled maps to `α →o β`:
5961
6062
* `order_embedding.to_order_hom`: convert `α ↪o β` to `α →o β`;
61-
* `rel_hom.to_order_hom`: conver a `rel_hom` between strict orders to a `order_hom`.
63+
* `rel_hom.to_order_hom`: convert a `rel_hom` between strict orders to a `order_hom`.
6264
6365
## Tags
6466
@@ -687,3 +689,23 @@ theorem order_iso.is_complemented_iff :
687689

688690
end bounded_order
689691
end lattice_isos
692+
693+
section boolean_algebra
694+
variables (α) [boolean_algebra α]
695+
696+
/-- Taking complements as an order isomorphism to the order dual. -/
697+
@[simps]
698+
def order_iso.compl : α ≃o order_dual α :=
699+
{ to_fun := order_dual.to_dual ∘ compl,
700+
inv_fun := compl ∘ order_dual.of_dual,
701+
left_inv := compl_compl,
702+
right_inv := compl_compl,
703+
map_rel_iff' := λ x y, compl_le_compl_iff_le }
704+
705+
theorem compl_strict_anti : strict_anti (compl : α → α) :=
706+
(order_iso.compl α).strict_mono
707+
708+
theorem compl_antitone : antitone (compl : α → α) :=
709+
(order_iso.compl α).monotone
710+
711+
end boolean_algebra

0 commit comments

Comments
 (0)