Skip to content

Commit 59e4a79

Browse files
urkudint-y1kim-emChrisHughes24
committed
feat: port Logic.Hydra (#2290)
Co-authored-by: Jason Yuen <[email protected]> Co-authored-by: int-y1 <[email protected]> Co-authored-by: Scott Morrison <[email protected]> Co-authored-by: ChrisHughes24 <[email protected]>
1 parent dfd333b commit 59e4a79

File tree

4 files changed

+166
-5
lines changed

4 files changed

+166
-5
lines changed

Diff for: Mathlib.lean

+1
Original file line numberDiff line numberDiff line change
@@ -1232,6 +1232,7 @@ import Mathlib.Logic.Equiv.Set
12321232
import Mathlib.Logic.Function.Basic
12331233
import Mathlib.Logic.Function.Conjugate
12341234
import Mathlib.Logic.Function.Iterate
1235+
import Mathlib.Logic.Hydra
12351236
import Mathlib.Logic.IsEmpty
12361237
import Mathlib.Logic.Lemmas
12371238
import Mathlib.Logic.Nonempty

Diff for: Mathlib/Data/Multiset/Basic.lean

+3-3
Original file line numberDiff line numberDiff line change
@@ -158,9 +158,9 @@ protected theorem induction {p : Multiset α → Prop} (empty : p 0)
158158
#align multiset.induction Multiset.induction
159159

160160
@[elab_as_elim]
161-
protected theorem induction_on {p : Multiset α → Prop} (s : Multiset α) (h₁ : p 0)
162-
(h₂ : ∀ ⦃a : α⦄ {s : Multiset α}, p s → p (a ::ₘ s)) : p s :=
163-
Multiset.induction h₁ h₂ s
161+
protected theorem induction_on {p : Multiset α → Prop} (s : Multiset α) (empty : p 0)
162+
(cons : ∀ ⦃a : α⦄ {s : Multiset α}, p s → p (a ::ₘ s)) : p s :=
163+
Multiset.induction empty cons s
164164
#align multiset.induction_on Multiset.induction_on
165165

166166
theorem cons_swap (a b : α) (s : Multiset α) : a ::ₘ b ::ₘ s = b ::ₘ a ::ₘ s :=

Diff for: Mathlib/Data/Multiset/Sections.lean

+2-2
Original file line numberDiff line numberDiff line change
@@ -62,8 +62,8 @@ theorem sections_add (s t : Multiset (Multiset α)) :
6262
theorem mem_sections {s : Multiset (Multiset α)} :
6363
∀ {a}, a ∈ Sections s ↔ s.Rel (fun s a => a ∈ s) a := by
6464
induction s using Multiset.induction_on
65-
case h₁ => simp
66-
case h₂ a a' ih => simp [ih, rel_cons_left, eq_comm]
65+
case empty => simp
66+
case cons a a' ih => simp [ih, rel_cons_left, eq_comm]
6767

6868
#align multiset.mem_sections Multiset.mem_sections
6969

Diff for: Mathlib/Logic/Hydra.lean

+160
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,160 @@
1+
/-
2+
Copyright (c) 2022 Junyan Xu. All rights reserved.
3+
Released under Apache 2.0 license as described in the file LICENSE.
4+
Authors: Junyan Xu
5+
6+
! This file was ported from Lean 3 source module logic.hydra
7+
! leanprover-community/mathlib commit 48085f140e684306f9e7da907cd5932056d1aded
8+
! Please do not edit these lines, except to modify the commit id
9+
! if you have ported upstream changes.
10+
-/
11+
import Mathlib.Data.Finsupp.Lex
12+
import Mathlib.Data.Finsupp.Multiset
13+
import Mathlib.Order.GameAdd
14+
15+
/-!
16+
# Termination of a hydra game
17+
18+
This file deals with the following version of the hydra game: each head of the hydra is
19+
labelled by an element in a type `α`, and when you cut off one head with label `a`, it
20+
grows back an arbitrary but finite number of heads, all labelled by elements smaller than
21+
`a` with respect to a well-founded relation `r` on `α`. We show that no matter how (in
22+
what order) you choose cut off the heads, the game always terminates, i.e. all heads will
23+
eventually be cut off (but of course it can last arbitrarily long, i.e. takes an
24+
arbitrary finite number of steps).
25+
26+
This result is stated as the well-foundedness of the `CutExpand` relation defined in
27+
this file: we model the heads of the hydra as a multiset of elements of `α`, and the
28+
valid "moves" of the game are modelled by the relation `CutExpand r` on `Multiset α`:
29+
`CutExpand r s' s` is true iff `s'` is obtained by removing one head `a ∈ s` and
30+
adding back an arbitrary multiset `t` of heads such that all `a' ∈ t` satisfy `r a' a`.
31+
32+
We follow the proof by Peter LeFanu Lumsdaine at https://mathoverflow.net/a/229084/3332.
33+
34+
TODO: formalize the relations corresponding to more powerful (e.g. Kirby–Paris and Buchholz)
35+
hydras, and prove their well-foundedness.
36+
-/
37+
38+
39+
namespace Relation
40+
41+
open Multiset Prod
42+
43+
variable {α : Type _}
44+
45+
/-- The relation that specifies valid moves in our hydra game. `CutExpand r s' s`
46+
means that `s'` is obtained by removing one head `a ∈ s` and adding back an arbitrary
47+
multiset `t` of heads such that all `a' ∈ t` satisfy `r a' a`.
48+
49+
This is most directly translated into `s' = s.erase a + t`, but `Multiset.erase` requires
50+
`DecidableEq α`, so we use the equivalent condition `s' + {a} = s + t` instead, which
51+
is also easier to verify for explicit multisets `s'`, `s` and `t`.
52+
53+
We also don't include the condition `a ∈ s` because `s' + {a} = s + t` already
54+
guarantees `a ∈ s + t`, and if `r` is irreflexive then `a ∉ t`, which is the
55+
case when `r` is well-founded, the case we are primarily interested in.
56+
57+
The lemma `Relation.cutExpand_iff` below converts between this convenient definition
58+
and the direct translation when `r` is irreflexive. -/
59+
def CutExpand (r : α → α → Prop) (s' s : Multiset α) : Prop :=
60+
∃ (t : Multiset α) (a : α), (∀ a' ∈ t, r a' a) ∧ s' + {a} = s + t
61+
#align relation.cut_expand Relation.CutExpand
62+
63+
variable {r : α → α → Prop}
64+
65+
theorem cutExpand_le_invImage_lex [IsIrrefl α r] :
66+
CutExpand r ≤ InvImage (Finsupp.Lex (rᶜ ⊓ (· ≠ ·)) (· < ·)) toFinsupp := by
67+
rintro s t ⟨u, a, hr, he⟩
68+
replace hr := fun a' ↦ mt (hr a')
69+
classical
70+
refine ⟨a, fun b h ↦ ?_, ?_⟩ <;> simp_rw [toFinsupp_apply]
71+
· apply_fun count b at he
72+
simpa only [count_add, count_singleton, if_neg h.2, add_zero, count_eq_zero.2 (hr b h.1)]
73+
using he
74+
· apply_fun count a at he
75+
simp only [count_add, count_singleton_self, count_eq_zero.2 (hr _ (irrefl_of r a)),
76+
add_zero] at he
77+
exact he ▸ Nat.lt_succ_self _
78+
#align relation.cut_expand_le_inv_image_lex Relation.cutExpand_le_invImage_lex
79+
80+
theorem cutExpand_singleton {s x} (h : ∀ x' ∈ s, r x' x) : CutExpand r s {x} :=
81+
⟨s, x, h, add_comm s _⟩
82+
#align relation.cut_expand_singleton Relation.cutExpand_singleton
83+
84+
theorem cutExpand_singleton_singleton {x' x} (h : r x' x) : CutExpand r {x'} {x} :=
85+
cutExpand_singleton fun a h ↦ by rwa [mem_singleton.1 h]
86+
#align relation.cut_expand_singleton_singleton Relation.cutExpand_singleton_singleton
87+
88+
theorem cutExpand_add_left {t u} (s) : CutExpand r (s + t) (s + u) ↔ CutExpand r t u :=
89+
exists₂_congr fun _ _ ↦ and_congr Iff.rfl <| by rw [add_assoc, add_assoc, add_left_cancel_iff]
90+
#align relation.cut_expand_add_left Relation.cutExpand_add_left
91+
92+
theorem cutExpand_iff [DecidableEq α] [IsIrrefl α r] {s' s : Multiset α} :
93+
CutExpand r s' s ↔
94+
∃ (t : Multiset α) (a : α), (∀ a' ∈ t, r a' a) ∧ a ∈ s ∧ s' = s.erase a + t := by
95+
simp_rw [CutExpand, add_singleton_eq_iff]
96+
refine' exists₂_congr fun t a ↦ ⟨_, _⟩
97+
· rintro ⟨ht, ha, rfl⟩
98+
obtain h | h := mem_add.1 ha
99+
exacts [⟨ht, h, erase_add_left_pos t h⟩, (@irrefl α r _ a (ht a h)).elim]
100+
· rintro ⟨ht, h, rfl⟩
101+
exact ⟨ht, mem_add.2 (Or.inl h), (erase_add_left_pos t h).symm⟩
102+
#align relation.cut_expand_iff Relation.cutExpand_iff
103+
104+
theorem not_cutExpand_zero [IsIrrefl α r] (s) : ¬CutExpand r s 0 := by
105+
classical
106+
rw [cutExpand_iff]
107+
rintro ⟨_, _, _, ⟨⟩, _⟩
108+
#align relation.not_cut_expand_zero Relation.not_cutExpand_zero
109+
110+
/-- For any relation `r` on `α`, multiset addition `Multiset α × Multiset α → Multiset α` is a
111+
fibration between the game sum of `CutExpand r` with itself and `CutExpand r` itself. -/
112+
theorem cutExpand_fibration (r : α → α → Prop) :
113+
Fibration (GameAdd (CutExpand r) (CutExpand r)) (CutExpand r) fun s ↦ s.1 + s.2 := by
114+
rintro ⟨s₁, s₂⟩ s ⟨t, a, hr, he⟩; dsimp at he ⊢
115+
classical
116+
-- Porting note: Originally `obtain ⟨ha, rfl⟩`
117+
-- This is https://github.com/leanprover/std4/issues/62
118+
obtain ⟨ha, hb⟩ := add_singleton_eq_iff.1 he
119+
rw [hb]
120+
rw [add_assoc, mem_add] at ha
121+
obtain h | h := ha
122+
· refine' ⟨(s₁.erase a + t, s₂), GameAdd.fst ⟨t, a, hr, _⟩, _⟩
123+
· rw [add_comm, ← add_assoc, singleton_add, cons_erase h]
124+
· rw [add_assoc s₁, erase_add_left_pos _ h, add_right_comm, add_assoc]
125+
· refine' ⟨(s₁, (s₂ + t).erase a), GameAdd.snd ⟨t, a, hr, _⟩, _⟩
126+
· rw [add_comm, singleton_add, cons_erase h]
127+
· rw [add_assoc, erase_add_right_pos _ h]
128+
#align relation.cut_expand_fibration Relation.cutExpand_fibration
129+
130+
/-- A multiset is accessible under `CutExpand` if all its singleton subsets are,
131+
assuming `r` is irreflexive. -/
132+
theorem acc_of_singleton [IsIrrefl α r] {s : Multiset α} (hs : ∀ a ∈ s, Acc (CutExpand r) {a}) :
133+
Acc (CutExpand r) s := by
134+
induction s using Multiset.induction
135+
case empty => exact Acc.intro 0 fun s h ↦ (not_cutExpand_zero s h).elim
136+
case cons a s ihs =>
137+
rw [← s.singleton_add a]
138+
rw [forall_mem_cons] at hs
139+
exact (hs.1.prod_gameAdd <| ihs fun a ha ↦ hs.2 a ha).of_fibration _ (cutExpand_fibration r)
140+
#align relation.acc_of_singleton Relation.acc_of_singleton
141+
142+
/-- A singleton `{a}` is accessible under `CutExpand r` if `a` is accessible under `r`,
143+
assuming `r` is irreflexive. -/
144+
theorem _root_.Acc.cutExpand [IsIrrefl α r] {a : α} (hacc : Acc r a) : Acc (CutExpand r) {a} := by
145+
induction' hacc with a h ih
146+
refine' Acc.intro _ fun s ↦ _
147+
classical
148+
simp only [cutExpand_iff, mem_singleton]
149+
rintro ⟨t, a, hr, rfl, rfl⟩
150+
refine' acc_of_singleton fun a' ↦ _
151+
rw [erase_singleton, zero_add]
152+
exact ih a' ∘ hr a'
153+
#align acc.cut_expand Acc.cutExpand
154+
155+
/-- `CutExpand r` is well-founded when `r` is. -/
156+
theorem _root_.WellFounded.cutExpand (hr : WellFounded r) : WellFounded (CutExpand r) :=
157+
have := hr.isIrrefl; fun _ ↦ acc_of_singleton fun a _ ↦ (hr.apply a).cutExpand⟩
158+
#align well_founded.cut_expand WellFounded.cutExpand
159+
160+
end Relation

0 commit comments

Comments
 (0)