@@ -64,28 +64,8 @@ end monotone
64
64
65
65
namespace function
66
66
67
- namespace commute
68
-
69
67
section preorder
70
-
71
- variables [preorder α] {f g : α → α}
72
-
73
- lemma iterate_le_of_map_le (h : commute f g) (hf : monotone f) (hg : monotone g)
74
- {x} (hx : f x ≤ g x) (n : ℕ) :
75
- f^[n] x ≤ (g^[n]) x :=
76
- by refine hf.seq_le_seq n _ (λ k hk, _) (λ k hk, _);
77
- simp [iterate_succ' f, h.iterate_right _ _, hg.iterate _ hx]
78
-
79
- lemma iterate_pos_lt_of_map_lt (h : commute f g) (hf : monotone f) (hg : strict_mono g)
80
- {x} (hx : f x < g x) {n} (hn : 0 < n) :
81
- f^[n] x < (g^[n]) x :=
82
- by refine hf.seq_pos_lt_seq_of_le_of_lt hn _ (λ k hk, _) (λ k hk, _);
83
- simp [iterate_succ' f, h.iterate_right _ _, hg.iterate _ hx]
84
-
85
- lemma iterate_pos_lt_of_map_lt' (h : commute f g) (hf : strict_mono f) (hg : monotone g)
86
- {x} (hx : f x < g x) {n} (hn : 0 < n) :
87
- f^[n] x < (g^[n]) x :=
88
- @iterate_pos_lt_of_map_lt (order_dual α) _ g f h.symm hg.order_dual hf.order_dual x hx n hn
68
+ variables [preorder α] {f : α → α}
89
69
90
70
lemma id_le_iterate_of_id_le (h : id ≤ f) :
91
71
∀ n, id ≤ (f^[n])
@@ -113,6 +93,30 @@ lemma iterate_le_iterate_of_le_id (h : f ≤ id) {m n : ℕ} (hmn : m ≤ n) :
113
93
114
94
end preorder
115
95
96
+ namespace commute
97
+
98
+ section preorder
99
+ variables [preorder α] {f g : α → α}
100
+
101
+ lemma iterate_le_of_map_le (h : commute f g) (hf : monotone f) (hg : monotone g)
102
+ {x} (hx : f x ≤ g x) (n : ℕ) :
103
+ f^[n] x ≤ (g^[n]) x :=
104
+ by refine hf.seq_le_seq n _ (λ k hk, _) (λ k hk, _);
105
+ simp [iterate_succ' f, h.iterate_right _ _, hg.iterate _ hx]
106
+
107
+ lemma iterate_pos_lt_of_map_lt (h : commute f g) (hf : monotone f) (hg : strict_mono g)
108
+ {x} (hx : f x < g x) {n} (hn : 0 < n) :
109
+ f^[n] x < (g^[n]) x :=
110
+ by refine hf.seq_pos_lt_seq_of_le_of_lt hn _ (λ k hk, _) (λ k hk, _);
111
+ simp [iterate_succ' f, h.iterate_right _ _, hg.iterate _ hx]
112
+
113
+ lemma iterate_pos_lt_of_map_lt' (h : commute f g) (hf : strict_mono f) (hg : monotone g)
114
+ {x} (hx : f x < g x) {n} (hn : 0 < n) :
115
+ f^[n] x < (g^[n]) x :=
116
+ @iterate_pos_lt_of_map_lt (order_dual α) _ g f h.symm hg.order_dual hf.order_dual x hx n hn
117
+
118
+ end preorder
119
+
116
120
variables [linear_order α] {f g : α → α}
117
121
118
122
lemma iterate_pos_lt_iff_map_lt (h : commute f g) (hf : monotone f)
0 commit comments