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

Commit 5ca6e9d

Browse files
committed
feat(logic/relation): characterization of the accessible predicate in terms of well_founded_on (#16720)
+ Show that `acc r a` (read: `a` is accessible under the relation `r`) if and only if `r` is well-founded on the subset of elements below `a` (transitively) under `r`, including (`refl_trans_gen`) or excluding (`trans_gen`) `a`. Stated as a `list.tfae`. + Move `relation.fibration` (used by the proof) to *logic/relation* to avoid importing *logic/hydra*. Add lemma `acc_trans_gen_iff` used in the proof.
1 parent 486cb2f commit 5ca6e9d

File tree

3 files changed

+59
-37
lines changed

3 files changed

+59
-37
lines changed

src/logic/hydra.lean

Lines changed: 3 additions & 34 deletions
Original file line numberDiff line numberDiff line change
@@ -5,7 +5,6 @@ Authors: Junyan Xu
55
-/
66
import data.multiset.basic
77
import order.game_add
8-
import order.well_founded
98

109
/-!
1110
# Termination of a hydra game
@@ -24,47 +23,18 @@ valid "moves" of the game are modelled by the relation `cut_expand r` on `multis
2423
`cut_expand r s' s` is true iff `s'` is obtained by removing one head `a ∈ s` and
2524
adding back an arbitrary multiset `t` of heads such that all `a' ∈ t` satisfy `r a' a`.
2625
27-
To prove this theorem, we follow the proof by Peter LeFanu Lumsdaine at
28-
https://mathoverflow.net/a/229084/3332, and along the way we introduce the notion of `fibration`
29-
of relations.
26+
We follow the proof by Peter LeFanu Lumsdaine at https://mathoverflow.net/a/229084/3332.
3027
3128
TODO: formalize the relations corresponding to more powerful (e.g. Kirby–Paris and Buchholz)
3229
hydras, and prove their well-foundedness.
3330
-/
3431

3532
namespace relation
3633

37-
variables {α β : Type*}
38-
39-
section fibration
40-
variables (rα : α → α → Prop) (rβ : β → β → Prop) (f : α → β)
41-
42-
/-- A function `f : α → β` is a fibration between the relation `rα` and `rβ` if for all
43-
`a : α` and `b : β`, whenever `b : β` and `f a` are related by `rβ`, `b` is the image
44-
of some `a' : α` under `f`, and `a'` and `a` are related by `rα`. -/
45-
def fibration := ∀ ⦃a b⦄, rβ b (f a) → ∃ a', rα a' a ∧ f a' = b
46-
47-
variables {rα rβ}
48-
49-
/-- If `f : α → β` is a fibration between relations `rα` and `rβ`, and `a : α` is
50-
accessible under `rα`, then `f a` is accessible under `rβ`. -/
51-
lemma _root_.acc.of_fibration (fib : fibration rα rβ f) {a} (ha : acc rα a) : acc rβ (f a) :=
52-
begin
53-
induction ha with a ha ih,
54-
refine acc.intro (f a) (λ b hr, _),
55-
obtain ⟨a', hr', rfl⟩ := fib hr,
56-
exact ih a' hr',
57-
end
58-
59-
lemma _root_.acc.of_downward_closed (dc : ∀ {a b}, rβ b (f a) → b ∈ set.range f)
60-
(a : α) (ha : acc (inv_image rβ f) a) : acc rβ (f a) :=
61-
ha.of_fibration f (λ a b h, let ⟨a', he⟩ := dc h in ⟨a', he.substr h, he⟩)
62-
63-
end fibration
64-
65-
section hydra
6634
open multiset prod
6735

36+
variables {α β : Type*}
37+
6838
/-- The relation that specifies valid moves in our hydra game. `cut_expand r s' s`
6939
means that `s'` is obtained by removing one head `a ∈ s` and adding back an arbitrary
7040
multiset `t` of heads such that all `a' ∈ t` satisfy `r a' a`.
@@ -152,6 +122,5 @@ end
152122
/-- `cut_expand r` is well-founded when `r` is. -/
153123
theorem _root_.well_founded.cut_expand (hr : well_founded r) : well_founded (cut_expand r) :=
154124
by { letI h := hr.is_irrefl, exact λ s, acc_of_singleton $ λ a _, (hr.apply a).cut_expand }⟩
155-
end hydra
156125

157126
end relation

src/logic/relation.lean

Lines changed: 33 additions & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -147,6 +147,33 @@ end
147147

148148
end comp
149149

150+
section fibration
151+
152+
variables (rα : α → α → Prop) (rβ : β → β → Prop) (f : α → β)
153+
154+
/-- A function `f : α → β` is a fibration between the relation `rα` and `rβ` if for all
155+
`a : α` and `b : β`, whenever `b : β` and `f a` are related by `rβ`, `b` is the image
156+
of some `a' : α` under `f`, and `a'` and `a` are related by `rα`. -/
157+
def fibration := ∀ ⦃a b⦄, rβ b (f a) → ∃ a', rα a' a ∧ f a' = b
158+
159+
variables {rα rβ}
160+
161+
/-- If `f : α → β` is a fibration between relations `rα` and `rβ`, and `a : α` is
162+
accessible under `rα`, then `f a` is accessible under `rβ`. -/
163+
lemma _root_.acc.of_fibration (fib : fibration rα rβ f) {a} (ha : acc rα a) : acc rβ (f a) :=
164+
begin
165+
induction ha with a ha ih,
166+
refine acc.intro (f a) (λ b hr, _),
167+
obtain ⟨a', hr', rfl⟩ := fib hr,
168+
exact ih a' hr',
169+
end
170+
171+
lemma _root_.acc.of_downward_closed (dc : ∀ {a b}, rβ b (f a) → ∃ c, f c = b)
172+
(a : α) (ha : acc (inv_image rβ f) a) : acc rβ (f a) :=
173+
ha.of_fibration f (λ a b h, let ⟨a', he⟩ := dc h in ⟨a', he.substr h, he⟩)
174+
175+
end fibration
176+
150177
/--
151178
The map of a relation `r` through a pair of functions pushes the
152179
relation to the codomains of the functions. The resulting relation is
@@ -373,16 +400,19 @@ end
373400

374401
end trans_gen
375402

376-
lemma _root_.acc.trans_gen {α} {r : α → α → Prop} {a : α} (h : acc r a) : acc (trans_gen r) a :=
403+
lemma _root_.acc.trans_gen (h : acc r a) : acc (trans_gen r) a :=
377404
begin
378405
induction h with x _ H,
379406
refine acc.intro x (λ y hy, _),
380407
cases hy with _ hyx z _ hyz hzx,
381408
exacts [H y hyx, (H z hzx).inv hyz],
382409
end
383410

384-
lemma _root_.well_founded.trans_gen {α} {r : α → α → Prop} (h : well_founded r) :
385-
well_founded (trans_gen r) := ⟨λ a, (h.apply a).trans_gen⟩
411+
lemma _root_.acc_trans_gen_iff : acc (trans_gen r) a ↔ acc r a :=
412+
⟨subrelation.accessible (λ _ _, trans_gen.single), acc.trans_gen⟩
413+
414+
lemma _root_.well_founded.trans_gen (h : well_founded r) : well_founded (trans_gen r) :=
415+
⟨λ a, (h.apply a).trans_gen⟩
386416

387417
section trans_gen
388418

src/order/well_founded_set.lean

Lines changed: 23 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -6,6 +6,7 @@ Authors: Aaron Anderson
66
import order.antichain
77
import order.order_iso_nat
88
import order.well_founded
9+
import tactic.tfae
910

1011
/-!
1112
# Well-founded sets
@@ -96,6 +97,28 @@ end
9697

9798
lemma subset (h : t.well_founded_on r) (hst : s ⊆ t) : s.well_founded_on r := h.mono le_rfl hst
9899

100+
open relation
101+
102+
/-- `a` is accessible under the relation `r` iff `r` is well-founded on the downward transitive
103+
closure of `a` under `r` (including `a` or not). -/
104+
lemma acc_iff_well_founded_on {α} {r : α → α → Prop} {a : α} :
105+
[ acc r a,
106+
{b | refl_trans_gen r b a}.well_founded_on r,
107+
{b | trans_gen r b a}.well_founded_on r ].tfae :=
108+
begin
109+
tfae_have : 12,
110+
{ refine λ h, ⟨λ b, _⟩, apply inv_image.accessible,
111+
rw ← acc_trans_gen_iff at h ⊢,
112+
obtain h'|h' := refl_trans_gen_iff_eq_or_trans_gen.1 b.2,
113+
{ rwa h' at h }, { exact h.inv h' } },
114+
tfae_have : 23,
115+
{ exact λ h, h.subset (λ _, trans_gen.to_refl) },
116+
tfae_have : 31,
117+
{ refine λ h, acc.intro _ (λ b hb, (h.apply ⟨b, trans_gen.single hb⟩).of_fibration subtype.val _),
118+
exact λ ⟨c, hc⟩ d h, ⟨⟨d, trans_gen.head h hc⟩, h, rfl⟩ },
119+
tfae_finish,
120+
end
121+
99122
end well_founded_on
100123
end any_rel
101124

0 commit comments

Comments
 (0)