section \ Quantities \ theory ISQ_Quantities imports ISQ_Dimensions begin section \ The Semantic SI-Unit-Type and its Operations \ record 'a Quantity = mag :: 'a \ \ Magnitude of the unit \ dim :: Dimension \ \ Dimension of the unit \ lemma Quantity_eq_intro: assumes "mag x = mag y" "dim x = dim y" "more x = more y" shows "x = y" by (simp add: assms) instantiation Quantity_ext :: (times, times) times begin definition [si_def]: "times_Quantity_ext x y = \ mag = mag x \ mag y, dim = dim x \ dim y, \ = more x \ more y \" instance .. end lemma mag_times [simp]: "mag (x \ y) = mag x \ mag y" by (simp add: times_Quantity_ext_def) lemma dim_times [simp]: "dim (x \ y) = dim x \ dim y" by (simp add: times_Quantity_ext_def) lemma more_times [simp]: "more (x \ y) = more x \ more y" by (simp add: times_Quantity_ext_def) instantiation Quantity_ext :: (zero, zero) zero begin definition "zero_Quantity_ext = \ mag = 0, dim = 1, \ = 0 \" instance .. end lemma mag_zero [simp]: "mag 0 = 0" by (simp add: zero_Quantity_ext_def) lemma dim_zero [simp]: "dim 0 = 1" by (simp add: zero_Quantity_ext_def) lemma more_zero [simp]: "more 0 = 0" by (simp add: zero_Quantity_ext_def) instantiation Quantity_ext :: (one, one) one begin definition [si_def]: "one_Quantity_ext = \ mag = 1, dim = 1, \ = 1 \" instance .. end lemma mag_one [simp]: "mag 1 = 1" by (simp add: one_Quantity_ext_def) lemma dim_one [simp]: "dim 1 = 1" by (simp add: one_Quantity_ext_def) lemma more_one [simp]: "more 1 = 1" by (simp add: one_Quantity_ext_def) instantiation Quantity_ext :: (inverse, inverse) inverse begin definition [si_def]: "inverse_Quantity_ext x = \ mag = inverse (mag x), dim = inverse (dim x), \ = inverse (more x) \" definition [si_def]: "divide_Quantity_ext x y = \ mag = mag x / mag y, dim = dim x / dim y, \ = more x / more y \" instance .. end lemma mag_inverse [simp]: "mag (inverse x) = inverse (mag x)" by (simp add: inverse_Quantity_ext_def) lemma dim_inverse [simp]: "dim (inverse x) = inverse (dim x)" by (simp add: inverse_Quantity_ext_def) lemma more_inverse [simp]: "more (inverse x) = inverse (more x)" by (simp add: inverse_Quantity_ext_def) lemma mag_divide [simp]: "mag (x / y) = mag x / mag y" by (simp add: divide_Quantity_ext_def) lemma dim_divide [simp]: "dim (x / y) = dim x / dim y" by (simp add: divide_Quantity_ext_def) lemma more_divide [simp]: "more (x / y) = more x / more y" by (simp add: divide_Quantity_ext_def) instance Quantity_ext :: (comm_monoid_mult, comm_monoid_mult) comm_monoid_mult by (intro_classes, simp_all add: one_Quantity_ext_def times_Quantity_ext_def mult.assoc, simp add: mult.commute) instance Quantity_ext :: (ab_group_mult, ab_group_mult) ab_group_mult by (intro_classes, rule Quantity_eq_intro, simp_all) instantiation Quantity_ext :: (ord, ord) ord begin definition less_eq_Quantity_ext :: "('a, 'b) Quantity_scheme \ ('a, 'b) Quantity_scheme \ bool" where "less_eq_Quantity_ext x y = (mag x \ mag y \ dim x = dim y \ more x \ more y)" definition less_Quantity_ext :: "('a, 'b) Quantity_scheme \ ('a, 'b) Quantity_scheme \ bool" where "less_Quantity_ext x y = (x \ y \ \ y \ x)" instance .. end instance Quantity_ext :: (order, order) order by (intro_classes, auto simp add: less_Quantity_ext_def less_eq_Quantity_ext_def) section \ The Abstract SI-\<^theory_text>\Quantity\-Type and its Operations \ text\We 'lift' SI type expressions to SI unit type expressions as follows:\ typedef (overloaded) ('n, 'u::dim_type) QuantT ("_[_]" [999,0] 999) = "{x :: 'n Quantity. dim x = QD('u)}" morphisms fromQ toQ by (rule_tac x="\ mag = undefined, dim = QD('u) \" in exI, simp) setup_lifting type_definition_QuantT text \ Coerce values when their dimensions are equivalent \ definition coerceQuantT :: "'u\<^sub>2 itself \ 'a['u\<^sub>1::dim_type] \ 'a['u\<^sub>2::dim_type]" where "QD('u\<^sub>1) = QD('u\<^sub>2) \ coerceQuantT t x = (toQ (fromQ x))" subsection\Predicates on Abstract SI-\<^theory_text>\Unit\-types\ text \ Two SI Unit types are orderable if their magnitude type is of class @{class "order"}, and have the same dimensions (not necessarily dimension types).\ lift_definition qless_eq :: "'n::order['a::dim_type] \ 'n['b::dim_type] \ bool" (infix "\\<^sub>Q" 50) is "(\)" . text\ Two SI Unit types are equivalent if they have the same dimensions (not necessarily dimension types). This equivalence the a vital point of the overall construction of SI Units. \ lift_definition qequiv :: "'n['a::dim_type] \ 'n['b::dim_type] \ bool" (infix "\\<^sub>Q" 50) is "(=)" . subsection\The Equivalence on Abstract SI-\<^theory_text>\Unit\-Types\ text\The latter predicate gives us an equivalence, but, unfortunately, not a congruence.\ lemma qequiv_refl [simp]: "a \\<^sub>Q a" by (simp add: qequiv_def) lemma qequiv_sym: "a \\<^sub>Q b \ b \\<^sub>Q a" by (simp add: qequiv_def) lemma qequiv_trans: "\ a \\<^sub>Q b; b \\<^sub>Q c \ \ a \\<^sub>Q c" by (simp add: qequiv_def) theorem qeq_iff_same_dim: fixes x y :: "'a['u::dim_type]" shows "x \\<^sub>Q y \ x = y" by (transfer, simp) (* the following series of equivalent statements ... *) lemma coerceQuant_eq_iff: fixes x :: "'a['u\<^sub>1::dim_type]" assumes "QD('u\<^sub>1) = QD('u\<^sub>2::dim_type)" shows "(coerceQuantT TYPE('u\<^sub>2) x) \\<^sub>Q x" by (metis qequiv.rep_eq assms coerceQuantT_def toQ_cases toQ_inverse) (* or equivalently *) lemma coerceQuant_eq_iff2: fixes x :: "'a['u\<^sub>1::dim_type]" assumes "QD('u\<^sub>1) = QD('u\<^sub>2::dim_type)" and "y = (coerceQuantT TYPE('u\<^sub>2) x)" shows "x \\<^sub>Q y" using qequiv_sym assms(1) assms(2) coerceQuant_eq_iff by blast lemma updown_eq_iff: fixes x :: "'a['u\<^sub>1::dim_type]" fixes y :: "'a['u\<^sub>2::dim_type]" assumes "QD('u\<^sub>1) = QD('u\<^sub>2::dim_type)" and "y = (toQ (fromQ x))" shows "x \\<^sub>Q y" by (simp add: assms(1) assms(2) coerceQuant_eq_iff2 coerceQuantT_def) text\This is more general that \y = x \ x \\<^sub>Q y\, since x and y may have different type.\ lemma qeq: fixes x :: "'a['u\<^sub>1::dim_type]" fixes y :: "'a['u\<^sub>2::dim_type]" assumes "x \\<^sub>Q y" shows "QD('u\<^sub>1) = QD('u\<^sub>2::dim_type)" by (metis (full_types) qequiv.rep_eq assms fromQ mem_Collect_eq) subsection\Operations on Abstract SI-Unit-Types\ lift_definition qtimes :: "('n::comm_ring_1)['a::dim_type] \ 'n['b::dim_type] \ 'n['a \'b]" (infixl "\<^bold>\" 69) is "(*)" by (simp add: dim_ty_sem_DimTimes_def times_Quantity_ext_def) lift_definition qinverse :: "('n::field)['a::dim_type] \ 'n['a\<^sup>-\<^sup>1]" ("(_\<^sup>-\<^sup>\)" [999] 999) is "inverse" by (simp add: inverse_Quantity_ext_def dim_ty_sem_DimInv_def) abbreviation qdivide :: "('n::field)['a::dim_type] \ 'n['b::dim_type] \ 'n['a/'b]" (infixl "\<^bold>'/" 70) where "qdivide x y \ x \<^bold>\ y\<^sup>-\<^sup>\" abbreviation qsq ("(_)\<^sup>\" [999] 999) where "u\<^sup>\ \ u\<^bold>\u" abbreviation qcube ("(_)\<^sup>\" [999] 999) where "u\<^sup>\ \ u\<^bold>\u\<^bold>\u" abbreviation qquart ("(_)\<^sup>\" [999] 999) where "u\<^sup>\ \ u\<^bold>\u\<^bold>\u\<^bold>\u" abbreviation qneq_sq ("(_)\<^sup>-\<^sup>\" [999] 999) where "u\<^sup>-\<^sup>\ \ (u\<^sup>\)\<^sup>-\<^sup>\" abbreviation qneq_cube ("(_)\<^sup>-\<^sup>\" [999] 999) where "u\<^sup>-\<^sup>\ \ (u\<^sup>\)\<^sup>-\<^sup>\" abbreviation qneq_quart ("(_)\<^sup>-\<^sup>\" [999] 999) where "u\<^sup>-\<^sup>\ \ (u\<^sup>\)\<^sup>-\<^sup>\" instantiation QuantT :: (zero,dim_type) zero begin lift_definition zero_QuantT :: "('a, 'b) QuantT" is "\ mag = 0, dim = QD('b) \" by simp instance .. end instantiation QuantT :: (one,dim_type) one begin lift_definition one_QuantT :: "('a, 'b) QuantT" is "\ mag = 1, dim = QD('b) \" by simp instance .. end instantiation QuantT :: (plus,dim_type) plus begin lift_definition plus_QuantT :: "'a['b] \ 'a['b] \ 'a['b]" is "\ x y. \ mag = mag x + mag y, dim = QD('b) \" by (simp) instance .. end instance QuantT :: (semigroup_add,dim_type) semigroup_add by (intro_classes, transfer, simp add: add.assoc) instance QuantT :: (ab_semigroup_add,dim_type) ab_semigroup_add by (intro_classes, transfer, simp add: add.commute) instance QuantT :: (monoid_add,dim_type) monoid_add by (intro_classes; (transfer, simp)) instance QuantT :: (comm_monoid_add,dim_type) comm_monoid_add by (intro_classes; transfer, simp) instantiation QuantT :: (uminus,dim_type) uminus begin lift_definition uminus_QuantT :: "'a['b] \ 'a['b]" is "\ x. \ mag = - mag x, dim = dim x \" by (simp) instance .. end instantiation QuantT :: (minus,dim_type) minus begin lift_definition minus_QuantT :: "'a['b] \ 'a['b] \ 'a['b]" is "\ x y. \ mag = mag x - mag y, dim = dim x \" by (simp) instance .. end instance QuantT :: (numeral,dim_type) numeral .. instance QuantT :: (ab_group_add,dim_type) ab_group_add by (intro_classes, (transfer, simp)+) instantiation QuantT :: (order,dim_type) order begin lift_definition less_eq_QuantT :: "'a['b] \ 'a['b] \ bool" is "\ x y. mag x \ mag y" . lift_definition less_QuantT :: "'a['b] \ 'a['b] \ bool" is "\ x y. mag x < mag y" . instance by (intro_classes, (transfer, simp add: less_le_not_le)+) end text\The scaling operator permits to multiply the magnitude of a unit; this scalar product produces what is called "prefixes" in the terminology of the SI system.\ lift_definition scaleQ :: "'a \ 'a::comm_ring_1['u::dim_type] \ 'a['u]" (infixr "*\<^sub>Q" 63) is "\ r x. \ mag = r * mag x, dim = QD('u) \" by simp notation scaleQ (infixr "\" 63) end