@@ -30,34 +30,61 @@ namespace rel_embedding
30
30
variables {r : α → α → Prop } [is_strict_order α r]
31
31
32
32
/-- If `f` is a strictly `r`-increasing sequence, then this returns `f` as an order embedding. -/
33
- def nat_lt (f : ℕ → α) (H : ∀ n : ℕ, r (f n) (f (n + 1 ))) :
34
- ((<) : ℕ → ℕ → Prop ) ↪r r :=
33
+ def nat_lt (f : ℕ → α) (H : ∀ n : ℕ, r (f n) (f (n + 1 ))) : ((<) : ℕ → ℕ → Prop ) ↪r r :=
35
34
of_monotone f $ nat.rel_of_forall_rel_succ_of_lt r H
36
35
37
- @[simp]
38
- lemma nat_lt_apply {f : ℕ → α} {H : ∀ n : ℕ, r (f n) (f (n + 1 ))} {n : ℕ} : nat_lt f H n = f n :=
39
- rfl
36
+ @[simp] lemma coe_nat_lt {f : ℕ → α} {H : ∀ n : ℕ, r (f n) (f (n + 1 ))} : ⇑(nat_lt f H) = f := rfl
40
37
41
38
/-- If `f` is a strictly `r`-decreasing sequence, then this returns `f` as an order embedding. -/
42
39
def nat_gt (f : ℕ → α) (H : ∀ n : ℕ, r (f (n + 1 )) (f n)) : ((>) : ℕ → ℕ → Prop ) ↪r r :=
43
- by haveI := is_strict_order.swap r; exact rel_embedding.swap (nat_lt f H)
40
+ by { haveI := is_strict_order.swap r, exact rel_embedding.swap (nat_lt f H) }
41
+
42
+ @[simp] lemma coe_nat_gt {f : ℕ → α} {H : ∀ n : ℕ, r (f (n + 1 )) (f n)} : ⇑(nat_gt f H) = f := rfl
43
+
44
+ theorem exists_not_acc_lt_of_not_acc {a : α} {r} (h : ¬ acc r a) : ∃ b, ¬ acc r b ∧ r b a :=
45
+ begin
46
+ contrapose! h,
47
+ refine ⟨_, λ b hr, _⟩,
48
+ by_contra hb,
49
+ exact h b hb hr
50
+ end
51
+
52
+ /-- A value is accessible iff it isn't contained in any infinite decreasing sequence. -/
53
+ theorem acc_iff_no_decreasing_seq {x} :
54
+ acc r x ↔ is_empty {f : ((>) : ℕ → ℕ → Prop ) ↪r r // x ∈ set.range f} :=
55
+ begin
56
+ split,
57
+ { refine λ h, h.rec_on (λ x h IH, _),
58
+ split,
59
+ rintro ⟨f, k, hf⟩,
60
+ exact is_empty.elim' (IH (f (k + 1 )) (hf ▸ f.map_rel_iff.2 (lt_add_one k))) ⟨f, _, rfl⟩ },
61
+ { have : ∀ x : {a // ¬ acc r a}, ∃ y : {a // ¬ acc r a}, r y.1 x.1 ,
62
+ { rintro ⟨x, hx⟩,
63
+ cases exists_not_acc_lt_of_not_acc hx,
64
+ exact ⟨⟨w, h.1 ⟩, h.2 ⟩ },
65
+ obtain ⟨f, h⟩ := classical.axiom_of_choice this ,
66
+ refine λ E, classical.by_contradiction (λ hx, E.elim'
67
+ ⟨(nat_gt (λ n, (f^[n] ⟨x, hx⟩).1 ) (λ n, _)), 0 , rfl⟩),
68
+ rw function.iterate_succ',
69
+ apply h }
70
+ end
71
+
72
+ theorem not_acc_of_decreasing_seq (f : ((>) : ℕ → ℕ → Prop ) ↪r r) (k : ℕ) : ¬ acc r (f k) :=
73
+ by { rw [acc_iff_no_decreasing_seq, not_is_empty_iff], exact ⟨⟨f, k, rfl⟩⟩ }
44
74
75
+ /-- A relation is well-founded iff it doesn't have any infinite decreasing sequence. -/
45
76
theorem well_founded_iff_no_descending_seq :
46
77
well_founded r ↔ is_empty (((>) : ℕ → ℕ → Prop ) ↪r r) :=
47
- ⟨λ ⟨h⟩, ⟨λ ⟨f, o⟩,
48
- suffices ∀ a, acc r a → ∀ n, a ≠ f n, from this (f 0 ) (h _) 0 rfl,
49
- λ a ac, begin
50
- induction ac with a _ IH,
51
- rintro n rfl,
52
- exact IH (f (n+1 )) (o.2 (nat.lt_succ_self _)) _ rfl
53
- end ⟩,
54
- λ E, ⟨λ a, classical.by_contradiction $ λ na,
55
- let ⟨f, h⟩ := classical.axiom_of_choice $
56
- show ∀ x : {a // ¬ acc r a}, ∃ y : {a // ¬ acc r a}, r y.1 x.1 ,
57
- from λ ⟨x, h⟩, classical.by_contradiction $ λ hn, h $
58
- ⟨_, λ y h, classical.by_contradiction $ λ na, hn ⟨⟨y, na⟩, h⟩⟩ in
59
- E.elim' (nat_gt (λ n, (f^[n] ⟨a, na⟩).1 ) $ λ n,
60
- by { rw [function.iterate_succ'], apply h })⟩⟩
78
+ begin
79
+ split,
80
+ { rintro ⟨h⟩,
81
+ exact ⟨λ f, not_acc_of_decreasing_seq f 0 (h _)⟩ },
82
+ { introI h,
83
+ exact ⟨λ x, acc_iff_no_decreasing_seq.2 infer_instance⟩ }
84
+ end
85
+
86
+ theorem not_well_founded_of_decreasing_seq (f : ((>) : ℕ → ℕ → Prop ) ↪r r) : ¬ well_founded r :=
87
+ by { rw [well_founded_iff_no_descending_seq, not_is_empty_iff], exact ⟨f⟩ }
61
88
62
89
end rel_embedding
63
90
0 commit comments