diff --git a/src/SI/ISQ_Proof.thy b/src/SI/ISQ_Proof.thy index 16209cf..9ce97b9 100644 --- a/src/SI/ISQ_Proof.thy +++ b/src/SI/ISQ_Proof.thy @@ -64,7 +64,7 @@ lemma magQuant_numeral [si_eq]: "\numeral n\\<^sub>Q = numeral n text \ The following tactic breaks an SI conjecture down to numeric and unit properties \ method si_simp uses add = - (simp add: add si_transfer si_eq) + (simp add: add si_transfer si_eq field_simps) text \ The next tactic additionally compiles the semantics of the underlying units \ diff --git a/src/SI/SI_Imperial.thy b/src/SI/SI_Imperial.thy index fca8011..9389db7 100644 --- a/src/SI/SI_Imperial.thy +++ b/src/SI/SI_Imperial.thy @@ -4,7 +4,7 @@ theory SI_Imperial imports SI_Accepted begin -subsection \ Definition \ +subsection \ Definitions \ default_sort field_char_0 @@ -28,7 +28,7 @@ definition mile :: "'a[L]" where default_sort type -subsection \ Properties \ +subsection \ Unit Equations \ lemma miles_to_yards: "mile = 1760 \ yard" by si_simp @@ -39,7 +39,7 @@ lemma miles_to_feet: "mile = 5280 \ foot" lemma mph_to_kmh: "1 \ (mile \<^bold>/ hour) = 1.609344 \ ((kilo \ meter) \<^bold>/ hour)" by si_simp -lemma celcius_to_farenheit: "(T::rat)\C = ((T - 32) \ 5/9)\F" - oops +lemma farenheit_to_celcius: "T\F = ((T - 32) \ 5/9)\C" + by si_simp end \ No newline at end of file diff --git a/src/SI/SI_Units.thy b/src/SI/SI_Units.thy index 342c175..27db167 100644 --- a/src/SI/SI_Units.thy +++ b/src/SI/SI_Units.thy @@ -1,7 +1,7 @@ section \ SI Units \ theory SI_Units - imports SI_Proof + imports ISQ_Proof begin lift_definition mk_base_unit :: "'u itself \ ('a::one)['u::dim_type]"