diff --git a/thys/Differential_Dynamic_Logic/Differential_Axioms.thy b/thys/Differential_Dynamic_Logic/Differential_Axioms.thy --- a/thys/Differential_Dynamic_Logic/Differential_Axioms.thy +++ b/thys/Differential_Dynamic_Logic/Differential_Axioms.thy @@ -1,2648 +1,2648 @@ theory "Differential_Axioms" imports Ordinary_Differential_Equations.ODE_Analysis "Ids" "Lib" "Syntax" "Denotational_Semantics" "Frechet_Correctness" "Axioms" "Coincidence" begin context ids begin section \Differential Axioms\ text \Differential axioms fall into two categories: Axioms for computing the derivatives of terms and axioms for proving properties of ODEs. The derivative axioms are all corollaries of the frechet correctness theorem. The ODE axioms are more involved, often requiring extensive use of the ODE libraries.\ subsection \Derivative Axioms\ definition diff_const_axiom :: "('sf, 'sc, 'sz) formula" where [axiom_defs]:"diff_const_axiom \ Equals (Differential ($f fid1 empty)) (Const 0)" definition diff_var_axiom :: "('sf, 'sc, 'sz) formula" where [axiom_defs]:"diff_var_axiom \ Equals (Differential (Var vid1)) (DiffVar vid1)" definition state_fun ::"'sf \ ('sf, 'sz) trm" where [axiom_defs]:"state_fun f = ($f f (\i. Var i))" definition diff_plus_axiom :: "('sf, 'sc, 'sz) formula" where [axiom_defs]:"diff_plus_axiom \ Equals (Differential (Plus (state_fun fid1) (state_fun fid2))) (Plus (Differential (state_fun fid1)) (Differential (state_fun fid2)))" definition diff_times_axiom :: "('sf, 'sc, 'sz) formula" where [axiom_defs]:"diff_times_axiom \ Equals (Differential (Times (state_fun fid1) (state_fun fid2))) (Plus (Times (Differential (state_fun fid1)) (state_fun fid2)) (Times (state_fun fid1) (Differential (state_fun fid2))))" \ \\[y=g(x)][y'=1](f(g(x))' = f(y)')\\ definition diff_chain_axiom::"('sf, 'sc, 'sz) formula" where [axiom_defs]:"diff_chain_axiom \ [[Assign vid2 (f1 fid2 vid1)]]([[DiffAssign vid2 (Const 1)]] (Equals (Differential ($f fid1 (singleton (f1 fid2 vid1)))) (Times (Differential (f1 fid1 vid2)) (Differential (f1 fid2 vid1)))))" subsection \ODE Axioms\ definition DWaxiom :: "('sf, 'sc, 'sz) formula" where [axiom_defs]:"DWaxiom = ([[EvolveODE (OVar vid1) (Predicational pid1)]](Predicational pid1))" definition DWaxiom' :: "('sf, 'sc, 'sz) formula" where [axiom_defs]:"DWaxiom' = ([[EvolveODE (OSing vid1 (Function fid1 (singleton (Var vid1)))) (Prop vid2 (singleton (Var vid1)))]](Prop vid2 (singleton (Var vid1))))" definition DCaxiom :: "('sf, 'sc, 'sz) formula" where [axiom_defs]:"DCaxiom = ( ([[EvolveODE (OVar vid1) (Predicational pid1)]]Predicational pid3) \ (([[EvolveODE (OVar vid1) (Predicational pid1)]](Predicational pid2)) \ ([[EvolveODE (OVar vid1) (And (Predicational pid1) (Predicational pid3))]]Predicational pid2)))" definition DEaxiom :: "('sf, 'sc, 'sz) formula" where [axiom_defs]:"DEaxiom = (([[EvolveODE (OSing vid1 (f1 fid1 vid1)) (p1 vid2 vid1)]] (P pid1)) \ ([[EvolveODE (OSing vid1 (f1 fid1 vid1)) (p1 vid2 vid1)]] [[DiffAssign vid1 (f1 fid1 vid1)]]P pid1))" definition DSaxiom :: "('sf, 'sc, 'sz) formula" where [axiom_defs]:"DSaxiom = (([[EvolveODE (OSing vid1 (f0 fid1)) (p1 vid2 vid1)]]p1 vid3 vid1) \ (Forall vid2 (Implies (Geq (Var vid2) (Const 0)) (Implies (Forall vid3 (Implies (And (Geq (Var vid3) (Const 0)) (Geq (Var vid2) (Var vid3))) (Prop vid2 (singleton (Plus (Var vid1) (Times (f0 fid1) (Var vid3))))))) ([[Assign vid1 (Plus (Var vid1) (Times (f0 fid1) (Var vid2)))]]p1 vid3 vid1)))))" \ \\(Q \ [c&Q](f(x)' \ g(x)'))\\ \ \\\\\ \ \\([c&Q](f(x) \ g(x))) --> (Q \ (f(x) \ g(x))\\ definition DIGeqaxiom :: "('sf, 'sc, 'sz) formula" where [axiom_defs]:"DIGeqaxiom = Implies (Implies (Prop vid1 empty) ([[EvolveODE (OVar vid1) (Prop vid1 empty)]](Geq (Differential (f1 fid1 vid1)) (Differential (f1 fid2 vid1))))) (Implies (Implies(Prop vid1 empty) (Geq (f1 fid1 vid1) (f1 fid2 vid1))) ([[EvolveODE (OVar vid1) (Prop vid1 empty)]](Geq (f1 fid1 vid1) (f1 fid2 vid1))))" \ \\g(x) > h(x) \ [x'=f(x), c & p(x)](g(x)' \ h(x)') \ [x'=f(x), c & p(x)]g(x) > h(x)\\ \ \\(Q \ [c&Q](f(x)' \ g(x)'))\\ \ \\\\\ \ \\([c&Q](f(x) > g(x))) <-> (Q \ (f(x) > g(x))\\ definition DIGraxiom :: "('sf, 'sc, 'sz) formula" where [axiom_defs]:"DIGraxiom = Implies (Implies (Prop vid1 empty) ([[EvolveODE (OVar vid1) (Prop vid1 empty)]](Geq (Differential (f1 fid1 vid1)) (Differential (f1 fid2 vid1))))) (Implies (Implies(Prop vid1 empty) (Greater (f1 fid1 vid1) (f1 fid2 vid1))) ([[EvolveODE (OVar vid1) (Prop vid1 empty)]](Greater (f1 fid1 vid1) (f1 fid2 vid1))))" \ \\[{1' = 1(1) & 1(1)}]2(1) <->\\ \ \\\2. [{1'=1(1), 2' = 2(1)*2 + 3(1) & 1(1)}]2(1)*)\\ definition DGaxiom :: "('sf, 'sc, 'sz) formula" where [axiom_defs]:"DGaxiom = (([[EvolveODE (OSing vid1 (f1 fid1 vid1)) (p1 vid1 vid1)]]p1 vid2 vid1) \ (Exists vid2 ([[EvolveODE (OProd (OSing vid1 (f1 fid1 vid1)) (OSing vid2 (Plus (Times (f1 fid2 vid1) (Var vid2)) (f1 fid3 vid1)))) (p1 vid1 vid1)]] p1 vid2 vid1)))" subsection \Proofs for Derivative Axioms\ lemma constant_deriv_inner: assumes interp:"\x i. (Functions I i has_derivative FunctionFrechet I i x) (at x)" shows "FunctionFrechet I id1 (vec_lambda (\i. sterm_sem I (empty i) (fst \))) (vec_lambda(\i. frechet I (empty i) (fst \) (snd \)))= 0" proof - have empty_zero:"(vec_lambda(\i. frechet I (empty i) (fst \) (snd \))) = 0" using local.empty_def Cart_lambda_cong frechet.simps(5) zero_vec_def apply auto apply(rule vec_extensionality) using local.empty_def Cart_lambda_cong frechet.simps(5) zero_vec_def by (simp add: local.empty_def) let ?x = "(vec_lambda (\i. sterm_sem I (empty i) (fst \)))" from interp have has_deriv:"(Functions I id1 has_derivative FunctionFrechet I id1 ?x) (at ?x)" by auto then have f_linear:"linear (FunctionFrechet I id1 ?x)" using Deriv.has_derivative_linear by auto then show ?thesis using empty_zero f_linear linear_0 by (auto) qed lemma constant_deriv_zero:"is_interp I \ directional_derivative I ($f id1 empty) \ = 0" apply(simp only: is_interp_def directional_derivative_def frechet.simps frechet_correctness) apply(rule constant_deriv_inner) apply(auto) done theorem diff_const_axiom_valid: "valid diff_const_axiom" apply(simp only: valid_def diff_const_axiom_def equals_sem) apply(rule allI | rule impI)+ apply(simp only: dterm_sem.simps constant_deriv_zero sterm_sem.simps) done theorem diff_var_axiom_valid: "valid diff_var_axiom" apply(auto simp add: diff_var_axiom_def valid_def directional_derivative_def) by (metis inner_prod_eq) theorem diff_plus_axiom_valid: "valid diff_plus_axiom" apply(auto simp add: diff_plus_axiom_def valid_def) subgoal for I a b using frechet_correctness[of I "(Plus (state_fun fid1) (state_fun fid2))" b] unfolding state_fun_def apply (auto intro: dfree.intros) unfolding directional_derivative_def by auto done theorem diff_times_axiom_valid: "valid diff_times_axiom" apply(auto simp add: diff_times_axiom_def valid_def) subgoal for I a b using frechet_correctness[of I "(Times (state_fun fid1) (state_fun fid2))" b] unfolding state_fun_def apply (auto intro: dfree.intros) unfolding directional_derivative_def by auto done subsection \Proofs for ODE Axioms\ lemma DW_valid:"valid DWaxiom" apply(unfold DWaxiom_def valid_def Let_def impl_sem ) apply(safe) apply(auto simp only: fml_sem.simps prog_sem.simps box_sem) subgoal for I aa ba ab bb sol t using mk_v_agree[of I "(OVar vid1)" "(ab,bb)" "sol t"] Vagree_univ[of "aa" "ba" "sol t" "ODEs I vid1 (sol t)"] solves_ode_domainD by (fastforce) done lemma DE_lemma: fixes ab bb::"'sz simple_state" and sol::"real \ 'sz simple_state" and I::"('sf, 'sc, 'sz) interp" shows "repd (mk_v I (OSing vid1 (f1 fid1 vid1)) (ab, bb) (sol t)) vid1 (dterm_sem I (f1 fid1 vid1) (mk_v I (OSing vid1 (f1 fid1 vid1)) (ab, bb) (sol t))) = mk_v I (OSing vid1 (f1 fid1 vid1)) (ab, bb) (sol t)" proof have set_eq:" {Inl vid1, Inr vid1} = {Inr vid1, Inl vid1}" by auto have agree:"Vagree (mk_v I (OSing vid1 (f1 fid1 vid1)) (ab, bb) (sol t)) (mk_xode I (OSing vid1 (f1 fid1 vid1)) (sol t)) {Inl vid1, Inr vid1}" using mk_v_agree[of I "(OSing vid1 (f1 fid1 vid1))" "(ab, bb)" "(sol t)"] unfolding semBV.simps using set_eq by auto have fact:"dterm_sem I (f1 fid1 vid1) (mk_v I (OSing vid1 (f1 fid1 vid1)) (ab, bb) (sol t)) = snd (mk_v I (OSing vid1 (f1 fid1 vid1)) (ab, bb) (sol t)) $ vid1" using agree unfolding Vagree_def dterm_sem.simps f1_def mk_xode.simps proof - assume alls:"(\i. Inl i \ {Inl vid1, Inr vid1} \ fst (mk_v I (OSing vid1 ($f fid1 (singleton (trm.Var vid1)))) (ab, bb) (sol t)) $ i = fst (sol t, ODE_sem I (OSing vid1 ($f fid1 (singleton (trm.Var vid1)))) (sol t)) $ i) \ (\i. Inr i \ {Inl vid1, Inr vid1} \ snd (mk_v I (OSing vid1 ($f fid1 (singleton (trm.Var vid1)))) (ab, bb) (sol t)) $ i = snd (sol t, ODE_sem I (OSing vid1 ($f fid1 (singleton (trm.Var vid1)))) (sol t)) $ i)" hence atVid'':"snd (mk_v I (OSing vid1 ($f fid1 (singleton (trm.Var vid1)))) (ab, bb) (sol t)) $ vid1 = sterm_sem I ($f fid1 (singleton (trm.Var vid1))) (sol t)" by auto have argsEq:"(\ i. dterm_sem I (singleton (trm.Var vid1) i) (mk_v I (OSing vid1 ($f fid1 (singleton (trm.Var vid1)))) (ab, bb) (sol t))) = (\ i. sterm_sem I (singleton (trm.Var vid1) i) (sol t))" using alls f1_def by auto thus "Functions I fid1 (\ i. dterm_sem I (singleton (trm.Var vid1) i) (mk_v I (OSing vid1 ($f fid1 (singleton (trm.Var vid1)))) (ab, bb) (sol t))) = snd (mk_v I (OSing vid1 ($f fid1 (singleton (trm.Var vid1)))) (ab, bb) (sol t)) $ vid1" by (simp only: atVid'' ODE_sem.simps sterm_sem.simps dterm_sem.simps) qed have eqSnd:"(\ y. if vid1 = y then snd (mk_v I (OSing vid1 (f1 fid1 vid1)) (ab, bb) (sol t)) $ vid1 else snd (mk_v I (OSing vid1 (f1 fid1 vid1)) (ab, bb) (sol t)) $ y) = snd (mk_v I (OSing vid1 (f1 fid1 vid1)) (ab, bb) (sol t))" by (simp add: vec_extensionality) have truth:"repd (mk_v I (OSing vid1 (f1 fid1 vid1)) (ab, bb) (sol t)) vid1 (dterm_sem I (f1 fid1 vid1) (mk_v I (OSing vid1 (f1 fid1 vid1)) (ab, bb) (sol t))) = mk_v I (OSing vid1 (f1 fid1 vid1)) (ab, bb) (sol t)" using fact by (auto simp only: eqSnd repd.simps fact prod.collapse split: if_split) thus "fst (repd (mk_v I (OSing vid1 (f1 fid1 vid1)) (ab, bb) (sol t)) vid1 (dterm_sem I (f1 fid1 vid1) (mk_v I (OSing vid1 (f1 fid1 vid1)) (ab, bb) (sol t)))) = fst (mk_v I (OSing vid1 (f1 fid1 vid1)) (ab, bb) (sol t))" "snd (repd (mk_v I (OSing vid1 (f1 fid1 vid1)) (ab, bb) (sol t)) vid1 (dterm_sem I (f1 fid1 vid1) (mk_v I (OSing vid1 (f1 fid1 vid1)) (ab, bb) (sol t)))) = snd (mk_v I (OSing vid1 (f1 fid1 vid1)) (ab, bb) (sol t)) " by auto qed lemma DE_valid:"valid DEaxiom" proof - have dsafe:"dsafe ($f fid1 (singleton (trm.Var vid1)))" unfolding singleton_def by(auto intro: dsafe.intros) have osafe:"osafe(OSing vid1 (f1 fid1 vid1))" unfolding f1_def empty_def singleton_def using dsafe osafe.intros dsafe.intros by (simp add: osafe_Sing dfree_Const) have fsafe:"fsafe (p1 vid2 vid1)" unfolding p1_def singleton_def using hpsafe_fsafe.intros(10) using dsafe dsafe_Fun_simps image_iff by (simp add: dfree_Const) show "valid DEaxiom" apply(auto simp only: DEaxiom_def valid_def Let_def iff_sem impl_sem) apply(auto simp only: fml_sem.simps prog_sem.simps mem_Collect_eq box_sem) proof - fix I::"('sf,'sc,'sz) interp" and aa ba ab bb sol and t::real and ac bc assume "is_interp I" assume allw:"\\. (\\ sol t. ((ab, bb), \) = (\, mk_v I (OSing vid1 (f1 fid1 vid1)) \ (sol t)) \ 0 \ t \ (sol solves_ode (\_. ODE_sem I (OSing vid1 (f1 fid1 vid1)))) {0..t} {x. mk_v I (OSing vid1 (f1 fid1 vid1)) \ x \ fml_sem I (p1 vid2 vid1)} \ (sol 0) = (fst \) ) \ \ \ fml_sem I (P pid1)" assume t:"0 \ t" assume aaba:"(aa, ba) = mk_v I (OSing vid1 (f1 fid1 vid1)) (ab, bb) (sol t)" assume solve:" (sol solves_ode (\_. ODE_sem I (OSing vid1 (f1 fid1 vid1)))) {0..t} {x. mk_v I (OSing vid1 (f1 fid1 vid1)) (ab, bb) x \ fml_sem I (p1 vid2 vid1)}" assume sol0:" (sol 0) = (fst (ab, bb)) " assume rep:" (ac, bc) = repd (mk_v I (OSing vid1 (f1 fid1 vid1)) (ab, bb) (sol t)) vid1 (dterm_sem I (f1 fid1 vid1) (mk_v I (OSing vid1 (f1 fid1 vid1)) (ab, bb) (sol t)))" have aaba_sem:"(aa,ba) \ fml_sem I (P pid1)" using allw t aaba solve sol0 rep by blast have truth:"repd (mk_v I (OSing vid1 (f1 fid1 vid1)) (ab, bb) (sol t)) vid1 (dterm_sem I (f1 fid1 vid1) (mk_v I (OSing vid1 (f1 fid1 vid1)) (ab, bb) (sol t))) = mk_v I (OSing vid1 (f1 fid1 vid1)) (ab, bb) (sol t)" using DE_lemma by auto show " repd (mk_v I (OSing vid1 (f1 fid1 vid1)) (ab, bb) (sol t)) vid1 (dterm_sem I (f1 fid1 vid1) (mk_v I (OSing vid1 (f1 fid1 vid1)) (ab, bb) (sol t))) \ fml_sem I (P pid1)" using aaba aaba_sem truth by (auto) next fix I::"('sf,'sc,'sz) interp" and aa ba ab bb sol and t::real assume "is_interp I" assume all:"\\. (\\ sol t. ((ab, bb), \) = (\, mk_v I (OSing vid1 (f1 fid1 vid1)) \ (sol t)) \ 0 \ t \ (sol solves_ode (\_. ODE_sem I (OSing vid1 (f1 fid1 vid1)))) {0..t} {x. mk_v I (OSing vid1 (f1 fid1 vid1)) \ x \ fml_sem I (p1 vid2 vid1)} \ (sol 0) = (fst \) ) \ (\\'. \' = repd \ vid1 (dterm_sem I (f1 fid1 vid1) \) \ \' \ fml_sem I (P pid1))" hence justW:"(\\ sol t. ((ab, bb), (aa, ba)) = (\, mk_v I (OSing vid1 (f1 fid1 vid1)) \ (sol t)) \ 0 \ t \ (sol solves_ode (\_. ODE_sem I (OSing vid1 (f1 fid1 vid1)))) {0..t} {x. mk_v I (OSing vid1 (f1 fid1 vid1)) \ x \ fml_sem I (p1 vid2 vid1)} \ (sol 0) = (fst \)) \ (\\'. \' = repd (aa, ba) vid1 (dterm_sem I (f1 fid1 vid1) (aa, ba)) \ \' \ fml_sem I (P pid1))" by (rule allE) assume t:"0 \ t" assume aaba:"(aa, ba) = mk_v I (OSing vid1 (f1 fid1 vid1)) (ab, bb) (sol t)" assume sol:"(sol solves_ode (\_. ODE_sem I (OSing vid1 (f1 fid1 vid1)))) {0..t} {x. mk_v I (OSing vid1 (f1 fid1 vid1)) (ab, bb) x \ fml_sem I (p1 vid2 vid1)}" assume sol0:" (sol 0) = (fst (ab, bb))" have "repd (aa, ba) vid1 (dterm_sem I (f1 fid1 vid1) (aa, ba)) \ fml_sem I (P pid1)" using justW t aaba sol sol0 by auto hence foo:"repd (mk_v I (OSing vid1 (f1 fid1 vid1)) (ab, bb) (sol t)) vid1 (dterm_sem I (f1 fid1 vid1) (mk_v I (OSing vid1 (f1 fid1 vid1)) (ab, bb) (sol t))) \ fml_sem I (P pid1)" using aaba by auto hence "repd (mk_v I (OSing vid1 (f1 fid1 vid1)) (ab, bb) (sol t)) vid1 (dterm_sem I (f1 fid1 vid1) (mk_v I (OSing vid1 (f1 fid1 vid1)) (ab, bb) (sol t))) = mk_v I (OSing vid1 (f1 fid1 vid1)) (ab, bb) (sol t)" using DE_lemma by auto thus "mk_v I (OSing vid1 (f1 fid1 vid1)) (ab, bb) (sol t) \ fml_sem I (P pid1)" using foo by auto qed qed lemma ODE_zero:"\i. Inl i \ BVO ODE \ Inr i \ BVO ODE \ ODE_sem I ODE \ $ i= 0" by(induction ODE, auto) lemma DE_sys_valid: assumes disj:"{Inl vid1, Inr vid1} \ BVO ODE = {}" shows "valid (([[EvolveODE (OProd (OSing vid1 (f1 fid1 vid1)) ODE) (p1 vid2 vid1)]] (P pid1)) \ ([[EvolveODE ((OProd (OSing vid1 (f1 fid1 vid1))ODE)) (p1 vid2 vid1)]] [[DiffAssign vid1 (f1 fid1 vid1)]]P pid1))" proof - have dsafe:"dsafe ($f fid1 (singleton (trm.Var vid1)))" unfolding singleton_def by(auto intro: dsafe.intros) have osafe:"osafe(OSing vid1 (f1 fid1 vid1))" unfolding f1_def empty_def singleton_def using dsafe osafe.intros dsafe.intros by (simp add: osafe_Sing dfree_Const) have fsafe:"fsafe (p1 vid2 vid1)" unfolding p1_def singleton_def using hpsafe_fsafe.intros(10) using dsafe dsafe_Fun_simps image_iff by (simp add: dfree_Const) show "valid (([[EvolveODE (OProd (OSing vid1 (f1 fid1 vid1)) ODE) (p1 vid2 vid1)]] (P pid1)) \ ([[EvolveODE ((OProd (OSing vid1 (f1 fid1 vid1)) ODE)) (p1 vid2 vid1)]] [[DiffAssign vid1 (f1 fid1 vid1)]]P pid1))" apply(auto simp only: DEaxiom_def valid_def Let_def iff_sem impl_sem) apply(auto simp only: fml_sem.simps prog_sem.simps mem_Collect_eq box_sem f1_def p1_def P_def expand_singleton) proof - fix I ::"('sf,'sc,'sz) interp" and aa ba ab bb sol and t::real and ac bc assume good:"is_interp I" assume bigAll:" \\. (\\ sol t. ((ab, bb), \) = (\, mk_v I (OProd (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0))) ODE) \ (sol t)) \ 0 \ t \ (sol solves_ode (\_. ODE_sem I (OProd(OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0))) ODE ))) {0..t} {x. Predicates I vid2 (\ i. dterm_sem I (if i = vid1 then trm.Var vid1 else Const 0) (mk_v I (OProd (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0)))ODE) \ x))} \ sol 0 = fst \) \ \ \ fml_sem I (Pc pid1)" let ?my\ = "mk_v I (OProd (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0)))ODE) (ab,bb) (sol t)" assume t:"0 \ t" assume aaba:"(aa, ba) = mk_v I (OProd (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0)))ODE) (ab, bb) (sol t)" assume sol:"(sol solves_ode (\_. ODE_sem I (OProd (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0)))ODE))) {0..t} {x. Predicates I vid2 (\ i. dterm_sem I (if i = vid1 then trm.Var vid1 else Const 0) (mk_v I (OProd (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0)))ODE) (ab, bb) x))}" assume sol0:"sol 0 = fst (ab, bb)" assume acbc:"(ac, bc) = repd (mk_v I (OProd (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0)))ODE) (ab, bb) (sol t)) vid1 (dterm_sem I ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0)) (mk_v I (OProd (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0)))ODE) (ab, bb) (sol t)))" have bigEx:"(\\ sol t. ((ab, bb), ?my\) = (\, mk_v I (OProd (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0)))ODE) \ (sol t)) \ 0 \ t \ (sol solves_ode (\_. ODE_sem I (OProd (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0)))ODE))) {0..t} {x. Predicates I vid2 (\ i. dterm_sem I (if i = vid1 then trm.Var vid1 else Const 0) (mk_v I (OProd (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0)))ODE) \ x))} \ sol 0 = fst \)" apply(rule exI[where x="(ab, bb)"]) apply(rule exI[where x="sol"]) apply(rule exI[where x="t"]) apply(rule conjI) apply(rule refl) apply(rule conjI) apply(rule t) apply(rule conjI) using sol apply blast by (rule sol0) have bigRes:"?my\ \ fml_sem I (Pc pid1)" using bigAll bigEx by blast have notin1:"Inl vid1 \ BVO ODE" using disj by auto have notin2:"Inr vid1 \ BVO ODE" using disj by auto have ODE_sem:"ODE_sem I ODE (sol t) $ vid1 = 0" using ODE_zero notin1 notin2 by blast have vec_eq:"(\ i. sterm_sem I (if i = vid1 then trm.Var vid1 else Const 0) (sol t)) = (\ i. dterm_sem I (if i = vid1 then trm.Var vid1 else Const 0) (mk_v I (OProd (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0)))ODE) (ab, bb) (sol t)))" apply(rule vec_extensionality) apply simp using mk_v_agree[of I "(OProd (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0)))ODE)" "(ab, bb)" "(sol t)"] by(simp add: Vagree_def) have sem_eq:"(?my\ \ fml_sem I (Pc pid1)) = ((repd (mk_v I (OProd (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0)))ODE) (ab, bb) (sol t)) vid1 (dterm_sem I ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0)) (mk_v I (OProd (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0)))ODE) (ab, bb) (sol t)))) \ fml_sem I (Pc pid1))" apply(rule coincidence_formula) subgoal by simp subgoal by (rule Iagree_refl) using mk_v_agree[of "I" "(OProd (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0)))ODE)" "(ab, bb)" "(sol t)"] unfolding Vagree_def apply simp apply(erule conjE)+ apply(erule allE[where x="vid1"])+ apply(simp add: ODE_sem) using vec_eq by simp show "repd (mk_v I (OProd (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0)))ODE) (ab, bb) (sol t)) vid1 (dterm_sem I ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0)) (mk_v I (OProd (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0)))ODE) (ab, bb) (sol t))) \ fml_sem I (Pc pid1)" using bigRes sem_eq by blast next fix I::"('sf,'sc,'sz)interp" and aa ba ab bb sol and t::real assume good_interp:"is_interp I" assume all:"\\. (\\ sol t. ((ab, bb), \) = (\, mk_v I (OProd (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0)))ODE) \ (sol t)) \ 0 \ t \ (sol solves_ode (\_. ODE_sem I (OProd (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0)))ODE))) {0..t} {x. Predicates I vid2 (\ i. dterm_sem I (if i = vid1 then trm.Var vid1 else Const 0) (mk_v I (OProd (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0)))ODE) \ x))} \ sol 0 = fst \) \ (\\'. \' = repd \ vid1 (dterm_sem I ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0)) \) \ \' \ fml_sem I (Pc pid1))" let ?my\ = "mk_v I (OProd (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0)))ODE) (ab, bb) (sol t)" assume t:"0 \ t" assume aaba:"(aa, ba) = mk_v I (OProd (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0)))ODE) (ab, bb) (sol t)" assume sol:" (sol solves_ode (\_. ODE_sem I (OProd (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0)))ODE))) {0..t} {x. Predicates I vid2 (\ i. dterm_sem I (if i = vid1 then trm.Var vid1 else Const 0) (mk_v I (OProd (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0)))ODE) (ab, bb) x))}" assume sol0:"sol 0 = fst (ab, bb)" have bigEx:"(\\ sol t. ((ab, bb), ?my\) = (\, mk_v I (OProd (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0)))ODE) \ (sol t)) \ 0 \ t \ (sol solves_ode (\_. ODE_sem I (OProd (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0)))ODE))) {0..t} {x. Predicates I vid2 (\ i. dterm_sem I (if i = vid1 then trm.Var vid1 else Const 0) (mk_v I (OProd (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0)))ODE) \ x))} \ sol 0 = fst \)" apply(rule exI[where x="(ab, bb)"]) apply(rule exI[where x=sol]) apply(rule exI[where x=t]) apply(rule conjI) apply(rule refl) apply(rule conjI) apply(rule t) apply(rule conjI) using sol sol0 by(blast)+ have rep_sem_eq:"repd (mk_v I (OProd (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0)))ODE) (ab, bb) (sol t)) vid1 (dterm_sem I ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0)) (mk_v I (OProd (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0)))ODE) (ab, bb) (sol t))) \ fml_sem I (Pc pid1) = (repd ?my\ vid1 (dterm_sem I ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0)) ?my\) \ fml_sem I (Pc pid1))" apply(rule coincidence_formula) subgoal by simp subgoal by (rule Iagree_refl) by(simp add: Vagree_def) have notin1:"Inl vid1 \ BVO ODE" using disj by auto have notin2:"Inr vid1 \ BVO ODE" using disj by auto have ODE_sem:"ODE_sem I ODE (sol t) $ vid1 = 0" using ODE_zero notin1 notin2 by blast have vec_eq:" (\ i. dterm_sem I (if i = vid1 then trm.Var vid1 else Const 0) (mk_v I (OProd (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0)))ODE) (ab, bb) (sol t))) = (\ i. sterm_sem I (if i = vid1 then trm.Var vid1 else Const 0) (sol t))" apply(rule vec_extensionality) using mk_v_agree[of I "(OProd (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0)))ODE)" "(ab, bb)" "(sol t)"] by (simp add: Vagree_def) have sem_eq: "(repd ?my\ vid1 (dterm_sem I ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0)) ?my\) \ fml_sem I (Pc pid1)) = (mk_v I (OProd (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0)))ODE) (ab, bb) (sol t) \ fml_sem I (Pc pid1)) " apply(rule coincidence_formula) subgoal by simp subgoal by (rule Iagree_refl) using mk_v_agree[of I "(OProd (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0)))ODE)" "(ab, bb)" "(sol t)"] unfolding Vagree_def apply simp apply(erule conjE)+ apply(erule allE[where x=vid1])+ by (simp add: ODE_sem vec_eq) have some_sem:"repd (mk_v I (OProd (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0)))ODE) (ab, bb) (sol t)) vid1 (dterm_sem I ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0)) (mk_v I (OProd (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0)))ODE) (ab, bb) (sol t))) \ fml_sem I (Pc pid1)" using rep_sem_eq using all bigEx by blast have bigImp:"(\\'. \' = repd ?my\ vid1 (dterm_sem I ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0)) ?my\) \ \' \ fml_sem I (Pc pid1))" apply(rule allI) apply(rule impI) apply auto using some_sem by auto have fml_sem:"repd ?my\ vid1 (dterm_sem I ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0)) ?my\) \ fml_sem I (Pc pid1)" using sem_eq bigImp by blast show "mk_v I (OProd (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0)))ODE) (ab, bb) (sol t) \ fml_sem I (Pc pid1)" using fml_sem sem_eq by blast qed qed lemma DC_valid:"valid DCaxiom" proof (auto simp only: fml_sem.simps prog_sem.simps DCaxiom_def valid_def iff_sem impl_sem box_sem, auto) fix I::"('sf,'sc,'sz) interp" and aa ba bb sol t assume "is_interp I" and all3:"\a b. (\sola. sol 0 = sola 0 \ (\t. (a, b) = mk_v I (OVar vid1) (sola 0, bb) (sola t) \ 0 \ t \ (sola solves_ode (\a. ODEs I vid1)) {0..t} {x. mk_v I (OVar vid1) (sola 0, bb) x \ Contexts I pid1 UNIV})) \ (a, b) \ Contexts I pid3 UNIV" and all2:"\a b. (\sola. sol 0 = sola 0 \ (\t. (a, b) = mk_v I (OVar vid1) (sola 0, bb) (sola t) \ 0 \ t \ (sola solves_ode (\a. ODEs I vid1)) {0..t} {x. mk_v I (OVar vid1) (sola 0, bb) x \ Contexts I pid1 UNIV})) \ (a, b) \ Contexts I pid2 UNIV" and t:"0 \ t" and aaba:"(aa, ba) = mk_v I (OVar vid1) (sol 0, bb) (sol t)" and sol:"(sol solves_ode (\a. ODEs I vid1)) {0..t} {x. mk_v I (OVar vid1) (sol 0, bb) x \ Contexts I pid1 UNIV \ mk_v I (OVar vid1) (sol 0, bb) x \ Contexts I pid3 UNIV}" from sol have sol1:"(sol solves_ode (\a. ODEs I vid1)) {0..t} {x. mk_v I (OVar vid1) (sol 0, bb) x \ Contexts I pid1 UNIV}" by (metis (mono_tags, lifting) Collect_mono solves_ode_supset_range) from all2 have all2':"\v. (\sola. sol 0 = sola 0 \ (\t. v = mk_v I (OVar vid1) (sola 0, bb) (sola t) \ 0 \ t \ (sola solves_ode (\a. ODEs I vid1)) {0..t} {x. mk_v I (OVar vid1) (sola 0, bb) x \ Contexts I pid1 UNIV})) \ v \ Contexts I pid2 UNIV" by auto show "mk_v I (OVar vid1) (sol 0, bb) (sol t) \ Contexts I pid2 UNIV" apply(rule all2'[of "mk_v I (OVar vid1) (sol 0, bb) (sol t)"]) apply(rule exI[where x=sol]) apply(rule conjI) subgoal by (rule refl) subgoal using t sol1 by auto done next fix I::"('sf,'sc,'sz) interp" and aa ba bb sol t assume "is_interp I" and all3:"\a b. (\sola. sol 0 = sola 0 \ (\t. (a, b) = mk_v I (OVar vid1) (sola 0, bb) (sola t) \ 0 \ t \ (sola solves_ode (\a. ODEs I vid1)) {0..t} {x. mk_v I (OVar vid1) (sola 0, bb) x \ Contexts I pid1 UNIV})) \ (a, b) \ Contexts I pid3 UNIV" and all2:"\a b. (\sola. sol 0 = sola 0 \ (\t. (a, b) = mk_v I (OVar vid1) (sola 0, bb) (sola t) \ 0 \ t \ (sola solves_ode (\a. ODEs I vid1)) {0..t} {x. mk_v I (OVar vid1) (sola 0, bb) x \ Contexts I pid1 UNIV \ mk_v I (OVar vid1) (sola 0, bb) x \ Contexts I pid3 UNIV})) \ (a, b) \ Contexts I pid2 UNIV" and t:"0 \ t" and aaba:"(aa, ba) = mk_v I (OVar vid1) (sol 0, bb) (sol t)" and sol:"(sol solves_ode (\a. ODEs I vid1)) {0..t} {x. mk_v I (OVar vid1) (sol 0, bb) x \ Contexts I pid1 UNIV}" from all2 have all2':"\v. (\sola. sol 0 = sola 0 \ (\t. v = mk_v I (OVar vid1) (sola 0, bb) (sola t) \ 0 \ t \ (sola solves_ode (\a. ODEs I vid1)) {0..t} {x. mk_v I (OVar vid1) (sola 0, bb) x \ Contexts I pid1 UNIV \ mk_v I (OVar vid1) (sola 0, bb) x \ Contexts I pid3 UNIV})) \ v \ Contexts I pid2 UNIV" by auto from all3 have all3':"\v. (\sola. sol 0 = sola 0 \ (\t. v = mk_v I (OVar vid1) (sola 0, bb) (sola t) \ 0 \ t \ (sola solves_ode (\a. ODEs I vid1)) {0..t} {x. mk_v I (OVar vid1) (sola 0, bb) x \ Contexts I pid1 UNIV})) \ v \ Contexts I pid3 UNIV" by auto have inp1:"\s. 0 \ s \ s \ t \ mk_v I (OVar vid1) (sol 0, bb) (sol s) \ Contexts I pid1 UNIV" using sol solves_odeD atLeastAtMost_iff by blast have inp3:"\s. 0 \ s \ s \ t \ mk_v I (OVar vid1) (sol 0, bb) (sol s) \ Contexts I pid3 UNIV" apply(rule all3') subgoal for s apply(rule exI [where x=sol]) apply(rule conjI) subgoal by (rule refl) apply(rule exI [where x=s]) apply(rule conjI) subgoal by (rule refl) apply(rule conjI) subgoal by assumption subgoal using sol by (meson atLeastatMost_subset_iff order_refl solves_ode_subset) done done have inp13:"\s. 0 \ s \ s \ t \ mk_v I (OVar vid1) (sol 0, bb) (sol s) \ Contexts I pid1 UNIV \ mk_v I (OVar vid1) (sol 0, bb) (sol s) \ Contexts I pid3 UNIV" using inp1 inp3 by auto have sol13:"(sol solves_ode (\a. ODEs I vid1)) {0..t} {x. mk_v I (OVar vid1) (sol 0, bb) x \ Contexts I pid1 UNIV \ mk_v I (OVar vid1) (sol 0, bb) x \ Contexts I pid3 UNIV}" apply(rule solves_odeI) subgoal using sol by (rule solves_odeD) subgoal for s using inp13[of s] by auto done show "mk_v I (OVar vid1) (sol 0, bb) (sol t) \ Contexts I pid2 UNIV" using t sol13 all2'[of "mk_v I (OVar vid1) (sol 0, bb) (sol t)"] by auto qed lemma DS_valid:"valid DSaxiom" proof - have dsafe:"dsafe($f fid1 (\i. Const 0))" using dsafe_Const by auto have osafe:"osafe(OSing vid1 (f0 fid1))" unfolding f0_def empty_def using dsafe osafe.intros by (simp add: osafe_Sing dfree_Const) have fsafe:"fsafe(p1 vid2 vid1)" unfolding p1_def apply(rule fsafe_Prop) using singleton.simps dsafe_Const by (auto intro: dfree.intros) show "valid DSaxiom" apply(auto simp only: DSaxiom_def valid_def Let_def iff_sem impl_sem box_sem) apply(auto simp only: fml_sem.simps prog_sem.simps mem_Collect_eq iff_sem impl_sem box_sem forall_sem) proof - fix I::"('sf,'sc,'sz) interp" and a b r aa ba assume good_interp:"is_interp I" assume allW:"\\. (\\ sol t. ((a, b), \) = (\, mk_v I (OSing vid1 (f0 fid1)) \ (sol t)) \ 0 \ t \ (sol solves_ode (\_. ODE_sem I (OSing vid1 (f0 fid1)))) {0..t} {x. mk_v I (OSing vid1 (f0 fid1)) \ x \ fml_sem I (p1 vid2 vid1)} \ (sol 0) = (fst \)) \ \ \ fml_sem I (p1 vid3 vid1)" assume "dterm_sem I (Const 0) (repv (a, b) vid2 r) \ dterm_sem I (trm.Var vid2) (repv (a, b) vid2 r)" hence leq:"0 \ r" by (auto) assume "\ra. repv (repv (a, b) vid2 r) vid3 ra \ {v. dterm_sem I (Const 0) v \ dterm_sem I (trm.Var vid3) v} \ {v. dterm_sem I (trm.Var vid3) v \ dterm_sem I (trm.Var vid2) v} \ Predicates I vid2 (\ i. dterm_sem I (singleton (Plus (trm.Var vid1) (Times (f0 fid1) (trm.Var vid3))) i) (repv (repv (a, b) vid2 r) vid3 ra))" hence constraint:"\ra. (0 \ ra \ ra \ r) \ (repv (repv (a, b) vid2 r) vid3 ra) \ fml_sem I (Prop vid2 (singleton (Plus (Var vid1) (Times (f0 fid1) (Var vid3)))))" using leq by auto assume aaba:" (aa, ba) = repv (repv (a, b) vid2 r) vid1 (dterm_sem I (Plus (trm.Var vid1) (Times (f0 fid1) (trm.Var vid2))) (repv (a, b) vid2 r))" let ?abba = "repv (repd (a, b) vid1 (Functions I fid1 (\ i. 0))) vid1 (dterm_sem I (Plus (trm.Var vid1) (Times (f0 fid1) (trm.Var vid2))) (repv (a, b) vid2 r))" from allW have thisW:"(\\ sol t. ((a, b), ?abba) = (\, mk_v I (OSing vid1 (f0 fid1)) \ (sol t)) \ 0 \ t \ (sol solves_ode (\_. ODE_sem I (OSing vid1 (f0 fid1)))) {0..t} {x. mk_v I (OSing vid1 (f0 fid1)) \ x \ fml_sem I (p1 vid2 vid1)} \ (sol 0) = (fst \)) \ ?abba \ fml_sem I (p1 vid3 vid1)" by blast let ?c = "Functions I fid1 (\ _. 0)" let ?sol = "(\t. \ i. if i = vid1 then (a $ i) + ?c * t else (a $ i))" have agrees:"Vagree (mk_v I (OSing vid1 (f0 fid1)) (a, b) (?sol r)) (a, b) (- semBV I (OSing vid1 (f0 fid1))) \ Vagree (mk_v I (OSing vid1 (f0 fid1)) (a, b) (?sol r)) (mk_xode I (OSing vid1 (f0 fid1)) (?sol r)) (semBV I (OSing vid1 (f0 fid1)))" using mk_v_agree[of "I" "(OSing vid1 (f0 fid1))" "(a,b)" "(?sol r)"] by auto have prereq1a:"fst ?abba = fst (mk_v I (OSing vid1 (f0 fid1)) (a,b) (?sol r))" using agrees aaba apply (auto simp add: aaba Vagree_def) apply (rule vec_extensionality) subgoal for i apply (cases "i = vid1") using vne12 agrees Vagree_def apply (auto simp add: aaba f0_def empty_def) done apply (rule vec_extensionality) subgoal for i apply (cases "i = vid1") apply(auto simp add: f0_def empty_def) done done have prereq1b:"snd (?abba) = snd (mk_v I (OSing vid1 (f0 fid1)) (a,b) (?sol r))" using agrees aaba apply (auto simp add: aaba Vagree_def) apply (rule vec_extensionality) subgoal for i apply (cases "i = vid1") using vne12 agrees Vagree_def apply (auto simp add: aaba f0_def empty_def ) done done have "?abba = mk_v I (OSing vid1 (f0 fid1)) (a,b) (?sol r)" using prod_eq_iff prereq1a prereq1b by blast hence req1:"((a, b), ?abba) = ((a, b), mk_v I (OSing vid1 (f0 fid1)) (a,b) (?sol r))" by auto have "sterm_sem I ($f fid1 (\i. Const 0)) b = Functions I fid1 (\ i. 0)" by auto hence vec_simp:"(\a b. \ i. if i = vid1 then sterm_sem I ($f fid1 (\i. Const 0)) b else 0) = (\a b. \ i. if i = vid1 then Functions I fid1 (\ i. 0) else 0)" by (auto simp add: vec_eq_iff cong: if_cong) have sub: "{0..r} \ UNIV" by auto have sub2:"{x. mk_v I (OSing vid1 (f0 fid1)) (a,b) x \ fml_sem I (p1 vid2 vid1)} \ UNIV" by auto have req3:"(?sol solves_ode (\_. ODE_sem I (OSing vid1 (f0 fid1)))) {0..r} {x. mk_v I (OSing vid1 (f0 fid1)) (a,b) x \ fml_sem I (p1 vid2 vid1)}" apply(auto simp add: f0_def empty_def vec_simp) apply(rule solves_odeI) apply(auto simp only: has_vderiv_on_def has_vector_derivative_def box_sem) apply (rule has_derivative_vec[THEN has_derivative_eq_rhs]) defer apply (rule ext) apply (subst scaleR_vec_def) apply (rule refl) apply (auto intro!: derivative_eq_intros) \ \Domain constraint satisfied\ using constraint apply (auto) subgoal for t apply(erule allE[where x="t"]) apply(auto simp add: p1_def) proof - have eq:"(\ i. dterm_sem I (if i = vid1 then Plus (trm.Var vid1) (Times (f0 fid1) (trm.Var vid3)) else Const 0) (\ y. if vid3 = y then t else fst (\ y. if vid2 = y then r else fst (a, b) $ y, b) $ y, b)) = (\ i. dterm_sem I (if i = vid1 then trm.Var vid1 else Const 0) (mk_v I (OSing vid1 ($f fid1 (\i. Const 0))) (a, b) (\ i. if i = vid1 then a $ i + Functions I fid1 (\ _. 0) * t else a $ i)))" using vne12 vne13 mk_v_agree[of "I" "(OSing vid1 ($f fid1 (\i. Const 0)))" "(a, b)" "(\ i. if i = vid1 then a $ i + Functions I fid1 (\ _. 0) * t else a $ i)"] by (auto simp add: vec_eq_iff f0_def empty_def Vagree_def) show "0 \ t \ t \ r \ Predicates I vid2 (\ i. dterm_sem I (if i = vid1 then Plus (trm.Var vid1) (Times (f0 fid1) (trm.Var vid3)) else Const 0) (\ y. if vid3 = y then t else fst (\ y. if vid2 = y then r else fst (a, b) $ y, b) $ y, b)) \ Predicates I vid2 (\ i. dterm_sem I (if i = vid1 then trm.Var vid1 else Const 0) (mk_v I (OSing vid1 ($f fid1 (\i. Const 0))) (a, b) (\ i. if i = vid1 then a $ i + Functions I fid1 (\ _. 0) * t else a $ i)))" using eq by auto qed done have req4':"?sol 0 = fst (a,b)" by (auto simp: vec_eq_iff) then have req4: " (?sol 0) = (fst (a,b))" using VSagree_refl[of a] req4' unfolding VSagree_def by auto have inPred:"?abba \ fml_sem I (p1 vid3 vid1)" using req1 leq req3 req4 thisW by fastforce have sem_eq:"?abba \ fml_sem I (p1 vid3 vid1) \ (aa,ba) \ fml_sem I (p1 vid3 vid1)" apply (rule coincidence_formula) apply (auto simp add: aaba Vagree_def p1_def f0_def empty_def) subgoal using Iagree_refl by auto done from inPred sem_eq have inPred':"(aa,ba) \ fml_sem I (p1 vid3 vid1)" by auto \ \thus by lemma 6 consequence for formulas\ show "repv (repv (a, b) vid2 r) vid1 (dterm_sem I (Plus (trm.Var vid1) (Times (f0 fid1) (trm.Var vid2))) (repv (a, b) vid2 r)) \ fml_sem I (p1 vid3 vid1)" using aaba inPred' by (auto) next fix I::"('sf,'sc,'sz) interp" and aa ba ab bb sol and t:: real assume good_interp:"is_interp I" assume all:" \r. dterm_sem I (Const 0) (repv (ab, bb) vid2 r) \ dterm_sem I (trm.Var vid2) (repv (ab, bb) vid2 r) \ (\ra. repv (repv (ab, bb) vid2 r) vid3 ra \ {v. dterm_sem I (Const 0) v \ dterm_sem I (trm.Var vid3) v} \ {v. dterm_sem I (trm.Var vid3) v \ dterm_sem I (trm.Var vid2) v} \ Predicates I vid2 (\ i. dterm_sem I (singleton (Plus (trm.Var vid1) (Times (f0 fid1) (trm.Var vid3))) i) (repv (repv (ab, bb) vid2 r) vid3 ra))) \ (\\. \ = repv (repv (ab, bb) vid2 r) vid1 (dterm_sem I (Plus (trm.Var vid1) (Times (f0 fid1) (trm.Var vid2))) (repv (ab, bb) vid2 r)) \ \ \ fml_sem I (p1 vid3 vid1))" assume t:"0 \ t" assume aaba:"(aa, ba) = mk_v I (OSing vid1 (f0 fid1)) (ab, bb) (sol t)" assume sol:"(sol solves_ode (\_. ODE_sem I (OSing vid1 (f0 fid1)))) {0..t} {x. mk_v I (OSing vid1 (f0 fid1)) (ab, bb) x \ fml_sem I (p1 vid2 vid1)}" hence constraint:"\s. s \ {0 .. t} \ sol s \ {x. mk_v I (OSing vid1 (f0 fid1)) (ab, bb) x \ fml_sem I (p1 vid2 vid1)}" using solves_ode_domainD by fastforce \ \\sol 0 = fst (ab, bb)\\ assume sol0:" (sol 0) = (fst (ab, bb)) " have impl:"dterm_sem I (Const 0) (repv (ab, bb) vid2 t) \ dterm_sem I (trm.Var vid2) (repv (ab, bb) vid2 t) \ (\ra. repv (repv (ab, bb) vid2 t) vid3 ra \ {v. dterm_sem I (Const 0) v \ dterm_sem I (trm.Var vid3) v} \ {v. dterm_sem I (trm.Var vid3) v \ dterm_sem I (trm.Var vid2) v} \ Predicates I vid2 (\ i. dterm_sem I (singleton (Plus (trm.Var vid1) (Times (f0 fid1) (trm.Var vid3))) i) (repv (repv (ab, bb) vid2 t) vid3 ra))) \ (\\. \ = repv (repv (ab, bb) vid2 t) vid1 (dterm_sem I (Plus (trm.Var vid1) (Times (f0 fid1) (trm.Var vid2))) (repv (ab, bb) vid2 t)) \ \ \ fml_sem I (p1 vid3 vid1))" using all by auto interpret ll:ll_on_open_it UNIV "(\_. ODE_sem I (OSing vid1 (f0 fid1)))" "UNIV" 0 apply(standard) apply(auto) unfolding local_lipschitz_def f0_def empty_def sterm_sem.simps using gt_ex lipschitz_on_constant by blast have eq_UNIV:"ll.existence_ivl 0 (sol 0) = UNIV" apply(rule ll.existence_ivl_eq_domain) apply(auto) subgoal for tm tM t apply(unfold f0_def empty_def sterm_sem.simps) by(metis add.right_neutral mult_zero_left order_refl) done \ \Combine with \flow_usolves_ode\ and \equals_flowI\ to get uniqueness of solution\ let ?f = "(\_. ODE_sem I (OSing vid1 (f0 fid1)))" have sol_UNIV: "\t x. (ll.flow 0 x usolves_ode ?f from 0) (ll.existence_ivl 0 x) UNIV" using ll.flow_usolves_ode by auto from sol have sol': "(sol solves_ode (\_. ODE_sem I (OSing vid1 (f0 fid1)))) {0..t} UNIV" apply (rule solves_ode_supset_range) by auto from sol' have sol'':"\s. s \ 0 \ s \ t \ (sol solves_ode (\_. ODE_sem I (OSing vid1 (f0 fid1)))) {0..s} UNIV" by (simp add: solves_ode_subset) have sol0_eq:"sol 0 = ll.flow 0 (sol 0) 0" using ll.general.flow_initial_time_if by auto have isFlow:"\s. s \ 0 \ s \ t \ sol s = ll.flow 0 (sol 0) s" apply(rule ll.equals_flowI) apply(auto) subgoal using eq_UNIV by auto subgoal using sol'' closed_segment_eq_real_ivl t by (auto simp add: solves_ode_singleton) subgoal using eq_UNIV sol sol0_eq by auto done let ?c = "Functions I fid1 (\ _. 0)" let ?sol = "(\t. \ i. if i = vid1 then (ab $ i) + ?c * t else (ab $ i))" have vec_simp:"(\a b. \ i. if i = vid1 then sterm_sem I ($f fid1 (\i. Const 0)) b else 0) = (\a b. \ i. if i = vid1 then Functions I fid1 (\ i. 0) else 0)" by (auto simp add: vec_eq_iff cong: if_cong) have exp_sol:"(?sol solves_ode (\_. ODE_sem I (OSing vid1 (f0 fid1)))) {0..t} UNIV" apply(auto simp add: f0_def empty_def vec_simp) apply(rule solves_odeI) apply(auto simp only: has_vderiv_on_def has_vector_derivative_def box_sem) apply (rule has_derivative_vec[THEN has_derivative_eq_rhs]) defer apply (rule ext) apply (subst scaleR_vec_def) apply (rule refl) apply (auto intro!: derivative_eq_intros) done from exp_sol have exp_sol':"\s. s \ 0 \ s \ t \ (?sol solves_ode (\_. ODE_sem I (OSing vid1 (f0 fid1)))) {0..s} UNIV" by (simp add: solves_ode_subset) have exp_sol0_eq:"?sol 0 = ll.flow 0 (?sol 0) 0" using ll.general.flow_initial_time_if by auto have more_eq:"(\ i. if i = vid1 then ab $ i + Functions I fid1 (\ _. 0) * 0 else ab $ i) = sol 0" using sol0 apply auto apply(rule vec_extensionality) by(auto) have exp_isFlow:"\s. s \ 0 \ s \ t \ ?sol s = ll.flow 0 (sol 0) s" apply(rule ll.equals_flowI) apply(auto) subgoal using eq_UNIV by auto defer subgoal for s using eq_UNIV apply auto subgoal using exp_sol exp_sol0_eq more_eq apply(auto) done done using exp_sol' closed_segment_eq_real_ivl t apply(auto) by (simp add: solves_ode_singleton) have sol_eq_exp:"\s. s \ 0 \ s \ t \ ?sol s = sol s" unfolding exp_isFlow isFlow by auto then have sol_eq_exp_t:"?sol t = sol t" using t by auto then have sol_eq_exp_t':"sol t $ vid1 = ?sol t $ vid1" by auto then have useful:"?sol t $ vid1 = ab $ vid1 + Functions I fid1 (\ i. 0) * t" by auto from sol_eq_exp_t' useful have useful':"sol t $ vid1 = ab $ vid1 + Functions I fid1 (\ i. 0) * t" by auto have sol_int:"((ll.flow 0 (sol 0)) usolves_ode ?f from 0) {0..t} {x. mk_v I (OSing vid1 (f0 fid1)) (ab, bb) x \ fml_sem I (p1 vid2 vid1)}" apply (rule usolves_ode_subset_range[of "(ll.flow 0 (sol 0))" "?f" "0" "{0..t}" "UNIV" "{x. mk_v I (OSing vid1 (f0 fid1)) (ab, bb) x \ fml_sem I (p1 vid2 vid1)}"]) subgoal using eq_UNIV sol_UNIV[of "(sol 0)"] apply (auto) apply (rule usolves_ode_subset) using t by(auto) apply(auto) using sol apply(auto dest!: solves_ode_domainD) subgoal for xa using isFlow[of xa] by(auto) done have thing:"\s. 0 \ s \ s \ t \ fst (mk_v I (OSing vid1 ($f fid1 (\i. Const 0))) (ab, bb) (?sol s)) $ vid1 = ab $ vid1 + Functions I fid1 (\ i. 0) * s" subgoal for s using mk_v_agree[of I "(OSing vid1 ($f fid1 (\i. Const 0)))" "(ab, bb)" "(?sol s)"] apply auto unfolding Vagree_def by auto done have thing':"\s. 0 \ s \ s \ t \ fst (mk_v I (OSing vid1 ($f fid1 (\i. Const 0))) (ab, bb) (sol s)) $ vid1 = ab $ vid1 + Functions I fid1 (\ i. 0) * s" subgoal for s using thing[of s] sol_eq_exp[of s] by auto done have another_eq:"\i s. 0 \ s \ s \ t \ dterm_sem I (if i = vid1 then trm.Var vid1 else Const 0) (mk_v I (OSing vid1 (f0 fid1)) (ab, bb) (sol s)) = dterm_sem I (if i = vid1 then Plus (trm.Var vid1) (Times (f0 fid1) (trm.Var vid3)) else Const 0) (\ y. if vid3 = y then s else fst (\ y. if vid2 = y then s else fst (ab, bb) $ y, bb) $ y, bb)" using mk_v_agree[of "I" "(OSing vid1 (f0 fid1))" "(ab, bb)" "(sol s)"] vne12 vne23 vne13 apply(auto simp add: f0_def p1_def empty_def) unfolding Vagree_def apply(simp add: f0_def empty_def) subgoal for s using thing' by auto done have allRa':"(\ra. repv (repv (ab, bb) vid2 t) vid3 ra \ {v. dterm_sem I (Const 0) v \ dterm_sem I (trm.Var vid3) v} \ {v. dterm_sem I (trm.Var vid3) v \ dterm_sem I (trm.Var vid2) v} \ Predicates I vid2 (\ i. dterm_sem I (if i = vid1 then trm.Var vid1 else Const 0) (mk_v I (OSing vid1 (f0 fid1)) (ab, bb) (sol ra))))" apply(rule allI) subgoal for ra using mk_v_agree[of "I" "(OSing vid1 (f0 fid1))" "(ab, bb)" "(sol ra)"] vne23 constraint[of ra] apply(auto simp add: Vagree_def p1_def) done done have anotherFact:"\ra. 0 \ ra \ ra \ t \ (\ i. if i = vid1 then ab $ i + Functions I fid1 (\ _. 0) * ra else ab $ i) $ vid1 = ab $ vid1 + dterm_sem I (f0 fid1) (\ y. if vid3 = y then ra else fst (\ y. if vid2 = y then t else fst (ab, bb) $ y, bb) $ y, bb) * ra " subgoal for ra apply simp apply(rule disjI2) by (auto simp add: f0_def empty_def) done have thing':"\ra i. 0 \ ra \ ra \ t \ dterm_sem I (if i = vid1 then trm.Var vid1 else Const 0) (mk_v I (OSing vid1 ($f fid1 (\i. Const 0))) (ab, bb) (sol ra)) = dterm_sem I (if i = vid1 then Plus (trm.Var vid1) (Times (f0 fid1) (trm.Var vid3)) else Const 0) (\ y. if vid3 = y then ra else fst (\ y. if vid2 = y then t else fst (ab, bb) $ y, bb) $ y, bb) " subgoal for ra i using vne12 vne13 mk_v_agree[of I "OSing vid1 ($f fid1 (\i. Const 0))" "(ab,bb)" "(sol ra)"] apply (auto) unfolding Vagree_def apply(safe) apply(erule allE[where x="vid1"])+ using sol_eq_exp[of ra] anotherFact[of ra] by auto done have allRa:"(\ra. repv (repv (ab, bb) vid2 t) vid3 ra \ {v. dterm_sem I (Const 0) v \ dterm_sem I (trm.Var vid3) v} \ {v. dterm_sem I (trm.Var vid3) v \ dterm_sem I (trm.Var vid2) v} \ Predicates I vid2 (\ i. dterm_sem I (singleton (Plus (trm.Var vid1) (Times (f0 fid1) (trm.Var vid3))) i) (repv (repv (ab, bb) vid2 t) vid3 ra)))" apply(rule allI) subgoal for ra using mk_v_agree[of "I" "(OSing vid1 (f0 fid1))" "(ab, bb)" "(sol ra)"] vne23 constraint[of ra] apply(auto simp add: Vagree_def p1_def) using sol_eq_exp[of ra] apply (auto simp add: f0_def empty_def Vagree_def vec_eq_iff) using thing' by auto done have fml3:"\ra. 0 \ ra \ ra \ t \ (\\. \ = repv (repv (ab, bb) vid2 t) vid1 (dterm_sem I (Plus (trm.Var vid1) (Times (f0 fid1) (trm.Var vid2))) (repv (ab, bb) vid2 t)) \ \ \ fml_sem I (p1 vid3 vid1))" using impl allRa by auto have someEq:"(\ i. dterm_sem I (if i = vid1 then trm.Var vid1 else Const 0) (\ y. if vid1 = y then (if vid2 = vid1 then t else fst (ab, bb) $ vid1) + Functions I fid1 (\ i. 0) * t else fst (\ y. if vid2 = y then t else fst (ab, bb) $ y, bb) $ y, bb)) = (\ i. dterm_sem I (if i = vid1 then trm.Var vid1 else Const 0) (mk_v I (OSing vid1 ($f fid1 (\i. Const 0))) (ab, bb) (sol t)))" apply(rule vec_extensionality) using vne12 sol_eq_exp t thing by auto show "mk_v I (OSing vid1 (f0 fid1)) (ab, bb) (sol t) \ fml_sem I (p1 vid3 vid1)" using mk_v_agree[of I "OSing vid1 (f0 fid1)" "(ab, bb)" "sol t"] fml3[of t] unfolding f0_def p1_def empty_def Vagree_def using someEq by(auto simp add: sol_eq_exp_t' t vec_extensionality vne12) qed qed lemma MVT0_within: fixes f ::"real \ real" and f'::"real \ real \ real" and s t :: real assumes f':"\x. x \ {0..t} \ (f has_derivative (f' x)) (at x within {0..t})" assumes geq':"\x. x \ {0..t} \ f' x s \ 0" assumes int_s:"s > 0 \ s \ t" assumes t: "0 < t" shows "f s \ f 0" proof - have "f 0 + 0 \ f s" apply (rule Lib.MVT_ivl'[OF f', of 0 s 0]) subgoal for x by assumption subgoal for x using geq' by auto using t int_s t apply auto subgoal for x by (metis int_s mult.commute mult.right_neutral order.trans real_mult_le_cancel_iff2) done then show "?thesis" by auto qed lemma MVT': fixes f g ::"real \ real" fixes f' g'::"real \ real \ real" fixes s t ::real assumes f':"\s. s \ {0..t} \ (f has_derivative (f' s)) (at s within {0..t})" assumes g':"\s. s \ {0..t} \ (g has_derivative (g' s)) (at s within {0..t})" assumes geq':"\x. x \ {0..t} \ f' x s \ g' x s" assumes geq0:"f 0 \ g 0" assumes int_s:"s > 0 \ s \ t" assumes t:"t > 0" shows "f s \ g s" proof - let ?h = "(\x. f x - g x)" let ?h' = "(\s x. f' s x - g' s x)" have "?h s \ ?h 0" apply(rule MVT0_within[of t ?h "?h'" s]) subgoal for s using f'[of s] g'[of s] by auto subgoal for sa using geq'[of sa] by auto subgoal using int_s by auto subgoal using t by auto done then show "?thesis" using geq0 by auto qed lemma MVT'_gr: fixes f g ::"real \ real" fixes f' g'::"real \ real \ real" fixes s t ::real assumes f':"\s. s \ {0..t} \ (f has_derivative (f' s)) (at s within {0..t})" assumes g':"\s. s \ {0..t} \ (g has_derivative (g' s)) (at s within {0..t})" assumes geq':"\x. x \ {0..t} \ f' x s \ g' x s" assumes geq0:"f 0 > g 0" assumes int_s:"s > 0 \ s \ t" assumes t:"t > 0" shows "f s > g s" proof - let ?h = "(\x. f x - g x)" let ?h' = "(\s x. f' s x - g' s x)" have "?h s \ ?h 0" apply(rule MVT0_within[of t ?h "?h'" s]) subgoal for s using f'[of s] g'[of s] by auto subgoal for sa using geq'[of sa] by auto subgoal using int_s by auto subgoal using t by auto done then show "?thesis" using geq0 by auto qed lemma frech_linear: fixes x \ \ \' I assumes good_interp:"is_interp I" assumes free:"dfree \" shows "x * frechet I \ \ \' = frechet I \ \ (x *\<^sub>R \')" using frechet_linear[OF good_interp free] by (simp add: linear_simps) lemma rift_in_space_time: fixes sol I ODE \ \ t s b assumes good_interp:"is_interp I" assumes free:"dfree \" assumes osafe:"osafe ODE" assumes sol:"(sol solves_ode (\_ \'. ODE_sem I ODE \')) {0..t} {x. mk_v I ODE (sol 0, b) x \ fml_sem I \}" assumes FVT:"FVT \ \ semBV I ODE" assumes ivl:"s \ {0..t}" shows "((\t. sterm_sem I \ (fst (mk_v I ODE (sol 0, b) (sol t)))) \ \This is Frechet derivative, so equivalent to:\ \ \\has_real_derivative frechet I \ (fst((mk_v I ODE (sol 0, b) (sol s)))) (snd (mk_v I ODE (sol 0, b) (sol s))))) (at s within {0..t})\\ has_derivative (\t'. t' * frechet I \ (fst((mk_v I ODE (sol 0, b) (sol s)))) (snd (mk_v I ODE (sol 0, b) (sol s))))) (at s within {0..t})" proof - let ?\ = "(\t. (mk_v I ODE (sol 0, b) (sol t)))" let ?\s = "(\t. fst (?\ t))" have sol_deriv:"\s. s \ {0..t} \ (sol has_derivative (\xa. xa *\<^sub>R ODE_sem I ODE (sol s))) (at s within {0..t})" using sol apply simp apply (drule solves_odeD(1)) unfolding has_vderiv_on_def has_vector_derivative_def by auto have sol_dom:"\s. s\ {0..t} \ ?\ s \ fml_sem I \" using sol apply simp apply (drule solves_odeD(2)) by auto let ?h = "(\t. sterm_sem I \ (?\s t))" let ?g = "(\\. sterm_sem I \ \)" let ?f = "?\s" let ?f' = "(\t'. t' *\<^sub>R (\ i. if i \ ODE_vars I ODE then ODE_sem I ODE (sol s) $ i else 0))" let ?g' = "(frechet I \ (?\s s))" have heq:"?h = ?g \ ?f" by (auto) have fact1:"\i. i \ ODE_vars I ODE \ (\t. ?\s(t) $ i) = (\t. sol t $ i)" subgoal for i apply(rule ext) subgoal for t using mk_v_agree[of I ODE "(sol 0, b)" "sol t"] unfolding Vagree_def by auto done done have fact2:"\i. i \ ODE_vars I ODE \ (\t. if i \ ODE_vars I ODE then ODE_sem I ODE (sol t) $ i else 0) = (\t. ODE_sem I ODE (sol t) $ i)" subgoal for i apply(rule ext) subgoal for t using mk_v_agree[of I ODE "(sol 0, b)" "sol t"] unfolding Vagree_def by auto done done have fact3:"\i. i \ (-ODE_vars I ODE) \ (\t. ?\s(t) $ i) = (\t. sol 0 $ i)" subgoal for i apply(rule ext) subgoal for t using mk_v_agree[of I ODE "(sol 0, b)" "sol t"] unfolding Vagree_def by auto done done have fact4:"\i. i \ (-ODE_vars I ODE) \ (\t. if i \ ODE_vars I ODE then ODE_sem I ODE (sol t) $ i else 0) = (\t. 0)" subgoal for i apply(rule ext) subgoal for t using mk_v_agree[of I ODE "(sol 0, b)" "sol t"] unfolding Vagree_def by auto done done have some_eq:"(\v'. \ i. v' *\<^sub>R ODE_sem I ODE (sol s) $ i) = (\v'. v' *\<^sub>R ODE_sem I ODE (sol s))" apply(rule ext) apply(rule vec_extensionality) by auto have some_sol:"(sol has_derivative (\v'. v' *\<^sub>R ODE_sem I ODE (sol s))) (at s within {0..t})" using sol ivl unfolding solves_ode_def has_vderiv_on_def has_vector_derivative_def by auto have some_eta:"(\t. \ i. sol t $ i) = sol" by (rule ext, rule vec_extensionality, auto) have ode_deriv:"\i. i \ ODE_vars I ODE \ ((\t. sol t $ i) has_derivative (\ v'. v' *\<^sub>R ODE_sem I ODE (sol s) $ i)) (at s within {0..t})" subgoal for i apply(rule has_derivative_proj) using some_eq some_sol some_eta by auto done have eta:"(\t. (\ i. ?f t $ i)) = ?f" by(rule ext, rule vec_extensionality, auto) have eta_esque:"(\t'. \ i. t' * (if i \ ODE_vars I ODE then ODE_sem I ODE (sol s) $ i else 0)) = (\t'. t' *\<^sub>R (\ i. if i \ ODE_vars I ODE then ODE_sem I ODE (sol s) $ i else 0))" apply(rule ext | rule vec_extensionality)+ subgoal for t' i by auto done have "((\t. (\ i. ?f t $ i)) has_derivative (\t'. (\ i. ?f' t' $ i))) (at s within {0..t})" apply (rule has_derivative_vec) subgoal for i apply(cases "i \ ODE_vars I ODE") subgoal using fact1[of i] fact2[of i] ode_deriv[of i] by auto subgoal using fact3[of i] fact4[of i] by auto done done then have fderiv:"(?f has_derivative ?f') (at s within {0..t})" using eta eta_esque by auto have gderiv:"(?g has_derivative ?g') (at (?f s) within ?f ` {0..t})" using has_derivative_at_withinI using frechet_correctness free good_interp by blast have chain:"((?g \ ?f) has_derivative (?g' \ ?f')) (at s within {0..t})" using fderiv gderiv diff_chain_within by blast let ?co\1 = "(fst (mk_v I ODE (sol 0, b) (sol s)), ODE_sem I ODE (fst (mk_v I ODE (sol 0, b) (sol s))))" let ?co\2 = "(fst (mk_v I ODE (sol 0, b) (sol s)), snd (mk_v I ODE (sol 0, b) (sol s)))" have sub_cont:"\a .a \ ODE_vars I ODE \ Inl a \ FVT \ \ False" using FVT by auto have sub_cont2:"\a .a \ ODE_vars I ODE \ Inr a \ FVT \ \ False" using FVT by auto have "Vagree (mk_v I ODE (sol 0, b) (sol s)) (sol s, b) (Inl ` ODE_vars I ODE)" using mk_v_agree[of I ODE "(sol 0, b)" "sol s"] unfolding Vagree_def by auto let ?co'\1 = "(\x. (fst (mk_v I ODE (sol 0, b) (sol s)), x *\<^sub>R (\ i. if i \ ODE_vars I ODE then ODE_sem I ODE (sol s) $ i else 0)))" let ?co'\2 = "(\x. (fst (mk_v I ODE (sol 0, b) (sol s)), x *\<^sub>R snd (mk_v I ODE (sol 0, b) (sol s))))" have co_agree_sem:"\s. Vagree (?co'\1 s) (?co'\2 s) (semBV I ODE)" subgoal for sa using mk_v_agree[of I ODE "(sol 0, b)" "sol s"] unfolding Vagree_def by auto done have co_agree_help:"\s. Vagree (?co'\1 s) (?co'\2 s) (FVT \)" using agree_sub[OF FVT co_agree_sem] by auto have co_agree':"\s. Vagree (?co'\1 s) (?co'\2 s) (FVDiff \)" subgoal for s using mk_v_agree[of I ODE "(sol 0, b)" "sol s"] unfolding Vagree_def apply auto subgoal for i x apply(cases x) subgoal for a apply(cases "a \ ODE_vars I ODE") by (simp | metis (no_types, lifting) FVT ODE_vars_lr Vagree_def mk_v_agree mk_xode.elims subsetD snd_conv)+ subgoal for a apply(cases "a \ ODE_vars I ODE") by (simp | metis (no_types, lifting) FVT Vagree_def mk_v_agree mk_xode.elims subsetD snd_conv)+ done subgoal for i x apply(cases x) subgoal for a apply(cases "a \ ODE_vars I ODE") using FVT ODE_vars_lr Vagree_def mk_v_agree mk_xode.elims subsetD snd_conv by auto subgoal for a apply(cases "a \ ODE_vars I ODE") apply(erule allE[where x=i])+ using FVT ODE_vars_lr Vagree_def mk_v_agree mk_xode.elims subsetD snd_conv by auto done done done have heq'':"(?g' \ ?f') = (\t'. t' *\<^sub>R frechet I \ (?\s s) (snd (?\ s)))" using mk_v_agree[of I ODE "(sol 0, b)" "sol s"] unfolding comp_def apply auto apply(rule ext | rule vec_extensionality)+ subgoal for x using frech_linear[of I \ x "(fst (mk_v I ODE (sol 0, b) (sol s)))" "(snd (mk_v I ODE (sol 0, b) (sol s)))", OF good_interp free] using coincidence_frechet[OF free, of "(?co'\1 x)" "(?co'\2 x)", OF co_agree'[of x], of I] by auto done have "((?g \ ?f) has_derivative (?g' \ ?f')) (at s within {0..t})" using chain by auto then have "((?g \ ?f) has_derivative (\t'. t' * frechet I \ (?\s s) (snd (?\ s)))) (at s within {0..t})" using heq'' by auto then have result:"((\t. sterm_sem I \ (?\s t)) has_derivative (\t. t * frechet I \ (?\s s) (snd (?\ s)))) (at s within {0..t})" using heq by auto then show "?thesis" by auto qed lemma dterm_sterm_dfree: "dfree \ \ (\\ \'. sterm_sem I \ \ = dterm_sem I \ (\, \'))" by(induction rule: dfree.induct, auto) \ \\g(x)\ h(x) \ [x'=f(x), c & p(x)](g(x)' \ h(x)') \ [x'=f(x), c]g(x) \ h(x)\\ lemma DIGeq_valid:"valid DIGeqaxiom" unfolding DIGeqaxiom_def apply(unfold DIGeqaxiom_def valid_def impl_sem iff_sem) apply(auto) (* 4 goals*) proof - fix I::"('sf,'sc,'sz) interp" and b aa ba and sol::"real \ 'sz simple_state" and t::real let ?ODE = "OVar vid1" let ?\ = "(\t. mk_v I (?ODE) (sol 0, b) (sol t))" assume t:"0 \ t" and sol:"(sol solves_ode (\a. ODEs I vid1)) {0..t} {x. Predicates I vid1 (\ i. dterm_sem I (empty i) (mk_v I ?ODE (sol 0, b) x))}" and notin:" \(Predicates I vid1 (\ i. dterm_sem I (empty i) (sol 0, b)))" have fsafe:"fsafe (Prop vid1 empty)" by (auto simp add: empty_def) from sol have "Predicates I vid1 (\ i. dterm_sem I (empty i) (?\ 0))" using t solves_ode_domainD[of sol "(\a. ODEs I vid1)" "{0..t}"] by auto then have incon:"Predicates I vid1 (\ i. dterm_sem I (empty i) ((sol 0, b)))" using coincidence_formula[OF fsafe Iagree_refl[of I], of "(sol 0, b)" "?\ 0"] unfolding Vagree_def by (auto simp add: empty_def) show "dterm_sem I (f1 fid2 vid1) (mk_v I (OVar vid1) (sol 0, b) (sol t)) \ dterm_sem I (f1 fid1 vid1) (mk_v I (OVar vid1) (sol 0, b) (sol t))" using notin incon by auto next fix I::"('sf,'sc,'sz) interp" and b aa ba and sol::"real \ 'sz simple_state" and t::real let ?ODE = "OVar vid1" let ?\ = "(\t. mk_v I (?ODE) (sol 0, b) (sol t))" assume t:"0 \ t" and sol:"(sol solves_ode (\a. ODEs I vid1)) {0..t} {x. Predicates I vid1 (\ i. dterm_sem I (empty i) (mk_v I ?ODE (sol 0, b) x))}" and notin:" \(Predicates I vid1 (\ i. dterm_sem I (empty i) (sol 0, b)))" have fsafe:"fsafe (Prop vid1 empty)" by (auto simp add: empty_def) from sol have "Predicates I vid1 (\ i. dterm_sem I (empty i) (?\ 0))" using t solves_ode_domainD[of sol "(\a. ODEs I vid1)" "{0..t}"] by auto then have incon:"Predicates I vid1 (\ i. dterm_sem I (empty i) ((sol 0, b)))" using coincidence_formula[OF fsafe Iagree_refl[of I], of "(sol 0, b)" "?\ 0"] unfolding Vagree_def by (auto simp add: empty_def) show "dterm_sem I (f1 fid2 vid1) (mk_v I (OVar vid1) (sol 0, b) (sol t)) \ dterm_sem I (f1 fid1 vid1) (mk_v I (OVar vid1) (sol 0, b) (sol t))" using notin incon by auto next fix I::"('sf,'sc,'sz) interp" and b aa ba and sol::"real \ 'sz simple_state" and t::real let ?ODE = "OVar vid1" let ?\ = "(\t. mk_v I (?ODE) (sol 0, b) (sol t))" assume t:"0 \ t" assume sol:"(sol solves_ode (\a. ODEs I vid1)) {0..t} {x. Predicates I vid1 (\ i. dterm_sem I (local.empty i) (mk_v I (OVar vid1) (sol 0, b) x))}" assume notin:"\ Predicates I vid1 (\ i. dterm_sem I (local.empty i) (sol 0, b))" have fsafe:"fsafe (Prop vid1 empty)" by (auto simp add: empty_def) from sol have "Predicates I vid1 (\ i. dterm_sem I (empty i) (?\ 0))" using t solves_ode_domainD[of sol "(\a. ODEs I vid1)" "{0..t}"] by auto then have incon:"Predicates I vid1 (\ i. dterm_sem I (empty i) ((sol 0, b)))" using coincidence_formula[OF fsafe Iagree_refl[of I], of "(sol 0, b)" "?\ 0"] unfolding Vagree_def by (auto simp add: empty_def) show "dterm_sem I (f1 fid2 vid1) (mk_v I (OVar vid1) (sol 0, b) (sol t)) \ dterm_sem I (f1 fid1 vid1) (mk_v I (OVar vid1) (sol 0, b) (sol t))" using incon notin by auto next fix I::"('sf,'sc,'sz) interp" and b aa ba and sol::"real \ 'sz simple_state" and t::real let ?ODE = "OVar vid1" let ?\ = "(\t. mk_v I (?ODE) (sol 0, b) (sol t))" assume good_interp:"is_interp I" assume aaba:"(aa, ba) = mk_v I (OVar vid1) (sol 0, b) (sol t)" assume t:"0 \ t" assume sol:"(sol solves_ode (\a. ODEs I vid1)) {0..t} {x. Predicates I vid1 (\ i. dterm_sem I (local.empty i) (mk_v I (OVar vid1) (sol 0, b) x))}" assume box:"\a ba. (\sola. sol 0 = sola 0 \ (\t. (a, ba) = mk_v I (OVar vid1) (sola 0, b) (sola t) \ 0 \ t \ (sola solves_ode (\a. ODEs I vid1)) {0..t} {x. Predicates I vid1 (\ i. dterm_sem I (local.empty i) (mk_v I (OVar vid1) (sola 0, b) x))})) \ directional_derivative I (f1 fid2 vid1) (a, ba) \ directional_derivative I (f1 fid1 vid1) (a, ba)" assume geq0:"dterm_sem I (f1 fid2 vid1) (sol 0, b) \ dterm_sem I (f1 fid1 vid1) (sol 0, b)" have free1:"dfree ($f fid2 (\i. if i = vid1 then trm.Var vid1 else Const 0))" by (auto intro: dfree.intros) have free2:"dfree ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0))" by (auto intro: dfree.intros) from geq0 have geq0':"sterm_sem I (f1 fid2 vid1) (sol 0) \ sterm_sem I (f1 fid1 vid1) (sol 0)" unfolding f1_def using dterm_sterm_dfree[OF free1, of I "sol 0" b] dterm_sterm_dfree[OF free2, of I "sol 0" b] by auto let ?\s = "\x. fst (?\ x)" let ?\t = "\x. snd (?\ x)" let ?df1 = "(\t. dterm_sem I (f1 fid2 vid1) (?\ t))" let ?f1 = "(\t. sterm_sem I (f1 fid2 vid1) (?\s t))" let ?f1' = "(\ s t'. t' * frechet I (f1 fid2 vid1) (?\s s) (?\t s))" have dfeq:"?df1 = ?f1" apply(rule ext) subgoal for t using dterm_sterm_dfree[OF free1, of I "?\s t" "snd (?\ t)"] unfolding f1_def expand_singleton by auto done have free3:"dfree (f1 fid2 vid1)" unfolding f1_def by (auto intro: dfree.intros) let ?df2 = "(\t. dterm_sem I (f1 fid1 vid1) (?\ t))" let ?f2 = "(\t. sterm_sem I (f1 fid1 vid1) (?\s t))" let ?f2' = "(\s t' . t' * frechet I (f1 fid1 vid1) (?\s s) (?\t s))" let ?int = "{0..t}" have bluh:"\x i. (Functions I i has_derivative (THE f'. \x. (Functions I i has_derivative f' x) (at x)) x) (at x)" using good_interp unfolding is_interp_def by auto have blah:"(Functions I fid2 has_derivative (THE f'. \x. (Functions I fid2 has_derivative f' x) (at x)) (\ i. if i = vid1 then sol t $ vid1 else 0)) (at (\ i. if i = vid1 then sol t $ vid1 else 0))" using bluh by auto have bigEx:"\s. s \ {0..t} \(\sola. sol 0 = sola 0 \ (\ta. (fst (?\ s), snd (?\ s)) = mk_v I (OVar vid1) (sola 0, b) (sola ta) \ 0 \ ta \ (sola solves_ode (\a. ODEs I vid1)) {0..ta} {x. Predicates I vid1 (\ i. dterm_sem I (local.empty i) (mk_v I (OVar vid1) (sol 0, b) x))}))" subgoal for s apply(rule exI[where x=sol]) apply(rule conjI) subgoal by (rule refl) apply(rule exI[where x=s]) apply(rule conjI) subgoal by auto apply(rule conjI) subgoal by auto using sol using atLeastAtMost_iff atLeastatMost_subset_iff order_refl solves_ode_on_subset by (metis (no_types, lifting) subsetI) done have box':"\s. s \ {0..t} \ directional_derivative I (f1 fid2 vid1) (?\s s, ?\t s) \ directional_derivative I (f1 fid1 vid1) (?\s s, ?\t s)" subgoal for s using box apply simp apply (erule allE[where x="?\s s"]) apply (erule allE[where x="?\t s"]) using bigEx[of s] by auto done have dsafe1:"dsafe (f1 fid2 vid1)" unfolding f1_def by (auto intro: dsafe.intros) have dsafe2:"dsafe (f1 fid1 vid1)" unfolding f1_def by (auto intro: dsafe.intros) have agree1:"Vagree (sol 0, b) (?\ 0) (FVT (f1 fid2 vid1))" using mk_v_agree[of I "(OVar vid1)" "(sol 0, b)" "fst (?\ 0)"] unfolding f1_def Vagree_def expand_singleton apply auto by (metis (no_types, lifting) Compl_iff Vagree_def fst_conv mk_v_agree mk_xode.simps semBV.simps) have agree2:"Vagree (sol 0, b) (?\ 0) (FVT (f1 fid1 vid1))" using mk_v_agree[of I "(OVar vid1)" "(sol 0, b)" "fst (?\ 0)"] unfolding f1_def Vagree_def expand_singleton apply auto by (metis (no_types, lifting) Compl_iff Vagree_def fst_conv mk_v_agree mk_xode.simps semBV.simps) have sem_eq1:"dterm_sem I (f1 fid2 vid1) (sol 0, b) = dterm_sem I (f1 fid2 vid1) (?\ 0)" using coincidence_dterm[OF dsafe1 agree1] by auto then have sem_eq1':"sterm_sem I (f1 fid2 vid1) (sol 0) = sterm_sem I (f1 fid2 vid1) (?\s 0)" using dterm_sterm_dfree[OF free1, of I "sol 0" "b"] dterm_sterm_dfree[OF free1, of I "(?\s 0)" "snd (?\ 0)"] unfolding f1_def expand_singleton by auto have sem_eq2:"dterm_sem I (f1 fid1 vid1) (sol 0, b) = dterm_sem I (f1 fid1 vid1) (?\ 0)" using coincidence_dterm[OF dsafe2 agree2] by auto then have sem_eq2':"sterm_sem I (f1 fid1 vid1) (sol 0) = sterm_sem I (f1 fid1 vid1) (?\s 0)" using dterm_sterm_dfree[OF free2, of I "sol 0" "b"] dterm_sterm_dfree[OF free2, of I "(?\s 0)" "snd (?\ 0)"] unfolding f1_def expand_singleton by auto have good_interp':"\i x. (Functions I i has_derivative (THE f'. \x. (Functions I i has_derivative f' x) (at x)) x) (at x)" using good_interp unfolding is_interp_def by auto have chain : "\f f' g g' x s. (f has_derivative f') (at x within s) \ (g has_derivative g') (at (f x) within f ` s) \ (g \ f has_derivative g' \ f') (at x within s)" by(auto intro: derivative_intros) have sol1:"(sol solves_ode (\_. ODE_sem I (OVar vid1))) {0..t} {x. mk_v I (OVar vid1) (sol 0, b) x \ fml_sem I (Prop vid1 empty)}" using sol unfolding p1_def singleton_def empty_def by auto have FVTsub1:"vid1 \ ODE_vars I (OVar vid1) \ FVT ($f fid2 (\i. if i = vid1 then trm.Var vid1 else Const 0)) \ semBV I ((OVar vid1))" apply auto subgoal for x xa apply(cases "xa = vid1") by auto done have FVTsub2:"vid1 \ ODE_vars I (OVar vid1) \ FVT ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0)) \ semBV I ((OVar vid1))" apply auto subgoal for x xa apply(cases "xa = vid1") by auto done have osafe:"osafe (OVar vid1)" by auto have deriv1:"\s. vid1 \ ODE_vars I (OVar vid1) \ s \ ?int \ (?f1 has_derivative (?f1' s)) (at s within {0..t})" subgoal for s using rift_in_space_time[OF good_interp free1 osafe sol1 FVTsub1, of s] unfolding f1_def expand_singleton directional_derivative_def by blast done have deriv2:"\s. vid1 \ ODE_vars I (OVar vid1) \ s \ ?int \ (?f2 has_derivative (?f2' s)) (at s within {0..t})" subgoal for s using rift_in_space_time[OF good_interp free2 osafe sol1 FVTsub2, of s] unfolding f1_def expand_singleton directional_derivative_def by blast done have leq:"\s . s \ ?int \ ?f1' s 1 \ ?f2' s 1" subgoal for s using box'[of s] by (simp add: directional_derivative_def) done have preserve_agree1:"vid1 \ ODE_vars I (OVar vid1) \ VSagree (sol 0) (fst (mk_v I (OVar vid1) (sol 0, b) (sol t))) {vid1}" using mk_v_agree[of I "OVar vid1" "(sol 0, b)" "sol t"] unfolding Vagree_def VSagree_def by auto have preserve_coincide1: "vid1 \ ODE_vars I (OVar vid1) \ sterm_sem I (f1 fid2 vid1) (fst (mk_v I (OVar vid1) (sol 0, b) (sol t))) = sterm_sem I (f1 fid2 vid1) (sol 0)" using coincidence_sterm[of "(sol 0, b)" "(mk_v I (OVar vid1) (sol 0, b) (sol t))" "f1 fid2 vid1" I] preserve_agree1 unfolding VSagree_def Vagree_def f1_def by auto have preserve_coincide2: "vid1 \ ODE_vars I (OVar vid1) \ sterm_sem I (f1 fid1 vid1) (fst (mk_v I (OVar vid1) (sol 0, b) (sol t))) = sterm_sem I (f1 fid1 vid1) (sol 0)" using coincidence_sterm[of "(sol 0, b)" "(mk_v I (OVar vid1) (sol 0, b) (sol t))" "f1 fid1 vid1" I] preserve_agree1 unfolding VSagree_def Vagree_def f1_def by auto have "?f1 t \ ?f2 t" apply(cases "t = 0") subgoal using geq0' sem_eq1' sem_eq2' by auto subgoal apply(cases "vid1 \ ODE_vars I (OVar vid1)") subgoal apply (rule MVT'[OF deriv2 deriv1, of t]) (* 8 subgoals *) subgoal by auto subgoal by auto subgoal for s using deriv2[of s] using leq by auto using t leq geq0' sem_eq1' sem_eq2' by auto subgoal using geq0 using dterm_sterm_dfree[OF free1, of I "sol 0" "b"] using dterm_sterm_dfree[OF free2, of I "sol 0" "b"] using preserve_coincide1 preserve_coincide2 by(simp add: f1_def) done done then show " dterm_sem I (f1 fid2 vid1) (mk_v I (OVar vid1) (sol 0, b) (sol t)) \ dterm_sem I (f1 fid1 vid1) (mk_v I (OVar vid1) (sol 0, b) (sol t)) " using t dterm_sterm_dfree[OF free2, of I "?\s t" "snd (?\ t)"] dterm_sterm_dfree[OF free1, of I "?\s t" "snd (?\ t)"] by (simp add: f1_def) qed lemma DIGr_valid:"valid DIGraxiom" unfolding DIGraxiom_def apply(unfold DIGraxiom_def valid_def impl_sem iff_sem) apply(auto) (* 4 subgoals*) proof - fix I::"('sf,'sc,'sz) interp" and b aa ba and sol::"real \ 'sz simple_state" and t::real let ?ODE = "OVar vid1" let ?\ = "(\t. mk_v I (?ODE) (sol 0, b) (sol t))" assume t:"0 \ t" and sol:"(sol solves_ode (\a. ODEs I vid1)) {0..t} {x. Predicates I vid1 (\ i. dterm_sem I (empty i) (mk_v I ?ODE (sol 0, b) x))}" and notin:" \(Predicates I vid1 (\ i. dterm_sem I (empty i) (sol 0, b)))" have fsafe:"fsafe (Prop vid1 empty)" by (auto simp add: empty_def) from sol have "Predicates I vid1 (\ i. dterm_sem I (empty i) (?\ 0))" using t solves_ode_domainD[of sol "(\a. ODEs I vid1)" "{0..t}"] by auto then have incon:"Predicates I vid1 (\ i. dterm_sem I (empty i) ((sol 0, b)))" using coincidence_formula[OF fsafe Iagree_refl[of I], of "(sol 0, b)" "?\ 0"] unfolding Vagree_def by (auto simp add: empty_def) show "dterm_sem I (f1 fid2 vid1) (mk_v I (OVar vid1) (sol 0, b) (sol t)) < dterm_sem I (f1 fid1 vid1) (mk_v I (OVar vid1) (sol 0, b) (sol t))" using notin incon by auto next fix I::"('sf,'sc,'sz) interp" and b aa ba and sol::"real \ 'sz simple_state" and t::real let ?ODE = "OVar vid1" let ?\ = "(\t. mk_v I (?ODE) (sol 0, b) (sol t))" assume t:"0 \ t" and sol:"(sol solves_ode (\a. ODEs I vid1)) {0..t} {x. Predicates I vid1 (\ i. dterm_sem I (empty i) (mk_v I ?ODE (sol 0, b) x))}" and notin:" \(Predicates I vid1 (\ i. dterm_sem I (empty i) (sol 0, b)))" have fsafe:"fsafe (Prop vid1 empty)" by (auto simp add: empty_def) from sol have "Predicates I vid1 (\ i. dterm_sem I (empty i) (?\ 0))" using t solves_ode_domainD[of sol "(\a. ODEs I vid1)" "{0..t}"] by auto then have incon:"Predicates I vid1 (\ i. dterm_sem I (empty i) ((sol 0, b)))" using coincidence_formula[OF fsafe Iagree_refl[of I], of "(sol 0, b)" "?\ 0"] unfolding Vagree_def by (auto simp add: empty_def) show "dterm_sem I (f1 fid2 vid1) (mk_v I (OVar vid1) (sol 0, b) (sol t)) < dterm_sem I (f1 fid1 vid1) (mk_v I (OVar vid1) (sol 0, b) (sol t))" using notin incon by auto next fix I::"('sf,'sc,'sz) interp" and b aa ba and sol::"real \ 'sz simple_state" and t::real let ?ODE = "OVar vid1" let ?\ = "(\t. mk_v I (?ODE) (sol 0, b) (sol t))" assume t:"0 \ t" assume sol:"(sol solves_ode (\a. ODEs I vid1)) {0..t} {x. Predicates I vid1 (\ i. dterm_sem I (local.empty i) (mk_v I (OVar vid1) (sol 0, b) x))}" assume notin:"\ Predicates I vid1 (\ i. dterm_sem I (local.empty i) (sol 0, b))" have fsafe:"fsafe (Prop vid1 empty)" by (auto simp add: empty_def) from sol have "Predicates I vid1 (\ i. dterm_sem I (empty i) (?\ 0))" using t solves_ode_domainD[of sol "(\a. ODEs I vid1)" "{0..t}"] by auto then have incon:"Predicates I vid1 (\ i. dterm_sem I (empty i) ((sol 0, b)))" using coincidence_formula[OF fsafe Iagree_refl[of I], of "(sol 0, b)" "?\ 0"] unfolding Vagree_def by (auto simp add: empty_def) show "dterm_sem I (f1 fid2 vid1) (mk_v I (OVar vid1) (sol 0, b) (sol t)) < dterm_sem I (f1 fid1 vid1) (mk_v I (OVar vid1) (sol 0, b) (sol t))" using incon notin by auto next fix I::"('sf,'sc,'sz) interp" and b aa ba and sol::"real \ 'sz simple_state" and t::real let ?ODE = "OVar vid1" let ?\ = "(\t. mk_v I (?ODE) (sol 0, b) (sol t))" assume good_interp:"is_interp I" assume aaba:"(aa, ba) = mk_v I (OVar vid1) (sol 0, b) (sol t)" assume t:"0 \ t" assume sol:"(sol solves_ode (\a. ODEs I vid1)) {0..t} {x. Predicates I vid1 (\ i. dterm_sem I (local.empty i) (mk_v I (OVar vid1) (sol 0, b) x))}" assume box:"\a ba. (\sola. sol 0 = sola 0 \ (\t. (a, ba) = mk_v I (OVar vid1) (sola 0, b) (sola t) \ 0 \ t \ (sola solves_ode (\a. ODEs I vid1)) {0..t} {x. Predicates I vid1 (\ i. dterm_sem I (local.empty i) (mk_v I (OVar vid1) (sola 0, b) x))})) \ directional_derivative I (f1 fid2 vid1) (a, ba) \ directional_derivative I (f1 fid1 vid1) (a, ba)" assume geq0:"dterm_sem I (f1 fid2 vid1) (sol 0, b) < dterm_sem I (f1 fid1 vid1) (sol 0, b)" have free1:"dfree ($f fid2 (\i. if i = vid1 then trm.Var vid1 else Const 0))" by (auto intro: dfree.intros) have free2:"dfree ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0))" by (auto intro: dfree.intros) from geq0 have geq0':"sterm_sem I (f1 fid2 vid1) (sol 0) < sterm_sem I (f1 fid1 vid1) (sol 0)" unfolding f1_def using dterm_sterm_dfree[OF free1, of I "sol 0" b] dterm_sterm_dfree[OF free2, of I "sol 0" b] by auto let ?\s = "\x. fst (?\ x)" let ?\t = "\x. snd (?\ x)" let ?df1 = "(\t. dterm_sem I (f1 fid2 vid1) (?\ t))" let ?f1 = "(\t. sterm_sem I (f1 fid2 vid1) (?\s t))" let ?f1' = "(\ s t'. t' * frechet I (f1 fid2 vid1) (?\s s) (?\t s))" have dfeq:"?df1 = ?f1" apply(rule ext) subgoal for t using dterm_sterm_dfree[OF free1, of I "?\s t" "snd (?\ t)"] unfolding f1_def expand_singleton by auto done have free3:"dfree (f1 fid2 vid1)" unfolding f1_def by (auto intro: dfree.intros) let ?df2 = "(\t. dterm_sem I (f1 fid1 vid1) (?\ t))" let ?f2 = "(\t. sterm_sem I (f1 fid1 vid1) (?\s t))" let ?f2' = "(\s t' . t' * frechet I (f1 fid1 vid1) (?\s s) (?\t s))" let ?int = "{0..t}" have bluh:"\x i. (Functions I i has_derivative (THE f'. \x. (Functions I i has_derivative f' x) (at x)) x) (at x)" using good_interp unfolding is_interp_def by auto have blah:"(Functions I fid2 has_derivative (THE f'. \x. (Functions I fid2 has_derivative f' x) (at x)) (\ i. if i = vid1 then sol t $ vid1 else 0)) (at (\ i. if i = vid1 then sol t $ vid1 else 0))" using bluh by auto have bigEx:"\s. s \ {0..t} \(\sola. sol 0 = sola 0 \ (\ta. (fst (?\ s), snd (?\ s)) = mk_v I (OVar vid1) (sola 0, b) (sola ta) \ 0 \ ta \ (sola solves_ode (\a. ODEs I vid1)) {0..ta} {x. Predicates I vid1 (\ i. dterm_sem I (local.empty i) (mk_v I (OVar vid1) (sol 0, b) x))}))" subgoal for s apply(rule exI[where x=sol]) apply(rule conjI) subgoal by (rule refl) apply(rule exI[where x=s]) apply(rule conjI) subgoal by auto apply(rule conjI) subgoal by auto using sol using atLeastAtMost_iff atLeastatMost_subset_iff order_refl solves_ode_on_subset by (metis (no_types, lifting) subsetI) done have box':"\s. s \ {0..t} \ directional_derivative I (f1 fid2 vid1) (?\s s, ?\t s) \ directional_derivative I (f1 fid1 vid1) (?\s s, ?\t s)" subgoal for s using box apply simp apply (erule allE[where x="?\s s"]) apply (erule allE[where x="?\t s"]) using bigEx[of s] by auto done have dsafe1:"dsafe (f1 fid2 vid1)" unfolding f1_def by (auto intro: dsafe.intros) have dsafe2:"dsafe (f1 fid1 vid1)" unfolding f1_def by (auto intro: dsafe.intros) have agree1:"Vagree (sol 0, b) (?\ 0) (FVT (f1 fid2 vid1))" using mk_v_agree[of I "(OVar vid1)" "(sol 0, b)" "fst (?\ 0)"] unfolding f1_def Vagree_def expand_singleton apply auto by (metis (no_types, lifting) Compl_iff Vagree_def fst_conv mk_v_agree mk_xode.simps semBV.simps) have agree2:"Vagree (sol 0, b) (?\ 0) (FVT (f1 fid1 vid1))" using mk_v_agree[of I "(OVar vid1)" "(sol 0, b)" "fst (?\ 0)"] unfolding f1_def Vagree_def expand_singleton apply auto by (metis (no_types, lifting) Compl_iff Vagree_def fst_conv mk_v_agree mk_xode.simps semBV.simps) have sem_eq1:"dterm_sem I (f1 fid2 vid1) (sol 0, b) = dterm_sem I (f1 fid2 vid1) (?\ 0)" using coincidence_dterm[OF dsafe1 agree1] by auto then have sem_eq1':"sterm_sem I (f1 fid2 vid1) (sol 0) = sterm_sem I (f1 fid2 vid1) (?\s 0)" using dterm_sterm_dfree[OF free1, of I "sol 0" "b"] dterm_sterm_dfree[OF free1, of I "(?\s 0)" "snd (?\ 0)"] unfolding f1_def expand_singleton by auto have sem_eq2:"dterm_sem I (f1 fid1 vid1) (sol 0, b) = dterm_sem I (f1 fid1 vid1) (?\ 0)" using coincidence_dterm[OF dsafe2 agree2] by auto then have sem_eq2':"sterm_sem I (f1 fid1 vid1) (sol 0) = sterm_sem I (f1 fid1 vid1) (?\s 0)" using dterm_sterm_dfree[OF free2, of I "sol 0" "b"] dterm_sterm_dfree[OF free2, of I "(?\s 0)" "snd (?\ 0)"] unfolding f1_def expand_singleton by auto have good_interp':"\i x. (Functions I i has_derivative (THE f'. \x. (Functions I i has_derivative f' x) (at x)) x) (at x)" using good_interp unfolding is_interp_def by auto have chain : "\f f' g g' x s. (f has_derivative f') (at x within s) \ (g has_derivative g') (at (f x) within f ` s) \ (g \ f has_derivative g' \ f') (at x within s)" by(auto intro: derivative_intros) have sol1:"(sol solves_ode (\_. ODE_sem I (OVar vid1))) {0..t} {x. mk_v I (OVar vid1) (sol 0, b) x \ fml_sem I (Prop vid1 empty)}" using sol unfolding p1_def singleton_def empty_def by auto have FVTsub1:"vid1 \ ODE_vars I (OVar vid1) \ FVT ($f fid2 (\i. if i = vid1 then trm.Var vid1 else Const 0)) \ semBV I ((OVar vid1))" apply auto subgoal for x xa apply(cases "xa = vid1") by auto done have FVTsub2:"vid1 \ ODE_vars I (OVar vid1) \ FVT ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0)) \ semBV I ((OVar vid1))" apply auto subgoal for x xa apply(cases "xa = vid1") by auto done have osafe:"osafe (OVar vid1)" by auto have deriv1:"\s. vid1 \ ODE_vars I (OVar vid1) \ s \ ?int \ (?f1 has_derivative (?f1' s)) (at s within {0..t})" subgoal for s using rift_in_space_time[OF good_interp free1 osafe sol1 FVTsub1, of s] unfolding f1_def expand_singleton directional_derivative_def by blast done have deriv2:"\s. vid1 \ ODE_vars I (OVar vid1) \ s \ ?int \ (?f2 has_derivative (?f2' s)) (at s within {0..t})" subgoal for s using rift_in_space_time[OF good_interp free2 osafe sol1 FVTsub2, of s] unfolding f1_def expand_singleton directional_derivative_def by blast done have leq:"\s . s \ ?int \ ?f1' s 1 \ ?f2' s 1" subgoal for s using box'[of s] by (simp add: directional_derivative_def) done have preserve_agree1:"vid1 \ ODE_vars I (OVar vid1) \ VSagree (sol 0) (fst (mk_v I (OVar vid1) (sol 0, b) (sol t))) {vid1}" using mk_v_agree[of I "OVar vid1" "(sol 0, b)" "sol t"] unfolding Vagree_def VSagree_def by auto have preserve_coincide1: "vid1 \ ODE_vars I (OVar vid1) \ sterm_sem I (f1 fid2 vid1) (fst (mk_v I (OVar vid1) (sol 0, b) (sol t))) = sterm_sem I (f1 fid2 vid1) (sol 0)" using coincidence_sterm[of "(sol 0, b)" "(mk_v I (OVar vid1) (sol 0, b) (sol t))" "f1 fid2 vid1" I] preserve_agree1 unfolding VSagree_def Vagree_def f1_def by auto have preserve_coincide2: "vid1 \ ODE_vars I (OVar vid1) \ sterm_sem I (f1 fid1 vid1) (fst (mk_v I (OVar vid1) (sol 0, b) (sol t))) = sterm_sem I (f1 fid1 vid1) (sol 0)" using coincidence_sterm[of "(sol 0, b)" "(mk_v I (OVar vid1) (sol 0, b) (sol t))" "f1 fid1 vid1" I] preserve_agree1 unfolding VSagree_def Vagree_def f1_def by auto have "?f1 t < ?f2 t" apply(cases "t = 0") subgoal using geq0' sem_eq1' sem_eq2' by auto subgoal apply(cases "vid1 \ ODE_vars I (OVar vid1)") subgoal apply (rule MVT'_gr[OF deriv2 deriv1, of t]) subgoal by auto subgoal by auto subgoal for s using deriv2[of s] using leq by auto using t leq geq0' sem_eq1' sem_eq2' by auto subgoal using geq0 using dterm_sterm_dfree[OF free1, of I "sol 0" "b"] using dterm_sterm_dfree[OF free2, of I "sol 0" "b"] using preserve_coincide1 preserve_coincide2 by(simp add: f1_def) done done then show "dterm_sem I (f1 fid2 vid1) (mk_v I (OVar vid1) (sol 0, b) (sol t)) < dterm_sem I (f1 fid1 vid1) (mk_v I (OVar vid1) (sol 0, b) (sol t))" using t dterm_sterm_dfree[OF free2, of I "?\s t" "snd (?\ t)"] dterm_sterm_dfree[OF free1, of I "?\s t" "snd (?\ t)"] using geq0 f1_def by (simp add: f1_def) qed lemma DG_valid:"valid DGaxiom" proof - have osafe:"osafe (OSing vid1 (f1 fid1 vid1))" by(auto simp add: osafe_Sing dfree_Fun dfree_Const f1_def expand_singleton) have fsafe:"fsafe (p1 vid1 vid1)" by(auto simp add: p1_def dfree_Const) have osafe2:"osafe (OProd (OSing vid1 (f1 fid1 vid1)) (OSing vid2 (Plus (Times (f1 fid2 vid1) (trm.Var vid2)) (f1 fid3 vid1))))" by(auto simp add: f1_def expand_singleton osafe.intros dfree.intros vne12) note sem = ode_alt_sem[OF osafe fsafe] note sem2 = ode_alt_sem[OF osafe2 fsafe] have p2safe:"fsafe (p1 vid2 vid1)" by(auto simp add: p1_def dfree_Const) show "valid DGaxiom" apply(auto simp del: prog_sem.simps(8) simp add: DGaxiom_def valid_def sem sem2) apply(rule exI[where x=0], auto simp add: f1_def p1_def expand_singleton) subgoal for I a b aa ba sol t proof - assume good_interp:"is_interp I" assume " \aa ba. (\sol t. (aa, ba) = mk_v I (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0))) (a, b) (sol t) \ 0 \ t \ (sol solves_ode (\a b. \ i. if i = vid1 then sterm_sem I (f1 fid1 vid1) b else 0)) {0..t} {x. Predicates I vid1 (\ i. dterm_sem I (if i = vid1 then trm.Var vid1 else Const 0) (mk_v I (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0))) (a, b) x))} \ VSagree (sol 0) a {uu. uu = vid1 \ Inl uu \ Inl ` {x. \xa. Inl x \ FVT (if xa = vid1 then trm.Var vid1 else Const 0)} \ (\x. Inl uu \ FVT (if x = vid1 then trm.Var vid1 else Const 0))}) \ Predicates I vid2 (\ i. dterm_sem I (if i = vid1 then trm.Var vid1 else Const 0) (aa, ba))" then have bigAll:" \aa ba. (\sol t. (aa, ba) = mk_v I (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0))) (a, b) (sol t) \ 0 \ t \ (sol solves_ode (\a b. \ i. if i = vid1 then sterm_sem I (f1 fid1 vid1) b else 0)) {0..t} {x. Predicates I vid1 (\ i. dterm_sem I (if i = vid1 then trm.Var vid1 else Const 0) (mk_v I (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0))) (a, b) x))} \ VSagree (sol 0) a {uu. uu = vid1 \ (\x. Inl uu \ FVT (if x = vid1 then trm.Var vid1 else Const 0))}) \ Predicates I vid2 (\ i. dterm_sem I (if i = vid1 then trm.Var vid1 else Const 0) (aa, ba))" by (auto) assume aaba:"(aa, ba) = mk_v I (OProd (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0))) (OSing vid2 (Plus (Times ($f fid2 (\i. if i = vid1 then trm.Var vid1 else Const 0)) (trm.Var vid2)) ($f fid3 (\i. if i = vid1 then trm.Var vid1 else Const 0))))) (\ y. if vid2 = y then 0 else fst (a, b) $ y, b) (sol t)" assume t:"0 \ t" assume sol:" (sol solves_ode (\a b. (\ i. if i = vid1 then sterm_sem I (f1 fid1 vid1) b else 0) + (\ i. if i = vid2 then sterm_sem I (Plus (Times (f1 fid2 vid1) (trm.Var vid2)) (f1 fid3 vid1)) b else 0))) {0..t} {x. Predicates I vid1 (\ i. dterm_sem I (if i = vid1 then trm.Var vid1 else Const 0) (mk_v I (OProd (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0))) (OSing vid2 (Plus (Times ($f fid2 (\i. if i = vid1 then trm.Var vid1 else Const 0)) (trm.Var vid2)) ($f fid3 (\i. if i = vid1 then trm.Var vid1 else Const 0))))) (\ y. if vid2 = y then 0 else fst (a, b) $ y, b) x))}" assume VSag:"VSagree (sol 0) (\ y. if vid2 = y then 0 else fst (a, b) $ y) {x. x = vid2 \ x = vid1 \ x = vid2 \ x = vid1 \ Inl x \ Inl ` {x. x = vid2 \ x = vid1} \ x = vid1}" let ?sol = "(\t. \ i. if i = vid1 then sol t $ vid1 else 0)" let ?aaba' = "mk_v I (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0))) (a, b) (?sol t)" from bigAll[of "fst ?aaba'" "snd ?aaba'"] have bigEx:"(\sol t. ?aaba' = mk_v I (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0))) (a, b) (sol t) \ 0 \ t \ (sol solves_ode (\a b. \ i. if i = vid1 then sterm_sem I (f1 fid1 vid1) b else 0)) {0..t} {x. Predicates I vid1 (\ i. dterm_sem I (if i = vid1 then trm.Var vid1 else Const 0) (mk_v I (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0))) (a, b) x))} \ VSagree (sol 0) a {uu. uu = vid1 \ (\x. Inl uu \ FVT (if x = vid1 then trm.Var vid1 else Const 0))}) \ Predicates I vid2 (\ i. dterm_sem I (if i = vid1 then trm.Var vid1 else Const 0) (?aaba'))" by simp have pre1:"?aaba' = mk_v I (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0))) (a, b) (?sol t)" by (rule refl) have agreeL:"\s. fst (mk_v I (OProd (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0))) (OSing vid2 (Plus (Times ($f fid2 (\i. if i = vid1 then trm.Var vid1 else Const 0)) (trm.Var vid2)) ($f fid3 (\i. if i = vid1 then trm.Var vid1 else Const 0))))) (\ y. if vid2 = y then 0 else fst (a, b) $ y, b) (sol s)) $ vid1 = sol s $ vid1" subgoal for s using mk_v_agree[of I "(OProd (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0))) (OSing vid2 (Plus (Times ($f fid2 (\i. if i = vid1 then trm.Var vid1 else Const 0)) (trm.Var vid2)) ($f fid3 (\i. if i = vid1 then trm.Var vid1 else Const 0)))))" "(\ y. if vid2 = y then 0 else fst (a, b) $ y, b)" "(sol s)"] unfolding Vagree_def by auto done have agreeR:"\s. fst (mk_v I (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0))) (a, b) (\ i. if i = vid1 then sol s $ vid1 else 0)) $ vid1 = sol s $ vid1" subgoal for s using mk_v_agree[of "I" "(OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0)))" "(a, b)" "(\ i. if i = vid1 then sol s $ vid1 else 0)"] unfolding Vagree_def by auto done have FV:"(FVF (p1 vid1 vid1)) = {Inl vid1}" unfolding p1_def expand_singleton apply auto subgoal for x xa apply(cases "xa = vid1") by auto done have agree:"\s. Vagree (mk_v I (OProd (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0))) (OSing vid2 (Plus (Times ($f fid2 (\i. if i = vid1 then trm.Var vid1 else Const 0)) (trm.Var vid2)) ($f fid3 (\i. if i = vid1 then trm.Var vid1 else Const 0))))) (\ y. if vid2 = y then 0 else fst (a, b) $ y, b) (sol s)) (mk_v I (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0))) (a, b) (\ i. if i = vid1 then sol s $ vid1 else 0)) (FVF (p1 vid1 vid1))" using agreeR agreeL unfolding Vagree_def FV by auto note con_sem_eq = coincidence_formula[OF fsafe Iagree_refl agree] have constraint:"\s. 0 \ s \ s \ t \ Predicates I vid1 (\ i. dterm_sem I (if i = vid1 then trm.Var vid1 else Const 0) (mk_v I (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0))) (a, b) (\ i. if i = vid1 then sol s $ vid1 else 0)))" using sol apply simp apply(drule solves_odeD(2)) apply auto[1] subgoal for s using con_sem_eq by (auto simp add: p1_def expand_singleton) done have eta:"sol = (\t. \ i. sol t $ i)" by (rule ext, rule vec_extensionality, simp) have yet_another_eq:"\x. (\xa. xa *\<^sub>R ((\ i. if i = vid1 then sterm_sem I (f1 fid1 vid1) (sol x) else 0) + (\ i. if i = vid2 then sterm_sem I (Plus (Times (f1 fid2 vid1) (trm.Var vid2)) (f1 fid3 vid1)) (sol x) else 0))) = (\xa. (\ i. (xa *\<^sub>R ((\ i. if i = vid1 then sterm_sem I (f1 fid1 vid1) (sol x) else 0) + (\ i. if i = vid2 then sterm_sem I (Plus (Times (f1 fid2 vid1) (trm.Var vid2)) (f1 fid3 vid1)) (sol x) else 0))) $ i))" subgoal for x by (rule ext, rule vec_extensionality, simp) done have sol_deriv:"\x. x \{0..t} \ (sol has_derivative (\xa. xa *\<^sub>R ((\ i. if i = vid1 then sterm_sem I (f1 fid1 vid1) (sol x) else 0) + (\ i. if i = vid2 then sterm_sem I (Plus (Times (f1 fid2 vid1) (trm.Var vid2)) (f1 fid3 vid1)) (sol x) else 0)))) (at x within {0..t})" using sol apply simp apply(drule solves_odeD(1)) unfolding has_vderiv_on_def has_vector_derivative_def by auto then have sol_deriv:"\x. x \ {0..t} \ ((\t. \ i. sol t $ i) has_derivative (\xa. (\ i. (xa *\<^sub>R ((\ i. if i = vid1 then sterm_sem I (f1 fid1 vid1) (sol x) else 0) + (\ i. if i = vid2 then sterm_sem I (Plus (Times (f1 fid2 vid1) (trm.Var vid2)) (f1 fid3 vid1)) (sol x) else 0))) $ i))) (at x within {0..t})" using yet_another_eq eta by auto have sol_deriv1: "\x. x \ {0..t} \ ((\t. sol t $ vid1) has_derivative (\xa. (xa *\<^sub>R ((\ i. if i = vid1 then sterm_sem I (f1 fid1 vid1) (sol x) else 0) + (\ i. if i = vid2 then sterm_sem I (Plus (Times (f1 fid2 vid1) (trm.Var vid2)) (f1 fid3 vid1)) (sol x) else 0)) $ vid1))) (at x within {0..t})" subgoal for s (* I heard higher-order unification is hard.*) apply(rule has_derivative_proj[of "(\ i t. sol t $ i)" "(\j xa. (xa *\<^sub>R ((\ i. if i = vid1 then sterm_sem I (f1 fid1 vid1) (sol s) else 0) + (\ i. if i = vid2 then sterm_sem I (Plus (Times (f1 fid2 vid1) (trm.Var vid2)) (f1 fid3 vid1)) (sol s) else 0)) $ j))" "at s within {0..t}""vid1"]) using sol_deriv[of s] by auto done have hmm:"\s. (\ i. sterm_sem I (if i = vid1 then trm.Var vid1 else Const 0) (sol s)) = (\ i. sterm_sem I (if i = vid1 then trm.Var vid1 else Const 0) (\ i. if i = vid1 then sol s $ vid1 else 0))" by(rule vec_extensionality, auto) have aha:"\s. (\xa. xa * sterm_sem I (f1 fid1 vid1) (sol s)) = (\xa. xa * sterm_sem I (f1 fid1 vid1) (\ i. if i = vid1 then sol s $ vid1 else 0))" subgoal for s apply(rule ext) subgoal for xa using hmm by (auto simp add: f1_def) done done let ?sol' = "(\s. (\xa. \ i. if i = vid1 then xa * sterm_sem I (f1 fid1 vid1) (\ i. if i = vid1 then sol s $ vid1 else 0) else 0))" let ?project_me_plz = "(\t. (\ i. if i = vid1 then ?sol t $ vid1 else 0))" have sol_deriv_eq:"\s. s \{0..t} \ ((\t. (\ i. if i = vid1 then ?sol t $ vid1 else 0)) has_derivative ?sol' s) (at s within {0..t})" subgoal for s apply(rule has_derivative_vec) subgoal for i apply (cases "i = vid1", cases "i = vid2", auto) using vne12 apply simp using sol_deriv1[of s] using aha by auto done done have yup:"(\t. (\ i. if i = vid1 then ?sol t $ vid1 else 0) $ vid1) = (\t. sol t $ vid1)" by(rule ext, auto) have maybe:"\s. (\xa. xa * sterm_sem I (f1 fid1 vid1) (\ i. if i = vid1 then sol s $ vid1 else 0)) = (\xa. (\ i. if i = vid1 then xa * sterm_sem I (f1 fid1 vid1) (\ i. if i = vid1 then sol s $ vid1 else 0) else 0) $ vid1) " by(rule ext, auto) have almost:"(\x. if vid1 = vid1 then (\ i. if i = vid1 then sol x $ vid1 else 0) $ vid1 else 0) = (\x. (\ i. if i = vid1 then sol x $ vid1 else 0) $ vid1)" by(rule ext, auto) have almost':"\s. (\h. if vid1 = vid1 then h * sterm_sem I (f1 fid1 vid1) (\ i. if i = vid1 then sol s $ vid1 else 0) else 0) = (\h. h * sterm_sem I (f1 fid1 vid1) (\ i. if i = vid1 then sol s $ vid1 else 0))" by(rule ext, auto) have deriv':" \x. x \ {0..t} \ ((\t. \ i. if i = vid1 then sol t $ vid1 else 0) has_derivative (\xa. (\ i. xa *\<^sub>R (if i = vid1 then sterm_sem I (f1 fid1 vid1) (\ i. if i = vid1 then sol x $ vid1 else 0) else 0)))) (at x within {0..t})" subgoal for s apply(rule has_derivative_vec) subgoal for i apply(cases "i = vid1") prefer 2 subgoal by auto apply auto using has_derivative_proj[OF sol_deriv_eq[of s], of vid1] using yup maybe[of s] almost almost'[of s] by fastforce done done have derEq:"\s. (\xa. (\ i. xa *\<^sub>R (if i = vid1 then sterm_sem I (f1 fid1 vid1) (\ i. if i = vid1 then sol s $ vid1 else 0) else 0))) = (\xa. xa *\<^sub>R (\ i. if i = vid1 then sterm_sem I (f1 fid1 vid1) (\ i. if i = vid1 then sol s $ vid1 else 0) else 0))" subgoal for s apply (rule ext, rule vec_extensionality) by auto done have "\x. x \ {0..t} \ ((\t. \ i. if i = vid1 then sol t $ vid1 else 0) has_derivative (\xa. xa *\<^sub>R (\ i. if i = vid1 then sterm_sem I (f1 fid1 vid1) (\ i. if i = vid1 then sol x $ vid1 else 0) else 0))) (at x within {0..t})" subgoal for s using deriv'[of s] derEq[of s] by auto done then have deriv:"((\t. \ i. if i = vid1 then sol t $ vid1 else 0) has_vderiv_on (\t. \ i. if i = vid1 then sterm_sem I (f1 fid1 vid1) (\ i. if i = vid1 then sol t $ vid1 else 0) else 0)) {0..t}" unfolding has_vderiv_on_def has_vector_derivative_def by auto have pre2:"(?sol solves_ode (\a b. \ i. if i = vid1 then sterm_sem I (f1 fid1 vid1) b else 0)) {0..t} {x. Predicates I vid1 (\ i. dterm_sem I (if i = vid1 then trm.Var vid1 else Const 0) (mk_v I (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0))) (a, b) x))}" apply(rule solves_odeI) subgoal by (rule deriv) subgoal for s using constraint by auto done have pre3:"VSagree (?sol 0) a {u. u = vid1 \ (\x. Inl u \ FVT (if x = vid1 then trm.Var vid1 else Const 0))}" using vne12 VSag unfolding VSagree_def by simp have bigPre:"(\sol t. ?aaba' = mk_v I (OSing vid1 ($f fid1 (\i. if i = vid1 then Var vid1 else Const 0))) (a, b) (sol t) \ 0 \ t \ (sol solves_ode (\a b. \ i. if i = vid1 then sterm_sem I (f1 fid1 vid1) b else 0)) {0..t} {x. Predicates I vid1 (\ i. dterm_sem I (if i = vid1 then trm.Var vid1 else Const 0) (mk_v I (OSing vid1 ($f fid1 (\i. if i = vid1 then Var vid1 else Const 0))) (a, b) x))} \ VSagree (sol 0) a {u. u = vid1 \ (\x. Inl u \ FVT (if x = vid1 then Var vid1 else Const 0))})" apply(rule exI[where x="?sol"]) apply(rule exI[where x=t]) apply(rule conjI) apply(rule pre1) apply(rule conjI) apply(rule t) apply(rule conjI) apply(rule pre2) by(rule pre3) have pred2:"Predicates I vid2 (\ i. dterm_sem I (if i = vid1 then trm.Var vid1 else Const 0) ?aaba')" using bigEx bigPre by auto then have pred2':"?aaba' \ fml_sem I (p1 vid2 vid1)" unfolding p1_def expand_singleton by auto let ?res_state = "(mk_v I (OProd (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0))) (OSing vid2 (Plus (Times ($f fid2 (\i. if i = vid1 then trm.Var vid1 else Const 0)) (trm.Var vid2)) ($f fid3 (\i. if i = vid1 then trm.Var vid1 else Const 0))))) (\ y. if vid2 = y then 0 else fst (a, b) $ y, b) (sol t))" have aabaX:"(fst ?aaba') $ vid1 = sol t $ vid1" using aaba mk_v_agree[of "I" "(OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0)))" "(a, b)" "(?sol t)"] proof - assume " Vagree (mk_v I (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0))) (a, b) (\ i. if i = vid1 then sol t $ vid1 else 0)) (a, b) (- semBV I (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0)))) \ Vagree (mk_v I (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0))) (a, b) (\ i. if i = vid1 then sol t $ vid1 else 0)) (mk_xode I (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0))) (\ i. if i = vid1 then sol t $ vid1 else 0)) (semBV I (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0))))" then have ag:" Vagree (mk_v I (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0))) (a, b) (?sol t)) (mk_xode I (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0))) (?sol t)) (semBV I (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0))))" by auto have sembv:"(semBV I (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0)))) = {Inl vid1, Inr vid1}" by auto have sub:"{Inl vid1} \ {Inl vid1, Inr vid1}" by auto have ag':"Vagree (mk_v I (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0))) (a, b) (?sol t)) (mk_xode I (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0))) (?sol t)) {Inl vid1}" using ag agree_sub[OF sub] sembv by auto then have eq1:"fst (mk_v I (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0))) (a, b) (?sol t)) $ vid1 = fst (mk_xode I (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0))) (?sol t)) $ vid1" unfolding Vagree_def by auto moreover have "... = sol t $ vid1" by auto ultimately show ?thesis by auto qed have res_stateX:"(fst ?res_state) $ vid1 = sol t $ vid1" using mk_v_agree[of I "(OProd (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0))) (OSing vid2 (Plus (Times ($f fid2 (\i. if i = vid1 then trm.Var vid1 else Const 0)) (trm.Var vid2)) ($f fid3 (\i. if i = vid1 then trm.Var vid1 else Const 0)))))" "(\ y. if vid2 = y then 0 else fst (a, b) $ y, b)" "(sol t)"] proof - assume "Vagree (mk_v I (OProd (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0))) (OSing vid2 (Plus (Times ($f fid2 (\i. if i = vid1 then trm.Var vid1 else Const 0)) (trm.Var vid2)) ($f fid3 (\i. if i = vid1 then trm.Var vid1 else Const 0))))) (\ y. if vid2 = y then 0 else fst (a, b) $ y, b) (sol t)) (\ y. if vid2 = y then 0 else fst (a, b) $ y, b) (- semBV I (OProd (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0))) (OSing vid2 (Plus (Times ($f fid2 (\i. if i = vid1 then trm.Var vid1 else Const 0)) (trm.Var vid2)) ($f fid3 (\i. if i = vid1 then trm.Var vid1 else Const 0)))))) \ Vagree (mk_v I (OProd (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0))) (OSing vid2 (Plus (Times ($f fid2 (\i. if i = vid1 then trm.Var vid1 else Const 0)) (trm.Var vid2)) ($f fid3 (\i. if i = vid1 then trm.Var vid1 else Const 0))))) (\ y. if vid2 = y then 0 else fst (a, b) $ y, b) (sol t)) (mk_xode I (OProd (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0))) (OSing vid2 (Plus (Times ($f fid2 (\i. if i = vid1 then trm.Var vid1 else Const 0)) (trm.Var vid2)) ($f fid3 (\i. if i = vid1 then trm.Var vid1 else Const 0))))) (sol t)) (semBV I (OProd (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0))) (OSing vid2 (Plus (Times ($f fid2 (\i. if i = vid1 then trm.Var vid1 else Const 0)) (trm.Var vid2)) ($f fid3 (\i. if i = vid1 then trm.Var vid1 else Const 0))))))" then have ag:" Vagree (mk_v I (OProd (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0))) (OSing vid2 (Plus (Times ($f fid2 (\i. if i = vid1 then trm.Var vid1 else Const 0)) (trm.Var vid2)) ($f fid3 (\i. if i = vid1 then trm.Var vid1 else Const 0))))) (\ y. if vid2 = y then 0 else fst (a, b) $ y, b) (sol t)) (mk_xode I (OProd (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0))) (OSing vid2 (Plus (Times ($f fid2 (\i. if i = vid1 then trm.Var vid1 else Const 0)) (trm.Var vid2)) ($f fid3 (\i. if i = vid1 then trm.Var vid1 else Const 0))))) (sol t)) (semBV I (OProd (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0))) (OSing vid2 (Plus (Times ($f fid2 (\i. if i = vid1 then trm.Var vid1 else Const 0)) (trm.Var vid2)) ($f fid3 (\i. if i = vid1 then trm.Var vid1 else Const 0))))))" by auto have sembv:"(semBV I (OProd (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0))) (OSing vid2 (Plus (Times ($f fid2 (\i. if i = vid1 then trm.Var vid1 else Const 0)) (trm.Var vid2)) ($f fid3 (\i. if i = vid1 then trm.Var vid1 else Const 0)))))) = {Inl vid1, Inr vid1, Inl vid2, Inr vid2}" by auto have sub:"{Inl vid1} \ {Inl vid1, Inr vid1, Inl vid2, Inr vid2}" by auto have ag':"Vagree (mk_v I (OProd (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0))) (OSing vid2 (Plus (Times ($f fid2 (\i. if i = vid1 then trm.Var vid1 else Const 0)) (trm.Var vid2)) ($f fid3 (\i. if i = vid1 then trm.Var vid1 else Const 0))))) (\ y. if vid2 = y then 0 else fst (a, b) $ y, b) (sol t)) (mk_xode I (OProd (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0))) (OSing vid2 (Plus (Times ($f fid2 (\i. if i = vid1 then trm.Var vid1 else Const 0)) (trm.Var vid2)) ($f fid3 (\i. if i = vid1 then trm.Var vid1 else Const 0))))) (sol t)) {Inl vid1}" using ag sembv agree_sub[OF sub] by auto then have "fst ?res_state $ vid1 = fst ((mk_xode I (OProd (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0))) (OSing vid2 (Plus (Times ($f fid2 (\i. if i = vid1 then trm.Var vid1 else Const 0)) (trm.Var vid2)) ($f fid3 (\i. if i = vid1 then trm.Var vid1 else Const 0))))) (sol t))) $ vid1" unfolding Vagree_def by blast moreover have "... = sol t $ vid1" by auto ultimately show "?thesis" by linarith qed have agree:"Vagree ?aaba' (?res_state) (FVF (p1 vid2 vid1))" unfolding p1_def Vagree_def using aabaX res_stateX by auto have fml_sem_eq:"(?res_state \ fml_sem I (p1 vid2 vid1)) = (?aaba' \ fml_sem I (p1 vid2 vid1))" using coincidence_formula[OF p2safe Iagree_refl agree, of I] by auto then show "Predicates I vid2 (\ i. dterm_sem I (if i = vid1 then trm.Var vid1 else Const 0) (mk_v I (OProd (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0))) (OSing vid2 (Plus (Times ($f fid2 (\i. if i = vid1 then trm.Var vid1 else Const 0)) (trm.Var vid2)) ($f fid3 (\i. if i = vid1 then trm.Var vid1 else Const 0))))) (\ y. if vid2 = y then 0 else fst (a, b) $ y, b) (sol t)))" using pred2 unfolding p1_def expand_singleton by auto qed subgoal for I a b r aa ba sol t proof - assume good_interp:"is_interp I" assume bigAll:" \aa ba. (\sol t. (aa, ba) = mk_v I (OProd (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0))) (OSing vid2 (Plus (Times ($f fid2 (\i. if i = vid1 then trm.Var vid1 else Const 0)) (trm.Var vid2)) ($f fid3 (\i. if i = vid1 then trm.Var vid1 else Const 0))))) (\ y. if vid2 = y then r else fst (a, b) $ y, b) (sol t) \ 0 \ t \ (sol solves_ode (\a b. (\ i. if i = vid1 then sterm_sem I (f1 fid1 vid1) b else 0) + (\ i. if i = vid2 then sterm_sem I (Plus (Times (f1 fid2 vid1) (trm.Var vid2)) (f1 fid3 vid1)) b else 0))) {0..t} {x. Predicates I vid1 (\ i. dterm_sem I (if i = vid1 then trm.Var vid1 else Const 0) (mk_v I (OProd (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0))) (OSing vid2 (Plus (Times ($f fid2 (\i. if i = vid1 then trm.Var vid1 else Const 0)) (trm.Var vid2)) ($f fid3 (\i. if i = vid1 then trm.Var vid1 else Const 0))))) (\ y. if vid2 = y then r else fst (a, b) $ y, b) x))} \ VSagree (sol 0) (\ y. if vid2 = y then r else fst (a, b) $ y) {uu. uu = vid2 \ uu = vid1 \ uu = vid2 \ uu = vid1 \ Inl uu \ Inl ` ({x. \xa. Inl x \ FVT (if xa = vid1 then trm.Var vid1 else Const 0)} \ {x. x = vid2 \ (\xa. Inl x \ FVT (if xa = vid1 then trm.Var vid1 else Const 0))}) \ (\x. Inl uu \ FVT (if x = vid1 then trm.Var vid1 else Const 0))}) \ Predicates I vid2 (\ i. dterm_sem I (if i = vid1 then trm.Var vid1 else Const 0) (aa, ba))" assume aaba:"(aa, ba) = mk_v I (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0))) (a, b) (sol t)" assume t:"0 \ t" assume sol:"(sol solves_ode (\a b. \ i. if i = vid1 then sterm_sem I (f1 fid1 vid1) b else 0)) {0..t} {x. Predicates I vid1 (\ i. dterm_sem I (if i = vid1 then trm.Var vid1 else Const 0) (mk_v I (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0))) (a, b) x))}" assume VSA:"VSagree (sol 0) a {uu. uu = vid1 \ Inl uu \ Inl ` {x. \xa. Inl x \ FVT (if xa = vid1 then trm.Var vid1 else Const 0)} \ (\x. Inl uu \ FVT (if x = vid1 then trm.Var vid1 else Const 0))}" let ?xode = "(\a b. \ i. if i = vid1 then sterm_sem I (f1 fid1 vid1) b else 0)" let ?xconstraint = UNIV let ?ivl = "ll_on_open.existence_ivl {0 .. t} ?xode ?xconstraint 0 (sol 0)" have freef1:"dfree ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0))" by(auto simp add: dfree_Fun dfree_Const) have simple_term_inverse':"\\. dfree \ \ raw_term (simple_term \) = \" using simple_term_inverse by auto have old_lipschitz:"local_lipschitz (UNIV::real set) UNIV (\a b. \ i. if i = vid1 then sterm_sem I (f1 fid1 vid1) b else 0)" apply(rule c1_implies_local_lipschitz[where f'="(\ (t,b). blinfun_vec(\ i. if i = vid1 then blin_frechet (good_interp I) (simple_term (Function fid1 (\ i. if i = vid1 then Var vid1 else Const 0))) b else Blinfun(\ _. 0)))"]) apply auto subgoal for x apply(rule has_derivative_vec) subgoal for i apply(auto simp add: bounded_linear_Blinfun_apply good_interp_inverse good_interp) apply(auto simp add: simple_term_inverse'[OF freef1]) apply(cases "i = vid1") apply(auto simp add: f1_def expand_singleton) proof - let ?h = "(\b. Functions I fid1 (\ i. sterm_sem I (if i = vid1 then trm.Var vid1 else Const 0) b))" let ?h' = "(\b'. FunctionFrechet I fid1 (\ i. sterm_sem I (if i = vid1 then trm.Var vid1 else Const 0) x) (\ i. frechet I (if i = vid1 then trm.Var vid1 else Const 0) x b'))" let ?f = "(\ b. (\ i. sterm_sem I (if i = vid1 then trm.Var vid1 else Const 0) b))" let ?f' = "(\ b'. (\ i. frechet I (if i = vid1 then trm.Var vid1 else Const 0) x b'))" let ?g = "Functions I fid1" let ?g'= "FunctionFrechet I fid1 (?f x)" have heq:"?h = ?g \ ?f" by(rule ext, auto) have heq':"?h' = ?g' \ ?f'" by(rule ext, auto) have fderiv:"(?f has_derivative ?f') (at x)" apply(rule has_derivative_vec) by (auto simp add: svar_deriv axis_def) have gderiv:"(?g has_derivative ?g') (at (?f x))" using good_interp unfolding is_interp_def by blast have gfderiv: "((?g \ ?f) has_derivative(?g' \ ?f')) (at x)" using fderiv gderiv diff_chain_at by blast have boring_eq:"(\b. Functions I fid1 (\ i. sterm_sem I (if i = vid1 then trm.Var vid1 else Const 0) b)) = sterm_sem I ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0))" by(rule ext, auto) have "(?h has_derivative ?h') (at x)" using gfderiv heq heq' by auto then show "(sterm_sem I ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0)) has_derivative (\v'. (THE f'. \x. (Functions I fid1 has_derivative f' x) (at x)) (\ i. sterm_sem I (if i = vid1 then trm.Var vid1 else Const 0) x) (\ i. frechet I (if i = vid1 then trm.Var vid1 else Const 0) x v'))) (at x)" using boring_eq by auto qed done proof - have the_thing:"continuous_on (UNIV::('sz Rvec set)) (\b. blinfun_vec (\i. if i = vid1 then blin_frechet (good_interp I) (simple_term ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0))) b else Blinfun (\_. 0)))" apply(rule continuous_blinfun_vec') subgoal for i apply(cases "i = vid1") apply(auto) using frechet_continuous[OF good_interp freef1] by (auto simp add: continuous_on_const) done have another_cont:"continuous_on (UNIV) (\x. blinfun_vec (\i. if i = vid1 then blin_frechet (good_interp I) (simple_term ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0))) (snd x) else Blinfun (\_. 0)))" apply(rule continuous_on_compose2[of UNIV "(\b. blinfun_vec (\i. if i = vid1 then blin_frechet (good_interp I) (simple_term ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0))) b else Blinfun (\_. 0)))"]) apply(rule the_thing) by (auto intro!: continuous_intros) have ext:"(\x. case x of (t, b) \ blinfun_vec (\i. if i = vid1 then blin_frechet (good_interp I) (simple_term ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0))) b else Blinfun (\_. 0))) =(\x. blinfun_vec (\i. if i = vid1 then blin_frechet (good_interp I) (simple_term ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0))) (snd x) else Blinfun (\_. 0))) " apply(rule ext, auto) by (metis snd_conv) then show "continuous_on (UNIV) (\x. case x of (t, b) \ blinfun_vec (\i. if i = vid1 then blin_frechet (good_interp I) (simple_term ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0))) b else Blinfun (\_. 0)))" using another_cont by (simp add: another_cont local.ext) qed have old_continuous:" \x. x \ UNIV \ continuous_on UNIV (\t. \ i. if i = vid1 then sterm_sem I (f1 fid1 vid1) x else 0)" by(rule continuous_on_const) interpret ll_old: ll_on_open_it "UNIV" ?xode ?xconstraint 0 apply(standard) subgoal by auto prefer 3 subgoal by auto prefer 3 subgoal by auto apply(rule old_lipschitz) by (rule old_continuous) let ?ivl = "(ll_old.existence_ivl 0 (sol 0))" let ?flow = "ll_old.flow 0 (sol 0)" have tclosed:"{0..t} = {0--t}" using t real_Icc_closed_segment by auto have "(sol solves_ode (\a b. \ i. if i = vid1 then sterm_sem I (f1 fid1 vid1) b else 0)) {0..t} UNIV" apply(rule solves_ode_supset_range) apply(rule sol) by auto then have sol':"(sol solves_ode (\a b. \ i. if i = vid1 then sterm_sem I (f1 fid1 vid1) b else 0)) {0--t} UNIV" using tclosed by auto have sub:"{0--t} \ ll_old.existence_ivl 0 (sol 0)" apply(rule ll_old.closed_segment_subset_existence_ivl) apply(rule ll_old.existence_ivl_maximal_segment) apply(rule sol') apply(rule refl) by auto have usol_old:"(?flow usolves_ode ?xode from 0) ?ivl UNIV" by(rule ll_old.flow_usolves_ode, auto) have sol_old:"(ll_old.flow 0 (sol 0) solves_ode ?xode) ?ivl UNIV" by(rule ll_old.flow_solves_ode, auto) have another_sub:"\s. s \ {0..t} \ {s--0} \ {0..t}" unfolding closed_segment_def apply auto by (metis diff_0_right diff_left_mono mult.commute mult_left_le order.trans) have sol_eq_flow:"\s. s \ {0..t} \ sol s = ?flow s" using usol_old apply simp apply(drule usolves_odeD(4)) (* 7 subgoals*) apply auto subgoal for s x proof - assume xs0:"x \ {s--0}" assume s0:"0 \ s" and st: "s \ t" have "{s--0} \ {0..t}" using another_sub[of s] s0 st by auto then have "x \ {0..t}" using xs0 by auto then have "x \ {0--t}" using tclosed by auto then show "x \ ll_old.existence_ivl 0 (sol 0)" using sub by auto qed apply(rule solves_ode_subset) using sol' apply auto[1] subgoal for s proof - assume s0:"0 \ s" and st:"s \ t" show "{s--0} \ {0--t}" using tclosed unfolding closed_segment using s0 st using another_sub intervalE by blast qed done have sol_deriv_orig:"\s. s\?ivl \ (?flow has_derivative (\xa. xa *\<^sub>R (\ i. if i = vid1 then sterm_sem I (f1 fid1 vid1) (?flow s) else 0))) (at s within ?ivl)" using sol_old apply simp apply(drule solves_odeD(1)) by (auto simp add: has_vderiv_on_def has_vector_derivative_def) have sol_eta:"(\t. \ i. ?flow t $ i) = ?flow" by(rule ext, rule vec_extensionality, auto) have sol_deriv_eq1:"\s i. (\xa. xa *\<^sub>R (\ i. if i = vid1 then sterm_sem I (f1 fid1 vid1) (?flow s) else 0)) = (\xa. \ i. xa * (if i = vid1 then sterm_sem I (f1 fid1 vid1) (?flow s) else 0))" by(rule ext, rule vec_extensionality, auto) have sol_deriv_proj:"\s i. s\?ivl \ ((\t. ?flow t $ i) has_derivative (\xa. (xa *\<^sub>R (\ i. if i = vid1 then sterm_sem I (f1 fid1 vid1) (?flow s) else 0)) $ i)) (at s within ?ivl)" subgoal for s i apply(rule has_derivative_proj[of "(\ i t. ?flow t $ i)" "(\ i t'. (t' *\<^sub>R (\ i. if i = vid1 then sterm_sem I (f1 fid1 vid1) (?flow s) else 0)) $ i)" "(at s within ?ivl)" "i"]) using sol_deriv_orig[of s] sol_eta sol_deriv_eq1 by auto done have sol_deriv_eq2:"\s i. (\xa. xa * (if i = vid1 then sterm_sem I (f1 fid1 vid1) (?flow s) else 0)) = (\xa. (xa *\<^sub>R (\ i. if i = vid1 then sterm_sem I (f1 fid1 vid1) (?flow s) else 0)) $ i)" by(rule ext, auto) have sol_deriv_proj':"\s i. s\?ivl \ ((\t. ?flow t $ i) has_derivative (\xa. xa * (if i = vid1 then sterm_sem I (f1 fid1 vid1) (?flow s) else 0))) (at s within ?ivl)" subgoal for s i using sol_deriv_proj[of s i] sol_deriv_eq2[of i s] by metis done have sol_deriv_proj_vid1:"\s. s\?ivl \ ((\t. ?flow t $ vid1) has_derivative (\xa. xa * (sterm_sem I (f1 fid1 vid1) (?flow s)))) (at s within ?ivl)" subgoal for s using sol_deriv_proj'[of s vid1] by auto done have deriv1_args:"\s. s \ ?ivl \ ((\ t. (\ i. sterm_sem I (if i = vid1 then trm.Var vid1 else Const 0) (?flow t))) has_derivative ((\ t'. \ i . t' * (if i = vid1 then sterm_sem I (f1 fid1 vid1) (?flow s) else 0)))) (at s within ?ivl)" apply(rule has_derivative_vec) by (auto simp add: sol_deriv_proj_vid1) have con_fid:"\fid. continuous_on ?ivl (\x. sterm_sem I (f1 fid vid1) (?flow x))" subgoal for fid apply(rule has_derivative_continuous_on[of "?ivl" "(\x. sterm_sem I (f1 fid vid1) (?flow x))" "(\t t'. FunctionFrechet I fid (\ i. sterm_sem I (if i = vid1 then trm.Var vid1 else Const 0) (?flow t)) (\ i . t' * (if i = vid1 then sterm_sem I (f1 fid1 vid1) (?flow t) else 0)))"]) proof - fix s assume ivl:"s \ ?ivl" let ?h = "(\x. sterm_sem I (f1 fid vid1) (?flow x))" let ?g = "Functions I fid" let ?f = "(\x. (\ i. sterm_sem I (if i = vid1 then trm.Var vid1 else Const 0) (?flow x)))" let ?h' = "(\t'. FunctionFrechet I fid (\ i. sterm_sem I (if i = vid1 then trm.Var vid1 else Const 0) (?flow s)) (\ i. t' * (if i = vid1 then sterm_sem I (f1 fid1 vid1) (?flow s) else 0)))" let ?g' = "FunctionFrechet I fid (?f s)" let ?f' = "(\ t'. \ i . t' * (if i = vid1 then sterm_sem I (f1 fid1 vid1) (?flow s) else 0))" have heq:"?h = ?g \ ?f" unfolding comp_def f1_def expand_singleton by auto have heq':"?h' = ?g' \ ?f'" unfolding comp_def by auto have fderiv:"(?f has_derivative ?f') (at s within ?ivl)" using deriv1_args[OF ivl] by auto have gderiv:"(?g has_derivative ?g') (at (?f s) within (?f ` ?ivl))" using good_interp unfolding is_interp_def - using has_derivative_within_subset by blast + using has_derivative_subset by blast have gfderiv:"((?g \ ?f) has_derivative (?g' \ ?f')) (at s within ?ivl)" using fderiv gderiv diff_chain_within by blast show "((\x. sterm_sem I (f1 fid vid1) (?flow x)) has_derivative (\t'. FunctionFrechet I fid (\ i. sterm_sem I (if i = vid1 then trm.Var vid1 else Const 0) (?flow s)) (\ i. t' * (if i = vid1 then sterm_sem I (f1 fid1 vid1) (?flow s) else 0)))) (at s within ?ivl)" using heq heq' gfderiv by auto qed done have con:"\x. continuous_on (?ivl) (\t. x * sterm_sem I (f1 fid2 vid1) (?flow t) + sterm_sem I (f1 fid3 vid1) (?flow t))" apply(rule continuous_on_add) apply(rule continuous_on_mult_left) apply(rule con_fid[of fid2]) by(rule con_fid[of fid3]) let ?axis = "(\ i. Blinfun(axis i))" have bounded_linear_deriv:"\t. bounded_linear (\y' . y' *\<^sub>R sterm_sem I (f1 fid2 vid1) (ll_old.flow 0 (sol 0) t))" using bounded_linear_scaleR_left by blast have ll:"local_lipschitz (ll_old.existence_ivl 0 (sol 0)) UNIV (\t y. y * sterm_sem I (f1 fid2 vid1) (?flow t) + sterm_sem I (f1 fid3 vid1) (?flow t))" apply(rule c1_implies_local_lipschitz[where f'="(\ (t,y). Blinfun(\y' . y' *\<^sub>R sterm_sem I (f1 fid2 vid1) (ll_old.flow 0 (sol 0) t)))"]) apply auto subgoal for t x apply(rule has_derivative_add_const) proof - have deriv:"((\x. x * sterm_sem I (f1 fid2 vid1) (ll_old.flow 0 (sol 0) t)) has_derivative (\x. x * sterm_sem I (f1 fid2 vid1) (ll_old.flow 0 (sol 0) t))) (at x)" by(auto intro: derivative_eq_intros) have eq:"(\x. x * sterm_sem I (f1 fid2 vid1) (ll_old.flow 0 (sol 0) t)) = blinfun_apply (Blinfun (\y'. y' * sterm_sem I (f1 fid2 vid1) (ll_old.flow 0 (sol 0) t)))" apply(rule ext) using bounded_linear_deriv[of t] by (auto simp add: bounded_linear_Blinfun_apply) show "((\x. x * sterm_sem I (f1 fid2 vid1) (ll_old.flow 0 (sol 0) t)) has_derivative blinfun_apply (Blinfun (\y'. y' * sterm_sem I (f1 fid2 vid1) (ll_old.flow 0 (sol 0) t)))) (at x)" using deriv eq by auto qed apply(auto intro: continuous_intros simp add: split_beta') proof - have bounded_linear:"\x. bounded_linear (\y'. y' * sterm_sem I (f1 fid2 vid1) x)" by (simp add: bounded_linear_mult_left) have eq:"(\x. Blinfun (\y'. y' * sterm_sem I (f1 fid2 vid1) x)) = (\x. (sterm_sem I (f1 fid2 vid1) x) *\<^sub>R id_blinfun)" apply(rule ext, rule blinfun_eqI) subgoal for x i using bounded_linear[of x] apply(auto simp add: bounded_linear_Blinfun_apply) by (simp add: blinfun.scaleR_left) done have conFlow:"continuous_on (ll_old.existence_ivl 0 (sol 0)) (ll_old.flow 0 (sol 0))" using ll_old.general.flow_continuous_on by blast have conF':"continuous_on (ll_old.flow 0 (sol 0) ` ll_old.existence_ivl 0 (sol 0)) (\x. (sterm_sem I (f1 fid2 vid1) x) *\<^sub>R id_blinfun)" apply(rule continuous_on_scaleR) apply(auto intro: continuous_intros) apply(rule sterm_continuous') apply(rule good_interp) by(auto simp add: f1_def intro: dfree.intros) have conF:"continuous_on (ll_old.flow 0 (sol 0) ` ll_old.existence_ivl 0 (sol 0)) (\x. Blinfun (\y'. y' * sterm_sem I (f1 fid2 vid1) x))" apply(rule continuous_on_compose2[of "UNIV" "(\x. Blinfun (\y'. y' * x))" "(ll_old.flow 0 (sol 0) ` ll_old.existence_ivl 0 (sol 0))" "sterm_sem I (f1 fid2 vid1)"]) subgoal by (metis blinfun_mult_left.abs_eq bounded_linear_blinfun_mult_left continuous_on_eq linear_continuous_on) apply(rule sterm_continuous') apply(rule good_interp) by(auto simp add: f1_def intro: dfree.intros) show "continuous_on (ll_old.existence_ivl 0 (sol 0) \ UNIV) (\x. Blinfun (\y'. y' * sterm_sem I (f1 fid2 vid1) (ll_old.flow 0 (sol 0) (fst x))))" apply(rule continuous_on_compose2[of "ll_old.existence_ivl 0 (sol 0)" "(\x. Blinfun (\y'. y' * sterm_sem I (f1 fid2 vid1) (ll_old.flow 0 (sol 0) x)))" "(ll_old.existence_ivl 0 (sol 0) \ UNIV)" "fst"]) apply(rule continuous_on_compose2[of "(ll_old.flow 0 (sol 0) ` ll_old.existence_ivl 0 (sol 0))" "(\x. Blinfun (\y'. y' * sterm_sem I (f1 fid2 vid1) x))" "(ll_old.existence_ivl 0 (sol 0))" "(ll_old.flow 0 (sol 0))"]) using conF conFlow by (auto intro!: continuous_intros) qed let ?ivl = "ll_old.existence_ivl 0 (sol 0)" \ \Construct solution to ODE for \y'\ here:\ let ?yode = "(\t y. y * sterm_sem I (f1 fid2 vid1) (?flow t) + sterm_sem I (f1 fid3 vid1) (?flow t))" let ?ysol0 = r interpret ll_new: ll_on_open_it "?ivl" "?yode" "UNIV" 0 apply(standard) apply(auto) apply(rule ll) by(rule con) have sol_new:"(ll_new.flow 0 r solves_ode ?yode) (ll_new.existence_ivl 0 r) UNIV" by(rule ll_new.flow_solves_ode, auto) have more_lipschitz:"\tm tM. tm \ ll_old.existence_ivl 0 (sol 0) \ tM \ ll_old.existence_ivl 0 (sol 0) \ \M L. \t\{tm..tM}. \x. \x * sterm_sem I (f1 fid2 vid1) (?flow t) + sterm_sem I (f1 fid3 vid1) (?flow t)\ \ M + L * \x\" proof - fix tm tM assume tm:"tm \ ll_old.existence_ivl 0 (sol 0)" assume tM:"tM \ ll_old.existence_ivl 0 (sol 0)" let ?f2 = "(\t. sterm_sem I (f1 fid2 vid1) (ll_old.flow 0 (sol 0) t))" let ?f3 = "(\t. sterm_sem I (f1 fid3 vid1) (ll_old.flow 0 (sol 0) t))" let ?boundLP = "(\L t . (tm \ t \ t \ tM \ \?f2 t\ \ L))" let ?boundL = "(SOME L. (\t. ?boundLP L t))" have compactT:"compact {tm..tM}" by auto have sub:"{tm..tM} \ ll_old.existence_ivl 0 (sol 0)" by (metis atLeastatMost_empty_iff empty_subsetI ll_old.general.segment_subset_existence_ivl real_Icc_closed_segment tM tm) let ?f2abs = "(\x. abs(?f2 x))" have neg_compact:"\S::real set. compact S \ compact ((\x. -x) ` S)" by(rule compact_continuous_image, auto intro: continuous_intros) have compactf2:"compact (?f2 ` {tm..tM})" apply(rule compact_continuous_image) apply(rule continuous_on_compose2[of UNIV "sterm_sem I (f1 fid2 vid1)" "{tm..tM}" "ll_old.flow 0 (sol 0)"]) apply(rule sterm_continuous) apply(rule good_interp) subgoal by (auto intro: dfree.intros simp add: f1_def) apply(rule continuous_on_subset) prefer 2 apply (rule sub) subgoal using ll_old.general.flow_continuous_on by blast by auto then have boundedf2:"bounded (?f2 ` {tm..tM})" using compact_imp_bounded by auto then have boundedf2neg:"bounded ((\x. -x) ` ?f2 ` {tm..tM})" using compact_imp_bounded neg_compact by auto then have bdd_above_f2neg:"bdd_above ((\x. -x) ` ?f2 ` {tm..tM})" by (rule bounded_imp_bdd_above) then have bdd_above_f2:"bdd_above ( ?f2 ` {tm..tM})" using bounded_imp_bdd_above boundedf2 by auto have bdd_above_f2_abs:"bdd_above (abs ` ?f2 ` {tm..tM})" using bdd_above_f2neg bdd_above_f2 unfolding bdd_above_def apply auto subgoal for M1 M2 apply(rule exI[where x="max M1 M2"]) by fastforce done then have theBound:"\L. (\t. ?boundLP L t)" unfolding bdd_above_def norm_conv_dist by (auto simp add: Ball_def Bex_def norm_conv_dist image_iff norm_bcontfun_def dist_blinfun_def) then have boundLP:"\t. ?boundLP (?boundL) t" using someI[of "(\ L. \t. ?boundLP L t)"] by blast let ?boundMP = "(\M t. (tm \ t \ t \ tM \ \?f3 t\ \ M))" let ?boundM = "(SOME M. (\t. ?boundMP M t))" have compactf3:"compact (?f3 ` {tm..tM})" apply(rule compact_continuous_image) apply(rule continuous_on_compose2[of UNIV "sterm_sem I (f1 fid3 vid1)" "{tm..tM}" "ll_old.flow 0 (sol 0)"]) apply(rule sterm_continuous) apply(rule good_interp) subgoal by (auto intro: dfree.intros simp add: f1_def) apply(rule continuous_on_subset) prefer 2 apply (rule sub) subgoal using ll_old.general.flow_continuous_on by blast by auto then have boundedf3:"bounded (?f3 ` {tm..tM})" using compact_imp_bounded by auto then have boundedf3neg:"bounded ((\x. -x) ` ?f3 ` {tm..tM})" using compact_imp_bounded neg_compact by auto then have bdd_above_f3neg:"bdd_above ((\x. -x) ` ?f3 ` {tm..tM})" by (rule bounded_imp_bdd_above) then have bdd_above_f3:"bdd_above ( ?f3 ` {tm..tM})" using bounded_imp_bdd_above boundedf3 by auto have bdd_above_f3_abs:"bdd_above (abs ` ?f3 ` {tm..tM})" using bdd_above_f3neg bdd_above_f3 unfolding bdd_above_def apply auto subgoal for M1 M2 apply(rule exI[where x="max M1 M2"]) by fastforce done then have theBound:"\L. (\t. ?boundMP L t)" unfolding bdd_above_def norm_conv_dist by (auto simp add: Ball_def Bex_def norm_conv_dist image_iff norm_bcontfun_def dist_blinfun_def) then have boundMP:"\t. ?boundMP (?boundM) t" using someI[of "(\ M. \t. ?boundMP M t)"] by blast show "\M L. \t\{tm..tM}. \x. \x * ?f2 t + ?f3 t\ \ M + L * \x\" apply(rule exI[where x="?boundM"]) apply(rule exI[where x="?boundL"]) apply auto proof - fix t and x :: real assume ttm:"tm \ t" assume ttM:"t \ tM" from ttm ttM have ttmM:"tm \ t \ t \ tM" by auto have leqf3:"\?f3 t\ \ ?boundM" using boundMP ttmM by auto have leqf2:"\?f2 t\ \ ?boundL" using boundLP ttmM by auto have gr0:" \x\ \ 0" by auto have leqf2x:"\?f2 t\ * \x\ \ ?boundL * \x\" using gr0 leqf2 by (metis (no_types, lifting) real_scaleR_def scaleR_right_mono) have "\x * ?f2 t + ?f3 t\ \ \x\ * \?f2 t\ + \?f3 t\" proof - have f1: "\r ra. \r::real\ * \ra\ = \r * ra\" by (metis norm_scaleR real_norm_def real_scaleR_def) have "\r ra. \(r::real) + ra\ \ \r\ + \ra\" by (metis norm_triangle_ineq real_norm_def) then show ?thesis using f1 by presburger qed moreover have "... = \?f3 t\ + \?f2 t\ * \x\" by auto moreover have "... \ ?boundM + \?f2 t\ * \x\" using leqf3 by linarith moreover have "... \ ?boundM + ?boundL * \x\" using leqf2x by linarith ultimately show "\x * ?f2 t + ?f3 t\ \ ?boundM + ?boundL * \x\" by linarith qed qed have ivls_eq:"(ll_new.existence_ivl 0 r) = (ll_old.existence_ivl 0 (sol 0))" apply(rule ll_new.existence_ivl_eq_domain) apply auto apply (rule more_lipschitz) by auto have sub':"{0--t} \ ll_new.existence_ivl 0 r" using sub ivls_eq by auto have sol_new':"(ll_new.flow 0 r solves_ode ?yode) {0--t} UNIV" by(rule solves_ode_subset, rule sol_new, rule sub') let ?soly = "ll_new.flow 0 r" let ?sol' = "(\t. \ i. if i = vid2 then ?soly t else sol t $ i)" let ?aaba' = "mk_v I (OProd (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0))) (OSing vid2 (Plus (Times ($f fid2 (\i. if i = vid1 then trm.Var vid1 else Const 0)) (trm.Var vid2)) ($f fid3 (\i. if i = vid1 then trm.Var vid1 else Const 0))))) (\ y. if vid2 = y then r else fst (a, b) $ y, b) (?sol' t)" have duh:"(fst ?aaba', snd ?aaba') = ?aaba'" by auto note bigEx = spec[OF spec[OF bigAll, where x="fst ?aaba'"], where x="snd ?aaba'"] have sol_deriv:"\s. s \ {0..t} \ (sol has_derivative (\xa. xa *\<^sub>R (\ i. if i = vid1 then sterm_sem I (f1 fid1 vid1) (sol s) else 0))) (at s within {0..t})" using sol apply simp by(drule solves_odeD(1), auto simp add: has_vderiv_on_def has_vector_derivative_def) have silly_eq1:"(\t. \ i. sol t $ i) = sol" by(rule ext, rule vec_extensionality, auto) have silly_eq2:"\s. (\xa. \ i. (xa *\<^sub>R (\ i. if i = vid1 then sterm_sem I (f1 fid1 vid1) (sol s) else 0)) $ i) = (\xa. xa *\<^sub>R (\ i. if i = vid1 then sterm_sem I (f1 fid1 vid1) (sol s) else 0))" by(rule ext, rule vec_extensionality, auto) have sol_proj_deriv:"\s i. s \ {0..t} \ ((\ t. sol t $ i) has_derivative (\xa. (xa *\<^sub>R (\ i. if i = vid1 then sterm_sem I (f1 fid1 vid1) (sol s) else 0)) $ i)) (at s within {0..t})" subgoal for s i apply(rule has_derivative_proj) using sol_deriv[of s] silly_eq1 silly_eq2[of s] by auto done have sol_proj_deriv_vid1:"\s. s \ {0..t} \ ((\ t. sol t $ vid1) has_derivative (\xa. xa * sterm_sem I (f1 fid1 vid1) (sol s))) (at s within {0..t})" subgoal for s using sol_proj_deriv[of s vid1] by auto done have sol_proj_deriv_other:"\s i. s \ {0..t} \ i \ vid1 \ ((\ t. sol t $ i) has_derivative (\xa. 0)) (at s within {0..t})" subgoal for s i using sol_proj_deriv[of s i] by auto done have fact:"\x. x \{0..t} \ (ll_new.flow 0 r has_derivative (\xa. xa *\<^sub>R (ll_new.flow 0 r x * sterm_sem I (f1 fid2 vid1) (ll_old.flow 0 (sol 0) x) + sterm_sem I (f1 fid3 vid1) (ll_old.flow 0 (sol 0) x)))) (at x within {0 .. t})" using sol_new' apply simp apply(drule solves_odeD(1)) using tclosed unfolding has_vderiv_on_def has_vector_derivative_def by auto have new_sol_deriv:"\s. s \ {0..t} \ (ll_new.flow 0 r has_derivative (\xa. xa *\<^sub>R (ll_new.flow 0 r s * sterm_sem I (f1 fid2 vid1) (sol s) + sterm_sem I (f1 fid3 vid1) (sol s)))) (at s within {0.. t})" subgoal for s using fact[of s] tclosed sol_eq_flow[of s] by auto done have sterm_agree:"\s. Vagree (\ i. if i = vid2 then ll_new.flow 0 r s else sol s $ i, undefined) (sol s, undefined) {Inl vid1}" subgoal for s unfolding Vagree_def using vne12 by auto done have FVF:"(FVT (f1 fid2 vid1)) = {Inl vid1}" unfolding f1_def expand_singleton apply auto subgoal for x xa by (cases "xa = vid1", auto) done have FVF2:"(FVT (f1 fid3 vid1)) = {Inl vid1}" unfolding f1_def expand_singleton apply auto subgoal for x xa by (cases "xa = vid1", auto) done have sterm_agree_FVF:"\s. Vagree (\ i. if i = vid2 then ll_new.flow 0 r s else sol s $ i, undefined) (sol s, undefined) (FVT (f1 fid2 vid1))" using sterm_agree FVF by auto have sterm_agree_FVF2:"\s. Vagree (\ i. if i = vid2 then ll_new.flow 0 r s else sol s $ i, undefined) (sol s, undefined) (FVT (f1 fid3 vid1))" using sterm_agree FVF2 by auto have y_component_sem_eq2:"\s. sterm_sem I (f1 fid2 vid1) (\ i. if i = vid2 then ll_new.flow 0 r s else sol s $ i) = sterm_sem I (f1 fid2 vid1) (sol s)" using coincidence_sterm[OF sterm_agree_FVF, of I] by auto have y_component_sem_eq3:"\s. sterm_sem I (f1 fid3 vid1) (\ i. if i = vid2 then ll_new.flow 0 r s else sol s $ i) = sterm_sem I (f1 fid3 vid1) (sol s)" using coincidence_sterm[OF sterm_agree_FVF2, of I] by auto have y_component_ode_eq:"\s. s \ {0..t} \ (\xa. xa * (ll_new.flow 0 r s * sterm_sem I (f1 fid2 vid1) (sol s) + sterm_sem I (f1 fid3 vid1) (sol s))) = (\xa. xa * (sterm_sem I (f1 fid2 vid1) (\ i. if i = vid2 then ll_new.flow 0 r s else sol s $ i) * ll_new.flow 0 r s + sterm_sem I (f1 fid3 vid1) (\ i. if i = vid2 then ll_new.flow 0 r s else sol s $ i)))" subgoal for s apply(rule ext) subgoal for xa using y_component_sem_eq2 y_component_sem_eq3 by auto done done have agree_vid1:"\s. Vagree (sol s, undefined) (\ i. if i = vid2 then ll_new.flow 0 r s else sol s $ i, undefined) {Inl vid1}" unfolding Vagree_def using vne12 by auto have FVT_vid1:"FVT(f1 fid1 vid1) = {Inl vid1}" apply(auto simp add: f1_def) subgoal for x xa apply(cases "xa = vid1") by auto done have agree_vid1_FVT:"\s. Vagree (sol s, undefined) (\ i. if i = vid2 then ll_new.flow 0 r s else sol s $ i, undefined) (FVT (f1 fid1 vid1))" using FVT_vid1 agree_vid1 by auto have sterm_eq_vid1:"\s. sterm_sem I (f1 fid1 vid1) (sol s) = sterm_sem I (f1 fid1 vid1) (\ i. if i = vid2 then ll_new.flow 0 r s else sol s $ i)" subgoal for s using coincidence_sterm[OF agree_vid1_FVT[of s], of I] by auto done have vid1_deriv_eq:"\s. (\xa. xa * sterm_sem I (f1 fid1 vid1) (sol s)) = (\xa. xa * sterm_sem I (f1 fid1 vid1) (\ i. if i = vid2 then ll_new.flow 0 r s else sol s $ i))" subgoal for s apply(rule ext) subgoal for x' using sterm_eq_vid1[of s] by auto done done have inner_deriv:"\s. s \ {0..t} \ ((\t. \ i. if i = vid2 then ll_new.flow 0 r t else sol t $ i) has_derivative (\xa. (\ i. xa * (if i = vid1 then sterm_sem I (f1 fid1 vid1) (\ i. if i = vid2 then ll_new.flow 0 r s else sol s $ i) else if i = vid2 then sterm_sem I (Plus (Times (f1 fid2 vid1) (trm.Var vid2)) (f1 fid3 vid1)) (\ i. if i = vid2 then ll_new.flow 0 r s else sol s $ i) else 0)))) (at s within {0..t})" subgoal for s apply(rule has_derivative_vec) subgoal for i apply(cases "i = vid2") subgoal using vne12 using new_sol_deriv[of s] using y_component_ode_eq by auto subgoal apply(cases "i = vid1") using sol_proj_deriv_vid1[of s] vid1_deriv_eq[of s] sol_proj_deriv_other[of s i] by auto done done done have deriv_eta:"\s. (\xa. xa *\<^sub>R ((\ i. if i = vid1 then sterm_sem I (f1 fid1 vid1) (\ i. if i = vid2 then ll_new.flow 0 r s else sol s $ i) else 0) + (\ i. if i = vid2 then sterm_sem I (Plus (Times (f1 fid2 vid1) (trm.Var vid2)) (f1 fid3 vid1)) (\ i. if i = vid2 then ll_new.flow 0 r s else sol s $ i) else 0))) = (\xa. (\ i. xa * (if i = vid1 then sterm_sem I (f1 fid1 vid1) (\ i. if i = vid2 then ll_new.flow 0 r s else sol s $ i) else if i = vid2 then sterm_sem I (Plus (Times (f1 fid2 vid1) (trm.Var vid2)) (f1 fid3 vid1)) (\ i. if i = vid2 then ll_new.flow 0 r s else sol s $ i) else 0))) " subgoal for s apply(rule ext) apply(rule vec_extensionality) using vne12 by auto done have sol'_deriv:"\s. s \ {0..t} \ ((\t. \ i. if i = vid2 then ll_new.flow 0 r t else sol t $ i) has_derivative (\xa. xa *\<^sub>R ((\ i. if i = vid1 then sterm_sem I (f1 fid1 vid1) (\ i. if i = vid2 then ll_new.flow 0 r s else sol s $ i) else 0) + (\ i. if i = vid2 then sterm_sem I (Plus (Times (f1 fid2 vid1) (trm.Var vid2)) (f1 fid3 vid1)) (\ i. if i = vid2 then ll_new.flow 0 r s else sol s $ i) else 0)))) (at s within {0..t})" subgoal for s using inner_deriv[of s] deriv_eta[of s] by auto done have FVT:"\i. FVT (if i = vid1 then trm.Var vid1 else Const 0) \ {Inl vid1}" by auto have agree:"\s. Vagree (mk_v I (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0))) (a, b) (sol s)) (mk_v I (OProd (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0))) (OSing vid2 (Plus (Times ($f fid2 (\i. if i = vid1 then trm.Var vid1 else Const 0)) (trm.Var vid2)) ($f fid3 (\i. if i = vid1 then trm.Var vid1 else Const 0))))) (\ y. if vid2 = y then r else fst (a, b) $ y, b) (\ i. if i = vid2 then ll_new.flow 0 r s else sol s $ i)) {Inl vid1}" subgoal for s using mk_v_agree [of "I" "(OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0)))" "(a, b)" "(sol s)"] using mk_v_agree [of I "(OProd (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0))) (OSing vid2 (Plus (Times ($f fid2 (\i. if i = vid1 then trm.Var vid1 else Const 0)) (trm.Var vid2)) ($f fid3 (\i. if i = vid1 then trm.Var vid1 else Const 0)))))" "(\ y. if vid2 = y then r else fst (a, b) $ y, b)" "(\ i. if i = vid2 then ll_new.flow 0 r s else sol s $ i)"] unfolding Vagree_def using vne12 by simp done have agree':"\s i. Vagree (mk_v I (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0))) (a, b) (sol s)) (mk_v I (OProd (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0))) (OSing vid2 (Plus (Times ($f fid2 (\i. if i = vid1 then trm.Var vid1 else Const 0)) (trm.Var vid2)) ($f fid3 (\i. if i = vid1 then trm.Var vid1 else Const 0))))) (\ y. if vid2 = y then r else fst (a, b) $ y, b) (\ i. if i = vid2 then ll_new.flow 0 r s else sol s $ i)) (FVT (if i = vid1 then trm.Var vid1 else Const 0))" subgoal for s i using agree_sub[OF FVT[of i] agree[of s]] by auto done have safe:"\i. dsafe (if i = vid1 then trm.Var vid1 else Const 0)" subgoal for i apply(cases "i = vid1", auto) done done have dterm_sem_eq:"\s i. dterm_sem I (if i = vid1 then trm.Var vid1 else Const 0) (mk_v I (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0))) (a, b) (sol s)) = dterm_sem I (if i = vid1 then trm.Var vid1 else Const 0) (mk_v I (OProd (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0))) (OSing vid2 (Plus (Times ($f fid2 (\i. if i = vid1 then trm.Var vid1 else Const 0)) (trm.Var vid2)) ($f fid3 (\i. if i = vid1 then trm.Var vid1 else Const 0))))) (\ y. if vid2 = y then r else fst (a, b) $ y, b) (\ i. if i = vid2 then ll_new.flow 0 r s else sol s $ i))" subgoal for s i using coincidence_dterm[OF safe[of i] agree'[of s i], of I] by auto done have dterm_vec_eq:"\s. (\ i. dterm_sem I (if i = vid1 then trm.Var vid1 else Const 0) (mk_v I (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0))) (a, b) (sol s))) = (\ i. dterm_sem I (if i = vid1 then trm.Var vid1 else Const 0) (mk_v I (OProd (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0))) (OSing vid2 (Plus (Times ($f fid2 (\i. if i = vid1 then trm.Var vid1 else Const 0)) (trm.Var vid2)) ($f fid3 (\i. if i = vid1 then trm.Var vid1 else Const 0))))) (\ y. if vid2 = y then r else fst (a, b) $ y, b) (\ i. if i = vid2 then ll_new.flow 0 r s else sol s $ i)))" subgoal for s apply(rule vec_extensionality) subgoal for i using dterm_sem_eq[of i s] by auto done done have pred_same:"\s. s \ {0..t} \ Predicates I vid1 (\ i. dterm_sem I (if i = vid1 then trm.Var vid1 else Const 0) (mk_v I (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0))) (a, b) (sol s))) \ Predicates I vid1 (\ i. dterm_sem I (if i = vid1 then trm.Var vid1 else Const 0) (mk_v I (OProd (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0))) (OSing vid2 (Plus (Times ($f fid2 (\i. if i = vid1 then trm.Var vid1 else Const 0)) (trm.Var vid2)) ($f fid3 (\i. if i = vid1 then trm.Var vid1 else Const 0))))) (\ y. if vid2 = y then r else fst (a, b) $ y, b) (\ i. if i = vid2 then ll_new.flow 0 r s else sol s $ i)))" subgoal for s using dterm_vec_eq[of s] by auto done have sol'_domain:"\s. 0 \ s \ s \ t \ Predicates I vid1 (\ i. dterm_sem I (if i = vid1 then trm.Var vid1 else Const 0) (mk_v I (OProd (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0))) (OSing vid2 (Plus (Times ($f fid2 (\i. if i = vid1 then trm.Var vid1 else Const 0)) (trm.Var vid2)) ($f fid3 (\i. if i = vid1 then trm.Var vid1 else Const 0))))) (\ y. if vid2 = y then r else fst (a, b) $ y, b) (\ i. if i = vid2 then ll_new.flow 0 r s else sol s $ i)))" subgoal for s using sol apply simp apply(drule solves_odeD(2)) using pred_same[of s] by auto done have sol':"(?sol' solves_ode (\a b. (\ i. if i = vid1 then sterm_sem I (f1 fid1 vid1) b else 0) + (\ i. if i = vid2 then sterm_sem I (Plus (Times (f1 fid2 vid1) (trm.Var vid2)) (f1 fid3 vid1)) b else 0))) {0..t} {x. Predicates I vid1 (\ i. dterm_sem I (if i = vid1 then trm.Var vid1 else Const 0) (mk_v I (OProd (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0))) (OSing vid2 (Plus (Times ($f fid2 (\i. if i = vid1 then trm.Var vid1 else Const 0)) (trm.Var vid2)) ($f fid3 (\i. if i = vid1 then trm.Var vid1 else Const 0))))) (\ y. if vid2 = y then r else fst (a, b) $ y, b) x))}" apply(rule solves_odeI) subgoal unfolding has_vderiv_on_def has_vector_derivative_def using sol'_deriv by auto by(auto, rule sol'_domain, auto) have set_eq:"{y. y = vid2 \ y = vid1 \ y = vid2 \ y = vid1 \ (\x. Inl y \ FVT (if x = vid1 then trm.Var vid1 else Const 0))} = {vid1, vid2}" by auto have "VSagree (?sol' 0) (\ y. if vid2 = y then r else fst (a, b) $ y) {vid1, vid2}" using VSA unfolding VSagree_def by simp then have VSA':" VSagree (?sol' 0) (\ y. if vid2 = y then r else fst (a, b) $ y) {y. y = vid2 \ y = vid1 \ y = vid2 \ y = vid1 \ (\x. Inl y \ FVT (if x = vid1 then trm.Var vid1 else Const 0))} " by (auto simp add: set_eq) have bigPre:"(\sol t. (fst ?aaba', snd ?aaba') = mk_v I (OProd (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0))) (OSing vid2 (Plus (Times ($f fid2 (\i. if i = vid1 then trm.Var vid1 else Const 0)) (trm.Var vid2)) ($f fid3 (\i. if i = vid1 then trm.Var vid1 else Const 0))))) ((\ y. if vid2 = y then r else fst (a,b) $ y), b) (sol t) \ 0 \ t \ (sol solves_ode (\a b. (\ i. if i = vid1 then sterm_sem I (f1 fid1 vid1) b else 0) + (\ i. if i = vid2 then sterm_sem I (Plus (Times (f1 fid2 vid1) (trm.Var vid2)) (f1 fid3 vid1)) b else 0))) {0..t} {x. Predicates I vid1 (\ i. dterm_sem I (if i = vid1 then trm.Var vid1 else Const 0) (mk_v I (OProd (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0))) (OSing vid2 (Plus (Times ($f fid2 (\i. if i = vid1 then trm.Var vid1 else Const 0)) (trm.Var vid2)) ($f fid3 (\i. if i = vid1 then trm.Var vid1 else Const 0))))) ((\ y. if vid2 = y then r else (fst (a,b)) $ y), b) x))} \ VSagree (sol 0) (\ y. if vid2 = y then r else fst (a,b) $ y) {uu. uu = vid2 \ uu = vid1 \ uu = vid2 \ uu = vid1 \ Inl uu \ Inl ` ({x. \xa. Inl x \ FVT (if xa = vid1 then trm.Var vid1 else Const 0)} \ {x. x = vid2 \ (\xa. Inl x \ FVT (if xa = vid1 then trm.Var vid1 else Const 0))}) \ (\x. Inl uu \ FVT (if x = vid1 then trm.Var vid1 else Const 0))})" apply(rule exI[where x="?sol'"]) apply(rule exI[where x=t]) apply(rule conjI) subgoal by simp apply(rule conjI) subgoal by (rule t) apply(rule conjI) apply(rule sol') using VSA' unfolding VSagree_def by auto have pred_sem:"Predicates I vid2 (\ i. dterm_sem I (if i = vid1 then trm.Var vid1 else Const 0) ?aaba')" using mp[OF bigEx bigPre] by auto let ?other_state = "(mk_v I (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0))) (a, b) (sol t))" have agree:"Vagree (?aaba') (?other_state) {Inl vid1} " using mk_v_agree [of "I" "(OProd (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0))) (OSing vid2 (Plus (Times ($f fid2 (\i. if i = vid1 then trm.Var vid1 else Const 0)) (trm.Var vid2)) ($f fid3 (\i. if i = vid1 then trm.Var vid1 else Const 0)))))" "(\ y. if vid2 = y then r else fst (a, b) $ y, b)" "(?sol' t)"] using mk_v_agree [of "I" "(OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0)))" "(a, b)" "(sol t)"] unfolding Vagree_def using vne12 by simp have sub:"\i. FVT (if i = vid1 then trm.Var vid1 else Const 0) \ {Inl vid1}" by auto have agree':"\i. Vagree (?aaba') (?other_state) (FVT (if i = vid1 then trm.Var vid1 else Const 0)) " subgoal for i using agree_sub[OF sub[of i] agree] by auto done have silly_safe:"\i. dsafe (if i = vid1 then trm.Var vid1 else Const 0)" subgoal for i apply(cases "i = vid1") by (auto simp add: dsafe_Var dsafe_Const) done have dsem_eq:"(\ i. dterm_sem I (if i = vid1 then trm.Var vid1 else Const 0) ?aaba') = (\ i. dterm_sem I (if i = vid1 then trm.Var vid1 else Const 0) ?other_state)" apply(rule vec_extensionality) subgoal for i using coincidence_dterm[OF silly_safe[of i] agree'[of i], of I] by auto done show "Predicates I vid2 (\ i. dterm_sem I (if i = vid1 then trm.Var vid1 else Const 0) (mk_v I (OSing vid1 ($f fid1 (\i. if i = vid1 then trm.Var vid1 else Const 0))) (a, b) (sol t)))" using pred_sem dsem_eq by auto qed done qed end end diff --git a/thys/Ordinary_Differential_Equations/IVP/Upper_Lower_Solution.thy b/thys/Ordinary_Differential_Equations/IVP/Upper_Lower_Solution.thy --- a/thys/Ordinary_Differential_Equations/IVP/Upper_Lower_Solution.thy +++ b/thys/Ordinary_Differential_Equations/IVP/Upper_Lower_Solution.thy @@ -1,491 +1,491 @@ section \Upper and Lower Solutions\ theory Upper_Lower_Solution imports Flow begin text \Following Walter~\cite{walter} in section 9\ lemma IVT_min: fixes f :: "real \ 'b :: {linorder_topology,real_normed_vector,ordered_real_vector}" \ \generalize?\ assumes y: "f a \ y" "y \ f b" "a \ b" assumes *: "continuous_on {a .. b} f" notes [continuous_intros] = *[THEN continuous_on_subset] obtains x where "a \ x" "x \ b" "f x = y" "\x'. a \ x' \ x' < x \ f x' < y" proof - let ?s = "((\x. f x - y) -` {0..}) \ {a..b}" have "?s \ {}" using assms by auto have "closed ?s" by (rule closed_vimage_Int) (auto intro!: continuous_intros) moreover have "bounded ?s" by (rule bounded_Int) (simp add: bounded_closed_interval) ultimately have "compact ?s" using compact_eq_bounded_closed by blast from compact_attains_inf[OF this \?s \ {}\] obtain x where x: "a \ x" "x \ b" "f x \ y" and min: "\z. a \ z \ z \ b \ f z \ y \ x \ z" by auto have "f x \ y" proof (rule ccontr) assume n: "\ f x \ y" then have "\z\a. z \ x \ (\x. f x - y) z = 0" using x by (intro IVT') (auto intro!: continuous_intros simp: assms) then obtain z where z: "a \ z" "z \ x" "f z = y" by auto then have "a \ z" "z \ b" "f z \ y" using x by auto from min [OF this] z n show False by auto qed then have "a \ x" "x \ b" "f x = y" using x by (auto ) moreover have "f x' < y" if "a \ x'" "x' < x" for x' apply (rule ccontr) using min[of x'] that x by (auto simp: not_less) ultimately show ?thesis .. qed lemma filtermap_at_left_shift: "filtermap (\x. x - d) (at_left a) = at_left (a - d::real)" by (simp add: filter_eq_iff eventually_filtermap eventually_at_filter filtermap_nhds_shift[symmetric]) context fixes v v' w w'::"real \ real" and t0 t1 e::real assumes v': "(v has_vderiv_on v') {t0 <.. t1}" and w': "(w has_vderiv_on w') {t0 <.. t1}" assumes pos_ivl: "t0 < t1" assumes e_pos: "e > 0" and e_in: "t0 + e \ t1" assumes less: "\t. t0 < t \ t < t0 + e \ v t < w t" begin lemma first_intersection_crossing_derivatives: assumes na: "t0 < tg" "tg \ t1" "v tg \ w tg" notes [continuous_intros] = vderiv_on_continuous_on[OF v', THEN continuous_on_subset] vderiv_on_continuous_on[OF w', THEN continuous_on_subset] obtains x0 where "t0 < x0" "x0 \ tg" "v' x0 \ w' x0" "v x0 = w x0" "\t. t0 < t \ t < x0 \ v t < w t" proof - have "(v - w) (min tg (t0 + e / 2)) \ 0" "0 \ (v - w) tg" "min tg (t0 + e / 2) \ tg" "continuous_on {min tg (t0 + e / 2)..tg} (v - w)" using less[of "t0 + e / 2"] less[of tg]na \e > 0\ by (auto simp: min_def intro!: continuous_intros) from IVT_min[OF this] obtain x0 where x0: "min tg (t0 + e / 2) \ x0" "x0 \ tg" "v x0 = w x0" "\x'. min tg (t0 + e / 2) \ x' \ x' < x0 \ v x' < w x'" by auto then have x0_in: "t0 < x0" "x0 \ t1" using \e > 0\ na(1,2) by (auto) note \t0 < x0\ \x0 \ tg\ moreover { from v' x0_in have "(v has_derivative (\x. x * v' x0)) (at x0 within {t0<..y. (v y - (v x0 + (y - x0) * v' x0)) / norm (y - x0)) \ 0) (at x0 within {t0<..x. x * w' x0)) (at x0 within {t0<..y. (w y - (w x0 + (y - x0) * w' x0)) / norm (y - x0)) \ 0) (at x0 within {t0<..\<^sub>F x in at x0 within {t0<..\<^sub>F x in at x0 within {t0<.. x < x0" using less na(1) na(3) x0(3) x0_in(1) by (force simp: min_def eventually_at_filter intro!: order_tendstoD[OF tendsto_ident_at])+ then have "\\<^sub>F x in at x0 within {t0<..t0 < t1\ sum_sqs_eq by (auto simp: divide_simps algebra_simps min_def intro!: eventuallyI split: if_split_asm) from this tendsto_diff[OF v w] have 1: "((\x. (v x - w x) / norm (x - x0) + (v' x0 - w' x0)) \ 0) (at x0 within {t0<..\<^sub>F x in at x0 within {t0<.. (v' x0 - w' x0)" by eventually_elim (auto simp: divide_simps intro!: less_imp_le x0(4)) moreover have "at x0 within {t0<.. bot" by (simp add: \t0 < x0\ at_within_eq_bot_iff less_imp_le) ultimately have "0 \ v' x0 - w' x0" by (rule tendsto_upperbound) then have "v' x0 \ w' x0" by simp } moreover note \v x0 = w x0\ moreover have "t0 < t \ t < x0 \ v t < w t" for t by (cases "min tg (t0 + e / 2) \ t") (auto intro: x0 less) ultimately show ?thesis .. qed lemma defect_less: assumes b: "\t. t0 < t \ t \ t1 \ v' t - f t (v t) < w' t - f t (w t)" notes [continuous_intros] = vderiv_on_continuous_on[OF v', THEN continuous_on_subset] vderiv_on_continuous_on[OF w', THEN continuous_on_subset] shows "\t \ {t0 <.. t1}. v t < w t" proof (rule ccontr) assume " \ (\t\{t0 <.. t1}. v t < w t)" then obtain tu where "t0 < tu" "tu \ t1" "v tu \ w tu" by auto from first_intersection_crossing_derivatives[OF this] obtain x0 where "t0 < x0" "x0 \ tu" "w' x0 \ v' x0" "v x0 = w x0" "\t. t0 < t \ t < x0 \ v t < w t" by metis with b[of x0] \tu \ t1\ show False by simp qed end lemma has_derivatives_less_lemma: fixes v v' ::"real \ real" assumes v': "(v has_vderiv_on v') T" assumes y': "(y has_vderiv_on y') T" assumes lu: "\t. t \ T \ t > t0 \ v' t - f t (v t) < y' t - f t (y t)" assumes lower: "v t0 \ y t0" assumes eq_imp: "v t0 = y t0 \ v' t0 < y' t0" assumes t: "t0 < t" "t0 \ T" "t \ T" "is_interval T" shows "v t < y t" proof - have subset: "{t0 .. t} \ T" by (rule atMostAtLeast_subset_convex) (auto simp: assms is_interval_convex) obtain d where "0 < d" "t0 < s \ s \ t \ s < t0 + d \ v s < y s" for s proof cases assume "v t0 = y t0" from this[THEN eq_imp] have *: "0 < y' t0 - v' t0" by (simp add: ) have "((\t. y t - v t) has_vderiv_on (\t0. y' t0 - v' t0)) {t0 .. t}" by (auto intro!: derivative_intros y' v' has_vderiv_on_subset[OF _ subset]) with \t0 < t\ have d: "((\t. y t - v t) has_real_derivative y' t0 - v' t0) (at t0 within {t0 .. t})" by (auto simp: has_vderiv_on_def has_field_derivative_iff_has_vector_derivative) from has_real_derivative_pos_inc_right[OF d *] \v t0 = y t0\ obtain d where "d > 0" and vy: "h > 0 \ t0 + h \ t \ h < d \ v (t0 + h) < y (t0 + h)" for h by auto have vy: "t0 < s \ s \ t \ s < t0 + d \ v s < y s" for s using vy[of "s - t0"] by simp with \d > 0\ show ?thesis .. next assume "v t0 \ y t0" then have "v t0 < y t0" using lower by simp moreover have "continuous_on {t0 .. t} v" "continuous_on {t0 .. t} y" by (auto intro!: vderiv_on_continuous_on assms has_vderiv_on_subset[OF _ subset]) then have "(v \ v t0) (at t0 within {t0 .. t})" "(y \ y t0) (at t0 within {t0 .. t})" by (auto simp: continuous_on) ultimately have "\\<^sub>F x in at t0 within {t0 .. t}. 0 < y x - v x" by (intro order_tendstoD) (auto intro!: tendsto_eq_intros) then obtain d where "d > 0" "\x. t0 < x \ x \ t \ x < t0 + d \ v x < y x" by atomize_elim (auto simp: eventually_at algebra_simps dist_real_def) then show ?thesis .. qed with \d > 0\ \t0 < t\ obtain e where "e > 0" "t0 + e \ t" "t0 < s \ s < t0 + e \ v s < y s" for s by atomize_elim (auto simp: min_def divide_simps intro!: exI[where x="min (d/2) ((t - t0) / 2)"] split: if_split_asm) from defect_less[ OF has_vderiv_on_subset[OF v'] has_vderiv_on_subset[OF y'] \t0 < t\ this lu] show "v t < y t" using \t0 < t\ subset by (auto simp: subset_iff assms) qed lemma strict_lower_solution: fixes v v' ::"real \ real" assumes sol: "(y solves_ode f) T X" assumes v': "(v has_vderiv_on v') T" assumes lower: "\t. t \ T \ t > t0 \ v' t < f t (v t)" assumes iv: "v t0 \ y t0" "v t0 = y t0 \ v' t0 < f t0 (y t0)" assumes t: "t0 < t" "t0 \ T" "t \ T" "is_interval T" shows "v t < y t" proof - note v' moreover note solves_odeD(1)[OF sol] moreover have 3: "v' t - f t (v t) < f t (y t) - f t (y t)" if "t \ T" "t > t0" for t using lower(1)[OF that] by arith moreover note iv moreover note t ultimately show "v t < y t" by (rule has_derivatives_less_lemma) qed lemma strict_upper_solution: fixes w w'::"real \ real" assumes sol: "(y solves_ode f) T X" assumes w': "(w has_vderiv_on w') T" and upper: "\t. t \ T \ t > t0 \ f t (w t) < w' t" and iv: "y t0 \ w t0" "y t0 = w t0 \ f t0 (y t0) < w' t0" assumes t: "t0 < t" "t0 \ T" "t \ T" "is_interval T" shows "y t < w t" proof - note solves_odeD(1)[OF sol] moreover note w' moreover have "f t (y t) - f t (y t) < w' t - f t (w t)" if "t \ T" "t > t0" for t using upper(1)[OF that] by arith moreover note iv moreover note t ultimately show "y t < w t" by (rule has_derivatives_less_lemma) qed lemma uniform_limit_at_within_subset: assumes "uniform_limit S x l (at t within T)" assumes "U \ T" shows "uniform_limit S x l (at t within U)" by (metis assms(1) assms(2) eventually_within_Un filterlim_iff subset_Un_eq) lemma uniform_limit_le: fixes f::"'c \ 'a \ 'b::{metric_space, linorder_topology}" assumes I: "I \ bot" assumes u: "uniform_limit X f g I" assumes u': "uniform_limit X f' g' I" assumes "\\<^sub>F i in I. \x \ X. f i x \ f' i x" assumes "x \ X" shows "g x \ g' x" proof - have "\\<^sub>F i in I. f i x \ f' i x" using assms by (simp add: eventually_mono) with I tendsto_uniform_limitI[OF u' \x \ X\] tendsto_uniform_limitI[OF u \x \ X\] show ?thesis by (rule tendsto_le) qed lemma uniform_limit_le_const: fixes f::"'c \ 'a \ 'b::{metric_space, linorder_topology}" assumes I: "I \ bot" assumes u: "uniform_limit X f g I" assumes "\\<^sub>F i in I. \x \ X. f i x \ h x" assumes "x \ X" shows "g x \ h x" proof - have "\\<^sub>F i in I. f i x \ h x" using assms by (simp add: eventually_mono) then show ?thesis by (metis tendsto_upperbound I tendsto_uniform_limitI[OF u \x \ X\]) qed lemma uniform_limit_ge_const: fixes f::"'c \ 'a \ 'b::{metric_space, linorder_topology}" assumes I: "I \ bot" assumes u: "uniform_limit X f g I" assumes "\\<^sub>F i in I. \x \ X. h x \ f i x" assumes "x \ X" shows "h x \ g x" proof - have "\\<^sub>F i in I. h x \ f i x" using assms by (simp add: eventually_mono) then show ?thesis by (metis tendsto_lowerbound I tendsto_uniform_limitI[OF u \x \ X\]) qed locale ll_on_open_real = ll_on_open T f X for T f and X::"real set" begin lemma lower_solution: fixes v v' ::"real \ real" assumes sol: "(y solves_ode f) S X" assumes v': "(v has_vderiv_on v') S" assumes lower: "\t. t \ S \ t > t0 \ v' t < f t (v t)" assumes iv: "v t0 \ y t0" assumes t: "t0 \ t" "t0 \ S" "t \ S" "is_interval S" "S \ T" shows "v t \ y t" proof cases assume "v t0 = y t0" have "{t0 -- t} \ S" using t by (simp add: closed_segment_subset is_interval_convex) with sol have "(y solves_ode f) {t0 -- t} X" using order_refl by (rule solves_ode_on_subset) moreover note refl moreover have "{t0 -- t} \ T" using \{t0 -- t} \ S\ \S \ T\ by (rule order_trans) ultimately have t_ex: "t \ existence_ivl t0 (y t0)" by (rule existence_ivl_maximal_segment) have t0_ex: "t0 \ existence_ivl t0 (y t0)" using in_existence_between_zeroI t_ex by blast have "t0 \ T" using assms(9) t(2) by blast from uniform_limit_flow[OF t0_ex t_ex] \t0 \ t\ have "uniform_limit {t0..t} (flow t0) (flow t0 (y t0)) (at (y t0))" by simp then have "uniform_limit {t0..t} (flow t0) (flow t0 (y t0)) (at_right (y t0))" by (rule uniform_limit_at_within_subset) simp moreover { have "\\<^sub>F i in at (y t0). t \ existence_ivl t0 i" by (rule eventually_mem_existence_ivl) fact then have "\\<^sub>F i in at_right (y t0). t \ existence_ivl t0 i" unfolding eventually_at_filter by eventually_elim simp moreover have "\\<^sub>F i in at_right (y t0). i \ X" proof - have f1: "\r ra rb. r \ existence_ivl ra rb \ rb \ X" by (metis existence_ivl_reverse flow_in_domain flows_reverse) obtain rr :: "(real \ bool) \ (real \ bool) \ real" where "\p f pa fa. (\ eventually p f \ eventually pa f \ p (rr p pa)) \ (\ eventually p fa \ \ pa (rr p pa) \ eventually pa fa)" by (metis (no_types) eventually_mono) then show ?thesis using f1 calculation by meson qed moreover have "\\<^sub>F i in at_right (y t0). y t0 < i" by (simp add: eventually_at_filter) ultimately have "\\<^sub>F i in at_right (y t0). \x\{t0..t}. v x \ flow t0 i x" proof eventually_elim case (elim y') show ?case proof safe fix s assume s: "s \ {t0..t}" show "v s \ flow t0 y' s" proof cases assume "s = t0" with elim iv show ?thesis by (simp add: \t0 \ T\ \y' \ X\) next assume "s \ t0" with s have "t0 < s" by simp have "{t0 -- s} \ S" using \{t0--t} \ S\ closed_segment_eq_real_ivl s by auto from s elim have "{t0..s} \ existence_ivl t0 y'" using ivl_subset_existence_ivl by blast with flow_solves_ode have sol: "(flow t0 y' solves_ode f) {t0 .. s} X" by (rule solves_ode_on_subset) (auto intro!: \y' \ X\ \t0 \ T\) have "{t0 .. s} \ S" using \{t0 -- s} \ S\ by (simp add: closed_segment_eq_real_ivl split: if_splits) with v' have v': "(v has_vderiv_on v') {t0 .. s}" by (rule has_vderiv_on_subset) from \y t0 < y'\ \v t0 = y t0\ have less_init: "v t0 < flow t0 y' t0" by (simp add: flow_initial_time_if \t0 \ T\ \y' \ X\) from strict_lower_solution[OF sol v' lower less_imp_le[OF less_init] _ \t0 < s\] \{t0 .. s} \ S\ less_init \t0 < s\ have "v s < flow t0 y' s" by (simp add: subset_iff is_interval_cc) then show ?thesis by simp qed qed qed } moreover have "t \ {t0 .. t}" using \t0 \ t\ by simp ultimately have "v t \ flow t0 (y t0) t" by (rule uniform_limit_ge_const[OF trivial_limit_at_right_real]) also have "flow t0 (y t0) t = y t" using sol t by (intro maximal_existence_flow) auto finally show ?thesis . next assume "v t0 \ y t0" then have less: "v t0 < y t0" using iv by simp show ?thesis apply (cases "t0 = t") subgoal using iv by blast subgoal using strict_lower_solution[OF sol v' lower iv] less t by force done qed lemma upper_solution: fixes v v' ::"real \ real" assumes sol: "(y solves_ode f) S X" assumes v': "(v has_vderiv_on v') S" assumes upper: "\t. t \ S \ t > t0 \ f t (v t) < v' t" assumes iv: "y t0 \ v t0" assumes t: "t0 \ t" "t0 \ S" "t \ S" "is_interval S" "S \ T" shows "y t \ v t" proof cases assume "v t0 = y t0" have "{t0 -- t} \ S" using t by (simp add: closed_segment_subset is_interval_convex) with sol have "(y solves_ode f) {t0 -- t} X" using order_refl by (rule solves_ode_on_subset) moreover note refl moreover have "{t0 -- t} \ T" using \{t0 -- t} \ S\ \S \ T\ by (rule order_trans) ultimately have t_ex: "t \ existence_ivl t0 (y t0)" by (rule existence_ivl_maximal_segment) have t0_ex: "t0 \ existence_ivl t0 (y t0)" using in_existence_between_zeroI t_ex by blast have "t0 \ T" using assms(9) t(2) by blast from uniform_limit_flow[OF t0_ex t_ex] \t0 \ t\ have "uniform_limit {t0..t} (flow t0) (flow t0 (y t0)) (at (y t0))" by simp then have "uniform_limit {t0..t} (flow t0) (flow t0 (y t0)) (at_left (y t0))" by (rule uniform_limit_at_within_subset) simp moreover { have "\\<^sub>F i in at (y t0). t \ existence_ivl t0 i" by (rule eventually_mem_existence_ivl) fact then have "\\<^sub>F i in at_left (y t0). t \ existence_ivl t0 i" unfolding eventually_at_filter by eventually_elim simp moreover have "\\<^sub>F i in at_left (y t0). i \ X" proof - have f1: "\r ra rb. r \ existence_ivl ra rb \ rb \ X" by (metis existence_ivl_reverse flow_in_domain flows_reverse) obtain rr :: "(real \ bool) \ (real \ bool) \ real" where "\p f pa fa. (\ eventually p f \ eventually pa f \ p (rr p pa)) \ (\ eventually p fa \ \ pa (rr p pa) \ eventually pa fa)" by (metis (no_types) eventually_mono) then show ?thesis using f1 calculation by meson qed moreover have "\\<^sub>F i in at_left (y t0). i < y t0" by (simp add: eventually_at_filter) ultimately have "\\<^sub>F i in at_left (y t0). \x\{t0..t}. flow t0 i x \ v x" proof eventually_elim case (elim y') show ?case proof safe fix s assume s: "s \ {t0..t}" show "flow t0 y' s \ v s" proof cases assume "s = t0" with elim iv show ?thesis by (simp add: \t0 \ T\ \y' \ X\) next assume "s \ t0" with s have "t0 < s" by simp have "{t0 -- s} \ S" using \{t0--t} \ S\ closed_segment_eq_real_ivl s by auto from s elim have "{t0..s} \ existence_ivl t0 y'" using ivl_subset_existence_ivl by blast with flow_solves_ode have sol: "(flow t0 y' solves_ode f) {t0 .. s} X" by (rule solves_ode_on_subset) (auto intro!: \y' \ X\ \t0 \ T\) have "{t0 .. s} \ S" using \{t0 -- s} \ S\ by (simp add: closed_segment_eq_real_ivl split: if_splits) with v' have v': "(v has_vderiv_on v') {t0 .. s}" by (rule has_vderiv_on_subset) from \y' < y t0\ \v t0 = y t0\ have less_init: "flow t0 y' t0 < v t0" by (simp add: flow_initial_time_if \t0 \ T\ \y' \ X\) from strict_upper_solution[OF sol v' upper less_imp_le[OF less_init] _ \t0 < s\] \{t0 .. s} \ S\ less_init \t0 < s\ have "flow t0 y' s < v s" by (simp add: subset_iff is_interval_cc) then show ?thesis by simp qed qed qed } moreover have "t \ {t0 .. t}" using \t0 \ t\ by simp ultimately have "flow t0 (y t0) t \ v t" by (rule uniform_limit_le_const[OF trivial_limit_at_left_real]) also have "flow t0 (y t0) t = y t" using sol t by (intro maximal_existence_flow) auto finally show ?thesis . next assume "v t0 \ y t0" then have less: "y t0 < v t0" using iv by simp show ?thesis apply (cases "t0 = t") subgoal using iv by blast subgoal using strict_upper_solution[OF sol v' upper iv] less t by force done qed end end diff --git a/thys/Ordinary_Differential_Equations/Library/Multivariate_Taylor.thy b/thys/Ordinary_Differential_Equations/Library/Multivariate_Taylor.thy --- a/thys/Ordinary_Differential_Equations/Library/Multivariate_Taylor.thy +++ b/thys/Ordinary_Differential_Equations/Library/Multivariate_Taylor.thy @@ -1,610 +1,610 @@ section \Multivariate Taylor\ theory Multivariate_Taylor imports "HOL-Analysis.Analysis" "../ODE_Auxiliarities" begin no_notation vec_nth (infixl "$" 90) notation blinfun_apply (infixl "$" 999) lemma fixes f::"'a::real_normed_vector \ 'b::banach" and Df::"'a \ nat \ 'a \ 'a \ 'b" assumes "n > 0" assumes Df_Nil: "\a x. Df a 0 H H = f a" assumes Df_Cons: "\a i d. a \ closed_segment X (X + H) \ i < n \ ((\a. Df a i H H) has_derivative (Df a (Suc i) H)) (at a within G)" assumes cs: "closed_segment X (X + H) \ G" defines "i \ \x. ((1 - x) ^ (n - 1) / fact (n - 1)) *\<^sub>R Df (X + x *\<^sub>R H) n H H" shows multivariate_Taylor_has_integral: "(i has_integral f (X + H) - (\iR Df X i H H)) {0..1}" and multivariate_Taylor: "f (X + H) = (\iR Df X i H H) + integral {0..1} i" and multivariate_Taylor_integrable: "i integrable_on {0..1}" proof goal_cases case 1 let ?G = "closed_segment X (X + H)" define line where "line t = X + t *\<^sub>R H" for t have segment_eq: "closed_segment X (X + H) = line ` {0 .. 1}" by (auto simp: line_def closed_segment_def algebra_simps) have line_deriv: "\x. (line has_derivative (\t. t *\<^sub>R H)) (at x)" by (auto intro!: derivative_eq_intros simp: line_def [abs_def]) define g where "g = f o line" define Dg where "Dg n t = Df (line t) n H H" for n :: nat and t :: real note \n > 0\ moreover have Dg0: "Dg 0 = g" by (auto simp add: Dg_def Df_Nil g_def) moreover have DgSuc: "(Dg m has_vector_derivative Dg (Suc m) t) (at t within {0..1})" if "m < n" "0 \ t" "t \ 1" for m::nat and t::real proof - from that have [intro]: "line t \ ?G" using assms by (auto simp: segment_eq) - note [derivative_intros] = has_derivative_in_compose[OF _ has_derivative_within_subset[OF Df_Cons]] + note [derivative_intros] = has_derivative_in_compose[OF _ has_derivative_subset[OF Df_Cons]] interpret Df: linear "(\d. Df (line t) (Suc m) H d)" by (auto intro!: has_derivative_linear derivative_intros \m < n\) note [derivative_intros] = has_derivative_compose[OF _ line_deriv] show ?thesis using Df.scaleR \m < n\ by (auto simp: Dg_def [abs_def] has_vector_derivative_def g_def segment_eq intro!: derivative_eq_intros subsetD[OF cs]) qed ultimately have g_Taylor: "(i has_integral g 1 - (\iR Dg i 0)) {0 .. 1}" unfolding i_def Dg_def [abs_def] line_def by (rule Taylor_has_integral) auto then show c: ?case using \n > 0\ by (auto simp: g_def line_def Dg_def) case 2 show ?case using c by (simp add: integral_unique add.commute) case 3 show ?case using c by force qed subsection \Symmetric second derivative\ lemma symmetric_second_derivative_aux: assumes first_fderiv[derivative_intros]: "\a. a \ G \ (f has_derivative (f' a)) (at a within G)" assumes second_fderiv[derivative_intros]: "\i. ((\x. f' x i) has_derivative (\j. f'' j i)) (at a within G)" assumes "i \ j" "i \ 0" "j \ 0" assumes "a \ G" assumes "\s t. s \ {0..1} \ t \ {0..1} \ a + s *\<^sub>R i + t *\<^sub>R j \ G" shows "f'' j i = f'' i j" proof - let ?F = "at_right (0::real)" define B where "B i j = {a + s *\<^sub>R i + t *\<^sub>R j |s t. s \ {0..1} \ t \ {0..1}}" for i j have "B i j \ G" using assms by (auto simp: B_def) { fix e::real and i j::'a assume "e > 0" assume "i \ j" "i \ 0" "j \ 0" assume "B i j \ G" let ?ij' = "\s t. \u. a + (s * u) *\<^sub>R i + (t * u) *\<^sub>R j" let ?ij = "\t. \u. a + (t * u) *\<^sub>R i + u *\<^sub>R j" let ?i = "\t. \u. a + (t * u) *\<^sub>R i" let ?g = "\u t. f (?ij t u) - f (?i t u)" have filter_ij'I: "\P. P a \ eventually P (at a within G) \ eventually (\x. \s\{0..1}. \t\{0..1}. P (?ij' s t x)) ?F" proof - fix P assume "P a" assume "eventually P (at a within G)" hence "eventually P (at a within B i j)" by (rule filter_leD[OF at_le[OF \B i j \ G\]]) then obtain d where d: "d > 0" and "\x d2. x \ B i j \ x \ a \ dist x a < d \ P x" by (auto simp: eventually_at) with \P a\ have P: "\x d2. x \ B i j \ dist x a < d \ P x" by (case_tac "x = a") auto let ?d = "min (min (d/norm i) (d/norm j) / 2) 1" show "eventually (\x. \s\{0..1}. \t\{0..1}. P (?ij' s t x)) (at_right 0)" unfolding eventually_at proof (rule exI[where x="?d"], safe) show "0 < ?d" using \0 < d\ \i \ 0\ \j \ 0\ by simp fix x s t :: real assume *: "s \ {0..1}" "t \ {0..1}" "0 < x" "dist x 0 < ?d" show "P (?ij' s t x)" proof (rule P) have "\x y::real. x \ {0..1} \ y \ {0..1} \ x * y \ {0..1}" by (auto intro!: order_trans[OF mult_left_le_one_le]) hence "s * x \ {0..1}" "t * x \ {0..1}" using * by (auto simp: dist_norm) thus "?ij' s t x \ B i j" by (auto simp: B_def) have "norm (s *\<^sub>R x *\<^sub>R i + t *\<^sub>R x *\<^sub>R j) \ norm (s *\<^sub>R x *\<^sub>R i) + norm (t *\<^sub>R x *\<^sub>R j)" by (rule norm_triangle_ineq) also have "\ < d / 2 + d / 2" using * \i \ 0\ \j \ 0\ by (intro add_strict_mono) (auto simp: ac_simps dist_norm pos_less_divide_eq le_less_trans[OF mult_left_le_one_le]) finally show "dist (?ij' s t x) a < d" by (simp add: dist_norm) qed qed qed have filter_ijI: "eventually (\x. \t\{0..1}. P (?ij t x)) ?F" if "P a" "eventually P (at a within G)" for P using filter_ij'I[OF that] by eventually_elim (force dest: bspec[where x=1]) have filter_iI: "eventually (\x. \t\{0..1}. P (?i t x)) ?F" if "P a" "eventually P (at a within G)" for P using filter_ij'I[OF that] by eventually_elim force { from second_fderiv[of i, simplified has_derivative_iff_norm, THEN conjunct2, THEN tendstoD, OF \0 < e\] have "eventually (\x. norm (f' x i - f' a i - f'' (x - a) i) / norm (x - a) \ e) (at a within G)" by eventually_elim (simp add: dist_norm) from filter_ijI[OF _ this] filter_iI[OF _ this] \0 < e\ have "eventually (\ij. \t\{0..1}. norm (f' (?ij t ij) i - f' a i - f'' (?ij t ij - a) i) / norm (?ij t ij - a) \ e) ?F" "eventually (\ij. \t\{0..1}. norm (f' (?i t ij) i - f' a i - f'' (?i t ij - a) i) / norm (?i t ij - a) \ e) ?F" by auto moreover have "eventually (\x. x \ G) (at a within G)" unfolding eventually_at_filter by simp hence eventually_in_ij: "eventually (\x. \t\{0..1}. ?ij t x \ G) ?F" and eventually_in_i: "eventually (\x. \t\{0..1}. ?i t x \ G) ?F" using \a \ G\ by (auto dest: filter_ijI filter_iI) ultimately have "eventually (\u. norm (?g u 1 - ?g u 0 - (u * u) *\<^sub>R f'' j i) \ u * u * e * (2 * norm i + 3 * norm j)) ?F" proof eventually_elim case (elim u) hence ijsub: "(\t. ?ij t u) ` {0..1} \ G" and isub: "(\t. ?i t u) ` {0..1} \ G" by auto note has_derivative_subset[OF _ ijsub, derivative_intros] note has_derivative_subset[OF _ isub, derivative_intros] let ?g' = "\t. (\ua. u *\<^sub>R ua *\<^sub>R (f' (?ij t u) i - (f' (?i t u) i)))" have g': "((?g u) has_derivative ?g' t) (at t within {0..1})" if "t \ {0..1}" for t::real proof - from elim that have linear_f': "\c x. f' (?ij t u) (c *\<^sub>R x) = c *\<^sub>R f' (?ij t u) x" "\c x. f' (?i t u) (c *\<^sub>R x) = c *\<^sub>R f' (?i t u) x" using linear_cmul[OF has_derivative_linear, OF first_fderiv] by auto show ?thesis using elim \t \ {0..1}\ by (auto intro!: derivative_eq_intros has_derivative_in_compose[of "\t. ?ij t u" _ _ _ f] has_derivative_in_compose[of "\t. ?i t u" _ _ _ f] simp: linear_f' scaleR_diff_right mult.commute) qed from elim(1) \i \ 0\ \j \ 0\ \0 < e\ have f'ij: "\t. t \ {0..1} \ norm (f' (a + (t * u) *\<^sub>R i + u *\<^sub>R j) i - f' a i - f'' ((t * u) *\<^sub>R i + u *\<^sub>R j) i) \ e * norm ((t * u) *\<^sub>R i + u *\<^sub>R j)" using linear_0[OF has_derivative_linear, OF second_fderiv] by (case_tac "u *\<^sub>R j + (t * u) *\<^sub>R i = 0") (auto simp: field_simps simp del: pos_divide_le_eq simp add: pos_divide_le_eq[symmetric]) from elim(2) have f'i: "\t. t \ {0..1} \ norm (f' (a + (t * u) *\<^sub>R i) i - f' a i - f'' ((t * u) *\<^sub>R i) i) \ e * abs (t * u) * norm i" using \i \ 0\ \j \ 0\ linear_0[OF has_derivative_linear, OF second_fderiv] by (case_tac "t * u = 0") (auto simp: field_simps simp del: pos_divide_le_eq simp add: pos_divide_le_eq[symmetric]) have "norm (?g u 1 - ?g u 0 - (u * u) *\<^sub>R f'' j i) = norm ((?g u 1 - ?g u 0 - u *\<^sub>R (f' (a + u *\<^sub>R j) i - (f' a i))) + u *\<^sub>R (f' (a + u *\<^sub>R j) i - f' a i - u *\<^sub>R f'' j i))" (is "_ = norm (?g10 + ?f'i)") by (simp add: algebra_simps linear_cmul[OF has_derivative_linear, OF second_fderiv] linear_add[OF has_derivative_linear, OF second_fderiv]) also have "\ \ norm ?g10 + norm ?f'i" by (blast intro: order_trans add_mono norm_triangle_le) also have "0 \ {0..1::real}" by simp have "\t \ {0..1}. onorm ((\ua. (u * ua) *\<^sub>R (f' (?ij t u) i - f' (?i t u) i)) - (\ua. (u * ua) *\<^sub>R (f' (a + u *\<^sub>R j) i - f' a i))) \ 2 * u * u * e * (norm i + norm j)" (is "\t \ _. onorm (?d t) \ _") proof fix t::real assume "t \ {0..1}" show "onorm (?d t) \ 2 * u * u * e * (norm i + norm j)" proof (rule onorm_le) fix x have "norm (?d t x) = norm ((u * x) *\<^sub>R (f' (?ij t u) i - f' (?i t u) i - f' (a + u *\<^sub>R j) i + f' a i))" by (simp add: algebra_simps) also have "\ = abs (u * x) * norm (f' (?ij t u) i - f' (?i t u) i - f' (a + u *\<^sub>R j) i + f' a i)" by simp also have "\ = abs (u * x) * norm ( f' (?ij t u) i - f' a i - f'' ((t * u) *\<^sub>R i + u *\<^sub>R j) i - (f' (?i t u) i - f' a i - f'' ((t * u) *\<^sub>R i) i) - (f' (a + u *\<^sub>R j) i - f' a i - f'' (u *\<^sub>R j) i))" (is "_ = _ * norm (?dij - ?di - ?dj)") using \a \ G\ by (simp add: algebra_simps linear_add[OF has_derivative_linear[OF second_fderiv]]) also have "\ \ abs (u * x) * (norm ?dij + norm ?di + norm ?dj)" by (rule mult_left_mono[OF _ abs_ge_zero]) norm also have "\ \ abs (u * x) * (e * norm ((t * u) *\<^sub>R i + u *\<^sub>R j) + e * abs (t * u) * norm i + e * (\u\ * norm j))" using f'ij f'i f'ij[OF \0 \ {0..1}\] \t \ {0..1}\ by (auto intro!: add_mono mult_left_mono) also have "\ = abs u * abs x * abs u * (e * norm (t *\<^sub>R i + j) + e * norm (t *\<^sub>R i) + e * (norm j))" by (simp add: algebra_simps norm_scaleR[symmetric] abs_mult del: norm_scaleR) also have "\ = u * u * abs x * (e * norm (t *\<^sub>R i + j) + e * norm (t *\<^sub>R i) + e * (norm j))" by (simp add: ac_simps) also have "\ = u * u * e * abs x * (norm (t *\<^sub>R i + j) + norm (t *\<^sub>R i) + norm j)" by (simp add: algebra_simps) also have "\ \ u * u * e * abs x * ((norm (1 *\<^sub>R i) + norm j) + norm (1 *\<^sub>R i) + norm j)" using \t \ {0..1}\ \0 < e\ by (intro mult_left_mono add_mono) (auto intro!: norm_triangle_le add_right_mono mult_left_le_one_le zero_le_square) finally show "norm (?d t x) \ 2 * u * u * e * (norm i + norm j) * norm x" by (simp add: ac_simps) qed qed with differentiable_bound_linearization[where f="?g u" and f'="?g'", of 0 1 _ 0, OF _ g'] have "norm ?g10 \ 2 * u * u * e * (norm i + norm j)" by simp also have "norm ?f'i \ abs u * norm ((f' (a + (u) *\<^sub>R j) i - f' a i - f'' (u *\<^sub>R j) i))" using linear_cmul[OF has_derivative_linear, OF second_fderiv] by simp also have "\ \ abs u * (e * norm ((u) *\<^sub>R j))" using f'ij[OF \0 \ {0..1}\] by (auto intro: mult_left_mono) also have "\ = u * u * e * norm j" by (simp add: algebra_simps abs_mult) finally show ?case by (simp add: algebra_simps) qed } } note wlog = this have e': "norm (f'' j i - f'' i j) \ e * (5 * norm j + 5 * norm i)" if "0 < e" for e t::real proof - have "B i j = B j i" using \i \ j\ by (force simp: B_def)+ with assms \B i j \ G\ have "j \ i" "B j i \ G" by (auto simp:) from wlog[OF \0 < e\ \i \ j\ \i \ 0\ \j \ 0\ \B i j \ G\] wlog[OF \0 < e\ \j \ i\ \j \ 0\ \i \ 0\ \B j i \ G\] have "eventually (\u. norm ((u * u) *\<^sub>R f'' j i - (u * u) *\<^sub>R f'' i j) \ u * u * e * (5 * norm j + 5 * norm i)) ?F" proof eventually_elim case (elim u) have "norm ((u * u) *\<^sub>R f'' j i - (u * u) *\<^sub>R f'' i j) = norm (f (a + u *\<^sub>R j + u *\<^sub>R i) - f (a + u *\<^sub>R j) - (f (a + u *\<^sub>R i) - f a) - (u * u) *\<^sub>R f'' i j - (f (a + u *\<^sub>R i + u *\<^sub>R j) - f (a + u *\<^sub>R i) - (f (a + u *\<^sub>R j) - f a) - (u * u) *\<^sub>R f'' j i))" by (simp add: field_simps) also have "\ \ u * u * e * (2 * norm j + 3 * norm i) + u * u * e * (3 * norm j + 2 * norm i)" using elim by (intro order_trans[OF norm_triangle_ineq4]) (auto simp: ac_simps intro: add_mono) finally show ?case by (simp add: algebra_simps) qed hence "eventually (\u. norm ((u * u) *\<^sub>R (f'' j i - f'' i j)) \ u * u * e * (5 * norm j + 5 * norm i)) ?F" by (simp add: algebra_simps) hence "eventually (\u. (u * u) * norm ((f'' j i - f'' i j)) \ (u * u) * (e * (5 * norm j + 5 * norm i))) ?F" by (simp add: ac_simps) hence "eventually (\u. norm ((f'' j i - f'' i j)) \ e * (5 * norm j + 5 * norm i)) ?F" unfolding mult_le_cancel_left eventually_at_filter by eventually_elim auto then show ?thesis by (auto simp add:eventually_at dist_norm dest!: bspec[where x="d/2" for d]) qed have e: "norm (f'' j i - f'' i j) < e" if "0 < e" for e::real proof - let ?e = "e/2/(5 * norm j + 5 * norm i)" have "?e > 0" using \0 < e\ \i \ 0\ \j \ 0\ by (auto intro!: divide_pos_pos add_pos_pos) from e'[OF this] have "norm (f'' j i - f'' i j) \ ?e * (5 * norm j + 5 * norm i)" . also have "\ = e / 2" using \i \ 0\ \j \ 0\ by (auto simp: ac_simps add_nonneg_eq_0_iff) also have "\ < e" using \0 < e\ by simp finally show ?thesis . qed have "norm (f'' j i - f'' i j) = 0" proof (rule ccontr) assume "norm (f'' j i - f'' i j) \ 0" hence "norm (f'' j i - f'' i j) > 0" by simp from e[OF this] show False by simp qed thus ?thesis by simp qed locale second_derivative_within = fixes f f' f'' a G assumes first_fderiv[derivative_intros]: "\a. a \ G \ (f has_derivative blinfun_apply (f' a)) (at a within G)" assumes in_G: "a \ G" assumes second_fderiv[derivative_intros]: "(f' has_derivative blinfun_apply f'') (at a within G)" begin lemma symmetric_second_derivative_within: assumes "a \ G" assumes "\s t. s \ {0..1} \ t \ {0..1} \ a + s *\<^sub>R i + t *\<^sub>R j \ G" shows "f'' i j = f'' j i" apply (cases "i = j \ i = 0 \ j = 0") apply (force simp add: blinfun.zero_right blinfun.zero_left) using first_fderiv _ _ _ _ assms by (rule symmetric_second_derivative_aux[symmetric]) (auto intro!: derivative_eq_intros simp: blinfun.bilinear_simps assms) end locale second_derivative = fixes f::"'a::real_normed_vector \ 'b::banach" and f' :: "'a \ 'a \\<^sub>L 'b" and f'' :: "'a \\<^sub>L 'a \\<^sub>L 'b" and a :: 'a and G :: "'a set" assumes first_fderiv[derivative_intros]: "\a. a \ G \ (f has_derivative f' a) (at a)" assumes in_G: "a \ interior G" assumes second_fderiv[derivative_intros]: "(f' has_derivative f'') (at a)" begin lemma symmetric_second_derivative: assumes "a \ interior G" shows "f'' i j = f'' j i" proof - from assms have "a \ G" using interior_subset by blast interpret second_derivative_within by unfold_locales (auto intro!: derivative_intros intro: has_derivative_at_withinI \a \ G\) from assms open_interior[of G] interior_subset[of G] obtain e where e: "e > 0" "\y. dist y a < e \ y \ G" by (force simp: open_dist) define e' where "e' = e / 3" define i' j' where "i' = e' *\<^sub>R i /\<^sub>R norm i" and "j' = e' *\<^sub>R j /\<^sub>R norm j" hence "norm i' \ e'" "norm j' \ e'" by (auto simp: field_simps e'_def \0 < e\ less_imp_le) hence "\s\ \ 1 \ \t\ \ 1 \ norm (s *\<^sub>R i' + t *\<^sub>R j') \ e' + e'" for s t by (intro norm_triangle_le[OF add_mono]) (auto intro!: order_trans[OF mult_left_le_one_le]) also have "\ < e" by (simp add: e'_def \0 < e\) finally have "f'' $ i' $ j' = f'' $ j' $ i'" by (intro symmetric_second_derivative_within \a \ G\ e) (auto simp add: dist_norm) thus ?thesis using e(1) by (auto simp: i'_def j'_def e'_def blinfun.zero_right blinfun.zero_left blinfun.scaleR_left blinfun.scaleR_right algebra_simps) qed end lemma uniform_explicit_remainder_Taylor_1: fixes f::"'a::{banach,heine_borel,perfect_space} \ 'b::banach" assumes f'[derivative_intros]: "\x. x \ G \ (f has_derivative blinfun_apply (f' x)) (at x)" assumes f'_cont: "\x. x \ G \ isCont f' x" assumes "open G" assumes "J \ {}" "compact J" "J \ G" assumes "e > 0" obtains d R where "d > 0" "\x z. f z = f x + f' x (z - x) + R x z" "\x y. x \ J \ y \ J \ dist x y < d \ norm (R x y) \ e * dist x y" "continuous_on (G \ G) (\(a, b). R a b)" proof - from assms have "continuous_on G f'" by (auto intro!: continuous_at_imp_continuous_on) note [continuous_intros] = continuous_on_compose2[OF this] define R where "R x z = f z - f x - f' x (z - x)" for x z from compact_in_open_separated[OF \J \ {}\ \compact J\ \open G\ \J \ G\] obtain \ where \: "0 < \" "{x. infdist x J \ \} \ G" (is "?J' \ _") by auto hence infdist_in_G: "infdist x J \ \ \ x \ G" for x by auto have dist_in_G: "\y. dist x y < \ \ y \ G" if "x \ J" for x by (auto intro!: infdist_in_G infdist_le2 that simp: dist_commute) have "compact ?J'" by (rule compact_infdist_le; fact) let ?seg = ?J' from \continuous_on G f'\ have ucont: "uniformly_continuous_on ?seg f'" using \?seg \ G\ by (auto intro!: compact_uniformly_continuous \compact ?seg\ intro: continuous_on_subset) define e' where "e' = e / 2" have "e' > 0" using \e > 0\ by (simp add: e'_def) from ucont[unfolded uniformly_continuous_on_def, rule_format, OF \0 < e'\] obtain du where du: "du > 0" "\x y. x \ ?seg \ y \ ?seg \ dist x y < du \ norm (f' x - f' y) < e'" by (auto simp: dist_norm) have "min \ du > 0" using \du > 0\ \\ > 0\ by simp moreover have "f z = f x + f' x (z - x) + R x z" for x z by (auto simp: R_def) moreover { fix x z::'a assume "x \ J" "z \ J" hence "x \ G" "z \ G" using assms by auto assume "dist x z < min \ du" hence d_eta: "dist x z < \" and d_du: "dist x z < du" by (auto simp add: min_def split: if_split_asm) from \dist x z < \\ have line_in: "\xa. 0 \ xa \ xa \ 1 \ x + xa *\<^sub>R (z - x) \ G" "(\xa. x + xa *\<^sub>R (z - x)) ` {0..1} \ G" by (auto intro!: dist_in_G \x \ J\ le_less_trans[OF mult_left_le_one_le] simp: dist_norm norm_minus_commute) have "R x z = f z - f x - f' x (z - x)" by (simp add: R_def) also have "f z - f x = f (x + (z - x)) - f x" by simp also have "f (x + (z - x)) - f x = integral {0..1} (\t. (f' (x + t *\<^sub>R (z - x))) (z - x))" using \dist x z < \\ by (intro mvt_integral[of "ball x \" f f' x "z - x"]) (auto simp: dist_norm norm_minus_commute at_within_ball \0 < \\ mem_ball intro!: le_less_trans[OF mult_left_le_one_le] derivative_eq_intros dist_in_G \x \ J\) also have "(integral {0..1} (\t. (f' (x + t *\<^sub>R (z - x))) (z - x)) - (f' x) (z - x)) = integral {0..1} (\t. f' (x + t *\<^sub>R (z - x)) - f' x) (z - x)" by (simp add: Henstock_Kurzweil_Integration.integral_diff integral_linear[where h="\y. blinfun_apply y (z - x)", simplified o_def] integrable_continuous_real continuous_intros line_in blinfun.bilinear_simps[symmetric]) finally have "R x z = integral {0..1} (\t. f' (x + t *\<^sub>R (z - x)) - f' x) (z - x)" . also have "norm \ \ norm (integral {0..1} (\t. f' (x + t *\<^sub>R (z - x)) - f' x)) * norm (z - x)" by (auto intro!: order_trans[OF norm_blinfun]) also have "\ \ e' * (1 - 0) * norm (z - x)" using d_eta d_du \0 < \\ by (intro mult_right_mono integral_bound) (auto simp: dist_norm norm_minus_commute intro!: line_in du[THEN less_imp_le] infdist_le2[OF \x \ J\] line_in continuous_intros order_trans[OF mult_left_le_one_le] le_less_trans[OF mult_left_le_one_le]) also have "\ \ e * dist x z" using \0 < e\ by (simp add: e'_def norm_minus_commute dist_norm) finally have "norm (R x z) \ e * dist x z" . } moreover { from f' have f_cont: "continuous_on G f" by (rule has_derivative_continuous_on[OF has_derivative_at_withinI]) note [continuous_intros] = continuous_on_compose2[OF this] from f'_cont have f'_cont: "continuous_on G f'" by (auto intro!: continuous_at_imp_continuous_on) note continuous_on_diff2=continuous_on_diff[OF continuous_on_compose[OF continuous_on_snd] continuous_on_compose[OF continuous_on_fst], where s="G \ G", simplified] have "continuous_on (G \ G) (\(a, b). f b - f a)" by (auto intro!: continuous_intros simp: split_beta) moreover have "continuous_on (G \ G) (\(a, b). f' a (b - a))" by (auto intro!: continuous_intros simp: split_beta') ultimately have "continuous_on (G \ G) (\(a, b). R a b)" by (rule iffD1[OF continuous_on_cong[OF refl] continuous_on_diff, rotated], auto simp: R_def) } ultimately show thesis .. qed text \TODO: rename, duplication?\ locale second_derivative_within' = fixes f f' f'' a G assumes first_fderiv[derivative_intros]: "\a. a \ G \ (f has_derivative f' a) (at a within G)" assumes in_G: "a \ G" assumes second_fderiv[derivative_intros]: "\i. ((\x. f' x i) has_derivative f'' i) (at a within G)" begin lemma symmetric_second_derivative_within: assumes "a \ G" "open G" assumes "\s t. s \ {0..1} \ t \ {0..1} \ a + s *\<^sub>R i + t *\<^sub>R j \ G" shows "f'' i j = f'' j i" proof (cases "i = j \ i = 0 \ j = 0") case True interpret bounded_linear "f'' k" for k by (rule has_derivative_bounded_linear) (rule second_fderiv) have z1: "f'' j 0 = 0" "f'' i 0 = 0" by (simp_all add: zero) have f'z: "f' x 0 = 0" if "x \ G" for x proof - interpret bounded_linear "f' x" by (rule has_derivative_bounded_linear) (rule first_fderiv that)+ show ?thesis by (simp add: zero) qed note aw = at_within_open[OF \a \ G\ \open G\] have "((\x. f' x 0) has_derivative (\_. 0)) (at a within G)" apply (rule has_derivative_transform_within) apply (rule has_derivative_const[where c=0]) apply (rule zero_less_one) apply fact by (simp add: f'z) from has_derivative_unique[OF second_fderiv[unfolded aw] this[unfolded aw]] have "f'' 0 = (\_. 0)" . with True z1 show ?thesis by (auto) next case False show ?thesis using first_fderiv _ _ _ _ assms(1,3-) by (rule symmetric_second_derivative_aux[]) (use False in \auto intro!: derivative_eq_intros simp: blinfun.bilinear_simps assms\) qed end locale second_derivative_on_open = fixes f::"'a::real_normed_vector \ 'b::banach" and f' :: "'a \ 'a \ 'b" and f'' :: "'a \ 'a \ 'b" and a :: 'a and G :: "'a set" assumes first_fderiv[derivative_intros]: "\a. a \ G \ (f has_derivative f' a) (at a)" assumes in_G: "a \ G" and open_G: "open G" assumes second_fderiv[derivative_intros]: "((\x. f' x i) has_derivative f'' i) (at a)" begin lemma symmetric_second_derivative: assumes "a \ G" shows "f'' i j = f'' j i" proof - interpret second_derivative_within' by unfold_locales (auto intro!: derivative_intros intro: has_derivative_at_withinI \a \ G\) from \a \ G\ open_G obtain e where e: "e > 0" "\y. dist y a < e \ y \ G" by (force simp: open_dist) define e' where "e' = e / 3" define i' j' where "i' = e' *\<^sub>R i /\<^sub>R norm i" and "j' = e' *\<^sub>R j /\<^sub>R norm j" hence "norm i' \ e'" "norm j' \ e'" by (auto simp: field_simps e'_def \0 < e\ less_imp_le) hence "\s\ \ 1 \ \t\ \ 1 \ norm (s *\<^sub>R i' + t *\<^sub>R j') \ e' + e'" for s t by (intro norm_triangle_le[OF add_mono]) (auto intro!: order_trans[OF mult_left_le_one_le]) also have "\ < e" by (simp add: e'_def \0 < e\) finally have "f'' i' j' = f'' j' i'" by (intro symmetric_second_derivative_within \a \ G\ e) (auto simp add: dist_norm open_G) moreover interpret f'': bounded_linear "f'' k" for k by (rule has_derivative_bounded_linear) (rule second_fderiv) note aw = at_within_open[OF \a \ G\ \open G\] have z1: "f'' j 0 = 0" "f'' i 0 = 0" by (simp_all add: f''.zero) have f'z: "f' x 0 = 0" if "x \ G" for x proof - interpret bounded_linear "f' x" by (rule has_derivative_bounded_linear) (rule first_fderiv that)+ show ?thesis by (simp add: zero) qed have "((\x. f' x 0) has_derivative (\_. 0)) (at a within G)" apply (rule has_derivative_transform_within) apply (rule has_derivative_const[where c=0]) apply (rule zero_less_one) apply fact by (simp add: f'z) from has_derivative_unique[OF second_fderiv[unfolded aw] this[unfolded aw]] have z2: "f'' 0 = (\_. 0)" . have "((\a. f' a (r *\<^sub>R x)) has_derivative f'' (r *\<^sub>R x)) (at a within G)" "((\a. f' a (r *\<^sub>R x)) has_derivative (\y. r *\<^sub>R f'' x y)) (at a within G)" for r x subgoal by (rule second_fderiv) subgoal proof - have "((\a. r *\<^sub>R f' a (x)) has_derivative (\y. r *\<^sub>R f'' x y)) (at a within G)" by (auto intro!: derivative_intros) then show ?thesis apply (rule has_derivative_transform[rotated 2]) apply (rule in_G) subgoal premises prems for a' proof - interpret bounded_linear "f' a'" apply (rule has_derivative_bounded_linear) by (rule first_fderiv[OF prems]) show ?thesis by (simp add: scaleR) qed done qed done then have "((\a. f' a (r *\<^sub>R x)) has_derivative f'' (r *\<^sub>R x)) (at a)" "((\a. f' a (r *\<^sub>R x)) has_derivative (\y. r *\<^sub>R f'' x y)) (at a)" for r x unfolding aw by auto then have f'z: "f'' (r *\<^sub>R x) = (\y. r *\<^sub>R f'' x y)" for r x by (rule has_derivative_unique[where f="(\a. f' a (r *\<^sub>R x))"]) ultimately show ?thesis using e(1) by (auto simp: i'_def j'_def e'_def f''.scaleR z1 z2 blinfun.zero_right blinfun.zero_left blinfun.scaleR_left blinfun.scaleR_right algebra_simps) qed end no_notation blinfun_apply (infixl "$" 999) notation vec_nth (infixl "$" 90) end diff --git a/thys/Ordinary_Differential_Equations/Numerics/Example_Utilities.thy b/thys/Ordinary_Differential_Equations/Numerics/Example_Utilities.thy --- a/thys/Ordinary_Differential_Equations/Numerics/Example_Utilities.thy +++ b/thys/Ordinary_Differential_Equations/Numerics/Example_Utilities.thy @@ -1,2708 +1,2708 @@ theory Example_Utilities imports Init_ODE_Solver begin definition "true_form = Less (floatarith.Num 0) (floatarith.Num 1)" lemma open_true_form[intro, simp]: "open_form true_form" by (auto simp: true_form_def) lemma max_Var_form_true_form[simp]: "max_Var_form true_form = 0" by (auto simp: true_form_def) lemma interpret_form_true_form[simp]: "interpret_form true_form = (\_. True)" by (auto simp: true_form_def) lemmas [simp] = length_aforms_of_ivls declare INF_cong_simp [cong] SUP_cong_simp [cong] image_cong_simp [cong del] declare [[ cd_patterns "_ = interpret_floatariths ?fas _" "_ = interpret_floatarith ?fa _"]] concrete_definition reify_example for i j k uses reify_example hide_const (open) Print.file_output definition "file_output s f = (if s = STR '''' then f (\_. ()) else if s = STR ''-'' then f print else Print.file_output s f)" definition "aforms_of_point xs = aforms_of_ivls xs xs" definition "unit_matrix_list D = concat (map (\i. map (\j. if i = j then 1 else 0::real) [0..) l x \ list_all2 (\) x u}" context includes lifting_syntax begin lemma list_interval_transfer[transfer_rule]: "((list_all2 A) ===> (list_all2 A) ===> rel_set (list_all2 A)) list_interval list_interval" if [transfer_rule]: "(A ===> A ===> (=)) (\) (\)" "bi_total A" unfolding list_interval_def by transfer_prover end lemma in_list_interval_lengthD: "x \ list_interval a b \ length x = length a" by (auto simp: list_interval_def list_all2_lengthD) context includes floatarith_notation begin definition "varvec_fas' D C = ((map Var [0..b. (map (\i. (Num (C i)) + Var (D + D * D) * (mvmult_fa D D (map Var [D..i. (map (\j. (Num (C i)) + Var (D + D * D) * Var (D + D * i + j)) [0.. \for illustration\ assumes[simp]: "D=3" "rf = real_of_float" shows "interpret_floatariths (varvec_fas D (\i. [a, b, c] ! i)) [a, b, c, d11, d12, d13, d21, d22, d23, d31, d32, d33, 2] = [rf a, rf b, rf c, rf a + 2 * rf d11, rf a + 2 * rf d12, rf a + 2 * rf d13, rf b + 2 * rf d21, rf b + 2 * rf d22, rf b + 2 * rf d23, rf c + 2 * rf d31, rf c + 2 * rf d32, rf c + 2 * rf d33]" by (simp add: varvec_fas_def mvmult_fa_def eval_nat_numeral) definition "vareq_projections n \ \dimension\ ps \ \pairs of coordinates to project onto\ ds \ \partial derivatives w.r.t. which variables\ cs \ \(color) coding for partial derivatives\ = [(i + n * (x + 1)::nat, i + n * (y + 1), c). (i, c) \ zip ds cs, (x, y) \ ps]" definition "varvec_aforms_line D X line = approx_floatariths 30 (varvec_fas D (\i. float_of (fst (X ! i)))) (take (D + D*D) X @ line)" definition "varvec_aforms_head D X s = varvec_aforms_line D X (aforms_of_point [s])" definition "varvec_aforms_vec D X s = varvec_aforms_line D (map (\x. (fst x, zero_pdevs)) X) [aform_of_ivl 0 s]" definition "shows_aforms_vareq n \ \dimension\ ps \ \pairs of coordinates to project onto\ ds \ \partial derivatives w.r.t. which variables\ csl \ \color coding for partial derivatives ('arrow' heads)\ csh \ \color coding for partial derivatives (lines)\ s \ \scale vectors for partial derivatives\ (no_str::string) \ \default string if no C1 info is present\ X \ \affine form with C1 info\ = (case (varvec_aforms_head n X s, varvec_aforms_vec n X s) of (Some X, Some Y) \ shows_sep (\(x, y, c). shows_segments_of_aform x y X c) shows_nl (vareq_projections n ps ds csl) o shows_nl o shows_sep (\(x, y, c). shows_segments_of_aform x y Y c) shows_nl (vareq_projections n ps ds csh) o shows_nl | _ \ shows_string no_str o shows_nl)" abbreviation "print_string s \ print (String.implode s)" abbreviation "print_show s \ print_string (s '''')" value [code] "print_show (shows_aforms_vareq 3 [(x, y). x \ [0..<3], y \ [0..<3], x < y] [0..<3] [''0x0000ff'', ''0x00ff00'', ''0xff0000''] [''0x0000ff'', ''0x00ff00'', ''0xff0000''] (FloatR 1 (-1)) ''# no C1 info'' ((((\(a, b). aforms_of_ivls a b) (with_unit_matrix 3 ([10, 20, 30], [12, 22, 32]))))))" method_setup guess_rhs = \ let fun compute ctxt var lhs = let val lhs' = Code_Evaluation.dynamic_value_strict ctxt lhs; val clhs' = Thm.cterm_of ctxt lhs'; val inst = Thm.instantiate ([], [(var, clhs')]); in PRIMITIVE inst end; fun eval_schematic_rhs ctxt t = (case try (HOLogic.dest_eq o HOLogic.dest_Trueprop) t of SOME (lhs, Var var) => compute ctxt var lhs | _ => no_tac); in Scan.succeed (fn ctxt => SIMPLE_METHOD (HEADGOAL (SUBGOAL (fn (t, _) => eval_schematic_rhs ctxt t)))) end \ lemma length_aforms_of_point[simp]: "length (aforms_of_point xs) = length xs" by (auto simp: aforms_of_point_def) definition "aform2d_plot_segments x y a = shows_segments_of_aform x y a ''0x000000''" lemma list_of_eucl_prod[simp]: "list_of_eucl (x, y) = list_of_eucl x @ list_of_eucl y" by (auto simp: list_of_eucl_def Basis_list_prod_def intro!: nth_equalityI) lemma list_of_eucl_real[simp]: "list_of_eucl (x::real) = [x]" by (auto simp: list_of_eucl_def Basis_list_real_def) lemma Joints_aforms_of_ivls_self[simp]: "xs \ Joints (aforms_of_ivls xs xs)" by (auto intro!: aforms_of_ivls) lemma Joints_aforms_of_ivls_self_eq[simp]: "Joints (aforms_of_ivls xs xs) = {xs}" apply (auto ) by (auto simp: aforms_of_ivls_def Joints_def valuate_def aform_val_def intro!: nth_equalityI) lemma local_lipschitz_c1_euclideanI: fixes T::"real set" and X::"'a::euclidean_space set" and f::"real \ 'a \ 'a" assumes f': "\t x. t \ T \ x \ X \ (f t has_derivative f' t x) (at x)" assumes cont_f': "\i. i \ Basis \ continuous_on (T \ X) (\(t, x). f' t x i)" assumes "open T" assumes "open X" shows "local_lipschitz T X f" using assms apply (intro c1_implies_local_lipschitz[where f'="\(t, x). Blinfun (f' t x)"]) apply (auto simp: bounded_linear_Blinfun_apply has_derivative_bounded_linear split_beta' intro!: has_derivative_Blinfun continuous_on_blinfun_componentwise) apply (subst continuous_on_cong[OF refl]) defer apply assumption apply auto apply (subst bounded_linear_Blinfun_apply) apply (rule has_derivative_bounded_linear) by auto definition "list_aform_of_aform (x::real aform) = (fst x, list_of_pdevs (snd x))" primrec split_aforms_list:: "(real aform) list list \ nat \ nat \ (real aform) list list" where "split_aforms_list Xs i 0 = Xs" | "split_aforms_list Xs i (Suc n) = split_aforms_list (concat (map (\x. (let (a, b) = split_aforms x i in [a, b])) Xs)) i n" definition "shows_aforms x y c X = shows_lines (map (\b. (shows_segments_of_aform x y b c ''\'')) X)" end definition "the_odo ode_fas safe_form = the(mk_ode_ops ode_fas safe_form)" locale ode_interpretation = fixes safe_form::form and safe_set::"'a::executable_euclidean_space set" and ode_fas::"floatarith list" and ode::"'a \ 'a" and finite::"'n::enum" assumes dims: "DIM('a) = CARD('n)" assumes len: "length ode_fas = CARD('n)" assumes safe_set_form: "safe_set = {x. interpret_form safe_form (list_of_eucl x)}" assumes interpret_fas: "\x. x \ safe_set \ einterpret ode_fas (list_of_eucl x) = ode x" assumes odo: "mk_ode_ops ode_fas safe_form \ None" assumes isFDERIV: "\xs. interpret_form safe_form xs \ isFDERIV (length ode_fas) [0.. the_odo ode_fas safe_form" lemmas odo_def = the_odo_def lemma odo_simps[simp]: "ode_expression odo = ode_fas" "safe_form_expr odo = safe_form" using odo by (auto simp: odo_def ode_expression_mk_ode_ops safe_form_expr_mk_ode_ops) lemma safe_set: "safe_set = aform.Csafe odo" using odo dims safe_set_form isFDERIV unfolding aform.Csafe_def aform.safe_def aform.safe_form_def aform.ode_e_def by (auto simp: mk_ode_ops_def safe_set_form len split: if_splits) lemma ode: "\x. x \ safe_set \ ode x = aform.ode odo x" by (auto simp: aform.ode_def aform.ode_e_def interpret_fas) sublocale auto_ll_on_open ode safe_set by (rule aform.auto_ll_on_open_congI[OF safe_set[symmetric] ode[symmetric]]) lemma ode_has_derivative_ode_d1: "(ode has_derivative blinfun_apply (aform.ode_d1 odo x)) (at x)" if "x \ safe_set" for x proof - from aform.fderiv[OF that[unfolded safe_set]] have "(aform.ode odo has_derivative blinfun_apply (aform.ode_d1 odo x)) (at x)" by simp moreover from topological_tendstoD[OF tendsto_ident_at open_domain(2) that] have "\\<^sub>F x' in at x. x' \ safe_set" . then have "\\<^sub>F x' in at x. aform.ode odo x' = ode x'" by eventually_elim (auto simp: ode) ultimately show ?thesis by (rule has_derivative_transform_eventually) (auto simp: ode that) qed sublocale c1_on_open_euclidean ode "aform.ode_d1 odo" safe_set apply unfold_locales subgoal by simp subgoal by (simp add: ode_has_derivative_ode_d1) subgoal by (rule aform.cont_fderiv') (auto intro!: continuous_intros simp: safe_set) done sublocale transfer_eucl_vec for a::'a and n::'n by unfold_locales (simp add: dims) lemma flow_eq: "t \ existence_ivl0 x \ aform.flow0 odo x t = flow0 x t" and Dflow_eq: "t \ existence_ivl0 x \ aform.Dflow odo x t = Dflow x t" and ex_ivl_eq: "t \ aform.existence_ivl0 odo x \ aform.existence_ivl0 odo x = existence_ivl0 x" and poincare_mapsto_eq: "closed a \ aform.poincare_mapsto odo a b c d e = poincare_mapsto a b c d e" and flowsto_eq: "aform.flowsto odo = flowsto" apply - subgoal by (rule flow0_cong[symmetric]) (auto simp: safe_set ode) subgoal by (rule Dflow_cong[symmetric]) (auto simp: safe_set ode) subgoal by (rule existence_ivl0_cong[symmetric]) (auto simp: safe_set ode) subgoal apply (subst aform.poincare_mapsto_cong[OF safe_set[symmetric]]) by (auto simp: ode) subgoal apply (intro ext) apply (subst flowsto_congI[OF safe_set ode]) by (auto simp: safe_set) done definition "avf \ \x::'n rvec. cast (aform.ode odo (cast x)::'a)::'n rvec" context includes lifting_syntax begin lemma aform_ode_transfer[transfer_rule]: "((=) ===> rel_ve ===> rel_ve) aform.ode aform.ode" unfolding aform.ode_def by transfer_prover lemma cast_aform_ode: "cast (aform.ode odo (cast (x::'n rvec))::'a) = aform.ode odo x" by transfer simp lemma aform_safe_transfer[transfer_rule]: "((=) ===> rel_ve ===> (=)) aform.safe aform.safe" unfolding aform.safe_def by transfer_prover lemma aform_Csafe_transfer[transfer_rule]: "((=) ===> rel_set rel_ve) aform.Csafe aform.Csafe" unfolding aform.Csafe_def by transfer_prover lemma cast_safe_set: "(cast ` safe_set::'n rvec set) = aform.Csafe odo" unfolding safe_set by transfer simp lemma aform_ode_d_raw_transfer[transfer_rule]: "((=) ===> (=) ===> rel_ve ===> rel_ve ===> rel_ve ===> rel_ve) aform.ode_d_raw aform.ode_d_raw" unfolding aform.ode_d_raw_def by transfer_prover lemma aform_ode_d_raw_aux_transfer: "((=) ===> rel_ve ===> rel_ve ===> rel_ve) (\x xb xa. if xb \ aform.Csafe x then aform.ode_d_raw x 0 xb 0 xa else 0) (\x xb xa. if xb \ aform.Csafe x then aform.ode_d_raw x 0 xb 0 xa else 0)" by transfer_prover lemma aform_ode_d1_transfer[transfer_rule]: "((=) ===> rel_ve ===> rel_blinfun rel_ve rel_ve) aform.ode_d1 aform.ode_d1" apply (auto simp: rel_blinfun_def aform.ode_d1_def intro!: rel_funI) unfolding aform.ode_d.rep_eq using aform_ode_d_raw_aux_transfer apply - apply (drule rel_funD, rule refl) apply (drule rel_funD, assumption) apply (drule rel_funD; assumption) done lemma cast_bl_transfer[transfer_rule]: "(rel_blinfun (=) (=) ===> rel_blinfun rel_ve rel_ve) id_blinfun cast_bl" by (auto simp: rel_ve_cast rel_blinfun_def intro!: rel_funI dest!: rel_funD) lemma cast_bl_transfer'[transfer_rule]: "(rel_blinfun rel_ve rel_ve ===> rel_blinfun (=) (=)) id_blinfun cast_bl" apply (auto simp: rel_ve_cast rel_blinfun_def cast_cast intro!: rel_funI dest!: rel_funD) by (subst cast_cast) auto lemma rel_blinfun_eq[relator_eq]: "rel_blinfun (=) (=) = (=)" by (auto simp: Rel_def rel_blinfun_def blinfun_ext rel_fun_eq intro!: rel_funI ext) lemma cast_aform_ode_D1: "cast_bl (aform.ode_d1 odo (cast (x::'n rvec))::'a\\<^sub>L'a) = (aform.ode_d1 odo x::'n rvec \\<^sub>L 'n rvec)" by transfer simp end definition "vf \ \x. cast (ode (cast x))" definition "vf' \ \x::'n rvec. cast_bl (aform.ode_d1 odo (cast x::'a)) ::'n rvec \\<^sub>L 'n rvec" definition "vX \ cast ` safe_set" sublocale a?: transfer_c1_on_open_euclidean a n ode "aform.ode_d1 odo" safe_set vf vf' vX for a::'a and n::'n by unfold_locales (simp_all add: dims vf_def vf'_def vX_def) sublocale av: transfer_c1_on_open_euclidean a n "aform.ode odo" "aform.ode_d1 odo" "(aform.Csafe odo)" avf vf' vX for a::'a and n::'n apply unfold_locales unfolding vX_def by (simp_all add: dims avf_def safe_set) lemma vflow_eq: "t \ v.existence_ivl0 x \ aform.flow0 odo x t = v.flow0 x t" thm flow_eq[of t "cast x"] flow_eq[of t "cast x", untransferred] apply (subst flow_eq[of t "cast x", untransferred, symmetric]) apply simp unfolding avf_def vX_def cast_aform_ode cast_safe_set .. lemma vf'_eq: "vf' = aform.ode_d1 odo" unfolding vf'_def cast_aform_ode_D1 .. lemma vDflow_eq: "t \ v.existence_ivl0 x \ aform.Dflow odo x t = v.Dflow x t" apply (subst Dflow_eq[of t "cast x", untransferred, symmetric]) apply simp unfolding avf_def vX_def cast_aform_ode cast_safe_set vf'_eq .. lemma vex_ivl_eq: "t \ aform.existence_ivl0 odo x \ aform.existence_ivl0 odo x = v.existence_ivl0 x" apply (subst ex_ivl_eq[of t "cast x", untransferred, symmetric]) unfolding avf_def vX_def cast_aform_ode cast_safe_set vf'_eq by auto context includes lifting_syntax begin lemma id_cast_eucl1_transfer_eq: "(\x. x) = (\x. (fst x, 1\<^sub>L o\<^sub>L snd x o\<^sub>L 1\<^sub>L))" by auto lemma cast_eucl1_transfer[transfer_rule]: "(rel_prod (=) (rel_blinfun (=) (=)) ===> rel_prod rel_ve (rel_blinfun rel_ve rel_ve)) (\x. x) cast_eucl1" unfolding cast_eucl1_def id_cast_eucl1_transfer_eq apply transfer_prover_start apply (transfer_step) apply (transfer_step) apply (transfer_step) apply (transfer_step) apply (transfer_step) apply simp done end lemma avpoincare_mapsto_eq: "aform.poincare_mapsto odo a (b::'n eucl1 set) c d e = av.v.poincare_mapsto a b c d e" if "closed a" unfolding avf_def vX_def cast_aform_ode cast_safe_set vf'_eq by auto lemma vpoincare_mapsto_eq: "aform.poincare_mapsto odo a (b::'n eucl1 set) c d e = v.poincare_mapsto a b c d e" if "closed a" proof - have "closed (cast ` a::'a set)" using that by transfer auto from poincare_mapsto_eq[of "cast ` a::'a set" "cast_eucl1 ` b::('a \ 'a \\<^sub>L 'a) set" "cast ` c::'a set" "cast ` d::'a set" "cast_eucl1 ` e::('a \ 'a \\<^sub>L 'a) set", OF this, untransferred] have "v.poincare_mapsto a b c d e = av.v.poincare_mapsto a b c d e" by auto also have "\ = aform.poincare_mapsto odo a (b::'n eucl1 set) c d e" unfolding avf_def vX_def cast_aform_ode cast_safe_set vf'_eq by auto finally show ?thesis by simp qed lemma avflowsto_eq: "aform.flowsto odo = (av.v.flowsto::'n eucl1 set \ _)" proof (intro ext, goal_cases) case (1 a b c d) have "av.v.flowsto a b c d = aform.flowsto odo a b c d" unfolding avf_def vX_def cast_aform_ode cast_safe_set vf'_eq by auto then show ?case by simp qed lemma vflowsto_eq: "aform.flowsto odo = (v.flowsto::'n eucl1 set \ _)" proof (intro ext, goal_cases) case (1 a b c d) have "aform.flowsto odo (cast_eucl1 ` a::'a c1_info set) b (cast_eucl1 ` c) (cast_eucl1 ` d) = flowsto (cast_eucl1 ` a::'a c1_info set) b (cast_eucl1 ` c) (cast_eucl1 ` d)" by (subst flowsto_eq) auto from this[untransferred] have "v.flowsto a b c d = av.v.flowsto a b c d" by auto also have "\ = aform.flowsto odo a b c d" unfolding avf_def vX_def cast_aform_ode cast_safe_set vf'_eq by auto finally show ?case by simp qed context includes lifting_syntax begin lemma flow1_of_list_transfer[transfer_rule]: "(list_all2 (=) ===> rel_prod rel_ve (rel_blinfun rel_ve rel_ve)) flow1_of_list flow1_of_list" unfolding flow1_of_list_def blinfun_of_list_def o_def flow1_of_vec1_def by transfer_prover lemma c1_info_of_appr_transfer[transfer_rule]: "(rel_prod (list_all2 (=)) (rel_option (list_all2 (=))) ===> rel_set (rel_prod rel_ve (rel_blinfun rel_ve rel_ve))) aform.c1_info_of_appr aform.c1_info_of_appr" unfolding aform.c1_info_of_appr_def by transfer_prover lemma c0_info_of_appr_transfer[transfer_rule]: "((list_all2 (=)) ===> rel_set rel_ve) aform.c0_info_of_appr aform.c0_info_of_appr" unfolding aform.c0_info_of_appr_def by transfer_prover lemma aform_scaleR2_transfer[transfer_rule]: "((=) ===> (=) ===> rel_set (rel_prod A B) ===> rel_set (rel_prod A B)) scaleR2 scaleR2" if [unfolded Rel_def, transfer_rule]: "((=) ===> B ===> B) (*\<^sub>R) (*\<^sub>R)" unfolding scaleR2_def by transfer_prover lemma scaleR_rel_blinfun_transfer[transfer_rule]: "((=) ===> rel_blinfun rel_ve rel_ve ===> rel_blinfun rel_ve rel_ve) (*\<^sub>R) (*\<^sub>R)" apply (auto intro!: rel_funI simp: rel_blinfun_def blinfun.bilinear_simps) apply (drule rel_funD) apply assumption apply (rule scaleR_transfer[THEN rel_funD, THEN rel_funD]) apply auto done lemma c1_info_of_appre_transfer[transfer_rule]: "(rel_prod (rel_prod (=) (=)) (rel_prod (list_all2 (=)) (rel_option (list_all2 (=)))) ===> rel_set (rel_prod rel_ve (rel_blinfun rel_ve rel_ve))) aform.c1_info_of_appre aform.c1_info_of_appre" unfolding aform.c1_info_of_appre_def by transfer_prover lemma c1_info_of_apprs_transfer[transfer_rule]: "((=) ===> rel_set (rel_prod rel_ve (rel_blinfun rel_ve rel_ve))) aform.c1_info_of_apprs aform.c1_info_of_apprs" unfolding aform.c1_info_of_apprs_def by transfer_prover lemma c1_info_of_appr'_transfer[transfer_rule]: "(rel_option (list_all2 (=)) ===> rel_set (rel_prod rel_ve (rel_blinfun rel_ve rel_ve))) aform.c1_info_of_appr' aform.c1_info_of_appr'" unfolding aform.c1_info_of_appr'_def by transfer_prover lemma c0_info_of_apprs_transfer[transfer_rule]: "((=) ===> rel_set rel_ve) aform.c0_info_of_apprs aform.c0_info_of_apprs" unfolding aform.c0_info_of_apprs_def by transfer_prover lemma c0_info_of_appr'_transfer[transfer_rule]: "(rel_option (list_all2 (=)) ===> rel_set rel_ve) aform.c0_info_of_appr' aform.c0_info_of_appr'" unfolding aform.c0_info_of_appr'_def by transfer_prover lemma aform_Csafe_vX[simp]: "aform.Csafe odo = (vX::'n rvec set)" by (simp add: vX_def cast_safe_set) definition blinfuns_of_lvivl::"real list \ real list \ ('b \\<^sub>L 'b::executable_euclidean_space) set" where "blinfuns_of_lvivl x = blinfun_of_list ` list_interval (fst x) (snd x)" lemma blinfun_of_list_transfer[transfer_rule]: "(list_all2 (=) ===> rel_blinfun rel_ve rel_ve) blinfun_of_list blinfun_of_list" unfolding blinfun_of_list_def by transfer_prover lemma blinfuns_of_lvivl_transfer[transfer_rule]: "(rel_prod (list_all2 (=)) (list_all2 (=)) ===> rel_set (rel_blinfun rel_ve rel_ve)) blinfuns_of_lvivl blinfuns_of_lvivl" unfolding blinfuns_of_lvivl_def by transfer_prover definition "blinfuns_of_lvivl' x = (case x of None \ UNIV | Some x \ blinfuns_of_lvivl x)" lemma blinfuns_of_lvivl'_transfer[transfer_rule]: "(rel_option (rel_prod (list_all2 (=)) (list_all2 (=))) ===> rel_set (rel_blinfun rel_ve rel_ve)) blinfuns_of_lvivl' blinfuns_of_lvivl'" unfolding blinfuns_of_lvivl'_def by transfer_prover lemma atLeastAtMost_transfer[transfer_rule]: "(A ===> A ===> rel_set A) atLeastAtMost atLeastAtMost" if [transfer_rule]: "(A ===> A ===> (=)) (\) (\)" "bi_total A" "bi_unique A" unfolding atLeastAtMost_def atLeast_def atMost_def by transfer_prover lemma set_of_ivl_transfer[transfer_rule]: "(rel_prod A A ===> rel_set A) set_of_ivl set_of_ivl" if [transfer_rule]: "(A ===> A ===> (=)) (\) (\)" "bi_total A" "bi_unique A" unfolding set_of_ivl_def by transfer_prover lemma set_of_lvivl_transfer[transfer_rule]: "(rel_prod (list_all2 (=)) (list_all2 (=)) ===> rel_set rel_ve) set_of_lvivl set_of_lvivl" unfolding set_of_lvivl_def by transfer_prover lemma set_of_lvivl_eq: "set_of_lvivl I = (eucl_of_list ` list_interval (fst I) (snd I)::'b::executable_euclidean_space set)" if [simp]: "length (fst I) = DIM('b)" "length (snd I) = DIM('b)" proof (auto simp: set_of_lvivl_def list_interval_def set_of_ivl_def, goal_cases) case (1 x) with lv_rel_le[where 'a='b, param_fo, OF lv_relI lv_relI, of "fst I" "list_of_eucl x"] lv_rel_le[where 'a='b, param_fo, OF lv_relI lv_relI, of "list_of_eucl x" "snd I"] show ?case by force next case (2 x) with lv_rel_le[where 'a='b, param_fo, OF lv_relI lv_relI, of "fst I" "x"] show ?case by (auto simp: list_all2_lengthD) next case (3 x) with lv_rel_le[where 'a='b, param_fo, OF lv_relI lv_relI, of "x" "snd I"] show ?case by (auto simp: list_all2_lengthD) qed lemma bounded_linear_matrix_vector_mul[THEN bounded_linear_compose, bounded_linear_intros]: "bounded_linear ((*v) x)" for x::"real^'x^'y" unfolding linear_linear by (rule matrix_vector_mul_linear) lemma blinfun_of_list_eq: "blinfun_of_list x = blinfun_of_vmatrix (eucl_of_list x::((real, 'b) vec, 'b) vec)" if "length x = CARD('b::enum)*CARD('b)" unfolding blinfun_of_list_def apply (transfer fixing: x) apply (rule linear_eq_stdbasis) unfolding linear_conv_bounded_linear apply (auto intro!: bounded_linear_intros) proof goal_cases case (1 b) have "(eucl_of_list x::((real, 'b) vec, 'b) vec) *v b = (eucl_of_list x::((real, 'b) vec, 'b) vec) *v eucl_of_list (list_of_eucl b)" by simp also have "\ = (\ij Basis_list ! j)) *\<^sub>R Basis_list ! i)" by (subst eucl_of_list_matrix_vector_mult_eq_sum_nth_Basis_list) (auto simp: that) also have "\ = (\i\Basis. \j\Basis. (b \ j * x ! (index Basis_list i * CARD('b) + index Basis_list j)) *\<^sub>R i)" apply (subst sum_list_Basis_list[symmetric])+ apply (subst sum_list_sum_nth)+ by (auto simp add: atLeast0LessThan scaleR_sum_left intro!: sum.cong) finally show ?case by simp qed lemma blinfuns_of_lvivl_eq: "blinfuns_of_lvivl x = (blinfun_of_vmatrix ` set_of_lvivl x::((real, 'b) vec \\<^sub>L (real, 'b) vec) set)" if "length (fst x) = CARD('b::enum)*CARD('b)" "length (snd x) = CARD('b)*CARD('b)" apply (subst set_of_lvivl_eq) subgoal by (simp add: that) subgoal by (simp add: that) unfolding blinfuns_of_lvivl_def image_image by (auto simp: that blinfun_of_list_eq[symmetric] in_list_interval_lengthD cong: image_cong) lemma range_blinfun_of_vmatrix[simp]: "range blinfun_of_vmatrix = UNIV" apply auto apply transfer subgoal for x by (rule image_eqI[where x="matrix x"]) auto done lemma blinfun_of_vmatrix_image: "blinfun_of_vmatrix ` aform.set_of_lvivl' x = (blinfuns_of_lvivl' x::((real, 'b) vec \\<^sub>L (real, 'b) vec) set)" if "aform.lvivl'_invar (CARD('b)*CARD('b::enum)) x" using that by (auto simp: aform.set_of_lvivl'_def blinfuns_of_lvivl'_def blinfuns_of_lvivl_eq aform.lvivl'_invar_def split: option.splits) lemma one_step_result123: "solves_one_step_until_time_aform optns odo X0i t1 t2 E dE \ (x0, d0) \ aform.c1_info_of_appre X0i \ t \ {t1 .. t2} \ set_of_lvivl E \ S \ blinfuns_of_lvivl' dE \ dS \ length (fst E) = CARD('n) \ length (snd E) = CARD('n) \ aform.lvivl'_invar (CARD('n) * CARD('n)) dE \ aform.c1_info_invare DIM('a) X0i \ aform.D odo = DIM('a) \ (t \ existence_ivl0 (x0::'a) \ flow0 x0 t \ S) \ Dflow x0 t o\<^sub>L d0 \ dS" apply (transfer fixing: optns X0i t1 t2 t E dE) subgoal premises prems for x0 d0 S dS proof - have "t \ aform.existence_ivl0 odo x0 \ aform.flow0 odo x0 t \ S \ aform.Dflow odo x0 t o\<^sub>L d0 \ dS" apply (rule one_step_in_ivl[of t t1 t2 x0 d0 X0i "fst E" "snd E" S dE dS odo optns]) using prems by (auto simp: eucl_of_list_prod set_of_lvivl_def set_of_ivl_def blinfun_of_vmatrix_image aform.D_def solves_one_step_until_time_aform_def) with vflow_eq[of t x0] vDflow_eq[of t x0] vex_ivl_eq[symmetric, of t x0] show ?thesis by simp qed done lemmas one_step_result12 = one_step_result123[THEN conjunct1] and one_step_result3 = one_step_result123[THEN conjunct2] lemmas one_step_result1 = one_step_result12[THEN conjunct1] and one_step_result2 = one_step_result12[THEN conjunct2] lemma plane_of_transfer[transfer_rule]: "(rel_sctn A ===> rel_set A) plane_of plane_of" if [transfer_rule]: "(A ===> A ===> (=)) (\) (\)" "bi_total A" unfolding plane_of_def by transfer_prover lemma below_halfspace_transfer[transfer_rule]: "(rel_sctn A ===> rel_set A) below_halfspace below_halfspace" if [transfer_rule]: "(A ===> A ===> (=)) (\) (\)" "bi_total A" unfolding below_halfspace_def le_halfspace_def by transfer_prover definition "rel_nres A a b \ (a, b) \ \{(a, b). A a b}\nres_rel" lemma FAILi_transfer[transfer_rule]: "(rel_nres B) FAILi FAILi" by (auto simp: rel_nres_def nres_rel_def) lemma RES_transfer[transfer_rule]: "(rel_set B ===> rel_nres B) RES RES" by (auto simp: rel_nres_def nres_rel_def rel_set_def intro!: rel_funI RES_refine) context includes autoref_syntax begin lemma RETURN_dres_nres_relI: "(fi, f) \ A \ B \ (\x. dRETURN (fi x), (\x. RETURN (f x))) \ A \ \B\dres_nres_rel" by (auto simp: dres_nres_rel_def dest: fun_relD) end lemma br_transfer[transfer_rule]: "((B ===> C) ===> (B ===> (=)) ===> rel_set (rel_prod B C)) br br" if [transfer_rule]: "bi_total B" "bi_unique C" "bi_total C" unfolding br_def by transfer_prover lemma aform_appr_rel_transfer[transfer_rule]: "(rel_set (rel_prod (list_all2 (=)) (rel_set rel_ve))) aform.appr_rel aform.appr_rel" unfolding aform.appr_rel_br by (transfer_prover) lemma appr1_rel_transfer[transfer_rule]: "(rel_set (rel_prod (rel_prod (list_all2 (=)) (rel_option (list_all2 (=)))) (rel_set (rel_prod rel_ve (rel_blinfun rel_ve rel_ve))))) aform.appr1_rel aform.appr1_rel" unfolding aform.appr1_rel_internal by transfer_prover lemma relAPP_transfer[transfer_rule]: "((rel_set (rel_prod B C) ===> D) ===> rel_set (rel_prod B C) ===> D) Relators.relAPP Relators.relAPP" unfolding relAPP_def by transfer_prover lemma prod_rel_transfer[transfer_rule]: "(rel_set (rel_prod B C) ===> rel_set (rel_prod D E) ===> rel_set (rel_prod (rel_prod B D) (rel_prod C E))) prod_rel prod_rel" if [transfer_rule]: "bi_total B" "bi_unique B" "bi_unique C" "bi_total C" "bi_unique D" "bi_total D" "bi_unique E" "bi_total E" unfolding prod_rel_def_internal by transfer_prover lemma Domain_transfer[transfer_rule]: "(rel_set (rel_prod A B) ===> rel_set A) Domain Domain" if [transfer_rule]: "bi_total A" "bi_unique A" "bi_total B" "bi_unique B" unfolding Domain_unfold by transfer_prover lemma set_rel_transfer[transfer_rule]: "(rel_set (rel_prod B C) ===> rel_set (rel_prod (rel_set B) (rel_set C))) set_rel set_rel" if [transfer_rule]: "bi_total B" "bi_unique B" "bi_unique C" "bi_total C" unfolding set_rel_def_internal by transfer_prover lemma relcomp_transfer[transfer_rule]: "(rel_set (rel_prod B C) ===> rel_set (rel_prod C D) ===> rel_set (rel_prod B D)) relcomp relcomp" if [transfer_rule]: "bi_total B" "bi_unique B" "bi_unique C" "bi_total C" "bi_unique D" "bi_total D" unfolding relcomp_unfold by transfer_prover lemma Union_rel_transfer[transfer_rule]: "(rel_set (rel_prod B (rel_set C)) ===> rel_set (rel_prod C (rel_set D)) ===> rel_set (rel_prod B (rel_set D))) Union_rel Union_rel" if [transfer_rule]: "bi_total B" "bi_unique B" "bi_unique C" "bi_total C" "bi_unique D" "bi_total D" unfolding Union_rel_internal top_fun_def top_bool_def by transfer_prover lemma fun_rel_transfer[transfer_rule]: "(rel_set (rel_prod B C) ===> rel_set (rel_prod D E) ===> rel_set (rel_prod (B ===> D) (C ===> E))) Relators.fun_rel Relators.fun_rel" if [transfer_rule]: "bi_unique B" "bi_unique C" "bi_unique D" "bi_total D" "bi_unique E" "bi_total E" unfolding fun_rel_def_internal by transfer_prover lemma c1_info_of_apprse_transfer[transfer_rule]: "(list_all2 (rel_prod (rel_prod (=) (=)) (rel_prod (list_all2 (=)) (rel_option (list_all2 (=))))) ===> rel_set (rel_prod rel_ve (rel_blinfun rel_ve rel_ve))) aform.c1_info_of_apprse aform.c1_info_of_apprse" unfolding aform.c1_info_of_apprse_def by transfer_prover term scaleR2_rel (* "scaleR2_rel" :: "('b \ ('c \ 'd) set) set \ (((ereal \ ereal) \ 'b) \ ('c \ 'd) set) set" *) lemma scaleR2_rel_transfer[transfer_rule]: "(rel_set (rel_prod (=) (rel_set (rel_prod (=) (=)))) ===> rel_set (rel_prod (rel_prod (rel_prod (=) (=)) (=)) (rel_set (rel_prod (=) (=))))) scaleR2_rel scaleR2_rel" unfolding scaleR2_rel_internal by transfer_prover lemma appr1_rele_transfer[transfer_rule]: "(rel_set (rel_prod (rel_prod (rel_prod (=) (=)) (rel_prod (list_all2 (=)) (rel_option (list_all2 (=))))) (rel_set (rel_prod rel_ve (rel_blinfun rel_ve rel_ve))))) aform.appr1e_rel aform.appr1e_rel" unfolding scaleR2_rel_internal by transfer_prover lemma flow1_of_vec1_times: "flow1_of_vec1 ` (A \ B) = A \ blinfun_of_vmatrix ` B" by (auto simp: flow1_of_vec1_def vec1_of_flow1_def) lemma stable_on_transfer[transfer_rule]: "(rel_set rel_ve ===> rel_set rel_ve ===> (=)) v.stable_on stable_on" unfolding stable_on_def v.stable_on_def by transfer_prover theorem solves_poincare_map_aform: "solves_poincare_map_aform optns odo (\x. dRETURN (symstart x)) [S] guards ivl sctn roi XS RET dRET \ (symstart, symstarta) \ fun_rel (aform.appr1e_rel) (clw_rel aform.appr_rel \\<^sub>r clw_rel aform.appr1e_rel) \ (\X0. (\(CX, X). flowsto (X0 - trap \ UNIV) {0..} (CX \ UNIV) X) (symstarta X0)) \ stable_on (aform.Csafe odo - set_of_lvivl ivl \ plane_of (map_sctn eucl_of_list sctn)) trap \ (\X. X \ set XS \ aform.c1_info_invare DIM('a) X) \ aform.D odo = DIM('a) \ length (normal sctn) = DIM('a) \ length (fst ivl) = DIM('a) \ length (snd ivl) = DIM('a) \ length (normal S) = DIM('a) \ (\a xs b ba ro. (xs, ro) \ set guards \ ((a, b), ba) \ set xs \ length a = DIM('a) \ length b = DIM('a) \ length (normal ba) = DIM('a)) \ length (fst RET) = CARD('n) \ length (snd RET) = CARD('n) \ aform.lvivl'_invar (CARD('n) * CARD('n)) dRET \ poincare_mapsto ((set_of_lvivl ivl::('a set)) \ plane_of (map_sctn eucl_of_list sctn)) (aform.c1_info_of_apprse XS - trap \ UNIV) (below_halfspace (map_sctn eucl_of_list S)) (aform.Csafe odo - set_of_lvivl ivl \ plane_of (map_sctn eucl_of_list sctn)) (set_of_lvivl RET \ blinfuns_of_lvivl' dRET)" apply (transfer fixing: optns symstart S guards ivl sctn roi XS RET dRET) subgoal premises prems for symstarta trap proof - have "aform.poincare_mapsto odo (set_of_lvivl ivl \ plane_of (map_sctn eucl_of_list sctn)) (aform.c1_info_of_apprse XS - trap \ UNIV) (below_halfspace (map_sctn eucl_of_list S)) (aform.Csafe odo - set_of_lvivl ivl \ plane_of (map_sctn eucl_of_list sctn)) (flow1_of_vec1 ` ({eucl_of_list (fst RET)..eucl_of_list (snd RET)} \ aform.set_of_lvivl' dRET))" apply (rule solves_poincare_map[OF _ RETURN_dres_nres_relI RETURN_rule, of optns odo symstart S guards ivl sctn roi XS "fst RET" "snd RET" dRET symstarta trap]) subgoal using prems(1) by (simp add: solves_poincare_map_aform_def) subgoal using prems(2) by (auto simp: fun_rel_def_internal) subgoal for X0 using prems(3)[of X0] vflowsto_eq by auto subgoal unfolding aform.stable_on_def proof (safe, goal_cases) case (1 t x0) from 1 have a: "t \ v.existence_ivl0 x0" using vex_ivl_eq by blast with 1 have b: "v.flow0 x0 t \ trap" using vflow_eq by simp have c: "v.flow0 x0 s \ vX - set_of_lvivl ivl \ plane_of (map_sctn eucl_of_list sctn)" if s: "s \ {0<..t}" for s using 1(4)[rule_format, OF s] apply (subst (asm) vflow_eq) unfolding aform_Csafe_vX[symmetric] using s a by (auto dest!: a.v.ivl_subset_existence_ivl) from a b c show ?case using prems(4)[unfolded v.stable_on_def, rule_format, OF b a 1(3) c] by simp qed subgoal using prems by auto subgoal using prems by (auto simp: aform.D_def) subgoal using prems by auto subgoal using prems by auto subgoal using prems by auto subgoal using prems by auto subgoal using prems by auto subgoal using prems by auto subgoal using prems by auto subgoal using prems by auto done then show ?thesis using vflow_eq vex_ivl_eq vflowsto_eq prems apply (subst vpoincare_mapsto_eq[symmetric]) by (auto simp: set_of_lvivl_def set_of_ivl_def blinfun_of_vmatrix_image flow1_of_vec1_times) qed done theorem solves_poincare_map_aform': "solves_poincare_map_aform' optns odo S guards ivl sctn roi XS RET dRET\ (\X. X \ set XS \ aform.c1_info_invare DIM('a) X) \ aform.D odo = DIM('a) \ length (normal sctn) = DIM('a) \ length (fst ivl) = DIM('a) \ length (snd ivl) = DIM('a) \ length (normal S) = DIM('a) \ (\a xs b ba ro. (xs, ro) \ set guards \ ((a, b), ba) \ set xs \ length a = DIM('a) \ length b = DIM('a) \ length (normal ba) = DIM('a)) \ length (fst RET) = CARD('n) \ length (snd RET) = CARD('n) \ aform.lvivl'_invar (CARD('n) * CARD('n)) dRET \ poincare_mapsto ((set_of_lvivl ivl::('a set)) \ plane_of (map_sctn eucl_of_list sctn)) (aform.c1_info_of_apprse XS) (below_halfspace (map_sctn eucl_of_list S)) (aform.Csafe odo - set_of_lvivl ivl \ plane_of (map_sctn eucl_of_list sctn)) (set_of_lvivl RET \ blinfuns_of_lvivl' dRET)" apply (transfer fixing: optns S guards ivl sctn roi XS RET dRET) subgoal using solves_poincare_map'[of optns odo S guards ivl sctn roi XS "fst RET" "snd RET" dRET] using vflow_eq vex_ivl_eq vflowsto_eq apply (subst vpoincare_mapsto_eq[symmetric]) by (auto intro!: closed_Int simp: set_of_lvivl_def set_of_ivl_def blinfun_of_vmatrix_image flow1_of_vec1_times aform.D_def solves_poincare_map_aform'_def) done theorem solves_poincare_map_onto_aform: "solves_poincare_map_onto_aform optns odo guards ivl sctn roi XS RET dRET\ (\X. X \ set XS \ aform.c1_info_invare DIM('a) X) \ aform.D odo = DIM('a) \ length (normal sctn) = DIM('a) \ length (fst ivl) = DIM('a) \ length (snd ivl) = DIM('a) \ (\a xs b ba ro. (xs, ro) \ set guards \ ((a, b), ba) \ set xs \ length a = DIM('a) \ length b = DIM('a) \ length (normal ba) = DIM('a)) \ length (fst RET) = CARD('n) \ length (snd RET) = CARD('n) \ aform.lvivl'_invar (CARD('n) * CARD('n)) dRET \ poincare_mapsto ((set_of_lvivl ivl::('a set)) \ plane_of (map_sctn eucl_of_list sctn)) (aform.c1_info_of_apprse XS) UNIV (aform.Csafe odo - set_of_lvivl ivl \ plane_of (map_sctn eucl_of_list sctn)) (set_of_lvivl RET \ blinfuns_of_lvivl' dRET)" apply (transfer fixing: optns guards ivl sctn roi XS RET dRET) subgoal using solves_poincare_map_onto[of optns odo guards ivl sctn roi XS "fst RET" "snd RET" dRET, where 'n='n, unfolded aform.poincare_maps_onto_def] using vflow_eq vex_ivl_eq vflowsto_eq apply (subst vpoincare_mapsto_eq[symmetric]) by (auto intro!: closed_Int simp: set_of_lvivl_def set_of_ivl_def blinfun_of_vmatrix_image flow1_of_vec1_times aform.D_def solves_poincare_map_onto_aform_def) done end end subsection \Example Utilities!\ hide_const floatarith.Max floatarith.Min lemma degree_sum_pdevs_scaleR_Basis: "degree (sum_pdevs (\i. pdevs_scaleR (a i) i) (Basis::'b::euclidean_space set)) = Max ((\i. degree (a i)) ` Basis)" apply (rule antisym) subgoal apply (rule degree_le) by (auto ) subgoal apply (rule Max.boundedI) apply simp apply simp apply (auto simp: intro!: degree_leI) by (auto simp: euclidean_eq_iff[where 'a='b]) done lemma Inf_aform_eucl_of_list_aform: assumes "length a = DIM('b::executable_euclidean_space)" shows "Inf_aform (eucl_of_list_aform a::'b aform) = eucl_of_list (map Inf_aform a)" using assms apply (auto simp: eucl_of_list_aform_def Inf_aform_def[abs_def] algebra_simps eucl_of_list_inner inner_sum_left intro!: euclidean_eqI[where 'a='b]) apply (auto simp: tdev_def inner_sum_left abs_inner inner_Basis if_distrib cong: if_cong) apply (rule sum.mono_neutral_cong_left) apply simp by (auto simp: degree_sum_pdevs_scaleR_Basis) lemma Sup_aform_eucl_of_list_aform: assumes "length a = DIM('b::executable_euclidean_space)" shows "Sup_aform (eucl_of_list_aform a::'b aform) = eucl_of_list (map Sup_aform a)" using assms apply (auto simp: eucl_of_list_aform_def Sup_aform_def[abs_def] algebra_simps eucl_of_list_inner inner_sum_left intro!: euclidean_eqI[where 'a='b]) apply (auto simp: tdev_def inner_sum_left abs_inner inner_Basis if_distrib cong: if_cong) apply (rule sum.mono_neutral_cong_right) apply simp by (auto simp: degree_sum_pdevs_scaleR_Basis) lemma eucl_of_list_map_Inf_aform_leI: assumes "x \ Affine (eucl_of_list_aform a::'b::executable_euclidean_space aform)" assumes "length a = DIM('b)" shows "eucl_of_list (map Inf_aform a) \ x" using Inf_aform_le_Affine[OF assms(1)] assms(2) by (auto simp: Inf_aform_eucl_of_list_aform) lemma eucl_of_list_map_Sup_aform_geI: assumes "x \ Affine (eucl_of_list_aform a::'b::executable_euclidean_space aform)" assumes "length a = DIM('b)" shows "x \ eucl_of_list (map Sup_aform a)" using Sup_aform_ge_Affine[OF assms(1)] assms(2) by (auto simp: Sup_aform_eucl_of_list_aform) lemma mem_Joints_appendE: assumes "x \ Joints (xs @ ys)" obtains x1 x2 where "x = x1 @ x2" "x1 \ Joints xs" "x2 \ Joints ys" using assms by (auto simp: Joints_def valuate_def) lemma c1_info_of_appr_subsetI1: fixes X1::"'b::executable_euclidean_space set" assumes subset: "{eucl_of_list (map Inf_aform (fst R)) .. eucl_of_list (map Sup_aform (fst R))} \ X1" assumes len: "length (fst R) = DIM('b)" shows "aform.c1_info_of_appr R \ X1 \ UNIV" using len apply (auto simp: aform.c1_info_of_appr_def flow1_of_list_def split: option.splits intro!: subsetD[OF subset] elim!: mem_Joints_appendE) subgoal by (auto intro!: eucl_of_list_mem_eucl_of_list_aform eucl_of_list_map_Inf_aform_leI eucl_of_list_map_Sup_aform_geI) subgoal by (auto intro!: eucl_of_list_mem_eucl_of_list_aform eucl_of_list_map_Inf_aform_leI eucl_of_list_map_Sup_aform_geI) subgoal apply (subst (2) eucl_of_list_take_DIM[symmetric, OF refl]) apply (auto simp: min_def) apply (simp add: Joints_imp_length_eq eucl_of_list_map_Inf_aform_leI eucl_of_list_mem_eucl_of_list_aform) apply (simp add: Joints_imp_length_eq eucl_of_list_map_Inf_aform_leI eucl_of_list_mem_eucl_of_list_aform) done subgoal apply (subst (2) eucl_of_list_take_DIM[symmetric, OF refl]) apply (auto simp: min_def) by (simp add: Joints_imp_length_eq eucl_of_list_map_Sup_aform_geI eucl_of_list_mem_eucl_of_list_aform) done lemmas [simp] = compute_tdev syntax product_aforms::"(real aform) list \ (real aform) list \ (real aform) list" (infixr "\\<^sub>a" 70) lemma matrix_inner_Basis_list: includes vec_syntax assumes "k < CARD('n) * CARD('m)" shows "(f::(('n::enum rvec, 'm::enum) vec)) \ Basis_list ! k = vec_nth (vec_nth f (enum_class.enum ! (k div CARD('n)))) (enum_class.enum ! (k mod CARD('n)))" proof - have "f \ Basis_list ! k = (\x\UNIV. \xa\UNIV. if enum_class.enum ! (k mod CARD('n)) = xa \ enum_class.enum ! (k div CARD('n)) = x then f $ x $ xa else 0)" using assms unfolding inner_vec_def apply (auto simp: Basis_list_vec_def concat_map_map_index) apply (subst (2) sum.cong[OF refl]) apply (subst sum.cong[OF refl]) apply (subst (2) vec_nth_Basis2) apply (force simp add: Basis_vec_def Basis_list_real_def) apply (rule refl) apply (rule refl) apply (auto simp: if_distribR if_distrib axis_eq_axis Basis_list_real_def cong: if_cong) done also have "\ = f $ enum_class.enum ! (k div CARD('n)) $ enum_class.enum ! (k mod CARD('n))" apply (subst if_conn) apply (subst sum.delta') apply simp by (simp add: sum.delta') finally show ?thesis by simp qed lemma list_of_eucl_matrix: includes vec_syntax shows "(list_of_eucl (M::(('n::enum rvec, 'm::enum) vec))) = concat (map (\i. map (\j. M $ (enum_class.enum ! i)$ (enum_class.enum ! j) ) [0.. real_of_float (lapprox_rat 20 i j))" definition "udec x = (case quotient_of x of (i, j) \ real_of_float (rapprox_rat 20 i j))" lemma ldec: "ldec x \ real_of_rat x" and udec: "real_of_rat x \ udec x" apply (auto simp: ldec_def udec_def split: prod.splits intro!: lapprox_rat[le] rapprox_rat[ge]) apply (metis Fract_of_int_quotient less_eq_real_def less_int_code(1) of_rat_rat quotient_of_denom_pos quotient_of_div) apply (metis Fract_of_int_quotient less_eq_real_def less_int_code(1) of_rat_rat quotient_of_denom_pos quotient_of_div) done context includes floatarith_notation begin definition "matrix_of_degrees\<^sub>e = (let ur = Rad_of (Var 0); vr = Rad_of (Var 1) in [Cos ur, Cos vr, 0, Sin ur, Sin vr, 0, 0, 0, 0])" definition "matrix_of_degrees u v = approx_floatariths 30 matrix_of_degrees\<^sub>e (aforms_of_point ([u, v, 0]))" lemma interpret_floatariths_matrix_of_degrees: "interpret_floatariths matrix_of_degrees\<^sub>e (([u::real, v::real, 0])) = [cos (rad_of u), cos (rad_of v), 0, sin (rad_of u), sin (rad_of v), 0, 0, 0, 0]" by (auto simp: matrix_of_degrees\<^sub>e_def Let_def inverse_eq_divide) definition "num_options p sstep m N a projs print_fun = \ precision = p, adaptive_atol = FloatR 1 (- a), adaptive_rtol = FloatR 1 (- a), method_id = 2, start_stepsize = FloatR 1 (- sstep), iterations = 40, halve_stepsizes = 40, widening_mod = 10, rk2_param = FloatR 1 0, default_reduce = correct_girard (p) (m) (N), printing_fun = (\a b. let _ = fold (\(x, y, c) _. print_fun (String.implode (shows_segments_of_aform (x) (y) b c ''\''))) projs (); _ = print_fun (String.implode (''# '' @ shows_box_of_aforms_hr (b) '''' @ ''\'')) in () ), tracing_fun = (\a b. let _ = print_fun (String.implode (''# '' @ a @ ''\'')) in case b of Some b \ (let _ = () in print_fun (String.implode (''# '' @ shows_box_of_aforms_hr (b) '''' @ ''\''))) | None \ ()) \" definition "num_options_c1 p sstep m N a projs dcolors print_fun = (let no = num_options p sstep m N a (map (\(x, y, c, ds). (x, y, c)) projs) print_fun; D = length dcolors in no \printing_fun:= (\a b. let _ = printing_fun no a b in if a then () else fold (\(x, y, c, ds) _. print_fun (String.implode (shows_aforms_vareq D [(x, y)] ds dcolors dcolors (FloatR 1 (-1)) ''# no C1 info'' b ''''))) projs () ) \)" definition "num_options_code p sstep m N a projs print_fun = num_options (nat_of_integer p) (int_of_integer sstep) (nat_of_integer m) (nat_of_integer N) (int_of_integer a) (map (\(i, j, k). (nat_of_integer i, nat_of_integer j, k)) projs) print_fun" definition "ro s n M g0 g1 inter_step = \max_tdev_thres = FloatR 1 s, pre_split_reduce = correct_girard 30 n M, pre_inter_granularity = FloatR 1 g0, post_inter_granularity = (FloatR 1 g1), pre_collect_granularity = FloatR 1 g0, max_intersection_step = FloatR 1 inter_step\" definition "code_ro s n m g0 g1 inter_step = ro (int_of_integer s) (nat_of_integer n) (nat_of_integer m) (int_of_integer g0) (int_of_integer g1) (int_of_integer inter_step)" fun xsec:: "real \ real \ real \ real \ real \ (real list \ real list) \ real list sctn" where "xsec x (y0, y1) (z0, z1) = (([x, y0, z0], [x, y1, z1]), Sctn [1,0,0] x)" fun xsec':: "real \ real \ real \ real \ real \ (real list \ real list) \ real list sctn" where "xsec' x (y0, y1) (z0, z1) = (([x, y0, z0], [x, y1, z1]), Sctn [-1,0,0] (-x))" fun ysec:: "real \ real \ real \ real \ real \ (real list \ real list) \ real list sctn" where "ysec (x0, x1) y (z0, z1) = (([x0, y, z0], [x1, y, z1]), Sctn [0, 1,0] y)" fun ysec':: "real \ real \ real \ real \ real \ (real list \ real list) \ real list sctn" where "ysec' (x0, x1) y (z0, z1) = (([x0, y, z0], [x1, y, z1]), Sctn [0, -1,0] (-y))" fun zsec:: "real \ real \ real \ real \ real \ (real list \ real list) \ real list sctn" where "zsec (x0, x1) (y0, y1) z = (([x0, y0, z], [x1, y1, z]), Sctn [0, 0, 1] z)" fun zsec':: "real \ real \ real \ real \ real \ (real list \ real list) \ real list sctn" where "zsec' (x0, x1) (y0, y1) z = (([x0, y0, z], [x1, y1, z]), Sctn [0, 0, -1] (-z))" fun xsec2:: "real \ real \ real \ (real list \ real list) \ real list sctn" where "xsec2 x (y0, y1) = (([x, y0], [x, y1]), Sctn [1,0] x)" fun xsec2':: "real \ real \ real \(real list \ real list) \ real list sctn" where "xsec2' x (y0, y1) = (([x, y0], [x, y1]), Sctn [-1,0] (-x))" fun ysec2:: "real \ real \ real \ (real list \ real list) \ real list sctn" where "ysec2 (x0, x1) y = (([x0, y], [x1, y]), Sctn [0, 1] y)" fun ysec2':: "real \ real \ real \ (real list \ real list) \ real list sctn" where "ysec2' (x0, x1) y = (([x0, y], [x1, y]), Sctn [0, -1] (-y))" fun ysec4:: "real \ real \ real \ real \ real \ real \ real \ (real list \ real list) \ real list sctn" where "ysec4 (x0, x1) y (z0, z1) (m0, m1) = (([x0, y, z0, m0], [x1, y, z1, m1]), Sctn [0, 1,0, 0] (y))" fun ysec4':: "real \ real \ real \ real \ real \ real \ real \ (real list \ real list) \ real list sctn" where "ysec4' (x0, x1) y (z0, z1) (m0, m1) = (([x0, y, z0, m0], [x1, y, z1, m1]), Sctn [0, -1,0, 0] (-y))" definition "code_sctn N n c = Sctn ((replicate (nat_of_integer N) (0::real))[nat_of_integer n := 1]) c" definition "code_sctn' N n c = Sctn ((replicate (nat_of_integer N) (0::real))[nat_of_integer n := -1]) (-c)" definition "lrat i j = real_of_float (lapprox_rat 30 (int_of_integer i) (int_of_integer j))" definition "urat i j = real_of_float (lapprox_rat 30 (int_of_integer i) (int_of_integer j))" definition [simp]: "TAG_optns (optns::string \ ((String.literal \ unit) \ (real aform) numeric_options)) = True" lemma TAG_optns: "P \ (TAG_optns optns \ P)" by (auto simp: ) definition [simp]: "TAG_reach_optns (roi::real aform reach_options) = True" lemma TAG_reach_optns: "P \ (TAG_reach_optns optns \ P)" by (auto simp: ) definition [simp]: "TAG_sctn (b::bool) = True" lemma TAG_sctn: "P \ (TAG_sctn optns \ P)" by (auto simp: ) subsection \Integrals and Computation\ lemma has_vderiv_on_PairD: assumes "((\t. (f t, g t)) has_vderiv_on fg') T" shows "(f has_vderiv_on (\t. fst (fg' t))) T" "(g has_vderiv_on (\t. snd (fg' t))) T" proof - from assms have "((\x. (f x, g x)) has_derivative (\xa. xa *\<^sub>R fg' t)) (at t within T)" if "t \ T" for t by (auto simp: has_vderiv_on_def has_vector_derivative_def that intro!: derivative_eq_intros) from diff_chain_within[OF this has_derivative_fst[OF has_derivative_ident]] diff_chain_within[OF this has_derivative_snd[OF has_derivative_ident]] show "(f has_vderiv_on (\t. fst (fg' t))) T" "(g has_vderiv_on (\t. snd (fg' t))) T" by (auto simp: has_vderiv_on_def has_vector_derivative_def o_def) qed lemma solves_autonomous_odeI: assumes "((\t. (t, phi t)) solves_ode (\t x. (1, f (fst x) (snd x)))) S (T \ X)" shows "(phi solves_ode f) S X" proof (rule solves_odeI) from solves_odeD[OF assms] have "((\t. (t, phi t)) has_vderiv_on (\t. (1, f (fst (t, phi t)) (snd (t, phi t))))) S" "\t. t \ S \ (phi t) \ X" by (auto simp: ) from has_vderiv_on_PairD(2)[OF this(1)] this(2) show "(phi has_vderiv_on (\t. f t (phi t))) S" "\t. t \ S \ phi t \ X" by auto qed lemma integral_solves_autonomous_odeI: fixes f::"real \ 'a::executable_euclidean_space" assumes "(phi solves_ode (\t _. f t)) {a .. b} X" "phi a = 0" assumes "a \ b" shows "(f has_integral phi b) {a .. b}" proof - have "(f has_integral phi b - phi a) {a..b}" apply (rule fundamental_theorem_of_calculus[of a b phi f]) unfolding has_vderiv_on_def[symmetric] apply fact using solves_odeD[OF assms(1)] by (simp add: has_vderiv_on_def) then show ?thesis by (simp add: assms) qed lemma zero_eq_eucl_of_list_rep_DIM: "(0::'a::executable_euclidean_space) = eucl_of_list (replicate (DIM('a)) 0)" by (auto intro!: euclidean_eqI[where 'a='a] simp: eucl_of_list_inner) lemma zero_eq_eucl_of_list_rep: "(0::'a::executable_euclidean_space) = eucl_of_list (replicate D 0)" if "D \ DIM('a)" proof - from that have "replicate D (0::real) = replicate (DIM('a)) 0 @ replicate (D - DIM('a)) 0" by (auto simp: replicate_add[symmetric]) also have "eucl_of_list \ = (eucl_of_list (replicate DIM('a) 0)::'a)" by (rule eucl_of_list_append_zeroes) also have "\ = 0" by (rule zero_eq_eucl_of_list_rep_DIM[symmetric]) finally show ?thesis by simp qed lemma one_has_ivl_integral: "((\x. 1::real) has_ivl_integral b - a) a b" using has_integral_const_real[of "1::real" a b] has_integral_const_real[of "1::real" b a] by (auto simp: has_ivl_integral_def split: if_splits) lemma Joints_aforms_of_point_self[simp]: "xs \ Joints (aforms_of_point xs)" by (simp add: aforms_of_point_def) lemma bind_eq_dRETURN_conv: "(f \ g = dRETURN S) \ (\R. f = dRETURN R \ g R = dRETURN S)" by (cases f) auto end lemma list_of_eucl_memI: "list_of_eucl (x::'x::executable_euclidean_space) \ S" if "x \ eucl_of_list ` S" "\x. x \ S \ length x = DIM('x)" using that by auto lemma Joints_aforms_of_ivls_append_point: "Joints (xs @ aforms_of_ivls p p) = (\x. x @ p) ` Joints xs" using aform.set_of_appr_of_ivl_append_point[unfolded aform_ops_def approximate_set_ops.simps] . context ode_interpretation begin theorem solves_one_step_ivl: assumes T: "T \ {t1 .. t2}" assumes X: "X \ {eucl_of_list lx .. eucl_of_list ux}" "length lx = DIM('a)" "length ux = DIM('a)" assumes S: "{eucl_of_list ls::'a .. eucl_of_list us} \ S" assumes lens: "length ls = DIM('a)" "length us = DIM('a)" \ \TODO: this could be verified\ assumes [simp]: "aform.D odo = DIM('a)" assumes r: "solves_one_step_until_time_aform optns odo ((1,1), aforms_of_ivls lx ux, None) t1 t2 (ls, us) None" shows "t \ T \ x0 \ X \ t \ existence_ivl0 x0 \ flow0 x0 t \ S" proof (intro impI) assume t: "t \ T" and x0: "x0 \ X" from S have S: "set_of_lvivl (ls, us) \ S" by (auto simp: set_of_lvivl_def set_of_ivl_def) have lens: "length (fst (ls, us)) = CARD('n)" "length (snd (ls, us)) = CARD('n)" by (auto simp: lens) have x0: "list_of_eucl x0 \ Joints (aforms_of_ivls lx ux)" apply (rule aforms_of_ivls) subgoal by (simp add: X) subgoal by (simp add: X) using subsetD[OF X(1) x0] apply (auto simp: eucl_le[where 'a='a] X) apply (metis assms(3) dim length_Basis_list list_of_eucl_eucl_of_list list_of_eucl_nth nth_Basis_list_in_Basis) apply (metis assms(4) dim length_Basis_list list_of_eucl_eucl_of_list list_of_eucl_nth nth_Basis_list_in_Basis) done from t T have t: "t \ {t1 .. t2}" by auto show "t \ existence_ivl0 x0 \ flow0 x0 t \ S" by (rule one_step_result12[OF r aform.c1_info_of_appre_c0_I[OF x0] t S subset_UNIV lens]) (auto simp: aform.c1_info_invare_None lens X) qed theorem solves_one_step_ivl': assumes T: "T \ {t1 .. t2}" assumes X: "X \ {eucl_of_list lx .. eucl_of_list ux}" "length lx = DIM('a)" "length ux = DIM('a)" assumes DS: "list_interval lds uds \ list_interval ld ud" and lends: "length lds = DIM('a)*DIM('a)" "length uds = DIM('a)*DIM('a)" assumes S: "{eucl_of_list ls::'a .. eucl_of_list us} \ S" assumes lens0: "length ls = DIM('a)" "length us = DIM('a)" \ \TODO: this could be verified\ "length dx0s = DIM('a)*DIM('a)" assumes [simp]: "aform.D odo = DIM('a)" assumes r: "solves_one_step_until_time_aform optns odo ((1,1), aforms_of_ivls lx ux, Some (aforms_of_point dx0s)) t1 t2 (ls, us) (Some (lds, uds))" shows "t \ T \ x0 \ X \ t \ existence_ivl0 x0 \ flow0 x0 t \ S \ Dflow x0 t o\<^sub>L blinfun_of_list dx0s \ blinfuns_of_lvivl (ld, ud)" proof (intro impI) assume t: "t \ T" and x0: "x0 \ X" from S have S: "set_of_lvivl (ls, us) \ S" by (auto simp: set_of_lvivl_def set_of_ivl_def) have lens: "length (fst (ls, us)) = CARD('n)" "length (snd (ls, us)) = CARD('n)" by (auto simp: lens0) have x0: "list_of_eucl x0 \ Joints (aforms_of_ivls lx ux)" apply (rule aforms_of_ivls) subgoal by (simp add: X) subgoal by (simp add: X) using subsetD[OF X(1) x0] apply (auto simp: eucl_le[where 'a='a] X) apply (metis assms(3) dim length_Basis_list list_of_eucl_eucl_of_list list_of_eucl_nth nth_Basis_list_in_Basis) apply (metis assms(4) dim length_Basis_list list_of_eucl_eucl_of_list list_of_eucl_nth nth_Basis_list_in_Basis) done have x0dx0: "(x0, blinfun_of_list dx0s) \ aform.c1_info_of_appre ((1, 1), aforms_of_ivls lx ux, Some (aforms_of_point dx0s))" apply (auto simp: aform.c1_info_of_appre_def aform.c1_info_of_appr_def) apply (rule image_eqI[where x="list_of_eucl x0@dx0s"]) using lens0 apply (auto simp: flow1_of_list_def aforms_of_point_def Joints_aforms_of_ivls_append_point) apply (rule imageI) apply (rule x0) done from t T have t: "t \ {t1 .. t2}" by auto have DS: "blinfuns_of_lvivl' (Some (lds, uds)) \ blinfun_of_list ` list_interval ld ud" using DS by (auto simp: blinfuns_of_lvivl'_def blinfuns_of_lvivl_def) have inv: "aform.lvivl'_invar (CARD('n) * CARD('n)) (Some (lds, uds))" "aform.c1_info_invare DIM('a) ((1::ereal, 1), aforms_of_ivls lx ux, Some (aforms_of_point dx0s))" by (auto simp: aform.lvivl'_invar_def lends aform.c1_info_invare_def X lens0 power2_eq_square aform.c1_info_invar_def) from one_step_result123[OF r x0dx0 t S DS lens inv \aform.D _ = _\] show "t \ existence_ivl0 x0 \ flow0 x0 t \ S \ Dflow x0 t o\<^sub>L blinfun_of_list dx0s \ blinfuns_of_lvivl (ld, ud)" by (auto simp: blinfuns_of_lvivl_def) qed end definition "zero_aforms D = map (\_. (0, zero_pdevs)) [0..pf. solves_one_step_until_time_aform (snd soptns pf) a b c d e f)" definition "solves_poincare_map_aform'_fo soptns a b c d e f g h i = file_output (String.implode (fst soptns)) (\pf. solves_poincare_map_aform' (snd soptns pf) a b c d e f g h i)" definition "solves_poincare_map_onto_aform_fo soptns a b c d e f g h = file_output (String.implode (fst soptns)) (\pf. solves_poincare_map_onto_aform (snd soptns pf) a b c d e f g h)" lemma solves_one_step_until_time_aform_foI: "solves_one_step_until_time_aform (snd optns (\_. ())) a b c d e f" if "solves_one_step_until_time_aform_fo optns a b c d e f" using that by (auto simp: solves_one_step_until_time_aform_fo_def file_output_def Print.file_output_def print_def[abs_def] split: if_splits) lemma solves_poincare_map_aform'_foI: "solves_poincare_map_aform' (snd optns (\_. ())) a b c d e f g h i" if "solves_poincare_map_aform'_fo optns a b c d e f g h i" using that by (auto simp: solves_poincare_map_aform'_fo_def file_output_def Print.file_output_def print_def[abs_def] split: if_splits) lemma solves_poincare_map_onto_aform_foI: "solves_poincare_map_onto_aform (snd optns (\_. ())) a b c d e f g h" if "solves_poincare_map_onto_aform_fo optns a b c d e f g h" using that by (auto simp: solves_poincare_map_onto_aform_fo_def file_output_def Print.file_output_def print_def[abs_def] split: if_splits) definition "can_mk_ode_ops fas safe_form \ mk_ode_ops fas safe_form \ None" theorem solve_one_step_until_time_aform_integral_bounds: fixes f::"real \ 'a::executable_euclidean_space" assumes "a \ b" assumes ba: "b - a \ {t1 .. t2}" assumes a: "a \ {a1 .. a2}" assumes ls_us_subset: "{eucl_of_list ls .. eucl_of_list us} \ {l .. u}" assumes fas: "\xs::real list. length xs > 0 \ interpret_form safe_form xs \ (1::real, f (xs ! 0)) = einterpret fas xs" assumes D: "D = DIM('a) + 1" "D = CARD('i::enum)" assumes lenlu: "length ls + 1 = D" "length us + 1 = D" assumes lfas: "length fas = D" assumes mv: "can_mk_ode_ops fas safe_form" assumes FDERIV: "\xs. interpret_form safe_form xs \ isFDERIV (length fas) [0.. {l .. u}" proof - have lens0: "length ((x::real) # replicate (D - 1) 0) = DIM(real \ 'a)" for x using assms by auto have a0: "(a, 0) \ {eucl_of_list (a1 # replicate (D - 1) 0)..eucl_of_list (a2 # replicate (D - 1) 0)}" using assms by (auto simp: eucl_of_list_prod) let ?U = "{x::real\'a. interpret_form safe_form (list_of_eucl x)}" interpret ode_interpretation safe_form ?U fas "\x. (1::real, f (fst x))" "undefined::'i" apply unfold_locales subgoal using assms by simp subgoal using assms by simp subgoal using mv by (simp add: D lfas) subgoal for x apply (cases x) by (rule HOL.trans[OF fas[symmetric]]) (auto simp: fas) subgoal using mv by (simp add: can_mk_ode_ops_def) subgoal by (rule FDERIV) done have lens: "length (0 # ls) = DIM(real \ 'a)" "length (t2 # us) = DIM(real \ 'a)" "aform.D odo = DIM(real \ 'a)" using lenlu by (simp_all add: lfas aform.D_def D aform.ode_e_def ) have D_odo: "aform.D odo = DIM(real \ 'a)" by (auto simp: aform.D_def aform.ode_e_def lfas D) from solves_one_step_ivl[rule_format, OF order_refl order_refl lens0 lens0 order_refl lens(1,2) D_odo, unfolded odo_def, OF sos ba a0] have lsus: "flow0 (a, 0) (b - a) \ {eucl_of_list (0#ls)..eucl_of_list (t2#us)}" and exivl: "b - a \ existence_ivl0 (a, 0)" by auto have flow: "flow0 (a, 0) (b - a) \ UNIV \ {l..u}" using lsus apply (rule rev_subsetD) using ls_us_subset by (auto simp: eucl_of_list_prod) from ivl_subset_existence_ivl[OF exivl] \a \ b\ exivl have "0 \ existence_ivl0 (a, 0)" by (auto simp del: existence_ivl_initial_time_iff) from mem_existence_ivl_iv_defined(2)[OF this] have safe: "(a, 0::'a) \ ?U" by simp from flow_solves_ode[OF UNIV_I this] have fs: "((\t. (fst (flow0 (a, 0) t), snd (flow0 (a, 0) t))) solves_ode (\_ x. (1, f (fst x)))) (existence_ivl0 (a, 0)) ?U" by simp with solves_odeD[OF fs] have vdp: "((\t. (fst (flow0 (a, 0) t), snd (flow0 (a, 0) t))) has_vderiv_on (\t. (1, f (fst (flow0 (a, 0) t))))) (existence_ivl0 (a, 0))" by simp have "fst (flow0 (a, 0) t) = a + t" if "t \ existence_ivl0 (a, 0)" for t proof - have "fst (flow0 (a, 0) 0) = a" using safe by (auto simp: ) have "((\t. fst (flow0 (a, 0) t)) has_vderiv_on (\x. 1)) (existence_ivl0 (a, 0))" using has_vderiv_on_PairD[OF vdp] by auto then have "((\t. fst (flow0 (a, 0) t)) has_vderiv_on (\x. 1)) {0--t}" apply (rule has_vderiv_on_subset) using closed_segment_subset_existence_ivl[OF that] by auto from fundamental_theorem_of_calculus_ivl_integral[OF this, THEN ivl_integral_unique] one_has_ivl_integral[of t 0, THEN ivl_integral_unique] safe show ?thesis by auto qed with vdp have "((\t. (t, snd (flow0 (a, 0) t))) solves_ode (\t x. (1, f (a + fst x)))) (existence_ivl0 (a, 0)) ((existence_ivl0 (a, 0)) \ UNIV)" apply (intro solves_odeI) apply auto apply (auto simp: has_vderiv_on_def has_vector_derivative_def) proof goal_cases case (1 x) have r: "((\(x, y). (x - a, y::'a)) has_derivative (\x. x)) (at x within t)" for x t by (auto intro!: derivative_eq_intros) from 1 have "((\x. (a + x, snd (flow0 (a, 0) x))) has_derivative (\xa. (xa, xa *\<^sub>R f (a + x)))) (at x within existence_ivl0 (a, 0))" by auto from has_derivative_in_compose2[OF r subset_UNIV _ this, simplified] 1 have "((\x. (x, snd (flow0 (a, 0) x))) has_derivative (\y. (y, y *\<^sub>R f (a + x)))) (at x within existence_ivl0 (a, 0))" by auto then show ?case by simp qed from solves_autonomous_odeI[OF this] have "((\t. snd (flow0 (a, 0) t)) solves_ode (\b c. f (a + b))) (existence_ivl0 (a, 0)) UNIV" by simp \ \TODO: do non-autonomous -- autonomous conversion automatically!\ then have "((\t. snd (flow0 (a, 0) t)) solves_ode (\b c. f (a + b))) {0 .. b - a} UNIV" apply (rule solves_ode_on_subset) using exivl by (rule ivl_subset_existence_ivl) (rule order_refl) from integral_solves_autonomous_odeI[OF this] have "((\b. f (a + b)) has_integral snd (flow0 (a, 0) (b - a))) (cbox 0 (b - a))" using \a \ b\ safe by auto from has_integral_affinity[OF this, where m=1 and c="-a"] have "(f has_integral snd (flow0 (a, 0) (b - a))) {a..b}" by auto then have "integral {a..b} f = snd (flow0 (a, 0) (b - a))" by blast also have "\ \ {l .. u}" using flow by auto finally show ?thesis by simp qed lemma [code_computation_unfold]: "numeral x = real_of_int (Code_Target_Int.positive x)" by simp lemma [code_computation_unfold]: "numeral x \ Float (Code_Target_Int.positive x) 0" by (simp add: Float_def float_of_numeral) definition no_print::"String.literal \ unit" where "no_print x = ()" lemmas [approximation_preproc] = list_of_eucl_real list_of_eucl_prod append.simps named_theorems DIM_simps lemmas [DIM_simps] = DIM_real DIM_prod length_nth_simps add_numeral_special add_numeral_special card_sum card_prod card_bit0 card_bit1 card_num0 card_num1 numeral_times_numeral numeral_mult mult_1_right mult_1 aform.D_def lemma numeral_refl: "numeral x = numeral x" by simp lemma plain_floatarith_approx_eq_SomeD: "approx prec fa [] = Some (the (approx prec fa []))" if "plain_floatarith 0 fa" using that by (auto dest!: plain_floatarith_approx_not_None[where p=prec and XS=Nil]) definition [simp]: "approx1 p f xs = real_of_float (lower (the (approx p f xs)))" definition [simp]: "approx2 p f xs = real_of_float (upper (the (approx p f xs)))" definition [simp]: "approx_defined p f xs \ approx p f xs \ None" definition "approxs p fs xs = those (map (\f. approx p f xs) fs)" definition [simp]: "approxs1 p f xs = (case approxs p f xs of Some y \ (map (real_of_float o lower) y) | None \ replicate (length f) 0)" definition [simp]: "approxs2 p f xs = (case approxs p f xs of Some y \ (map (real_of_float o upper) y) | None \ replicate (length f) 0)" definition "approxs_defined p fs xs \ (those (map (\f. approx p f xs) fs) \ None)" lemma length_approxs: "length (approxs1 p f xs) = length f" "length (approxs2 p f xs) = length f" by (auto simp: approxs_def dest!: those_eq_SomeD split: option.splits) lemma real_in_approxI: "x \ {(approx1 prec fa []) .. (approx2 prec fa [])}" if "x = interpret_floatarith fa []" "approx_defined prec fa []" using that by (force dest: approx_emptyD simp: set_of_eq) lemma real_subset_approxI: "{a .. b} \ {(approx1 prec fa []) .. (approx2 prec fb [])}" if "a = interpret_floatarith fa []" "b = interpret_floatarith fb []" "approx_defined prec fa []" "approx_defined prec fb []" using that by (force dest: approx_emptyD simp: set_of_eq) lemma approxs_eq_Some_lengthD: "length ys = length fas" if "approxs prec fas XS = Some ys" using that by (auto simp: approxs_def dest!: those_eq_SomeD) lemma approxs_pointwise: "interpret_floatarith (fas ! ia) xs \ {real_of_float (lower (ys ! ia)) .. (upper (ys ! ia))}" if "approxs prec fas XS = Some ys" "bounded_by xs XS" "ia < length fas" proof - from those_eq_SomeD[OF that(1)[unfolded approxs_def]] have ys: "ys = map (the \ (\f. approx prec f XS)) fas" and ex: "\y. i < length fas \ approx prec (fas ! i) XS = Some y" for i by auto from ex[of ia] that obtain ivl where ivl: "approx prec (fas ! ia) XS = Some ivl" by auto from approx[OF that(2) this] have "interpret_floatarith (fas ! ia) xs \\<^sub>r ivl" by auto moreover have "ys ! ia = ivl" unfolding ys apply (auto simp: o_def) apply (subst nth_map) apply (simp add: that) using ivl by simp ultimately show ?thesis using that by (auto simp: approxs_eq_Some_lengthD set_of_eq split: prod.splits) qed lemmas approxs_pointwise_le = approxs_pointwise[simplified, THEN conjunct1] and approxs_pointwise_ge = approxs_pointwise[simplified, THEN conjunct2] lemma approxs_eucl: "eucl_of_list (interpret_floatariths fas xs) \ {eucl_of_list (map lower ys) .. eucl_of_list (map upper ys)::'a::executable_euclidean_space}" if "approxs prec fas XS = Some ys" "length fas = DIM('a)" "bounded_by xs XS" using that by (auto simp: eucl_le[where 'a='a] eucl_of_list_inner o_def approxs_eq_Some_lengthD intro!: approxs_pointwise_le approxs_pointwise_ge) lemma plain_floatariths_approx_eq_SomeD: "approxs prec fas [] = Some (the (approxs prec fas []))" if "list_all (plain_floatarith 0) fas" using that apply (induction fas) apply (auto simp: approxs_def split: option.splits dest!: plain_floatarith_approx_eq_SomeD) subgoal for a fas aa apply (cases "those (map (\f. approx prec f []) fas)") by auto done lemma approxs_definedD: "approxs prec fas xs = Some (the (approxs prec fas xs))" if "approxs_defined prec fas xs" using that by (auto simp: approxs_defined_def approxs_def) lemma approxs_defined_ne_None[simp]: "approxs prec fas xs \ None" if "approxs_defined prec fas xs" using that by (auto simp: approxs_defined_def approxs_def) lemma approx_subset_euclI: "{eucl_of_list (approxs2 prec fals [])::'a .. eucl_of_list (approxs1 prec faus [])} \ {l .. u}" if "list_of_eucl l = interpret_floatariths fals []" and "list_of_eucl u = interpret_floatariths faus []" and "length fals = DIM('a::executable_euclidean_space)" and "length faus = DIM('a::executable_euclidean_space)" and "approxs_defined prec fals []" and "approxs_defined prec faus []" using that by (auto intro!: bounded_by_Nil dest!: approxs_eucl[where 'a='a] list_of_eucl_eqD plain_floatariths_approx_eq_SomeD[where prec=prec] split: option.splits) lemma eucl_subset_approxI: "{l .. u} \ {eucl_of_list (approxs1 prec fals [])::'a .. eucl_of_list (approxs2 prec faus [])}" if "list_of_eucl l = interpret_floatariths fals []" and "list_of_eucl u = interpret_floatariths faus []" and "length fals = DIM('a::executable_euclidean_space)" and "length faus = DIM('a::executable_euclidean_space)" and "approxs_defined prec fals []" and "approxs_defined prec faus []" using that by (auto intro!: bounded_by_Nil dest!: approxs_eucl[where 'a='a] list_of_eucl_eqD plain_floatariths_approx_eq_SomeD[where prec=prec] split: option.splits) lemma approx_subset_listI: "list_interval (approxs2 prec fals []) (approxs1 prec faus []) \ list_interval l u" if "l = interpret_floatariths fals []" and "u = interpret_floatariths faus []" and "length fals = length l" and "length faus = length u" and "approxs_defined prec fals []" and "approxs_defined prec faus []" using that apply (auto simp: list_interval_def list_all2_conv_all_nth dest: approxs_eq_Some_lengthD intro!: bounded_by_Nil dest!: plain_floatariths_approx_eq_SomeD[where prec=prec] split: option.splits) apply (rule order_trans) apply (rule approxs_pointwise_ge) apply assumption apply (rule bounded_by_Nil) apply (auto dest: approxs_eq_Some_lengthD) apply (subst interpret_floatariths_nth) apply (auto dest: approxs_eq_Some_lengthD) apply (rule approxs_pointwise_le[ge]) apply assumption apply (rule bounded_by_Nil) apply (auto dest: approxs_eq_Some_lengthD) done definition "unit_list D n = (replicate D 0)[n:=1]" definition "mirror_sctn (sctn::real list sctn) = Sctn (map uminus (normal sctn)) (- pstn sctn)" definition "mirrored_sctn b (sctn::real list sctn) = (if b then mirror_sctn sctn else sctn)" lemma mirror_sctn_simps[simp]: "pstn (mirror_sctn sctn) = - pstn sctn" "normal (mirror_sctn sctn) = map uminus (normal sctn)" by (cases sctn) (auto simp: mirror_sctn_def) lemma length_unit_list[simp]: "length (unit_list D n) = D" by (auto simp: unit_list_def) lemma eucl_of_list_unit_list[simp]: "(eucl_of_list (unit_list D n)::'a::executable_euclidean_space) = Basis_list ! n" if "D = DIM('a)" "n < D" using that by (auto simp: unit_list_def eucl_of_list_inner inner_Basis nth_list_update' intro!: euclidean_eqI[where 'a='a]) lemma le_eucl_of_list_iff: "(t::'a::executable_euclidean_space) \ eucl_of_list uX0 \ (\i. i < DIM('a) \ t \ Basis_list ! i \ uX0 ! i)" if "length uX0 = DIM('a)" using that apply (auto simp: eucl_le[where 'a='a] eucl_of_list_inner) using nth_Basis_list_in_Basis apply force by (metis Basis_list in_Basis_index_Basis_list index_less_size_conv length_Basis_list) lemma eucl_of_list_le_iff: "eucl_of_list uX0 \ (t::'a::executable_euclidean_space) \ (\i. i < DIM('a) \ uX0 ! i \ t \ Basis_list ! i)" if "length uX0 = DIM('a)" using that apply (auto simp: eucl_le[where 'a='a] eucl_of_list_inner) using nth_Basis_list_in_Basis apply force by (metis Basis_list in_Basis_index_Basis_list index_less_size_conv length_Basis_list) lemma Joints_aforms_of_ivls: "Joints (aforms_of_ivls lX0 uX0) = list_interval lX0 uX0" if "list_all2 (\) lX0 uX0" using that apply (auto simp: list_interval_def dest: Joints_aforms_of_ivlsD1[OF _ that] Joints_aforms_of_ivlsD2[OF _ that] list_all2_lengthD intro!: aforms_of_ivls) by (auto simp: list_all2_conv_all_nth) lemma list_of_eucl_in_list_interval_iff: "list_of_eucl x0 \ list_interval lX0 uX0 \ x0 \ {eucl_of_list lX0 .. eucl_of_list uX0::'a}" if "length lX0 = DIM('a::executable_euclidean_space)" "length uX0 = DIM('a::executable_euclidean_space)" using that by (auto simp: list_interval_def eucl_of_list_le_iff le_eucl_of_list_iff list_all2_conv_all_nth) text \TODO: make a tactic out of this?!\ lemma file_output_iff: "file_output s f = f (\_. ())" by (auto simp: file_output_def print_def[abs_def] Print.file_output_def) context ode_interpretation begin lemma poincare_mapsto_subset: "poincare_mapsto P X0 SS CX R" if "poincare_mapsto P' Y0 RR CZ S" "X0 \ Y0" "CZ \ CX" "S \ R" "RR = SS" "P = P'" using that by (force simp: poincare_mapsto_def) theorem solves_poincare_map_aform'_derivI: assumes solves: "solves_poincare_map_aform'_fo optns odo (Sctn (unit_list D n) (lP ! n)) guards (lP, uP) (Sctn (unit_list D n) (lP ! n)) roi [((1,1), aforms_of_ivls lX0 uX0, Some (aforms_of_point DX0))] (lR, uR) (Some (lDR, uDR))" and D: "D = DIM('a)" assumes DS: "list_interval lDR uDR \ list_interval lDS uDS" and dim: "aform.D odo = DIM('a)" and lens: "length (lP) = DIM('a)" "length (uP) = DIM('a)" "length (lX0) = DIM('a)" "length (uX0) = DIM('a)" "length (lR) = DIM('a)" "length (uR) = DIM('a)" "length DX0 = DIM('a)*DIM('a)" "length lDR = CARD('n) * CARD('n)" "length uDR = CARD('n) * CARD('n)" and guards: "(\a xs b ba ro. (xs, ro) \ set guards \ ((a, b), ba) \ set xs \ length a = DIM('a) \ length b = DIM('a) \ length (normal ba) = DIM('a))" and P: "P = {eucl_of_list lP .. eucl_of_list uP}" and plane: "uP ! n = lP ! n" and X0: "X0 \ {eucl_of_list lX0 .. eucl_of_list uX0}" and nD: "n < DIM('a)" and SS: "SS = {x::'a. x \ Basis_list ! n \ lP ! n}" and R: "{eucl_of_list lR .. eucl_of_list uR} \ R" shows "\x\X0. returns_to P x \ return_time P differentiable at x within SS \ (\D. (poincare_map P has_derivative blinfun_apply D) (at x within SS) \ poincare_map P x \ R \ D o\<^sub>L blinfun_of_list DX0 \ blinfuns_of_lvivl (lDS, uDS))" proof (rule ballI) fix x assume "x \ X0" then have la2: "list_all2 (\) lX0 uX0" using X0 by (force simp: subset_iff eucl_of_list_le_iff le_eucl_of_list_iff lens list_all2_conv_all_nth) have 1: "\X. X \ set [((1::ereal, 1::ereal), aforms_of_ivls lX0 uX0, Some (aforms_of_point DX0))] \ aform.c1_info_invare DIM('a) X" for X by (auto simp: aform.c1_info_invare_def aform.c1_info_invar_def lens power2_eq_square) have 2: "length (normal (Sctn (unit_list D n) (lP ! n))) = DIM('a)" by (auto simp: D) have 3: "length (fst (lP, uP)) = DIM('a)" "length (snd (lP, uP)) = DIM('a)" by (auto simp: lens) have 4: "length (normal (Sctn (unit_list D n) (lP ! n))) = DIM('a)" by (auto simp: D) have 5: "length (fst (lR, uR)) = CARD('n)" "length (snd (lR, uR)) = CARD('n)" "aform.lvivl'_invar (CARD('n) * CARD('n)) (Some (lDR, uDR))" by (auto simp: lens aform.lvivl'_invar_def) note solves = solves[unfolded solves_poincare_map_aform'_fo_def file_output_iff] have "poincare_mapsto (set_of_lvivl (lP, uP) \ plane_of (map_sctn eucl_of_list (Sctn (unit_list D n) (lP ! n)))) (aform.c1_info_of_apprse [((1, 1), aforms_of_ivls lX0 uX0, Some (aforms_of_point DX0))]) (below_halfspace (map_sctn eucl_of_list (Sctn (unit_list D n) (lP ! n)))) (aform.Csafe odo - set_of_lvivl (lP, uP) \ plane_of (map_sctn eucl_of_list (Sctn (unit_list D n) (lP ! n)))) (set_of_lvivl (lR, uR) \ blinfuns_of_lvivl' (Some (lDR, uDR)))" by (rule solves_poincare_map_aform'[OF solves, OF 1 dim 2 3 4 guards 5]) auto then have "poincare_mapsto P (X0 \ {blinfun_of_list DX0}::('a \ ('a \\<^sub>L 'a)) set) SS UNIV (R \ blinfuns_of_lvivl (lDS, uDS))" apply (rule poincare_mapsto_subset) subgoal using X0 apply (auto simp: aform.c1_info_of_appre_def aform.c1_info_of_appr_def aform.c1_info_of_apprse_def) subgoal for x0 apply (rule image_eqI[where x="list_of_eucl x0@DX0"]) using lens apply (auto simp: flow1_of_list_def aforms_of_point_def Joints_aforms_of_ivls_append_point) apply (rule imageI) using X0 by (auto simp: Joints_aforms_of_ivls la2 list_of_eucl_in_list_interval_iff) done subgoal by simp subgoal using R DS by (auto simp: set_of_lvivl_def set_of_ivl_def blinfuns_of_lvivl'_def blinfuns_of_lvivl_def lens) subgoal using assms by (simp add: below_halfspace_def le_halfspace_def[abs_def]) subgoal using assms by (fastforce simp add: P set_of_lvivl_def set_of_ivl_def plane_of_def le_eucl_of_list_iff eucl_of_list_le_iff) done then show "returns_to P x \ return_time P differentiable at x within SS \ (\D. (poincare_map P has_derivative blinfun_apply D) (at x within SS) \ poincare_map P x \ R \ D o\<^sub>L blinfun_of_list DX0 \ blinfuns_of_lvivl (lDS, uDS))" using \x \ X0\ by (auto simp: poincare_mapsto_def) qed definition guards_invar::"nat \ (((real list \ real list) \ real list sctn) list \ (real \ real pdevs) reach_options) list \ bool" where "guards_invar D guards = (\(xs, ro) \ set guards. \((a, b), ba) \ set xs. length a = D \ length b = D \ length (normal ba) = D)" theorem solves_poincare_map_aform'I: assumes "TAG_optns optns" assumes "TAG_reach_optns roi" assumes "TAG_sctn mirrored" assumes D: "D = DIM('a)" assumes guards: "guards_invar DIM('a) guards" and P: "P = {eucl_of_list lP .. eucl_of_list uP}" and plane: "uP ! n = lP ! n" and dim: "aform.D odo = DIM('a)" and X0: "X0 \ {eucl_of_list lX0 .. eucl_of_list uX0}" and nD: "n < DIM('a)" and R: "{eucl_of_list lR .. eucl_of_list uR} \ R" and lens: "length (lP) = DIM('a)" "length (uP) = DIM('a)" "length (lX0) = DIM('a)" "length (uX0) = DIM('a)" "length (lR) = DIM('a)" "length (uR) = DIM('a)" and solves: "solves_poincare_map_aform'_fo optns odo (Sctn (unit_list D n) (lP ! n)) guards (lP, uP) (mirrored_sctn mirrored (Sctn (unit_list D n) (lP ! n))) roi [((1,1), aforms_of_ivls lX0 uX0, None)] (lR, uR) None" shows "\x\X0. returns_to P x \ poincare_map P x \ R" proof - note solves = solves[unfolded solves_poincare_map_aform'_fo_def file_output_iff] have 1: "\X. X \ set [((1::ereal, 1::ereal), aforms_of_ivls lX0 uX0, None)] \ aform.c1_info_invare DIM('a) X" for X by (auto simp: aform.c1_info_invare_def aform.c1_info_invar_def lens) have 2: "length (normal ((mirrored_sctn mirrored (Sctn (unit_list D n) (lP ! n))))) = DIM('a)" by (auto simp: D mirrored_sctn_def) have 3: "length (fst (lP, uP)) = DIM('a)" "length (snd (lP, uP)) = DIM('a)" by (auto simp: lens) have 4: "length (normal (((Sctn (unit_list D n) (lP ! n))))) = DIM('a)" by (auto simp: D mirrored_sctn_def) from guards have guards: "(xs, ro) \ set guards \ ((a, b), ba) \ set xs \ length a = DIM('a) \ length b = DIM('a) \ length (normal ba) = DIM('a)" for xs ro a b ba by (auto simp: guards_invar_def) have 5: "length (fst (lR, uR)) = CARD('n)" "length (snd (lR, uR)) = CARD('n)" "aform.lvivl'_invar (CARD('n) * CARD('n)) None" by (auto simp: lens) have "poincare_mapsto (set_of_lvivl (lP, uP) \ plane_of (map_sctn eucl_of_list (mirrored_sctn mirrored (Sctn (unit_list D n) (lP ! n))))) (aform.c1_info_of_apprse [((1, 1), aforms_of_ivls lX0 uX0, None)]) (below_halfspace (map_sctn eucl_of_list (Sctn (unit_list D n) (lP ! n)))) (aform.Csafe odo - set_of_lvivl (lP, uP) \ plane_of (map_sctn eucl_of_list (mirrored_sctn mirrored (Sctn (unit_list D n) (lP ! n))))) (set_of_lvivl (lR, uR) \ blinfuns_of_lvivl' None)" by (rule solves_poincare_map_aform'[OF solves, OF 1 dim 2 3 4 guards 5]) then have "poincare_mapsto P (X0 \ UNIV::('a \ ('a \\<^sub>L 'a)) set) (below_halfspace (map_sctn eucl_of_list (((Sctn (unit_list D n) (lP ! n)))))) UNIV (R \ UNIV)" apply (rule poincare_mapsto_subset) subgoal using X0 apply (auto simp: aform.c1_info_of_apprse_def aform.c1_info_of_appre_def aform.c1_info_of_appr_def) apply (rule image_eqI) apply (rule eucl_of_list_list_of_eucl[symmetric]) apply (rule aforms_of_ivls) by (auto simp add: lens subset_iff le_eucl_of_list_iff eucl_of_list_le_iff) subgoal by simp subgoal using R by (auto simp: set_of_lvivl_def set_of_ivl_def) subgoal using assms by (simp add: below_halfspace_def le_halfspace_def[abs_def]) subgoal using assms by (fastforce simp add: P set_of_lvivl_def set_of_ivl_def plane_of_def le_eucl_of_list_iff eucl_of_list_le_iff mirrored_sctn_def mirror_sctn_def) done then show "\x\X0. returns_to P x \ poincare_map P x \ R" by (auto simp: poincare_mapsto_def) qed definition "poincare_map_from_outside = poincare_map" theorem poincare_maps_onto_aformI: assumes "TAG_optns optns" assumes "TAG_reach_optns roi" assumes "TAG_sctn mirrored" assumes D: "D = DIM('a)" assumes guards: "guards_invar DIM('a) guards" and P: "P = {eucl_of_list lP .. eucl_of_list uP}" and plane: "uP ! n = lP ! n" and dim: "aform.D odo = DIM('a)" and X0: "X0 \ {eucl_of_list lX0 .. eucl_of_list uX0}" and nD: "n < DIM('a)" and R: "{eucl_of_list lR .. eucl_of_list uR} \ R" and lens: "length (lP) = DIM('a)" "length (uP) = DIM('a)" "length (lX0) = DIM('a)" "length (uX0) = DIM('a)" "length (lR) = DIM('a)" "length (uR) = DIM('a)" and solves: "solves_poincare_map_onto_aform_fo optns odo guards (lP, uP) (mirrored_sctn mirrored (Sctn (unit_list D n) (lP ! n))) roi [((1,1), aforms_of_ivls lX0 uX0, None)] (lR, uR) None" shows "\x\X0. returns_to P x \ poincare_map_from_outside P x \ R" proof - note solves = solves[unfolded solves_poincare_map_onto_aform_fo_def file_output_iff] have 1: "\X. X \ set [((1::ereal, 1::ereal), aforms_of_ivls lX0 uX0, None)] \ aform.c1_info_invare DIM('a) X" for X by (auto simp: aform.c1_info_invare_def aform.c1_info_invar_def lens) have 2: "length (normal ((mirrored_sctn mirrored (Sctn (unit_list D n) (lP ! n))))) = DIM('a)" by (auto simp: D mirrored_sctn_def) have 3: "length (fst (lP, uP)) = DIM('a)" "length (snd (lP, uP)) = DIM('a)" by (auto simp: lens) from guards have guards: "(xs, ro) \ set guards \ ((a, b), ba) \ set xs \ length a = DIM('a) \ length b = DIM('a) \ length (normal ba) = DIM('a)" for xs ro a b ba by (auto simp: guards_invar_def) have 5: "length (fst (lR, uR)) = CARD('n)" "length (snd (lR, uR)) = CARD('n)" "aform.lvivl'_invar (CARD('n) * CARD('n)) None" by (auto simp: lens) have "poincare_mapsto (set_of_lvivl (lP, uP) \ plane_of (map_sctn eucl_of_list (mirrored_sctn mirrored (Sctn (unit_list D n) (lP ! n))))) (aform.c1_info_of_apprse [((1, 1), aforms_of_ivls lX0 uX0, None)]) UNIV (aform.Csafe odo - set_of_lvivl (lP, uP) \ plane_of (map_sctn eucl_of_list (mirrored_sctn mirrored (Sctn (unit_list D n) (lP ! n))))) (set_of_lvivl (lR, uR) \ blinfuns_of_lvivl' None)" by (rule solves_poincare_map_onto_aform[OF solves, OF 1 dim 2 3 guards 5]) then have "poincare_mapsto P (X0 \ UNIV::('a \ ('a \\<^sub>L 'a)) set) UNIV UNIV (R \ UNIV)" apply (rule poincare_mapsto_subset) subgoal using X0 apply (auto simp: aform.c1_info_of_apprse_def aform.c1_info_of_appre_def aform.c1_info_of_appr_def) apply (rule image_eqI) apply (rule eucl_of_list_list_of_eucl[symmetric]) apply (rule aforms_of_ivls) by (auto simp add: lens subset_iff le_eucl_of_list_iff eucl_of_list_le_iff) subgoal by simp subgoal using R by (auto simp: set_of_lvivl_def set_of_ivl_def) subgoal by simp subgoal using assms by (fastforce simp add: P set_of_lvivl_def set_of_ivl_def plane_of_def le_eucl_of_list_iff eucl_of_list_le_iff mirrored_sctn_def mirror_sctn_def) done then show "\x\X0. returns_to P x \ poincare_map_from_outside P x \ R" by (auto simp: poincare_mapsto_def poincare_map_from_outside_def) qed end lemmas [simp] = length_approxs context includes ode_ops.lifting begin lift_definition empty_ode_ops::"ode_ops" is "([], true_form)" by (auto simp: ) end ML \val ode_numerics_conv = @{computation_check terms: Trueprop Not solves_one_step_until_time_aform_fo solves_poincare_map_aform'_fo solves_poincare_map_onto_aform_fo num_options num_options_c1 ro (* nat *) Suc "0::nat" "1::nat" "(+)::nat \ nat \ nat" "(-) ::nat \ nat \ nat" "(=)::nat\nat\bool" "(^)::nat\nat\nat" (* int / integer*) "(=)::int\int\bool" "(+)::int\int\int" "uminus::_\int" "uminus::_\integer" int_of_integer integer_of_int "0::int" "1::int" "(^)::int\nat\int" (* real *) "(=)::real\real\bool" "real_of_float" "(/)::real\real\real" "(^)::real\nat\real" "uminus::real\_" "(+)::real\real\real" "(-)::real\real\real" "(*)::real\real\real" real_divl real_divr real_of_int "0::real" "1::real" (* rat *) Fract "0::rat" "1::rat" "(+)::rat\rat\rat" "(-)::rat\rat\rat" "(*)::rat\rat\rat" "uminus::rat\_" "(/)::rat\rat\rat" "(^)::rat\nat\rat" (* ereal *) "1::ereal" (* lists: *) "replicate::_\real\_" "unit_list::_\_\real list" "Nil::(nat \ nat \ string) list" "Cons::_\_\(nat \ nat \ string) list" "Nil::(nat \ nat \ string \ nat list) list" "Cons::_\_\(nat \ nat \ string \ nat list) list" "Nil::real list" "Cons::_\_\real list" "Nil::nat list" "Cons::_\_\nat list" "Nil::string list" "Cons::_\_\string list" "Nil::real aform list" "Cons::_\_\real aform list" "Nil::(float interval) option list" "Cons::_\_\(float interval) option list" "nth::_\_\real" "upt" (* products: *) "Pair::_\_\(nat \ string)" "Pair::_\_\(nat \ nat \ string)" "Pair::_\_\char list \ nat list" "Pair::_\_\nat \ char list \ nat list" "Pair::_\_\nat \ nat \ char list \ nat list" "Pair::_\_\char list \ ((String.literal \ unit) \ (real \ real pdevs) numeric_options)" "Pair::_\_\ereal\ereal" "Pair::_\_\real aform list \ real aform list option" "Pair::_\_\(ereal \ ereal) \ real aform list \ real aform list option" "Pair::_\_\real aform" "Pair::_\_\real list \ real list" "Nil::(((real list \ real list) \ real list sctn) list \ (real aform) reach_options) list" "Cons::_\_\(((real list \ real list) \ real list sctn) list \ (real aform) reach_options) list" "Nil::((real list \ real list) \ real list sctn) list" "Cons::_\_\((real list \ real list) \ real list sctn) list" "Pair::_\_\((real list \ real list) \ real list sctn) list \ real aform reach_options" "Nil::((ereal \ ereal) \ (real aform) list \ (real aform) list option) list" "Cons::_\_\((ereal \ ereal) \ (real aform) list \ (real aform) list option) list" (* option *) "None::(real aform) list option" "Some::_\(real aform) list option" "None::(real list \ real list) option" "Some::_\(real list \ real list) option" (* aforms *) "aform_of_ivl::_\_\real aform" aforms_of_ivls aforms_of_point (* pdevs*) "zero_pdevs::real pdevs" "zero_aforms::_ \ real aform list" (* ode_ops *) mk_ode_ops init_ode_ops empty_ode_ops can_mk_ode_ops "the::ode_ops option \ ode_ops" the_odo (* Characters/Strings *) String.Char String.implode "Nil::string" "Cons::_\_\string" (* float *) "(=)::float\float\bool" "(+)::float\float\float" "uminus::_\float" "(-)::_\_\float" Float float_of_int float_of_nat (* approximation... *) approx approx1 approx2 approxs1 approxs2 approx_defined approxs_defined (* floatarith *) "0::floatarith" "1::floatarith" "(+)::_\_\floatarith" "(-)::_\_\floatarith" "(*)::_\_\floatarith" "(/)::_\_\floatarith" "inverse::_\floatarith" "uminus::_\floatarith" "Sum\<^sub>e::_\nat list\floatarith" Sin Half Tan R\<^sub>e Norm (* form *) true_form Half (* Printing *) print no_print (* sections *) xsec xsec' ysec ysec' zsec zsec' xsec2 xsec2' ysec2 ysec2' ysec4 ysec4' mirrored_sctn Code_Target_Nat.natural Code_Target_Int.positive Code_Target_Int.negative Code_Numeral.positive Code_Numeral.negative datatypes: bool num floatarith "floatarith list" form "real list sctn" "real \ real" } \ ML \fun ode_numerics_tac ctxt = CONVERSION (ode_numerics_conv ctxt) THEN' (resolve_tac ctxt [TrueI])\ lemma eq_einterpretI: assumes "list_of_eucl (VS::'a::executable_euclidean_space) = interpret_floatariths fas xs" assumes "length fas = DIM('a)" shows "VS = eucl_of_list (interpret_floatariths fas xs)" using assms apply (subst (asm) list_of_eucl_eucl_of_list[symmetric]) apply (auto intro!: ) by (metis eucl_of_list_list_of_eucl) lemma one_add_square_ne_zero[simp]: "1 + t * t \ 0" for t::real by (metis semiring_normalization_rules(12) sum_squares_eq_zero_iff zero_neq_one) lemma abs_rat_bound: "abs (x - y) \ e / f" if "y \ {yl .. yu}" "x \ {yu - real_divl p e f.. yl + real_divl p e f}" for x y e::real proof - note \x \ _\ also have "{yu - real_divl p e f.. yl + real_divl p e f} \ {yu - e / f .. yl + e / f}" by (auto intro!: diff_mono real_divl) finally show ?thesis using that unfolding abs_diff_le_iff by auto qed lemma in_ivl_selfI: "a \ {a .. a}" for a::real by auto lemma pi4_bnds: "pi / 4 \ {real_divl 80 (lb_pi 80) 4 .. real_divr 80 (ub_pi 80) 4}" using pi_boundaries[of 80] unfolding atLeastAtMost_iff by (intro conjI real_divl[le] real_divr[ge] divide_right_mono) auto lemma abs_minus_leI: "\x - x'\ \ e" if "x \ {x' - e .. x' + e}" for x e::real using that by (auto simp: ) lemmas [DIM_simps] = Suc_numeral One_nat_def[symmetric] TrueI Suc_1 length_approxs arith_simps lemma (in ode_interpretation) length_ode_e[DIM_simps]: "length (ode_expression odo) = DIM('a)" by (auto simp: len) named_theorems solves_one_step_ivl_thms context ode_interpretation begin lemmas [solves_one_step_ivl_thms] = TAG_optns[OF solves_one_step_ivl[OF _ _ _ _ _ _ _ _ solves_one_step_until_time_aform_foI], rotated -1, of optns _ _ _ _ _ _ _ _ _ optns for optns] lemmas [solves_one_step_ivl_thms] = TAG_optns[OF solves_one_step_ivl'[OF _ _ _ _ _ _ _ _ _ _ _ _ solves_one_step_until_time_aform_foI], rotated -1, of optns _ _ _ _ _ _ _ _ _ _ _ _ _ _ optns for optns] lemmas [solves_one_step_ivl_thms] = solves_poincare_map_aform'I poincare_maps_onto_aformI end lemma TAG_optnsI: "TAG_optns optns" by simp named_theorems poincare_tac_theorems lemmas [DIM_simps] = one_less_numeral_iff rel_simps abbreviation "point_ivl a \ {a .. a}" lemma isFDERIV_compute: "isFDERIV D vs fas xs \ (list_all (\i. list_all (\j. isDERIV (vs ! i) (fas ! j) xs) [0.. length fas = D \ length vs = D" unfolding isFDERIV_def by (auto simp: list.pred_set) theorem (in ode_interpretation) solves_poincare_map_aform'_derivI'[solves_one_step_ivl_thms]: \ \TODO: replace @{thm solves_poincare_map_aform'_derivI}\ assumes "TAG_optns optns" assumes "TAG_reach_optns roi" assumes "TAG_sctn mirrored" and D: "D = DIM('a)" assumes DS: "list_interval lDR uDR \ list_interval lDS uDS" and ode_fas: "aform.D odo = DIM('a)" and guards: "guards_invar DIM('a) guards" and P: "P = {eucl_of_list lP .. eucl_of_list uP}" and plane: "uP ! n = lP ! n" and X0: "X0 \ {eucl_of_list lX0 .. eucl_of_list uX0}" and nD: "n < DIM('a)" and R: "{eucl_of_list lR .. eucl_of_list uR} \ R" and lens: "length (lP) = DIM('a)" "length (uP) = DIM('a)" "length (lX0) = DIM('a)" "length (uX0) = DIM('a)" "length (lR) = DIM('a)" "length (uR) = DIM('a)" "length DX0 = DIM('a)*DIM('a)" "length lDR = CARD('n) * CARD('n)" "length uDR = CARD('n) * CARD('n)" and SS: "SS = {x::'a. if mirrored then x \ Basis_list ! n \ lP ! n else x \ Basis_list ! n \ lP ! n}" assumes solves: "solves_poincare_map_aform'_fo optns odo (mirrored_sctn (\mirrored) (Sctn (unit_list D n) (lP ! n))) guards (lP, uP) (mirrored_sctn mirrored (Sctn (unit_list D n) (lP ! n))) roi [((1,1), aforms_of_ivls lX0 uX0, Some (aforms_of_point DX0))] (lR, uR) (Some (lDR, uDR))" shows "\x\X0. returns_to P x \ return_time P differentiable at x within SS \ (\D. (poincare_map P has_derivative blinfun_apply D) (at x within SS) \ poincare_map P x \ R \ D o\<^sub>L blinfun_of_list DX0 \ blinfuns_of_lvivl (lDS, uDS))" proof (rule ballI) fix x assume "x \ X0" then have la2: "list_all2 (\) lX0 uX0" using X0 by (force simp: subset_iff eucl_of_list_le_iff le_eucl_of_list_iff lens list_all2_conv_all_nth) have 1: "\X. X \ set [((1::ereal, 1::ereal), aforms_of_ivls lX0 uX0, Some (aforms_of_point DX0))] \ aform.c1_info_invare DIM('a) X" for X by (auto simp: aform.c1_info_invare_def aform.c1_info_invar_def lens power2_eq_square) have 2: "length (normal (mirrored_sctn (\mirrored) (Sctn (unit_list D n) (lP ! n)))) = DIM('a)" by (auto simp: D mirrored_sctn_def) have 3: "length (fst (lP, uP)) = DIM('a)" "length (snd (lP, uP)) = DIM('a)" by (auto simp: lens) have 4: "length (normal (mirrored_sctn mirrored (Sctn (unit_list D n) (lP ! n)))) = DIM('a)" by (auto simp: D mirrored_sctn_def) have 5: "length (fst (lR, uR)) = CARD('n)" "length (snd (lR, uR)) = CARD('n)" "aform.lvivl'_invar (CARD('n) * CARD('n)) (Some (lDR, uDR))" by (auto simp: lens aform.lvivl'_invar_def) note solves = solves[unfolded solves_poincare_map_aform'_fo_def file_output_iff] have "poincare_mapsto (set_of_lvivl (lP, uP) \ plane_of (map_sctn eucl_of_list (mirrored_sctn mirrored (Sctn (unit_list D n) (lP ! n))))) (aform.c1_info_of_apprse [((1, 1), aforms_of_ivls lX0 uX0, Some (aforms_of_point DX0))]) (below_halfspace (map_sctn eucl_of_list (mirrored_sctn (\mirrored) (Sctn (unit_list D n) (lP ! n))))) (aform.Csafe odo - set_of_lvivl (lP, uP) \ plane_of (map_sctn eucl_of_list (mirrored_sctn mirrored (Sctn (unit_list D n) (lP ! n))))) (set_of_lvivl (lR, uR) \ blinfuns_of_lvivl' (Some (lDR, uDR)))" by (rule solves_poincare_map_aform'[OF solves, OF 1 ode_fas 4 3 2 _ 5]) (use guards in \auto simp: guards_invar_def\) then have "poincare_mapsto P (X0 \ {blinfun_of_list DX0}::('a \ ('a \\<^sub>L 'a)) set) SS UNIV (R \ blinfuns_of_lvivl (lDS, uDS))" apply (rule poincare_mapsto_subset) subgoal using X0 apply (auto simp: aform.c1_info_of_appre_def aform.c1_info_of_appr_def aform.c1_info_of_apprse_def) subgoal for x0 apply (rule image_eqI[where x="list_of_eucl x0@DX0"]) using lens apply (auto simp: flow1_of_list_def aforms_of_point_def Joints_aforms_of_ivls_append_point) apply (rule imageI) using X0 by (auto simp: Joints_aforms_of_ivls la2 list_of_eucl_in_list_interval_iff) done subgoal by simp subgoal using R DS by (auto simp: set_of_lvivl_def set_of_ivl_def blinfuns_of_lvivl'_def blinfuns_of_lvivl_def lens) subgoal using assms by (auto simp: below_halfspace_def le_halfspace_def[abs_def] mirrored_sctn_def mirror_sctn_def) subgoal using assms by (fastforce simp add: P set_of_lvivl_def set_of_ivl_def plane_of_def le_eucl_of_list_iff eucl_of_list_le_iff mirrored_sctn_def mirror_sctn_def) done then show "returns_to P x \ return_time P differentiable at x within SS \ (\D. (poincare_map P has_derivative blinfun_apply D) (at x within SS) \ poincare_map P x \ R \ D o\<^sub>L blinfun_of_list DX0 \ blinfuns_of_lvivl (lDS, uDS))" using \x \ X0\ by (auto simp: poincare_mapsto_def) qed lemmas [DIM_simps] = aform.ode_e_def ML \ structure ODE_Numerics_Tac = struct fun mk_nat n = HOLogic.mk_number @{typ nat} n fun mk_int n = HOLogic.mk_number @{typ int} n fun mk_integer n = @{term integer_of_int} $ (HOLogic.mk_number @{typ int} n) fun mk_bool b = if b then @{term True} else @{term False} fun mk_numeralT n = let fun mk_bit 0 ty = Type (@{type_name bit0}, [ty]) | mk_bit 1 ty = Type (@{type_name bit1}, [ty]); fun bin_of n = if n = 1 then @{typ num1} else if n = 0 then @{typ num0} else if n = ~1 then raise TERM ("negative type numeral", []) else let val (q, r) = Integer.div_mod n 2; in mk_bit r (bin_of q) end; in bin_of n end; fun print_tac' ctxt s = K (print_tac ctxt s) val using_master_directory = File.full_path o Resources.master_directory o Proof_Context.theory_of; fun using_master_directory_term ctxt s = (if s = "-" orelse s = "" then s else Path.explode s |> using_master_directory ctxt |> Path.implode) |> HOLogic.mk_string fun real_in_approx_tac ctxt p = let val inst_approx = ([], [((("prec", 0), @{typ nat}), mk_nat p |> Thm.cterm_of ctxt)]) val approx_thm = Thm.instantiate inst_approx @{thm real_in_approxI} in resolve_tac ctxt [approx_thm] THEN' SOLVED' (reify_floatariths_tac ctxt) THEN' ode_numerics_tac ctxt end fun real_subset_approx_tac ctxt p = let val inst_approx = ([], [((("prec", 0), @{typ nat}), mk_nat p |> Thm.cterm_of ctxt)]) val approx_thm = Thm.instantiate inst_approx @{thm real_subset_approxI} in resolve_tac ctxt [approx_thm] THEN' SOLVED' (reify_floatariths_tac ctxt) THEN' SOLVED' (reify_floatariths_tac ctxt) THEN' ode_numerics_tac ctxt THEN' ode_numerics_tac ctxt end fun basic_nt_ss ctxt nt = put_simpset HOL_basic_ss ctxt addsimps Named_Theorems.get ctxt nt fun DIM_tac defs ctxt = (Simplifier.simp_tac (basic_nt_ss ctxt @{named_theorems DIM_simps} addsimps defs)) fun subset_approx_preconds_tac ctxt p thm = let val inst_approx = ([], [((("prec", 0), @{typ nat}), mk_nat p |> Thm.cterm_of ctxt)]) in resolve_tac ctxt [Thm.instantiate inst_approx thm] THEN' SOLVED' (reify_floatariths_tac ctxt) THEN' SOLVED' (reify_floatariths_tac ctxt) THEN' SOLVED' (DIM_tac [] ctxt) THEN' SOLVED' (DIM_tac [] ctxt) THEN' SOLVED' (ode_numerics_tac ctxt) THEN' SOLVED' (ode_numerics_tac ctxt) end val cfg_trace = Attrib.setup_config_bool @{binding ode_numerics_trace} (K false) fun tracing_tac ctxt = if Config.get ctxt cfg_trace then print_tac ctxt else K all_tac fun tracing_tac' ctxt = fn s => K (tracing_tac ctxt s) fun eucl_subset_approx_tac ctxt p = subset_approx_preconds_tac ctxt p @{thm eucl_subset_approxI} fun approx_subset_eucl_tac ctxt p = subset_approx_preconds_tac ctxt p @{thm approx_subset_euclI} fun approx_subset_list_tac ctxt p = subset_approx_preconds_tac ctxt p @{thm approx_subset_listI} val static_simpset = Simplifier.simpset_of @{context} fun nth_tac ctxt = Simplifier.simp_tac (put_simpset HOL_basic_ss ctxt addsimps @{thms nth_Cons_0 nth_Cons_Suc numeral_nat}) fun nth_list_eq_tac ctxt n = Subgoal.FOCUS_PARAMS (fn {context, concl, ...} => case try (Thm.term_of #> HOLogic.dest_Trueprop #> HOLogic.dest_eq) concl of SOME (@{const List.nth(real)} $ xs $ Var _, @{const List.nth(real)} $ ys $ Var _) => let val i = find_index (op=) (HOLogic.dest_list xs ~~ HOLogic.dest_list ys) val thm = Goal.prove context [] [] (HOLogic.mk_eq (@{const List.nth(real)} $ xs $ HOLogic.mk_number @{typ nat} i, @{const List.nth(real)} $ ys $ HOLogic.mk_number @{typ nat} i) |> HOLogic.mk_Trueprop) (fn {context, ...} => HEADGOAL (nth_tac context)) in SOLVE (HEADGOAL (resolve_tac context [thm])) end | _ => no_tac ) ctxt n fun numeric_precond_step_tac defs thms p = Subgoal.FOCUS_PARAMS (fn {context, concl, ...} => let val prems = Logic.strip_imp_prems (Thm.term_of concl) val conclusion = Logic.strip_imp_concl (Thm.term_of concl) in (case conclusion |> HOLogic.dest_Trueprop of @{const Set.member(real)} $ _ $ _ => tracing_tac context "numeric_precond_step: real in approx" THEN HEADGOAL (real_in_approx_tac context p) | Const(@{const_name less_eq}, _) $ (Const (@{const_name atLeastAtMost}, _) $ _ $ _) $ (Const (@{const_name atLeastAtMost}, _) $ Var _ $ Var _) => tracing_tac context "numeric_precond_step: approx subset eucl" THEN HEADGOAL (real_subset_approx_tac context p) | Const (@{const_name less_eq}, _) $ (Const (@{const_name atLeastAtMost}, _) $ (Const (@{const_name eucl_of_list}, _) $ Var _) $ _) $ _ => tracing_tac context "numeric_precond_step: approx subset eucl" THEN HEADGOAL (approx_subset_eucl_tac context p) | Const (@{const_name less_eq}, _) $ _ $ (Const (@{const_name atLeastAtMost}, _) $ (Const (@{const_name eucl_of_list}, _) $ Var _) $ _) => tracing_tac context "numeric_precond_step: eucl subset approx" THEN HEADGOAL (eucl_subset_approx_tac context p) | Const (@{const_name less_eq}, _) $ (@{const list_interval(real)} $ _ $ _) $ (@{const list_interval(real)} $ _ $ _) => tracing_tac context "numeric_precond_step: approx subset list" THEN HEADGOAL (approx_subset_list_tac context p) | @{const HOL.eq(nat)} $ _ $ _ => tracing_tac context "numeric_precond_step: DIM_tac" THEN HEADGOAL (SOLVED' (DIM_tac [] context)) | @{const less(nat)} $ _ $ _ => tracing_tac context "numeric_precond_step: DIM_tac" THEN HEADGOAL (SOLVED' (DIM_tac [] context)) | @{const HOL.eq(real)} $ (@{const nth(real)} $ _ $ _) $ (@{const nth(real)} $ _ $ _) => tracing_tac context "numeric_precond_step: nth_list_eq_tac" THEN HEADGOAL (SOLVED' (nth_list_eq_tac context)) | Const (@{const_name "HOL.eq"}, _) $ _ $ (Const (@{const_name eucl_of_list}, _) $ (@{const interpret_floatariths} $ _ $ _)) => tracing_tac context "numeric_precond_step: reify floatariths" THEN HEADGOAL (resolve_tac context @{thms eq_einterpretI} THEN' reify_floatariths_tac context) | t as _ $ _ => let val (c, args) = strip_comb t in if member (op=) [@{const "solves_one_step_until_time_aform_fo"}, @{const "solves_poincare_map_aform'_fo"}, @{const "solves_poincare_map_onto_aform_fo"}, @{const "can_mk_ode_ops"} ] c then tracing_tac context "numeric_precond_step: ode_numerics_tac" THEN HEADGOAL ( CONVERSION (Simplifier.rewrite (put_simpset HOL_basic_ss context addsimps defs)) THEN' tracing_tac' context "numeric_precond_step: ode_numerics_tac (unfolded)" THEN' ode_numerics_tac context) else if member (op=) [@{const "isFDERIV"}] c then tracing_tac context "numeric_precond_step: isFDERIV" THEN HEADGOAL (SOLVED'(Simplifier.asm_full_simp_tac (put_simpset static_simpset context addsimps (@{thms isFDERIV_def less_Suc_eq_0_disj isDERIV_Power_iff} @ thms @ defs)) THEN' tracing_tac' context "numeric_precond_step: simplified isFDERIV" )) else tracing_tac context "numeric_precond_step: boolean, try thms" THEN HEADGOAL (SOLVED' (resolve_tac context thms)) end | _ => tracing_tac context "numeric_precond_step: boolean constant" THEN no_tac ) end) fun integral_bnds_tac_gen_start sstep d p m N atol filename ctxt i = let val insts = ([((("'i", 0), @{sort "{enum}"}), mk_numeralT (d + 1) |> Thm.ctyp_of ctxt)], [((("optns", 0), @{typ "string \ ((String.literal \ unit) \(real aform) numeric_options)"}), HOLogic.mk_prod (using_master_directory_term ctxt filename, (@{term num_options} $ mk_nat p $ mk_int sstep $ mk_nat m $ mk_nat N $ mk_int atol $ @{term "[(0::nat, 1::nat, ''0x000000'')]"})) |> Thm.cterm_of ctxt), ((("safe_form", 0), @{typ form}), @{cterm true_form}) ]) in resolve_tac ctxt [Thm.instantiate insts @{thm solve_one_step_until_time_aform_integral_bounds}] i THEN (Lin_Arith.tac ctxt i ORELSE Simplifier.simp_tac ctxt i) end fun integral_bnds_tac_gen sstep d p m N atol filename thms ctxt = integral_bnds_tac_gen_start sstep d p m N atol filename ctxt THEN_ALL_NEW_FWD REPEAT_ALL_NEW_FWD (numeric_precond_step_tac [] thms p ctxt) val integral_bnds_tac = integral_bnds_tac_gen 5 fun mk_proj (m, n, s) = HOLogic.mk_tuple [mk_nat m, mk_nat n, HOLogic.mk_string s] fun mk_projs projs = HOLogic.mk_list @{typ "nat \ nat \ string"} (map mk_proj projs) fun mk_string_list ds = HOLogic.mk_list @{typ "string"} (map HOLogic.mk_string ds) fun mk_nat_list ds = HOLogic.mk_list @{typ "nat"} (map mk_nat ds) fun mk_proj_c1 (m, n, s, ds) = HOLogic.mk_tuple [mk_nat m, mk_nat n, HOLogic.mk_string s, mk_nat_list ds] fun mk_projs_c1 projs = HOLogic.mk_list @{typ "nat \ nat \ string \ nat list"} (map mk_proj_c1 projs) fun TAG_optns_thm p sstep m N atol projs filename ctxt = Thm.instantiate ([], [((("optns", 0), @{typ "string \ ((String.literal \ unit) \(real aform) numeric_options)"}), HOLogic.mk_prod (using_master_directory_term ctxt filename, @{term num_options} $ mk_nat p $ mk_int sstep $ mk_nat m $ mk_nat N $ mk_int atol $ mk_projs projs) |> Thm.cterm_of ctxt)]) @{thm TAG_optnsI} fun TAG_optns_c1_thm p sstep m N atol projs ds filename ctxt = Thm.instantiate ([], [((("optns", 0), @{typ "string \ ((String.literal \ unit) \(real aform) numeric_options)"}), HOLogic.mk_prod (using_master_directory_term ctxt filename, @{term num_options_c1} $ mk_nat p $ mk_int sstep $ mk_nat m $ mk_nat N $ mk_int atol $ mk_projs_c1 projs $ mk_string_list ds) |> Thm.cterm_of ctxt)]) @{thm TAG_optnsI} fun ode_bnds_tac_gen_start sstep p m N atol projs filename ctxt = tracing_tac' ctxt "solves_one_step_ivl_thms" THEN' resolve_tac ctxt (Named_Theorems.get ctxt @{named_theorems solves_one_step_ivl_thms}) THEN' tracing_tac' ctxt "resolved solves_one_step_ivl_thms" THEN' resolve_tac ctxt [TAG_optns_thm p sstep m N atol projs filename ctxt] fun ode_bnds_tac_gen sstep ode_defs p m N atol projs filename ctxt = ode_bnds_tac_gen_start sstep p m N atol projs filename ctxt THEN_ALL_NEW_FWD REPEAT_ALL_NEW_FWD (numeric_precond_step_tac ode_defs [] p ctxt) val ode_bnds_tac = ode_bnds_tac_gen 5 fun ode'_bnds_tac_gen_start sstep p m N atol projs ds filename ctxt = tracing_tac' ctxt "solves_one_step_ivl_thms" THEN' resolve_tac ctxt (Named_Theorems.get ctxt @{named_theorems solves_one_step_ivl_thms}) THEN' tracing_tac' ctxt "resolved solves_one_step_ivl_thms" THEN' resolve_tac ctxt [TAG_optns_c1_thm p sstep m N atol projs ds filename ctxt] fun ode'_bnds_tac_gen sstep ode_defs p m N atol projs ds filename ctxt = ode'_bnds_tac_gen_start sstep p m N atol projs ds filename ctxt THEN_ALL_NEW_FWD REPEAT_ALL_NEW_FWD (numeric_precond_step_tac ode_defs [] p ctxt) val ode'_bnds_tac = ode'_bnds_tac_gen 5 fun poincare_bnds_tac_gen_start sstep p m N atol projs filename ctxt = tracing_tac' ctxt "solves_one_step_ivl_thms" THEN' resolve_tac ctxt (Named_Theorems.get ctxt @{named_theorems solves_one_step_ivl_thms}) THEN' tracing_tac' ctxt "resolved solves_one_step_ivl_thms" THEN' resolve_tac ctxt [TAG_optns_thm p sstep m N atol projs filename ctxt] fun poincare_bnds_tac_gen sstep ode_defs p m N atol projs filename ctxt = poincare_bnds_tac_gen_start sstep p m N atol projs filename ctxt THEN_ALL_NEW_FWD REPEAT_ALL_NEW_FWD ( numeric_precond_step_tac ode_defs (Named_Theorems.get ctxt @{named_theorems poincare_tac_theorems}) p ctxt) val poincare_bnds_tac = poincare_bnds_tac_gen 5 fun poincare'_bnds_tac_gen_start sstep p m N atol projs filename ctxt = resolve_tac ctxt (Named_Theorems.get ctxt @{named_theorems solves_one_step_ivl_thms}) THEN' resolve_tac ctxt [TAG_optns_thm p sstep m N atol projs filename ctxt] fun poincare'_bnds_tac_gen sstep ode_defs p m N atol projs filename ctxt = poincare'_bnds_tac_gen_start sstep p m N atol projs filename ctxt THEN_ALL_NEW_FWD REPEAT_ALL_NEW_FWD ( numeric_precond_step_tac ode_defs (Named_Theorems.get ctxt @{named_theorems poincare_tac_theorems}) p ctxt) val poincare'_bnds_tac = poincare'_bnds_tac_gen 5 end \ lemma (in auto_ll_on_open) Poincare_Banach_fixed_pointI: assumes "convex S" and c: "complete S" "S \ {}" and "S \ T" assumes derivative_bounded: "\x\S. poincare_map \ x \ S \ (\D. (poincare_map \ has_derivative D) (at x within T) \ onorm D \ B)" assumes B: "B < 1" shows "\!x. x \ S \ poincare_map \ x = x" using c _ B proof (rule banach_fix) from derivative_bounded c show "0 \ B" by (auto dest!: has_derivative_bounded_linear onorm_pos_le) from derivative_bounded show "poincare_map \ ` S \ S" by auto obtain D where D: "\x \ S. (poincare_map \ has_derivative D x) (at x within T) \ onorm (D x) \ B" apply atomize_elim apply (rule bchoice) using derivative_bounded by auto with \S \ T\ have "(\x. x \ S \ (poincare_map \ has_derivative D x) (at x within S))" - by (auto intro: has_derivative_within_subset) + by (auto intro: has_derivative_subset) from bounded_derivative_imp_lipschitz[of S "poincare_map \" D B, OF this] \convex S\ D c \0 \ B\ have "B-lipschitz_on S (poincare_map \)" by auto then show "\x\S. \y\S. dist (poincare_map \ x) (poincare_map \ y) \ B * dist x y" by (auto simp: lipschitz_on_def) qed ML \open ODE_Numerics_Tac\ lemma isFDERIV_product: "isFDERIV n xs fas vs \ length fas = n \ length xs = n \ list_all (\(x, f). isDERIV x f vs) (List.product xs fas)" apply (auto simp: isFDERIV_def list_all2_iff in_set_zip list_all_length product_nth) apply auto apply (metis gr_implies_not_zero gr_zeroI less_mult_imp_div_less pos_mod_bound) done end diff --git a/thys/Ordinary_Differential_Equations/Numerics/Refine_Reachability_Analysis.thy b/thys/Ordinary_Differential_Equations/Numerics/Refine_Reachability_Analysis.thy --- a/thys/Ordinary_Differential_Equations/Numerics/Refine_Reachability_Analysis.thy +++ b/thys/Ordinary_Differential_Equations/Numerics/Refine_Reachability_Analysis.thy @@ -1,1701 +1,1701 @@ theory Refine_Reachability_Analysis imports Abstract_Reachability_Analysis Refine_Rigorous_Numerics begin lemma isDERIV_simps[simp]: "isDERIV i 1 xs" "isDERIV i 0 xs" "isDERIV i (a + b) xs \ isDERIV i a xs \ isDERIV i b xs" "isDERIV i (a - b) xs \ isDERIV i a xs \ isDERIV i b xs" "isDERIV i (a * b) xs \ isDERIV i a xs \ isDERIV i b xs" "isDERIV i (a / b) xs \ isDERIV i a xs \ isDERIV i b xs \ interpret_floatarith b xs \ 0" "isDERIV i (-a) xs = isDERIV i a xs" by (auto simp: one_floatarith_def zero_floatarith_def plus_floatarith_def minus_floatarith_def times_floatarith_def divide_floatarith_def uminus_floatarith_def) lemma list_of_eucl_of_env: assumes [simp]: "length xs = DIM('a)" shows "(list_of_eucl (eucl_of_env xs vs::'a)) = (map (\i. vs ! (xs ! i)) [0.. DIM('a)" "length ode_e = DIM('a)" and mV: "max_Var_floatariths ode_e \ DIM('a)" shows "(eucl_of_list (interpret_floatariths (ode_fa xs) vs)::'a) = ode (eucl_of_list (interpret_floatariths xs vs))" unfolding ode_fa_def apply (auto intro!: euclidean_eqI[where 'a='a] simp: eucl_of_list_inner ode_def) apply (subst interpret_floatarith_subst_floatarith[where D="length vs"]) apply (auto intro!: max_Var_floatarith_le_max_Var_floatariths_nth[le] mV[le]) apply (rule interpret_floatarith_max_Var_cong) apply (drule max_Var_floatariths_lessI) apply simp apply (drule less_le_trans[OF _ mV]) apply auto apply (subst nth_map) apply simp using assms(2) order.strict_trans2 apply blast apply (subst nth_upt) apply simp apply (rule less_le_trans, assumption, simp) apply auto done lemma length_ode_fa[simp]: "length (ode_fa xs) = length ode_e" by (auto simp: ode_fa_def) lemma max_Var_ode_fa[le]: assumes "max_Var_floatariths ode_e \ length xs" shows "max_Var_floatariths (ode_fa xs) \ max_Var_floatariths xs" by (auto simp: ode_fa_def intro!: assms max_Var_floatariths_subst_floatarith_le) lemma max_Var_floatariths_ode_d_expr[le]: "max_Var_floatariths ode_e \ d \ d > 0 \ max_Var_floatariths (ode_d_expr d m) \ 3 * d" by (auto simp: ode_d_expr_def intro!: max_Var_floatarith_FDERIV_n_floatariths[le] max_Var_floatarith_FDERIV_floatariths[le]) lemma interpret_ode_d_fa: assumes FDERIV: "(eucl_of_list (interpret_floatariths xs vs)::'a::executable_euclidean_space) \ Csafe" assumes [simp]: "length ds = DIM('a)" "length xs = DIM('a)" notes [simp] = safe_length[OF FDERIV] shows "(eucl_of_list (interpret_floatariths (ode_d_fa n xs ds) vs)::'a) = ode_d n (eucl_of_list (interpret_floatariths xs vs)) (eucl_of_list (interpret_floatariths ds vs)) (eucl_of_list (interpret_floatariths ds vs))" apply (transfer fixing: xs ds vs n) using FDERIV apply (auto simp del: isnFDERIV.simps simp: interpret_floatariths_append) apply (auto simp add: list_of_eucl_of_env ode_def ode_d_raw_def eucl_of_list_inner intro!: euclidean_eqI[where 'a='a]) apply (auto simp: ode_d_fa_def ) apply (subst interpret_floatarith_subst_floatarith[OF max_Var_floatarith_le_max_Var_floatariths_nth], simp) apply (rule interpret_floatarith_max_Var_cong) subgoal premises prems for b i proof - from prems have i: "i < max_Var_floatariths (ode_d_expr DIM('a) n)" by (auto dest!: max_Var_floatariths_lessI) also have "\ \ 3 * DIM('a)" by (auto intro!: max_Var_floatariths_ode_d_expr safe_max_Var[OF FDERIV]) finally have "i < 3 * DIM('a)" . then show ?thesis using prems i by (auto simp: nth_append) qed done lemma safe_at_within: "x \ Csafe \ at x = at x within Csafe" by (subst (2) at_within_open) (auto simp: open_safe) lemma ivlflowsD: assumes "ivlflows stops stopcont trap rsctn" "ivl \ \(plane_of ` stops) \ UNIV " shows "ivl \ (snd (stopcont ivl))" "(fst (stopcont ivl)) \ sbelow_halfspace rsctn \ UNIV" "fst (stopcont ivl) \ snd (stopcont ivl)" "(snd (stopcont ivl)) \ sbelow_halfspace rsctn \ UNIV" "flowsto (ivl) {0..} ((snd (stopcont ivl))) ((fst (stopcont ivl)) \ trap)" using assms(1)[unfolded ivlflows_def, rule_format, OF assms(2)] by auto lemma ivlflows_flowsto: assumes "ivlflows stops stopcont trap rsctn" "ivl \ \(plane_of ` stops) \ UNIV" assumes "stopcont ivl = (x, y)" shows "flowsto (ivl) {0..} y (x \ trap)" using ivlflowsD[OF assms(1,2)] assms(3) by auto lemma ivlflows_emptyI: "ivlflows {} (\x. (x, x)) J K" by (auto simp: ivlflows_def set_of_ivl_def) lemma plane_of_neg[simp]: "plane_of (Sctn (- normal sctn) (- pstn sctn)) = plane_of sctn" by (auto simp: plane_of_def) lemmas safe_max_Var_le[intro] = safe_max_Var[le] lemmas [simp] = safe_length lemma continuous_on_ode_d2: "continuous_on (Csafe) ode_d2" proof - have isn: "isnFDERIV DIM('a) ode_e [0.. Csafe" for x i::'a by (rule safe_isnFDERIV) fact have "continuous_on (Csafe::'a set) (\x. ode_d_raw (Suc 0) x i j)" if "i \ Basis" "j \ Basis" for i j apply (rule has_derivative_continuous_on) apply (auto simp: ode_d_raw_def at_within_open[OF _ open_safe]) apply (rule interpret_floatarith_FDERIV_floatariths_append) apply (auto simp: ode_d_expr_def intro!: isDERIV_FDERIV_floatariths safe_isFDERIV_append that isFDERIV_map_Var safe_max_Var_le max_Var_floatarith_FDERIV_floatariths[le]) apply assumption apply arith done then show ?thesis apply (auto intro!: continuous_on_blinfun_componentwise) subgoal for i j apply (rule continuous_on_eq[where f="(\x. ode_d_raw (Suc 0) x i j)"]) apply force apply (subst ode_d2.rep_eq) apply simp apply (subst ode_d.rep_eq) apply (split if_splits) apply (rule conjI) apply simp using isn apply force done done qed lemmas continuous_on_ode_d2_comp[continuous_intros] = continuous_on_compose2[OF continuous_on_ode_d2] lemma map_ode_fa_nth[simp]: "d \ length ode_e \ map (ode_fa_nth CX) [0.. length ode_e \ map (ode_d_fa_nth i CX X) [0.. length vs" "max_Var_floatariths ode_e \ DIM('a)" shows "(einterpret (euler_incr_fas X0 h CX) vs::'a::executable_euclidean_space) = einterpret X0 vs + (interpret_floatarith h vs) *\<^sub>R ode (einterpret CX vs)" by (simp add: euler_incr_fas_def euler_incr_fas_nth_def assms ode_fa_nth cong: map_cong) lemma einterpret_euler_err_fas: assumes safe: "(einterpret CX vs::'a) \ Csafe" assumes [simp]: "length X0 = DIM('a)" "length CX = DIM('a)" "DIM('a) \ length vs" shows "(einterpret (euler_err_fas X0 h CX) vs::'a::executable_euclidean_space) = (((interpret_floatarith h vs))\<^sup>2/2) *\<^sub>R ode_d 0 (einterpret CX vs) (ode (einterpret CX vs)) (ode (einterpret CX vs))" using safe_length[OF safe] safe_max_Var[OF safe] apply (simp add: euler_err_fas_def euler_err_fas_nth_def[abs_def] euler_incr_fas_def) apply (subst interpret_ode_d_fa) by (auto simp: safe) lemma einterpret_euler_fas1: assumes safe[simp]: "(einterpret CX vs::'a) \ Csafe" assumes [simp]: "length X0 = DIM('a)" "length CX = DIM('a)" "DIM('a) \ length vs" shows "(einterpret (take DIM('a) (euler_fas X0 h CX)) vs::'a::executable_euclidean_space) = einterpret X0 vs + (interpret_floatarith h vs) *\<^sub>R ode (einterpret X0 vs) + (((interpret_floatarith h vs))\<^sup>2/2) *\<^sub>R ode_d 0 (einterpret CX vs) (ode (einterpret CX vs)) (ode (einterpret CX vs))" using safe_length[OF safe] safe_max_Var[OF safe] by (simp add: euler_fas_def euler_incr_fas_def euler_incr_fas_nth_def[abs_def] einterpret_euler_err_fas euler_err_fas_nth_def[abs_def] interpret_ode_d_fa) lemma einterpret_euler_fas2: assumes [simp]: "(einterpret CX vs::'a) \ Csafe" assumes [simp]: "length X0 = DIM('a)" "length CX = DIM('a)" "DIM('a) \ length vs" shows "(einterpret (drop DIM('a) (euler_fas X0 h CX)) vs::'a::executable_euclidean_space) = (((interpret_floatarith h vs))\<^sup>2/2) *\<^sub>R ode_d 0 (einterpret CX vs) (ode (einterpret CX vs)) (ode (einterpret CX vs))" by (simp add: euler_fas_def euler_incr_fas_def einterpret_euler_err_fas) lemma ode_d_Suc_0_eq_ode_d2: "x \ Csafe \ ode_d (Suc 0) x = ode_d2 x" unfolding ode_d2.rep_eq by auto lemma rk2_increment_rk2_fas_err: fixes h s1 s2 rkp x0 cx vs defines "h' \ interpret_floatarith h vs" defines "s2' \ interpret_floatarith s2 vs" defines "rkp' \ interpret_floatarith rkp vs" defines "x0' \ einterpret x0 vs" defines "cx' \ einterpret cx vs" assumes cx_flow: "cx' = flow0 x0' (h' * s1')" assumes [simp]: "length x0 = DIM('a)" "length cx = DIM('a)" "DIM('a) \ length vs" assumes safes: "x0' \ Csafe" "cx' \ Csafe" "(x0' + (s2' * h' * rkp') *\<^sub>R ode x0')\ Csafe" shows "(einterpret (rk2_fas_err rkp x0 h cx s2) vs::'a::executable_euclidean_space) = heun_remainder1 (flow0 x0') ode_na ode_d_na ode_d2_na 0 h' s1' - heun_remainder2 rkp' (flow0 x0') ode_na ode_d2_na 0 h' s2'" using safes using safe_length[OF safes(1)] safe_max_Var[OF safes(1)] apply (auto simp: heun_remainder1_def heun_remainder2_def discrete_evolution_def ode_na_def ode_d_na_def ode_d2_na_def rk2_increment x0'_def rkp'_def h'_def s2'_def cx'_def euler_incr_fas_def rk2_fas_err_def rk2_fas_err_nth_def[abs_def] euler_incr_fas_nth_def[abs_def] interpret_ode_d_fa) apply (simp add: ode_d1_eq[symmetric] ode_d_Suc_0_eq_ode_d2 inverse_eq_divide) apply (simp add: algebra_simps field_simps divide_simps) unfolding cx'_def[symmetric] cx_flow x0'_def h'_def apply (simp add: algebra_simps) done lemma map_rk2_fas_err_nth[simp]: "d = length ode_e \ length b = length ode_e \ map (rk2_fas_err_nth a b c e f) [0.. interpret_floatarith h vs" defines "s2' \ interpret_floatarith s2 vs" defines "rkp' \ interpret_floatarith rkp vs" defines "x0' \ einterpret x0 vs" defines "cx' \ einterpret cx vs" assumes cx_flow: "cx' = flow0 x0' (h' * s1')" assumes [simp]: "length x0 = DIM('a)" "length cx = DIM('a)" "DIM('a) \ length vs" assumes safes: "(x0'::'a)\ Csafe" "(cx'::'a)\ Csafe" "(x0' + (s2' * h' * rkp') *\<^sub>R ode x0'::'a)\ Csafe" shows "(einterpret (take DIM('a) (rk2_fas rkp x0 h cx s2)) vs::'a::executable_euclidean_space) = discrete_evolution (rk2_increment rkp' (\_. ode)) h' 0 x0' + (heun_remainder1 (flow0 x0') ode_na ode_d_na ode_d2_na 0 h' s1' - heun_remainder2 rkp' (flow0 x0') ode_na ode_d2_na 0 h' s2')" using safes using safe_length[OF safes(1)] safe_max_Var[OF safes(1)] apply (auto simp: discrete_evolution_def rk2_fas_def) apply (subst rk2_increment_rk2_fas_err[OF cx_flow[unfolded cx'_def x0'_def h'_def]]) subgoal by simp subgoal by simp subgoal by simp subgoal by (simp add: x0'_def) subgoal by (simp add: cx'_def) subgoal by (simp add: x0'_def s2'_def h'_def rkp'_def) subgoal using [[show_consts, show_sorts, show_types]] by (auto simp: x0'_def s2'_def h'_def rkp'_def rk2_increment euler_incr_fas_def euler_incr_fas_nth_def[abs_def] inverse_eq_divide) done lemma rk2_increment_rk2_fas2: fixes h s1 s2 rkp x0 cx vs defines "h' \ interpret_floatarith h vs" defines "s2' \ interpret_floatarith s2 vs" defines "rkp' \ interpret_floatarith rkp vs" defines "x0' \ einterpret x0 vs" defines "cx' \ einterpret cx vs" assumes cx_flow: "cx' = flow0 x0' (h' * s1')" assumes [simp]: "length x0 = DIM('a)" "length cx = DIM('a)" "DIM('a) \ length vs" assumes safes: "x0'\ Csafe" "cx'\ Csafe" "(x0' + (s2' * h' * rkp') *\<^sub>R ode x0')\ Csafe" shows "(einterpret (drop DIM('a) (rk2_fas rkp x0 h cx s2)) vs::'a::executable_euclidean_space) = (heun_remainder1 (flow0 x0') ode_na ode_d_na ode_d2_na 0 h' s1' - heun_remainder2 rkp' (flow0 x0') ode_na ode_d2_na 0 h' s2')" using safes apply (auto simp: discrete_evolution_def rk2_fas_def) apply (subst rk2_increment_rk2_fas_err[OF cx_flow[unfolded cx'_def x0'_def h'_def]]) subgoal by simp subgoal by simp subgoal by simp subgoal by (simp add: x0'_def) subgoal by (simp add: cx'_def) subgoal by (simp add: x0'_def s2'_def h'_def rkp'_def) subgoal by (auto simp: x0'_def s2'_def h'_def rkp'_def rk2_increment euler_incr_fas_def inverse_eq_divide) done subsubsection \safe set relation\ lemma mk_safe[le, refine_vcg]: "wd TYPE('a::executable_euclidean_space)\ mk_safe X \ SPEC (\R::'a set. R = X \ X \ Csafe)" unfolding mk_safe_def by refine_vcg lemma mk_safe_coll[le, refine_vcg]: "wd TYPE('a::executable_euclidean_space) \ mk_safe_coll X \ SPEC (\R::'a set. R = X \ X \ Csafe)" unfolding mk_safe_coll_def autoref_tag_defs by (refine_vcg FORWEAK_invarI[where I="\a b. X = \b \ a \ a \ Csafe"]) auto lemma ode_set_spec[THEN order.trans, refine_vcg]: assumes [refine_vcg]: "wd TYPE('a::executable_euclidean_space)" shows "ode_set X \ SPEC (\r. ode ` X \ (r::'a set))" using assms wdD[OF assms(1)] unfolding ode_set_def apply (refine_vcg) subgoal by (auto simp: env_len_def ode_slp_def) subgoal premises prems using prems(1,2,4-) by (auto simp: env_len_def eucl_of_list_prod ode_def) done lemmas fderiv[derivative_intros] = ode_has_derivative_safeI lemma fderiv2[derivative_intros]: "x \ Csafe \ (ode_d1 has_derivative ode_d2 x) (at x)" by (frule ode_d1_has_derivative_safeI) (simp add: ode_d_Suc_0_eq_ode_d2) lemma derivative_within_safe[derivative_intros]: "(g has_derivative h) (at x) \ (g has_derivative h) (at x within Csafe)" by (rule has_derivative_at_withinI) lemma cont_fderiv: "continuous_on (Csafe) ode_d1" by (rule has_derivative_continuous_on) (auto intro!: derivative_intros) lemmas cont_fderiv'[continuous_intros] = continuous_on_compose2[OF cont_fderiv] lemma continuous_on_ode1: "continuous_on (Csafe) (ode)" using fderiv by (auto intro!: has_derivative_continuous_on derivative_intros) lemma continuous_on_ode[continuous_intros]: "continuous_on s g \ (\x. x \ s \ (g x) \ Csafe) \ continuous_on s (\x. ode (g x))" using continuous_on_ode1 by (rule continuous_on_compose2) auto lemma fderiv'[derivative_intros]: assumes "(g has_derivative g' y) (at y within X)" assumes "(g y) \ Csafe" shows "((\y. ode (g y)) has_derivative (blinfun_apply (ode_d1 (g y)) \\ g') y) (at y within X)" - using diff_chain_within[OF assms(1) has_derivative_within_subset[OF fderiv]] assms(2) + using diff_chain_within[OF assms(1) has_derivative_subset[OF fderiv]] assms(2) by (simp add: o_def) lemma fderiv2'[derivative_intros]: assumes "(g has_derivative g' y) (at y within X)" assumes "(g y) \ Csafe" shows "((\y. ode_d1 (g y)) has_derivative (blinfun_apply (ode_d2 (g y)) \\ g') y) (at y within X)" - using diff_chain_within[OF assms(1) has_derivative_within_subset[OF fderiv2]] assms(2) + using diff_chain_within[OF assms(1) has_derivative_subset[OF fderiv2]] assms(2) by (simp add: o_def) subsubsection \step of Picard iteration\ lemma ncc_interval: "ncc {a .. b::'a::executable_euclidean_space} \ a \ b" by (auto simp: ncc_def) lemma nonempty_interval: "nonempty {a .. b::'a::executable_euclidean_space} \ a \ b" by (auto simp: nonempty_def) lemmas [refine_vcg] = Picard_step_def[THEN eq_refl, THEN order.trans] lemma Basis_list_zero_mem_Basis[simp]: "Basis_list ! 0 \ Basis" unfolding Basis_list[symmetric] apply (rule nth_mem) apply (rule length_Basis_list_pos) done lemma cfuncset_empty_iff: fixes l u::"'d::ordered_euclidean_space" shows "l \ u \ cfuncset l u X = {} \ X = {}" unfolding cfuncset_def Pi_def proof (safe, goal_cases) case hyps: (1 x) from \x \ X\ have "(\_. x) \ {f. \x. x \ {l..u} \ f x \ X} \ Collect (continuous_on {l..u})" by (auto intro!: continuous_on_const) then show ?case using hyps by auto qed auto lemma lv_ivl_sings: "lv_ivl [x] [y] = (\x. [x]) ` {x .. y}" apply (auto simp: lv_ivl_def) subgoal for x by (cases x) auto done lemma Picard_step_ivl_refine[le, refine_vcg]: assumes [refine_vcg]: "wd TYPE('a::executable_euclidean_space)" assumes "(X::'a set) \ Csafe" assumes "0 \ h" shows "Picard_step_ivl X0 t0 h X \ Picard_step X0 t0 h X" proof - have "h' \ {t0..t0 + h} \ t0 \ h'" for h' by simp then show ?thesis unfolding Picard_step_ivl_def Picard_step_def apply (refine_vcg, clarsimp_all simp del: atLeastAtMost_iff) subgoal using \0 \ h\ by simp subgoal by (auto simp: euler_incr_slp_def wdD) subgoal by (auto simp: euler_incr_fas'_def) subgoal for XS l u apply (auto simp: lv_ivl_sings nonempty_interval simp del: atLeastAtMost_iff intro!: add_integral_ivl_bound) subgoal for x0 h' phi x h'' apply (drule bspec, assumption) apply (drule bspec[where x="h'' - t0"], force) proof goal_cases case (1) have *: "map ((!) (list_of_eucl b)) [0.. Basis_list ! (DIM('a) - Suc 0)] = list_of_eucl b" for b::'a apply (auto intro!: nth_equalityI simp: nth_append not_less) using Intersection.le_less_Suc_eq by blast have "phi x \ X" if "x \ {t0 .. h''}" for x using 1 that by (auto simp: cfuncset_iff) have "x0 + (h'' - t0) *\<^sub>R ode b \ {l .. u}" if "b \ X" for b proof - from 1(17)[rule_format, OF that] assms(1) have "einterpret (euler_incr_fas' D) (list_of_eucl x0 @ (h'' - t0) # list_of_eucl b) \ eucl_of_list ` XS" by (auto simp: wdD) also have "eucl_of_list ` XS \ {l .. u}" by fact finally show ?thesis by (simp add: euler_incr_fas'_def einterpret_euler_incr_fas map_nth_append1 nth_append wdD[OF \wd _\] *) qed then have *: "(h'' - t0) *\<^sub>R ode b \ {l - x0..u - x0}" if "b \ X" for b using that by (auto simp: algebra_simps) show ?case apply (rule *) using 1 by (auto simp: cfuncset_iff) qed subgoal using assms(2) by (auto intro!: integrable_continuous_real continuous_intros simp: cfuncset_iff) done done qed subsubsection \Picard iteration\ lemma inf_le_supI[simp]: fixes a b c d::"'d::ordered_euclidean_space" shows "a \ c \ inf a b \ sup c d" "a \ d \ inf a b \ sup c d" "b \ c \ inf a b \ sup c d" "b \ d \ inf a b \ sup c d" by (auto simp: eucl_le[where 'a='d] eucl_inf[where 'a='d] eucl_sup[where 'a='d] inf_real_def sup_real_def intro!: sum_mono scaleR_right_mono) lemmas [refine_vcg_def] = do_widening_spec_def lemma P_iter_spec[le, refine_vcg]: assumes "PHI \ Csafe" assumes "0 \ h" assumes [refine_vcg]: "wd TYPE('a::executable_euclidean_space)" shows "P_iter X0 h i PHI \ SPEC (\r. case r of None \ True | Some (PHI'::'a set) \ nonempty PHI' \ compact PHI' \ (\PHI'' \ PHI'. RETURN (Some PHI'') \ Picard_step X0 0 h PHI'))" using assms[unfolded autoref_tag_defs] proof (induction i arbitrary: PHI) case 0 then show ?case unfolding P_iter.simps by refine_vcg next case (Suc i) show ?case unfolding P_iter.simps apply (refine_vcg Suc) subgoal by (auto simp: cfuncset_iff Picard_step_def algebra_simps add_increasing2) subgoal for lu l u b CX CX' lu' l' u' b' apply (simp add: nonempty_interval Picard_step_def) apply (safe intro!: exI[where x="{l .. u}"] compact_interval) subgoal by (auto simp: nonempty_interval) apply (rule subsetD[of CX' "{l .. u}"]) subgoal apply (rule order_trans, assumption) unfolding atLeastatMost_subset_iff apply (rule disjI2) apply (rule conjI) subgoal apply (rule order_trans[where y="inf l' l - (if b' then \l' - l\ else 0)"]) by (simp_all split: if_split_asm add: algebra_simps add_increasing2) subgoal apply (split if_split_asm) apply (rule order_trans[where y="sup u' u + \u' - u\"]) by (auto split: if_split_asm simp add: algebra_simps add_increasing2) done subgoal by force done done qed subsubsection \iterate step size\ lemma Ball_cfuncset_continuous_on: "\f\cfuncset a b X. continuous_on {a..b} f" by (simp add: cfuncset_iff) lemmas one_step_methodD = one_step_method_def[THEN iffD1, rule_format, le] lemmas one_step_methodI = one_step_method_def[THEN iffD2, rule_format] lemma cert_stepsize_lemma: assumes prems: " 0 < h" "X0 \ {l..u}" "Res \ Csafe" "Res_ivl \ Csafe" "{l..u} \ Csafe" "nonempty PHI'" "nonempty Res" "\x0\X0. x0 \ Csafe \ h \ existence_ivl0 x0 \ (\h'\{0..h}. flow0 x0 h' \ PHI') \ x0 + h *\<^sub>R ode x0 \ PHI' \ flow0 x0 h \ Res" "nonempty Res_ivl" "\x0\X0. x0 \ Csafe \ (\h\{0..h}. h \ existence_ivl0 x0 \ (\h'\{0..h}. flow0 x0 h' \ PHI') \ x0 + h *\<^sub>R ode x0 \ PHI' \ flow0 x0 h \ Res_ivl)" "compact PHI'" "PHI'' \ PHI'" "RETURN (Some PHI'') \ Picard_step X0 0 h PHI'" shows "flowpipe0 X0 h h Res_ivl Res" proof - from prems have Ps: "RETURN (Some PHI'') \ Picard_step X0 0 h PHI'" by simp from Ps have PHI': "compact PHI''" "PHI'' \ Csafe" "\x\PHI''. x \ Csafe" "\x0 h'' phi. x0 \ X0 \ 0 \ h'' \ h'' \ h \ phi \ cfuncset 0 h'' PHI' \ x0 + integral {0..h''} (\t. ode (phi t)) \ PHI''" by (auto simp: Picard_step_def) then obtain cx where cx: "(\t::real. cx) \ cfuncset 0 0 PHI'" using \PHI'' \ PHI'\ \nonempty PHI'\ by (auto simp: cfuncset_def continuous_on_const nonempty_def) show "flowpipe0 X0 h h Res_ivl Res" unfolding flowpipe0_def atLeastAtMost_singleton proof safe show "0 \ h" using \0 < h\ by simp show safe_X0: "x \ Csafe" if "x \ X0" for x using that \{l..u} \ Csafe\ \X0 \ {l..u}\ by blast show "x \ Csafe" if "x \ Res_ivl" for x using prems that by (auto simp: ) show "x \ Csafe" if "x \ Res" for x using prems that by (auto simp:) fix x0 assume "x0 \ X0" from PHI'(4)[OF \x0 \ X0\ order_refl \0 \ h\ cx] have "x0 \ PHI''" by simp have *: "h \ existence_ivl0 x0" "s \ {0 .. h} \ flow0 x0 s \ PHI''" for s using \x0 \ X0\ \PHI'' \ PHI'\ \0 \ h\ PHI'(3) \x0 \ PHI''\ by (auto simp: cfuncset_def Pi_iff closed_segment_eq_real_ivl ivl_integral_def intro!: Picard_iterate_mem_existence_ivlI[OF UNIV_I _ UNIV_I \compact PHI''\ \x0 \ PHI''\ \PHI'' \ Csafe\] PHI'(4)) force+ show h_ex: "h \ existence_ivl0 x0" by fact have cf: "(\t::real. x0) \ cfuncset 0 h PHI'" for h using \x0 \ PHI''\ \PHI'' \ PHI'\ by (auto simp: cfuncset_def continuous_intros) have mem_PHI': "x0 + h' *\<^sub>R ode x0 \ PHI'" if "0 \ h'" "h' \ h" for h' using that \PHI'' \ PHI'\ PHI'(4)[OF \x0 \ X0\ that cf] by auto from this prems safe_X0 show "flow0 x0 h \ Res" using \0 \ h\ h_ex * \PHI'' \ PHI'\ \x0 \ X0\ by (auto simp: subset_iff dest!: bspec[where x=x0]) fix h' assume h': "h' \ {0..h}" then have "h' \ existence_ivl0 x0" by (meson atLeastAtMost_iff existence_ivl_zero h_ex is_interval_1 local.is_interval_existence_ivl local.mem_existence_ivl_iv_defined(2)) from h' this prems safe_X0 show "flow0 x0 h' \ Res_ivl" using \0 < h\ h_ex * \PHI'' \ PHI'\ \x0 \ X0\ mem_PHI' \x0 \ PHI''\ by (auto simp: subset_iff dest!: bspec[where x=x0]) qed qed lemma cert_stepsize_spec[le,refine_vcg]: assumes "h > 0" assumes "one_step_method m" assumes [refine_vcg]: "wd TYPE('a::executable_euclidean_space)" shows "cert_stepsize m X0 h i n \ SPEC (\(h', RES::'a set, RES_ivl, _). nonempty RES \ nonempty RES_ivl \ 0 < h' \ h' \ h \ flowpipe0 X0 h' h' RES_ivl RES)" using assms(1)[unfolded autoref_tag_defs] proof (induction n arbitrary: h) case 0 then show ?case by simp next case (Suc n) note [refine_vcg] = Suc.IH[of "h/2", le] show ?case unfolding cert_stepsize.simps using Suc.prems by (refine_vcg Ball_cfuncset_continuous_on one_step_methodD[OF assms(2)]) (clarsimp_all simp: cert_stepsize_lemma) qed subsubsection \Euler step\ lemma embed_real_ivl_iff[simp]: "(\x \ {0 .. h *\<^sub>R (One::'a::executable_euclidean_space)}. P (x \ hd Basis_list)) \ (\x \ {0 .. h}. P x)" proof (auto simp: eucl_le[where 'a='a], goal_cases) case hyps: (1 x) have "x = x *\<^sub>R (One::'a) \ hd Basis_list" by auto also have "P \" apply (rule hyps[rule_format]) using hyps by (auto simp: eucl_le[where 'a='a]) finally show ?case . qed lemma convex_on_segmentI: assumes mem_convex: "convex C" "a \ C" "a + j *\<^sub>R b \ C" assumes "0 \ i" "i \ j" shows "a + i *\<^sub>R b \ C" proof - have "a + i *\<^sub>R b = (1 - i / j) *\<^sub>R a + (i / j) *\<^sub>R (a + j *\<^sub>R b)" using assms by (auto simp: algebra_simps diff_divide_distrib) also have "\ \ C" using assms by (auto simp: divide_simps intro!: convexD[OF mem_convex]) finally show ?thesis . qed lemma one_step_flowpipe: assumes [THEN one_step_methodD, refine_vcg]: "one_step_method m" assumes [refine_vcg]: "wd TYPE('a::executable_euclidean_space)" shows "one_step X0 h m \ SPEC (\(h', _, RES_ivl, RES::'a set). 0 < h' \ h' \ h \ flowpipe0 X0 h' h' RES_ivl RES)" using assms unfolding one_step_def by refine_vcg lemma ncc_imageD: assumes "ncc ((\x. x ! i) ` env)" assumes "nth_image_precond env i" shows "compact ((\x. x ! i) ` env::real set)" "closed ((\x. x ! i) ` env)" "bounded ((\x. x ! i) ` env)" "((\x. x ! i) ` env) \ {}" "convex ((\x. x ! i) ` env)" using assms by (auto simp: ncc_def nth_image_precond_def compact_eq_bounded_closed) lemma max_Var_floatariths_ode_d_fa[le]: assumes [simp]: "length ode_e > 0" "max_Var_floatariths ode_e \ length ode_e" "length cxs = length ode_e" "length ys = length ode_e" shows "max_Var_floatariths (ode_d_fa i cxs ys) \ max (max_Var_floatariths (cxs)) (max_Var_floatariths ys)" apply (auto simp: ode_d_fa_def max_Var_floatariths_Max ) using assms apply auto[1] apply (auto intro!: max_Var_floatarith_subst_floatarith_le max_Var_floatariths_ode_d_expr max_Var_floatarith_le_max_Var_floatariths_nthI max_Var_ode_fa simp: in_set_conv_nth) apply (auto simp: max_Var_floatariths_Max in_set_conv_nth) done lemma max_Var_floatariths_euler_err_fas[le]: assumes nz: "0 < length ode_e" and [simp]: "max_Var_floatariths ode_e \ length ode_e" "length xs = length ode_e" "length cxs = length ode_e" shows "max_Var_floatariths (euler_err_fas xs h cxs) \ max (max_Var_floatariths xs) (max (max_Var_floatarith h) (max_Var_floatariths cxs))" using nz by (auto simp: euler_err_fas_def[abs_def] euler_err_fas_nth_def[abs_def] map_nth_eq_self simp del: length_0_conv intro!: max_Var_floatariths_ode_d_fa max_Var_floatariths_map_times max_Var_floatariths_map_const max_Var_ode_fa; arith) lemma max_Var_floatariths_euler_incr_fas[le]: assumes [simp]: "max_Var_floatariths ode_e \ length ode_e" "length xs = length ode_e" "length cxs = length ode_e" shows "max_Var_floatariths (euler_incr_fas xs h cxs) \ max (max_Var_floatariths xs) (max (max_Var_floatarith h) (max_Var_floatariths cxs))" using length_ode_fa by (auto simp: euler_incr_fas_def euler_incr_fas_nth_def[abs_def] simp del: length_ode_fa intro!: max_Var_floatariths_ode_d_fa max_Var_floatariths_map_plus max_Var_floatariths_map_times max_Var_floatariths_map_const max_Var_ode_fa) lemma map_euler_incr_fas_nth: "length X0 = d \ map (euler_incr_fas_nth X0 h CX) [0.. map (euler_err_fas_nth X0 h CX) [0.. length ode_e" "length xs = length ode_e" "length cxs = length ode_e" assumes nz: "0 < length ode_e" shows "max_Var_floatariths (euler_fas xs h cxs) \ Max {max_Var_floatariths xs, max_Var_floatarith h, max_Var_floatariths cxs}" using nz by (auto simp: euler_fas_def map_euler_incr_fas_nth map_euler_err_fas_nth intro!: max_Var_floatariths_map_plus max_Var_floatariths_euler_incr_fas max_Var_floatariths_euler_err_fas) lemma take_interpret_floatariths: "d < length fas \ take d (interpret_floatariths fas vs) = interpret_floatariths (take d fas) vs" by (auto intro!: nth_equalityI) lemma length_euler_slp_le: "2 * D \ length euler_slp" by (auto simp: euler_fas'_def euler_slp_def intro!: order_trans[OF _ length_slp_of_fas_le]) lemma ncc_nonempty[simp]: "ncc x \ nonempty x" by (simp add: ncc_def nonempty_def) lemma nccD: assumes "ncc X" shows "compact X" "closed X" "bounded X" "X \ {}" "convex X" using assms by (auto simp: ncc_def nth_image_precond_def compact_eq_bounded_closed) lemma D_DIM_wdD[simp]: "wd TYPE('a::executable_euclidean_space) \ D = DIM('a)" by (auto simp: wdD) lemma euler_step_flowpipe: includes floatarith_notation assumes [refine_vcg]: "wd TYPE('a::executable_euclidean_space)" shows "euler_step X0 h \ SPEC (\(h', _, RES_ivl, RES::'a set). 0 < h' \ h' \ h \ flowpipe0 X0 h' h' RES_ivl RES)" unfolding euler_step_def THE_NRES_def apply (intro SPEC_rule_conjI one_step_flowpipe one_step_methodI) apply (refine_vcg, clarsimp_all) subgoal using assms by (auto simp: euler_slp_def euler_fas'_def) subgoal by (auto simp: euler_slp_def euler_fas'_def) subgoal using length_euler_slp_le assms by (auto simp: env_len_def wdD[OF \wd _\]) subgoal using length_euler_slp_le assms by (auto simp: env_len_def wdD[OF \wd _\]) proof (goal_cases) case hyps: (1 X0 CX hl hu env res b x0 enve h) then interpret derivative_on_prod "{0 .. h}" CX "\_. ode" "\(t, x). ode_d1 x o\<^sub>L snd_blinfun" by unfold_locales (auto intro!: continuous_intros derivative_eq_intros simp: split_beta' subset_iff wdD[OF \wd _\]) from \h \ existence_ivl0 x0\ have s_ex: "s \ existence_ivl0 x0" if "0 \ s" "s \ h" for s by (metis (no_types, lifting) atLeastAtMost_iff ivl_subset_existence_ivl subset_iff that) have "flow0 x0 (h) = flow0 x0 (0 + (h))" by simp also have "\ \ eucl_of_list ` take D ` env" using hyps apply (intro euler_consistent_traj_set[where x="flow0 x0" and u = "h"]) apply (auto intro!: \0 \ h\ flow_has_vector_derivative[THEN has_vector_derivative_at_within] simp: nccD discrete_evolution_def euler_increment subset_iff wdD[OF \wd _\] Let_def s_ex min_def max_def lv_ivl_sings) subgoal premises prems for s proof - have "interpret_floatariths (euler_fas' DIM('a)) (list_of_eucl x0 @ list_of_eucl (flow0 x0 s) @ [h]) \ env" using prems by (auto intro!: prems(1)[rule_format]) then have "eucl_of_list (take D (interpret_floatariths (euler_fas' DIM('a)) (list_of_eucl x0 @ list_of_eucl (flow0 x0 s) @ [h]))) \ eucl_of_list ` take D ` env" (is "eucl_of_list (take _ (interpret_floatariths _ ?vs)) \ _") by auto also have "take (2 * D) (interpret_floatariths (euler_fas' DIM('a)) ?vs) = interpret_floatariths (map fold_const_fa (euler_fas (map floatarith.Var [0..wd _\] simp del: map_map intro!: max_Var_floatariths_map_plus max_Var_floatariths_euler_incr_fas max_Var_floatariths_euler_err_fas \wd _\ max_Var_floatariths_fold_const_fa[le]) then have "take D (take (2 * D) (interpret_floatariths (euler_fas' DIM('a)) ?vs)) = take D (interpret_floatariths (euler_fas (map floatarith.Var [0..wd _\]) also have "eucl_of_list \ = x0 + h *\<^sub>R ode x0 + (h\<^sup>2 / 2) *\<^sub>R (ode_d 0 (flow0 x0 s) (ode (flow0 x0 s))) (ode (flow0 x0 s))" by (auto simp: take_interpret_floatariths einterpret_euler_fas1 map_nth_append1 prems nth_append wdD[OF \wd _\]) finally show ?thesis by (simp add: prems(10) prems(13) prems(14) prems(5) ode_d1_eq[symmetric] wdD[OF \wd _\]) qed done also have "\ \ res" using assms hyps by auto finally show ?case by simp qed (auto simp: assms) lemma length_rk2_slp_le: "2 * D \ length rk2_slp" by (auto simp: rk2_slp_def rk2_fas'_def intro!: order_trans[OF _ length_slp_of_fas_le]) lemma max_Var_floatarith_R\<^sub>e[simp]: "max_Var_floatarith (R\<^sub>e x) = 0" by (auto simp: R\<^sub>e_def split: prod.splits) lemma max_Var_floatariths_rk2_fas_err[le]: assumes nz: "0 < length ode_e" and [simp]: "max_Var_floatariths ode_e \ length ode_e" "length x0 = length ode_e" "length cx = length ode_e" shows "max_Var_floatariths (rk2_fas_err rkp x0 h cx s2) \ Max {max_Var_floatarith rkp, max_Var_floatariths x0, max_Var_floatarith h, max_Var_floatariths cx, max_Var_floatarith s2}" using nz unfolding rk2_fas_err_def rk2_fas_err_nth_def by (auto simp: rk2_fas_err_def intro!: max_Var_floatariths_append max_Var_floatariths_map_plus max_Var_floatariths_map_times max_Var_floatariths_map_const max_Var_ode_fa max_Var_floatariths_euler_incr_fas max_Var_floatariths_ode_d_fa; arith) lemma max_Var_floatarith_one[simp]: "max_Var_floatarith 1 = 0" and max_Var_floatarith_zero[simp]: "max_Var_floatarith 0 = 0" by (auto simp: one_floatarith_def zero_floatarith_def) lemma max_Var_floatariths_rk2_fas[le]: assumes nz: "0 < length ode_e" and [simp]: "max_Var_floatariths ode_e \ length ode_e" "length x0 = length ode_e" "length cx = length ode_e" shows "max_Var_floatariths (rk2_fas rkp x0 h cx s2) \ Max {max_Var_floatarith rkp, max_Var_floatariths x0, max_Var_floatarith h, max_Var_floatariths cx, max_Var_floatarith s2}" using nz by (auto simp: rk2_fas_def intro!: max_Var_floatariths_append max_Var_floatariths_map_plus max_Var_floatariths_map_times max_Var_floatariths_map_const max_Var_ode_fa max_Var_floatariths_euler_incr_fas max_Var_floatariths_rk2_fas_err) lemma rk2_step_flowpipe: includes floatarith_notation assumes [refine_vcg]: "wd TYPE('a::executable_euclidean_space)" shows "rk2_step X0 h \ SPEC (\(h', _, RES_ivl, RES::'a set). 0 < h' \ h' \ h \ flowpipe0 X0 h' h' RES_ivl RES)" unfolding rk2_step_def THE_NRES_def apply (intro one_step_flowpipe assms one_step_methodI) apply (refine_vcg, clarsimp_all) subgoal using assms by (auto simp: rk2_slp_def rk2_fas'_def) subgoal by (auto simp: rk2_slp_def rk2_fas'_def) subgoal using length_rk2_slp_le by (auto simp: env_len_def wdD[OF \wd _\]) subgoal using length_rk2_slp_le by (auto simp: env_len_def wdD[OF \wd _\]) proof (goal_cases) case hyps: (1 X0 CX hl hu rk2_param env res b x0 el h) from assms have "D = DIM('a)" by simp have aux: "ode (flow0 x0 s) = ode (snd (s, flow0 x0 s))" for s by simp from hyps interpret derivative_on_prod "{0 .. h}" CX "\_ x. ode x" "\(t, x). ode_d1 x o\<^sub>L snd_blinfun" by unfold_locales (auto intro!: continuous_intros derivative_eq_intros simp: split_beta' subset_iff) have aux2: "blinfun_apply (ode_d1 (snd tx)) \ snd = blinfun_apply (ode_d1 (snd tx) o\<^sub>L snd_blinfun)" for tx::"real\'a" by (auto intro!: blinfun_eqI) have aux3: "blinfun_apply (ode_d2 (snd tx)) (snd h) o\<^sub>L snd_blinfun = (flip_blinfun (flip_blinfun (ode_d2 (snd tx) o\<^sub>L snd_blinfun) o\<^sub>L snd_blinfun)) h" for tx h::"real\'a" by (auto intro!: blinfun_eqI) have "flow0 x0 (h) = flow0 x0 (0 + (h))" by simp also have "\ \ eucl_of_list ` take D ` env" using hyps assms apply (intro rk2_consistent_traj_set[where x="flow0 x0" and u = "h" and T="{0..h}" and X="CX" and p="rk2_param" and f = "ode_na" and f' = ode_d_na and g' = ode_d_na and f'' = ode_d2_na and g'' = ode_d2_na]) subgoal by (simp add: \0 \ h\) subgoal by simp subgoal by simp subgoal by auto subgoal by (auto simp add: ncc_def nonempty_def) subgoal apply (rule flow_has_vector_derivative[THEN has_vector_derivative_at_within, THEN has_vector_derivative_eq_rhs]) subgoal by (metis (no_types, lifting) ivl_subset_existence_ivl subset_iff) subgoal by (force simp: ode_na_def[abs_def] ode_d_na_def[abs_def] ode_d2_na_def[abs_def]) done subgoal unfolding ode_na_def ode_d_na_def ode_d2_na_def apply (rule derivative_eq_intros) apply (rule derivative_intros) apply (rule derivative_intros) subgoal by (auto simp: ncc_def nonempty_def) subgoal by force done subgoal unfolding ode_na_def ode_d_na_def ode_d2_na_def apply (rule derivative_eq_intros) apply (rule derivative_intros) apply (rule derivative_intros) apply (rule derivative_intros) subgoal by (force simp: nonempty_def) apply (rule derivative_intros) subgoal by (auto intro!: aux3) done subgoal by (rule refl) subgoal by (rule refl) subgoal apply (rule compact_imp_bounded) apply (rule compact_continuous_image) subgoal by (auto intro!: continuous_intros simp: ode_na_def ode_d_na_def ode_d2_na_def) subgoal by (auto simp: ncc_def intro!: compact_Times) done subgoal by auto subgoal by simp subgoal by simp subgoal apply (rule convex_on_segmentI[where j=h]) using mult_left_le_one_le[of h "rk2_param"] by (auto simp: ncc_def mult_left_le_one_le mult_le_one ac_simps ode_na_def ode_d_na_def ode_d2_na_def dest: bspec[where x=0]) subgoal by (simp add: ncc_def) subgoal by (simp add: ncc_def compact_imp_closed) subgoal for s1 s2 apply (clarsimp simp add: lv_ivl_sings) subgoal premises prems proof - have "s2 * rk2_param * h \ h" apply (rule mult_left_le_one_le) using assms prems by (auto intro!: mult_le_one) then have s2: "(s2 * h * rk2_param) \ {0 .. h}" using prems assms by (auto simp: ac_simps) have s1: "h * s1 \ {0 .. h}" using prems by (auto intro!: mult_right_le_one_le) then have "interpret_floatariths (rk2_fas' D) (list_of_eucl x0 @ list_of_eucl (flow0 x0 (h * s1)) @ [rk2_param, h, s2]) \ env" unfolding \D = _\ using prems by (intro prems(17)[rule_format]) auto then have "take (2 * D) (interpret_floatariths (rk2_fas' D) (list_of_eucl x0 @ list_of_eucl (flow0 x0 (h * s1)) @ [rk2_param, h, s2])) \ take (2 * D) ` env" (is "?l \ _") by auto also have "?l = interpret_floatariths (map fold_const_fa (rk2_fas (Var (2 * D)) (map floatarith.Var [0..wd _\]) finally have "take D (interpret_floatariths ?fas ?xs) \ take D ` take (2 * D) ` env" by auto also have "\ = take D ` env" by (auto simp: image_image wdD[OF \wd _\]) finally have "eucl_of_list (take D (interpret_floatariths ?fas ?xs)) \ eucl_of_list ` take D ` env" by simp then have "einterpret (take D ?fas) ?xs \ eucl_of_list ` take D ` env" by (simp add: take_interpret_floatariths wdD[OF \wd _\]) also have "einterpret (take D ?fas) ?xs = discrete_evolution (rk2_increment (rk2_param) (\t x. ode_na (t, x))) h 0 x0 + heun_remainder1 (flow0 x0) ode_na ode_d_na ode_d2_na 0 h s1 - heun_remainder2 (rk2_param) (flow0 x0) ode_na ode_d2_na 0 h s2" apply (simp add: wdD[OF \wd _\]) apply (subst rk2_increment_rk2_fas1[where ?s1'.0 = s1]) subgoal by (auto simp: nth_append map_nth_append1) subgoal by auto subgoal by auto subgoal by auto subgoal by (auto simp: nth_append map_nth_append1 \x0 \ Csafe\) subgoal apply (auto simp: nth_append map_nth_append1 \x0 \ Csafe\) by (meson connectedD_interval existence_ivl_zero flow0_defined hyps mult_right_le_one_le mult_sign_intros(1) mvar.connected prems) subgoal proof - have "x0 + ((rk2_param * s2) * h) *\<^sub>R ode x0 \ CX" by (rule convex_on_segmentI[where j=h]) (use prems in \auto simp: ncc_def mult_left_le_one_le mult_le_one dest: bspec[where x=0]\) also note \\ \ Csafe\ finally show ?thesis by (auto simp: nth_append map_nth_append1 \x0 \ Csafe\ ac_simps) qed subgoal by (auto simp: nth_append map_nth_append1 ode_na_def) done finally show ?thesis by (simp add: \D = _\) qed done done also have "\ \ res" using hyps(6) by (simp add: \D = _\) finally show ?case by (simp add: \D = _\) qed lemma interpret_adapt_stepsize_fa: "interpret_floatarith (adapt_stepsize_fa rtol m_id e h') [] = float_of h' * (float_of(rtol) / float_of e) powr (1 / (float_of (m_id) + 1))" by (auto simp: inverse_eq_divide adapt_stepsize_fa_def) lemma choose_step_flowpipe[le, refine_vcg]: assumes "wd TYPE('a::executable_euclidean_space)" shows "choose_step X0 h \ SPEC (\(h', _, RES_ivl, (RES::'a set)). 0 < h' \ h' \ h \ flowpipe0 X0 h' h' RES_ivl RES)" using assms unfolding choose_step_def by (refine_vcg rk2_step_flowpipe euler_step_flowpipe) lemma CsafeI: "t \ existence_ivl0 x \ x \ Csafe" using local.mem_existence_ivl_iv_defined(2) by blast lemma apply_vareq: "blinfun_apply (vareq x t) = ode_d1 (flow0 x t)" by (auto simp: vareq_def) lemma Dflow_has_derivative: "t \ existence_ivl0 x \ (Dflow x has_derivative blinfun_scaleR_left (ode_d1 (flow0 x t) o\<^sub>L Dflow x t)) (at t)" by (auto simp: Dflow_def blinfun.bilinear_simps scaleR_blinfun_compose_left apply_vareq CsafeI intro!: derivative_eq_intros mvar.flow_has_derivative[THEN has_derivative_eq_rhs] ext blinfun_eqI) lemma matrix_scaleR: "matrix (blinfun_apply (h *\<^sub>R X)) = h *\<^sub>R matrix X" by (vector matrix_def blinfun.bilinear_simps) lemma blinfun_of_vmatrix_matrix_matrix_mult[simp]: "blinfun_of_vmatrix (A ** B) = blinfun_of_vmatrix A o\<^sub>L blinfun_of_vmatrix B" including blinfun.lifting by transfer (auto simp: o_def matrix_vector_mul_assoc) lemma blinfun_of_vmatrix_mat_1[simp]: "blinfun_of_vmatrix (mat 1) = 1\<^sub>L" including blinfun.lifting by transfer (auto simp: matrix_vector_mul_lid) lemma blinfun_of_vmatrix_matrix[simp]: "blinfun_of_vmatrix (matrix (blinfun_apply A)) = A" including blinfun.lifting by transfer (auto simp: bounded_linear.linear matrix_works) lemma inner_Basis_eq_vec_nth: "b \ Basis \ v \ b = vec_nth v (enum_class.enum ! index Basis_list b)" by (auto simp: inner_vec_def vec_nth_Basis if_distrib Basis_vec_def axis_eq_axis index_Basis_list_axis1 cong: if_cong) lemma intersects_sctns_spec_nres[le, refine_vcg]: "intersects_sctns X' sctns \ intersects_sctns_spec X' sctns" unfolding intersects_sctns_spec_def intersects_sctns_def by refine_vcg auto lemma intersects_sections_spec_clw_ref[le, refine_vcg]: "intersects_sctns_spec_clw R sctns \ intersects_sctns_spec R sctns" unfolding intersects_sctns_spec_def intersects_sctns_spec_clw_def by (refine_vcg FORWEAK_mono_rule[where I="\S b. \b \ \S \ \(plane_of ` sctns) = {}"]) auto lemma eq_nth_iff_index: "distinct xs \ n < length xs \ i = xs ! n \ index xs i = n" using index_nth_id by fastforce lemma max_Var_floatariths_ode_e_wd: assumes wd: "wd (TYPE('n::enum rvec))" assumes "CARD('n) \ K" shows "max_Var_floatariths ode_e \ K" using wdD[OF wd] assms by auto lemma nonzero_component[le, refine_vcg]: "nonzero_component s X n \ SPEC (\_. \b\X. b \ n \ 0)" unfolding nonzero_component_def by refine_vcg auto lemma interpret_slp_env_lenD: assumes "\cx\CX. interpret_slp (slp_of_fas (fas)) (env cx) \ R" assumes "cx \ CX" shows "interpret_floatariths fas (env cx) \ take (length fas) ` R" proof - from slp_of_fas have "take (length fas) (interpret_slp (slp_of_fas fas) (env cx)) = interpret_floatariths fas (env cx)" by auto moreover from assms(1)[rule_format, OF \cx \ CX\] have "interpret_slp (slp_of_fas fas) (env cx) \ R" by auto ultimately show ?thesis by force qed lemma flowpipe0_imp_flowpipe: assumes "flowpipe0 (fst ` X0) x1 x1 aba bba" shows "flowpipe X0 x1 x1 (aba \ UNIV) (bba \ UNIV)" using assms by (auto simp: flowpipe0_def flowpipe_def) lemma disjoints_spec[le, refine_vcg]: "disjoints_spec X Y \ SPEC (\b. b \ (X \ Y = {}))" unfolding disjoints_spec_def autoref_tag_defs by refine_vcg auto lemma inner_eq_zero_abs_BasisI: "\y\ \ Basis \ b \ Basis \ \y\ \ b \ y \ b = 0" for b y::"'a::executable_euclidean_space" by (metis abs_inner inner_Basis linorder_not_le order_refl zero_less_abs_iff) lemma abs_in_Basis_absE: fixes x y::"'a::executable_euclidean_space" assumes "abs y \ Basis" obtains "abs y = y" | "abs y = -y" proof - have "abs y = (\i\Basis. (abs (y \ i)) *\<^sub>R i)" by (simp add: euclidean_representation abs_inner[symmetric] assms) also have "Basis = insert (abs y) (Basis - {abs y})" using assms by auto also have "(\i\insert \y\ (Basis - {\y\}). \y \ i\ *\<^sub>R i) = \y \ \y\\ *\<^sub>R \y\" apply (subst sum.insert) using assms by (auto simp: abs_inner[symmetric] inner_Basis if_distribR if_distrib cong: if_cong) finally have "\y\ = \y \ \y\\ *\<^sub>R \y\" by simp moreover have "\ = y \ \ = - y" using assms by (auto simp: abs_real_def algebra_simps intro!: euclidean_eqI[where 'a='a] simp: inner_Basis inner_eq_zero_abs_BasisI split: if_splits) ultimately consider "\y\ = y" | "\y\ = - y" by auto then show ?thesis by (cases; rule that) qed lemma abs_in_BasisE: fixes x y::"'a::executable_euclidean_space" assumes "abs y \ Basis" obtains i where "i \ Basis" "y = i" | i where "i \ Basis" "y = -i" proof - from abs_in_Basis_absE[OF assms] consider "\y\ = y" | "\y\ = - y" by auto then show ?thesis proof cases case 1 with assms have "abs y \ Basis" "y = abs y" by auto then show ?thesis .. next case 2 with assms have "abs y \ Basis" "y = - abs y" by auto then show ?thesis .. qed qed lemma subset_spec_plane[le, refine_vcg]: "subset_spec_plane X sctn \ SPEC (\b. b \ X \ plane_of sctn)" unfolding subset_spec_plane_def by (refine_vcg) (auto simp: plane_of_def eucl_le[where 'a='a] dest!: bspec elim!: abs_in_BasisE) lemma subset_spec_coll_refine[le, refine_vcg]: "subset_spec_coll X Y \ subset_spec X Y" unfolding subset_spec_coll_def autoref_tag_defs subset_spec_def by (refine_vcg FORWEAK_mono_rule[where I="\X b. b \ \X \ Y"]) auto lemma eventually_in_planerectI: fixes n::"'a::executable_euclidean_space" assumes "abs n \ Basis" assumes "{l .. u} \ plane n c" "l \ u" assumes "\i. i \ Basis \ i \ abs n \ l \ i < x \ i" assumes "\i. i \ Basis \ i \ abs n \ x \ i < u \ i" shows "\\<^sub>F x in at x within plane n c. x \ {l .. u}" proof - have "\\<^sub>F x in at x within plane n c. x \ plane n c" unfolding eventually_at_filter by simp then have "\\<^sub>F x in at x within plane n c. l \ abs n \ x \ abs n \ x \ abs n \ u \ abs n" apply eventually_elim using assms(1,2,3) by (auto simp: elim!: abs_in_BasisE) moreover { fix i assume that: "i \ Basis" "i \ abs n" have "\\<^sub>F x in at x within plane n c. l \ i < x \ i" "\\<^sub>F x in at x within plane n c. x \ i < u \ i" by (auto intro!: order_tendstoD assms tendsto_eq_intros that) then have "\\<^sub>F x in at x within plane n c. l \ i < x \ i \ x \ i < u \ i" by eventually_elim auto } then have "\\<^sub>F x in at x within plane n c. \i \ Basis - {abs n}. l \ i < x \ i \ x \ i < u \ i" by (auto intro!: eventually_ball_finite) then have "\\<^sub>F x in at x within plane n c. \i \ Basis - {abs n}. l \ i \ x \ i \ x \ i \ u \ i" by eventually_elim (auto intro!: less_imp_le) ultimately have "\\<^sub>F x in at x within plane n c. \i\Basis. l \ i \ x \ i \ x \ i \ u \ i" by eventually_elim auto then show ?thesis by eventually_elim (auto simp: eucl_le[where 'a='a]) qed lemma mem_ivl_euclI: "k \ {c..d::'x::ordered_euclidean_space}" if "\i. i \ Basis \ c \ i \ k \ i" "\i. i \ Basis \ k \ i \ d \ i" using that by (auto simp: eucl_le[where 'a='x]) lemma op_eventually_within_sctn[le, refine_vcg]: "op_eventually_within_sctn X sctn S \ SPEC (\b. b \ (\x \ X. x \ S \ (\\<^sub>F x in at x within plane_of sctn. x \ S)))" unfolding op_eventually_within_sctn_def apply refine_vcg unfolding plane_of_def autoref_tag_defs apply (safe intro!: eventually_in_planerectI mem_ivl_euclI) subgoal premises prems for a b c d e f g h i j k B proof cases assume "B = \normal sctn\" moreover have "c \ plane (normal sctn) (pstn sctn)" "k \ plane (normal sctn) (pstn sctn)" using prems by auto ultimately show "c \ B \ k \ B" using \\normal sctn\ \ set Basis_list\ by (auto simp: elim!: abs_in_Basis_absE) next assume B: "B \ \normal sctn\" have "k \ B \ {g \ B .. h \ B}" using \k \ X\ \X \ {g..h}\ \B \ Basis\ by (auto simp: eucl_le[where 'a='a]) with B prems show ?thesis by (auto dest!: bspec elim!: abs_in_Basis_absE) qed subgoal premises prems for a b c d e f g h i j k B proof cases assume "B = \normal sctn\" moreover have "d \ plane (normal sctn) (pstn sctn)" "k \ plane (normal sctn) (pstn sctn)" using prems by auto ultimately show "d \ B \ k \ B" using \\normal sctn\ \ set Basis_list\ by (auto simp: elim!: abs_in_Basis_absE) qed (use prems in \auto elim!: abs_in_BasisE simp: eucl_le[where 'a='a] dest!: bspec subsetD\) subgoal by simp subgoal using [[simproc del: defined_all]] by (auto elim!: abs_in_BasisE simp: eucl_le[where 'a='a] dest!: bspec subsetD cong del: image_cong_simp) subgoal using [[simproc del: defined_all]] by (auto elim!: abs_in_BasisE simp: eucl_le[where 'a='a] dest!: bspec subsetD cong del: image_cong_simp) done lemma Let_unit: "Let (x::unit) f = f ()" by auto lemma CHECK_no_text: "CHECKs (x#ys) a = CHECKs [] a" by auto lemma frontier_above_halfspace: "normal sctn \ 0 \ frontier (above_halfspace sctn) = plane_of sctn" using frontier_halfspace_ge[of "normal sctn" "pstn sctn"] by (auto simp: halfspace_simps plane_of_def inner_commute) lemma flowpipe_subset: assumes "flowpipe X0 hl hu CX X1" and subs: "Y0 \ X0" "hl \ il" "il \ iu" "iu \ hu" "CX \ CY" "X1 \ Y1" and safe: "fst ` CY \ fst ` Y1 \ Csafe" shows "flowpipe Y0 il iu CY Y1" proof - from assms(1) have fp: "0 \ hl" "hl \ hu" "fst ` X0 \ Csafe" "fst ` CX \ Csafe" "fst ` X1 \ Csafe" "\(x0, d0)\X0. \h\{hl..hu}. h \ existence_ivl0 x0 \ (flow0 x0 h, Dflow x0 h o\<^sub>L d0) \ X1 \ (\h'\{0..h}. (flow0 x0 h', Dflow x0 h' o\<^sub>L d0) \ CX)" by (auto simp: flowpipe_def) then show ?thesis unfolding flowpipe_def apply safe subgoal using subs by auto subgoal using subs by auto subgoal using subs by fastforce subgoal using safe by auto subgoal using safe by auto subgoal using subs by fastforce subgoal using subs fp by fastforce subgoal for x0 d0 h h' using subs fp apply - apply (rule subsetD, assumption) apply (drule bspec) apply (rule subsetD; assumption) apply safe apply (drule bspec[where x=h], force) apply auto done done qed lemma poincare_mapsto_unionI: assumes "poincare_mapsto P r U t d" assumes "poincare_mapsto P s U u e" shows "poincare_mapsto P (r \ s) U (t \ u) (d \ e)" using assms apply (auto simp: poincare_mapsto_def) subgoal apply (drule bspec, assumption) by auto subgoal by fastforce done lemma sabove_not_le_halfspace: "x \ sabove_halfspace sctn \ \ le_halfspace sctn x" by (auto simp: sabove_halfspace_def le_halfspace_def gt_halfspace_def) lemma (in c1_on_open_euclidean) flowsto_self:\ \TODO: move!\ "0 \ T \ X0 \ Z \ fst ` X0 \ X \ flowsto X0 T Y Z" by (force simp: flowsto_def intro!: bexI[where x=0]) lemma (in c1_on_open_euclidean) flowpipe_imp_flowsto:\ \TODO: move!\ assumes "flowpipe X0 hl hu CX Y" "hl > 0" shows "flowsto X0 {0<..hl} CX Y" using assms by (fastforce simp: flowsto_def flowpipe_def open_segment_eq_real_ivl dest: bspec[where x=hl] intro!: bexI[where x=hl]) lemma flowsto_source_unionI: "flowsto X0 T A B \ flowsto Z T A B \ flowsto (X0 \ Z) T A B" by (fastforce simp: flowsto_def dest: bspec) lemma poincare_mapsto_subset: "poincare_mapsto P X0 U CX1 X1 \ X0' \ X0 \ X1 \ X2 \ CX1 \ CX2 \ fst ` X2 \ Csafe \ poincare_mapsto P X0' U CX2 X2" by (force simp: poincare_mapsto_def) lemma PDP_abs_lemma: fixes n::"'a::executable_euclidean_space" assumes "abs n \ Basis" shows "(x, d - (blinfun_scaleR_left (f (x)) o\<^sub>L (blinfun_scaleR_left (inverse (f x \ n)) o\<^sub>L (blinfun_inner_left n o\<^sub>L d)))) = (x, d - (blinfun_scaleR_left (f (x)) o\<^sub>L (blinfun_scaleR_left (inverse (f x \ (abs n))) o\<^sub>L (blinfun_inner_left (abs n) o\<^sub>L d))))" proof - consider "n \ Basis" | "- n \ Basis" using abs_in_Basis_absE[OF assms] assms by metis then show ?thesis proof cases case 1 then show ?thesis by simp next case 2 define i where "i = -n" with 2 have "i \ Basis" "n = -i" by (auto simp: ) then show ?thesis by (auto simp: inverse_eq_divide intro!: blinfun_eqI blinfun.bilinear_simps euclidean_eqI[where 'a='a]) qed qed lemma abs_in_BasisI: "\n\ \ Basis" if n: "n \ Basis \ - n \ Basis" for n::"'a::executable_euclidean_space" proof - consider "n \ Basis" | "- n \ Basis" using n by auto then show ?thesis proof cases case 1 then show ?thesis by simp next case 2 define i where "i = -n" with 2 have "i \ Basis" "n = -i" by (auto simp: ) then show ?thesis by (auto simp: inverse_eq_divide intro!: blinfun_eqI blinfun.bilinear_simps euclidean_eqI[where 'a='a]) qed qed lemma flowsto_poincareD: assumes f: "flowsto X0 T CX X1" assumes X1: "fst ` X1 \ P" assumes P: "(P \ UNIV) \ CX = {}" "closed P" assumes pos: "\t. t \ T \ t > 0" assumes x0: "x0 \ fst ` X0" assumes "fst ` X1 \ K" shows returns_to_flowstoI: "returns_to P x0" and poincare_map_mem_flowstoI: "poincare_map P x0 \ K" proof - from x0 obtain d where x0d: "(x0, d) \ X0" by auto from flowstoE[OF f x0d] obtain h where h: "h \ T" "h \ existence_ivl0 x0" "(flow0 x0 h, Dflow x0 h o\<^sub>L d) \ X1" and CX: "\h'. h' \ {0<-- (flow0 x0 h', Dflow x0 h' o\<^sub>L d) \ CX" by auto have "h > 0" by (auto intro!: pos h) have "flow0 x0 h \ P" using X1 h by auto have "\\<^sub>F t in at_right 0. t \ {0<..0 < h\, of "{0<..}"] by (auto simp: eventually_at_filter) then have "\\<^sub>F t in at_right 0. flow0 x0 t \ fst ` CX" by eventually_elim (use CX \0 < h\ open_segment_eq_real_ivl in auto) then have evnP: "\\<^sub>F t in at_right 0. flow0 x0 t \ P" by eventually_elim (use P in force) from \h > 0\ h(2) \flow0 x0 h \ P\ evnP P(2) show "returns_to P x0" by (rule returns_toI) have nin_P: "0 < s \ s < h \ flow0 x0 s \ P" for s using CX[of s] P by (auto simp: open_segment_eq_real_ivl) have "return_time P x0 = h" using h X1 by (auto intro!: return_time_eqI \0 < h\ h assms simp: nin_P) then have "poincare_map P x0 = flow0 x0 h" by (auto simp: poincare_map_def) also have "\ \ fst ` X1" using h by auto also note \_ \ K\ finally show "poincare_map P x0 \ K" . qed lemma inner_abs_Basis_eq_zero_iff: "abs n \ Basis \ x \ \n\ = 0 \ x \ n = 0" for n::"'a::executable_euclidean_space" by (auto simp: elim!: abs_in_BasisE) lemmas [simp] = sbelow_halfspaces_insert lemma Int_Un_eq_emptyI: "a \ (b \ c) = {}" if "a \ b = {}" "a \ c = {}" using that by auto lemma cancel_times_UNIV_subset: "A \ UNIV \ B \ UNIV \ A \ B" by auto lemma split_spec_coll_spec[le,refine_vcg]: "split_spec_coll X \ SPEC (\(A, B). X \ A \ B)" unfolding split_spec_coll_def by (refine_vcg) lemma Un_snd_sing_Pair_eq: "(e, f) \ a \ f \ (\x\a - {(e, f)}. snd x) = (\x\a. snd x)" by force lemma let_unit: "Let X y = y ()" by simp lemma (in c1_on_open_euclidean) flowpipe_imp_flowsto_nonneg:\ \TODO: move!\ assumes "flowpipe X0 hl hu CX Y" shows "flowsto X0 {0..} CX Y" using assms by (fastforce simp: flowsto_def flowpipe_def open_segment_eq_real_ivl dest: bspec[where x=hl] intro!: bexI[where x=hl]) lemma subset_DiffI: "A \ B \ A \ C = {} \ A \ B - C" by auto lemma flowsto_source_Union: "flowsto (\x\R. X0 x) T CX X1" if "\x. x \ R \ flowsto (X0 x) T CX X1" using that by (auto simp: flowsto_def) lemma times_subset_iff: "A \ B \ C \ E \ A = {} \ B = {} \ A \ C \ B \ E" by auto lemma flowsto_UnionE: assumes "finite Gs" assumes "flowsto X T CX (\Gs)" obtains XGs where "\X G. (X, G) \ XGs \ flowsto X T CX G" "Gs = snd ` XGs" "X = \(fst ` XGs)" apply atomize_elim using assms proof (induction arbitrary: X) case empty then show ?case by auto next case (insert x F) from insert.prems obtain X1 X2 where X: "X = X1 \ X2" and X1: "flowsto X1 T CX x" and X2: "flowsto X2 T CX (\F)" by (auto elim!: flowsto_unionE) from insert.IH[OF X2] obtain XGs where XGs: "\X G. (X, G) \ XGs \ flowsto X T CX G" "F = snd ` XGs" "X2 = (\a\XGs. fst a)" by auto then show ?case using X X1 X2 by (intro exI[where x="insert (X1, x) XGs"]) auto qed lemma flowsto_Union_funE: assumes "finite Gs" assumes "flowsto X T CX (\Gs)" obtains f where "\G. G \ Gs \ flowsto (f G) T CX G" "X = \(f ` Gs)" apply atomize_elim using assms proof (induction arbitrary: X) case empty then show ?case by auto next case (insert x F) from insert.prems obtain X1 X2 where X: "X = X1 \ X2" and X1: "flowsto X1 T CX x" and X2: "flowsto X2 T CX (\F)" by (auto elim!: flowsto_unionE) from insert.IH[OF X2] obtain f where f: "\G. G \ F \ flowsto (f G) T CX G" "X2 = (\a\F. f a)" by auto then show ?case using X X1 X2 insert.hyps by (intro exI[where x="f (x:=X1)"]) (auto split: if_splits) qed lemma flowsto_Union_Un_funE: assumes "flowsto X T CX (\Gs \ trap)" assumes "finite Gs" "Gs \ {}" obtains f where "\G. G \ Gs \ flowsto (f G) T CX (G \ trap)" "X = \(f ` Gs)" proof - from assms have "flowsto X T CX (\g \ Gs. g \ trap)" by auto from flowsto_Union_funE[OF finite_imageI[OF \finite Gs\] this] obtain f where f: "\G. G \ (\g. g \ trap) ` Gs \ flowsto (f G) T CX G" "X = \ (f ` ((\g. g \ trap) ` Gs))" by auto define f' where "f' g = f (g \ trap)" for g have "G \ Gs \ flowsto (f' G) T CX (G \ trap)" for G using f(1)[of "G \ trap"] by (auto simp: f'_def) moreover have "X = \(f' ` Gs)" using f by (auto simp: f'_def) ultimately show ?thesis .. qed lemma flowsto_Diff_to_Union_funE: assumes "flowsto (X - trap) T CX (\Gs)" assumes "finite Gs" obtains f where "\G. G \ Gs \ flowsto (f G - trap) T CX G" "Gs \ {} \ X = \(f ` Gs)" apply atomize_elim using assms(2,1) proof (induction arbitrary: X) case empty then show ?case by simp next case (insert x F) from insert.prems obtain X1 X2 where X: "X - trap = X1 \ X2" and X1: "flowsto (X1) T CX x" and X2: "flowsto X2 T CX (\F)" by (auto elim!: flowsto_unionE) then have "X1 = X1 - trap" "X2 = X2 - trap" by auto then have X1': "flowsto (X1 - trap) T CX x" and X2': "flowsto (X2 - trap) T CX (\F)" using X1 X2 by auto from insert.IH[OF X2'] obtain f where f: "\G. G \ F \ flowsto (f G - trap) T CX G" "F \ {} \ X2 = (\a\F. f a)" by auto show ?case apply (cases "F = {}") subgoal using f X X1 X2 X1' X2' insert.hyps insert.prems by auto subgoal apply (rule exI[where x="f (x:=X1 \ (trap \ X))"]) apply auto subgoal using X1 by (rule flowsto_subset) auto subgoal using X X1 X2 insert.hyps by auto subgoal using f X X1 X2 insert.hyps by auto subgoal using f X X1 X2 insert.hyps by auto subgoal using f X X1 X2 X1' X2' insert.hyps insert.prems by auto subgoal using f X X1 X2 X1' X2' insert.hyps insert.prems insert by auto subgoal using f X X1 X2 insert.hyps by (auto split: if_splits) done done qed lemma refine_case_list[refine_vcg]: assumes "xs = [] \ f \ SPEC P" assumes "\y ys. xs = y # ys \ g y ys \ SPEC P" shows "(case xs of [] \ f | (x#xs) \ g x xs) \ SPEC P" using assms by (auto split: list.splits) lemma flowsto_stays_sbelow: assumes "flowsto X0 {0<..} CXS X1" assumes "fst ` X0 \ below_halfspace sctn" assumes "\x d. (x, d) \ CXS \ ode x \ normal sctn < 0" shows "flowsto X0 {0<..} (CXS \ sbelow_halfspace sctn \ UNIV) X1" unfolding flowsto_def proof safe fix x d assume "(x, d) \ X0" with assms obtain t where "t>0" "t \ existence_ivl0 x" "(\s\{0<..L d) \ CXS)" "(flow0 x t, Dflow x t o\<^sub>L d) \ X1" by (auto simp: flowsto_def subset_iff open_segment_eq_real_ivl) moreover have "\s\{0<.. sbelow_halfspace sctn" proof (rule ccontr, clarsimp) fix s assume s: "flow0 x s \ sbelow_halfspace sctn" "0 < s" "s < t" let ?f = "\t. flow0 x t \ normal sctn - pstn sctn" let ?f' = "\t dx. dx * (ode (flow0 x t) \ normal sctn)" have "\xa\{0<..0 < s\, of ?f ?f']) (use ivl_subset_existence_ivl[OF \t \ existence_ivl0 x\] \0 < s\ \s < t\ in \auto intro!: continuous_intros derivative_eq_intros flow_has_derivative simp: flowderiv_def blinfun.bilinear_simps\) then obtain s' where "?f s - ?f 0 = ?f' s' (s - 0)" "0 < s'" "s' < s" by (auto simp: algebra_simps) note this(1) also have "(flow0 x s', Dflow x s' o\<^sub>L d )\ CXS" using \0 < s'\ \\s\{0<.. CXS\ \s < t\ \s' < s\ by auto then have "?f' s' (s - 0) < 0" using assms \(x, d) \ X0\ \0 < s\ by (auto simp: flowsto_def halfspace_simps algebra_simps subset_iff intro!: mult_pos_neg) finally have 1: "?f s < ?f 0" by simp also have "?f 0 \ 0" using \(x, d) \ X0\ assms mem_existence_ivl_iv_defined[OF \t \ existence_ivl0 _\] by (auto simp: halfspace_simps subset_iff) finally have "?f s < 0" . moreover from s have "0 \ ?f s" by (auto simp: halfspace_simps not_less) ultimately show False by simp qed ultimately show "\t\{0<..}. t \ existence_ivl0 x \ (flow0 x t, Dflow x t o\<^sub>L d) \ X1 \ (\s\{0<--L d) \ CXS \ sbelow_halfspace sctn \ UNIV)" by (auto intro!: simp: open_segment_eq_real_ivl) qed lemma poincare_mapsto_Union: "poincare_mapsto P (\xa) S CXS PS" if "\x. x \ xa \ poincare_mapsto P x S CXS PS" by (force simp: poincare_mapsto_def dest!: that) lemma diff_subset: "(\x\xa. f x) - (\x\xa. g x) \ (\x\xa. f x - g x)" by auto lemma poincare_mapsto_imp_flowsto: assumes "poincare_mapsto P X0 S CX X1" assumes "closed P" shows "flowsto X0 {0<..} (CX \ UNIV) (fst ` X1 \ UNIV)" unfolding flowsto_def proof safe fix x0 d0 assume "(x0, d0) \ X0" with assms obtain D where D: "returns_to P x0" "fst ` X0 \ S" "return_time P differentiable at x0 within S" "(poincare_map P has_derivative blinfun_apply D) (at x0 within S)" "(flow0 x0 (return_time P x0), D o\<^sub>L d0) \ X1" "\t. 0 < t \ t < return_time P x0 \ flow0 x0 t \ CX" by (auto simp: poincare_mapsto_def poincare_map_def) show "\h\{0<..}. h \ existence_ivl0 x0 \ (flow0 x0 h, Dflow x0 h o\<^sub>L d0) \ fst ` X1 \ UNIV \ (\h'\{0<--L d0) \ CX \ UNIV)" by (auto intro!: bexI[where x="return_time P x0"] return_time_exivl D assms return_time_pos simp: open_segment_eq_real_ivl) qed lemma flowsto_poincare_mapsto_trans_flowsto: assumes "flowsto X0 T CX X1'" assumes "poincare_mapsto P X1 S CY X2" assumes "closed P" assumes "fst ` X1' \ fst ` X1" assumes "X1' \ CX \ CY \ UNIV \ CZ" assumes "\t. t \ T \ t \ 0" shows "flowsto X0 {0<..} CZ (fst ` X2 \ UNIV)" proof - have X1D: "(a, b) \ X1' \ \c. (a, c) \ X1" for a b using assms(4) by force from poincare_mapsto_imp_flowsto[OF assms(2,3)] have "flowsto X1 {0<..} (CY \ UNIV) (fst ` X2 \ UNIV)" . then have "flowsto X1' {0<..} (CY \ UNIV) (fst ` X2 \ UNIV)" by (auto simp: flowsto_def dest!: X1D) from flowsto_trans[OF assms(1) this] show ?thesis apply (rule flowsto_subset) using assms by (auto intro!: add_nonneg_pos) qed lemma eq_blinfun_inner_left[intro]: "(\x. x \ n) = blinfun_apply (blinfun_inner_left n)" by force lemma flowsto_union_DiffE: assumes "flowsto X0 T CX (Y \ Z)" obtains X1 where "X1 \ X0" "flowsto X1 T CX Y" "flowsto (X0 - X1) T CX Z" proof - let ?X1 = "{x\X0. flowsto {x} T CX Y}" from assms have "?X1 \ X0" "flowsto ?X1 T CX Y" "flowsto (X0 - ?X1) T CX Z" by (auto simp: flowsto_def) thus ?thesis .. qed lemma eucl_less_le_trans: fixes a b::"'a::executable_euclidean_space" shows "eucl_less a b \ b \ c \ eucl_less a c" by (force simp: eucl_less_def[where 'a='a] eucl_le[where 'a='a]) lemma le_eucl_less_trans: fixes a b::"'a::executable_euclidean_space" shows "a \ b \ eucl_less b c \ eucl_less a c" by (force simp: eucl_less_def[where 'a='a] eucl_le[where 'a='a]) lemma flowsto_source_UnionI: assumes "\i. i \ I \ flowsto i T CXS (f i)" shows "flowsto (\I) T CXS (\(f ` I))" apply (auto simp: flowsto_def) subgoal premises prems for y a b using assms[unfolded flowsto_def, OF \y \ I\, rule_format, OF \_ \ y\] prems by auto done lemma poincare_mapsto_UnionI: assumes pm[unfolded poincare_mapsto_def, rule_format]: "\i. i \ I \ poincare_mapsto p (X0 i) S (CX i) (X1 i)" assumes R: "\i x. i \ I \ x \ X1 i \ x \ R" shows "poincare_mapsto p (\x\I. X0 x) S ((\x\I. CX x)) R" unfolding poincare_mapsto_def proof (safe del: conjI, goal_cases) case (1 x0 d0 i) moreover have "fst ` \(X0 ` I) \ S" proof (safe, goal_cases) case (1 _ x0 d0 i) from this pm[OF 1] show ?case by auto qed ultimately show ?case using pm[OF 1] by (auto intro!: R) qed lemma tendsto_at_top_translateI: assumes "(f \ l) (at_top::real filter)" shows "((\x. f (x + t)::'a::topological_space) \ l) at_top" proof (rule topological_tendstoI) fix S::"'a set" assume "open S" "l \ S" from topological_tendstoD[OF assms this] obtain N where "\n. n \ N \ f n \ S" by (auto simp: eventually_at_top_linorder) then have "\n. n \ N - t \ f (n + t) \ S" by auto then show "\\<^sub>F x in at_top. f (x + t) \ S" unfolding eventually_at_top_linorder by blast qed lemma tendsto_at_top_translate_iff: "((\x. f (x + t)::'a::topological_space) \ l) at_top \ (f \ l) (at_top::real filter)" using tendsto_at_top_translateI[of f l t] tendsto_at_top_translateI[of "\x. f (x + t)" l "- t"] by auto lemma stable_on_mono: "stable_on A B" if "stable_on C B" "A \ C" using that unfolding stable_on_def by fastforce lemma flowsto_mapsto_avoid_trap: assumes "flowsto (X0 - trap \ UNIV) {0<..} CX P" assumes trapprop: "stable_on (fst ` (CX \ P)) trap" shows "flowsto (X0 - trap \ UNIV) {0<..} CX (P - trap \ UNIV)" unfolding flowsto_def proof (rule, goal_cases) case (1 xd) from assms(1)[unfolded flowsto_def, rule_format, OF this] obtain h x0 d0 where "xd = (x0, d0)" "0 < h" "h \ existence_ivl0 (x0)" "(flow0 x0 h, Dflow x0 h o\<^sub>L d0) \ P" "(\h'\{0<--L d0) \ CX)" by auto then show ?case using 1 trapprop apply (auto intro!: bexI[where x=h] dest!: stable_onD simp: open_segment_eq_real_ivl image_Un) subgoal for s by (cases "s = h") auto done qed end lemma map_prod_def': "map_prod f g x = (f (fst x), g (snd x))" by (cases x) auto lemmas rel_prod_br = br_rel_prod lemmas lvivl_relI = lv_relivl_relI end diff --git a/thys/Ordinary_Differential_Equations/Numerics/Runge_Kutta.thy b/thys/Ordinary_Differential_Equations/Numerics/Runge_Kutta.thy --- a/thys/Ordinary_Differential_Equations/Numerics/Runge_Kutta.thy +++ b/thys/Ordinary_Differential_Equations/Numerics/Runge_Kutta.thy @@ -1,1372 +1,1372 @@ section\Runge-Kutta methods\ theory Runge_Kutta imports "HOL-Analysis.Analysis" One_Step_Method "HOL-Library.Float" Affine_Arithmetic.Executable_Euclidean_Space Ordinary_Differential_Equations.Multivariate_Taylor begin subsection \aux\ lemma scale_back: "(r, r *\<^sub>R x) = r *\<^sub>R (1, x)" "(0, r *\<^sub>R x) = r *\<^sub>R (0, x)" by simp_all lemma integral_normalize_bounds: fixes s t::real assumes "t \ s" assumes "f integrable_on {t .. s}" shows [symmetric]: "(s - t) *\<^sub>R integral {0 .. 1} (\x. f ((s - t) *\<^sub>R x + t)) = integral {t..s} f" proof cases assume "s > t" hence "s - t \ 0" "0 \ s - t" by simp_all from assms have "(f has_integral integral {t .. s} f) (cbox t s)" by (auto simp: integrable_integral) from has_integral_affinity[OF this \s - t \ 0\, of t] have "((\x. f ((s - t) * x + t)) has_integral (1 / \s - t\) *\<^sub>R integral {t..s} f) ((\x. (x - t) / (s - t)) ` {t..s})" using \s > t\ by (simp add: divide_simps) also have "t < s \ 0 \ x \ x \ 1 \ x * (s - t) + t \ s" for x by (auto simp add: algebra_simps dest: mult_left_le_one_le[OF \0 \ s - t\]) then have "((\x. (x - t) / (s - t)) ` {t..s}) = {0 .. 1}" using \s > t\ by (auto intro!: image_eqI[where x="x * (s - t) + t" for x] simp: divide_simps) finally have "integral {0..1} (\x. f ((s - t) * x + t)) = (1 / \s - t\) *\<^sub>R integral {t..s} f" by (rule integral_unique) then show ?thesis using \s > t\ by simp qed (insert assms, simp) lemma has_integral_integral_eqI: "f integrable_on s \ integral s f = k \ (f has_integral k) s" by (simp add: has_integral_integral) lemma convex_scaleR_sum2: assumes "x \ G" "y \ G" "convex G" assumes "a \ 0" "b \ 0" "a + b \ 0" shows "(a *\<^sub>R x + b *\<^sub>R y) /\<^sub>R (a + b) \ G" proof - have "(a / (a + b)) *\<^sub>R x + (b / (a + b)) *\<^sub>R y \ G" using assms by (intro convexD) (auto simp: divide_simps) then show ?thesis by (auto simp: algebra_simps divide_simps) qed lemma sum_by_parts_ivt: assumes "finite X" assumes "convex G" assumes "\i. i \ X \ g i \ G" assumes "\i. i \ X \ 0 \ c i" obtains y where "y \ G" "(\x\X. c x *\<^sub>R g x) = sum c X *\<^sub>R y" | "G = {}" proof (atomize_elim, cases "sum c X = 0", goal_cases) case pos: 2 let ?y = "(\x\X. (c x / sum c X) *\<^sub>R g x)" have "?y \ G" using pos by (intro convex_sum) (auto simp: sum_divide_distrib[symmetric] intro!: divide_nonneg_nonneg assms sum_nonneg) thus ?case by (auto intro!: exI[where x = ?y] simp: scaleR_right.sum pos) qed (insert assms, auto simp: sum_nonneg_eq_0_iff) lemma integral_by_parts_near_bounded_convex_set: assumes f: "(f has_integral I) (cbox a b)" assumes s: "((\x. f x *\<^sub>R g x) has_integral P) (cbox a b)" assumes G: "\x. x \ cbox a b \ g x \ G" assumes nonneg: "\x. x \ cbox a b \ f x \ 0" assumes convex: "convex G" assumes bounded: "bounded G" shows "infdist P ((*\<^sub>R) I ` G) = 0" proof (rule dense_eq0_I, cases) fix e'::real assume e0: "0 < e'" assume "G \ {}" from bounded obtain bnd where bnd: "\y. y \ G \ norm y < bnd" "bnd > 0" by (meson bounded_pos gt_ex le_less_trans norm_ge_zero) define e where "e = min (e' / 2) (e' / 2 / bnd)" have e: "e > 0" using e0 by (auto simp add: e_def intro!: divide_pos_pos \bnd > 0\) from has_integral[of f I a b, THEN iffD1, OF f, rule_format, OF e] has_integral[of "\x. f x *\<^sub>R g x" P a b, THEN iffD1, OF s, rule_format, OF e] obtain d1 d3 where d1: "gauge d1" "\p. p tagged_division_of cbox a b \ d1 fine p \ norm ((\(x, k)\p. content k *\<^sub>R f x) - I) < e" and d3: "gauge d3" "\p. p tagged_division_of cbox a b \ d3 fine p \ norm ((\(x, k)\p. content k *\<^sub>R f x *\<^sub>R g x) - P) < e" by auto define d where "d x = d1 x \ d3 x" for x from d1(1) d3(1) have "gauge d" by (auto simp add: d_def [abs_def]) from fine_division_exists[OF this, of a b] obtain p where p: "p tagged_division_of cbox a b" "d fine p" by metis from tagged_division_of_finite[OF p(1)] have "finite p" . from \d fine p\ have "d1 fine p" "d3 fine p" by (auto simp: d_def [abs_def] fine_Int) have f_less: "norm ((\(x, k)\p. content k *\<^sub>R f x) - I) < e" (is "norm (?f - I) < _") by (rule d1(2)[OF p(1)]) fact have "norm ((\(x, k)\p. content k *\<^sub>R f x *\<^sub>R g x) - P) < e" (is "norm (?s - P) < _") by (rule d3(2)[OF p(1)]) fact hence "dist (\(x, k)\p. content k *\<^sub>R f x *\<^sub>R g x) P < e" by (simp add: dist_norm) also let ?h = "(\x k y. (content k * f x) *\<^sub>R y)" let ?s' = "\y. sum (\(x, k). ?h x k y) p" let ?g = "\(x, k). g x" let ?c = "\(x, k). content k * f x" have Pi: "\x. x \ p \ ?g x \ G" "\x. x \ p \ ?c x \ 0" using nonneg G p using tag_in_interval[OF p(1)] by (auto simp: intro!: mult_nonneg_nonneg) obtain y where y: "y \ G" "?s = ?s' y" by (rule sum_by_parts_ivt[OF \finite p\ \convex G\ Pi]) (auto simp: split_beta' scaleR_sum_left \G \ {}\) note this(2) also have "(\(x, k)\p. (content k * f x) *\<^sub>R y) = ?f *\<^sub>R y" by (auto simp: scaleR_left.sum intro!: sum.cong) finally have "dist P ((\(x, k)\p. content k *\<^sub>R f x) *\<^sub>R y) \ e" by (simp add: dist_commute) moreover have "dist (I *\<^sub>R y) ((\(x, k)\p. content k *\<^sub>R f x) *\<^sub>R y) \ norm y * e" using f_less by (auto simp add: dist_real_def mult.commute [of _ "norm y"] intro!: mult_left_mono) ultimately have "dist P (I *\<^sub>R y) \ e + norm y * e" by (rule dist_triangle_le[OF add_mono]) with _ have "infdist P ((*\<^sub>R) I ` G) \ e + norm y * e" using y(1) by (intro infdist_le2) auto also have "norm y * e < bnd * e" by (rule mult_strict_right_mono) (auto simp: \e > 0\ less_imp_le intro!: bnd \y \ G\) also have "bnd * e \ e' / 2" using \e' > 0\ \bnd > 0\ by (auto simp: e_def min_def divide_simps) also have "e \ e' / 2" by (simp add: e_def) also have "e' / 2 + e' / 2 = e'" by simp finally show "\infdist P ((*\<^sub>R) I ` G)\ \ e'" by (auto simp: infdist_nonneg) qed (simp add: infdist_def) lemma integral_by_parts_in_bounded_closed_convex_set: assumes f: "(f has_integral I) (cbox a b)" assumes s: "((\x. f x *\<^sub>R g x) has_integral P) (cbox a b)" assumes G: "\x. x \ cbox a b \ g x \ G" assumes nonneg: "\x. x \ cbox a b \ f x \ 0" assumes bounded: "bounded G" assumes closed: "closed G" assumes convex: "convex G" assumes nonempty: "cbox a b \ {}" shows "P \ (*\<^sub>R) I ` G" proof - let ?IG = "(*\<^sub>R) I ` G" from bounded closed have "bounded ?IG" "closed ?IG" by (simp_all add: bounded_scaling closed_scaling) have "G \ {}" using nonempty G by auto then show ?thesis using \closed ?IG\ by (subst in_closed_iff_infdist_zero) (auto intro!: assms compact_imp_bounded integral_by_parts_near_bounded_convex_set) qed lemma integral_by_parts_in_bounded_set: assumes f: "(f has_integral I) (cbox a b)" assumes s: "((\x. f x *\<^sub>R g x) has_integral P) (cbox a b)" assumes nonneg: "\x. x \ cbox a b \ f x \ 0" assumes bounded: "bounded (g ` cbox a b)" assumes nonempty: "cbox a b \ {}" shows "P \ (*\<^sub>R) I ` closure (convex hull (g ` cbox a b))" proof - have "x \ cbox a b \ g x \ closure (convex hull g ` cbox a b)" for x by (meson closure_subset hull_subset imageI subsetCE) then show ?thesis by (intro integral_by_parts_in_bounded_closed_convex_set[OF f s _ nonneg _ _ _ nonempty]) (auto intro!: bounded_closure bounded_convex_hull bounded convex_closure simp: convex_convex_hull) qed lemma snd_imageI: "(a, b) \ R \ b \ snd ` R" by force lemma in_minus_Collect: "a \ A \ b \ B \ a - b \ {x - y|x y. x \ A \ y \ B}" by blast lemma closure_minus_Collect: fixes A B::"'a::real_normed_vector set" shows "{x - y|x y. x \ closure A \ y \ closure B} \ closure {x - y|x y. x \ A \ y \ B}" proof - have image: "(\(x, y). x - y) ` (A \ B) = {x - y|x y. x \ A \ y \ B}" for A B::"'a set" by auto have "{x - y|x y. x \ closure A \ y \ closure B} = (\(x, y). x - y) ` closure (A \ B)" unfolding closure_Times by (rule image[symmetric]) also have "\ \ closure ((\(x, y). x - y) ` (A \ B))" by (rule image_closure_subset) (auto simp: split_beta' intro!: subsetD[OF closure_subset] continuous_at_imp_continuous_on) also note image finally show ?thesis . qed lemma convex_hull_minus_Collect: fixes A B::"'a::real_normed_vector set" shows "{x - y|x y. x \ convex hull A \ y \ convex hull B} = convex hull {x - y|x y. x \ A \ y \ B}" proof - have image: "(\(x, y). x - y) ` (A \ B) = {x - y|x y. x \ A \ y \ B}" for A B::"'a set" by auto have "{x - y|x y. x \ convex hull A \ y \ convex hull B} = (\(x, y). x - y) ` (convex hull (A \ B))" unfolding convex_hull_Times by (rule image[symmetric]) also have "\ = convex hull ((\(x, y). x - y) ` (A \ B))" apply (rule convex_hull_linear_image) by unfold_locales (auto simp: algebra_simps) also note image finally show ?thesis . qed lemma set_minus_subset: "A \ C \ B \ D \ {a - b |a b. a \ A \ b \ B} \ {a - b |a b. a \ C \ b \ D}" by auto lemma (in bounded_bilinear) bounded_image: assumes "bounded (f ` s)" assumes "bounded (g ` s)" shows "bounded ((\x. prod (f x) (g x)) ` s)" proof - from nonneg_bounded obtain K where K: "\a b. norm (prod a b) \ norm a * norm b * K" and "0 \ K" by auto from assms obtain F G where F: "\x. x \ s \ norm (f x) \ F" and G: "\x. x \ s \ norm (g x) \ G" and nonneg: "0 \ F" "0 \ G" by (auto simp: bounded_pos intro: less_imp_le) have "norm (prod (f x) (g x)) \ F * G * K" if x: "x \ s" for x using F[OF x] G[OF x] nonneg \0 \ K\ by (auto intro!: mult_mono mult_nonneg_nonneg order_trans[OF K]) thus ?thesis by (auto simp: bounded_iff) qed lemmas bounded_scaleR_image = bounded_bilinear.bounded_image[OF bounded_bilinear_scaleR] and bounded_blinfun_apply_image = bounded_bilinear.bounded_image[OF bounded_bilinear_blinfun_apply] lemma bounded_plus_image: fixes f::"'a \ 'b::real_normed_vector" assumes "bounded (f ` s)" assumes "bounded (g ` s)" shows "bounded ((\x. f x + g x) ` s)" proof - from assms obtain F G where F: "\x. x \ s \ norm (f x) \ F" and G: "\x. x \ s \ norm (g x) \ G" by (auto simp: bounded_iff) have "norm (f x + g x) \ F + G" if x: "x \ s" for x using F[OF x] G[OF x] by norm thus ?thesis by (auto simp: bounded_iff) qed lemma bounded_uminus_image[simp]: fixes f::"'a \ 'b::real_normed_vector" shows "bounded ((\x. - f x) ` s) \ bounded (f ` s)" apply (subst (2) bounded_uminus[symmetric]) unfolding image_image by simp lemma bounded_minus_image: fixes f::"'a \ 'b::real_normed_vector" assumes "bounded (f ` s)" assumes "bounded (g ` s)" shows "bounded ((\x. f x - g x) ` s)" using bounded_plus_image[of f s "\x. - g x"] assms by simp lemma bounded_Pair_image: fixes f::"'a \ 'b::real_normed_vector" fixes g::"'a \ 'c::real_normed_vector" assumes "bounded (f ` s)" assumes "bounded (g ` s)" shows "bounded ((\x. (f x, g x)) ` s)" proof - from assms obtain F G where F: "\x. x \ s \ norm (f x) \ F" and G: "\x. x \ s \ norm (g x) \ G" by (auto simp: bounded_iff) have "norm (f x, g x) \ F + G" if x: "x \ s" for x using F[OF x] G[OF x] by (intro order_trans[OF norm_Pair_le]) norm thus ?thesis by (auto simp: bounded_iff) qed lemma closed_scaleR_image_iff: fixes f::"'a \ 'b::real_normed_vector" shows "closed ((\x. r *\<^sub>R (f x)) ` R) \ (r = 0 \ closed (f ` R))" (is "?l \ _ \ ?r") proof safe assume ?r from closed_scaling[OF this] show ?l by (auto simp: image_image) next assume l: ?l { assume "r \ 0" with closed_scaling[OF l, of "inverse r"] have "?r" by (auto simp: image_image) } then show "\ ?r \ r = 0" by blast qed (simp add: image_constant_conv) lemma closed_translation_iff[simp]: fixes y::"'a::real_normed_vector" shows "closed ((\x. f x + y) ` S) \ closed (f ` S)" "closed ((\x. y + f x) ` S) \ closed (f ` S)" using closed_translation[of "f ` S" y] closed_translation[of "(+) y ` f ` S" "- y"] by (auto simp: add.commute image_image cong: image_cong) lemma closed_minus_translation_iff[simp]: fixes y::"'a::real_normed_vector" shows "closed ((\x. f x - y) ` S) \ closed (f ` S)" using closed_translation_iff(1)[of f "- y" S] by simp lemma convex_scaleR_image_iff: fixes f::"'a \ 'b::real_normed_vector" shows "convex ((\x. r *\<^sub>R (f x)) ` R) \ (r = 0 \ convex (f ` R))" (is "?l \ _ \ ?r") proof safe assume ?r from convex_scaling[OF this] show ?l by (auto simp: image_image) next assume l: ?l { assume "r \ 0" with convex_scaling[OF l, of "inverse r"] have "?r" by (auto simp: image_image) } then show "\ ?r \ r = 0" by blast qed (simp add: image_constant_conv) lemma convex_translation_iff[simp]: fixes y::"'a::real_normed_vector" shows "convex ((\x. f x + y) ` S) \ convex (f ` S)" "convex ((\x. y + f x) ` S) \ convex (f ` S)" using convex_translation[of "f ` S" y] convex_translation[of "(+) y ` f ` S" "- y"] by (auto simp: add.commute image_image cong: image_cong) lemma convex_minus_translation_iff[simp]: fixes y::"'a::real_normed_vector" shows "convex ((\x. f x - y) ` S) \ convex (f ` S)" using convex_translation_iff(1)[of f "- y" S] by simp text\\label{sec:rk}\ subsection \Definitions\ text\\label{sec:rk-definition}\ declare sum.cong[fundef_cong] fun rk_eval :: "(nat\nat\real) \ (nat\real) \ (real \ 'a::real_vector \ 'a) \ real \ real \ 'a \ nat \ 'a" where "rk_eval A c f t h x j = f (t + h * c j) (x + h *\<^sub>R (\l=1 ..< j. A j l *\<^sub>R rk_eval A c f t h x l))" primrec rk_eval_dynamic :: "(nat\nat\real) \ (nat\real) \ (real\'a::{comm_monoid_add, scaleR} \ 'a) \ real \ real \ 'a \ nat \ (nat \ 'a)" where "rk_eval_dynamic A c f t h x 0 = (\i. 0)" | "rk_eval_dynamic A c f t h x (Suc j) = (let K = rk_eval_dynamic A c f t h x j in K(Suc j:=f (t + h * c (Suc j), x + h *\<^sub>R (\l=1..j. A (Suc j) l *\<^sub>R K l))))" definition rk_increment where "rk_increment f s A b c h t x = (\j=1..s. b j *\<^sub>R rk_eval A c f t h x j)" definition rk_increment' where "rk_increment' error f s A b c h t x = eucl_down error (\j=1..s. b j *\<^sub>R rk_eval A c f t h x j)" definition euler_increment where "euler_increment f = rk_increment f 1 (\i j. 0) (\i. 1) (\i. 0)" definition euler where "euler f = grid_function (discrete_evolution (euler_increment f))" definition euler_increment' where "euler_increment' e f = rk_increment' e f 1 (\i j. 0) (\i. 1) (\i. 0)" definition euler' where "euler' e f = grid_function (discrete_evolution (euler_increment' e f))" definition rk2_increment where "rk2_increment x f = rk_increment f 2 (\i j. if i = 2 \ j = 1 then x else 0) (\i. if i = 1 then 1 - 1 / (2 * x) else 1 / (2 * x)) (\i. if i = 2 then x else 0)" definition rk2 where "rk2 x f = grid_function (discrete_evolution (rk2_increment x f))" subsection \Euler method is consistent \label{sec:rk-euler-cons}\ lemma euler_increment: shows "euler_increment f h t x = f t x" unfolding euler_increment_def rk_increment_def by (subst rk_eval.simps) (simp del: rk_eval.simps) lemma euler_float_increment: shows "euler_increment' e f h t x = eucl_down e (f t x)" unfolding euler_increment'_def rk_increment'_def by (subst rk_eval.simps) (simp del: rk_eval.simps) lemma euler_lipschitz: assumes t: "t \ {t0..T}" assumes lipschitz: "\t\{t0..T}. L-lipschitz_on D' (\x. f t x)" shows "L-lipschitz_on D' (euler_increment f h t)" using t lipschitz by (simp add: lipschitz_on_def euler_increment del: One_nat_def) lemma rk2_increment: shows "rk2_increment p f h t x = (1 - 1 / (p * 2)) *\<^sub>R f t x + (1 / (p * 2)) *\<^sub>R f (t + h * p) (x + (h * p) *\<^sub>R f t x)" unfolding rk2_increment_def rk_increment_def apply (subst rk_eval.simps) apply (simp del: rk_eval.simps add: numeral_2_eq_2) apply (subst rk_eval.simps) apply (simp del: rk_eval.simps add: field_simps) done subsection \Set-Based Consistency of Euler Method \label{sec:setconsistent}\ context derivative_on_prod begin lemma euler_consistent_traj_set: fixes t assumes ht: "0 \ h" "t + h \ u" assumes T: "{t..u} \ T" assumes x': "\s. s \ {t..u} \ (x has_vector_derivative f s (x s)) (at s within {t..u})" assumes x: "\s. s \ {t..u} \ x s \ X" assumes R: "\s. s \ {t..u} \ discrete_evolution (euler_increment f) (t + h) t (x t) + (h\<^sup>2 / 2) *\<^sub>R (f' (s, x s)) (1, f s (x s)) \ R" assumes bcc: "bounded R" "closed R" "convex R" shows "x (t + h) \ R" proof cases assume "h = 0" with assms show ?thesis by (auto simp: discrete_evolution_def) next assume "h \ 0" from this ht have "t < u" by simp from ht have line_subset: "(\ta. t + ta * h) ` {0..1} \ {t..u}" by (auto intro!: order_trans[OF add_left_mono[OF mult_left_le_one_le]]) hence line_in: "\s. 0 \ s \ s \ 1 \ t + s * h \ {t..u}" by (rule subsetD) auto from ht have subset: "{t .. t + h} \ {t .. u}" by simp let ?T = "{t..u}" from ht have subset: "{t .. t + h} \ {t .. u}" by simp from \t < u\ have "t \ ?T" by auto from \t < u\ have tx: "t \ T" "x t \ X" using assms by auto from tx assms have "0 \ norm (f t (x t))" by simp have x_diff: "\s. s \ ?T \ x differentiable at s within ?T" by (rule differentiableI, rule x'[simplified has_vector_derivative_def]) note [continuous_intros] = continuous_on_compose2[OF has_derivative_continuous_on[OF f'] continuous_on_Pair, simplified] continuous_on_compose2[OF has_derivative_continuous_on[OF x'[unfolded has_vector_derivative_def]], simplified] let ?p = "(\t. f' (t, x t) (1, f t (x t)))" define diff where "diff \ \n::nat. if n = 0 then x else if n = 1 then \t. f t (x t) else ?p" have diff_0[simp]: "diff 0 = x" by (simp add: diff_def) have diff: "(diff m has_vector_derivative diff (Suc m) ta) (at ta within {t..t + h})" if mta: "m < 2" "t \ ta" "ta \ t + h" for m::nat and ta::real proof - have image_subset: "(\xa. (xa, x xa)) ` {t..u} \ {t..u} \ X" using assms by auto note has_derivative_subset[OF _ image_subset, derivative_intros] note f'[derivative_intros] note x'[simplified has_vector_derivative_def, derivative_intros] have [simp]: "\c x'. c *\<^sub>R f' (ta, x ta) x' = f' (ta, x ta) (c *\<^sub>R x')" using mta ht assms by (force intro!: f' linear_cmul[symmetric] has_derivative_linear) have f_comp': "((\t. f t (x t)) has_vector_derivative f' (ta, x ta) (1, f ta (x ta))) (at ta within {t..u})" unfolding has_vector_derivative_def using assms ht mta by (auto intro!: derivative_eq_intros) then show ?thesis using mta ht by (auto simp: diff_def intro!: has_vector_derivative_within_subset[OF _ subset] x') qed from Taylor_has_integral[of 2 diff x t "t + h", OF _ _ diff] \0 \ h\ have Taylor: "((\xa. (t + h - xa) *\<^sub>R f' (xa, x xa) (1, f xa (x xa))) has_integral x (t + h) - (x t + h *\<^sub>R f t (x t))) {t..t + h}" by (simp add: eval_nat_numeral diff_def) have *: "h\<^sup>2 / 2 = content {t..t + h} *\<^sub>R (t + h) - (if t \ t + h then (t + h)\<^sup>2 / 2 - t\<^sup>2 / 2 else 0)" using \0 \ h\ by (simp add: algebra_simps power2_eq_square divide_simps) have integral: "((-) (t + h) has_integral h\<^sup>2 / 2) (cbox t (t + h))" unfolding * cbox_interval using \0 \ h\ by (auto intro!: has_integral_diff ident_has_integral[THEN has_integral_eq_rhs] has_integral_const_real[THEN has_integral_eq_rhs]) from Taylor_has_integral[of 2 diff x t "t + h", OF _ _ diff] \0 \ h\ have Taylor: "((\xa. (t + h - xa) *\<^sub>R f' (xa, x xa) (1, f xa (x xa))) has_integral x (t + h) - (x t + h *\<^sub>R f t (x t))) {t..t + h}" by (simp add: eval_nat_numeral diff_def) define F' where "F' \ (\y. (2 / h\<^sup>2) *\<^sub>R (y - discrete_evolution (euler_increment f) (t + h) t (x t))) ` R" have "x (t + h) - (x t + h *\<^sub>R f t (x t)) \ (*\<^sub>R) (h\<^sup>2 / 2) ` F'" apply (rule integral_by_parts_in_bounded_closed_convex_set[OF integral Taylor[unfolded interval_cbox]]) subgoal using R \h \ 0\ \0 \ h\ subset by (force simp: F'_def) by (auto intro!: bounded_scaleR_image bounded_minus_image closed_injective_image_subspace bcc \0 \ h\ simp: F'_def image_constant_conv closed_scaleR_image_iff convex_scaleR_image_iff \h \ 0\) then show ?thesis using \h \ 0\ by (auto simp add: discrete_evolution_def euler_increment F'_def) qed end lemma numeral_6_eq_6: "6 = Suc (Suc (Suc (Suc (Suc (Suc 0)))))" by linarith context begin interpretation blinfun_syntax . definition "heun_remainder1 x f f' f'' t h s = (h ^ 3 / 6) *\<^sub>R (f'' (h * s + t, x (h * s + t)) $ (1::real, f (h * s + t, x (h * s + t))) $ (1::real, f (h * s + t, x (h * s + t))) + f' (h * s + t, x (h * s + t)) $ (0::real, f' (h * s + t, x (h * s + t)) $ (1, f (h * s + t, x (h * s + t)))))" definition "heun_remainder2 p x f f'' t h s = (h ^ 3 * p / 4) *\<^sub>R f'' (t + s * h * p, x t + (s * h * p) *\<^sub>R f (t, (x t))) $ (1::real, f (t, (x t))) $ (1::real, f (t, (x t)))" lemma rk2_consistent_traj_set: fixes x ::"real \ 'a::banach" and t and f' :: "real \ 'a \ (real \ 'a) \\<^sub>L 'a" and g' :: "real \ 'a \ (real \ 'a) \ 'a" and f'' :: "real \ 'a \ (real \ 'a) \\<^sub>L (real \ 'a) \\<^sub>L 'a" and g'' :: "real \ 'a \ (real \ 'a) \ (real \ 'a) \\<^sub>L 'a" assumes ht: "0 \ h" "t + h \ u" assumes T: "{t..u} \ T" and X_nonempty: "X \ {}" and convex_X: "convex X" assumes x': "\s. s \ {t..u} \ (x has_vector_derivative f (s, x s)) (at s within {t..u})" assumes f': "\tx. tx \ T \ X \ (f has_derivative g' tx) (at tx)" assumes f'': "\tx. tx \ T \ X \ (f' has_derivative g'' tx) (at tx)" assumes g': "\tx. tx \ T \ X \ g' tx = f' tx" assumes g'': "\tx. tx \ T \ X \ g'' tx = f'' tx" assumes f''_bounded: "bounded (f'' ` (T \ X))" assumes x: "\s. s \ {t..u} \ x s \ X" assumes p: "0 < p" "p \ 1" assumes step_in: "x t + (h * p) *\<^sub>R f (t, (x t)) \ X" assumes ccR: "convex R" "closed R" assumes R: "\s1 s2. 0 \ s1 \ s1 \ 1 \ 0 \ s2 \ s2 \ 1 \ discrete_evolution (rk2_increment p (\t x. f (t, x))) (t + h) t (x t) + heun_remainder1 x f f' f'' t h s1 - heun_remainder2 p x f f'' t h s2 \ R" shows "x (t + h) \ R" proof cases assume "h = 0" with R[of 0 0] show ?thesis by (auto simp: discrete_evolution_def heun_remainder1_def heun_remainder2_def) next have f': "\tx. tx \ T \ X \ (f has_derivative blinfun_apply (f' tx)) (at tx)" using f' g' by simp have f'': "\tx. tx \ T \ X \ (f' has_derivative blinfun_apply (f'' tx)) (at tx)" using f'' g'' by simp assume "h \ 0" from this ht have "t < u" by simp have [simp]: "p \ 0" using p by simp from \h \ 0\ \h \ 0\ have "h > 0" by simp let ?r = "\a. f'' (t + a, x t + a *\<^sub>R f (t, x t)) (1, f (t, x t)) (1, f (t, x t))" let ?q = "\s. f'' (s, x s) (1, f (s, x s)) (1, f (s, x s)) + f' (s, x s) (0, f' (s, x s) (1, f (s, x s)))" let ?d = "\tq tr. (h^3) *\<^sub>R ((1/6)*\<^sub>R ?q tq - (p / 4) *\<^sub>R ?r tr)" from ht have line_subset: "(\ta. t + ta * h) ` {0..1} \ {t..u}" by (auto intro!: order_trans[OF add_left_mono[OF mult_left_le_one_le]]) hence line_in: "\s. 0 \ s \ s \ 1 \ t + s * h \ {t..u}" by (rule subsetD) auto from ht have subset: "{t .. t + h} \ {t .. u}" by simp let ?T = "{t..u}" from ht have subset: "{t .. t + h} \ {t .. u}" by simp from \t < u\ have "t \ ?T" by auto from \t < u\ have tx: "t \ T" "x t \ X" using T ht x by auto from tx assms have "0 \ norm (f (t, x t))" by simp have x_diff: "\s. s \ ?T \ x differentiable at s within ?T" by (rule differentiableI, rule x'[simplified has_vector_derivative_def]) let ?p = "(\t. f' (t, x t) (1, f (t, x t)))" note f'[derivative_intros] note f''[derivative_intros] note x'[simplified has_vector_derivative_def, derivative_intros] have x_cont: "continuous_on {t..u} x" by (rule continuous_on_vector_derivative) (rule x') have f_cont: "continuous_on (T \ X) f" apply (rule has_derivative_continuous_on) apply (rule has_derivative_at_withinI) by (rule assms) have f'_cont: "continuous_on (T \ X) f'" apply (rule has_derivative_continuous_on) apply (rule has_derivative_at_withinI) by (rule assms) note [continuous_intros] = continuous_on_compose2[OF x_cont] continuous_on_compose2[OF f_cont] continuous_on_compose2[OF f'_cont] from f' f'' have f'_within: "tx \ T \ X \ (f has_derivative f' tx) (at tx within T \ X)" and f''_within: "tx \ T \ X \ (f' has_derivative f'' tx) (at tx within T \ X)" for tx by (auto intro: has_derivative_at_withinI) from f'' have f''_within: "tx \ T \ X \ (f' has_derivative ($) (f'' tx)) (at tx within T \ X)" for tx by (auto intro: has_derivative_at_withinI) note [derivative_intros] = has_derivative_in_compose2[OF f'_within] has_derivative_in_compose2[OF f''_within] have p': "\s. s \ {t .. u} \ (?p has_vector_derivative ?q s) (at s within ?T)" unfolding has_vector_derivative_def using T x by (auto intro!: derivative_eq_intros simp: scale_back blinfun.bilinear_simps algebra_simps simp del: scaleR_Pair) define diff where "diff n = (if n = 0 then x else if n = 1 then \t. f (t, x t) else if n = 2 then ?p else ?q)" for n :: nat have diff_0[simp]: "diff 0 = x" by (simp add: diff_def) have diff: "(diff m has_vector_derivative diff (Suc m) ta) (at ta within {t..t + h})" if mta: "m < 3" "t \ ta" "ta \ t + h" for m::nat and ta::real proof - have image_subset: "(\xa. (xa, x xa)) ` {t..u} \ {t..u} \ X" using assms by auto note has_derivative_in_compose[where f="(\xa. (xa, x xa))" and g = f, derivative_intros] note has_derivative_subset[OF _ image_subset, derivative_intros] note f'[derivative_intros] note x'[simplified has_vector_derivative_def, derivative_intros] have [simp]: "\c x'. c *\<^sub>R f' (ta, x ta) x' = f' (ta, x ta) (c *\<^sub>R x')" using mta ht assms T x by (force intro!: f' linear_cmul[symmetric] has_derivative_linear) have "((\t. f (t, x t)) has_vector_derivative f' (ta, x ta) (1, f (ta, x ta))) (at ta within {t..u})" unfolding has_vector_derivative_def using assms ht mta T x - by (force intro!: derivative_eq_intros has_derivative_within_subset[OF f']) + by (force intro!: derivative_eq_intros has_derivative_subset[OF f']) then show ?thesis using mta ht by (auto simp: diff_def intro!: has_vector_derivative_within_subset[OF _ subset] x' p') qed from Taylor_has_integral[of 3 diff x t "t + h", OF _ _ diff] have "((\x. ((t + h - x) ^ 2 / 2) *\<^sub>R diff 3 x) has_integral x (t + h) - x t - h *\<^sub>R (f (t, x t)) - (h ^ 2 / (2::nat)) *\<^sub>R (?p t)) (cbox t (t + h))" using ht \h\0\ by (auto simp: field_simps of_nat_Suc Pi_iff numeral_2_eq_2 numeral_3_eq_3 numeral_6_eq_6 power2_eq_square diff_def scaleR_sum_right) from has_integral_affinity[OF this \h \ 0\, of t, simplified] have "((\x. ((h - h * x)\<^sup>2 / 2) *\<^sub>R diff 3 (h * x + t)) has_integral (1 / \h\) *\<^sub>R (x (t + h) - x t - h *\<^sub>R f (t, x t) - (h\<^sup>2 / 2) *\<^sub>R f' (t, x t) $ (1, f (t, x t)))) ((\x. x / h - t / h) ` {t..t + h})" by simp also have "((\x. x / h - t / h) ` {t..t + h}) = {0 .. 1}" using \h \ 0\ \h \ 0\ by (auto simp: divide_simps intro!: image_eqI[where x="x * h + t" for x]) finally have "((\x. ((h - h * x)\<^sup>2 / 2) *\<^sub>R diff 3 (h * x + t)) has_integral (1 / \h\) *\<^sub>R (x (t + h) - x t - h *\<^sub>R f (t, x t) - (h\<^sup>2 / 2) *\<^sub>R f' (t, x t) $ (1, f (t, x t)))) {0..1}" . from has_integral_cmul[OF this, of h] have Taylor: "((\x. (1 - x)\<^sup>2 *\<^sub>R ((h^3 / 2) *\<^sub>R ?q (h * x + t))) has_integral (x (t + h) - x t - h *\<^sub>R f (t, x t) - (h\<^sup>2 / 2) *\<^sub>R f' (t, x t) $ (1, f (t, x t)))) {0..1}" (is "(?i_Taylor has_integral _) _") using \h \ 0\ \h \ 0\ by (simp add: diff_def divide_simps algebra_simps power2_eq_square power3_eq_cube) have line_in': "h * y + t \ T" "x (h * y + t) \ X" "t \ h * y + t" "h * y + t \ u" if "y \ cbox 0 1" for y using line_in[of y] that T by (auto simp: algebra_simps x) let ?integral = "\x. x^3/3 - x^2 + x" have intsquare: "((\x. (1 - x)\<^sup>2) has_integral ?integral 1 - ?integral 0) (cbox 0 (1::real))" unfolding cbox_interval by (rule fundamental_theorem_of_calculus) (auto intro!: derivative_eq_intros simp: has_vector_derivative_def power2_eq_square algebra_simps) have f_Taylor: "((\s. (1 - s) *\<^sub>R f'' (x + s *\<^sub>R h) h h) has_integral f (x + h) - f x - f' x $ h) {0..1}" if line_in: "(\s. x + s *\<^sub>R h) ` {0..1} \ T \ X" for x h::"real*'a" proof - from that have *: "y \ T \ X" if "y \ closed_segment x (x + h)" for y using that by (force simp: closed_segment_def algebra_simps intro: image_eqI[where x = "1 - x" for x]) define Df where "Df x i h1 h2 = (if i = 0 then f x else if i = 1 then f' x h2 else f'' x h2 h1)" for x h1 h2 and i::nat have "((\y. ((1 - y) ^ (2 - 1) / fact (2 - 1)) *\<^sub>R Df (x + y *\<^sub>R h) 2 h h) has_integral f (x + h) - (\i<2. (1 / fact i) *\<^sub>R Df x i h h)) {0..1}" apply (rule multivariate_Taylor_has_integral[of 2 Df h f x "T \ X"]) subgoal by simp subgoal by (simp add: Df_def) subgoal premises prems for a i d proof - consider "i = 0" | "i = 1" | "i = 2" using prems by arith then show ?thesis by cases (use prems in \auto simp: Df_def[abs_def] blinfun.bilinear_simps intro!: derivative_eq_intros intro: *\) qed subgoal using "*" by blast done then show ?thesis by (simp add: Df_def eval_nat_numeral algebra_simps) qed let ?k = "\t. f ((t, x t) + (h * p) *\<^sub>R (1, f (t, x t)))" have line_in: "(\s. (t, x t) + s *\<^sub>R ((h * p) *\<^sub>R (1, f (t, x t)))) ` {0..1} \ T \ X" proof (clarsimp, safe) fix s::real assume s: "0 \ s" "s \ 1" have "t + s * (h * p) = t + s * p * h" by (simp add: ac_simps) also have "\ \ {t .. u}" using \0 < p\ \p \ 1\ s by (intro line_in) (auto intro!: mult_nonneg_nonneg mult_left_le_one_le mult_le_one) also note \\ \ T\ finally show "t + s * (h * p) \ T" . show "x t + (s * (h * p)) *\<^sub>R f (t, x t) \ X" using convexD_alt[OF \convex X\ tx(2) step_in s] by (simp add: algebra_simps) qed from f_Taylor[OF line_in, simplified] have k: "((\s. (1 - s) *\<^sub>R ((h\<^sup>2 * p\<^sup>2) *\<^sub>R f'' (t + s * (h * p), x t + (s * (h * p)) *\<^sub>R f (t, x t)) $ (1, f (t, x t)) $ (1, f (t, x t)))) has_integral ?k t - f (t, x t) - f' (t, x t) $ (h * p, (h * p) *\<^sub>R f (t, x t))) {0..1}" (is "(?i has_integral _) _") unfolding scale_back blinfun.bilinear_simps by (simp add: power2_eq_square algebra_simps) have rk2: "discrete_evolution (rk2_increment p (\t x. f (t, x))) (t + h) t (x t) = x t + h *\<^sub>R f (t, x t) - (h / (2 * p)) *\<^sub>R f (t, x t) + (h / (p * 2)) *\<^sub>R ?k t" (is "_ = ?rk2 t") unfolding rk2_increment_def discrete_evolution_def rk_increment_def apply (subst rk_eval.simps) supply rk_eval.simps[simp del] apply (simp add: eval_nat_numeral) apply (subst rk_eval.simps) apply (simp add: algebra_simps) done also have "\ = x t + h *\<^sub>R f (t, x t) + (h / (2 * p)) *\<^sub>R (f' (t, x t) ((h * p), (h * p) *\<^sub>R f (t, x t))) + (h / (p * 2)) *\<^sub>R integral {0 .. 1} ?i" unfolding integral_unique[OF k] by (simp add: algebra_simps) also have "(h / (2 * p)) *\<^sub>R f' (t, x t) (h * p, (h * p) *\<^sub>R f (t, x t)) = (h\<^sup>2 / 2) *\<^sub>R ?p t" by (simp add: scale_back blinfun.bilinear_simps power2_eq_square del: scaleR_Pair) finally have "integral {0 .. 1} ?i = (discrete_evolution (rk2_increment p (\t x. f (t, x))) (t + h) t (x t) - x t - h *\<^sub>R f (t, x t) - (h\<^sup>2 / 2) *\<^sub>R ?p t) /\<^sub>R (h / (p * 2))" by (simp add: blinfun.bilinear_simps zero_prod_def[symmetric]) with _ have "(?i has_integral (discrete_evolution (rk2_increment p (\t x. f (t, x))) (t + h) t (x t) - x t - h *\<^sub>R f (t, x t) - (h\<^sup>2 / 2) *\<^sub>R ?p t) /\<^sub>R (h / (p * 2))) {0 .. 1}" using k by (intro has_integral_integral_eqI) (rule has_integral_integrable) from has_integral_cmul[OF this, of "h / (p * 2)"] have discrete_Taylor: "((\s. (1 - s) *\<^sub>R ((h^3 * p / 2) *\<^sub>R f'' (t + s * (h * p), x t + (s * (h * p)) *\<^sub>R f (t, x t)) $ (1, f (t, x t)) $ (1, f (t, x t)))) has_integral (discrete_evolution (rk2_increment p (\t x. f (t, x))) (t + h) t (x t) - x t - h *\<^sub>R f (t, x t) - (h\<^sup>2 / 2) *\<^sub>R f' (t, x t) (1, f (t, x t)))) {0 .. 1}" (is "(?i_dTaylor has_integral _) _") using \h > 0\ by (simp add: algebra_simps diff_divide_distrib power2_eq_square power3_eq_cube) have integral_minus: "((-) 1 has_integral 1/2) (cbox 0 (1::real))" by (auto intro!: has_integral_eq_rhs[OF has_integral_diff] ident_has_integral) have bounded_f: "bounded ((\xa. f (h * xa + t, x (h * xa + t))) ` {0..1})" using \0 \ h\ by (auto intro!: compact_imp_bounded compact_continuous_image continuous_intros mult_nonneg_nonneg simp: line_in') have bounded_f': "bounded ((\xa. f' (h * xa + t, x (h * xa + t))) ` {0..1})" using \0 \ h\ by (auto intro!: compact_imp_bounded compact_continuous_image continuous_intros simp: line_in') have bounded_f'': "bounded ((\xa. f'' (h * xa + t, x (h * xa + t))) ` {0..1})" apply (subst o_def[of f'', symmetric]) apply (subst image_comp[symmetric]) apply (rule bounded_subset[OF f''_bounded]) by (auto intro!: image_eqI line_in') have bounded_f''_2: "bounded ((\xa. f'' (t + xa * (h * p), x t + (xa * (h * p)) *\<^sub>R f (t, x t))) ` {0..1})" apply (subst o_def[of f'', symmetric]) apply (subst image_comp[symmetric]) apply (rule bounded_subset[OF f''_bounded]) using line_in by auto have 1: "x (t + h) - x t - h *\<^sub>R f (t, x t) - (h\<^sup>2 / 2) *\<^sub>R f' (t, x t) $ (1, f (t, x t)) \ (*\<^sub>R) (1 / 3) ` closure (convex hull (\xa. (h ^ 3 / 2) *\<^sub>R (f'' (h * xa + t, x (h * xa + t)) $ (1, f (h * xa + t, x (h * xa + t))) $ (1, f (h * xa + t, x (h * xa + t))) + f' (h * xa + t, x (h * xa + t)) $ (0, f' (h * xa + t, x (h * xa + t)) $ (1, f (h * xa + t, x (h * xa + t)))))) ` cbox 0 1)" by (rule rev_subsetD[OF integral_by_parts_in_bounded_set[OF intsquare Taylor[unfolded interval_cbox]]]) (auto intro!: bounded_scaleR_image bounded_plus_image bounded_blinfun_apply_image bounded_Pair_image bounded_f'' bounded_f' bounded_f simp: image_constant[of 0]) have 2: "discrete_evolution (rk2_increment p (\t x. f (t, x))) (t + h) t (x t) - x t - h *\<^sub>R f (t, x t) - (h\<^sup>2 / 2) *\<^sub>R f' (t, x t) $ (1, f (t, x t)) \ (*\<^sub>R) (1 / 2) ` closure (convex hull (\s. (h ^ 3 * p / 2) *\<^sub>R f'' (t + s * (h * p), x t + (s * (h * p)) *\<^sub>R f (t, x t)) $ (1, f (t, x t)) $ (1, f (t, x t))) ` cbox 0 1)" by (rule integral_by_parts_in_bounded_set[OF integral_minus discrete_Taylor[unfolded interval_cbox]]) (auto intro!: bounded_scaleR_image bounded_blinfun_apply_image bounded_f''_2 simp: image_constant[of 0]) have "x (t + h) - discrete_evolution (rk2_increment p (\t x. f (t, x))) (t + h) t (x t) \ {a - b|a b. a \ closure (convex hull (*\<^sub>R) (1/3) ` (\xa. (h ^ 3 / 2) *\<^sub>R (f'' (h * xa + t, x (h * xa + t)) $ (1, f (h * xa + t, x (h * xa + t))) $ (1, f (h * xa + t, x (h * xa + t))) + f' (h * xa + t, x (h * xa + t)) $ (0, f' (h * xa + t, x (h * xa + t)) $ (1, f (h * xa + t, x (h * xa + t)))))) ` cbox 0 1) \ b \ closure (convex hull (*\<^sub>R) (1 / 2) ` (\s. (h ^ 3 * p / 2) *\<^sub>R f'' (t + s * (h * p), x t + (s * (h * p)) *\<^sub>R f (t, x t)) $ (1, f (t, x t)) $ (1, f (t, x t))) ` cbox 0 1)}" using in_minus_Collect[OF 1 2] unfolding closure_scaleR convex_hull_scaling by auto also note closure_minus_Collect also note convex_hull_minus_Collect also define F' where "F' \ (\y. y - discrete_evolution (rk2_increment p (\t x. f (t, x))) (t + h) t (x t)) ` R" have "closure (convex hull {xa - y |xa y. xa \ (*\<^sub>R) (1 / 3) ` (\xa. (h ^ 3 / 2) *\<^sub>R (f'' (h * xa + t, x (h * xa + t)) $ (1, f (h * xa + t, x (h * xa + t))) $ (1, f (h * xa + t, x (h * xa + t))) + f' (h * xa + t, x (h * xa + t)) $ (0, f' (h * xa + t, x (h * xa + t)) $ (1, f (h * xa + t, x (h * xa + t)))))) ` cbox 0 1 \ y \ (*\<^sub>R) (1 / 2) ` (\s. (h ^ 3 * p / 2) *\<^sub>R f'' (t + s * (h * p), x t + (s * (h * p)) *\<^sub>R f (t, x t)) $ (1, f (t, x t)) $ (1, f (t, x t))) ` cbox 0 1}) \ F'" apply (rule closure_minimal) apply (rule hull_minimal) apply (safe) subgoal for _ _ _ _ _ s1 s2 using R[of s1 s2] by (force simp: F'_def heun_remainder1_def heun_remainder2_def algebra_simps) by (auto intro!: ccR simp: F'_def) finally show "x (t + h) \ R" by (auto simp: F'_def) qed end locale derivative_norm_bounded = derivative_on_prod T X f f' for T and X::"'a::euclidean_space set" and f f' + fixes B B' assumes nonempty: "T \ {}" "X \ {}" assumes X_bounded: "bounded X" assumes convex: "convex T" "convex X" assumes f_bounded: "\t x. t\T \ x\X \ norm (f t x) \ B" assumes f'_bounded: "\t x. t\T \ x\X \ onorm (f' (t,x)) \ B'" begin lemma f_bound_nonneg: "0 \ B" proof - from nonempty obtain t x where "t \ T" "x \ X" by auto have "0 \ norm (f t x)" by simp also have "\ \ B" by (rule f_bounded) fact+ finally show ?thesis . qed lemma f'_bound_nonneg: "0 \ B'" proof - from nonempty f_bounded ex_norm_eq_1[where 'a="real*'a"] obtain t x and d::"real*'a" where tx: "t \ T" "x \ X" "norm d = 1" by auto have "0 \ norm (f' (t, x) d)" by simp also have "... \ B'" using tx by (intro order_trans[OF onorm[OF has_derivative_bounded_linear[OF f']]]) (auto intro!: f'_bounded f' has_derivative_linear) finally show ?thesis . qed sublocale g?: global_lipschitz _ _ _ B' proof fix t assume "t \ T" show "B'-lipschitz_on X (f t)" proof (rule lipschitz_onI) show "0 \ B'" using f'_bound_nonneg . fix x y let ?I = "T \ X" have "convex ?I" by (intro convex convex_Times) moreover have "\x. x \ ?I \ ((\(t, x). f t x) has_derivative f' x) (at x within ?I)" "\x. x \ ?I \ onorm (f' x) \ B'" using f' f'_bounded by (auto simp add: intro!: f'_bounded has_derivative_linear) moreover assume "x \ X" "y \ X" with \t \ T\ have "(t, x) \ ?I" "(t, y) \ ?I" by simp_all ultimately have "norm ((\(t, x). f t x) (t, x) - (\(t, x). f t x) (t, y)) \ B' * norm ((t, x) - (t, y))" by (rule differentiable_bound) then show "dist (f t x) (f t y) \ B' * dist x y" by (simp add: dist_norm norm_Pair) qed qed definition euler_C::"'a itself \ real" where "euler_C (TYPE('a)) = (sqrt DIM('a) * (B' * (B + 1) / 2))" lemma euler_C_nonneg: "euler_C TYPE('a) \ 0" using f_bounded f_bound_nonneg f'_bound_nonneg by (simp add: euler_C_def) lemma euler_consistent_traj: fixes t assumes T: "{t..u} \ T" assumes x': "(x has_vderiv_on (\s. f s (x s))) {t..u}" assumes x: "\s. s \ {t..u} \ x s \ X" shows "consistent x t u (euler_C (TYPE('a))) 1 (euler_increment f)" proof from x' have x': "\s. s \ {t..u} \ (x has_vector_derivative f s (x s)) (at s within {t..u})" by (simp add: has_vderiv_on_def) fix h::real assume ht: "0 < h" "t + h \ u" hence "t < u" "0 < h\<^sup>2 / 2" by simp_all let ?d = "discrete_evolution (euler_increment f) (t + h) t (x t)" have "x (t + h) \ (\b. ?d + (h\<^sup>2 / 2) *\<^sub>R b) ` cbox (- (B' * (B + 1)) *\<^sub>R One) ((B' * (B + 1)) *\<^sub>R One)" proof (rule euler_consistent_traj_set[OF _ \t + h \ u\ T x' x]) fix s assume "s \ {t .. u}" then have "?d + (h\<^sup>2 / 2) *\<^sub>R (f' (s, x s)) (1, f s (x s)) \ (\b. ?d + (h\<^sup>2 / 2) *\<^sub>R b) ` ((\s. (f' (s, x s)) (1, f s (x s)))` {t .. u})" by auto also have "\ \ (\b. ?d + (h\<^sup>2 / 2) *\<^sub>R b) ` cbox (- (B' * (B + 1)) *\<^sub>R One) ((B' * (B + 1)) *\<^sub>R One)" proof (rule image_mono, safe) fix s assume "s \ {t .. u}" with T have "norm (f' (s, x s) (1, f s (x s))) \ onorm (f' (s, x s)) * norm (1::real, f s (x s))" by (force intro!: onorm has_derivative_bounded_linear f' x) also have "\ \ B' * (B + 1)" using T \s \ _\ by (force intro!: mult_mono f'_bounded f_bounded f'_bound_nonneg x order_trans[OF norm_Pair_le]) finally have "f' (s, x s) (1, f s (x s)) \ cball 0 (B' * (B + 1))" by (auto simp: dist_norm mem_cball) also note cball_in_cbox finally show "f' (s, x s) (1, f s (x s)) \ cbox (- (B' * (B + 1)) *\<^sub>R One) ((B' * (B + 1)) *\<^sub>R One)" by simp qed finally show "?d + (h\<^sup>2 / 2) *\<^sub>R (f' (s, x s)) (1, f s (x s)) \ (\b. ?d + (h\<^sup>2 / 2) *\<^sub>R b) ` cbox (- (B' * (B + 1)) *\<^sub>R One) ((B' * (B + 1)) *\<^sub>R One)" . qed (auto intro!: less_imp_le[OF \0 < h\] bounded_plus_image bounded_scaleR_image bounded_cbox closed_scaling convex_scaling convex_box simp: image_constant_conv closed_translation_iff) then have "x (t + h) - discrete_evolution (euler_increment f) (t + h) t (x t) \ (*\<^sub>R) (h\<^sup>2 / 2) ` cbox (- (B' * (B + 1)) *\<^sub>R One) ((B' * (B + 1)) *\<^sub>R One)" by auto also have "\ = cbox (- ((h\<^sup>2 / 2) * (B' * (B + 1))) *\<^sub>R One) (((h\<^sup>2 / 2) * (B' * (B + 1))) *\<^sub>R One)" using f_bound_nonneg f'_bound_nonneg by (auto simp add: image_smult_cbox box_eq_empty mult_less_0_iff) also note centered_cbox_in_cball finally show "dist (x (t + h)) (discrete_evolution (euler_increment f) (t + h) t (x t)) \ euler_C(TYPE('a)) * h ^ (1 + 1)" by (auto simp: euler_C_def dist_norm algebra_simps norm_minus_commute power2_eq_square mem_cball) qed lemma derivative_norm_bounded_subset: assumes X'_ne: "X' \ {}" and X'_subset: "X' \ X" and "convex X'" shows "derivative_norm_bounded T X' f f' B B'" proof - interpret derivative_on_prod T X' f f' using X'_subset by (rule derivative_on_prod_subset) show ?thesis using X'_subset by unfold_locales (auto intro!: nonempty X'_ne bounded_subset[OF X_bounded X'_subset] convex f_bounded f'_bounded \convex X'\) qed end locale grid_from = grid + fixes t0 assumes grid_min: "t0 = t 0" locale euler_consistent = derivative_norm_bounded T X' f f' B B' for T f X X' B f' B' + fixes solution t0 x0 r e assumes sol: "(solution solves_ode f) T X" and sol_iv: "solution t0 = x0" and iv_defined: "t0 \ T" assumes domain_subset: "X \ X'" assumes interval: "T = {t0 - e .. t0 + e}" assumes lipschitz_area: "\t. t \ T \ cball (solution t) \r\ \ X'" begin lemma euler_consistent_solution: fixes t' assumes t': "t' \ {t0 .. t0 + e}" shows "consistent solution t' (t0 + e) (euler_C(TYPE('a))) 1 (euler_increment f)" proof (rule euler_consistent_traj) show "{t'..t0 + e} \ T" using t' interval by simp with solves_odeD(1)[OF sol] show "(solution has_vderiv_on (\s. f s (solution s))) {t'..t0 + e}" by (rule has_vderiv_on_subset) fix s assume "s \ {t'..t0 + e}" with solves_odeD(2)[OF sol, of s] show "solution s \ X'" using domain_subset \{t' .. t0 + e} \ T\ by (auto simp: subset_iff) qed end subsection \Euler method is convergent\ text\\label{sec:rk-euler-conv}\ locale max_step1 = grid + fixes t1 L B r assumes max_step: "\j. t j \ t1 \ max_stepsize j \ \r\ * L / B / (exp (L * (t1 - t 0) + 1) - 1)" sublocale max_step1 < max_step?: max_step t t1 1 L B r using max_step by unfold_locales simp_all locale euler_convergent = euler_consistent T f "X::'a::euclidean_space set" + max_step1 t "t0 + e" B' "euler_C(TYPE('a))" r for T f X t + assumes grid_from: "t0 = t 0" subsection \Euler method on Rectangle is convergent\ text\\label{sec:rk-euler-conv-on-rect}\ locale ivp_rectangle_bounded_derivative = solution_in_cylinder t0 T x0 b f X B + outer?: derivative_norm_bounded T "cball x0 r" f f' B B' for t0 T x0 b f X B f' B' r e + assumes T_def: "T \ cball t0 e" assumes subset_cylinders: "b < r" assumes positive_time: "0 < e" sublocale ivp_rectangle_bounded_derivative \ unique_on_cylinder t0 T x0 b X f B B' by unfold_locales (insert subset_cylinders positive_time, auto intro!: lipschitz_on_subset[OF lipschitz] simp: mem_cball) sublocale ivp_rectangle_bounded_derivative \ euler_consistent T f X "cball x0 r" B f' B' solution t0 x0 "r - b" e proof - interpret derivative_norm_bounded T X f f' B B' using b_pos subset_cylinders by (intro outer.derivative_norm_bounded_subset) (auto simp: X_def mem_cball) show "euler_consistent T f X (cball x0 r) B f' B' solution t0 x0 (r - b) e" apply unfold_locales subgoal by (rule solution_solves_ode) subgoal by (rule solution_iv) subgoal by (rule initial_time_in) subgoal using subset_cylinders by (auto simp: X_def mem_cball) subgoal by (auto simp add: T_def dist_real_def mem_cball) subgoal premises prems for t proof safe fix x have "dist x x0 \ dist x (solution t) + dist x0 (solution t)" by (rule dist_triangle2) also assume "x \ cball (solution t) \r - b\" then have "dist x (solution t) \ r - b" using subset_cylinders by (simp add: dist_commute mem_cball) also have "{t0--t} \ T" by (rule subset_T[OF prems]) then have "dist x0 (solution t) \ B * \t - t0\" using subset_cylinders is_solution_in_cone[of t, OF prems solves_ode_on_subset[OF solution_solves_ode _ order_refl] solution_iv] by (simp add: mem_cball) also have "B * abs (t - t0) \ b" using e_bounded[OF prems] b_pos B_nonneg by (auto simp: dist_real_def divide_simps ac_simps split: if_splits) finally show "x \ cball x0 r" by (simp add: dist_commute mem_cball) qed done qed locale euler_on_rectangle = ivp_rectangle_bounded_derivative t0 T x0 b f X B f' B' r e + grid_from t t0 + max_step1 t "t0 + e" B' "outer.euler_C(TYPE('a::euclidean_space))" "r - b" for t0 T x0 f X t e b r B f' B' sublocale euler_consistent \ consistent_one_step t0 "t0 + e" solution "euler_increment f" 1 "euler_C(TYPE('a))" r B' proof show "0 < (1::nat)" by simp show "0 \ euler_C(TYPE('a))" using euler_C_nonneg by simp show "0 \ B'" using lipschitz_on_nonneg[OF lipschitz] iv_defined by simp fix s x assume s: "s \ {t0 .. t0 + e}" show "consistent solution s (t0 + e) (euler_C(TYPE('a))) 1 (euler_increment f)" using interval s f_bounded f'_bounded f' strip by (intro euler_consistent_solution) auto fix h assume "h \ {0 .. t0 + e - s}" have "B'-lipschitz_on X' (euler_increment f h s)" using s lipschitz interval strip by (auto intro!: euler_lipschitz) thus "B'-lipschitz_on (cball (solution s) \r\) (euler_increment f h s)" using s interval by (auto intro: lipschitz_on_subset[OF _ lipschitz_area]) qed sublocale euler_convergent \ convergent_one_step t0 "t0 + e" solution "euler_increment f" 1 "euler_C(TYPE('a))" r B' t by unfold_locales (simp add: grid_from) sublocale euler_on_rectangle \ convergent?: euler_convergent "cball x0 r" B f' B' solution t0 x0 "r - b" e T f X t proof unfold_locales qed (rule grid_min) lemma "B \ (0::real) \ 0 \ (exp (B + 1) - 1)" by (simp add: algebra_simps) context euler_on_rectangle begin lemma convergence: assumes "t j \ t0 + e" shows "dist (solution (t j)) (euler f x0 t j) \ sqrt DIM('a) * (B + 1) / 2 * (exp (B' * e + 1) - 1) * max_stepsize j" proof - have "dist (solution (t j)) (euler f x0 t j) \ sqrt DIM('a) * (B + 1) / 2 * B' / B' * ((exp (B' * e + 1) - 1) * max_stepsize j)" using assms convergence[OF assms] f'_bound_nonneg unfolding euler_C_def by (simp add: euler_def grid_min[symmetric] ac_simps) also have "\ \ sqrt DIM('a) * (B + 1) / 2 * ((exp (B' * e + 1) - 1) * max_stepsize j)" using f_bound_nonneg f'_bound_nonneg by (auto intro!: mult_right_mono mult_nonneg_nonneg max_stepsize_nonneg add_nonneg_nonneg simp: le_diff_eq) finally show ?thesis by simp qed end subsection \Stability and Convergence of Approximate Euler \label{sec:rk-euler-stable}\ locale euler_rounded_on_rectangle = ivp_rectangle_bounded_derivative t0 T x0 b f X B f' B' r e1' + grid?: grid_from t t0' + max_step_r_2?: max_step1 t "t0 + e2'" B' "euler_C(TYPE('a))" "(r - b)/2" for t0 T x0 f and X::"'a::executable_euclidean_space set" and t :: "nat \ real" and t0' e1' e2'::real and x0' :: "'a" and b r B f' B' + fixes g::"real \ 'a \ 'a" and e::int assumes t0_float: "t0 = t0'" assumes ordered_bounds: "e1' \ e2'" assumes approx_f_e: "\j x. t j \ t0 + e1' \ dist (f (t j) x) ((g (t j) x)) \ sqrt (DIM('a)) * 2 powr -e" assumes initial_error: "dist x0 (x0') \ euler_C(TYPE('a)) / B' * (exp 1 - 1) * stepsize 0" assumes rounding_error: "\j. t j \ t0 + e1' \ sqrt (DIM('a)) * 2 powr -e \ euler_C(TYPE('a)) / 2 * stepsize j" begin lemma approx_f: "t j \ t0 + e1' \ dist (f (t j) (x)) ((g (t j) (x))) \ euler_C(TYPE('a)) / 2 * stepsize j" using approx_f_e[of j x] rounding_error[of j] by auto lemma t0_le: "t 0 \ t0 + e1'" unfolding grid_min[symmetric] t0_float[symmetric] by (metis atLeastAtMost_iff interval iv_defined(1)) end sublocale euler_rounded_on_rectangle \ grid'?: grid_from t t0' using grid t0_float grid_min by unfold_locales auto sublocale euler_rounded_on_rectangle \ max_step_r?: max_step1 t "t0 + e2'" B' "euler_C(TYPE('a))" "r - b" proof unfold_locales fix j assume *: "(t j) \ t0 + e2'" moreover from * grid_mono[of 0 j] have "t 0 \ t0 + e2'" by (simp add: less_eq_float_def) ultimately show "max_stepsize j \ \r - b\ * B' / euler_C(TYPE('a)) / (exp (B' * (t0 + e2' - (t 0)) + 1) - 1)" using max_step_mono_r lipschitz B_nonneg f'_bound_nonneg by (auto simp: less_eq_float_def euler_C_def) qed lemma max_step1_mono: assumes "t 0 \ t1" assumes "t1 \ t2" assumes "0 \ a" assumes "0 \ b" assumes ms2: "max_step1 t t2 a b c" shows "max_step1 t t1 a b c" proof - interpret t2: max_step1 t t2 a b c using ms2 . show ?thesis proof fix j assume "t j \ t1" hence "t j \ t2" using assms by simp hence "t2.max_stepsize j \ \c\ * a / b / (exp (a * (t2 - t 0) + 1) - 1)" (is "_ \ ?x t2") by (rule t2.max_step) also have "\ \ ?x t1" using assms by (cases "b = 0") (auto intro!: divide_left_mono mult_mono abs_ge_zero add_increasing mult_pos_pos add_strict_increasing2 simp: le_diff_eq less_diff_eq) finally show "t2.max_stepsize j \ ?x t1" . qed qed sublocale euler_rounded_on_rectangle \ max_step_r1?: max_step1 t "t0 + e1'" B' "euler_C(TYPE('a))" "r - b" by (rule max_step1_mono[of t, OF t0_le add_left_mono[OF ordered_bounds] f'_bound_nonneg euler_C_nonneg]) unfold_locales sublocale euler_rounded_on_rectangle \ c?: euler_on_rectangle t0 T x0 f X t e1' b r B f' B' using t0_float grid_min by unfold_locales simp sublocale euler_rounded_on_rectangle \ consistent_one_step "t 0" "t0 + e1'" solution "euler_increment f" 1 "euler_C(TYPE('a))" "r - b" B' using consistent_nonneg consistent lipschitz_nonneg lipschitz_incr t0_float grid_min by unfold_locales simp_all sublocale euler_rounded_on_rectangle \ max_step1 t "t0 + e1'" B' "euler_C(TYPE('a))" "(r - b) / 2" by (rule max_step1_mono[of t, OF t0_le add_left_mono[OF ordered_bounds] f'_bound_nonneg euler_C_nonneg]) unfold_locales sublocale euler_rounded_on_rectangle \ one_step?: rounded_one_step t "t0 + e1'" solution "euler_increment f" 1 "euler_C(TYPE('a))" "r - b" B' "euler_increment' e g" x0' proof fix h j x assume "t j \ t0 + e1'" have "dist (euler_increment f (h) (t j) (x)) ((euler_increment' e g h (t j) x)) = dist (f (t j) (x)) ((eucl_down e (g (t j) (x))))" by (simp add: euler_increment euler_float_increment) also have "... \ dist (f (t j) (x)) ((g (t j) (x))) + dist ((g (t j) (x))) ((eucl_down e (g (t j) (x))))" by (rule dist_triangle) also from approx_f[OF \t j \ t0 + e1'\] have "dist (f (t j) (x)) ((g (t j) (x))) \ euler_C(TYPE('a)) / 2 * stepsize j" . also from eucl_truncate_down_correct[of "g (t j) (x)" e] have "dist ((g (t j) (x))) ((eucl_down e (g (t j) (x)))) \ sqrt (DIM('a)) * 2 powr - e" by simp also have "sqrt (DIM('a)) * 2 powr -e \ euler_C(TYPE('a)) / 2 * stepsize j" using rounding_error \t j \ t0 + e1'\ . finally have "dist (euler_increment f (h) (t j) (x)) ((euler_increment' e g h (t j) x)) \ euler_C(TYPE('a)) * stepsize j" by arith thus "dist (euler_increment f h (t j) (x)) ((euler_increment' e g h (t j) x)) \ euler_C(TYPE('a)) * stepsize j ^ 1" by simp qed (insert initial_error, simp add: grid_min[symmetric]) context euler_rounded_on_rectangle begin lemma stability: assumes "t j \ t0 + e1'" shows "dist (euler' e g x0' t j) (euler f x0 t j) \ sqrt DIM('a) * (B + 1) / 2 * (exp (B' * e1' + 1) - 1) * max_stepsize j" proof - have "dist ((euler' e g x0' t j)) (euler f x0 t j) \ sqrt DIM('a) * (B + 1) / 2 * B' / B' * (exp (B' * e1' + 1) - 1) * max_stepsize j" using assms stability[OF assms] unfolding grid_min[symmetric] euler_C_def sol_iv by (auto simp add: euler_def euler'_def) also have "\ \ sqrt DIM('a) * (B + 1) / 2 * ((exp (B' * e1' + 1) - 1) * max_stepsize j)" using f_bound_nonneg f'_bound_nonneg by (auto intro!: mult_right_mono mult_nonneg_nonneg max_stepsize_nonneg add_nonneg_nonneg simp: le_diff_eq) finally show ?thesis by simp qed lemma convergence_float: assumes "t j \ t0 + e1'" shows "dist (solution (t j)) (euler' e g x0' t j) \ sqrt DIM('a) * (B + 1) * (exp (B' * e1' + 1) - 1) * max_stepsize j" proof - have "dist (solution ((t j))) ((euler' e g x0' t j)) \ dist (solution ((t j))) (euler f x0 t j) + dist ((euler' e g x0' t j)) (euler f x0 t j)" by (rule dist_triangle2) also have "dist (solution ((t j))) (euler f x0 t j) \ sqrt DIM('a) * (B + 1) / 2 * (exp (B' * e1' + 1) - 1) * max_stepsize j" using assms convergence[OF assms] t0_float by simp also have "dist ((euler' e g x0' t j)) (euler f x0 t j) \ sqrt DIM('a) * (B + 1) / 2 * (exp (B' * e1' + 1) - 1) * max_stepsize j" using assms stability by simp finally have "dist (solution ((t j))) ((euler' e g x0' t j)) \ sqrt DIM('a) * (B + 1) / 2 * (exp (B' * (e1') + 1) - 1) * max_stepsize j + sqrt DIM('a) * (B + 1) / 2 * (exp (B' * (e1') + 1) - 1) * max_stepsize j" by simp thus ?thesis by (simp add: field_simps) qed end end diff --git a/thys/Ordinary_Differential_Equations/ODE_Auxiliarities.thy b/thys/Ordinary_Differential_Equations/ODE_Auxiliarities.thy --- a/thys/Ordinary_Differential_Equations/ODE_Auxiliarities.thy +++ b/thys/Ordinary_Differential_Equations/ODE_Auxiliarities.thy @@ -1,620 +1,620 @@ section \Auxiliary Lemmas\ theory ODE_Auxiliarities imports "HOL-Analysis.Analysis" "HOL-Library.Float" "List-Index.List_Index" Affine_Arithmetic.Affine_Arithmetic_Auxiliarities Affine_Arithmetic.Executable_Euclidean_Space begin instantiation prod :: (zero_neq_one, zero_neq_one) zero_neq_one begin definition "1 = (1, 1)" instance by standard (simp add: zero_prod_def one_prod_def) end subsection \there is no inner product for type @{typ "'a \\<^sub>L 'b"}\ lemma (in real_inner) parallelogram_law: "(norm (x + y))\<^sup>2 + (norm (x - y))\<^sup>2 = 2 * (norm x)\<^sup>2 + 2 * (norm y)\<^sup>2" proof - have "(norm (x + y))\<^sup>2 + (norm (x - y))\<^sup>2 = inner (x + y) (x + y) + inner (x - y) (x - y)" by (simp add: norm_eq_sqrt_inner) also have "\ = 2 * (norm x)\<^sup>2 + 2 * (norm y)\<^sup>2" by (simp add: algebra_simps norm_eq_sqrt_inner) finally show ?thesis . qed locale no_real_inner begin lift_definition fstzero::"(real*real) \\<^sub>L (real*real)" is "\(x, y). (x, 0)" by (auto intro!: bounded_linearI') lemma [simp]: "fstzero (a, b) = (a, 0)" by transfer simp lift_definition zerosnd::"(real*real) \\<^sub>L (real*real)" is "\(x, y). (0, y)" by (auto intro!: bounded_linearI') lemma [simp]: "zerosnd (a, b) = (0, b)" by transfer simp lemma fstzero_add_zerosnd: "fstzero + zerosnd = id_blinfun" by transfer auto lemma norm_fstzero_zerosnd: "norm fstzero = 1" "norm zerosnd = 1" "norm (fstzero - zerosnd) = 1" by (rule norm_blinfun_eqI[where x="(1, 0)"]) (auto simp: norm_Pair blinfun.bilinear_simps intro: norm_blinfun_eqI[where x="(0, 1)"] norm_blinfun_eqI[where x="(1, 0)"]) text \compare with @{thm parallelogram_law}\ lemma "(norm (fstzero + zerosnd))\<^sup>2 + (norm (fstzero - zerosnd))\<^sup>2 \ 2 * (norm fstzero)\<^sup>2 + 2 * (norm zerosnd)\<^sup>2" by (simp add: fstzero_add_zerosnd norm_fstzero_zerosnd) end subsection \Topology\ subsection \Vector Spaces\ lemma ex_norm_eq_1: "\x. norm (x::'a::{real_normed_vector, perfect_space}) = 1" by (metis vector_choose_size zero_le_one) subsection \Reals\ subsection \Balls\ text \sometimes @{thm mem_ball} etc. are not good \[simp]\ rules (although they are often useful): not sure that inequalities are ``simpler'' than set membership (distorts automatic reasoning when only sets are involved)\ lemmas [simp del] = mem_ball mem_cball mem_sphere mem_ball_0 mem_cball_0 subsection \Boundedness\ lemma bounded_subset_cboxE: assumes "\i. i \ Basis \ bounded ((\x. x \ i) ` X)" obtains a b where "X \ cbox a b" proof - have "\i. i \ Basis \ \a b. ((\x. x \ i) ` X) \ {a..b}" by (metis box_real(2) box_subset_cbox subset_trans bounded_subset_box_symmetric[OF assms] ) then obtain a b where bnds: "\i. i \ Basis \ ((\x. x \ i) ` X) \ {a i .. b i}" by metis then have "X \ {x. \i\Basis. x \ i \ {a i .. b i}}" by force also have "\ = cbox (\i\Basis. a i *\<^sub>R i) (\i\Basis. b i *\<^sub>R i)" by (auto simp: cbox_def) finally show ?thesis .. qed lemma bounded_euclideanI: assumes "\i. i \ Basis \ bounded ((\x. x \ i) ` X)" shows "bounded X" proof - from bounded_subset_cboxE[OF assms] obtain a b where "X \ cbox a b" . with bounded_cbox show ?thesis by (rule bounded_subset) qed subsection \Intervals\ notation closed_segment ("(1{_--_})") notation open_segment ("(1{_<--<_})") lemma min_zero_mult_nonneg_le: "0 \ h' \ h' \ h \ min 0 (h * k::real) \ h' * k" by (metis dual_order.antisym le_cases min_le_iff_disj mult_eq_0_iff mult_le_0_iff mult_right_mono_neg) lemma max_zero_mult_nonneg_le: "0 \ h' \ h' \ h \ h' * k \ max 0 (h * k::real)" by (metis dual_order.antisym le_cases le_max_iff_disj mult_eq_0_iff mult_right_mono zero_le_mult_iff) lemmas closed_segment_eq_real_ivl = closed_segment_eq_real_ivl lemma bdd_above_is_intervalI: "bdd_above I" if "is_interval I" "a \ b" "a \ I" "b \ I" for I::"real set" by (meson bdd_above_def is_interval_1 le_cases that) lemma bdd_below_is_intervalI: "bdd_below I" if "is_interval I" "a \ b" "a \ I" "b \ I" for I::"real set" by (meson bdd_below_def is_interval_1 le_cases that) subsection \Extended Real Intervals\ subsection \Euclidean Components\ subsection \Operator Norm\ subsection \Limits\ lemma eventually_open_cball: assumes "open X" assumes "x \ X" shows "eventually (\e. cball x e \ X) (at_right 0)" proof - from open_contains_cball_eq[OF assms(1)] assms(2) obtain e where "e > 0" "cball x e \ X" by auto thus ?thesis by (auto simp: eventually_at dist_real_def mem_cball intro!: exI[where x=e]) qed subsection \Continuity\ subsection \Derivatives\ lemma if_eventually_has_derivative: assumes "(f has_derivative F') (at x within S)" assumes "\\<^sub>F x in at x within S. P x" "P x" "x \ S" shows "((\x. if P x then f x else g x) has_derivative F') (at x within S)" using assms(1) apply (rule has_derivative_transform_eventually) subgoal using assms(2) by eventually_elim auto by (auto simp: assms) lemma norm_le_in_cubeI: "norm x \ norm y" if "\i. i \ Basis \ abs (x \ i) \ abs (y \ i)" for x y unfolding norm_eq_sqrt_inner apply (subst euclidean_inner) apply (subst (3) euclidean_inner) using that by (auto intro!: sum_mono simp: abs_le_square_iff power2_eq_square[symmetric]) lemma has_derivative_partials_euclidean_convexI: fixes f::"'a::euclidean_space \ 'b::real_normed_vector" assumes f': "\i x xi. i \ Basis \ (\j\Basis. x \ j \ X j) \ xi = x \ i \ ((\p. f (x + (p - x \ i) *\<^sub>R i)) has_vector_derivative f' i x) (at xi within X i)" assumes df_cont: "\i. i \ Basis \ (f' i \ (f' i x)) (at x within {x. \j\Basis. x \ j \ X j})" assumes "\i. i \ Basis \ x \ i \ X i" assumes "\i. i \ Basis \ convex (X i)" shows "(f has_derivative (\h. \j\Basis. (h \ j) *\<^sub>R f' j x)) (at x within {x. \j\Basis. x \ j \ X j})" (is "_ (at x within ?S)") proof (rule has_derivativeI) show "bounded_linear (\h. \j\Basis. (h \ j) *\<^sub>R f' j x)" by (auto intro!: bounded_linear_intros) obtain E where [simp]: "set E = (Basis::'a set)" "distinct E" using finite_distinct_list[OF finite_Basis] by blast have [simp]: "length E = DIM('a)" using \distinct E\ distinct_card by fastforce have [simp]: "E ! j \ Basis" if "j < DIM('a)" for j by (metis \length E = DIM('a)\ \set E = Basis\ nth_mem that) have "convex ?S" by (rule convex_prod) (use assms in auto) have sum_Basis_E: "sum g Basis = (\jdistinct E\ \length E = DIM('a)\ \set E = Basis\ bij_betw_def bij_betw_nth) have segment: "\\<^sub>F x' in at x within ?S. x' \ ?S" "\\<^sub>F x' in at x within ?S. x' \ x" unfolding eventually_at_filter by auto show "((\y. (f y - f x - (\j\Basis. ((y - x) \ j) *\<^sub>R f' j x)) /\<^sub>R norm (y - x)) \ 0) (at x within {x. \j\Basis. x \ j \ X j})" apply (rule tendstoI) unfolding norm_conv_dist[symmetric] proof - fix e::real assume "e > 0" define B where "B = e / norm (2*DIM('a) + 1)" with \e > 0\ have B_thms: "B > 0" "2 * DIM('a) * B < e" "B \ 0" by (auto simp: divide_simps algebra_simps B_def) define B' where "B' = B / 2" have "B' > 0" by (simp add: B'_def \0 < B\) have "\i \ Basis. \\<^sub>F xa in at x within {x. \j\Basis. x \ j \ X j}. dist (f' i xa) (f' i x) < B'" apply (rule ballI) subgoal premises prems using df_cont[OF prems, THEN tendstoD, OF \0 < B'\] . done from eventually_ball_finite[OF finite_Basis this] have "\\<^sub>F x' in at x within {x. \j\Basis. x \ j \ X j}. \j\Basis. dist (f' j x') (f' j x) < B'" . then obtain d where "d > 0" and "\x' j. x' \ {x. \j\Basis. x \ j \ X j} \ x' \ x \ dist x' x < d \ j \ Basis \ dist (f' j x') (f' j x) < B'" using \0 < B'\ by (auto simp: eventually_at) then have B': "x' \ {x. \j\Basis. x \ j \ X j} \ dist x' x < d \ j \ Basis \ dist (f' j x') (f' j x) < B'" for x' j by (cases "x' = x", auto simp add: \0 < B'\) then have B: "norm (f' j x' - f' j y) < B" if "(\j. j \ Basis \ x' \ j \ X j)" "(\j. j \ Basis \ y \ j \ X j)" "dist x' x < d" "dist y x < d" "j \ Basis" for x' y j proof - have "dist (f' j x') (f' j x) < B'" "dist (f' j y) (f' j x) < B'" using that by (auto intro!: B') then have "dist (f' j x') (f' j y) < B' + B'" by norm also have "\ = B" by (simp add: B'_def) finally show ?thesis by (simp add: dist_norm) qed have "\\<^sub>F x' in at x within {x. \j\Basis. x \ j \ X j}. dist x' x < d" by (rule tendstoD[OF tendsto_ident_at \d > 0\]) with segment show "\\<^sub>F x' in at x within {x. \j\Basis. x \ j \ X j}. norm ((f x' - f x - (\j\Basis. ((x' - x) \ j) *\<^sub>R f' j x)) /\<^sub>R norm (x' - x)) < e" proof eventually_elim case (elim x') then have os_subset: "open_segment x x' \ ?S" using \convex ?S\ assms(3) unfolding convex_contains_open_segment by auto then have cs_subset: "closed_segment x x' \ ?S" using elim assms(3) by (auto simp add: open_segment_def) have csc_subset: "closed_segment (x' \ i) (x \ i) \ X i" if i: "i \ Basis" for i apply (rule closed_segment_subset) using cs_subset elim assms(3,4) that by (auto ) have osc_subset: "open_segment (x' \ i) (x \ i) \ X i" if i: "i \ Basis" for i using segment_open_subset_closed csc_subset[OF i] by (rule order_trans) define h where "h = x' - x" define z where "z j = (\k E ! k) *\<^sub>R (E ! k))" for j define g where "g j t = (f (x + z j + (t - x \ E ! j) *\<^sub>R E ! j))" for j t have z: "z j \ E ! j = 0" if "j < DIM('a)" for j using that by (auto simp: z_def h_def algebra_simps inner_sum_left inner_Basis if_distrib nth_eq_iff_index_eq sum.delta intro!: euclidean_eqI[where 'a='a] cong: if_cong) from distinct_Ex1[OF \distinct E\, unfolded \set E = Basis\ Ex1_def \length E = _\] obtain index where index: "\i. i \ Basis \ i = E ! index i" "\i. i \ Basis \ index i < DIM('a)" and unique: "\i j. i \ Basis \ j < DIM('a) \ E ! j = i \ j = index i" by metis have nth_eq_iff_index: "E ! k = i \ index i = k" if "i \ Basis" "k < DIM('a)" for k i using unique[OF that] index[OF \i \ Basis\] by auto have z_inner: "z j \ i = (if j \ index i then 0 else h \ i)" if "j < DIM('a)" "i \ Basis" for j i using that index[of i] by (auto simp: z_def h_def algebra_simps inner_sum_left inner_Basis if_distrib sum.delta nth_eq_iff_index intro!: euclidean_eqI[where 'a='a] cong: if_cong) have z_mem: "j < DIM('a) \ ja \ Basis \ x \ ja + z j \ ja \ X ja" for j ja using csc_subset by (auto simp: z_inner h_def algebra_simps) have "norm (x' - x) < d" using elim by (simp add: dist_norm) have norm_z': "y \ closed_segment (x \ E ! j) (x' \ E ! j) \ norm (z j + y *\<^sub>R E ! j - (x \ E ! j) *\<^sub>R E ! j) < d" if "j < DIM('a)" for j y apply (rule le_less_trans[OF _ \norm (x' - x) < d\]) apply (rule norm_le_in_cubeI) apply (auto simp: inner_diff_left inner_add_left inner_Basis that z) subgoal by (auto simp: closed_segment_eq_real_ivl split: if_splits) subgoal for i using that by (auto simp: z_inner h_def algebra_simps) done have norm_z: "norm (z j) < d" if "j < DIM('a)" for j using norm_z'[OF that ends_in_segment(1)] by (auto simp: z_def) { fix j assume j: "j < DIM('a)" have eq: "(x + z j + ((y - (x + z j)) \ E ! j) *\<^sub>R E ! j + (p - (x + z j + ((y - (x + z j)) \ E ! j) *\<^sub>R E ! j) \ E ! j) *\<^sub>R E ! j) = (x + z j + (p - (x \ E ! j)) *\<^sub>R E ! j)" for y p by (auto simp: algebra_simps j z) have f_has_derivative: "((\p. f (x + z j + (p - x \ E ! j) *\<^sub>R E ! j)) has_derivative (\xa. xa *\<^sub>R f' (E ! j) (x + z j + ((y *\<^sub>R E ! j - (x + z j)) \ E ! j) *\<^sub>R E ! j))) (at y within closed_segment (x \ E ! j) (x' \ E ! j))" if "y \ closed_segment (x \ E ! j) (x' \ E ! j)" for y - apply (rule has_derivative_within_subset) + apply (rule has_derivative_subset) apply (rule f'[unfolded has_vector_derivative_def, where x= "x + z j + ((y *\<^sub>R E!j - (x + z j))\ E!j) *\<^sub>R E ! j" and i="E ! j", unfolded eq]) subgoal by (simp add: j) subgoal using that apply (auto simp: algebra_simps z j inner_Basis) using closed_segment_commute \E ! j \ Basis\ csc_subset apply blast by (simp add: z_mem j) subgoal by (auto simp: algebra_simps z j inner_Basis) subgoal apply (auto simp: algebra_simps z j inner_Basis) using closed_segment_commute \\j. j < DIM('a) \ E ! j \ Basis\ csc_subset j apply blast done done have *: "((xa *\<^sub>R E ! j - (x + z j)) \ E ! j) = xa - x \ E ! j" for xa by (auto simp: algebra_simps z j) have g': "(g j has_vector_derivative (f' (E ! j) (x + z j + (xa - x \ E ! j) *\<^sub>R E ! j))) (at xa within (closed_segment (x\E!j) (x'\E!j)))" (is "(_ has_vector_derivative ?g' j xa) _") if "xa \ closed_segment (x\E!j) (x'\E!j)" for xa using that by (auto simp: has_vector_derivative_def g_def[abs_def] * intro!: derivative_eq_intros f_has_derivative[THEN has_derivative_eq_rhs]) define g' where "g' j = ?g' j" for j with g' have g': "(g j has_vector_derivative g' j t) (at t within closed_segment (x\E!j) (x'\E!j))" if "t \ closed_segment (x\E!j) (x'\E!j)" for t by (simp add: that) have cont_bound: "\y. y\closed_segment (x \ E ! j) (x' \ E ! j) \ norm (g' j y - g' j (x \ E ! j)) \ B" apply (auto simp add: g'_def j algebra_simps inner_Basis z dist_norm intro!: less_imp_le B z_mem norm_z norm_z') using closed_segment_commute \\j. j < DIM('a) \ E ! j \ Basis\ csc_subset j apply blast done from vector_differentiable_bound_linearization[OF g' order_refl cont_bound ends_in_segment(1)] have n: "norm (g j (x' \ E ! j) - g j (x \ E ! j) - (x' \ E ! j - x \ E ! j) *\<^sub>R g' j (x \ E ! j)) \ norm (x' \ E ! j - x \ E ! j) * B" . have "z (Suc j) = z j + (x' \ E ! j - x \ E ! j) *\<^sub>R E ! j" by (auto simp: z_def h_def algebra_simps) then have "f (x + z (Suc j)) = f (x + z j + (x' \ E ! j - x \ E ! j) *\<^sub>R E ! j) " by (simp add: ac_simps) with n have "norm (f (x + z (Suc j)) - f (x + z j) - (x' \ E ! j - x \ E ! j) *\<^sub>R f' (E ! j) (x + z j)) \ \x' \ E ! j - x \ E ! j\ * B" by (simp add: g_def g'_def) } note B_le = this have B': "norm (f' (E ! j) (x + z j) - f' (E ! j) x) \ B" if "j < DIM('a)" for j using that assms(3) by (auto simp add: algebra_simps inner_Basis z dist_norm \0 < d\ intro!: less_imp_le B z_mem norm_z) have "(\j\DIM('a) - 1. f (x + z j) - f (x + z (Suc j))) = f (x + z 0) - f (x + z (Suc (DIM('a) - 1)))" by (rule sum_telescope) moreover have "z DIM('a) = h" using index by (auto simp: z_def h_def algebra_simps inner_sum_left inner_Basis if_distrib nth_eq_iff_index sum.delta intro!: euclidean_eqI[where 'a='a] cong: if_cong) moreover have "z 0 = 0" by (auto simp: z_def) moreover have "{..DIM('a) - 1} = {..jj\Basis. ((x' - x) \ j) *\<^sub>R f' j x)) = norm( (\j E ! j - x \ E ! j) *\<^sub>R f' (E ! j) (x + z j)) + (\j E ! j - x \ E ! j) *\<^sub>R (f' (E ! j) (x + z j) - f' (E ! j) x)))" (is "_ = norm (sum ?a ?E + sum ?b ?E)") by (intro arg_cong[where f=norm]) (simp add: sum_negf sum_subtractf sum.distrib algebra_simps sum_Basis_E) also have "\ \ norm (sum ?a ?E) + norm (sum ?b ?E)" by (norm) also have "norm (sum ?a ?E) \ sum (\x. norm (?a x)) ?E" by (rule norm_sum) also have "\ \ sum (\j. norm \x' \ E ! j - x \ E ! j\ * B) ?E" by (auto intro!: sum_mono B_le) also have "\ \ sum (\j. norm (x' - x) * B) ?E" apply (rule sum_mono) apply (auto intro!: mult_right_mono \0 \ B\) by (metis (full_types) \\j. j < DIM('a) \ E ! j \ Basis\ inner_diff_left norm_bound_Basis_le order_refl) also have "\ = norm (x' - x) * DIM('a) * B" by simp also have "norm (sum ?b ?E) \ sum (\x. norm (?b x)) ?E" by (rule norm_sum) also have "\ \ sum (\j. norm (x' - x) * B) ?E" apply (intro sum_mono) apply (auto intro!: mult_mono B') apply (metis (full_types) \\j. j < DIM('a) \ E ! j \ Basis\ inner_diff_left norm_bound_Basis_le order_refl) done also have "\ = norm (x' - x) * DIM('a) * B" by simp finally have "norm (f (x + h) - f x - (\j\Basis. ((x' - x) \ j) *\<^sub>R f' j x)) \ norm (x' - x) * real DIM('a) * B + norm (x' - x) * real DIM('a) * B" by arith also have "\ /\<^sub>R norm (x' - x) \ norm (2 * DIM('a) * B)" using \B \ 0\ by (simp add: divide_simps abs_mult) also have "\ < e" using B_thms by simp finally show ?case by (auto simp: divide_simps abs_mult h_def) qed qed qed lemma frechet_derivative_equals_partial_derivative: fixes f::"'a::euclidean_space \ 'a" assumes Df: "\x. (f has_derivative Df x) (at x)" assumes f': "((\p. f (x + (p - x \ i) *\<^sub>R i) \ b) has_real_derivative f' x i b) (at (x \ i))" shows "Df x i \ b = f' x i b" proof - define Dfb where "Dfb x = Blinfun (Df x)" for x have Dfb_apply: "blinfun_apply (Dfb x) = Df x" for x unfolding Dfb_def apply (rule bounded_linear_Blinfun_apply) apply (rule has_derivative_bounded_linear) apply (rule assms) done have "Dfb x = blinfun_of_matrix (\i b. Dfb x b \ i)" for x using blinfun_of_matrix_works[of "Dfb x"] by auto have Dfb: "\x. (f has_derivative Dfb x) (at x)" by (auto simp: Dfb_apply Df) note [derivative_intros] = diff_chain_at[OF _ Dfb, unfolded o_def] have "((\p. f (x + (p - x \ i) *\<^sub>R i) \ b) has_real_derivative Dfb x i \ b) (at (x \ i))" by (auto intro!: derivative_eq_intros ext simp: has_field_derivative_def blinfun.bilinear_simps) from DERIV_unique[OF f' this] show ?thesis by (simp add: Dfb_apply) qed subsection \Integration\ lemmas content_real[simp] lemmas integrable_continuous[intro, simp] and integrable_continuous_real[intro, simp] lemma integral_eucl_le: fixes f g::"'a::euclidean_space \ 'b::ordered_euclidean_space" assumes "f integrable_on s" and "g integrable_on s" and "\x. x \ s \ f x \ g x" shows "integral s f \ integral s g" using assms by (auto intro!: integral_component_le simp: eucl_le[where 'a='b]) lemma integral_ivl_bound: fixes l u::"'a::ordered_euclidean_space" assumes "\x h'. h' \ {t0 .. h} \ x \ {t0 .. h} \ (h' - t0) *\<^sub>R f x \ {l .. u}" assumes "t0 \ h" assumes f_int: "f integrable_on {t0 .. h}" shows "integral {t0 .. h} f \ {l .. u}" proof - from assms(1)[of t0 t0] assms(2) have "0 \ {l .. u}" by auto have "integral {t0 .. h} f = integral {t0 .. h} (\t. if t \ {t0, h} then 0 else f t)" by (rule integral_spike[where S="{t0, h}"]) auto also { fix x assume 1: "x \ {t0 <..< h}" with assms have "(h - t0) *\<^sub>R f x \ {l .. u}" by auto then have "(if x \ {t0, h} then 0 else f x) \ {l /\<^sub>R (h - t0) .. u /\<^sub>R (h - t0)}" using \x \ _\ by (auto simp: inverse_eq_divide simp: eucl_le[where 'a='a] field_simps algebra_simps) } then have "\ \ {integral {t0..h} (\_. l /\<^sub>R (h - t0)) .. integral {t0..h} (\_. u /\<^sub>R (h - t0))}" unfolding atLeastAtMost_iff apply (safe intro!: integral_eucl_le) using \0 \ {l .. u}\ apply (auto intro!: assms intro: integrable_continuous_real integrable_spike[where S="{t0, h}", OF f_int] simp: eucl_le[where 'a='a] divide_simps split: if_split_asm) done also have "\ \ {l .. u}" using assms \0 \ {l .. u}\ by (auto simp: inverse_eq_divide) finally show ?thesis . qed lemma add_integral_ivl_bound: fixes l u::"'a::ordered_euclidean_space" assumes "\x h'. h' \ {t0 .. h} \ x \ {t0 .. h} \ (h' - t0) *\<^sub>R f x \ {l - x0 .. u - x0}" assumes "t0 \ h" assumes f_int: "f integrable_on {t0 .. h}" shows "x0 + integral {t0 .. h} f \ {l .. u}" using integral_ivl_bound[OF assms] by (auto simp: algebra_simps) subsection \conditionally complete lattice\ subsection \Lists\ lemma Ball_set_Cons[simp]: "(\a\set_Cons x y. P a) \ (\a\x. \b\y. P (a#b))" by (auto simp: set_Cons_def) lemma set_cons_eq_empty[iff]: "set_Cons a b = {} \ a = {} \ b = {}" by (auto simp: set_Cons_def) lemma listset_eq_empty_iff[iff]: "listset XS = {} \ {} \ set XS" by (induction XS) auto lemma sing_in_sings[simp]: "[x] \ (\x. [x]) ` xd \ x \ xd" by auto lemma those_eq_None_set_iff: "those xs = None \ None \ set xs" by (induction xs) (auto split: option.split) lemma those_eq_Some_lengthD: "those xs = Some ys \ length xs = length ys" by (induction xs arbitrary: ys) (auto split: option.splits) lemma those_eq_Some_map_Some_iff: "those xs = Some ys \ (xs = map Some ys)" (is "?l \ ?r") proof safe assume ?l then have "length xs = length ys" by (rule those_eq_Some_lengthD) then show ?r using \?l\ by (induction xs ys rule: list_induct2) (auto split: option.splits) next assume ?r then have "length xs = length ys" by simp then show "those (map Some ys) = Some ys" using \?r\ by (induction xs ys rule: list_induct2) (auto split: option.splits) qed subsection \Set(sum)\ subsection \Max\ subsection \Uniform Limit\ subsection \Bounded Linear Functions\ lift_definition comp3::\ \TODO: name?\ "('c::real_normed_vector \\<^sub>L 'd::real_normed_vector) \ ('b::real_normed_vector \\<^sub>L 'c) \\<^sub>L 'b \\<^sub>L 'd" is "\(cd::('c \\<^sub>L 'd)) (bc::'b \\<^sub>L 'c). (cd o\<^sub>L bc)" by (rule bounded_bilinear.bounded_linear_right[OF bounded_bilinear_blinfun_compose]) lemma blinfun_apply_comp3[simp]: "blinfun_apply (comp3 a) b = (a o\<^sub>L b)" by (simp add: comp3.rep_eq) lemma bounded_linear_comp3[bounded_linear]: "bounded_linear comp3" by transfer (rule bounded_bilinear_blinfun_compose) lift_definition comp12::\ \TODO: name?\ "('a::real_normed_vector \\<^sub>L 'c::real_normed_vector) \ ('b::real_normed_vector \\<^sub>L 'c) \ ('a \ 'b) \\<^sub>L 'c" is "\f g (a, b). f a + g b" by (auto intro!: bounded_linear_intros intro: bounded_linear_compose simp: split_beta') lemma blinfun_apply_comp12[simp]: "blinfun_apply (comp12 f g) b = f (fst b) + g (snd b)" by (simp add: comp12.rep_eq split_beta) subsection \Order Transitivity Attributes\ attribute_setup le = \Scan.succeed (Thm.rule_attribute [] (fn context => fn thm => thm RS @{thm order_trans}))\ "transitive version of inequality (useful for intro)" attribute_setup ge = \Scan.succeed (Thm.rule_attribute [] (fn context => fn thm => thm RS @{thm order_trans[rotated]}))\ "transitive version of inequality (useful for intro)" subsection \point reflection\ definition preflect::"'a::real_vector \ 'a \ 'a" where "preflect \ \t0 t. 2 *\<^sub>R t0 - t" lemma preflect_preflect[simp]: "preflect t0 (preflect t0 t) = t" by (simp add: preflect_def algebra_simps) lemma preflect_preflect_image[simp]: "preflect t0 ` preflect t0 ` S = S" by (simp add: image_image) lemma is_interval_preflect[simp]: "is_interval (preflect t0 ` S) \ is_interval S" by (auto simp: preflect_def[abs_def]) lemma iv_in_preflect_image[intro, simp]: "t0 \ T \ t0 \ preflect t0 ` T" by (auto intro!: image_eqI simp: preflect_def algebra_simps scaleR_2) lemma preflect_tendsto[tendsto_intros]: fixes l::"'a::real_normed_vector" shows "(g \ l) F \ (h \ m) F \ ((\x. preflect (g x) (h x)) \ preflect l m) F" by (auto intro!: tendsto_eq_intros simp: preflect_def) lemma continuous_preflect[continuous_intros]: fixes a::"'a::real_normed_vector" shows "continuous (at a within A) (preflect t0)" by (auto simp: continuous_within intro!: tendsto_intros) lemma fixes t0::"'a::ordered_real_vector" shows preflect_le[simp]: "t0 \ preflect t0 b \ b \ t0" and le_preflect[simp]: "preflect t0 b \ t0 \ t0 \ b" and antimono_preflect: "antimono (preflect t0)" and preflect_le_preflect[simp]: "preflect t0 a \ preflect t0 b \ b \ a" and preflect_eq_cancel[simp]: "preflect t0 a = preflect t0 b \ a = b" by (auto intro!: antimonoI simp: preflect_def scaleR_2) lemma preflect_eq_point_iff[simp]: "t0 = preflect t0 s \ t0 = s" "preflect t0 s = t0 \ t0 = s" by (auto simp: preflect_def algebra_simps scaleR_2) lemma preflect_minus_self[simp]: "preflect t0 s - t0 = t0 - s" by (simp add: preflect_def scaleR_2) end diff --git a/thys/Smooth_Manifolds/Bump_Function.thy b/thys/Smooth_Manifolds/Bump_Function.thy --- a/thys/Smooth_Manifolds/Bump_Function.thy +++ b/thys/Smooth_Manifolds/Bump_Function.thy @@ -1,521 +1,521 @@ section \Bump Functions\ theory Bump_Function imports Smooth "HOL-Analysis.Weierstrass_Theorems" begin subsection \Construction\ context begin qualified definition f :: "real \ real" where "f t = (if t > 0 then exp(-inverse t) else 0)" lemma f_nonpos[simp]: "x \ 0 \ f x = 0" by (auto simp: f_def) lemma exp_inv_limit_0_right: "((\(t::real). exp(-inverse t)) \ 0) (at_right 0)" apply (rule filterlim_compose[where g = exp]) apply (rule exp_at_bot) apply (rule filterlim_compose[where g = uminus]) apply (rule filterlim_uminus_at_bot_at_top) by (rule filterlim_inverse_at_top_right) lemma "\\<^sub>F t in at_right 0. ((\x. inverse (x ^ Suc k)) has_real_derivative - (inverse (t ^ Suc k) * ((1 + real k) * t ^ k) * inverse (t ^ Suc k))) (at t)" unfolding eventually_at_filter by (auto simp del: power_Suc intro!: derivative_eq_intros eventuallyI) lemma exp_inv_limit_0_right_gen': "((\(t::real). inverse (t ^ k) / exp(inverse t)) \ 0) (at_right 0)" proof (induct k) case 0 then show ?case using exp_inv_limit_0_right by (auto simp: exp_minus inverse_eq_divide) next case (Suc k) have df: "\\<^sub>F t in at_right 0. ((\x. inverse (x ^ Suc k)) has_real_derivative - (inverse (t ^ k) * ((1 + real k)) * (inverse t ^ 2))) (at t)" unfolding eventually_at_filter apply (auto simp del: power_Suc intro!: derivative_eq_intros eventuallyI) by (auto simp: power2_eq_square) have dg: "\\<^sub>F t in at_right 0. ((\x. exp (inverse x)) has_real_derivative - (exp (inverse t) * (inverse t ^ 2))) (at t)" unfolding eventually_at_filter by (auto simp del: power_Suc intro!: derivative_eq_intros eventuallyI simp: power2_eq_square) show ?case apply (rule lhopital_right_0_at_top [OF _ _ df dg]) apply (rule filterlim_compose[where g = exp]) apply (rule exp_at_top) apply (rule filterlim_inverse_at_top_right) subgoal by (auto simp: eventually_at_filter) subgoal apply (rule Lim_transform_eventually[where f = "\x. (1 + real k) * (inverse (x ^ k) / exp (inverse x))"]) using Suc.hyps tendsto_mult_right_zero apply blast by (auto simp: eventually_at_filter) done qed lemma exp_inv_limit_0_right_gen: "((\(t::real). exp(-inverse t) / t ^ k) \ 0) (at_right 0)" using exp_inv_limit_0_right_gen'[of k] by (metis (no_types, lifting) Groups.mult_ac(2) Lim_cong_within divide_inverse exp_minus) lemma f_limit_0_right: "(f \ 0) (at_right 0)" proof - have "\\<^sub>F t in at_right 0. (t::real) > 0" by (rule eventually_at_right_less) then have "\\<^sub>F t in at_right 0. exp(-inverse t) = f t" by (eventually_elim) (auto simp: f_def) moreover have "((\(t::real). exp(-inverse t)) \ 0) (at_right 0)" by (rule exp_inv_limit_0_right) ultimately show ?thesis by (blast intro: Lim_transform_eventually) qed lemma f_limit_0: "(f \ 0) (at 0)" using _ f_limit_0_right proof (rule filterlim_split_at_real) have "\\<^sub>F t in at_left 0. 0 = f t" by (auto simp: f_def eventually_at_filter) then show "(f \ 0) (at_left 0)" by (blast intro: Lim_transform_eventually) qed lemma f_tendsto: "(f \ f x) (at x)" proof - consider "x = 0" | "x < 0" | "x > 0" by arith then show ?thesis proof cases case 1 then show ?thesis by (auto simp: f_limit_0 f_def) next case 2 have "\\<^sub>F t in at x. t < 0" apply (rule order_tendstoD) by (rule tendsto_intros) fact then have "\\<^sub>F t in at x. 0 = f t" by (eventually_elim) (auto simp: f_def) then show ?thesis using \x < 0\ by (auto simp: f_def intro: Lim_transform_eventually) next case 3 have "\\<^sub>F t in at x. t > 0" apply (rule order_tendstoD) by (rule tendsto_intros) fact then have "\\<^sub>F t in at x. exp(-inverse t) = f t" by (eventually_elim) (auto simp: f_def) moreover have "(\t. exp (- inverse t)) \x\ f x" using \x > 0\ by (auto simp: f_def tendsto_intros ) ultimately show ?thesis by (blast intro: Lim_transform_eventually) qed qed lemma f_continuous: "continuous_on S f" using f_tendsto continuous_on continuous_on_subset subset_UNIV by metis lemma continuous_on_real_polynomial_function: "continuous_on S p" if "real_polynomial_function p" using that by induction (auto intro: continuous_intros linear_continuous_on) lemma f_nth_derivative_is_poly: "higher_differentiable_on {0<..} f k \ (\p. real_polynomial_function p \ (\t>0. nth_derivative k f t 1 = p t / (t ^ (2 * k)) * exp(-inverse t)))" proof (induction k) case 0 then show ?case apply (auto simp: higher_differentiable_on.simps f_continuous) by (auto simp: f_def) next case (Suc k) obtain p where fk: "higher_differentiable_on {0<..} f k" and p1: "real_polynomial_function p" and p2: "\t>0. nth_derivative k f t 1 = p t / t ^ (2 * k) * exp (- inverse t)" using Suc by auto obtain p' where p'1: "real_polynomial_function p'" and p'2: "\t. (p has_real_derivative (p' t)) (at t)" using has_real_derivative_polynomial_function[of p] p1 by auto define rp where "rp t = (t\<^sup>2 * p' t - 2 * real k * t * p t + p t)" for t have rp: "real_polynomial_function rp" unfolding rp_def by (auto intro!: real_polynomial_function.intros(2-) real_polynomial_function_diff p1 p'1 simp: power2_eq_square) moreover have fk': "(\x. nth_derivative k f x 1) differentiable at t" (is ?a) "frechet_derivative (\x. nth_derivative k f x 1) (at t) 1 = rp t * (exp (-inverse t) / t^(2*k+2))" (is ?b) if "0 < t" for t proof - from p'2 that have dp: "(p has_derivative ((*) (p' t))) (at t within {0<..})" by (auto simp: at_within_open[of _ "{0<..}"] has_field_derivative_def ac_simps) have "((\t. p t / t ^ (2 * k) * exp (- inverse t)) has_derivative (\v. v * rp t * (exp (-inverse t) / t^(2*k+2)))) (at t within {0<..})" using that apply (auto intro!: derivative_eq_intros dp ext) apply (simp add: divide_simps algebra_simps rp_def power2_eq_square) by (metis Suc_pred mult_is_0 neq0_conv power_Suc zero_neq_numeral) then have "((\x. nth_derivative k f x 1) has_derivative (\v. v * rp t * (exp (-inverse t) / t^(2*k+2)))) (at t within {0<..})" apply (rule has_derivative_transform_within[OF _ zero_less_one]) using that p2 by (auto simp: ) then have "((\x. nth_derivative k f x 1) has_derivative (\v. v * rp t * (exp (-inverse t) / t^(2*k+2)))) (at t)" using that by (auto simp: at_within_open[of _ "{0<..}"]) from frechet_derivative_at'[OF this] this show ?a ?b by (auto simp: differentiable_def) qed have hdS: "higher_differentiable_on {0<..} f (Suc k)" apply (subst higher_differentiable_on_real_Suc') apply (auto simp: fk fk' frechet_derivative_nth_derivative_commute[symmetric]) apply (subst continuous_on_cong[OF refl]) apply (rule fk') by (auto intro!: continuous_intros p'1 p1 rp intro: continuous_on_real_polynomial_function) moreover have "nth_derivative (Suc k) f t 1 = rp t / t ^ (2 * (Suc k)) * exp (- inverse t)" if "t > 0" for t proof - have "nth_derivative (Suc k) f t 1 = frechet_derivative (\x. nth_derivative k f x 1) (at t) 1" by (simp add: frechet_derivative_nth_derivative_commute) also have "\ = rp t / t^(2*k+2) * (exp (-inverse t))" using fk'[OF \t > 0\] by simp finally show ?thesis by simp qed ultimately show ?case by blast qed lemma f_has_derivative_at_neg: " x < 0 \ (f has_derivative (\x. 0)) (at x)" by (rule has_derivative_transform_within_open[where f="\x. 0" and s="{..<0}"]) (auto simp: f_def) lemma f_differentiable_at_neg: "x < 0 \ f differentiable at x" using f_has_derivative_at_neg by (auto simp: differentiable_def) lemma frechet_derivative_f_at_neg: "x \ {..<0} \ frechet_derivative f (at x) = (\x. 0)" by (rule frechet_derivative_at') (rule f_has_derivative_at_neg, simp) lemma f_nth_derivative_lt_0: "higher_differentiable_on {..<0} f k \ (\t<0. nth_derivative k f t 1 = 0)" proof (induction k) case 0 have rewr: "a \ {..<0} \ \0 < a" for a::real by simp show ?case by (auto simp: higher_differentiable_on.simps f_def rewr simp del: lessThan_iff cong: continuous_on_cong) next case (Suc k) have "t < 0 \ (\x. nth_derivative k f x 1) differentiable at t" for t by (rule differentiable_eqI[where g=0 and X="{..<0}"]) (auto simp: zero_fun_def frechet_derivative_const Suc.IH) then have "frechet_derivative (\x. nth_derivative k f x 1) (at t) 1 = 0" if "t < 0" for t using that Suc.IH by (subst frechet_derivative_transform_within_open[where X="{..<0}" and g =0]) (auto simp: frechet_derivative_zero_fun) with Suc show ?case by (auto simp: higher_differentiable_on.simps f_differentiable_at_neg frechet_derivative_f_at_neg zero_fun_def simp flip: frechet_derivative_nth_derivative_commute simp del: lessThan_iff intro!: higher_differentiable_on_const cong: higher_differentiable_on_cong) qed lemma netlimit_at_left: "netlimit (at_left x) = x" for x::real by (rule Lim_ident_at) simp lemma netlimit_at_right: "netlimit (at_right x) = x" for x::real by (rule Lim_ident_at) simp lemma has_derivative_split_at: "(g has_derivative g') (at x)" if "(g has_derivative g') (at_left x)" "(g has_derivative g') (at_right x)" for x::real using that unfolding has_derivative_def netlimit_at netlimit_at_right netlimit_at_left by (auto intro: filterlim_split_at) lemma has_derivative_at_left_at_right': "(g has_derivative g') (at x)" if "(g has_derivative g') (at x within {..x})" "(g has_derivative g') (at x within {x..})" for x::real apply (rule has_derivative_split_at) - subgoal by (rule has_derivative_within_subset) (fact, auto) - subgoal by (rule has_derivative_within_subset) (fact, auto) + subgoal by (rule has_derivative_subset) (fact, auto) + subgoal by (rule has_derivative_subset) (fact, auto) done lemma real_polynomial_function_tendsto: "(p \ p x) (at x within X)" if "real_polynomial_function p" using that by (induction p) (auto intro!: tendsto_eq_intros intro: bounded_linear.tendsto) lemma f_nth_derivative_cases: "higher_differentiable_on UNIV f k \ (\t\0. nth_derivative k f t 1 = 0) \ (\p. real_polynomial_function p \ (\t>0. nth_derivative k f t 1 = p t / (t ^ (2 * k)) * exp(-inverse t)))" proof (induction k) case 0 then show ?case apply (auto simp: higher_differentiable_on.simps f_continuous) by (auto simp: f_def) next case (Suc k) from Suc.IH obtain pk where IH: "higher_differentiable_on UNIV f k" "\t. t \ 0 \ nth_derivative k f t 1 = 0" "real_polynomial_function pk" "\t. t > 0 \ nth_derivative k f t 1 = pk t / t ^ (2 * k) * exp (- inverse t)" by auto from f_nth_derivative_lt_0[of "Suc k"] local.f_nth_derivative_is_poly[of "Suc k"] obtain p where neg: "higher_differentiable_on {..<0} f (Suc k)" and neg0: "(\t<0. nth_derivative (Suc k) f t 1 = 0)" and pos: "higher_differentiable_on {0<..} f (Suc k)" and p: "real_polynomial_function p" "\t. t > 0 \ nth_derivative (Suc k) f t 1 = p t / t ^ (2 * Suc k) * exp (- inverse t)" by auto moreover have at_within_eq_at_right: "(at 0 within {0..}) = at_right (0::real)" apply (auto simp: filter_eq_iff eventually_at_filter ) apply (simp add: eventually_mono) apply (simp add: eventually_mono) done have [simp]: "{0..} - {0} = {0::real<..}" by auto have [simp]: "(at (0::real) within {0..}) \ bot" by (auto simp: at_within_eq_bot_iff) have k_th_has_derivative_at_left: "((\x. nth_derivative k f x 1) has_derivative (\x. 0)) (at 0 within {..0})" apply (rule has_derivative_transform_within[OF _ zero_less_one]) prefer 2 apply force prefer 2 apply (simp add: IH) by (rule derivative_intros) have k_th_has_derivative_at_right: "((\x. nth_derivative k f x 1) has_derivative (\x. 0)) (at 0 within {0..})" apply (rule has_derivative_transform_within[where f="\x'. if x' = 0 then 0 else pk x' / x' ^ (2 * k) * exp (- inverse x')", OF _ zero_less_one]) subgoal unfolding has_derivative_def apply (auto simp: Lim_ident_at) apply (rule Lim_transform_eventually[where f="\x. (pk x * (exp (- inverse x) / x ^ (2 * k + 1)))"]) apply (rule tendsto_eq_intros) apply (rule real_polynomial_function_tendsto[THEN tendsto_eq_rhs]) apply fact apply (rule refl) apply (subst at_within_eq_at_right) apply (rule exp_inv_limit_0_right_gen) apply (auto simp add: eventually_at_filter divide_simps) done subgoal by force subgoal by (auto simp: IH(2) IH(4)) done have k_th_has_derivative: "((\x. nth_derivative k f x 1) has_derivative (\x. 0)) (at 0)" apply (rule has_derivative_at_left_at_right') apply (rule k_th_has_derivative_at_left) apply (rule k_th_has_derivative_at_right) done have nth_Suc_zero: "nth_derivative (Suc k) f 0 1 = 0" apply (auto simp: frechet_derivative_nth_derivative_commute[symmetric]) apply (subst frechet_derivative_at') apply (rule k_th_has_derivative) by simp moreover have "higher_differentiable_on UNIV f (Suc k)" proof - have "continuous_on UNIV (\x. nth_derivative (Suc k) f x 1)" unfolding continuous_on_eq_continuous_within proof fix x::real consider "x < 0" | "x > 0" | "x = 0" by arith then show "isCont (\x. nth_derivative (Suc k) f x 1) x" proof cases case 1 then have at_eq: "at x = at x within {..<0}" using at_within_open[of x "{..<0}"] by auto show ?thesis unfolding at_eq apply (rule continuous_transform_within[OF _ zero_less_one]) using neg0 1 by (auto simp: at_eq) next case 2 then have at_eq: "at x = at x within {0<..}" using at_within_open[of x "{0<..}"] by auto show ?thesis unfolding at_eq apply (rule continuous_transform_within[OF _ zero_less_one]) using p 2 by (auto intro!: continuous_intros intro: continuous_real_polymonial_function continuous_at_imp_continuous_within) next case 3 have "((\x. nth_derivative (Suc k) f x 1) \ 0) (at_left 0)" proof - have "\\<^sub>F x in at_left 0. 0 = nth_derivative (Suc k) f x 1" using neg0 by (auto simp: eventually_at_filter) then show ?thesis by (blast intro: Lim_transform_eventually) qed moreover have "((\x. nth_derivative (Suc k) f x 1) \ 0) (at_right 0)" proof - have "((\x. p x * (exp (- inverse x) / x ^ (2 * Suc k))) \ 0) (at_right 0)" apply (rule tendsto_eq_intros) apply (rule real_polynomial_function_tendsto) apply fact apply (rule exp_inv_limit_0_right_gen) by simp moreover have "\\<^sub>F x in at_right 0. p x * (exp (- inverse x) / x ^ (2 * Suc k)) = nth_derivative (Suc k) f x 1" using p by (auto simp: eventually_at_filter) ultimately show ?thesis by (rule Lim_transform_eventually) qed ultimately show ?thesis by (auto simp: continuous_def nth_Suc_zero 3 filterlim_split_at simp del: nth_derivative.simps) qed qed moreover have "(\x. nth_derivative k f x 1) differentiable at x" for x proof - consider "x = 0" | "x < 0" | "x > 0"by arith then show ?thesis proof cases case 1 then show ?thesis using k_th_has_derivative by (auto simp: differentiable_def) next case 2 with neg show ?thesis by (subst (asm) higher_differentiable_on_real_Suc') (auto simp: ) next case 3 with pos show ?thesis by (subst (asm) higher_differentiable_on_real_Suc') (auto simp: ) qed qed moreover have "higher_differentiable_on UNIV f k" by fact ultimately show ?thesis by (subst higher_differentiable_on_real_Suc'[OF open_UNIV]) auto qed ultimately show ?case by (auto simp: less_eq_real_def) qed lemma f_smooth_on: "k-smooth_on S f" and f_higher_differentiable_on: "higher_differentiable_on S f n" using f_nth_derivative_cases by (auto simp: smooth_on_def higher_differentiable_on_subset[OF _ subset_UNIV]) lemma f_compose_smooth_on: "k-smooth_on S (\x. f (g x))" if "k-smooth_on S g" "open S" using smooth_on_compose[OF f_smooth_on that open_UNIV subset_UNIV] by (auto simp: o_def) lemma f_nonneg: "f x \ 0" by (auto simp: f_def) lemma f_pos_iff: "f x > 0 \ x > 0" by (auto simp: f_def) lemma f_eq_zero_iff: "f x = 0 \ x \ 0" by (auto simp: f_def) subsection \Cutoff function\ definition "h t = f (2 - t) / (f (2 - t) + f (t - 1))" lemma denominator_pos: "f (2 - t) + f (t - 1) > 0" by (auto simp: f_def add_pos_pos) lemma denominator_nonzero: "f (2 - t) + f (t - 1) = 0 \ False" using denominator_pos[of t] by auto lemma h_range: "0 \ h t" "h t \ 1" by (auto simp: h_def f_nonneg denominator_pos) lemma h_pos: "t < 2 \ 0 < h t" and h_less_one: "1 < t \ h t < 1" by (auto simp: h_def f_pos_iff denominator_pos) lemma h_eq_0: "h t = 0" if "t \ 2" using that by (auto simp: h_def) lemma h_eq_1: "h t = 1" if "t \ 1" using that by (auto simp: h_def f_eq_zero_iff) lemma h_compose_smooth_on: "k-smooth_on S (\x. h (g x))" if "k-smooth_on S g" "open S" by (auto simp: h_def[abs_def] denominator_nonzero intro!: smooth_on_divide f_compose_smooth_on smooth_on_minus smooth_on_add that) subsection \Bump function\ definition H::"_::real_inner \ real" where "H x = h (norm x)" lemma H_range: "0 \ H x" "H x \ 1" by (auto simp: H_def h_range) lemma H_eq_one: "H x = 1" if "x \ cball 0 1" using that by (auto simp: H_def h_eq_1) lemma H_pos: "H x > 0" if "x \ ball 0 2" using that by (auto simp: H_def h_pos) lemma H_eq_zero: "H x = 0" if "x \ ball 0 2" using that by (auto simp: H_def h_eq_0) lemma H_neq_zeroD: "H x \ 0 \ x \ ball 0 2" using H_eq_zero by blast lemma H_smooth_on: "k-smooth_on UNIV H" proof - have 1: "k-smooth_on (ball 0 1) H" by (rule smooth_on_cong[where g="\x. 1"]) (auto simp: H_eq_one) have 2: "k-smooth_on (UNIV - cball 0 (1/2)) H" by (auto simp: H_def[abs_def] intro!: h_compose_smooth_on smooth_on_norm) have O: "open (ball 0 1)" "open (UNIV - cball 0 (1 / 2))" by auto have *: "ball 0 1 \ (UNIV - cball 0 (1 / 2)) = UNIV" by (auto simp: mem_ball) from smooth_on_open_Un[OF 1 2 O, unfolded *] show ?thesis by (rule smooth_on_subset) auto qed lemma H_compose_smooth_on: "k-smooth_on S (\x. H (g x))" if "k-smooth_on S g" "open S" for g :: "_ \ _::euclidean_space" using smooth_on_compose[OF H_smooth_on that] by (auto simp: o_def) end end \ No newline at end of file diff --git a/thys/Smooth_Manifolds/Smooth.thy b/thys/Smooth_Manifolds/Smooth.thy --- a/thys/Smooth_Manifolds/Smooth.thy +++ b/thys/Smooth_Manifolds/Smooth.thy @@ -1,1340 +1,1340 @@ section \Smooth Functions between Normed Vector Spaces\ theory Smooth imports Analysis_More begin subsection \From/To \Multivariate_Taylor.thy\\ lemma multivariate_Taylor_integral: fixes f::"'a::real_normed_vector \ 'b::banach" and Df::"'a \ nat \ 'a \ 'b" assumes "n > 0" assumes Df_Nil: "\a x. Df a 0 H = f a" assumes Df_Cons: "\a i d. a \ closed_segment X (X + H) \ i < n \ ((\a. Df a i H) has_derivative (Df a (Suc i))) (at a within G)" assumes cs: "closed_segment X (X + H) \ G" defines "i \ \x. ((1 - x) ^ (n - 1) / fact (n - 1)) *\<^sub>R Df (X + x *\<^sub>R H) n H" shows multivariate_Taylor_has_integral: "(i has_integral f (X + H) - (\iR Df X i H)) {0..1}" and multivariate_Taylor: "f (X + H) = (\iR Df X i H) + integral {0..1} i" and multivariate_Taylor_integrable: "i integrable_on {0..1}" proof goal_cases case 1 let ?G = "closed_segment X (X + H)" define line where "line t = X + t *\<^sub>R H" for t have segment_eq: "closed_segment X (X + H) = line ` {0 .. 1}" by (auto simp: line_def closed_segment_def algebra_simps) have line_deriv: "\x. (line has_derivative (\t. t *\<^sub>R H)) (at x)" by (auto intro!: derivative_eq_intros simp: line_def [abs_def]) define g where "g = f o line" define Dg where "Dg n t = Df (line t) n H" for n :: nat and t :: real note \n > 0\ moreover have Dg0: "Dg 0 = g" by (auto simp add: Dg_def Df_Nil g_def) moreover have DgSuc: "(Dg m has_vector_derivative Dg (Suc m) t) (at t within {0..1})" if "m < n" "0 \ t" "t \ 1" for m::nat and t::real proof - from that have [intro]: "line t \ ?G" using assms by (auto simp: segment_eq) - note [derivative_intros] = has_derivative_in_compose[OF _ has_derivative_within_subset[OF Df_Cons]] + note [derivative_intros] = has_derivative_in_compose[OF _ has_derivative_subset[OF Df_Cons]] interpret Df: linear "(\d. Df (line t) (Suc m) d)" by (auto intro!: has_derivative_linear derivative_intros \m < n\) note [derivative_intros] = has_derivative_compose[OF _ line_deriv] show ?thesis using Df.scaleR \m < n\ by (auto simp: Dg_def [abs_def] has_vector_derivative_def g_def segment_eq intro!: derivative_eq_intros subsetD[OF cs]) qed ultimately have g_Taylor: "(i has_integral g 1 - (\iR Dg i 0)) {0 .. 1}" unfolding i_def Dg_def [abs_def] line_def by (rule Taylor_has_integral) auto then show c: ?case using \n > 0\ by (auto simp: g_def line_def Dg_def) case 2 show ?case using c by (simp add: integral_unique add.commute) case 3 show ?case using c by force qed subsection \Higher-order differentiable\ fun higher_differentiable_on :: "'a::real_normed_vector set \ ('a \ 'b::real_normed_vector) \ nat \ bool" where "higher_differentiable_on S f 0 \ continuous_on S f" | "higher_differentiable_on S f (Suc n) \ (\x\S. f differentiable (at x)) \ (\v. higher_differentiable_on S (\x. frechet_derivative f (at x) v) n)" lemma ball_differentiable_atD: "\x\S. f differentiable at x \ f differentiable_on S" by (auto simp: differentiable_on_def differentiable_at_withinI) lemma higher_differentiable_on_imp_continuous_on: "continuous_on S f" if "higher_differentiable_on S f n" using that by (cases n) (auto simp: differentiable_imp_continuous_on ball_differentiable_atD) lemma higher_differentiable_on_imp_differentiable_on: "f differentiable_on S" if "higher_differentiable_on S f k" "k > 0" using that by (cases k) (auto simp: ball_differentiable_atD) lemma higher_differentiable_on_congI: assumes "open S" "higher_differentiable_on S g n" and "\x. x \ S \ f x = g x" shows "higher_differentiable_on S f n" using assms(2,3) proof (induction n arbitrary: f g) case 0 then show ?case by auto next case (Suc n) have 1: "\x\S. g differentiable (at x)" and 2: "higher_differentiable_on S (\x. frechet_derivative g (at x) v) n" for v using Suc by auto have 3: "\x\S. f differentiable (at x)" using 1 Suc(3) assms(1) by (metis differentiable_eqI) have 4: "frechet_derivative f (at x) v = frechet_derivative g (at x) v" if "x \ S" for x v using "3" Suc.prems(2) assms(1) frechet_derivative_transform_within_open_ext that by blast from 2 3 4 show ?case using Suc.IH[OF 2 4] by auto qed lemma higher_differentiable_on_cong: assumes "open S" "S = T" and "\x. x \ T \ f x = g x" shows "higher_differentiable_on S f n \ higher_differentiable_on T g n" using higher_differentiable_on_congI assms by auto lemma higher_differentiable_on_SucD: "higher_differentiable_on S f n" if "higher_differentiable_on S f (Suc n)" using that by (induction n arbitrary: f) (auto simp: differentiable_imp_continuous_on ball_differentiable_atD) lemma higher_differentiable_on_addD: "higher_differentiable_on S f n" if "higher_differentiable_on S f (n + m)" using that by (induction m arbitrary: f n) (auto simp del: higher_differentiable_on.simps dest!: higher_differentiable_on_SucD) lemma higher_differentiable_on_le: "higher_differentiable_on S f n" if "higher_differentiable_on S f m" "n \ m" using higher_differentiable_on_addD[of S f n "m - n"] that by auto lemma higher_differentiable_on_open_subsetsI: "higher_differentiable_on S f n" if "\x. x \ S \ \T. x \ T \ open T \ higher_differentiable_on T f n" using that proof (induction n arbitrary: f) case 0 show ?case by (force simp: continuous_on_def dest!: 0 dest: at_within_open intro!: Lim_at_imp_Lim_at_within[where S=S]) next case (Suc n) have "f differentiable at x" if "x \ S" for x using Suc.prems[OF \x \ S\] by (auto simp: differentiable_on_def dest: at_within_open dest!: bspec) then have "f differentiable_on S" by (auto simp: differentiable_on_def intro!: differentiable_at_withinI[where s=S]) with Suc show ?case by fastforce qed lemma higher_differentiable_on_const: "higher_differentiable_on S (\x. c) n" by (induction n arbitrary: c) (auto simp: continuous_intros frechet_derivative_const) lemma higher_differentiable_on_id: "higher_differentiable_on S (\x. x) n" by (cases n) (auto simp: frechet_derivative_works higher_differentiable_on_const) lemma higher_differentiable_on_add: "higher_differentiable_on S (\x. f x + g x) n" if "higher_differentiable_on S f n" "higher_differentiable_on S g n" "open S" using that proof (induction n arbitrary: f g) case 0 then show ?case by (auto intro!: continuous_intros) next case (Suc n) from Suc.prems have f: "\x. x\S \ f differentiable (at x)" and hf: "higher_differentiable_on S (\x. frechet_derivative f (at x) v) n" and g: "\x. x\S \ g differentiable (at x)" and hg: "higher_differentiable_on S (\x. frechet_derivative g (at x) v) n" for v by auto show ?case using f g \open S\ by (auto simp: frechet_derivative_plus intro!: derivative_intros f g Suc.IH hf hg cong: higher_differentiable_on_cong) qed lemma (in bounded_bilinear) differentiable: "(\x. prod (f x) (g x)) differentiable at x within S" if "f differentiable at x within S" "g differentiable at x within S" by (blast intro: differentiableI frechet_derivative_worksI that FDERIV) context begin private lemmas d = bounded_bilinear.differentiable lemmas differentiable_inner = bounded_bilinear_inner[THEN d] and differentiable_scaleR = bounded_bilinear_scaleR[THEN d] and differentiable_mult = bounded_bilinear_mult[THEN d] end lemma (in bounded_bilinear) differentiable_on: "(\x. prod (f x) (g x)) differentiable_on S" if "f differentiable_on S" "g differentiable_on S" using that by (auto simp: differentiable_on_def differentiable) context begin private lemmas do = bounded_bilinear.differentiable_on lemmas differentiable_on_inner = bounded_bilinear_inner[THEN do] and differentiable_on_scaleR = bounded_bilinear_scaleR[THEN do] and differentiable_on_mult = bounded_bilinear_mult[THEN do] end lemma (in bounded_bilinear) higher_differentiable_on: "higher_differentiable_on S (\x. prod (f x) (g x)) n" if "higher_differentiable_on S f n" "higher_differentiable_on S g n" "open S" using that proof (induction n arbitrary: f g) case 0 then show ?case by (auto intro!: continuous_intros continuous_on) next case (Suc n) from Suc.prems have f: "\x. x\S \ f differentiable (at x)" and hf: "higher_differentiable_on S (\x. frechet_derivative f (at x) v) n" and g: "\x. x\S \ g differentiable (at x)" and hg: "higher_differentiable_on S (\x. frechet_derivative g (at x) v) n" for v by auto show ?case using f g \open S\ Suc by (auto simp: frechet_derivative intro!: derivative_intros f g differentiable higher_differentiable_on_add Suc.IH intro: higher_differentiable_on_SucD cong: higher_differentiable_on_cong) qed context begin private lemmas hdo = bounded_bilinear.higher_differentiable_on lemmas higher_differentiable_on_inner = bounded_bilinear_inner[THEN hdo] and higher_differentiable_on_scaleR = bounded_bilinear_scaleR[THEN hdo] and higher_differentiable_on_mult = bounded_bilinear_mult[THEN hdo] end lemma higher_differentiable_on_sum: "higher_differentiable_on S (\x. \i\F. f i x) n" if "\i. i \ F \ finite F \ higher_differentiable_on S (f i) n" "open S" using that by (induction F rule: infinite_finite_induct) (auto intro!: higher_differentiable_on_const higher_differentiable_on_add) lemma higher_differentiable_on_subset: "higher_differentiable_on S f n" if "higher_differentiable_on T f n" "S \ T" using that by (induction n arbitrary: f) (auto intro: continuous_on_subset differentiable_on_subset) lemma higher_differentiable_on_compose: "higher_differentiable_on S (f o g) n" if "higher_differentiable_on T f n" "higher_differentiable_on S g n" "g ` S \ T" "open S" "open T" for g::"_\_::euclidean_space"\\TODO: can we get around this restriction?\ using that(1-3) proof (induction n arbitrary: f g) case 0 then show ?case using that(4-) by (auto simp: continuous_on_compose2) next case (Suc n) from Suc.prems have f: "\x. x \ T \ f differentiable (at x)" and g: "\x. x \ S \ g differentiable (at x)" and hf: "higher_differentiable_on T (\x. frechet_derivative f (at x) v) n" and hg: "higher_differentiable_on S (\x. frechet_derivative g (at x) w) n" for v w by auto show ?case using \g ` _ \ _\ \open S\ f g \open T\ Suc Suc.IH[where f="\x. frechet_derivative f (at x) v" and g = "\x. g x" for v, unfolded o_def] higher_differentiable_on_SucD[OF Suc.prems(2)] by (auto simp: frechet_derivative_compose_eucl subset_iff simp del: o_apply intro!: differentiable_chain_within higher_differentiable_on_sum higher_differentiable_on_scaleR higher_differentiable_on_inner higher_differentiable_on_const intro: differentiable_at_withinI cong: higher_differentiable_on_cong) qed lemma higher_differentiable_on_uminus: "higher_differentiable_on S (\x. - f x) n" if "higher_differentiable_on S f n" "open S" using higher_differentiable_on_scaleR[of S "\x. -1" n f] that by (auto simp: higher_differentiable_on_const) lemma higher_differentiable_on_minus: "higher_differentiable_on S (\x. f x - g x) n" if "higher_differentiable_on S f n" "higher_differentiable_on S g n" "open S" using higher_differentiable_on_add[OF _ higher_differentiable_on_uminus, OF that(1,2,3,3)] by simp lemma higher_differentiable_on_inverse: "higher_differentiable_on S (\x. inverse (f x)) n" if "higher_differentiable_on S f n" "0 \ f ` S" "open S" for f::"_\_::real_normed_field" using that proof (induction n arbitrary: f) case 0 then show ?case by (auto simp: continuous_on_inverse) next case (Suc n) from Suc.prems(1) have fn: "higher_differentiable_on S f n" by (rule higher_differentiable_on_SucD) from Suc show ?case by (auto simp: continuous_on_inverse image_iff power2_eq_square frechet_derivative_inverse divide_inverse intro!: differentiable_inverse higher_differentiable_on_uminus higher_differentiable_on_mult Suc.IH fn cong: higher_differentiable_on_cong) qed lemma higher_differentiable_on_divide: "higher_differentiable_on S (\x. f x / g x) n" if "higher_differentiable_on S f n" "higher_differentiable_on S g n" "\x. x \ S \ g x \ 0" "open S" for f::"_\_::real_normed_field" using higher_differentiable_on_mult[OF _ higher_differentiable_on_inverse, OF that(1,2) _ that(4,4)] that(3) by (auto simp: divide_inverse image_iff) lemma differentiable_on_open_Union: "f differentiable_on \S" if "\s. s \ S \ f differentiable_on s" "\s. s \ S \ open s" using that unfolding differentiable_on_def by (metis Union_iff at_within_open open_Union) lemma higher_differentiable_on_open_Union: "higher_differentiable_on (\S) f n" if "\s. s \ S \ higher_differentiable_on s f n" "\s. s \ S \ open s" using that proof (induction n arbitrary: f) case 0 then show ?case by (auto simp: continuous_on_open_Union) next case (Suc n) then show ?case by (auto simp: differentiable_on_open_Union) qed lemma differentiable_on_open_Un: "f differentiable_on S \ T" if "f differentiable_on S" "f differentiable_on T" "open S" "open T" using that differentiable_on_open_Union[of "{S, T}" f] by auto lemma higher_differentiable_on_open_Un: "higher_differentiable_on (S \ T) f n" if "higher_differentiable_on S f n" "higher_differentiable_on T f n" "open S" "open T" using higher_differentiable_on_open_Union[of "{S, T}" f n] that by auto lemma higher_differentiable_on_sqrt: "higher_differentiable_on S (\x. sqrt (f x)) n" if "higher_differentiable_on S f n" "0 \ f ` S" "open S" using that proof (induction n arbitrary: f) case 0 then show ?case by (auto simp: continuous_intros) next case (Suc n) from Suc.prems(1) have fn: "higher_differentiable_on S f n" by (rule higher_differentiable_on_SucD) then have "continuous_on S f" by (rule higher_differentiable_on_imp_continuous_on) with \open S\ have op: "open (S \ f -` {0<..})" (is "open ?op") by (rule open_continuous_vimage') simp from \open S\ \continuous_on S f\ have on: "open (S \ f -` {..<0})" (is "open ?on") by (rule open_continuous_vimage') simp have op': "higher_differentiable_on ?op (\x. 1) n" and on': "higher_differentiable_on ?on (\x. -1) n" by (rule higher_differentiable_on_const)+ then have i: "higher_differentiable_on (?op \ ?on) (\x. if 0 < f x then 1::real else - 1) n" by (auto intro!: higher_differentiable_on_open_Un op on higher_differentiable_on_congI[OF _ op'] higher_differentiable_on_congI[OF _ on']) also have "?op \ ?on = S" using Suc by auto finally have i: "higher_differentiable_on S (\x. if 0 < f x then 1::real else - 1) n" . from fn i Suc show ?case by (auto simp: sqrt_differentiable_on image_iff frechet_derivative_sqrt intro!: sqrt_differentiable higher_differentiable_on_mult higher_differentiable_on_inverse higher_differentiable_on_divide higher_differentiable_on_const cong: higher_differentiable_on_cong) qed lemma higher_differentiable_on_frechet_derivativeI: "higher_differentiable_on X (\x. frechet_derivative f (at x) h) i" if "higher_differentiable_on X f (Suc i)" "open X" "x \ X" using that(1) by (induction i arbitrary: f h) auto lemma higher_differentiable_on_norm: "higher_differentiable_on S (\x. norm (f x)) n" if "higher_differentiable_on S f n" "0 \ f ` S" "open S" for f::"_\_::real_inner" using that proof (induction n arbitrary: f) case 0 then show ?case by (auto simp: continuous_on_norm) next case (Suc n) from Suc.prems(1) have fn: "higher_differentiable_on S f n" by (rule higher_differentiable_on_SucD) from Suc show ?case by (auto simp: continuous_on_norm frechet_derivative_norm image_iff sgn_div_norm cong: higher_differentiable_on_cong intro!: differentiable_norm_compose_at higher_differentiable_on_inner higher_differentiable_on_inverse higher_differentiable_on_mult Suc.IH fn) qed declare higher_differentiable_on.simps [simp del] lemma higher_differentiable_on_Pair: "higher_differentiable_on S f k \ higher_differentiable_on S g k \ higher_differentiable_on S (\x. (f x, g x)) k" if "open S" proof (induction k arbitrary: f g) case 0 then show ?case unfolding higher_differentiable_on.simps by (auto intro!: continuous_intros) next case (Suc k) then show ?case using that unfolding higher_differentiable_on.simps by (auto simp: frechet_derivative_pair[of f _ g] cong: higher_differentiable_on_cong) qed lemma higher_differentiable_on_compose': "higher_differentiable_on S (\x. f (g x)) n" if "higher_differentiable_on T f n" "higher_differentiable_on S g n" "g ` S \ T" "open S" "open T" for g::"_\_::euclidean_space" using higher_differentiable_on_compose[of T f n S g] comp_def that by (metis (no_types, lifting) higher_differentiable_on_cong) lemma higher_differentiable_on_fst: "higher_differentiable_on (S \ T) fst k" proof (induction k) case (Suc k) then show ?case unfolding higher_differentiable_on.simps by (auto simp: differentiable_at_fst frechet_derivative_fst frechet_derivative_id higher_differentiable_on_const) qed (auto simp: higher_differentiable_on.simps continuous_on_fst) lemma higher_differentiable_on_snd: "higher_differentiable_on (S \ T) snd k" proof (induction k) case (Suc k) then show ?case unfolding higher_differentiable_on.simps by (auto intro!: continuous_intros simp: differentiable_at_snd frechet_derivative_snd frechet_derivative_id higher_differentiable_on_const) qed (auto simp: higher_differentiable_on.simps continuous_on_snd) lemma higher_differentiable_on_fst_comp: "higher_differentiable_on S (\x. fst (f x)) k" if "higher_differentiable_on S f k" "open S" using that by (induction k arbitrary: f) (auto intro!: continuous_intros differentiable_at_fst cong: higher_differentiable_on_cong simp: higher_differentiable_on.simps frechet_derivative_fst) lemma higher_differentiable_on_snd_comp: "higher_differentiable_on S (\x. snd (f x)) k" if "higher_differentiable_on S f k" "open S" using that by (induction k arbitrary: f) (auto intro!: continuous_intros differentiable_at_snd cong: higher_differentiable_on_cong simp: higher_differentiable_on.simps frechet_derivative_snd) lemma higher_differentiable_on_Pair': "higher_differentiable_on S f k \ higher_differentiable_on T g k \ higher_differentiable_on (S \ T) (\x. (f (fst x), g (snd x))) k" if S: "open S" and T: "open T" for f::"_::euclidean_space\_" and g::"_::euclidean_space\_" by (auto intro!: higher_differentiable_on_Pair open_Times S T higher_differentiable_on_fst higher_differentiable_on_snd higher_differentiable_on_compose'[where f=f and T=S] higher_differentiable_on_compose'[where f=g and T=T]) lemma higher_differentiable_on_sin: "higher_differentiable_on S (\x. sin (f x::real)) n" and higher_differentiable_on_cos: "higher_differentiable_on S (\x. cos (f x::real)) n" if f: "higher_differentiable_on S f n" and S: "open S" unfolding atomize_conj using f proof (induction n) case (Suc n) then have "higher_differentiable_on S f n" "higher_differentiable_on S (\x. sin (f x)) n" "higher_differentiable_on S (\x. cos (f x)) n" "\x. x \ S \ f differentiable at x" using higher_differentiable_on_SucD by (auto simp: higher_differentiable_on.simps) with Suc show ?case by (auto simp: higher_differentiable_on.simps sin_differentiable_at cos_differentiable_at frechet_derivative_sin frechet_derivative_cos S intro!: higher_differentiable_on_mult higher_differentiable_on_uminus cong: higher_differentiable_on_cong[OF S]) qed (auto simp: higher_differentiable_on.simps intro!: continuous_intros) subsection \Higher directional derivatives\ primrec nth_derivative :: "nat \ ('a::real_normed_vector \ 'b::real_normed_vector) \ 'a \ 'a \ 'b" where "nth_derivative 0 f x h = f x" | "nth_derivative (Suc i) f x h = nth_derivative i (\x. frechet_derivative f (at x) h) x h" lemma frechet_derivative_nth_derivative_commute: "frechet_derivative (\x. nth_derivative i f x h) (at x) h = nth_derivative i (\x. frechet_derivative f (at x) h) x h" by (induction i arbitrary: f) auto lemma nth_derivative_funpow: "nth_derivative i f x h = ((\f x. frechet_derivative f (at x) h) ^^ i) f x" by (induction i arbitrary: f) (auto simp del: funpow.simps simp: funpow_Suc_right) lemma nth_derivative_exists: "\f'. ((\x. nth_derivative i f x h) has_derivative f') (at x) \ f' h = nth_derivative (Suc i) f x h" if "higher_differentiable_on X f (Suc i)" "open X" "x \ X" using that(1) proof (induction i arbitrary: f) case 0 with that show ?case by (auto simp: higher_differentiable_on.simps that dest!: frechet_derivative_worksI) next case (Suc i) from Suc.prems have "\x. x \ X \ f differentiable at x" "higher_differentiable_on X (\x. frechet_derivative f (at x) h) (Suc i)" unfolding higher_differentiable_on.simps(2)[where n = "Suc i"] by auto from Suc.IH[OF this(2)] show ?case by auto qed lemma higher_derivatives_exists: assumes "higher_differentiable_on X f n" "open X" obtains Df where "\a h. Df a 0 h = f a" "\a h i. i < n \ a \ X \ ((\a. Df a i H) has_derivative Df a (Suc i)) (at a)" "\a i. i \ n \ a \ X \ Df a i H = nth_derivative i f a H" proof - have "higher_differentiable_on X f (Suc i)" if "i < n" for i apply (rule higher_differentiable_on_le[OF assms(1)]) using that by simp from nth_derivative_exists[OF this assms(2)] have "\i\{..x \ X. \f'. ((\x. nth_derivative i f x H) has_derivative f') (at x) \ f' H = nth_derivative (Suc i) f x H" by blast from this[unfolded bchoice_iff] obtain f' where f': "i < n \ x \ X \ ((\x. nth_derivative i f x H) has_derivative f' x i) (at x)" "i < n \ x \ X \ f' x i H = nth_derivative (Suc i) f x H" for x i by force define Df where "Df a i h = (if i = 0 then f a else f' a (i - 1) h)" for a i h have "Df a 0 h = f a" for a h by (auto simp: Df_def) moreover have "i < n \ a \ X \ ((\a. Df a i H) has_derivative Df a (Suc i)) (at a)" for i a apply (auto simp: Df_def[abs_def]) using _ \open X\ apply (rule has_derivative_transform_within_open) apply (rule f') apply (auto simp: f') done moreover have "i \ n \ a \ X \ Df a i H = nth_derivative i f a H" for i a by (auto simp: Df_def f') ultimately show ?thesis .. qed lemma nth_derivative_differentiable: assumes "higher_differentiable_on S f (Suc n)" "x \ S" shows "(\x. nth_derivative n f x v) differentiable at x" using assms by (induction n arbitrary: f) (auto simp: higher_differentiable_on.simps) lemma higher_differentiable_on_imp_continuous_nth_derivative: assumes "higher_differentiable_on S f n" shows "continuous_on S (\x. nth_derivative n f x v)" using assms by (induction n arbitrary: f) (auto simp: higher_differentiable_on.simps) lemma frechet_derivative_at_real_eq_scaleR: "frechet_derivative f (at x) v = v *\<^sub>R frechet_derivative f (at x) 1" if "f differentiable (at x)" "NO_MATCH 1 v" by (simp add: frechet_derivative_eq_vector_derivative that) lemma higher_differentiable_on_real_Suc: "higher_differentiable_on S f (Suc n) \ (\x\S. f differentiable (at x)) \ (higher_differentiable_on S (\x. frechet_derivative f (at x) 1) n)" if "open S" for S::"real set" using \open S\ by (auto simp: higher_differentiable_on.simps frechet_derivative_at_real_eq_scaleR intro!: higher_differentiable_on_scaleR higher_differentiable_on_const cong: higher_differentiable_on_cong) lemma higher_differentiable_on_real_SucI: fixes S::"real set" assumes "\x. x \ S \ (\x. nth_derivative n f x 1) differentiable at x" "continuous_on S (\x. nth_derivative (Suc n) f x 1)" "higher_differentiable_on S f n" and o: "open S" shows "higher_differentiable_on S f (Suc n)" using assms proof (induction n arbitrary: f) case 0 then show ?case by (auto simp: higher_differentiable_on_real_Suc higher_differentiable_on.simps(1) o) qed (fastforce simp: higher_differentiable_on_real_Suc) lemma higher_differentiable_on_real_Suc': "open S \ higher_differentiable_on S f (Suc n) \ (\v. continuous_on S (\x. nth_derivative (Suc n) f x 1)) \ (\x\S. \v. (\x. nth_derivative n f x 1) differentiable (at x)) \ higher_differentiable_on S f n" for S::"real set" apply (auto simp: nth_derivative_differentiable dest: higher_differentiable_on_SucD intro: higher_differentiable_on_real_SucI) by (auto simp: higher_differentiable_on_real_Suc higher_differentiable_on.simps(1) higher_differentiable_on_imp_continuous_nth_derivative) lemma closed_segment_subsetD: "0 \ x \ x \ 1 \ (X + x *\<^sub>R H) \ S" if "closed_segment X (X + H) \ S" using that by (rule subsetD) (auto simp: closed_segment_def algebra_simps intro!: exI[where x=x]) lemma higher_differentiable_Taylor: fixes f::"'a::real_normed_vector \ 'b::banach" and H::"'a" and Df::"'a \ nat \ 'a \ 'a \ 'b" assumes "n > 0" assumes hd: "higher_differentiable_on S f n" "open S" assumes cs: "closed_segment X (X + H) \ S" defines "i \ \x. ((1 - x) ^ (n - 1) / fact (n - 1)) *\<^sub>R nth_derivative n f (X + x *\<^sub>R H) H" shows "(i has_integral f (X + H) - (\iR nth_derivative i f X H)) {0..1}" (is ?th1) and "f (X + H) = (\iR nth_derivative i f X H) + integral {0..1} i" (is ?th2) and "i integrable_on {0..1}" (is ?th3) proof - from higher_derivatives_exists[OF hd] obtain Df where Df: "(\a h. Df a 0 h = f a)" "(\a h i. i < n \ a \ S \ ((\a. Df a i H) has_derivative Df a (Suc i)) (at a))" "\a i. i \ n \ a \ S \ Df a i H = nth_derivative i f a H" by blast from multivariate_Taylor_integral[OF \n > 0\, of Df H f X, OF Df(1,2)] cs have mt: "((\x. ((1 - x) ^ (n - 1) / fact (n - 1)) *\<^sub>R Df (X + x *\<^sub>R H) n H) has_integral f (X + H) - (\iR Df X i H)) {0..1}" by force from cs have "X \ S" by auto show ?th1 apply (rule has_integral_eq_rhs) unfolding i_def using negligible_empty _ mt apply (rule has_integral_spike) using closed_segment_subsetD[OF cs] by (auto simp: Df \X \ S\) then show ?th2 ?th3 unfolding has_integral_iff by auto qed lemma frechet_derivative_componentwise: "frechet_derivative f (at a) v = (\i\Basis. (v \ i) * (frechet_derivative f (at a) i))" if "f differentiable at a" for f::"'a::euclidean_space \ real" proof - have "linear (frechet_derivative f (at a))" using that by (rule linear_frechet_derivative) from Linear_Algebra.linear_componentwise[OF this, of v 1] show ?thesis by simp qed lemma second_derivative_componentwise: "nth_derivative 2 f a v = (\i\Basis. (\j\Basis. frechet_derivative (\a. frechet_derivative f (at a) j) (at a) i * (v \ j)) * (v \ i))" if "higher_differentiable_on S f 2" and S: "open S" "a \ S" for f::"'a::euclidean_space \ real" proof - have diff: "(\x. frechet_derivative f (at x) v) differentiable at a" for v using that by (auto simp: numeral_2_eq_2 higher_differentiable_on.simps differentiable_on_openD dest!: spec[where x=v]) have d1: "x \ S \ f differentiable at x" for x using that by (auto simp: numeral_2_eq_2 higher_differentiable_on.simps dest!: differentiable_on_openD) have eq: "\x. x \ Basis \ frechet_derivative (\x. \i\Basis. v \ i * frechet_derivative f (at x) i) (at a) x = (\j\Basis. frechet_derivative (\a. frechet_derivative f (at a) j) (at a) x * (v \ j))" apply (subst frechet_derivative_sum) subgoal by (auto intro!: differentiable_mult diff) apply (rule sum.cong) apply simp apply (subst frechet_derivative_times) subgoal by simp subgoal by (rule diff) by (simp add: frechet_derivative_const) show ?thesis apply (simp add: numeral_2_eq_2) apply (subst frechet_derivative_componentwise[OF diff]) apply (rule sum.cong) apply simp apply simp apply (rule disjI2) apply (rule trans) apply (rule frechet_derivative_transform_within_open_ext [OF _ S frechet_derivative_componentwise]) apply (simp add: diff) apply (rule d1, assumption) apply (simp add: eq) done qed lemma higher_differentiable_Taylor1: fixes f::"'a::real_normed_vector \ 'b::banach" assumes hd: "higher_differentiable_on S f 2" "open S" assumes cs: "closed_segment X (X + H) \ S" defines "i \ \x. ((1 - x)) *\<^sub>R nth_derivative 2 f (X + x *\<^sub>R H) H" shows "(i has_integral f (X + H) - (f X + nth_derivative 1 f X H)) {0..1}" and "f (X + H) = f X + nth_derivative 1 f X H + integral {0..1} i" and "i integrable_on {0..1}" using higher_differentiable_Taylor[OF _ hd cs] by (auto simp: numeral_2_eq_2 i_def) lemma differentiable_on_open_blinfunE: assumes "f differentiable_on S" "open S" obtains f' where "\x. x \ S \ (f has_derivative blinfun_apply (f' x)) (at x)" proof - { fix x assume "x \ S" with assms obtain f' where f': "(f has_derivative f') (at x)" by (auto dest!: differentiable_on_openD simp: differentiable_def) then have "bounded_linear f'" by (rule has_derivative_bounded_linear) then obtain bf' where "f' = blinfun_apply bf'" by (metis bounded_linear_Blinfun_apply) then have "\bf'. (f has_derivative blinfun_apply bf') (at x)" using f' by blast } then obtain f' where "\x. x \ S \ (f has_derivative blinfun_apply (f' x)) (at x)" by metis then show ?thesis .. qed lemma continuous_on_blinfunI1: "continuous_on X f" if "\i. i \ Basis \ continuous_on X (\x. blinfun_apply (f x) i)" using that by (auto simp: continuous_on_def intro: tendsto_componentwise1) lemma c1_euclidean_blinfunE: fixes f::"'a::euclidean_space\'b::real_normed_vector" assumes "\x. x \ S \ (f has_derivative f' x) (at x within S)" assumes "\i. i \ Basis \ continuous_on S (\x. f' x i)" obtains bf' where "\x. x \ S \ (f has_derivative blinfun_apply (bf' x)) (at x within S)" "continuous_on S bf'" "\x. x \ S \ blinfun_apply (bf' x) = f' x" proof - from assms have "bounded_linear (f' x)" if "x \ S" for x by (auto intro!: has_derivative_bounded_linear that) then obtain bf' where bf': "\x \ S. f' x = blinfun_apply (bf' x)" apply atomize_elim apply (rule bchoice) apply auto by (metis bounded_linear_Blinfun_apply) with assms have "\x. x \ S \ (f has_derivative blinfun_apply (bf' x)) (at x within S)" by simp moreover have f_tendsto: "((\n. f' n j) \ f' x j) (at x within S)" if "x \ S" "j \ Basis" for x j using assms by (auto simp: continuous_on_def that) have "continuous_on S bf'" by (rule continuous_on_blinfunI1) (simp add: bf'[rule_format, symmetric] assms) moreover have "\x. x \ S \ blinfun_apply (bf' x) = f' x" using bf' by auto ultimately show ?thesis .. qed lemma continuous_Sigma: assumes defined: "y \ Pi T X" assumes f_cont: "continuous_on (Sigma T X) (\(t, x). f t x)" assumes y_cont: "continuous_on T y" shows "continuous_on T (\x. f x (y x))" using defined continuous_on_compose2[OF continuous_on_subset[where t="(\x. (x, y x)) ` T", OF f_cont] continuous_on_Pair[OF continuous_on_id y_cont]] by auto lemma continuous_on_Times_swap: "continuous_on (X \ Y) (\(x, y). f x y)" if "continuous_on (Y \ X) (\(y, x). f x y)" using continuous_on_compose2[OF that continuous_on_swap, where s="X \ Y"] by (auto simp: split_beta' product_swap) lemma leibniz_rule': "\x. x \ S \ ((\x. integral (cbox a b) (f x)) has_derivative (\v. integral (cbox a b) (\t. fx x t v))) (at x within S)" "(\x. integral (cbox a b) (f x)) differentiable_on S" if "convex S" and c1: "\t x. t \ cbox a b \ x \ S \ ((\x. f x t) has_derivative fx x t) (at x within S)" "\i. i \ Basis \ continuous_on (S \ cbox a b) (\(x, t). fx x t i)" and i: "\x. x \ S \ f x integrable_on cbox a b" for S::"'a::euclidean_space set" and f::"'a \ 'b::euclidean_space \ 'c::euclidean_space" proof - have fx1: "continuous_on S (\x. fx x t i)" if "t \ cbox a b" "i \ Basis" for i t by (rule continuous_Sigma[OF _ c1(2), where y="\_. t"]) (auto simp: continuous_intros that) { fix x assume "x \ S" have fx2: "continuous_on (cbox a b) (\t. fx x t i)" if "i \ Basis" for i by (rule continuous_Sigma[OF _ continuous_on_Times_swap[OF c1(2)]]) (auto simp: continuous_intros that \x \ S\) { fix t assume "t \ cbox a b" have "\f'. (\x \ S. ((\x. f x t) has_derivative blinfun_apply (f' x)) (at x within S) \ blinfun_apply (f' x) = fx x t) \ continuous_on S f'" by (rule c1_euclidean_blinfunE[OF c1(1)[OF \t \ _\] fx1[OF \t \ _\]]) (auto, metis) } then obtain fx' where fx': "\t x. t \ cbox a b \ x \ S \ ((\x. f x t) has_derivative blinfun_apply (fx' t x)) (at x within S)" "\t x. t \ cbox a b \ x \ S \ blinfun_apply (fx' t x) = fx x t" "\t. t \ cbox a b \ continuous_on S (fx' t)" by metis have c: "continuous_on (S \ cbox a b) (\(x, t). fx' t x)" apply (rule continuous_on_blinfunI1) using c1(2) apply (rule continuous_on_eq) apply assumption by (auto simp: fx' split_beta') from leibniz_rule[of S a b f "\x t. fx' t x" x, OF fx'(1) i c \x \ S\ \convex S\] have "((\x. integral (cbox a b) (f x)) has_derivative blinfun_apply (integral (cbox a b) (\t. fx' t x))) (at x within S)" by auto then have "((\x. integral (cbox a b) (f x)) has_derivative blinfun_apply (integral (cbox a b) (\t. fx' t x))) (at x within S)" by auto also have fx'xi: "(\t. fx' t x) integrable_on cbox a b" apply (rule integrable_continuous) apply (rule continuous_on_blinfunI1) by (simp add: fx' \x \ S\ fx2) have "blinfun_apply (integral (cbox a b) (\t. fx' t x)) = (\v. integral (cbox a b) (\xb. fx x xb v))" apply (rule ext) apply (subst blinfun_apply_integral) apply (rule fx'xi) by (simp add: \x \ S\ fx' cong: integral_cong) finally show "((\x. integral (cbox a b) (f x)) has_derivative (\c. integral (cbox a b) (\xb. fx x xb c))) (at x within S)" by simp } then show "(\x. integral (cbox a b) (f x)) differentiable_on S" by (auto simp: differentiable_on_def differentiable_def) qed lemmas leibniz_rule'_interval = leibniz_rule'[where 'b="_::ordered_euclidean_space", unfolded cbox_interval] lemma leibniz_rule'_higher: "higher_differentiable_on S (\x. integral (cbox a b) (f x)) k" if "convex S" "open S" and c1: "higher_differentiable_on (S\cbox a b) (\(x, t). f x t) k" \\this condition is actually too strong: it would suffice if higher partial derivatives (w.r.t. x) are continuous w.r.t. t. but it makes the statement short and no need to introduce new constants\ for S::"'a::euclidean_space set" and f::"'a \ 'b::euclidean_space \ 'c::euclidean_space" using c1 proof (induction k arbitrary: f) case 0 then show ?case by (auto simp: higher_differentiable_on.simps intro!: integral_continuous_on_param) next case (Suc k) define D where "D x = frechet_derivative (\(x, y). f x y) (at x)" for x note [continuous_intros] = Suc.prems[THEN higher_differentiable_on_imp_continuous_on, THEN continuous_on_compose2, of _ "\x. (f x, g x)" for f g, unfolded split_beta' fst_conv snd_conv] from Suc.prems have prems: "\xt. xt \ S \ cbox a b \ (\(x, y). f x y) differentiable at xt" "higher_differentiable_on (S \ cbox a b) (\x. D x (dx, dt)) k" for dx dt by (auto simp: higher_differentiable_on.simps D_def) from frechet_derivative_worksI[OF this(1), folded D_def] have D: "x \ S \ t \ cbox a b \ ((\(x, y). f x y) has_derivative D (x, t)) (at (x, t))" for x t by auto have p1: "((\x. (x, t::'b)) has_derivative (\h. (h, 0))) (at x within S)" for x t by (auto intro!: derivative_eq_intros) have Dx: "x \ S \ t \ cbox a b \ ((\x. f x t) has_derivative (\dx. D (x, t) (dx, 0))) (at x within S)" for x t by (drule has_derivative_compose[OF p1 D], assumption) auto have cD: "continuous_on (S \ cbox a b) (\(x, t). D (x, t) v)" for v apply (rule higher_differentiable_on_imp_continuous_on[where n=k]) using prems(2)[of "fst v" "snd v"] by (auto simp: split_beta') have fi: "x \ S \ f x integrable_on cbox a b" for x by (rule integrable_continuous) (auto intro!: continuous_intros) from leibniz_rule'[OF \convex S\ Dx cD fi] have ihd: "x \ S \ ((\x. integral (cbox a b) (f x)) has_derivative (\v. integral (cbox a b) (\t. D (x, t) (v, 0)))) (at x within S)" and "(\x. integral (cbox a b) (f x)) differentiable_on S" for x by auto then have "x \ S \ (\x. integral (cbox a b) (f x)) differentiable at x" for x using \open S\ by (auto simp: differentiable_on_openD) moreover have "higher_differentiable_on S (\x. frechet_derivative (\x. integral (cbox a b) (f x)) (at x) v) k" for v proof - have *: "frechet_derivative (\x. integral (cbox a b) (f x)) (at x) = (\v. integral (cbox a b) (\t. D (x, t) (v, 0)))" if "x \ S" for x apply (rule frechet_derivative_at') using ihd(1)[OF that] at_within_open[OF that \open S\] by auto have **: "higher_differentiable_on S (\x. integral (cbox a b) (\t. D (x, t) (v, 0))) k" apply (rule Suc.IH) using prems by auto show ?thesis using \open S\ by (auto simp: * ** cong: higher_differentiable_on_cong) qed ultimately show ?case by (auto simp: higher_differentiable_on.simps) qed lemmas leibniz_rule'_higher_interval = leibniz_rule'_higher[where 'b="_::ordered_euclidean_space", unfolded cbox_interval] subsection \Smoothness\ definition k_smooth_on :: "enat \'a::real_normed_vector set \ ('a \ 'b::real_normed_vector) \ bool" ("_-smooth'_on" [1000]) where smooth_on_def: "k-smooth_on S f = (\n\k. higher_differentiable_on S f n)" abbreviation "smooth_on S f \ \-smooth_on S f" lemma derivative_is_smooth': assumes "(k+1)-smooth_on S f" shows "k-smooth_on S (\x. frechet_derivative f (at x) v)" unfolding smooth_on_def proof (intro allI impI) fix n assume "enat n \ k" then have "Suc n \ k + 1" unfolding plus_1_eSuc by (auto simp: eSuc_def split: enat.splits) then have "higher_differentiable_on S f (Suc n)" using assms(1) by (auto simp: smooth_on_def) then show "higher_differentiable_on S (\x. frechet_derivative f (at x) v) n" by (auto simp: higher_differentiable_on.simps(2)) qed lemma derivative_is_smooth: "smooth_on S f \ smooth_on S (\x. frechet_derivative f (at x) v)" using derivative_is_smooth'[of \ S f] by simp lemma smooth_on_imp_continuous_on: "continuous_on S f" if "k-smooth_on S f" apply (rule higher_differentiable_on_imp_continuous_on[where n=0]) using that by (simp add: smooth_on_def enat_0) lemma smooth_on_imp_differentiable_on[simp]: "f differentiable_on S" if "k-smooth_on S f" "k \ 0" using that by (auto simp: smooth_on_def Suc_ile_eq enat_0 dest!: spec[where x=1] intro!: higher_differentiable_on_imp_differentiable_on[where k=1]) lemma smooth_on_cong: assumes "k-smooth_on S g" "open S" and "\x. x \ S \ f x = g x" shows "k-smooth_on S f" using assms unfolding smooth_on_def by (auto cong: higher_differentiable_on_cong) lemma smooth_on_open_Un: "k-smooth_on S f \ k-smooth_on T f \ open S \ open T \ k-smooth_on (S \ T) f" by (auto simp: smooth_on_def higher_differentiable_on_open_Un) lemma smooth_on_open_subsetsI: "k-smooth_on S f" if "\x. x \ S \ \T. x \ T \ open T \ k-smooth_on T f" using that unfolding smooth_on_def by (force intro: higher_differentiable_on_open_subsetsI) lemma smooth_on_const[intro]: "k-smooth_on S (\x. c)" by (auto simp: smooth_on_def higher_differentiable_on_const) lemma smooth_on_id[intro]: "k-smooth_on S (\x. x)" by (auto simp: smooth_on_def higher_differentiable_on_id) lemma smooth_on_add_fun: "k-smooth_on S f \ k-smooth_on S g \ open S \ k-smooth_on S (f + g)" by (auto simp: smooth_on_def higher_differentiable_on_add plus_fun_def) lemmas smooth_on_add = smooth_on_add_fun[unfolded plus_fun_def] lemma smooth_on_sum: "n-smooth_on S (\x. \i\F. f i x)" if "\i. i \ F \ finite F \ n-smooth_on S (f i)" "open S" using that by (auto simp: smooth_on_def higher_differentiable_on_sum) lemma (in bounded_bilinear) smooth_on: includes no_matrix_mult assumes "k-smooth_on S f" "k-smooth_on S g" "open S" shows "k-smooth_on S (\x. (f x) ** (g x))" using assms by (auto simp: smooth_on_def higher_differentiable_on) lemma smooth_on_compose2: fixes g::"_\_::euclidean_space" assumes "k-smooth_on T f" "k-smooth_on S g" "open U" "open T" "g ` U \ T" "U \ S" shows "k-smooth_on U (f o g)" using assms by (auto simp: smooth_on_def intro!: higher_differentiable_on_compose intro: higher_differentiable_on_subset) lemma smooth_on_compose: fixes g::"_\_::euclidean_space" assumes "k-smooth_on T f" "k-smooth_on S g" "open S" "open T" "g ` S \ T" shows "k-smooth_on S (f o g)" using assms by (rule smooth_on_compose2) auto lemma smooth_on_subset: "k-smooth_on S f" if "k-smooth_on T f" "S \ T" using higher_differentiable_on_subset[of T f _ S] that by (auto simp: smooth_on_def) context begin private lemmas s = bounded_bilinear.smooth_on lemmas smooth_on_inner = bounded_bilinear_inner[THEN s] and smooth_on_scaleR = bounded_bilinear_scaleR[THEN s] and smooth_on_mult = bounded_bilinear_mult[THEN s] end lemma smooth_on_divide:"k-smooth_on S f \ k-smooth_on S g \ open S \(\x. x \ S \ g x \ 0) \ k-smooth_on S (\x. f x / g x)" for f::"_\_::real_normed_field" by (auto simp: smooth_on_def higher_differentiable_on_divide) lemma smooth_on_scaleR_fun: "k-smooth_on S g \ open S \ k-smooth_on S (c *\<^sub>R g)" by (auto simp: scaleR_fun_def intro!: smooth_on_scaleR ) lemma smooth_on_uminus_fun: "k-smooth_on S g \ open S \ k-smooth_on S (- g)" using smooth_on_scaleR_fun[where c="-1", of k S g] by auto lemmas smooth_on_uminus = smooth_on_uminus_fun[unfolded fun_Compl_def] lemma smooth_on_minus_fun: "k-smooth_on S f \ k-smooth_on S g \ open S \ k-smooth_on S (f - g)" unfolding diff_conv_add_uminus apply (rule smooth_on_add_fun) apply assumption apply (rule smooth_on_uminus_fun) by auto lemmas smooth_on_minus = smooth_on_minus_fun[unfolded fun_diff_def] lemma smooth_on_times_fun: "k-smooth_on S f \ k-smooth_on S g \ open S \ k-smooth_on S (f * g)" for f g::"_ \_::real_normed_algebra" by (auto simp: times_fun_def intro!: smooth_on_mult) lemma smooth_on_le: "l-smooth_on S f" if "k-smooth_on S f" "l \ k" using that by (auto simp: smooth_on_def) lemma smooth_on_inverse: "k-smooth_on S (\x. inverse (f x))" if "k-smooth_on S f" "0 \ f ` S" "open S" for f::"_ \_::real_normed_field" using that by (auto simp: smooth_on_def intro!: higher_differentiable_on_inverse) lemma smooth_on_norm: "k-smooth_on S (\x. norm (f x))" if "k-smooth_on S f" "0 \ f ` S" "open S" for f::"_ \_::real_inner" using that by (auto simp: smooth_on_def intro!: higher_differentiable_on_norm) lemma smooth_on_sqrt: "k-smooth_on S (\x. sqrt (f x))" if "k-smooth_on S f" "0 \ f ` S" "open S" using that by (auto simp: smooth_on_def intro!: higher_differentiable_on_sqrt) lemma smooth_on_frechet_derivative: "\-smooth_on UNIV (\x. frechet_derivative f (at x) v)" if "\-smooth_on UNIV f" \\TODO: generalize\ using that apply (auto simp: smooth_on_def) apply (rule higher_differentiable_on_frechet_derivativeI) by auto lemmas smooth_on_frechet_derivivative_comp = smooth_on_compose2[OF smooth_on_frechet_derivative, unfolded o_def] lemma smooth_onD: "higher_differentiable_on S f n" if "m-smooth_on S f" "enat n \ m" using that by (auto simp: smooth_on_def) lemma (in bounded_linear) higher_differentiable_on: "higher_differentiable_on S f n" proof (induction n) case 0 then show ?case by (auto simp: higher_differentiable_on.simps linear_continuous_on bounded_linear_axioms) next case (Suc n) then show ?case apply (auto simp: higher_differentiable_on.simps frechet_derivative higher_differentiable_on_const) using bounded_linear_axioms apply (rule bounded_linear_imp_differentiable) done qed lemma (in bounded_linear) smooth_on: "k-smooth_on S f" by (auto simp: smooth_on_def higher_differentiable_on) lemma smooth_on_snd: "k-smooth_on S (\x. snd (f x))" if "k-smooth_on S f" "open S" using higher_differentiable_on_snd_comp that by (auto simp: smooth_on_def) lemma smooth_on_fst: "k-smooth_on S (\x. fst (f x))" if "k-smooth_on S f" "open S" using higher_differentiable_on_fst_comp that by (auto simp: smooth_on_def) lemma smooth_on_sin: "n-smooth_on S (\x. sin (f x::real))" if "n-smooth_on S f" "open S" using that by (auto simp: smooth_on_def intro!: higher_differentiable_on_sin) lemma smooth_on_cos: "n-smooth_on S (\x. cos (f x::real))" if "n-smooth_on S f" "open S" using that by (auto simp: smooth_on_def intro!: higher_differentiable_on_cos) lemma smooth_on_Taylor2E: fixes f::"'a::euclidean_space \ real" assumes hd: "\-smooth_on UNIV f" obtains g where "\Y. f Y = f X + frechet_derivative f (at X) (Y - X) + (\i\Basis. (\j\Basis. ((Y - X) \ j) * ((Y - X) \ i) * g i j Y))" "\i j. i \ Basis \ j \ Basis \ \-smooth_on UNIV (g i j)" \\TODO: generalize\ proof - define S::"'a set" where "S = UNIV" have "open S" and "convex S" "X \ S" by (auto simp: S_def) have hd: "\-smooth_on S f" using hd by (auto simp: S_def) define i where "i H x = ((1 - x)) *\<^sub>R nth_derivative 2 f (X + x *\<^sub>R H) H" for x H define d2 where "d2 v v' = (\x. frechet_derivative (\x. frechet_derivative f (at x) v') (at x) v)" for v v' define g where "g H x i j = d2 i j (X + x *\<^sub>R H)" for i j x H define g' where "g' i j H = integral {0 .. 1} (\x. (1 - x) * g H x i j)" for i j H have "higher_differentiable_on S f 2" using hd(1) by (auto simp: smooth_on_def dest!: spec[where x=2]) note hd2 = this \open S\ have d2_cont: "continuous_on S (d2 i j)" for i j using hd2(1) by (auto simp: g_def numeral_2_eq_2 higher_differentiable_on.simps d2_def) note [continuous_intros] = continuous_on_compose2[OF d2_cont] have hdiff2: "\-smooth_on S (d2 v v')" for v v' apply (auto simp: d2_def) apply (rule smooth_on_frechet_derivivative_comp) apply (rule smooth_on_frechet_derivivative_comp) by (auto simp: S_def assms) { fix Y assume "Y \ S" define H where "H = Y - X" from \Y \ S\ have "X + H \ S" by (simp add: H_def) with \X \ S\ have cs: "closed_segment X (X + H) \ S" using \convex S\ by (rule closed_segment_subset) have i: "(i H has_integral f (X + H) - (f X + nth_derivative 1 f X H)) {0..1}" "f (X + H) = f X + nth_derivative 1 f X H + integral {0..1} (i H)" "i H integrable_on {0..1}" unfolding i_def using higher_differentiable_Taylor1[OF hd2 cs] by auto note i(2) also have integrable: "(\x. \n\Basis. (1 - x) * (g H x a n * (H \ n) * (H \ a))) integrable_on {0..1}" "(\x. (1 - x) * (g H x n a * (H \ a) * (H \ n))) integrable_on {0..1}" for a n by (auto intro!: integrable_continuous_interval continuous_intros closed_segment_subsetD cs simp: g_def) have i_eq: "i H x = (1 - x) *\<^sub>R (\i\Basis. (\j\Basis. g H x i j * (H \ j)) * (H \ i))" if "0 \ x" "x \ 1" for x unfolding i_def apply (subst second_derivative_componentwise[OF hd2]) apply (rule closed_segment_subsetD, rule cs, rule that, rule that) by (simp add: g_def d2_def) have "integral {0 .. 1} (i H) = integral {0..1} (\x. (1 - x) * (\i\Basis. (\j\Basis. g H x i j * (H \ j)) * (H \ i)))" apply (subst integral_spike[OF negligible_empty]) apply (rule sym) apply (rule i_eq) by (auto simp: that) also have "\ = (\i\Basis. (\j\Basis. (H \ j) * (H \ i) * g' i j H))" apply (simp add: sum_distrib_left sum_distrib_right integral_sum integrable g'_def) apply (simp add: integral_mult_right[symmetric] del: integral_mult_right) by (simp only: ac_simps) finally have "f (X + H) = f X + nth_derivative 1 f X H + (\i\Basis. \j\Basis. H \ j * (H \ i) * g' i j H)" . } note * = this have "f Y = f X + frechet_derivative f (at X) (Y - X) + (\i\Basis. \j\Basis. (Y - X) \ j * ((Y - X) \ i) * g' i j (Y - X))" for Y using *[of Y] by (auto simp: S_def) moreover define T::"real set" where "T = {- 1<..<2}" have "open T" by (auto simp: T_def) have "{0 .. 1} \ T" by (auto simp: T_def) have T_small: "a \ S \ b \ T \ X + b *\<^sub>R (a - X) \ S" for a b by (auto simp: S_def) have "open (S \ T)" by (auto intro: open_Times \open S\ \open T\) have "smooth_on S (\Y. g' i j (Y - X))" if i: "i \ Basis" and j: "j \ Basis" for i j unfolding smooth_on_def apply safe apply (simp add: g'_def) apply (rule leibniz_rule'_higher_interval) apply fact apply fact apply (rule higher_differentiable_on_subset[where T="S \ T"]) apply (auto intro!: higher_differentiable_on_mult simp: split_beta') apply (subst diff_conv_add_uminus) apply (rule higher_differentiable_on_add) apply (rule higher_differentiable_on_const) apply (subst scaleR_minus1_left[symmetric]) apply (rule higher_differentiable_on_scaleR) apply (rule higher_differentiable_on_const) apply (rule higher_differentiable_on_snd_comp) apply (rule higher_differentiable_on_id) apply fact apply fact apply fact apply (auto simp: g_def) apply (rule smooth_onD) apply (rule smooth_on_compose2[OF hdiff2, unfolded o_def]) using \open S\ \open T\ using T_small \_ \ T\ by (auto intro!: open_Times smooth_on_add smooth_on_scaleR smooth_on_snd smooth_on_minus smooth_on_fst) ultimately show ?thesis unfolding S_def .. qed lemma smooth_on_Pair: "k-smooth_on S (\x. (f x, g x))" if "open S" "k-smooth_on S f" "k-smooth_on S g" proof (auto simp: smooth_on_def) fix n assume n: "enat n \ k" have 1: "higher_differentiable_on S f n" using that(2) n unfolding smooth_on_def by auto have 2: "higher_differentiable_on S g n" using that(3) n unfolding smooth_on_def by auto show "higher_differentiable_on S (\x. (f x, g x)) n" by (rule higher_differentiable_on_Pair [OF that(1) 1 2]) qed lemma smooth_on_Pair': "k-smooth_on (S \ T) (\x. (f (fst x), g (snd x)))" if "open S" "open T" "k-smooth_on S f" "k-smooth_on T g" for f::"_::euclidean_space\_" and g::"_::euclidean_space\_" proof (auto simp: smooth_on_def) fix n assume n: "enat n \ k" have 1: "higher_differentiable_on S f n" using that(3) n unfolding smooth_on_def by auto have 2: "higher_differentiable_on T g n" using that(4) n unfolding smooth_on_def by auto show "higher_differentiable_on (S \ T) (\x. (f (fst x), g (snd x))) n" by (rule higher_differentiable_on_Pair'[OF that(1,2) 1 2]) qed subsection \Diffeomorphism\ definition "diffeomorphism k S T p p' \ k-smooth_on S p \ k-smooth_on T p' \ homeomorphism S T p p'" lemma diffeomorphism_imp_homeomorphism: assumes "diffeomorphism k s t p p'" shows "homeomorphism s t p p'" using assms by (auto simp: diffeomorphism_def) lemma diffeomorphismD: assumes "diffeomorphism k S T f g" shows diffeomorphism_smoothD: "k-smooth_on S f" "k-smooth_on T g" and diffeomorphism_inverseD: "\x. x \ S \ g (f x) = x" "\y. y \ T \ f (g y) = y" and diffeomorphism_image_eq: "(f ` S = T)" "(g ` T = S)" using assms by (auto simp: diffeomorphism_def homeomorphism_def) lemma diffeomorphism_compose: "diffeomorphism n S T f g \ diffeomorphism n T U h k \ open S \ open T \ open U \ diffeomorphism n S U (h \ f) (g \ k)" for f::"_\_::euclidean_space" by (auto simp: diffeomorphism_def intro!: smooth_on_compose homeomorphism_compose) (auto simp: homeomorphism_def) lemma diffeomorphism_add: "diffeomorphism k UNIV UNIV (\x. x + c) (\x. x - c)" by (auto simp: diffeomorphism_def homeomorphism_add intro!: smooth_on_minus smooth_on_add) lemma diffeomorphism_scaleR: "diffeomorphism k UNIV UNIV (\x. c *\<^sub>R x) (\x. x /\<^sub>R c)" if "c \ 0" by (auto simp: that diffeomorphism_def homeomorphism_scaleR intro!: smooth_on_minus smooth_on_scaleR) end