Some renaming, cleaning, and addition of astronomical units

This commit is contained in:
Simon Foster 2020-03-18 16:02:09 +00:00
parent 25f2c76f1b
commit 31c0dd56e8
6 changed files with 58 additions and 32 deletions

View File

@ -1,3 +1,5 @@
section \<open> Meta-Theory for ISQ \<close>
theory ISQ
imports ISQ_Dimensions ISQ_Quantities ISQ_Proof ISQ_Algebra
begin end

View File

@ -6,59 +6,62 @@ begin
named_theorems si_transfer
definition magQuant :: "'a['u::dim_type] \<Rightarrow> 'a" ("\<lbrakk>_\<rbrakk>\<^sub>Q") where
[si_def]: "magQuant x = mag (fromQ x)"
definition magQ :: "'a['u::dim_type] \<Rightarrow> 'a" ("\<lbrakk>_\<rbrakk>\<^sub>Q") where
[si_def]: "magQ x = mag (fromQ x)"
definition dimQ :: "'a['u::dim_type] \<Rightarrow> Dimension" where
[si_def]: "dimQ x = dim (fromQ x)"
lemma unit_eq_iff_mag_eq [si_transfer]:
"x = y \<longleftrightarrow> \<lbrakk>x\<rbrakk>\<^sub>Q = \<lbrakk>y\<rbrakk>\<^sub>Q"
by (auto simp add: magQuant_def, transfer, simp)
by (auto simp add: magQ_def, transfer, simp)
lemma unit_equiv_iff [si_transfer]:
fixes x :: "'a['u\<^sub>1::dim_type]" and y :: "'a['u\<^sub>2::dim_type]"
shows "x \<cong>\<^sub>Q y \<longleftrightarrow> \<lbrakk>x\<rbrakk>\<^sub>Q = \<lbrakk>y\<rbrakk>\<^sub>Q \<and> QD('u\<^sub>1) = QD('u\<^sub>2)"
proof -
have "\<forall>t ta. (ta::'a['u\<^sub>2]) = t \<or> mag (fromQ ta) \<noteq> mag (fromQ t)"
by (simp add: magQuant_def unit_eq_iff_mag_eq)
by (simp add: magQ_def unit_eq_iff_mag_eq)
then show ?thesis
by (metis (full_types) qequiv.rep_eq coerceQuant_eq_iff2 qeq magQuant_def)
by (metis (full_types) qequiv.rep_eq coerceQuant_eq_iff2 qeq magQ_def)
qed
lemma unit_le_iff_magn_le [si_transfer]:
"x \<le> y \<longleftrightarrow> \<lbrakk>x\<rbrakk>\<^sub>Q \<le> \<lbrakk>y\<rbrakk>\<^sub>Q"
by (auto simp add: magQuant_def; (transfer, simp))
by (auto simp add: magQ_def; (transfer, simp))
lemma magQuant_zero [si_eq]: "\<lbrakk>0\<rbrakk>\<^sub>Q = 0"
by (simp add: magQuant_def, transfer, simp)
lemma magQ_zero [si_eq]: "\<lbrakk>0\<rbrakk>\<^sub>Q = 0"
by (simp add: magQ_def, transfer, simp)
lemma magQuant_one [si_eq]: "\<lbrakk>1\<rbrakk>\<^sub>Q = 1"
by (simp add: magQuant_def, transfer, simp)
lemma magQ_one [si_eq]: "\<lbrakk>1\<rbrakk>\<^sub>Q = 1"
by (simp add: magQ_def, transfer, simp)
lemma magQuant_plus [si_eq]: "\<lbrakk>x + y\<rbrakk>\<^sub>Q = \<lbrakk>x\<rbrakk>\<^sub>Q + \<lbrakk>y\<rbrakk>\<^sub>Q"
by (simp add: magQuant_def, transfer, simp)
lemma magQ_plus [si_eq]: "\<lbrakk>x + y\<rbrakk>\<^sub>Q = \<lbrakk>x\<rbrakk>\<^sub>Q + \<lbrakk>y\<rbrakk>\<^sub>Q"
by (simp add: magQ_def, transfer, simp)
lemma magQuant_minus [si_eq]: "\<lbrakk>x - y\<rbrakk>\<^sub>Q = \<lbrakk>x\<rbrakk>\<^sub>Q - \<lbrakk>y\<rbrakk>\<^sub>Q"
by (simp add: magQuant_def, transfer, simp)
lemma magQ_minus [si_eq]: "\<lbrakk>x - y\<rbrakk>\<^sub>Q = \<lbrakk>x\<rbrakk>\<^sub>Q - \<lbrakk>y\<rbrakk>\<^sub>Q"
by (simp add: magQ_def, transfer, simp)
lemma magQuant_uminus [si_eq]: "\<lbrakk>- x\<rbrakk>\<^sub>Q = - \<lbrakk>x\<rbrakk>\<^sub>Q"
by (simp add: magQuant_def, transfer, simp)
lemma magQ_uminus [si_eq]: "\<lbrakk>- x\<rbrakk>\<^sub>Q = - \<lbrakk>x\<rbrakk>\<^sub>Q"
by (simp add: magQ_def, transfer, simp)
lemma magQuant_scaleQ [si_eq]: "\<lbrakk>x *\<^sub>Q y\<rbrakk>\<^sub>Q = x * \<lbrakk>y\<rbrakk>\<^sub>Q"
by (simp add: magQuant_def, transfer, simp)
lemma magQ_scaleQ [si_eq]: "\<lbrakk>x *\<^sub>Q y\<rbrakk>\<^sub>Q = x * \<lbrakk>y\<rbrakk>\<^sub>Q"
by (simp add: magQ_def, transfer, simp)
lemma magQuant_qtimes [si_eq]: "\<lbrakk>x \<^bold>\<cdot> y\<rbrakk>\<^sub>Q = \<lbrakk>x\<rbrakk>\<^sub>Q \<cdot> \<lbrakk>y\<rbrakk>\<^sub>Q"
by (simp add: magQuant_def, transfer, simp)
lemma magQ_qtimes [si_eq]: "\<lbrakk>x \<^bold>\<cdot> y\<rbrakk>\<^sub>Q = \<lbrakk>x\<rbrakk>\<^sub>Q \<cdot> \<lbrakk>y\<rbrakk>\<^sub>Q"
by (simp add: magQ_def, transfer, simp)
lemma magQuant_qinverse [si_eq]: "\<lbrakk>x\<^sup>-\<^sup>\<one>\<rbrakk>\<^sub>Q = inverse \<lbrakk>x\<rbrakk>\<^sub>Q"
by (simp add: magQuant_def, transfer, simp)
lemma magQ_qinverse [si_eq]: "\<lbrakk>x\<^sup>-\<^sup>\<one>\<rbrakk>\<^sub>Q = inverse \<lbrakk>x\<rbrakk>\<^sub>Q"
by (simp add: magQ_def, transfer, simp)
lemma magQuant_qdivivide [si_eq]: "\<lbrakk>(x::('a::field)[_]) \<^bold>/ y\<rbrakk>\<^sub>Q = \<lbrakk>x\<rbrakk>\<^sub>Q / \<lbrakk>y\<rbrakk>\<^sub>Q"
by (simp add: magQuant_def, transfer, simp add: field_class.field_divide_inverse)
lemma magQ_qdivivide [si_eq]: "\<lbrakk>(x::('a::field)[_]) \<^bold>/ y\<rbrakk>\<^sub>Q = \<lbrakk>x\<rbrakk>\<^sub>Q / \<lbrakk>y\<rbrakk>\<^sub>Q"
by (simp add: magQ_def, transfer, simp add: field_class.field_divide_inverse)
lemma magQuant_numeral [si_eq]: "\<lbrakk>numeral n\<rbrakk>\<^sub>Q = numeral n"
lemma magQ_numeral [si_eq]: "\<lbrakk>numeral n\<rbrakk>\<^sub>Q = numeral n"
apply (induct n, simp_all add: si_def)
apply (metis magQuant_def magQuant_one)
apply (metis magQuant_def magQuant_plus numeral_code(2))
apply (metis magQuant_def magQuant_one magQuant_plus numeral_code(3))
apply (metis magQ_def magQ_one)
apply (metis magQ_def magQ_plus numeral_code(2))
apply (metis magQ_def magQ_one magQ_plus numeral_code(3))
done
text \<open> The following tactic breaks an SI conjecture down to numeric and unit properties \<close>

View File

@ -166,6 +166,12 @@ text \<open> Since quantities can have dimension type expressions that are disti
definition coerceQuantT :: "'d\<^sub>2 itself \<Rightarrow> 'a['d\<^sub>1::dim_type] \<Rightarrow> 'a['d\<^sub>2::dim_type]" where
"QD('d\<^sub>1) = QD('d\<^sub>2) \<Longrightarrow> coerceQuantT t x = (toQ (fromQ x))"
syntax
"_QCOERCE" :: "type \<Rightarrow> logic \<Rightarrow> logic" ("QCOERCE[_]")
translations
"QCOERCE['t]" == "CONST coerceQuantT TYPE('t)"
subsection \<open> Predicates on Typed Quantities \<close>
text \<open> The standard HOL order \<^term>\<open>(\<le>)\<close> and equality \<^term>\<open>(=)\<close> have the homogeneous type
@ -237,7 +243,7 @@ lift_definition
qinverse :: "('n::field)['a::dim_type] \<Rightarrow> 'n['a\<^sup>-\<^sup>1]" ("(_\<^sup>-\<^sup>\<one>)" [999] 999)
is "inverse" by (simp add: inverse_Quantity_ext_def dim_ty_sem_DimInv_def)
abbreviation
abbreviation (input)
qdivide :: "('n::field)['a::dim_type] \<Rightarrow> 'n['b::dim_type] \<Rightarrow> 'n['a/'b]" (infixl "\<^bold>'/" 70) where
"qdivide x y \<equiv> x \<^bold>\<cdot> y\<^sup>-\<^sup>\<one>"

View File

@ -1,3 +1,5 @@
section \<open> Meta-Theory for SI Units \<close>
theory SI
imports SI_Units SI_Constants SI_Prefix SI_Derived SI_Accepted SI_Imperial
begin end

View File

@ -0,0 +1,13 @@
section \<open> Astronomical Constants \<close>
theory SI_Astronomical
imports SI_Accepted
begin
definition julian_year :: "'a::field[T]" where
"julian_year = 365.25 \<odot> day"
definition light_year :: "'a::field_char_0[L]" where
"light_year = QCOERCE[L] (\<^bold>c \<^bold>\<cdot> julian_year)"
end

View File

@ -3,7 +3,7 @@ chapter \<open> International System of Units \<close>
section \<open> SI Units Semantics \<close>
theory SI_Units
imports ISQ_Proof
imports ISQ
begin
text \<open> An SI unit is simply a particular kind of quantity. \<close>
@ -27,8 +27,8 @@ 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]: "\<lbrakk>BUNIT('u::basedim_type)\<rbrakk>\<^sub>Q = 1"
by (simp add: magQuant_def si_eq, transfer, simp)
lemma magQ_mk [si_eq]: "\<lbrakk>BUNIT('u::basedim_type)\<rbrakk>\<^sub>Q = 1"
by (simp add: magQ_def si_eq, transfer, simp)
text \<open> We now define the seven base units. Effectively, these definitions axiomatise given names
for the \<^term>\<open>1\<close> elements of the base quantities. \<close>