@@ -1048,15 +1048,15 @@ noncomputable def gcdMonoidOfGCD [DecidableEq α] (gcd : α → α → α)
1048
1048
lcm := fun a b =>
1049
1049
if a = 0 then 0 else Classical.choose ((gcd_dvd_left a b).trans (Dvd.intro b rfl))
1050
1050
gcd_mul_lcm := fun a b => by
1051
- -- Porting note (#10971 ): need `dsimp only` before `split_ifs`
1052
- dsimp only
1051
+ -- Porting note(#12129 ): additional beta reduction needed
1052
+ beta_reduce
1053
1053
split_ifs with a0
1054
1054
· rw [mul_zero, a0, zero_mul]
1055
1055
· rw [← Classical.choose_spec ((gcd_dvd_left a b).trans (Dvd.intro b rfl))]
1056
1056
lcm_zero_left := fun a => if_pos rfl
1057
1057
lcm_zero_right := fun a => by
1058
- -- Porting note (#10971 ): need `dsimp only` before `split_ifs`
1059
- dsimp only
1058
+ -- Porting note(#12129 ): additional beta reduction needed
1059
+ beta_reduce
1060
1060
split_ifs with a0
1061
1061
· rfl
1062
1062
have h := (Classical.choose_spec ((gcd_dvd_left a 0 ).trans (Dvd.intro 0 rfl))).symm
@@ -1108,17 +1108,17 @@ noncomputable def normalizedGCDMonoidOfGCD [NormalizationMonoid α] [DecidableEq
1108
1108
rw [← normalize_gcd] at this
1109
1109
rwa [normalize.map_mul, normalize_gcd, mul_right_inj' h1] at h2
1110
1110
gcd_mul_lcm := fun a b => by
1111
- -- Porting note (#10971 ): need `dsimp only`
1112
- dsimp only
1111
+ -- Porting note(#12129 ): additional beta reduction needed
1112
+ beta_reduce
1113
1113
split_ifs with a0
1114
1114
· rw [mul_zero, a0, zero_mul]
1115
1115
· rw [←
1116
1116
Classical.choose_spec (dvd_normalize_iff.2 ((gcd_dvd_left a b).trans (Dvd.intro b rfl)))]
1117
1117
exact normalize_associated (a * b)
1118
1118
lcm_zero_left := fun a => if_pos rfl
1119
1119
lcm_zero_right := fun a => by
1120
- -- Porting note (#10971 ): need `dsimp only`
1121
- dsimp only
1120
+ -- Porting note(#12129 ): additional beta reduction needed
1121
+ beta_reduce
1122
1122
split_ifs with a0
1123
1123
· rfl
1124
1124
rw [← normalize_eq_zero] at a0
@@ -1141,17 +1141,17 @@ noncomputable def gcdMonoidOfLCM [DecidableEq α] (lcm : α → α → α)
1141
1141
{ lcm
1142
1142
gcd := fun a b => if a = 0 then b else if b = 0 then a else Classical.choose (exists_gcd a b)
1143
1143
gcd_mul_lcm := fun a b => by
1144
- -- Porting note (#10971 ): need `dsimp only`
1145
- dsimp only
1144
+ -- Porting note(#12129 ): additional beta reduction needed
1145
+ beta_reduce
1146
1146
split_ifs with h h_1
1147
1147
· rw [h, eq_zero_of_zero_dvd (dvd_lcm_left _ _), mul_zero, zero_mul]
1148
1148
· rw [h_1, eq_zero_of_zero_dvd (dvd_lcm_right _ _)]
1149
1149
rw [mul_comm, ← Classical.choose_spec (exists_gcd a b)]
1150
1150
lcm_zero_left := fun a => eq_zero_of_zero_dvd (dvd_lcm_left _ _)
1151
1151
lcm_zero_right := fun a => eq_zero_of_zero_dvd (dvd_lcm_right _ _)
1152
1152
gcd_dvd_left := fun a b => by
1153
- -- Porting note (#10971 ): need `dsimp only`
1154
- dsimp only
1153
+ -- Porting note(#12129 ): additional beta reduction needed
1154
+ beta_reduce
1155
1155
split_ifs with h h_1
1156
1156
· rw [h]
1157
1157
apply dvd_zero
@@ -1167,8 +1167,8 @@ noncomputable def gcdMonoidOfLCM [DecidableEq α] (lcm : α → α → α)
1167
1167
mul_dvd_mul_iff_right h]
1168
1168
apply dvd_lcm_right
1169
1169
gcd_dvd_right := fun a b => by
1170
- -- Porting note (#10971 ): need `dsimp only`
1171
- dsimp only
1170
+ -- Porting note(#12129 ): additional beta reduction needed
1171
+ beta_reduce
1172
1172
split_ifs with h h_1
1173
1173
· exact dvd_rfl
1174
1174
· rw [h_1]
@@ -1184,8 +1184,8 @@ noncomputable def gcdMonoidOfLCM [DecidableEq α] (lcm : α → α → α)
1184
1184
mul_dvd_mul_iff_right h_1]
1185
1185
apply dvd_lcm_left
1186
1186
dvd_gcd := fun {a b c} ac ab => by
1187
- -- Porting note (#10971 ): need `dsimp only`
1188
- dsimp only
1187
+ -- Porting note(#12129 ): additional beta reduction needed
1188
+ beta_reduce
1189
1189
split_ifs with h h_1
1190
1190
· exact ab
1191
1191
· exact ac
@@ -1219,7 +1219,7 @@ noncomputable def normalizedGCDMonoidOfLCM [NormalizationMonoid α] [DecidableEq
1219
1219
if a = 0 then normalize b
1220
1220
else if b = 0 then normalize a else Classical.choose (exists_gcd a b)
1221
1221
gcd_mul_lcm := fun a b => by
1222
- dsimp only
1222
+ beta_reduce
1223
1223
split_ifs with h h_1
1224
1224
· rw [h, eq_zero_of_zero_dvd (dvd_lcm_left _ _), mul_zero, zero_mul]
1225
1225
· rw [h_1, eq_zero_of_zero_dvd (dvd_lcm_right _ _), mul_zero, mul_zero]
@@ -1247,7 +1247,7 @@ noncomputable def normalizedGCDMonoidOfLCM [NormalizationMonoid α] [DecidableEq
1247
1247
lcm_zero_left := fun a => eq_zero_of_zero_dvd (dvd_lcm_left _ _)
1248
1248
lcm_zero_right := fun a => eq_zero_of_zero_dvd (dvd_lcm_right _ _)
1249
1249
gcd_dvd_left := fun a b => by
1250
- dsimp only
1250
+ beta_reduce
1251
1251
split_ifs with h h_1
1252
1252
· rw [h]
1253
1253
apply dvd_zero
@@ -1263,7 +1263,7 @@ noncomputable def normalizedGCDMonoidOfLCM [NormalizationMonoid α] [DecidableEq
1263
1263
mul_comm, mul_dvd_mul_iff_right h]
1264
1264
apply dvd_lcm_right
1265
1265
gcd_dvd_right := fun a b => by
1266
- dsimp only
1266
+ beta_reduce
1267
1267
split_ifs with h h_1
1268
1268
· exact (normalize_associated _).dvd
1269
1269
· rw [h_1]
@@ -1279,7 +1279,7 @@ noncomputable def normalizedGCDMonoidOfLCM [NormalizationMonoid α] [DecidableEq
1279
1279
mul_dvd_mul_iff_right h_1]
1280
1280
apply dvd_lcm_left
1281
1281
dvd_gcd := fun {a b c} ac ab => by
1282
- dsimp only
1282
+ beta_reduce
1283
1283
split_ifs with h h_1
1284
1284
· apply dvd_normalize_iff.2 ab
1285
1285
· apply dvd_normalize_iff.2 ac
@@ -1357,36 +1357,37 @@ instance (priority := 100) : NormalizedGCDMonoid G₀ where
1357
1357
normUnit x := if h : x = 0 then 1 else (Units.mk0 x h)⁻¹
1358
1358
normUnit_zero := dif_pos rfl
1359
1359
normUnit_mul := fun {x y} x0 y0 => Units.eq_iff.1 (by
1360
- -- Porting note (#10971): need `dsimp only`, also `simp` reaches maximum heartbeat
1360
+ -- Porting note(#12129): additional beta reduction needed
1361
+ -- Porting note: `simp` reaches maximum heartbeat
1361
1362
-- by Units.eq_iff.mp (by simp only [x0, y0, mul_comm])
1362
- dsimp only
1363
+ beta_reduce
1363
1364
split_ifs with h
1364
1365
· rw [mul_eq_zero] at h
1365
1366
cases h
1366
1367
· exact absurd ‹x = 0 › x0
1367
1368
· exact absurd ‹y = 0 › y0
1368
1369
· rw [Units.mk0_mul, mul_inv_rev, mul_comm] )
1369
1370
normUnit_coe_units u := by
1370
- -- Porting note (#10971 ): need `dsimp only`
1371
- dsimp only
1371
+ -- Porting note(#12129 ): additional beta reduction needed
1372
+ beta_reduce
1372
1373
rw [dif_neg (Units.ne_zero _), Units.mk0_val]
1373
1374
gcd a b := if a = 0 ∧ b = 0 then 0 else 1
1374
1375
lcm a b := if a = 0 ∨ b = 0 then 0 else 1
1375
1376
gcd_dvd_left a b := by
1376
- -- Porting note (#10971 ): need `dsimp only`
1377
- dsimp only
1377
+ -- Porting note(#12129 ): additional beta reduction needed
1378
+ beta_reduce
1378
1379
split_ifs with h
1379
1380
· rw [h.1 ]
1380
1381
· exact one_dvd _
1381
1382
gcd_dvd_right a b := by
1382
- -- Porting note (#10971 ): need `dsimp only`
1383
- dsimp only
1383
+ -- Porting note(#12129 ): additional beta reduction needed
1384
+ beta_reduce
1384
1385
split_ifs with h
1385
1386
· rw [h.2 ]
1386
1387
· exact one_dvd _
1387
1388
dvd_gcd := fun {a b c} hac hab => by
1388
- -- Porting note (#10971 ): need `dsimp only`
1389
- dsimp only
1389
+ -- Porting note(#12129 ): additional beta reduction needed
1390
+ beta_reduce
1390
1391
split_ifs with h
1391
1392
· apply dvd_zero
1392
1393
· rw [not_and_or] at h
@@ -1402,8 +1403,8 @@ instance (priority := 100) : NormalizedGCDMonoid G₀ where
1402
1403
· by_cases hb : b = 0
1403
1404
· simp only [hb, and_true, or_true, ite_true, mul_zero]
1404
1405
exact Associated.refl _
1405
- -- Porting note (#10971 ): need `dsimp only`
1406
- · dsimp only
1406
+ -- Porting note(#12129 ): additional beta reduction needed
1407
+ · beta_reduce
1407
1408
rw [if_neg (not_and_of_not_left _ ha), one_mul, if_neg (not_or_of_not ha hb)]
1408
1409
exact (associated_one_iff_isUnit.mpr ((IsUnit.mk0 _ ha).mul (IsUnit.mk0 _ hb))).symm
1409
1410
lcm_zero_left b := if_pos (Or.inl rfl)
0 commit comments