| author | Fabian Huch <huch@in.tum.de> | 
| Tue, 13 Feb 2024 11:57:41 +0100 | |
| changeset 79593 | 587a7dfeb03c | 
| parent 71398 | e0237f2eb49d | 
| permissions | -rw-r--r-- | 
| 65435 | 1 | (* Title: HOL/Computational_Algebra/Field_as_Ring.thy | 
| 65417 | 2 | Author: Manuel Eberl | 
| 3 | *) | |
| 4 | ||
| 5 | theory Field_as_Ring | |
| 6 | imports | |
| 7 | Complex_Main | |
| 8 | Euclidean_Algorithm | |
| 9 | begin | |
| 10 | ||
| 11 | context field | |
| 12 | begin | |
| 13 | ||
| 14 | subclass idom_divide .. | |
| 15 | ||
| 16 | definition normalize_field :: "'a \<Rightarrow> 'a" | |
| 17 | where [simp]: "normalize_field x = (if x = 0 then 0 else 1)" | |
| 18 | definition unit_factor_field :: "'a \<Rightarrow> 'a" | |
| 19 | where [simp]: "unit_factor_field x = x" | |
| 20 | definition euclidean_size_field :: "'a \<Rightarrow> nat" | |
| 21 | where [simp]: "euclidean_size_field x = (if x = 0 then 0 else 1)" | |
| 22 | definition mod_field :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" | |
| 23 | where [simp]: "mod_field x y = (if y = 0 then x else 0)" | |
| 24 | ||
| 25 | end | |
| 26 | ||
| 71398 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 Manuel Eberl <eberlm@in.tum.de> parents: 
70817diff
changeset | 27 | instantiation real :: | 
| 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 Manuel Eberl <eberlm@in.tum.de> parents: 
70817diff
changeset | 28 |   "{unique_euclidean_ring, normalization_euclidean_semiring, normalization_semidom_multiplicative}"
 | 
| 65417 | 29 | begin | 
| 30 | ||
| 31 | definition [simp]: "normalize_real = (normalize_field :: real \<Rightarrow> _)" | |
| 32 | definition [simp]: "unit_factor_real = (unit_factor_field :: real \<Rightarrow> _)" | |
| 66838 
17989f6bc7b2
clarified uniqueness criterion for euclidean rings
 haftmann parents: 
66817diff
changeset | 33 | definition [simp]: "modulo_real = (mod_field :: real \<Rightarrow> _)" | 
| 65417 | 34 | definition [simp]: "euclidean_size_real = (euclidean_size_field :: real \<Rightarrow> _)" | 
| 66838 
17989f6bc7b2
clarified uniqueness criterion for euclidean rings
 haftmann parents: 
66817diff
changeset | 35 | definition [simp]: "division_segment (x :: real) = 1" | 
| 65417 | 36 | |
| 37 | instance | |
| 38 | by standard | |
| 70817 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
66838diff
changeset | 39 | (simp_all add: dvd_field_iff field_split_simps split: if_splits) | 
| 65417 | 40 | |
| 41 | end | |
| 42 | ||
| 43 | instantiation real :: euclidean_ring_gcd | |
| 44 | begin | |
| 45 | ||
| 46 | definition gcd_real :: "real \<Rightarrow> real \<Rightarrow> real" where | |
| 47 | "gcd_real = Euclidean_Algorithm.gcd" | |
| 48 | definition lcm_real :: "real \<Rightarrow> real \<Rightarrow> real" where | |
| 49 | "lcm_real = Euclidean_Algorithm.lcm" | |
| 50 | definition Gcd_real :: "real set \<Rightarrow> real" where | |
| 51 | "Gcd_real = Euclidean_Algorithm.Gcd" | |
| 52 | definition Lcm_real :: "real set \<Rightarrow> real" where | |
| 53 | "Lcm_real = Euclidean_Algorithm.Lcm" | |
| 54 | ||
| 55 | instance by standard (simp_all add: gcd_real_def lcm_real_def Gcd_real_def Lcm_real_def) | |
| 56 | ||
| 57 | end | |
| 58 | ||
| 71398 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 Manuel Eberl <eberlm@in.tum.de> parents: 
70817diff
changeset | 59 | instance real :: field_gcd .. | 
| 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 Manuel Eberl <eberlm@in.tum.de> parents: 
70817diff
changeset | 60 | |
| 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 Manuel Eberl <eberlm@in.tum.de> parents: 
70817diff
changeset | 61 | |
| 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 Manuel Eberl <eberlm@in.tum.de> parents: 
70817diff
changeset | 62 | instantiation rat :: | 
| 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 Manuel Eberl <eberlm@in.tum.de> parents: 
70817diff
changeset | 63 |   "{unique_euclidean_ring, normalization_euclidean_semiring, normalization_semidom_multiplicative}"
 | 
| 65417 | 64 | begin | 
| 65 | ||
| 66 | definition [simp]: "normalize_rat = (normalize_field :: rat \<Rightarrow> _)" | |
| 67 | definition [simp]: "unit_factor_rat = (unit_factor_field :: rat \<Rightarrow> _)" | |
| 66838 
17989f6bc7b2
clarified uniqueness criterion for euclidean rings
 haftmann parents: 
66817diff
changeset | 68 | definition [simp]: "modulo_rat = (mod_field :: rat \<Rightarrow> _)" | 
| 65417 | 69 | definition [simp]: "euclidean_size_rat = (euclidean_size_field :: rat \<Rightarrow> _)" | 
| 66838 
17989f6bc7b2
clarified uniqueness criterion for euclidean rings
 haftmann parents: 
66817diff
changeset | 70 | definition [simp]: "division_segment (x :: rat) = 1" | 
| 65417 | 71 | |
| 72 | instance | |
| 73 | by standard | |
| 70817 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
66838diff
changeset | 74 | (simp_all add: dvd_field_iff field_split_simps split: if_splits) | 
| 65417 | 75 | |
| 76 | end | |
| 77 | ||
| 78 | instantiation rat :: euclidean_ring_gcd | |
| 79 | begin | |
| 80 | ||
| 81 | definition gcd_rat :: "rat \<Rightarrow> rat \<Rightarrow> rat" where | |
| 82 | "gcd_rat = Euclidean_Algorithm.gcd" | |
| 83 | definition lcm_rat :: "rat \<Rightarrow> rat \<Rightarrow> rat" where | |
| 84 | "lcm_rat = Euclidean_Algorithm.lcm" | |
| 85 | definition Gcd_rat :: "rat set \<Rightarrow> rat" where | |
| 86 | "Gcd_rat = Euclidean_Algorithm.Gcd" | |
| 87 | definition Lcm_rat :: "rat set \<Rightarrow> rat" where | |
| 88 | "Lcm_rat = Euclidean_Algorithm.Lcm" | |
| 89 | ||
| 90 | instance by standard (simp_all add: gcd_rat_def lcm_rat_def Gcd_rat_def Lcm_rat_def) | |
| 91 | ||
| 92 | end | |
| 93 | ||
| 71398 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 Manuel Eberl <eberlm@in.tum.de> parents: 
70817diff
changeset | 94 | instance rat :: field_gcd .. | 
| 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 Manuel Eberl <eberlm@in.tum.de> parents: 
70817diff
changeset | 95 | |
| 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 Manuel Eberl <eberlm@in.tum.de> parents: 
70817diff
changeset | 96 | |
| 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 Manuel Eberl <eberlm@in.tum.de> parents: 
70817diff
changeset | 97 | instantiation complex :: | 
| 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 Manuel Eberl <eberlm@in.tum.de> parents: 
70817diff
changeset | 98 |   "{unique_euclidean_ring, normalization_euclidean_semiring, normalization_semidom_multiplicative}"
 | 
| 65417 | 99 | begin | 
| 100 | ||
| 101 | definition [simp]: "normalize_complex = (normalize_field :: complex \<Rightarrow> _)" | |
| 102 | definition [simp]: "unit_factor_complex = (unit_factor_field :: complex \<Rightarrow> _)" | |
| 66838 
17989f6bc7b2
clarified uniqueness criterion for euclidean rings
 haftmann parents: 
66817diff
changeset | 103 | definition [simp]: "modulo_complex = (mod_field :: complex \<Rightarrow> _)" | 
| 65417 | 104 | definition [simp]: "euclidean_size_complex = (euclidean_size_field :: complex \<Rightarrow> _)" | 
| 66838 
17989f6bc7b2
clarified uniqueness criterion for euclidean rings
 haftmann parents: 
66817diff
changeset | 105 | definition [simp]: "division_segment (x :: complex) = 1" | 
| 65417 | 106 | |
| 107 | instance | |
| 108 | by standard | |
| 70817 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
66838diff
changeset | 109 | (simp_all add: dvd_field_iff field_split_simps split: if_splits) | 
| 65417 | 110 | |
| 111 | end | |
| 112 | ||
| 113 | instantiation complex :: euclidean_ring_gcd | |
| 114 | begin | |
| 115 | ||
| 116 | definition gcd_complex :: "complex \<Rightarrow> complex \<Rightarrow> complex" where | |
| 117 | "gcd_complex = Euclidean_Algorithm.gcd" | |
| 118 | definition lcm_complex :: "complex \<Rightarrow> complex \<Rightarrow> complex" where | |
| 119 | "lcm_complex = Euclidean_Algorithm.lcm" | |
| 120 | definition Gcd_complex :: "complex set \<Rightarrow> complex" where | |
| 121 | "Gcd_complex = Euclidean_Algorithm.Gcd" | |
| 122 | definition Lcm_complex :: "complex set \<Rightarrow> complex" where | |
| 123 | "Lcm_complex = Euclidean_Algorithm.Lcm" | |
| 124 | ||
| 125 | instance by standard (simp_all add: gcd_complex_def lcm_complex_def Gcd_complex_def Lcm_complex_def) | |
| 126 | ||
| 127 | end | |
| 128 | ||
| 71398 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 Manuel Eberl <eberlm@in.tum.de> parents: 
70817diff
changeset | 129 | instance complex :: field_gcd .. | 
| 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 Manuel Eberl <eberlm@in.tum.de> parents: 
70817diff
changeset | 130 | |
| 65417 | 131 | end |