chapter \ International System of Units \ section \ SI Units Semantics \ theory SI_Units imports ISQ_Proof begin text \ An SI unit is simply a particular kind of quantity. \ type_synonym ('n, 'd) UnitT = "('n, 'd) QuantT" text \ Parallel to the seven base quantities, there are seven base units. In the implementation of the SI unit system, we fix these to be precisely those quantities that have a base dimension and a magnitude of \<^term>\1\. Consequently, a base unit corresponds to a unit in the algebraic sense. \ lift_definition is_base_unit :: "'a::one['d::dim_type] \ bool" is "\ x. mag x = 1 \ is_BaseDim (dim x)" . definition mk_base_unit :: "'u itself \ ('a::one)['u::basedim_type]" where [si_eq]: "mk_base_unit t = 1" syntax "_mk_base_unit" :: "type \ logic" ("BUNIT'(_')") translations "BUNIT('a)" == "CONST mk_base_unit TYPE('a)" lemma mk_base_unit: "is_base_unit (mk_base_unit a)" by (simp add: si_eq, transfer, simp add: is_BaseDim) lemma magQuant_mk [si_eq]: "\BUNIT('u::basedim_type)\\<^sub>Q = 1" by (simp add: magQuant_def si_eq, transfer, simp) text \ We now define the seven base units. Effectively, these definitions axiomatise given names for the \<^term>\1\ elements of the base quantities. \ definition [si_eq]: "meter = BUNIT(L)" definition [si_eq]: "second = BUNIT(T)" definition [si_eq]: "kilogram = BUNIT(M)" definition [si_eq]: "ampere = BUNIT(I)" definition [si_eq]: "kelvin = BUNIT(\)" definition [si_eq]: "mole = BUNIT(N)" definition [si_eq]: "candela = BUNIT(J)" text\Note that as a consequence of our construction, the term \<^term>\meter\ is a SI Unit constant of SI-type \<^typ>\'a[L]\, so a unit of dimension \<^typ>\Length\ with the magnitude of type \<^typ>\'a\. A magnitude instantiation can be, e.g., an integer, a rational number, a real number, or a vector of type \<^typ>\real\<^sup>3\. Note than when considering vectors, dimensions refer to the \<^emph>\norm\ of the vector, not to its components. \ lemma BaseUnits: "is_base_unit meter" "is_base_unit second" "is_base_unit kilogram" "is_base_unit ampere" "is_base_unit kelvin" "is_base_unit mole" "is_base_unit candela" by (simp add: si_eq, transfer, simp)+ text \ The effect of the above encoding is that we can use the SI base units as synonyms for their corresponding dimensions at the type level. \ type_synonym meter = Length type_synonym second = Time type_synonym kilogram = Mass type_synonym ampere = Current type_synonym kelvin = Temperature type_synonym mole = Amount type_synonym candela = Intensity text \ We can therefore construct a quantity such as \<^term>\5 :: rat[meter]\, which unambiguously identifies that the unit of $5$ is meters using the type system. This works because each base unit it the one element. \ subsection \ Example Unit Equations \ lemma "(meter \<^bold>\ second\<^sup>-\<^sup>\) \<^bold>\ second \\<^sub>Q meter" by (si_calc) end