Added core physical constants of the 2019 SI standard, dimensionless units, and various proof facilities for support.

This commit is contained in:
Simon Foster 2020-02-18 17:46:13 +00:00
parent 5c303a7192
commit 2c63ef07e9
6 changed files with 163 additions and 45 deletions

View File

@ -1,13 +1,67 @@
section \<open> Physical Constants \<close>
theory SI_Constants
imports SI_Quantities
imports SI_Proof
begin
abbreviation speed_of_light :: "'a::division_ring[L \<cdot> T\<^sup>-\<^sup>1]" where
abbreviation "hertz \<equiv> second\<^sup>-\<^sup>\<one>"
abbreviation "radian \<equiv> meter \<^bold>\<cdot> meter\<^sup>-\<^sup>\<one>"
abbreviation "steradian \<equiv> meter\<^sup>\<two> \<^bold>\<cdot> meter\<^sup>-\<^sup>\<two>"
abbreviation "joule \<equiv> kilogram \<^bold>\<cdot> meter\<^sup>\<two> \<^bold>\<cdot> second\<^sup>-\<^sup>\<two>"
abbreviation "watt \<equiv> kilogram \<^bold>\<cdot> meter\<^sup>\<two> \<^bold>\<cdot> second\<^sup>-\<^sup>\<three>"
abbreviation "coulomb \<equiv> ampere \<^bold>\<cdot> second"
abbreviation "lumen \<equiv> candela \<^bold>\<cdot> steradian"
text \<open> The most general types we support must form a field into which the natural numbers can
be injected. \<close>
default_sort field_char_0
abbreviation (input) caesium_frequency:: "'a[T\<^sup>-\<^sup>1]" ("\<Delta>v\<^sub>C\<^sub>s") where
"caesium_frequency \<equiv> 9192631770\<cdot>hertz"
abbreviation speed_of_light :: "'a[L \<cdot> T\<^sup>-\<^sup>1]" where
"speed_of_light \<equiv> 299792458\<cdot>(meter\<^bold>\<cdot>second\<^sup>-\<^sup>\<one>)"
abbreviation gravitational_constant where
abbreviation Planck :: "'a[M \<cdot> L\<^sup>2 \<cdot> T\<^sup>-\<^sup>2 \<cdot> T]" where
"Planck \<equiv> (6.62607015 \<cdot> 1/(10^34))\<cdot>(joule\<^bold>\<cdot>second)"
abbreviation elementary_charge :: "'a[I \<cdot> T]" where
"elementary_charge \<equiv> (1.602176634 \<cdot> 1/(10^19))\<cdot>coulomb"
abbreviation Boltzmann :: "'a[M \<cdot> L\<^sup>2 \<cdot> T\<^sup>-\<^sup>2 \<cdot> \<Theta>\<^sup>-\<^sup>1]" where
"Boltzmann \<equiv> (1.380649\<cdot>1/(10^23))\<cdot>(joule \<^bold>/ kelvin)"
abbreviation Avogadro :: "'a[N\<^sup>-\<^sup>1]" where
"Avogadro \<equiv> 6.02214076\<cdot>(10^23)\<cdot>(mole\<^sup>-\<^sup>\<one>)"
abbreviation max_luminous_frequency :: "'a[T\<^sup>-\<^sup>1]" where
"max_luminous_frequency \<equiv> (540\<cdot>10^12)\<cdot>hertz"
abbreviation luminous_efficacy :: "'a[J \<cdot> (L\<^sup>2 \<cdot> L\<^sup>-\<^sup>2) \<cdot> (M \<cdot> L\<^sup>2 \<cdot> T\<^sup>-\<^sup>3)\<^sup>-\<^sup>1]" where
"luminous_efficacy \<equiv> 683\<cdot>(lumen\<^bold>/watt)"
abbreviation gravitational_constant :: "'a[L\<^sup>3 \<cdot> M\<^sup>-\<^sup>1 \<cdot> T\<^sup>-\<^sup>2]" where
"gravitational_constant \<equiv> (6.6743015 \<cdot> 1/(10 ^ 11)) \<cdot> (meter\<^sup>\<three>\<^bold>\<cdot>kilogram\<^sup>-\<^sup>\<one>\<^bold>\<cdot>second\<^sup>-\<^sup>\<two>)"
thm si_def
theorem Quant_eq_iff_same_dim:
"x \<approx>\<^sub>Q y \<longleftrightarrow> x = y"
by (transfer, simp)
theorem hertz_definition: "1\<cdot>hertz = \<Delta>v\<^sub>C\<^sub>s / 9192631770"
by (simp add: unit_eq_iff_magn_eq si_def)
theorem second_definition: "1\<cdot>second \<approx>\<^sub>Q (9192631770::_[\<one>]) \<^bold>/ \<Delta>v\<^sub>C\<^sub>s"
by (simp add: unit_equiv_iff, simp add: Quant_equiv_def unit_eq_iff_magn_eq si_def)
default_sort type
end

View File

@ -2,14 +2,18 @@ section \<open> Derived Units\<close>
theory SI_Derived
imports SI_Prefix
begin
begin
definition "radian = 1 \<cdot> (meter \<^bold>\<cdot> meter\<^sup>-\<^sup>\<one>)"
definition degree :: "'a::{inverse,real_algebra_1}[L/L]" where
[si_def]: "degree = (2\<cdot>(UNIT(of_real pi,_)) / 180)\<cdot>radian"
definition degree :: "real[L/L]" where
[si_def]: "degree = (2\<cdot>(UNIT(pi,_)) / 180)\<cdot>radian"
abbreviation degrees ("_\<degree>" [999] 999) where "n\<degree> \<equiv> n\<cdot>degree"
abbreviation degrees ("_\<degree>" [999] 999) where "n\<degree> \<equiv> n\<cdot>degree"
definition degrees_celcius :: "'a::division_ring[\<Theta>] \<Rightarrow> 'a[\<Theta>]" ("_\<degree>C" [999] 999)
where [si_def]: "degrees_celcius x = (x + 273.151)\<cdot>kelvin"
definition degrees_farenheit :: "'a::division_ring[\<Theta>] \<Rightarrow> 'a[\<Theta>]" ("_\<degree>F" [999] 999)
where [si_def]: "degrees_farenheit x = (x + 459.67)\<cdot>5/9\<cdot>kelvin"
definition [si_def]: "litre = 1/1000 \<cdot> meter\<^sup>\<three>"
@ -17,15 +21,20 @@ definition [si_def]: "pint = 0.56826125 \<cdot> litre"
definition [si_def]: "hour = 3600 \<cdot> second"
definition [si_def]: "gram = milli\<cdot>kilogram"
abbreviation "tonne \<equiv> kilo\<cdot>kilogram"
abbreviation "newton \<equiv> (kilogram \<^bold>\<cdot> meter) \<^bold>/ second\<^sup>\<two>"
abbreviation "volt \<equiv> kilogram \<^bold>\<cdot> meter\<^sup>\<two> \<^bold>\<cdot> second\<^sup>-\<^sup>\<three> \<^bold>\<cdot> ampere\<^sup>-\<^sup>\<one>"
abbreviation "watt \<equiv> kilogram \<^bold>\<cdot> meter\<^sup>\<two> \<^bold>\<cdot> second\<^sup>-\<^sup>\<three>"
definition "inch = 25.5\<cdot>milli\<cdot>meter"
definition "foot = 0.3048\<cdot>meter"
definition "yard = 0.9144\<cdot>meter"
abbreviation "joule \<equiv> kilogram \<^bold>\<cdot> meter\<^sup>\<two> \<^bold>\<cdot> second\<^sup>-\<^sup>\<two>"
text\<open>The full beauty of the approach is perhaps revealed here, with the
type of a classical three-dimensional gravitation field:\<close>

View File

@ -1,7 +1,7 @@
section \<open> SI Prefixes \<close>
theory SI_Prefix
imports SI_Quantities
imports SI_Constants
begin
definition [si_def]: "giga = UNIT(1000000000, _)"

View File

@ -1,12 +1,12 @@
theory SI_Proof
imports SI_Derived
imports SI_Quantities
begin
section \<open> Tactic Support for SI type expressions. \<close>
lemmas [si_def] = si_sem_Length_def si_sem_Mass_def si_sem_Time_def
si_sem_Current_def si_sem_Temperature_def si_sem_Amount_def
si_sem_Intensity_def
si_sem_Intensity_def si_sem_NoDimension_def
si_sem_UnitTimes_def si_sem_UnitInv_def
times_Unit_ext_def one_Unit_ext_def
@ -19,48 +19,49 @@ lemma "SI(I) = 1\<lparr>Amperes := 1\<rparr>" by(simp add: si_def)
lemma "SI(\<Theta>) = 1\<lparr>Kelvins := 1\<rparr> " by(simp add: si_def)
lemma "SI(N) = 1\<lparr>Moles := 1\<rparr>" by(simp add: si_def)
lemma "SI(J) = 1\<lparr>Candelas := 1\<rparr>" by(simp add: si_def)
lemma "SI(\<one>) = 1" by(simp add: si_def)
lemma "SI(N \<cdot> \<Theta> \<cdot> N) = SI(\<Theta> \<cdot> N\<^sup>2)" by(simp add: si_def)
lemma [si_def]:"fromUnit UNIT(x::'a::one, Time) =
\<lparr>magn = x,
unit = \<lparr>Seconds = 1, Meters = 0, Kilograms = 0, Amperes = 0,
unit = \<lparr>Meters = 0, Kilograms = 0, Seconds = 1, Amperes = 0,
Kelvins = 0, Moles = 0, Candelas = 0\<rparr>\<rparr>"
by (simp add: mk_unit.rep_eq one_Unit_ext_def si_sem_Time_def)
lemma [si_def]:"fromUnit UNIT(x::'a::one, Length) =
\<lparr>magn = x,
unit = \<lparr>Seconds = 0, Meters = 1, Kilograms = 0, Amperes = 0,
unit = \<lparr>Meters = 1, Kilograms = 0, Seconds = 0, Amperes = 0,
Kelvins = 0, Moles = 0, Candelas = 0\<rparr>\<rparr>"
by (simp add: mk_unit.rep_eq one_Unit_ext_def si_sem_Length_def)
lemma [si_def]:"fromUnit UNIT(x::'a::one, Mass) =
\<lparr>magn = x,
unit = \<lparr>Seconds = 0, Meters = 0, Kilograms = 1, Amperes = 0,
unit = \<lparr>Meters = 0, Kilograms = 1, Seconds = 0, Amperes = 0,
Kelvins = 0, Moles = 0, Candelas = 0\<rparr>\<rparr>"
by (simp add: mk_unit.rep_eq one_Unit_ext_def si_sem_Mass_def)
lemma [si_def]:"fromUnit UNIT(x::'a::one, Current) =
\<lparr>magn = x,
unit = \<lparr>Seconds = 0, Meters = 0, Kilograms = 0, Amperes = 1,
unit = \<lparr>Meters = 0, Kilograms = 0, Seconds = 0, Amperes = 1,
Kelvins = 0, Moles = 0, Candelas = 0\<rparr>\<rparr>"
by (simp add: mk_unit.rep_eq one_Unit_ext_def si_sem_Current_def)
lemma [si_def]:"fromUnit UNIT(x::'a::one, Temperature) =
\<lparr>magn = x,
unit = \<lparr>Seconds = 0, Meters = 0, Kilograms = 0, Amperes = 0,
unit = \<lparr>Meters = 0, Kilograms = 0, Seconds = 0, Amperes = 0,
Kelvins = 1, Moles = 0, Candelas = 0\<rparr>\<rparr>"
by (simp add: mk_unit.rep_eq one_Unit_ext_def si_sem_Temperature_def)
lemma [si_def]:"fromUnit UNIT(x::'a::one, Amount) =
\<lparr>magn = x,
unit = \<lparr>Seconds = 0, Meters = 0, Kilograms = 0, Amperes = 0,
unit = \<lparr>Meters = 0, Kilograms = 0, Seconds = 0, Amperes = 0,
Kelvins = 0, Moles = 1, Candelas = 0\<rparr>\<rparr>"
by (simp add: mk_unit.rep_eq one_Unit_ext_def si_sem_Amount_def)
lemma [si_def]:"fromUnit UNIT(x::'a::one, Intensity) =
\<lparr>magn = x,
unit = \<lparr>Seconds = 0, Meters = 0, Kilograms = 0, Amperes = 0,
unit = \<lparr>Meters = 0, Kilograms = 0, Seconds = 0, Amperes = 0,
Kelvins = 0, Moles = 0, Candelas = 1\<rparr>\<rparr>"
by (simp add: mk_unit.rep_eq one_Unit_ext_def si_sem_Intensity_def)
@ -109,6 +110,7 @@ lemma Unit_inverse_weak_cong:
shows "(X)\<^sup>-\<^sup>\<one> \<approx>\<^sub>Q (Y)\<^sup>-\<^sup>\<one>"
using assms by (transfer, simp)
(*
text\<open>In order to compute a normal form, we would additionally need these three:\<close>
(* field ? *)
lemma Unit_inverse_distrib:
@ -154,6 +156,6 @@ lemma "watt \<^bold>\<cdot> hour \<approx>\<^sub>Q 3600 \<^bold>\<cdot> joule"
lemma "watt \<^bold>\<cdot> hour \<approx>\<^sub>Q 3.6 \<^bold>\<cdot> kilo \<^bold>\<cdot> joule"
oops
*)
end

View File

@ -18,7 +18,8 @@ lemma Quantity_eq_intro:
instantiation Quantity_ext :: (times, times) times
begin
definition "times_Quantity_ext x y = \<lparr> magn = magn x \<cdot> magn y, unit = unit x \<cdot> unit y, \<dots> = more x \<cdot> more y \<rparr>"
definition [si_def]:
"times_Quantity_ext x y = \<lparr> magn = magn x \<cdot> magn y, unit = unit x \<cdot> unit y, \<dots> = more x \<cdot> more y \<rparr>"
instance ..
end
@ -38,7 +39,7 @@ lemma more_zero [simp]: "more 0 = 0" by (simp add: zero_Quantity_ext_def)
instantiation Quantity_ext :: (one, one) one
begin
definition "one_Quantity_ext = \<lparr> magn = 1, unit = 1, \<dots> = 1 \<rparr>"
definition [si_def]: "one_Quantity_ext = \<lparr> magn = 1, unit = 1, \<dots> = 1 \<rparr>"
instance ..
end
@ -48,8 +49,8 @@ lemma more_one [simp]: "more 1 = 1" by (simp add: one_Quantity_ext_def)
instantiation Quantity_ext :: (inverse, inverse) inverse
begin
definition "inverse_Quantity_ext x = \<lparr> magn = inverse (magn x), unit = inverse (unit x), \<dots> = inverse (more x) \<rparr>"
definition "divide_Quantity_ext x y = \<lparr> magn = magn x / magn y, unit = unit x / unit y, \<dots> = more x / more y \<rparr>"
definition [si_def]: "inverse_Quantity_ext x = \<lparr> magn = inverse (magn x), unit = inverse (unit x), \<dots> = inverse (more x) \<rparr>"
definition [si_def]: "divide_Quantity_ext x y = \<lparr> magn = magn x / magn y, unit = unit x / unit y, \<dots> = more x / more y \<rparr>"
instance ..
end
@ -300,8 +301,6 @@ translations "UNIT(n, 'a)" == "CONST mk_unit n TYPE('a)"
subsection \<open>Polymorphic Operations for Elementary SI Units \<close>
named_theorems si_def
definition [si_def]: "meter = UNIT(1, Length)"
definition [si_def]: "second = UNIT(1, Time)"
definition [si_def]: "kilogram = UNIT(1, Mass)"
@ -321,6 +320,16 @@ lemma unit_eq_iff_magn_eq:
"x = y \<longleftrightarrow> \<lbrakk>x\<rbrakk>\<^sub>Q = \<lbrakk>y\<rbrakk>\<^sub>Q"
by (auto simp add: magnQuant_def, transfer, simp)
lemma unit_equiv_iff:
fixes x :: "'a['u\<^sub>1::si_type]" and y :: "'a['u\<^sub>2::si_type]"
shows "x \<approx>\<^sub>Q y \<longleftrightarrow> \<lbrakk>x\<rbrakk>\<^sub>Q = \<lbrakk>y\<rbrakk>\<^sub>Q \<and> SI('u\<^sub>1) = SI('u\<^sub>2)"
proof -
have "\<forall>t ta. (ta::'a['u\<^sub>2]) = t \<or> magn (fromUnit ta) \<noteq> magn (fromUnit t)"
by (simp add: magnQuant_def unit_eq_iff_magn_eq)
then show ?thesis
by (metis (full_types) Quant_equiv.rep_eq coerceQuant_eq_iff2 eq_ magnQuant_def)
qed
lemma unit_le_iff_magn_le:
"x \<le> y \<longleftrightarrow> \<lbrakk>x\<rbrakk>\<^sub>Q \<le> \<lbrakk>y\<rbrakk>\<^sub>Q"
by (auto simp add: magnQuant_def; (transfer, simp))
@ -340,6 +349,12 @@ lemma magnQuant_times [si_def]: "\<lbrakk>x * y\<rbrakk>\<^sub>Q = \<lbrakk>x\<r
lemma magnQuant_div [si_def]: "\<lbrakk>x / y\<rbrakk>\<^sub>Q = \<lbrakk>x\<rbrakk>\<^sub>Q / \<lbrakk>y\<rbrakk>\<^sub>Q"
by (simp add: magnQuant_def, transfer, simp)
lemma magnQuant_qinv [si_def]: "\<lbrakk>x\<^sup>-\<^sup>\<one>\<rbrakk>\<^sub>Q = inverse \<lbrakk>x\<rbrakk>\<^sub>Q"
by (simp add: magnQuant_def, transfer, simp)
lemma magnQuant_qdiv [si_def]: "\<lbrakk>(x::('a::field)[_]) \<^bold>/ y\<rbrakk>\<^sub>Q = \<lbrakk>x\<rbrakk>\<^sub>Q / \<lbrakk>y\<rbrakk>\<^sub>Q"
by (simp add: magnQuant_def, transfer, simp add: field_class.field_divide_inverse)
lemma magnQuant_numeral [si_def]: "\<lbrakk>numeral n\<rbrakk>\<^sub>Q = numeral n"
apply (induct n, simp_all add: si_def)
apply (metis magnQuant_plus numeral_code(2))

View File

@ -6,6 +6,8 @@ theory SI_Units
"HOL-Eisbach.Eisbach"
begin
named_theorems si_def
text\<open>
The International System of Units (SI, abbreviated from the French
Système international (d'unités)) is the modern form of the metric
@ -46,21 +48,28 @@ text \<open> An SI unit associates with each of the seven base unit an integer t
to which it is raised. We use a record to represent this 7-tuple, to enable code generation. \<close>
record Unit =
Seconds :: int
Meters :: int
Kilograms :: int
Meters :: int
Kilograms :: int
Seconds :: int
Amperes :: int
Kelvins :: int
Moles :: int
Candelas :: int
text \<open> Units scaled by a power of 10 \<close>
type_synonym SUnit = "int Unit_ext"
abbreviation scale :: "SUnit \<Rightarrow> int" where
"scale \<equiv> more"
text \<open> We define a commutative monoid for SI units. \<close>
instantiation Unit_ext :: (one) one
begin
\<comment> \<open> Here, $1$ is the dimensionless unit \<close>
definition one_Unit_ext :: "'a Unit_ext"
where [code_unfold]: "1 = \<lparr> Seconds = 0, Meters = 0, Kilograms = 0, Amperes = 0
where [code_unfold, si_def]: "1 = \<lparr> Meters = 0, Kilograms = 0, Seconds = 0, Amperes = 0
, Kelvins = 0, Moles = 0, Candelas = 0, \<dots> = 1 \<rparr>"
instance ..
end
@ -69,9 +78,9 @@ instantiation Unit_ext :: (times) times
begin
\<comment> \<open> Multiplication is defined by adding together the powers \<close>
definition times_Unit_ext :: "'a Unit_ext \<Rightarrow> 'a Unit_ext \<Rightarrow> 'a Unit_ext"
where [code_unfold]:
"x * y = \<lparr> Seconds = Seconds x + Seconds y, Meters = Meters x + Meters y
, Kilograms = Kilograms x + Kilograms y, Amperes = Amperes x + Amperes y
where [code_unfold, si_def]:
"x * y = \<lparr> Meters = Meters x + Meters y, Kilograms = Kilograms x + Kilograms y
, Seconds = Seconds x + Seconds y, Amperes = Amperes x + Amperes y
, Kelvins = Kelvins x + Kelvins y, Moles = Moles x + Moles y
, Candelas = Candelas x + Candelas y, \<dots> = more x * more y \<rparr>"
instance ..
@ -93,9 +102,9 @@ text \<open> We also define the inverse and division operations, and an abelian
instantiation Unit_ext :: ("{times,inverse}") inverse
begin
definition inverse_Unit_ext :: "'a Unit_ext \<Rightarrow> 'a Unit_ext"
where [code_unfold]:
"inverse x = \<lparr> Seconds = - Seconds x , Meters = - Meters x
, Kilograms = - Kilograms x, Amperes = - Amperes x
where [code_unfold, si_def]:
"inverse x = \<lparr> Meters = - Meters x, Kilograms = - Kilograms x
, Seconds = - Seconds x, Amperes = - Amperes x
, Kelvins = - Kelvins x, Moles = - Moles x
, Candelas = - Candelas x, \<dots> = inverse (more x) \<rparr>"
@ -155,13 +164,32 @@ subsection \<open> Basic SI-types \<close>
text \<open> We provide a syntax for type-expressions; The definition of
the basic type constructors is straight-forward via a one-elementary set. \<close>
typedef Length ("L") = "UNIV :: unit set" .. setup_lifting type_definition_Length
typedef Mass ("M") = "UNIV :: unit set" .. setup_lifting type_definition_Mass
typedef Time ("T") = "UNIV :: unit set" .. setup_lifting type_definition_Time
typedef Current ("I") = "UNIV :: unit set" .. setup_lifting type_definition_Current
typedef Temperature ("\<Theta>") = "UNIV :: unit set" .. setup_lifting type_definition_Temperature
typedef Amount ("N") = "UNIV :: unit set" .. setup_lifting type_definition_Amount
typedef Intensity ("J") = "UNIV :: unit set" .. setup_lifting type_definition_Intensity
typedef Length = "UNIV :: unit set" .. setup_lifting type_definition_Length
typedef Mass = "UNIV :: unit set" .. setup_lifting type_definition_Mass
typedef Time = "UNIV :: unit set" .. setup_lifting type_definition_Time
typedef Current = "UNIV :: unit set" .. setup_lifting type_definition_Current
typedef Temperature = "UNIV :: unit set" .. setup_lifting type_definition_Temperature
typedef Amount = "UNIV :: unit set" .. setup_lifting type_definition_Amount
typedef Intensity = "UNIV :: unit set" .. setup_lifting type_definition_Intensity
typedef NoDimension = "UNIV :: unit set" .. setup_lifting type_definition_NoDimension
type_synonym M = Mass
type_synonym L = Length
type_synonym T = Time
type_synonym I = Current
type_synonym \<Theta> = Temperature
type_synonym N = Amount
type_synonym J = Intensity
type_notation NoDimension ("\<one>")
translations
(type) "M" <= (type) "Mass"
(type) "L" <= (type) "Length"
(type) "T" <= (type) "Time"
(type) "I" <= (type) "Current"
(type) "\<Theta>" <= (type) "Temperature"
(type) "N" <= (type) "Amount"
(type) "J" <= (type) "Intensity"
subsection \<open> SI-type expressions and SI-type interpretation \<close>
@ -238,6 +266,12 @@ begin
instance by (intro_classes, auto simp add: si_sem_Intensity_def is_BaseUnit_def, (transfer, simp)+)
end
instantiation NoDimension :: si_type
begin
definition si_sem_NoDimension :: "NoDimension itself \<Rightarrow> Unit" where "si_sem_NoDimension x = 1"
instance by (intro_classes, auto simp add: si_sem_NoDimension_def is_BaseUnit_def, (transfer, simp)+)
end
lemma base_units [simp]:
"is_BaseUnit SI(Length)" "is_BaseUnit SI(Mass)" "is_BaseUnit SI(Time)"
"is_BaseUnit SI(Current)" "is_BaseUnit SI(Temperature)" "is_BaseUnit SI(Amount)"
@ -313,5 +347,9 @@ type_synonym Frequency = "T\<^sup>-\<^sup>1"
type_synonym Energy = "L\<^sup>2\<cdot>M\<cdot>T\<^sup>-\<^sup>2"
type_synonym Power = "L\<^sup>2\<cdot>M\<cdot>T\<^sup>-\<^sup>3"
type_synonym Force = "L\<cdot>M\<cdot>T\<^sup>-\<^sup>2"
type_synonym Pressure = "L\<^sup>-\<^sup>1\<cdot>M\<cdot>T\<^sup>-\<^sup>2"
type_synonym Charge = "I\<cdot>T"
type_synonym PotentialDifference = "L\<^sup>2\<cdot>M\<cdot>T\<^sup>-\<^sup>3\<cdot>I\<^sup>-\<^sup>1"
type_synonym Capacitance = "L\<^sup>-\<^sup>2\<cdot>M\<^sup>-\<^sup>1\<cdot>T\<^sup>4\<cdot>I\<^sup>2"
end