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

Commit 8af7e08

Browse files
feat(data/fintype/basic): make subtype_of_fintype computable (#5919)
This smokes out a few places downstream that are missing decidability hypotheses needed for the fintype instance to exist. Co-authored-by: Eric Wieser <wieser.eric@gmail.com>
1 parent f45dee4 commit 8af7e08

File tree

8 files changed

+36
-13
lines changed

8 files changed

+36
-13
lines changed

src/combinatorics/simple_graph/basic.lean

Lines changed: 10 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -101,6 +101,9 @@ variables {V : Type u} (G : simple_graph V)
101101
/-- `G.neighbor_set v` is the set of vertices adjacent to `v` in `G`. -/
102102
def neighbor_set (v : V) : set V := set_of (G.adj v)
103103

104+
instance neighbor_set.mem_decidable (v : V) [decidable_rel G.adj] :
105+
decidable_pred (∈ G.neighbor_set v) := by { unfold neighbor_set, apply_instance }
106+
104107
lemma ne_of_adj {a b : V} (hab : G.adj a b) : a ≠ b :=
105108
by { rintro rfl, exact G.loopless a hab }
106109

@@ -360,6 +363,11 @@ The maximum degree of all vertices
360363
def max_degree (G : simple_graph V) [nonempty V] [decidable_rel G.adj] : ℕ :=
361364
finset.max' (univ.image (λ (v : V), G.degree v)) (nonempty.image univ_nonempty _)
362365

366+
/-! The following lemmas about `fintype.card` use noncomputable decidable instances to get fintype
367+
assumptions. -/
368+
section
369+
open_locale classical
370+
363371
lemma degree_lt_card_verts (G : simple_graph V) (v : V) : G.degree v < fintype.card V :=
364372
begin
365373
classical,
@@ -408,6 +416,8 @@ begin
408416
apply common_neighbors_subset_neighbor_set } },
409417
end
410418

419+
end
420+
411421
end finite
412422

413423
section complement

src/combinatorics/simple_graph/degree_sum.lean

Lines changed: 3 additions & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -177,7 +177,7 @@ G.dart_card_eq_sum_degrees.symm.trans G.dart_card_eq_twice_card_edges
177177
end degree_sum
178178

179179
/-- The handshaking lemma. See also `simple_graph.sum_degrees_eq_twice_card_edges`. -/
180-
theorem even_card_odd_degree_vertices [fintype V] :
180+
theorem even_card_odd_degree_vertices [fintype V] [decidable_rel G.adj] :
181181
even (univ.filter (λ v, odd (G.degree v))).card :=
182182
begin
183183
classical,
@@ -197,7 +197,7 @@ begin
197197
trivial }
198198
end
199199

200-
lemma odd_card_odd_degree_vertices_ne [fintype V] [decidable_eq V]
200+
lemma odd_card_odd_degree_vertices_ne [fintype V] [decidable_eq V] [decidable_rel G.adj]
201201
(v : V) (h : odd (G.degree v)) :
202202
odd (univ.filter (λ w, w ≠ v ∧ odd (G.degree w))).card :=
203203
begin
@@ -220,7 +220,7 @@ begin
220220
{ simpa only [true_and, mem_filter, mem_univ] },
221221
end
222222

223-
lemma exists_ne_odd_degree_of_exists_odd_degree [fintype V]
223+
lemma exists_ne_odd_degree_of_exists_odd_degree [fintype V] [decidable_rel G.adj]
224224
(v : V) (h : odd (G.degree v)) :
225225
∃ (w : V), w ≠ v ∧ odd (G.degree w) :=
226226
begin

src/data/fintype/basic.lean

Lines changed: 6 additions & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -243,6 +243,12 @@ def of_surjective [decidable_eq β] [fintype α] (f : α → β) (H : function.s
243243
fintype β :=
244244
⟨univ.image f, λ b, let ⟨a, e⟩ := H b in e ▸ mem_image_of_mem _ (mem_univ _)⟩
245245

246+
instance subtype_of_fintype {α : Sort*} [fintype α] (p : α → Prop) [decidable_pred p] :
247+
fintype (subtype p) :=
248+
{ elems := ⟨((finset.univ : finset α).filter p).1.pmap subtype.mk (by simp),
249+
multiset.nodup_pmap (λ _ _ _ _, subtype.mk.inj) (multiset.nodup_filter p finset.univ.nodup)⟩,
250+
complete := λ ⟨x, h⟩, by simp [h] }
251+
246252
/-- Given an injective function to a fintype, the domain is also a
247253
fintype. This is noncomputable because injectivity alone cannot be
248254
used to construct preimages. -/
@@ -252,9 +258,6 @@ if hα : nonempty α then by letI := classical.inhabited_of_nonempty hα;
252258
exact of_surjective (inv_fun f) (inv_fun_surjective H)
253259
else ⟨∅, λ x, (hα ⟨x⟩).elim⟩
254260

255-
noncomputable instance subtype_of_fintype [fintype α] (p : α → Prop) : fintype (subtype p) :=
256-
fintype.of_injective coe subtype.coe_injective
257-
258261
/-- If `f : α ≃ β` and `α` is a fintype, then `β` is also a fintype. -/
259262
def of_equiv (α : Type*) [fintype α] (f : α ≃ β) : fintype β := of_bijective _ f.bijective
260263

src/data/fintype/card.lean

Lines changed: 2 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -285,7 +285,8 @@ end
285285

286286
@[to_additive]
287287
lemma finset.prod_to_finset_eq_subtype {M : Type*} [comm_monoid M] [fintype α]
288-
(p : α → Prop) (f : α → M) : ∏ a in {x | p x}.to_finset, f a = ∏ a : subtype p, f a :=
288+
(p : α → Prop) [decidable_pred p] (f : α → M) :
289+
∏ a in {x | p x}.to_finset, f a = ∏ a : subtype p, f a :=
289290
by { rw ← finset.prod_subtype, simp }
290291

291292
@[to_additive] lemma finset.prod_fiberwise [decidable_eq β] [fintype β] [comm_monoid γ]

src/field_theory/galois.lean

Lines changed: 3 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -321,8 +321,9 @@ begin
321321
have p : 0 < findim (intermediate_field.fixed_field (⊤ : subgroup (E ≃ₐ[F] E))) E := findim_pos,
322322
rw [←intermediate_field.findim_eq_one_iff, ←mul_left_inj' (ne_of_lt p).symm, findim_mul_findim,
323323
←h, one_mul, intermediate_field.findim_fixed_field_eq_card],
324-
exact fintype.card_congr { to_fun := λ g, ⟨g, subgroup.mem_top g⟩, inv_fun := coe,
325-
left_inv := λ g, rfl, right_inv := λ _, by { ext, refl } },
324+
apply fintype.card_congr,
325+
exact { to_fun := λ g, ⟨g, subgroup.mem_top g⟩, inv_fun := coe,
326+
left_inv := λ g, rfl, right_inv := λ _, by { ext, refl } },
326327
end
327328

328329
variables {F} {E} {p : polynomial F}

src/group_theory/order_of_element.lean

Lines changed: 1 addition & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -301,6 +301,7 @@ lemma is_cyclic_of_prime_card [group α] [fintype α] {p : ℕ} [hp : fact p.pri
301301
begin
302302
obtain ⟨g, hg⟩ : ∃ g : α, g ≠ 1,
303303
from fintype.exists_ne_of_one_lt_card (by { rw h, exact nat.prime.one_lt hp }) 1,
304+
classical, -- for fintype (subgroup.gpowers g)
304305
have : fintype.card (subgroup.gpowers g) ∣ p,
305306
{ rw ←h,
306307
apply card_subgroup_dvd_card },

src/group_theory/perm/subgroup.lean

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -21,13 +21,13 @@ universes u
2121

2222
@[simp]
2323
lemma sum_congr_hom.card_range {α β : Type*}
24-
[decidable_eq α] [decidable_eq β] [fintype α] [fintype β] :
24+
[fintype (sum_congr_hom α β).range] [fintype (perm α × perm β)] :
2525
fintype.card (sum_congr_hom α β).range = fintype.card (perm α × perm β) :=
2626
fintype.card_eq.mpr ⟨(set.range (sum_congr_hom α β) sum_congr_hom_injective).symm⟩
2727

2828
@[simp]
2929
lemma sigma_congr_right_hom.card_range {α : Type*} {β : α → Type*}
30-
[decidable_eq α] [∀ a, decidable_eq (β a)] [fintype α] [∀ a, fintype (β a)] :
30+
[fintype (sigma_congr_right_hom β).range] [fintype a, perm (β a))] :
3131
fintype.card (sigma_congr_right_hom β).range = fintype.card (Π a, perm (β a)) :=
3232
fintype.card_eq.mpr ⟨(set.range (sigma_congr_right_hom β) sigma_congr_right_hom_injective).symm⟩
3333

src/group_theory/subgroup.lean

Lines changed: 9 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -160,6 +160,10 @@ lemma coe_coe (K : subgroup G) : ↥(K : set G) = K := rfl
160160
attribute [norm_cast] add_subgroup.mem_coe
161161
attribute [norm_cast] add_subgroup.coe_coe
162162

163+
@[to_additive]
164+
instance (K : subgroup G) [d : decidable_pred K.carrier] [fintype G] : fintype K :=
165+
show fintype {g : G // g ∈ K.carrier}, from infer_instance
166+
163167
end subgroup
164168

165169
@[to_additive]
@@ -362,10 +366,13 @@ begin
362366
exact ⟨h x, by { rintros rfl, exact H.one_mem }⟩ },
363367
end
364368

365-
@[to_additive] lemma eq_top_of_card_eq [fintype G] (h : fintype.card H = fintype.card G) : H = ⊤ :=
369+
@[to_additive] lemma eq_top_of_card_eq [fintype H] [fintype G]
370+
(h : fintype.card H = fintype.card G) : H = ⊤ :=
366371
begin
372+
classical,
373+
rw fintype.card_congr (equiv.refl _) at h, -- this swaps the fintype instance to classical
367374
change fintype.card H.carrier = _ at h,
368-
cases H with S hS1 hS2 hS3,
375+
unfreezingI { cases H with S hS1 hS2 hS3, },
369376
have : S = set.univ,
370377
{ suffices : S.to_finset = finset.univ,
371378
{ rwa [←set.to_finset_univ, set.to_finset_inj] at this, },

0 commit comments

Comments
 (0)