Theory HyperdualFunctionExtension
section ‹Hyperdual Extension of Functions›
theory HyperdualFunctionExtension
imports Hyperdual TwiceFieldDifferentiable
begin
text‹The following is an important fact in the derivation of the hyperdual extension.›
lemma
fixes x :: "('a :: comm_ring_1) hyperdual" and n :: nat
assumes "Base x = 0"
shows "x ^ (n + 3) = 0"
proof (induct n)
case 0
then show ?case
using assms hyperdual_power[of x 3] by simp
next
case (Suc n)
then show ?case
using assms power_Suc[of x "n + 3"] mult_zero_right add_Suc by simp
qed
text‹We define the extension of a function to the hyperdual numbers.›
primcorec hypext :: "(('a :: real_normed_field) ⇒ 'a) ⇒ 'a hyperdual ⇒ 'a hyperdual" (‹*h* _› [80] 80)
where
"Base ((*h* f) x) = f (Base x)"
| "Eps1 ((*h* f) x) = Eps1 x * deriv f (Base x)"
| "Eps2 ((*h* f) x) = Eps2 x * deriv f (Base x)"
| "Eps12 ((*h* f) x) = Eps12 x * deriv f (Base x) + Eps1 x * Eps2 x * deriv (deriv f) (Base x)"
text‹This has the expected behaviour when expressed in terms of the units.›
lemma hypext_Hyperdual_eq:
"(*h* f) (Hyperdual a b c d) =
Hyperdual (f a) (b * deriv f a) (c * deriv f a) (d * deriv f a + b * c * deriv (deriv f) a)"
by (simp add: hypext.code)
lemma hypext_Hyperdual_eq_parts:
"(*h* f) (Hyperdual a b c d) =
f a *⇩H ba + (b * deriv f a) *⇩H e1 + (c * deriv f a) *⇩H e2 +
(d * deriv f a + b * c * deriv (deriv f) a) *⇩H e12 "
by (metis Hyperdual_eq hypext_Hyperdual_eq)
text‹
The extension can be used to extract the function value, and first and second derivatives at x
when applied to @{term "x *⇩H re + e1 + e2 + 0 *⇩H e12"}, which we denote by @{term "β x"}.
›
definition hyperdualx :: "('a :: real_normed_field) ⇒ 'a hyperdual" ("β")
where "β x = (Hyperdual x 1 1 0)"
lemma hyperdualx_sel [simp]:
shows "Base (β x) = x"
and "Eps1 (β x) = 1"
and "Eps2 (β x) = 1"
and "Eps12 (β x) = 0"
by (simp_all add: hyperdualx_def)
lemma :
"(*h* f) (β x) = f x *⇩H ba + deriv f x *⇩H e1 + deriv f x *⇩H e2 + deriv (deriv f) x *⇩H e12"
by (simp add: hypext_Hyperdual_eq_parts hyperdualx_def)
lemma Base_hypext:
"Base ((*h* f) (β x)) = f x"
by (simp add: hyperdualx_def)
lemma Eps1_hypext:
"Eps1 ((*h* f) (β x)) = deriv f x"
by (simp add: hyperdualx_def)
lemma Eps2_hypext:
"Eps2 ((*h* f) (β x)) = deriv f x"
by (simp add: hyperdualx_def)
lemma Eps12_hypext:
"Eps12 ((*h* f) (β x)) = deriv (deriv f) x"
by (simp add: hyperdualx_def)
subsubsection‹Convenience Interface›
text‹Define a datatype to hold the function value, and the first and second derivative values.›