section \ Imperial Units via SI-Units\ theory SI_Imperial imports SI_Accepted begin subsection \ Definitions \ default_sort field_char_0 definition inch :: "'a[L]" where [si_eq]: "inch = 25.5 \ milli \ meter" definition foot :: "'a[L]" where [si_eq]: "foot = 0.3048 \ meter" definition yard :: "'a[L]" where [si_eq]: "yard = 0.9144 \ meter" definition mile :: "'a[L]" where [si_eq]: "mile = 1609.344 \ meter" definition nautical_mile :: "'a[L]" where [si_eq]: "nautical_mile = 1852 \ meter" definition knot :: "'a[L \ T\<^sup>-\<^sup>1]" where [si_eq]: "knot = 1 \ (nautical_mile \<^bold>/ hour)" definition pint :: "'a[Volume]" where [si_eq]: "pint = 0.56826125 \ litre" definition gallon :: "'a[Volume]" where [si_eq]: "gallon = 8 \ pint" definition pound :: "'a[M]" where [si_eq]: "pound = 0.45359237 \ kilogram" definition ounce :: "'a[M]" where [si_eq]: "ounce = 1/16 \ pound" definition stone :: "'a[M]" where [si_eq]: "stone = 14 \ pound" definition degrees_farenheit :: "'a \ 'a[\]" ("_\F" [999] 999) where [si_eq]: "degrees_farenheit x = (x + 459.67)\5/9 \ kelvin" default_sort type subsection \ Unit Equations \ lemma miles_to_yards: "mile = 1760 \ yard" by si_simp lemma miles_to_feet: "mile = 5280 \ foot" by si_simp lemma mph_to_kmh: "1 \ (mile \<^bold>/ hour) = 1.609344 \ ((kilo \ meter) \<^bold>/ hour)" by si_simp lemma farenheit_to_celcius: "T\F = ((T - 32) \ 5/9)\C" by si_simp end