mirror of
https://github.com/GTBarkley/comm_alg.git
synced 2024-12-26 07:38:36 -06:00
golf attempt
This commit is contained in:
parent
4b9f84f7c1
commit
c702c535b5
1 changed files with 8 additions and 18 deletions
|
@ -98,33 +98,23 @@ lemma map_lt_adjoin_x (I : PrimeSpectrum R) : map_prime I < adjoin_x I := by
|
||||||
simp
|
simp
|
||||||
rw [←Ideal.eq_top_iff_one]
|
rw [←Ideal.eq_top_iff_one]
|
||||||
exact I.IsPrime.ne_top'
|
exact I.IsPrime.ne_top'
|
||||||
|
|
||||||
/- Given an ideal p in R, define the ideal p[x] in R[x] -/
|
|
||||||
lemma ht_adjoin_x_eq_ht_add_one [Nontrivial R] (I : PrimeSpectrum R) : height I + 1 ≤ height (adjoin_x I) := by
|
lemma ht_adjoin_x_eq_ht_add_one [Nontrivial R] (I : PrimeSpectrum R) : height I + 1 ≤ height (adjoin_x I) := by
|
||||||
suffices H : height I + (1 : ℕ) ≤ height (adjoin_x I) + (0 : ℕ)
|
suffices H : height I + (1 : ℕ) ≤ height (adjoin_x I) + (0 : ℕ)
|
||||||
. norm_cast at H; rw [add_zero] at H; exact H
|
. norm_cast at H; rw [add_zero] at H; exact H
|
||||||
rw [height, height, Set.chainHeight_add_le_chainHeight_add {J | J < I} _ 1 0]
|
rw [height, height, Set.chainHeight_add_le_chainHeight_add {J | J < I} _ 1 0]
|
||||||
intro l hl
|
intro l hl
|
||||||
use ((l.map map_prime) ++ [map_prime I])
|
use ((l.map map_prime) ++ [map_prime I])
|
||||||
constructor
|
refine' ⟨_, by simp⟩
|
||||||
. simp [Set.append_mem_subchain_iff]
|
. simp [Set.append_mem_subchain_iff]
|
||||||
refine' ⟨_,_,_⟩
|
refine' ⟨_, map_lt_adjoin_x I, fun a ha => _⟩
|
||||||
. show (List.map map_prime l).Chain' (· < ·) ∧ ∀ i ∈ _, i ∈ _
|
. refine' ⟨_, fun i hi => _⟩
|
||||||
constructor
|
. apply List.chain'_map_of_chain' map_prime (fun a b hab => map_strictmono hab) hl.1
|
||||||
. apply List.chain'_map_of_chain' map_prime
|
. rw [List.mem_map] at hi
|
||||||
intro a b hab
|
|
||||||
apply map_strictmono
|
|
||||||
exact hab
|
|
||||||
exact hl.1
|
|
||||||
. intro i hi
|
|
||||||
rw [List.mem_map] at hi
|
|
||||||
obtain ⟨a, ha, rfl⟩ := hi
|
obtain ⟨a, ha, rfl⟩ := hi
|
||||||
show map_prime a < adjoin_x I
|
|
||||||
calc map_prime a < map_prime I := by apply map_strictmono; apply hl.2; apply ha
|
calc map_prime a < map_prime I := by apply map_strictmono; apply hl.2; apply ha
|
||||||
_ < adjoin_x I := by apply map_lt_adjoin_x
|
_ < adjoin_x I := by apply map_lt_adjoin_x
|
||||||
. apply map_lt_adjoin_x
|
. have H : ∃ b : PrimeSpectrum R, b ∈ l ∧ map_prime b = a
|
||||||
. intro a ha
|
|
||||||
have H : ∃ b : PrimeSpectrum R, b ∈ l ∧ map_prime b = a
|
|
||||||
. have H2 : l ≠ []
|
. have H2 : l ≠ []
|
||||||
. intro h
|
. intro h
|
||||||
rw [h] at ha
|
rw [h] at ha
|
||||||
|
@ -141,7 +131,7 @@ lemma ht_adjoin_x_eq_ht_add_one [Nontrivial R] (I : PrimeSpectrum R) : height I
|
||||||
apply map_strictmono
|
apply map_strictmono
|
||||||
apply hl.2
|
apply hl.2
|
||||||
exact hb
|
exact hb
|
||||||
. simp
|
|
||||||
|
|
||||||
/-
|
/-
|
||||||
dim R + 1 ≤ dim R[X]
|
dim R + 1 ≤ dim R[X]
|
||||||
|
|
Loading…
Reference in a new issue