diff --git a/thys/Affine_Arithmetic/Floatarith_Expression.thy b/thys/Affine_Arithmetic/Floatarith_Expression.thy --- a/thys/Affine_Arithmetic/Floatarith_Expression.thy +++ b/thys/Affine_Arithmetic/Floatarith_Expression.thy @@ -1,2232 +1,2232 @@ section \Operations on Expressions\ theory Floatarith_Expression imports "HOL-Decision_Procs.Approximation" Affine_Arithmetic_Auxiliarities Executable_Euclidean_Space begin text \Much of this could move to the distribution...\ subsection \Approximating Expression*s*\ unbundle floatarith_notation text \\label{sec:affineexpr}\ primrec interpret_floatariths :: "floatarith list \ real list \ real list" where "interpret_floatariths [] vs = []" | "interpret_floatariths (a#bs) vs = interpret_floatarith a vs#interpret_floatariths bs vs" lemma length_interpret_floatariths[simp]: "length (interpret_floatariths fas xs) = length fas" by (induction fas) auto lemma interpret_floatariths_nth[simp]: "interpret_floatariths fas xs ! n = interpret_floatarith (fas ! n) xs" if "n < length fas" using that by (induction fas arbitrary: n) (auto simp: nth_Cons split: nat.splits) abbreviation "einterpret \ \fas vs. eucl_of_list (interpret_floatariths fas vs)" subsection \Syntax\ syntax interpret_floatarith::"floatarith \ real list \ real" instantiation floatarith :: "{plus, minus, uminus, times, inverse, zero, one}" begin definition "- f = Minus f" lemma interpret_floatarith_uminus[simp]: "interpret_floatarith (- f) xs = - interpret_floatarith f xs" by (auto simp: uminus_floatarith_def) definition "f + g = Add f g" lemma interpret_floatarith_plus[simp]: "interpret_floatarith (f + g) xs = interpret_floatarith f xs + interpret_floatarith g xs" by (auto simp: plus_floatarith_def) definition "f - g = Add f (Minus g)" lemma interpret_floatarith_minus[simp]: "interpret_floatarith (f - g) xs = interpret_floatarith f xs - interpret_floatarith g xs" by (auto simp: minus_floatarith_def) definition "inverse f = Inverse f" lemma interpret_floatarith_inverse[simp]: "interpret_floatarith (inverse f) xs = inverse (interpret_floatarith f xs)" by (auto simp: inverse_floatarith_def) definition "f * g = Mult f g" lemma interpret_floatarith_times[simp]: "interpret_floatarith (f * g) xs = interpret_floatarith f xs * interpret_floatarith g xs" by (auto simp: times_floatarith_def) definition "f div g = f * Inverse g" lemma interpret_floatarith_divide[simp]: "interpret_floatarith (f div g) xs = interpret_floatarith f xs / interpret_floatarith g xs" by (auto simp: divide_floatarith_def inverse_eq_divide) definition "1 = Num 1" lemma interpret_floatarith_one[simp]: "interpret_floatarith 1 xs = 1" by (auto simp: one_floatarith_def) definition "0 = Num 0" lemma interpret_floatarith_zero[simp]: "interpret_floatarith 0 xs = 0" by (auto simp: zero_floatarith_def) instance proof qed end subsection \Derived symbols\ definition "R\<^sub>e r = (case quotient_of r of (n, d) \ Num (of_int n) / Num (of_int d))" declare [[coercion R\<^sub>e ]] lemma interpret_R\<^sub>e[simp]: "interpret_floatarith (R\<^sub>e x) xs = real_of_rat x" by (auto simp: R\<^sub>e_def of_rat_divide dest!: quotient_of_div split: prod.splits) definition "Sin x = Cos ((Pi * (Num (Float 1 (-1)))) - x)" lemma interpret_floatarith_Sin[simp]: "interpret_floatarith (Sin x) vs = sin (interpret_floatarith x vs)" by (auto simp: Sin_def approximation_preproc_floatarith(11)) definition "Half x = Mult (Num (Float 1 (-1))) x" lemma interpret_Half[simp]: "interpret_floatarith (Half x) xs = interpret_floatarith x xs / 2" by (auto simp: Half_def) definition "Tan x = (Sin x) / (Cos x)" lemma interpret_floatarith_Tan[simp]: "interpret_floatarith (Tan x) vs = tan (interpret_floatarith x vs)" by (auto simp: Tan_def approximation_preproc_floatarith(12) inverse_eq_divide) primrec Sum\<^sub>e where "Sum\<^sub>e f [] = 0" | "Sum\<^sub>e f (x#xs) = f x + Sum\<^sub>e f xs" lemma interpret_floatarith_Sum\<^sub>e[simp]: "interpret_floatarith (Sum\<^sub>e f x) vs = (\i\x. interpret_floatarith (f i) vs)" by (induction x) auto definition Norm where "Norm is = Sqrt (Sum\<^sub>e (\i. i * i) is)" lemma interpret_floatarith_norm[simp]: assumes [simp]: "length x = DIM('a)" shows "interpret_floatarith (Norm x) vs = norm (einterpret x vs::'a::executable_euclidean_space)" apply (auto simp: Norm_def norm_eq_sqrt_inner) apply (subst euclidean_inner[where 'a='a]) apply (auto simp: power2_eq_square[symmetric] ) apply (subst sum_list_Basis_list[symmetric]) apply (rule sum_list_nth_eqI) by (auto simp: in_set_zip eucl_of_list_inner) notation floatarith.Power (infixr "^\<^sub>e" 80) subsection \Constant Folding\ fun dest_Num_fa where "dest_Num_fa (floatarith.Num x) = Some x" | "dest_Num_fa _ = None" fun_cases dest_Num_fa_None: "dest_Num_fa fa = None" and dest_Num_fa_Some: "dest_Num_fa fa = Some x" fun fold_const_fa where "fold_const_fa (Add fa1 fa2) = (let (ffa1, ffa2) = (fold_const_fa fa1, fold_const_fa fa2) in case (dest_Num_fa ffa1, dest_Num_fa (ffa2)) of (Some a, Some b) \ Num (a + b) | (Some a, None) \ (if a = 0 then ffa2 else Add (Num a) ffa2) | (None, Some a) \ (if a = 0 then ffa1 else Add ffa1 (Num a)) | (None, None) \ Add ffa1 ffa2)" | "fold_const_fa (Minus a) = (case (fold_const_fa a) of (Num x) \ Num (-x) | x \ Minus x)" | "fold_const_fa (Mult fa1 fa2) = (let (ffa1, ffa2) = (fold_const_fa fa1, fold_const_fa fa2) in case (dest_Num_fa ffa1, dest_Num_fa (ffa2)) of (Some a, Some b) \ Num (a * b) | (Some a, None) \ (if a = 0 then Num 0 else if a = 1 then ffa2 else Mult (Num a) ffa2) | (None, Some a) \ (if a = 0 then Num 0 else if a = 1 then ffa1 else Mult ffa1 (Num a)) | (None, None) \ Mult ffa1 ffa2)" | "fold_const_fa (Inverse a) = Inverse (fold_const_fa a)" | "fold_const_fa (Abs a) = (case (fold_const_fa a) of (Num x) \ Num (abs x) | x \ Abs x)" | "fold_const_fa (Max a b) = (case (fold_const_fa a, fold_const_fa b) of (Num x, Num y) \ Num (max x y) | (x, y) \ Max x y)" | "fold_const_fa (Min a b) = (case (fold_const_fa a, fold_const_fa b) of (Num x, Num y) \ Num (min x y) | (x, y) \ Min x y)" | "fold_const_fa (Floor a) = (case (fold_const_fa a) of (Num x) \ Num (floor_fl x) | x \ Floor x)" | "fold_const_fa (Power a b) = (case (fold_const_fa a) of (Num x) \ Num (x ^ b) | x \ Power x b)" | "fold_const_fa (Cos a) = Cos (fold_const_fa a)" | "fold_const_fa (Arctan a) = Arctan (fold_const_fa a)" | "fold_const_fa (Sqrt a) = Sqrt (fold_const_fa a)" | "fold_const_fa (Exp a) = Exp (fold_const_fa a)" | "fold_const_fa (Ln a) = Ln (fold_const_fa a)" | "fold_const_fa (Powr a b) = Powr (fold_const_fa a) (fold_const_fa b)" | "fold_const_fa Pi = Pi" | "fold_const_fa (Var v) = Var v" | "fold_const_fa (Num x) = Num x" fun_cases fold_const_fa_Num: "fold_const_fa fa = Num y" and fold_const_fa_Add: "fold_const_fa fa = Add x y" and fold_const_fa_Minus: "fold_const_fa fa = Minus y" lemma fold_const_fa[simp]: "interpret_floatarith (fold_const_fa fa) xs = interpret_floatarith fa xs" by (induction fa) (auto split!: prod.splits floatarith.splits option.splits elim!: dest_Num_fa_None dest_Num_fa_Some simp: max_def min_def floor_fl_def) subsection \Free Variables\ primrec max_Var_floatarith where\ \TODO: include bound in predicate\ "max_Var_floatarith (Add a b) = max (max_Var_floatarith a) (max_Var_floatarith b)" | "max_Var_floatarith (Mult a b) = max (max_Var_floatarith a) (max_Var_floatarith b)" | "max_Var_floatarith (Inverse a) = max_Var_floatarith a" | "max_Var_floatarith (Minus a) = max_Var_floatarith a" | "max_Var_floatarith (Num a) = 0" | "max_Var_floatarith (Var i) = Suc i" | "max_Var_floatarith (Cos a) = max_Var_floatarith a" | "max_Var_floatarith (floatarith.Arctan a) = max_Var_floatarith a" | "max_Var_floatarith (Abs a) = max_Var_floatarith a" | "max_Var_floatarith (floatarith.Max a b) = max (max_Var_floatarith a) (max_Var_floatarith b)" | "max_Var_floatarith (floatarith.Min a b) = max (max_Var_floatarith a) (max_Var_floatarith b)" | "max_Var_floatarith (floatarith.Pi) = 0" | "max_Var_floatarith (Sqrt a) = max_Var_floatarith a" | "max_Var_floatarith (Exp a) = max_Var_floatarith a" | "max_Var_floatarith (Powr a b) = max (max_Var_floatarith a) (max_Var_floatarith b)" | "max_Var_floatarith (floatarith.Ln a) = max_Var_floatarith a" | "max_Var_floatarith (Power a n) = max_Var_floatarith a" | "max_Var_floatarith (Floor a) = max_Var_floatarith a" primrec max_Var_floatariths where "max_Var_floatariths [] = 0" | "max_Var_floatariths (x#xs) = max (max_Var_floatarith x) (max_Var_floatariths xs)" primrec max_Var_form where "max_Var_form (Conj a b) = max (max_Var_form a) (max_Var_form b)" | "max_Var_form (Disj a b) = max (max_Var_form a) (max_Var_form b)" | "max_Var_form (Less a b) = max (max_Var_floatarith a) (max_Var_floatarith b)" | "max_Var_form (LessEqual a b) = max (max_Var_floatarith a) (max_Var_floatarith b)" | "max_Var_form (Bound a b c d) = linorder_class.Max {max_Var_floatarith a,max_Var_floatarith b, max_Var_floatarith c, max_Var_form d}" | "max_Var_form (AtLeastAtMost a b c) = linorder_class.Max {max_Var_floatarith a,max_Var_floatarith b, max_Var_floatarith c}" | "max_Var_form (Assign a b c) = linorder_class.Max {max_Var_floatarith a,max_Var_floatarith b, max_Var_form c}" lemma interpret_floatarith_eq_take_max_VarI: assumes "take (max_Var_floatarith ra) ys = take (max_Var_floatarith ra) zs" shows "interpret_floatarith ra ys = interpret_floatarith ra zs" using assms by (induct ra) (auto dest!: take_max_eqD simp: take_Suc_eq split: if_split_asm) lemma interpret_floatariths_eq_take_max_VarI: assumes "take (max_Var_floatariths ea) ys = take (max_Var_floatariths ea) zs" shows "interpret_floatariths ea ys = interpret_floatariths ea zs" using assms apply (induction ea) subgoal by simp subgoal by (clarsimp) (metis interpret_floatarith_eq_take_max_VarI take_map take_max_eqD) done lemma Max_Image_distrib: includes no_floatarith_notation assumes "finite X" "X \ {}" shows "Max ((\x. max (f1 x) (f2 x)) ` X) = max (Max (f1 ` X)) (Max (f2 ` X))" apply (rule Max_eqI) subgoal using assms by simp subgoal for y using assms by (force intro: max.coboundedI1 max.coboundedI2 Max_ge) subgoal proof - have "Max (f1 ` X) \ f1 ` X" using assms by auto then obtain x1 where x1: "x1 \ X" "Max (f1 ` X) = f1 x1" by auto have "Max (f2 ` X) \ f2 ` X" using assms by auto then obtain x2 where x2: "x2 \ X" "Max (f2 ` X) = f2 x2" by auto show ?thesis apply (rule image_eqI[where x="if f1 x1 \ f2 x2 then x2 else x1"]) using x1 x2 assms apply (auto simp: max_def) apply (metis Max_ge dual_order.trans finite_imageI image_eqI assms(1)) apply (metis Max_ge dual_order.trans finite_imageI image_eqI assms(1)) done qed done lemma max_Var_floatarith_simps[simp]: "max_Var_floatarith (a / b) = max (max_Var_floatarith a) (max_Var_floatarith b)" "max_Var_floatarith (a * b) = max (max_Var_floatarith a) (max_Var_floatarith b)" "max_Var_floatarith (a + b) = max (max_Var_floatarith a) (max_Var_floatarith b)" "max_Var_floatarith (a - b) = max (max_Var_floatarith a) (max_Var_floatarith b)" "max_Var_floatarith (- b) = (max_Var_floatarith b)" by (auto simp: divide_floatarith_def times_floatarith_def plus_floatarith_def minus_floatarith_def uminus_floatarith_def) lemma max_Var_floatariths_Max: "max_Var_floatariths xs = (if set xs = {} then 0 else linorder_class.Max (max_Var_floatarith ` set xs))" by (induct xs) auto lemma max_Var_floatariths_map_plus[simp]: "max_Var_floatariths (map (\i. fa1 i + fa2 i) xs) = max (max_Var_floatariths (map fa1 xs)) (max_Var_floatariths (map fa2 xs))" by (auto simp: max_Var_floatariths_Max image_image Max_Image_distrib) lemma max_Var_floatariths_map_times[simp]: "max_Var_floatariths (map (\i. fa1 i * fa2 i) xs) = max (max_Var_floatariths (map fa1 xs)) (max_Var_floatariths (map fa2 xs))" by (auto simp: max_Var_floatariths_Max image_image Max_Image_distrib) lemma max_Var_floatariths_map_divide[simp]: "max_Var_floatariths (map (\i. fa1 i / fa2 i) xs) = max (max_Var_floatariths (map fa1 xs)) (max_Var_floatariths (map fa2 xs))" by (auto simp: max_Var_floatariths_Max image_image Max_Image_distrib) lemma max_Var_floatariths_map_uminus[simp]: "max_Var_floatariths (map (\i. - fa1 i) xs) = max_Var_floatariths (map fa1 xs)" by (auto simp: max_Var_floatariths_Max image_image Max_Image_distrib) lemma max_Var_floatariths_map_const[simp]: "max_Var_floatariths (map (\i. fa) xs) = (if xs = [] then 0 else max_Var_floatarith fa)" by (auto simp: max_Var_floatariths_Max image_image image_constant_conv) lemma max_Var_floatariths_map_minus[simp]: "max_Var_floatariths (map (\i. fa1 i - fa2 i) xs) = max (max_Var_floatariths (map fa1 xs)) (max_Var_floatariths (map fa2 xs))" by (auto simp: max_Var_floatariths_Max image_image Max_Image_distrib) primrec fresh_floatarith where "fresh_floatarith (Var y) x \ (x \ y)" | "fresh_floatarith (Num a) x \ True" | "fresh_floatarith Pi x \ True" | "fresh_floatarith (Cos a) x \ fresh_floatarith a x" | "fresh_floatarith (Abs a) x \ fresh_floatarith a x" | "fresh_floatarith (Arctan a) x \ fresh_floatarith a x" | "fresh_floatarith (Sqrt a) x \ fresh_floatarith a x" | "fresh_floatarith (Exp a) x \ fresh_floatarith a x" | "fresh_floatarith (Floor a) x \ fresh_floatarith a x" | "fresh_floatarith (Power a n) x \ fresh_floatarith a x" | "fresh_floatarith (Minus a) x \ fresh_floatarith a x" | "fresh_floatarith (Ln a) x \ fresh_floatarith a x" | "fresh_floatarith (Inverse a) x \ fresh_floatarith a x" | "fresh_floatarith (Add a b) x \ fresh_floatarith a x \ fresh_floatarith b x" | "fresh_floatarith (Mult a b) x \ fresh_floatarith a x \ fresh_floatarith b x" | "fresh_floatarith (Max a b) x \ fresh_floatarith a x \ fresh_floatarith b x" | "fresh_floatarith (Min a b) x \ fresh_floatarith a x \ fresh_floatarith b x" | "fresh_floatarith (Powr a b) x \ fresh_floatarith a x \ fresh_floatarith b x" lemma fresh_floatarith_subst: fixes v::float assumes "fresh_floatarith e x" assumes "x < length vs" shows "interpret_floatarith e (vs[x:=v]) = interpret_floatarith e vs" using assms by (induction e) (auto simp: map_update) lemma fresh_floatarith_max_Var: assumes "max_Var_floatarith ea \ i" shows "fresh_floatarith ea i" using assms by (induction ea) auto primrec fresh_floatariths where "fresh_floatariths [] x \ True" | "fresh_floatariths (a#as) x \ fresh_floatarith a x \ fresh_floatariths as x" lemma fresh_floatariths_max_Var: assumes "max_Var_floatariths ea \ i" shows "fresh_floatariths ea i" using assms by (induction ea) (auto simp: fresh_floatarith_max_Var) lemma interpret_floatariths_take_eqI: assumes "take n ys = take n zs" assumes "max_Var_floatariths ea \ n" shows "interpret_floatariths ea ys = interpret_floatariths ea zs" by (rule interpret_floatariths_eq_take_max_VarI) (rule take_greater_eqI[OF assms]) lemma interpret_floatarith_fresh_eqI: assumes "\i. fresh_floatarith ea i \ (i < length ys \ i < length zs \ ys ! i = zs ! i)" shows "interpret_floatarith ea ys = interpret_floatarith ea zs" using assms by (induction ea) force+ lemma interpret_floatariths_fresh_eqI: assumes "\i. fresh_floatariths ea i \ (i < length ys \ i < length zs \ ys ! i = zs ! i)" shows "interpret_floatariths ea ys = interpret_floatariths ea zs" using assms apply (induction ea) subgoal by (force simp: interpret_floatarith_fresh_eqI intro: interpret_floatarith_fresh_eqI) subgoal for e ea apply clarsimp apply (auto simp: list_eq_iff_nth_eq) using interpret_floatarith_fresh_eqI by blast done lemma interpret_floatarith_max_Var_cong: assumes "\i. i < max_Var_floatarith f \ xs ! i = ys ! i" shows "interpret_floatarith f ys = interpret_floatarith f xs" using assms by (induction f) auto lemma interpret_floatarith_fresh_cong: assumes "\i. \fresh_floatarith f i \ xs ! i = ys ! i" shows "interpret_floatarith f ys = interpret_floatarith f xs" using assms by (induction f) auto lemma max_Var_floatarith_le_max_Var_floatariths: "fa \ set fas \ max_Var_floatarith fa \ max_Var_floatariths fas" by (induction fas) (auto simp: nth_Cons max_def split: nat.splits) lemma max_Var_floatarith_le_max_Var_floatariths_nth: "n < length fas \ max_Var_floatarith (fas ! n) \ max_Var_floatariths fas" by (rule max_Var_floatarith_le_max_Var_floatariths) auto lemma max_Var_floatariths_leI: assumes "\i. i < length xs \ max_Var_floatarith (xs ! i) \ F" shows "max_Var_floatariths xs \ F" using assms by (auto simp: max_Var_floatariths_Max in_set_conv_nth) lemma fresh_floatariths_map_Var[simp]: "fresh_floatariths (map floatarith.Var xs) i \ i \ set xs" by (induction xs) auto lemma max_Var_floatarith_fold_const_fa: "max_Var_floatarith (fold_const_fa fa) \ max_Var_floatarith fa" by (induction fa) (auto simp: fold_const_fa.simps split!: option.splits floatarith.splits) lemma max_Var_floatariths_fold_const_fa: "max_Var_floatariths (map fold_const_fa xs) \ max_Var_floatariths xs" by (auto simp: intro!: max_Var_floatariths_leI max_Var_floatarith_le_max_Var_floatariths_nth max_Var_floatarith_fold_const_fa[THEN order_trans]) lemma interpret_form_max_Var_cong: assumes "\i. i < max_Var_form f \ xs ! i = ys ! i" shows "interpret_form f xs = interpret_form f ys" using assms by (induction f) (auto simp: interpret_floatarith_max_Var_cong[where xs=xs and ys=ys]) lemma max_Var_floatariths_lessI: "i < max_Var_floatarith (fas ! j) \ j < length fas \ i < max_Var_floatariths fas" by (metis leD le_trans less_le max_Var_floatarith_le_max_Var_floatariths nth_mem) lemma interpret_floatariths_max_Var_cong: assumes "\i. i < max_Var_floatariths f \ xs ! i = ys ! i" shows "interpret_floatariths f ys = interpret_floatariths f xs" by (auto intro!: nth_equalityI interpret_floatarith_max_Var_cong assms max_Var_floatariths_lessI) lemma max_Var_floatarithimage_Var[simp]: "max_Var_floatarith ` Var ` X = Suc ` X" by force lemma max_Var_floatariths_map_Var[simp]: "max_Var_floatariths (map Var xs) = (if xs = [] then 0 else Suc (linorder_class.Max (set xs)))" by (auto simp: max_Var_floatariths_Max hom_Max_commute split: if_splits) lemma Max_atLeastLessThan_nat[simp]: "a < b \ linorder_class.Max {a..Derivatives\ lemma isDERIV_Power_iff: "isDERIV j (Power fa n) xs = (if n = 0 then True else isDERIV j fa xs)" by (cases n) auto lemma isDERIV_max_Var_floatarithI: assumes "isDERIV n f ys" assumes "\i. i < max_Var_floatarith f \ xs ! i = ys ! i" shows "isDERIV n f xs" using assms proof (induction f) case (Power f n) then show ?case by (cases n) auto qed (auto simp: max_def interpret_floatarith_max_Var_cong[of _ xs ys] split: if_splits) definition isFDERIV where "isFDERIV n xs fas vs \ (\ij length fas = n \ length xs = n" lemma isFDERIV_I: "(\i j. i < n \ j < n \ isDERIV (xs ! i) (fas ! j) vs) \ length fas = n \ length xs = n \ isFDERIV n xs fas vs" by (auto simp: isFDERIV_def) lemma isFDERIV_isDERIV_D: "isFDERIV n xs fas vs \ i < n \ j < n \ isDERIV (xs ! i) (fas ! j) vs" by (auto simp: isFDERIV_def) lemma isFDERIV_lengthD: "length fas = n" "length xs = n" if "isFDERIV n xs fas vs" using that by (auto simp: isFDERIV_def) lemma isFDERIV_uptD: "isFDERIV n [0.. i < n \ j < n \ isDERIV i (fas ! j) vs" by (auto simp: isFDERIV_def) lemma isFDERIV_max_Var_congI: "isFDERIV n xs fas ws" if f: "isFDERIV n xs fas vs" and c: "(\i. i < max_Var_floatariths fas \ vs ! i = ws ! i)" using c f by (auto simp: isFDERIV_def max_Var_floatariths_lessI intro!: isFDERIV_I isDERIV_max_Var_floatarithI[OF isFDERIV_isDERIV_D[OF f]]) lemma isFDERIV_max_Var_cong: "isFDERIV n xs fas ws \ isFDERIV n xs fas vs" if c: "(\i. i < max_Var_floatariths fas \ vs ! i = ws ! i)" using c by (auto intro: isFDERIV_max_Var_congI) lemma isDERIV_max_VarI: "i \ max_Var_floatarith fa \ isDERIV j fa xs \ isDERIV i fa xs" by (induction fa) (auto simp: isDERIV_Power_iff) lemmas max_Var_floatarith_le_max_Var_floatariths_nthI = max_Var_floatarith_le_max_Var_floatariths_nth[THEN order_trans] lemma isFDERIV_appendD1: assumes "isFDERIV (J + K) [0.. J" shows "isFDERIV J [0.. length xs" shows "map ((!) (xs @ ys)) [a.. 0" shows "map ((!) (x # ys)) [a.. (map ((!) fas) [0..i j. i < J + K \ j < K \ isDERIV i (rs ! j) xs" assumes "length es = J" assumes "length rs = K" assumes "max_Var_floatariths es \ J" shows "isFDERIV (J + K) [0..R (A o\<^sub>L B) = i *\<^sub>R A o\<^sub>L B" and scaleR_blinfun_compose_right: "i *\<^sub>R (A o\<^sub>L B) = A o\<^sub>L i *\<^sub>R B" by (auto intro!: blinfun_eqI simp: blinfun.bilinear_simps) lemma matrix_blinfun_compose: fixes A B::"(real ^ 'n) \\<^sub>L (real ^ 'n)" shows "matrix (A o\<^sub>L B) = (matrix A) ** (matrix B)" by transfer (auto simp: matrix_compose linear_linear) lemma matrix_add_rdistrib: "((B + C) ** A) = (B ** A) + (C ** A)" by (vector matrix_matrix_mult_def sum.distrib[symmetric] field_simps) lemma matrix_scaleR_right: "r *\<^sub>R (a::'a::real_algebra_1^'n^'m) ** b = r *\<^sub>R (a ** b)" by (vector matrix_matrix_mult_def algebra_simps scaleR_sum_right) lemma matrix_scaleR_left: "(a::'a::real_algebra_1^'n^'m) ** r *\<^sub>R b = r *\<^sub>R (a ** b)" by (vector matrix_matrix_mult_def algebra_simps scaleR_sum_right) lemma bounded_bilinear_matrix_matrix_mult[bounded_bilinear]: "bounded_bilinear ((**):: ('a::{euclidean_space, real_normed_algebra_1}^'n^'m) \ ('a::{euclidean_space, real_normed_algebra_1}^'p^'n) \ ('a::{euclidean_space, real_normed_algebra_1}^'p^'m))" unfolding bilinear_conv_bounded_bilinear[symmetric] unfolding bilinear_def apply safe by unfold_locales (auto simp: matrix_add_ldistrib matrix_add_rdistrib matrix_scaleR_right matrix_scaleR_left) lemma norm_axis: "norm (axis ia 1::'a::{real_normed_algebra_1}^'n) = 1" by (auto simp: axis_def norm_vec_def L2_set_def if_distrib if_distribR sum.delta cong: if_cong) lemma abs_vec_nth_blinfun_apply_lemma: fixes x::"(real^'n) \\<^sub>L (real^'m)" shows "abs (vec_nth (blinfun_apply x (axis ia 1)) i) \ norm x" apply (rule component_le_norm_cart[THEN order_trans]) apply (rule norm_blinfun[THEN order_trans]) by (auto simp: norm_axis) lemma bounded_linear_matrix_blinfun_apply: "bounded_linear (\x::(real^'n) \\<^sub>L (real^'m). matrix (blinfun_apply x))" apply standard subgoal by (vector blinfun.bilinear_simps matrix_def) subgoal by (vector blinfun.bilinear_simps matrix_def) apply (rule exI[where x="real (CARD('n) * CARD('m))"]) apply (auto simp: matrix_def) apply (subst norm_vec_def) apply (rule L2_set_le_sum[THEN order_trans]) apply simp apply auto apply (rule sum_mono[THEN order_trans]) apply (subst norm_vec_def) apply (rule L2_set_le_sum) apply simp apply (rule sum_mono[THEN order_trans]) apply (rule sum_mono) apply simp apply (rule abs_vec_nth_blinfun_apply_lemma) apply (simp add: abs_vec_nth_blinfun_apply_lemma) done lemma matrix_has_derivative: shows "((\x::(real^'n)\\<^sub>L(real^'n). matrix (blinfun_apply x)) has_derivative (\h. matrix (blinfun_apply h))) (at x)" apply (auto simp: has_derivative_at2) unfolding linear_linear subgoal by (rule bounded_linear_matrix_blinfun_apply) subgoal by (auto simp: blinfun.bilinear_simps matrix_def) vector done lemma matrix_comp_has_derivative[derivative_intros]: fixes f::"'a::real_normed_vector \ ((real^'n)\\<^sub>L(real^'n))" assumes "(f has_derivative f') (at x within S)" shows "((\x. matrix (blinfun_apply (f x))) has_derivative (\x. matrix (f' x))) (at x within S)" using has_derivative_compose[OF assms matrix_has_derivative] by auto fun inner_floatariths where "inner_floatariths [] _ = Num 0" | "inner_floatariths _ [] = Num 0" | "inner_floatariths (x#xs) (y#ys) = Add (Mult x y) (inner_floatariths xs ys)" lemma interpret_floatarith_inner_eq: assumes "length xs = length ys" shows "interpret_floatarith (inner_floatariths xs ys) vs = (\i y" using assms by (subst euclidean_inner) (auto simp: interpret_floatarith_inner_eq sum_Basis_sum_nth_Basis_list eucl_of_list_inner index_nth_id intro!: euclidean_eqI[where 'a='a] sum.cong) lemma max_Var_floatarith_inner_floatariths[simp]: assumes "length f = length g" shows "max_Var_floatarith (inner_floatariths f g) = max (max_Var_floatariths f) (max_Var_floatariths g)" using assms by (induction f g rule: list_induct2) auto definition FDERIV_floatarith where "FDERIV_floatarith fa xs d = inner_floatariths (map (\x. fold_const_fa (DERIV_floatarith x fa)) xs) d" \ \TODO: specialize to \FDERIV_floatarith fa [0.. and do the rest with @{term subst_floatarith}? TODO: introduce approximation on type @{typ "real^'i^'j"} and use @{term jacobian}?\ lemma interpret_floatariths_map: "interpret_floatariths (map f xs) vs = map (\x. interpret_floatarith (f x) vs) xs" by (induct xs) (auto simp: ) lemma max_Var_floatarith_DERIV_floatarith: "max_Var_floatarith (DERIV_floatarith x fa) \ max_Var_floatarith fa" by (induction x fa rule: DERIV_floatarith.induct) (auto) lemma max_Var_floatarith_FDERIV_floatarith: "length xs = length d \ max_Var_floatarith (FDERIV_floatarith fa xs d) \ max (max_Var_floatarith fa) (max_Var_floatariths d)" unfolding FDERIV_floatarith_def by (auto simp: max_Var_floatariths_Max intro!: max_Var_floatarith_DERIV_floatarith[THEN order_trans] max_Var_floatarith_fold_const_fa[THEN order_trans]) definition FDERIV_floatariths where "FDERIV_floatariths fas xs d = map (\fa. FDERIV_floatarith fa xs d) fas" lemma max_Var_floatarith_FDERIV_floatariths: "length xs = length d \ max_Var_floatariths (FDERIV_floatariths fa xs d) \ max (max_Var_floatariths fa) (max_Var_floatariths d)" by (auto simp: FDERIV_floatariths_def max_Var_floatariths_Max intro!: max_Var_floatarith_FDERIV_floatarith[THEN order_trans]) (auto simp: max_def) lemma length_FDERIV_floatariths[simp]: "length (FDERIV_floatariths fas xs ds) = length fas" by (auto simp: FDERIV_floatariths_def) lemma FDERIV_floatariths_nth[simp]: "i < length fas \ FDERIV_floatariths fas xs ds ! i = FDERIV_floatarith (fas ! i) xs ds" by (auto simp: FDERIV_floatariths_def) definition "FDERIV_n_floatariths fas xs ds n = ((\x. FDERIV_floatariths x xs ds)^^n) fas" lemma FDERIV_n_floatariths_Suc[simp]: "FDERIV_n_floatariths fa xs ds 0 = fa" "FDERIV_n_floatariths fa xs ds (Suc n) = FDERIV_floatariths (FDERIV_n_floatariths fa xs ds n) xs ds" by (auto simp: FDERIV_n_floatariths_def) lemma length_FDERIV_n_floatariths[simp]: "length (FDERIV_n_floatariths fa xs ds n) = length fa" by (induction n) (auto simp: FDERIV_n_floatariths_def) lemma max_Var_floatarith_FDERIV_n_floatariths: "length xs = length d \ max_Var_floatariths (FDERIV_n_floatariths fa xs d n) \ max (max_Var_floatariths fa) (max_Var_floatariths d)" by (induction n) (auto intro!: max_Var_floatarith_FDERIV_floatariths[THEN order_trans] simp: FDERIV_n_floatariths_def) lemma interpret_floatarith_FDERIV_floatarith_cong: assumes rq: "\i. i < max_Var_floatarith f \ rs ! i = qs ! i" assumes [simp]: "length ds = length xs" "length es = length xs" assumes "interpret_floatariths ds qs = interpret_floatariths es rs" shows "interpret_floatarith (FDERIV_floatarith f xs ds) qs = interpret_floatarith (FDERIV_floatarith f xs es) rs" apply (auto simp: FDERIV_floatarith_def interpret_floatarith_inner_eq) apply (rule sum.cong[OF refl]) subgoal premises prems for i proof - have "interpret_floatarith (DERIV_floatarith (xs ! i) f) qs = interpret_floatarith (DERIV_floatarith (xs ! i) f) rs" apply (rule interpret_floatarith_max_Var_cong) apply (auto simp: intro!: rq) by (metis leD le_trans max_Var_floatarith_DERIV_floatarith nat_less_le) moreover have "interpret_floatarith (ds ! i) qs = interpret_floatarith (es ! i) rs" using assms by (metis \i \ {.. interpret_floatariths_nth lessThan_iff) ultimately show ?thesis by auto qed done theorem matrix_vector_mult_eq_list_of_eucl_nth: "(M::real^'n::enum^'m::enum) *v v = (\ijR Basis_list ! i)" using eucl_of_list_matrix_vector_mult_eq_sum_nth_Basis_list[of "list_of_eucl M" "list_of_eucl v", where 'n='n and 'm = 'm] by auto definition "mmult_fa l m n AS BS = concat (map (\i. map (\k. inner_floatariths (map (\j. AS ! (i * m + j)) [0..j. BS ! (j * n + k)) [0.. max (max_Var_floatariths A) (max_Var_floatariths B)" apply (auto simp: mmult_fa_def concat_map_map_index intro!: max_Var_floatariths_leI) apply (rule max.coboundedI1) apply (auto intro!: max_Var_floatarith_le_max_Var_floatariths_nth max.coboundedI2) apply (cases "F = 0") apply simp_all done lemma isDERIV_inner_iff: assumes "length xs = length ys" shows "isDERIV i (inner_floatariths xs ys) vs \ (\k < length xs. isDERIV i (xs ! k) vs) \ (\k < length ys. isDERIV i (ys ! k) vs)" using assms by (induction xs ys rule: list_induct2) (auto simp: nth_Cons split: nat.splits) lemma isDERIV_Power: "isDERIV x (fa) vs \ isDERIV x (fa ^\<^sub>e n) vs" by (induction n) (auto simp: FDERIV_floatarith_def isDERIV_inner_iff) lemma isDERIV_mmult_fa_nth: assumes "\j. j < D * E \ isDERIV i (A ! j) xs" assumes "\j. j < E * F \ isDERIV i (B ! j) xs" assumes [simp]: "length A = D * E" "length B = E * F" "j < D * F" shows "isDERIV i (mmult_fa D E F A B ! j) xs" using assms apply (cases "F = 0") apply (auto simp: mmult_fa_def concat_map_map_index isDERIV_inner_iff ac_simps) apply (metis add.commute assms(5) in_square_lemma less_square_imp_div_less mult.commute) done definition "mvmult_fa n m AS B = map (\i. inner_floatariths (map (\j. AS ! (i * m + j)) [0..j. B ! j) [0.. max (max_Var_floatariths A) (max_Var_floatariths B)" apply (auto simp: mvmult_fa_def concat_map_map_index intro!: max_Var_floatariths_leI) apply (rule max.coboundedI1) by (auto intro!: max_Var_floatarith_le_max_Var_floatariths_nth max.coboundedI2) lemma isDERIV_mvmult_fa_nth: assumes "\j. j < D * E \ isDERIV i (A ! j) xs" assumes "\j. j < E \ isDERIV i (B ! j) xs" assumes [simp]: "length A = D * E" "length B = E" "j < D" shows "isDERIV i (mvmult_fa D E A B ! j) xs" using assms apply (auto simp: mvmult_fa_def concat_map_map_index isDERIV_inner_iff ac_simps) by (metis assms(5) in_square_lemma semiring_normalization_rules(24) semiring_normalization_rules(7)) lemma max_Var_floatariths_mapI: assumes "\x. x \ set xs \ max_Var_floatarith (f x) \ m" shows "max_Var_floatariths (map f xs) \ m" using assms by (force intro!: max_Var_floatariths_leI simp: in_set_conv_nth) lemma max_Var_floatariths_list_updateI: assumes "max_Var_floatariths xs \ m" assumes "max_Var_floatarith v \ m" assumes "i < length xs" shows "max_Var_floatariths (xs[i := v]) \ m" using assms apply (auto simp: nth_list_update intro!: max_Var_floatariths_leI ) using max_Var_floatarith_le_max_Var_floatariths_nthI by blast lemma max_Var_floatariths_replicateI: assumes "max_Var_floatarith v \ m" shows "max_Var_floatariths (replicate n v) \ m" using assms by (auto intro!: max_Var_floatariths_leI ) definition "FDERIV_n_floatarith fa xs ds n = ((\x. FDERIV_floatarith x xs ds)^^n) fa" lemma FDERIV_n_floatariths_nth: "i < length fas \ FDERIV_n_floatariths fas xs ds n ! i = FDERIV_n_floatarith (fas ! i) xs ds n" by (induction n) (auto simp: FDERIV_n_floatarith_def FDERIV_floatariths_nth) lemma einterpret_fold_const_fa[simp]: "(einterpret (map (\i. fold_const_fa (fa i)) xs) vs::'a::executable_euclidean_space) = einterpret (map fa xs) vs" if "length xs = DIM('a)" using that by (auto intro!: euclidean_eqI[where 'a='a] simp: algebra_simps eucl_of_list_inner) lemma einterpret_plus[simp]: shows "(einterpret (map (\i. fa1 i + fa2 i) [0..i. fa1 * fa2 i) [0..R einterpret (map fa2 [0.. (Basis_list ! n)" proof - have "interpret_floatarith (mvmult_fa D E xs ys ! n) vs = einterpret (mvmult_fa D E xs ys) vs \ (Basis_list ! n::'n rvec)" using assms by (auto simp: eucl_of_list_inner) also from einterpret_mvmult_fa[OF assms(1,2), of xs ys vs] have "einterpret (mvmult_fa D E xs ys) vs = (einterpret xs vs::((real, 'm) vec, 'n) vec) *v einterpret ys vs" using assms by simp finally show ?thesis by simp qed lemmas [simp del] = fold_const_fa.simps lemma take_eq_map_nth: "n < length xs \ take n xs = map ((!) xs) [0..Definition of Approximating Function using Affine Arithmetic\ lemma interpret_Floatreal: "interpret_floatarith (floatarith.Num f) vs = (real_of_float f)" by simp ML \ (* Make a congruence rule out of a defining equation for the interpretation th is one defining equation of f, i.e. th is "f (Cp ?t1 ... ?tn) = P(f ?t1, .., f ?tn)" Cp is a constructor pattern and P is a pattern The result is: [|?A1 = f ?t1 ; .. ; ?An= f ?tn |] ==> P (?A1, .., ?An) = f (Cp ?t1 .. ?tn) + the a list of names of the A1 .. An, Those are fresh in the ctxt *) fun mk_congeq ctxt fs th = let val Const (fN, _) = th |> Thm.prop_of |> HOLogic.dest_Trueprop |> HOLogic.dest_eq |> fst |> strip_comb |> fst; val ((_, [th']), ctxt') = Variable.import true [th] ctxt; val (lhs, rhs) = HOLogic.dest_eq (HOLogic.dest_Trueprop (Thm.prop_of th')); fun add_fterms (t as t1 $ t2) = if exists (fn f => Term.could_unify (t |> strip_comb |> fst, f)) fs then insert (op aconv) t else add_fterms t1 #> add_fterms t2 | add_fterms (t as Abs _) = if exists_Const (fn (c, _) => c = fN) t then K [t] else K [] | add_fterms _ = I; val fterms = add_fterms rhs []; val (xs, ctxt'') = Variable.variant_fixes (replicate (length fterms) "x") ctxt'; val tys = map fastype_of fterms; val vs = map Free (xs ~~ tys); val env = fterms ~~ vs; (*FIXME*) fun replace_fterms (t as t1 $ t2) = (case AList.lookup (op aconv) env t of SOME v => v | NONE => replace_fterms t1 $ replace_fterms t2) | replace_fterms t = (case AList.lookup (op aconv) env t of SOME v => v | NONE => t); fun mk_def (Abs (x, xT, t), v) = HOLogic.mk_Trueprop (HOLogic.all_const xT $ Abs (x, xT, HOLogic.mk_eq (v $ Bound 0, t))) | mk_def (t, v) = HOLogic.mk_Trueprop (HOLogic.mk_eq (v, t)); fun tryext x = (x RS @{lemma "(\x. f x = g x) \ f = g" by blast} handle THM _ => x); val cong = (Goal.prove ctxt'' [] (map mk_def env) (HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, replace_fterms rhs))) (fn {context, prems, ...} => Local_Defs.unfold0_tac context (map tryext prems) THEN resolve_tac ctxt'' [th'] 1)) RS sym; val (cong' :: vars') = Variable.export ctxt'' ctxt (cong :: map (Drule.mk_term o Thm.cterm_of ctxt'') vs); val vs' = map (fst o fst o Term.dest_Var o Thm.term_of o Drule.dest_term) vars'; in (vs', cong') end; fun mk_congs ctxt eqs = let val fs = fold_rev (fn eq => insert (op =) (eq |> Thm.prop_of |> HOLogic.dest_Trueprop |> HOLogic.dest_eq |> fst |> strip_comb |> fst)) eqs []; val tys = fold_rev (fn f => fold (insert (op =)) (f |> fastype_of |> binder_types |> tl)) fs []; val (vs, ctxt') = Variable.variant_fixes (replicate (length tys) "vs") ctxt; val subst = the o AList.lookup (op =) (map2 (fn T => fn v => (T, Thm.cterm_of ctxt' (Free (v, T)))) tys vs); fun prep_eq eq = let val (_, _ :: vs) = eq |> Thm.prop_of |> HOLogic.dest_Trueprop |> HOLogic.dest_eq |> fst |> strip_comb; val subst = map_filter (fn Var v => SOME (v, subst (#2 v)) | _ => NONE) vs; - in Thm.instantiate ([], subst) eq end; + in Thm.instantiate (TVars.empty, Vars.make subst) eq end; val (ps, congs) = map_split (mk_congeq ctxt' fs o prep_eq) eqs; val bds = AList.make (K ([], [])) tys; in (ps ~~ Variable.export ctxt' ctxt congs, bds) end \ ML \ fun interpret_floatariths_congs ctxt = mk_congs ctxt @{thms interpret_floatarith.simps interpret_floatariths.simps} |> fst |> map snd \ ML \ fun preproc_form_conv ctxt = Simplifier.rewrite (put_simpset HOL_basic_ss ctxt addsimps (Named_Theorems.get ctxt @{named_theorems approximation_preproc}))\ ML \fun reify_floatariths_tac ctxt i = CONVERSION (preproc_form_conv ctxt) i THEN REPEAT_ALL_NEW (fn i => resolve_tac ctxt (interpret_floatariths_congs ctxt) i) i\ method_setup reify_floatariths = \ Scan.succeed (fn ctxt => SIMPLE_METHOD' (reify_floatariths_tac ctxt)) \ "reification of floatariths expression" schematic_goal reify_example: "[xs!i * xs!j, xs!i + xs!j powr (sin (xs!0)), xs!k + (2 / 3 * xs!i * xs!j)] = interpret_floatariths ?fas xs" by (reify_floatariths) ML \fun interpret_floatariths_step_tac ctxt i = resolve_tac ctxt (interpret_floatariths_congs ctxt) i\ method_setup reify_floatariths_step = \ Scan.succeed (fn ctxt => SIMPLE_METHOD' (interpret_floatariths_step_tac ctxt)) \ "reification of floatariths expression (step)" lemma eucl_of_list_interpret_floatariths_cong: fixes y::"'a::executable_euclidean_space" assumes "\b. b \ Basis \ interpret_floatarith (fa (index Basis_list b)) vs = y \ b" assumes "length xs = DIM('a)" shows "eucl_of_list (interpret_floatariths (map fa [0.. D" shows "interpret_floatarith (subst_floatarith s fa) vs = interpret_floatarith fa (map (\i. interpret_floatarith (s i) vs) [0.. max_Var_floatarith fa" shows "max_Var_floatarith (subst_floatarith ((!) xs) fa) \ max_Var_floatariths xs" using assms by (induction fa) (auto intro!: max_Var_floatarith_le_max_Var_floatariths_nth) lemma max_Var_floatariths_subst_floatarith_le[THEN order_trans]: assumes "length xs \ max_Var_floatariths fas" shows "max_Var_floatariths (map (subst_floatarith ((!) xs)) fas) \ max_Var_floatariths xs" using assms by (induction fas) (auto simp: max_Var_floatarith_subst_floatarith_le) fun continuous_on_floatarith :: "floatarith \ bool" where "continuous_on_floatarith (Add a b) = (continuous_on_floatarith a \ continuous_on_floatarith b)" | "continuous_on_floatarith (Mult a b) = (continuous_on_floatarith a \ continuous_on_floatarith b)" | "continuous_on_floatarith (Minus a) = continuous_on_floatarith a" | "continuous_on_floatarith (Inverse a) = False" | "continuous_on_floatarith (Cos a) = continuous_on_floatarith a" | "continuous_on_floatarith (Arctan a) = continuous_on_floatarith a" | "continuous_on_floatarith (Min a b) = (continuous_on_floatarith a \ continuous_on_floatarith b)" | "continuous_on_floatarith (Max a b) = (continuous_on_floatarith a \ continuous_on_floatarith b)" | "continuous_on_floatarith (Abs a) = continuous_on_floatarith a" | "continuous_on_floatarith Pi = True" | "continuous_on_floatarith (Sqrt a) = False" | "continuous_on_floatarith (Exp a) = continuous_on_floatarith a" | "continuous_on_floatarith (Powr a b) = False" | "continuous_on_floatarith (Ln a) = False" | "continuous_on_floatarith (Floor a) = False" | "continuous_on_floatarith (Power a n) = (if n = 0 then True else continuous_on_floatarith a)" | "continuous_on_floatarith (Num f) = True" | "continuous_on_floatarith (Var n) = True" definition "Maxs\<^sub>e xs = fold (\a b. floatarith.Max a b) xs" definition "norm2\<^sub>e n = Maxs\<^sub>e (map (\j. Norm (map (\i. Var (Suc j * n + i)) [0..r l = Num (float_of l)" lemma interpret_floatarith_Norm: "interpret_floatarith (Norm xs) vs = L2_set (\i. interpret_floatarith (xs ! i) vs) {0..r U) vs = real_of_float (float_of U)" by (auto simp: N\<^sub>r_def) fun list_updates where "list_updates [] _ xs = xs" | "list_updates _ [] xs = xs" | "list_updates (i#is) (u#us) xs = list_updates is us (xs[i:=u])" lemma list_updates_nth_notmem: assumes "length xs = length ys" assumes "i \ set xs" shows "list_updates xs ys vs ! i = vs ! i" using assms by (induction xs ys arbitrary: i vs rule: list_induct2) auto lemma list_updates_nth_less: assumes "length xs = length ys" "distinct xs" assumes "i < length vs" shows "list_updates xs ys vs ! i = (if i \ set xs then ys ! (index xs i) else vs ! i)" using assms by (induction xs ys arbitrary: i vs rule: list_induct2) (auto simp: list_updates_nth_notmem) lemma length_list_updates[simp]: "length (list_updates xs ys vs) = length vs" by (induction xs ys vs rule: list_updates.induct) simp_all lemma list_updates_nth_ge[simp]: "x \ length vs \ length xs = length ys \ list_updates xs ys vs ! x = vs ! x" apply (induction xs ys vs rule: list_updates.induct) apply (auto simp: nth_list_update) by (metis list_update_beyond nth_list_update_neq) lemma list_updates_nth: assumes [simp]: "length xs = length ys" "distinct xs" shows "list_updates xs ys vs ! i = (if i < length vs \ i \ set xs then ys ! index xs i else vs ! i)" by (auto simp: list_updates_nth_less list_updates_nth_notmem) lemma list_of_eucl_coord_update: assumes [simp]: "length xs = DIM('a::executable_euclidean_space)" assumes [simp]: "distinct xs" assumes [simp]: "i \ Basis" assumes [simp]: "\n. n \ set xs \ n < length vs" shows "list_updates xs (list_of_eucl (x + (p - x \ i) *\<^sub>R i::'a)) vs = (list_updates xs (list_of_eucl x) vs)[xs ! index Basis_list i := p]" apply (auto intro!: nth_equalityI simp: list_updates_nth nth_list_update) apply (simp add: algebra_simps inner_Basis index_nth_id) apply (auto simp add: algebra_simps inner_Basis index_nth_id) done definition "eucl_of_env is vs = eucl_of_list (map (nth vs) is)" lemma list_updates_list_of_eucl_of_env[simp]: assumes [simp]: "length xs = DIM('a::executable_euclidean_space)" "distinct xs" shows "list_updates xs (list_of_eucl (eucl_of_env xs vs::'a)) vs = vs" by (auto intro!: nth_equalityI simp: list_updates_nth nth_list_update eucl_of_env_def) lemma nth_nth_eucl_of_env_inner: "b \ Basis \ length is = DIM('a) \ vs ! (is ! index Basis_list b) = eucl_of_env is vs \ b" for b::"'a::executable_euclidean_space" by (auto simp: eucl_of_env_def eucl_of_list_inner) lemma list_updates_idem[simp]: assumes "(\i. i \ set X0 \ i < length vs)" shows "(list_updates X0 (map ((!) vs) X0) vs) = vs" using assms by (induction X0) auto lemma pairwise_orthogonal_Basis[intro, simp]: "pairwise orthogonal Basis" by (auto simp: pairwise_alt orthogonal_def inner_Basis) primrec freshs_floatarith where "freshs_floatarith (Var y) x \ (y \ set x)" | "freshs_floatarith (Num a) x \ True" | "freshs_floatarith Pi x \ True" | "freshs_floatarith (Cos a) x \ freshs_floatarith a x" | "freshs_floatarith (Abs a) x \ freshs_floatarith a x" | "freshs_floatarith (Arctan a) x \ freshs_floatarith a x" | "freshs_floatarith (Sqrt a) x \ freshs_floatarith a x" | "freshs_floatarith (Exp a) x \ freshs_floatarith a x" | "freshs_floatarith (Floor a) x \ freshs_floatarith a x" | "freshs_floatarith (Power a n) x \ freshs_floatarith a x" | "freshs_floatarith (Minus a) x \ freshs_floatarith a x" | "freshs_floatarith (Ln a) x \ freshs_floatarith a x" | "freshs_floatarith (Inverse a) x \ freshs_floatarith a x" | "freshs_floatarith (Add a b) x \ freshs_floatarith a x \ freshs_floatarith b x" | "freshs_floatarith (Mult a b) x \ freshs_floatarith a x \ freshs_floatarith b x" | "freshs_floatarith (floatarith.Max a b) x \ freshs_floatarith a x \ freshs_floatarith b x" | "freshs_floatarith (floatarith.Min a b) x \ freshs_floatarith a x \ freshs_floatarith b x" | "freshs_floatarith (Powr a b) x \ freshs_floatarith a x \ freshs_floatarith b x" lemma freshs_floatarith[simp]: assumes "freshs_floatarith fa ds" "length ds = length xs" shows "interpret_floatarith fa (list_updates ds xs vs) = interpret_floatarith fa vs" using assms by (induction fa) (auto simp: list_updates_nth_notmem) lemma freshs_floatarith_max_Var_floatarithI: assumes "\x. x \ set xs \ max_Var_floatarith f \ x" shows "freshs_floatarith f xs" using assms Suc_n_not_le_n by (induction f; force) definition "freshs_floatariths fas xs = (\fa\set fas. freshs_floatarith fa xs)" lemma freshs_floatariths_max_Var_floatarithsI: assumes "\x. x \ set xs \ max_Var_floatariths f \ x" shows "freshs_floatariths f xs" using assms le_trans max_Var_floatarith_le_max_Var_floatariths by (force simp: freshs_floatariths_def intro!: freshs_floatarith_max_Var_floatarithI) lemma freshs_floatariths_freshs_floatarithI: assumes "\fa. fa \ set fas \ freshs_floatarith fa xs" shows "freshs_floatariths fas xs" by (auto simp: freshs_floatariths_def assms) lemma fresh_floatariths_fresh_floatarithI: assumes "freshs_floatariths fas xs" assumes "fa \ set fas" shows "freshs_floatarith fa xs" using assms by (auto simp: freshs_floatariths_def) lemma fresh_floatariths_fresh_floatarith[simp]: "fresh_floatariths (fas) i \ fa \ set fas \ fresh_floatarith fa i" by (induction fas) auto lemma interpret_floatariths_fresh_cong: assumes "\i. \fresh_floatariths f i \ xs ! i = ys ! i" shows "interpret_floatariths f ys = interpret_floatariths f xs" by (auto intro!: nth_equalityI assms interpret_floatarith_fresh_cong simp: ) fun subterms :: "floatarith \ floatarith set" where "subterms (Add a b) = insert (Add a b) (subterms a \ subterms b)" | "subterms (Mult a b) = insert (Mult a b) (subterms a \ subterms b)" | "subterms (Min a b) = insert (Min a b) (subterms a \ subterms b)" | "subterms (floatarith.Max a b) = insert (floatarith.Max a b) (subterms a \ subterms b)" | "subterms (Powr a b) = insert (Powr a b) (subterms a \ subterms b)" | "subterms (Inverse a) = insert (Inverse a) (subterms a)" | "subterms (Cos a) = insert (Cos a) (subterms a)" | "subterms (Arctan a) = insert (Arctan a) (subterms a)" | "subterms (Abs a) = insert (Abs a) (subterms a)" | "subterms (Sqrt a) = insert (Sqrt a) (subterms a)" | "subterms (Exp a) = insert (Exp a) (subterms a)" | "subterms (Ln a) = insert (Ln a) (subterms a)" | "subterms (Power a n) = insert (Power a n) (subterms a)" | "subterms (Floor a) = insert (Floor a) (subterms a)" | "subterms (Minus a) = insert (Minus a) (subterms a)" | "subterms Pi = {Pi}" | "subterms (Var v) = {Var v}" | "subterms (Num n) = {Num n}" lemma subterms_self[simp]: "fa2 \ subterms fa2" by (induction fa2) auto lemma interpret_floatarith_FDERIV_floatarith_eucl_of_env:\ \TODO: cleanup, reduce to DERIV?!\ assumes iD: "\i. i < DIM('a) \ isDERIV (xs ! i) fa vs" assumes ds_fresh: "freshs_floatarith fa ds" assumes [simp]: "length xs = DIM ('a)" "length ds = DIM ('a)" "\i. i \ set xs \ i < length vs" "distinct xs" "\i. i \ set ds \ i < length vs" "distinct ds" shows "((\x::'a::executable_euclidean_space. (interpret_floatarith fa (list_updates xs (list_of_eucl x) vs))) has_derivative (\d. interpret_floatarith (FDERIV_floatarith fa xs (map Var ds)) (list_updates ds (list_of_eucl d) vs) ) ) (at (eucl_of_env xs vs))" using iD ds_fresh proof (induction fa) case (Add fa1 fa2) then show ?case by (auto intro!: derivative_eq_intros simp: FDERIV_floatarith_def interpret_floatarith_inner_floatariths interpret_floatariths_map algebra_simps list_of_eucl_coord_update[symmetric]) next case (Minus fa) then show ?case by (auto intro!: derivative_eq_intros simp: FDERIV_floatarith_def interpret_floatarith_inner_floatariths interpret_floatariths_map algebra_simps list_of_eucl_coord_update[symmetric]) next case (Mult fa1 fa2) then show ?case by (auto intro!: derivative_eq_intros simp: FDERIV_floatarith_def interpret_floatarith_inner_floatariths interpret_floatariths_map algebra_simps list_of_eucl_coord_update[symmetric]) next case (Inverse fa) then show ?case by (force intro!: derivative_eq_intros DIM_positive simp: FDERIV_floatarith_def interpret_floatarith_inner_floatariths interpret_floatariths_map algebra_simps list_of_eucl_coord_update[symmetric] power2_eq_square) next case (Cos fa) then show ?case by (auto intro!: derivative_eq_intros ext simp: FDERIV_floatarith_def interpret_floatarith_inner_floatariths interpret_floatariths_map add.commute minus_sin_cos_eq simp flip: mult_minus_left list_of_eucl_coord_update cos_pi_minus) next case (Arctan fa) then show ?case by (auto intro!: derivative_eq_intros simp: FDERIV_floatarith_def interpret_floatarith_inner_floatariths interpret_floatariths_map algebra_simps list_of_eucl_coord_update[symmetric]) next case (Abs fa) then show ?case by (auto intro!: derivative_eq_intros DIM_positive simp: FDERIV_floatarith_def interpret_floatarith_inner_floatariths interpret_floatariths_map algebra_simps list_of_eucl_coord_update[symmetric] ) next case (Max fa1 fa2) then show ?case by (auto intro!: derivative_eq_intros DIM_positive simp: FDERIV_floatarith_def interpret_floatarith_inner_floatariths interpret_floatariths_map algebra_simps list_of_eucl_coord_update[symmetric] ) next case (Min fa1 fa2) then show ?case by (auto intro!: derivative_eq_intros DIM_positive simp: FDERIV_floatarith_def interpret_floatarith_inner_floatariths interpret_floatariths_map algebra_simps list_of_eucl_coord_update[symmetric] ) next case Pi then show ?case by (auto intro!: derivative_eq_intros DIM_positive simp: FDERIV_floatarith_def interpret_floatarith_inner_floatariths interpret_floatariths_map algebra_simps list_of_eucl_coord_update[symmetric] ) next case (Sqrt fa) then show ?case by (force intro!: derivative_eq_intros DIM_positive simp: FDERIV_floatarith_def interpret_floatarith_inner_floatariths interpret_floatariths_map algebra_simps list_of_eucl_coord_update[symmetric] ) next case (Exp fa) then show ?case by (auto intro!: derivative_eq_intros DIM_positive simp: FDERIV_floatarith_def interpret_floatarith_inner_floatariths interpret_floatariths_map algebra_simps list_of_eucl_coord_update[symmetric] ) next case (Powr fa1 fa2) then show ?case by (force intro!: derivative_eq_intros DIM_positive simp: FDERIV_floatarith_def interpret_floatarith_inner_floatariths interpret_floatariths_map algebra_simps divide_simps list_of_eucl_coord_update[symmetric] ) next case (Ln fa) then show ?case by (force intro!: derivative_eq_intros DIM_positive simp: FDERIV_floatarith_def interpret_floatarith_inner_floatariths interpret_floatariths_map algebra_simps list_of_eucl_coord_update[symmetric] ) next case (Power fa x2a) then show ?case apply (cases x2a) apply (auto intro!: DIM_positive derivative_eq_intros simp: FDERIV_floatarith_def interpret_floatarith_inner_floatariths interpret_floatariths_map algebra_simps list_of_eucl_coord_update[symmetric]) apply (auto intro!: ext simp: ) by (simp add: semiring_normalization_rules(27)) next case (Floor fa) then show ?case by (force intro!: derivative_eq_intros DIM_positive simp: FDERIV_floatarith_def interpret_floatarith_inner_floatariths interpret_floatariths_map algebra_simps list_of_eucl_coord_update[symmetric] ) next case (Var x) then show ?case apply (auto intro!: derivative_eq_intros DIM_positive simp: FDERIV_floatarith_def interpret_floatarith_inner_floatariths interpret_floatariths_map algebra_simps list_of_eucl_coord_update[symmetric] if_distrib) apply (subst list_updates_nth) apply (auto intro!: derivative_eq_intros ext split: if_splits cong: if_cong simp: if_distribR eucl_of_list_if) apply (subst inner_commute) apply (rule arg_cong[where f="\b. a \ b" for a]) apply (auto intro!: euclidean_eqI[where 'a='a] simp: eucl_of_list_inner list_updates_nth index_nth_id) done next case (Num x) then show ?case by (auto intro!: derivative_eq_intros DIM_positive simp: FDERIV_floatarith_def interpret_floatarith_inner_floatariths interpret_floatariths_map algebra_simps list_of_eucl_coord_update[symmetric] ) qed lemma interpret_floatarith_FDERIV_floatarith_append: assumes iD: "\i j. i < DIM('a) \ isDERIV i (fa) (list_of_eucl x @ params)" assumes m: "max_Var_floatarith fa \ DIM('a) + length params" shows "((\x::'a::executable_euclidean_space. interpret_floatarith fa (list_of_eucl x @ params)) has_derivative (\d. interpret_floatarith (FDERIV_floatarith fa [0.. ia < DIM('a) + length params" for ia using less_le_trans m by blast have "((\xa::'a. interpret_floatarith fa (list_updates [0..d. interpret_floatarith (FDERIV_floatarith fa [0.. isDERIV i (fa) (list_of_eucl x)" assumes m: "max_Var_floatarith fa \ DIM('a)" shows "((\x::'a::executable_euclidean_space. interpret_floatarith fa (list_of_eucl x)) has_derivative (\d. interpret_floatarith (FDERIV_floatarith fa [0..i j. i < DIM('a) \ isDERIV i fa (list_of_eucl x @ params)" assumes m: "max_Var_floatarith fa \ DIM('a::executable_euclidean_space) + length params" shows "\i < DIM('a). \\<^sub>F (x::'a) in at x. isDERIV i fa (list_of_eucl x @ params)" using iD m proof (induction fa) case (Inverse fa) then have "\i\<^sub>F x in at x. isDERIV i fa (list_of_eucl x @ params)" by auto moreover have iD: "i < DIM('a) \ isDERIV i fa (list_of_eucl x @ params)" "interpret_floatarith fa (list_of_eucl x @ params) \ 0" for i using Inverse.prems(1)[OF ] by force+ from Inverse have m: "max_Var_floatarith fa \ DIM('a) + length params" by simp from has_derivative_continuous[OF interpret_floatarith_FDERIV_floatarith_append, OF iD(1) m] have "isCont (\x. interpret_floatarith fa (list_of_eucl x @ params)) x" by simp then have "\\<^sub>F x in at x. interpret_floatarith fa (list_of_eucl x @ params) \ 0" using iD(2) tendsto_imp_eventually_ne by (auto simp: isCont_def) ultimately show ?case by (auto elim: eventually_elim2) next case (Sqrt fa) then have "\i\<^sub>F x in at x. isDERIV i fa (list_of_eucl x @ params)" by auto moreover have iD: "i < DIM('a) \ isDERIV i fa (list_of_eucl x @ params)" "interpret_floatarith fa (list_of_eucl x @ params) > 0" for i using Sqrt.prems(1)[OF ] by force+ from Sqrt have m: "max_Var_floatarith fa \ DIM('a) + length params" by simp from has_derivative_continuous[OF interpret_floatarith_FDERIV_floatarith_append, OF iD(1) m] have "isCont (\x. interpret_floatarith fa (list_of_eucl x @ params)) x" by simp then have "\\<^sub>F x in at x. interpret_floatarith fa (list_of_eucl x @ params) > 0" using iD(2) order_tendstoD by (auto simp: isCont_def) ultimately show ?case by (auto elim: eventually_elim2) next case (Powr fa1 fa2) then have "\i\<^sub>F x in at x. isDERIV i fa1 (list_of_eucl x @ params)" "\i\<^sub>F x in at x. isDERIV i fa2 (list_of_eucl x @ params)" by auto moreover have iD: "i < DIM('a) \ isDERIV i fa1 (list_of_eucl x @ params)" "interpret_floatarith fa1 (list_of_eucl x @ params) > 0" for i using Powr.prems(1) by force+ from Powr have m: "max_Var_floatarith fa1 \ DIM('a) + length params" by simp from has_derivative_continuous[OF interpret_floatarith_FDERIV_floatarith_append, OF iD(1) m] have "isCont (\x. interpret_floatarith fa1 (list_of_eucl x @ params)) x" by simp then have "\\<^sub>F x in at x. interpret_floatarith fa1 (list_of_eucl x @ params) > 0" using iD(2) order_tendstoD by (auto simp: isCont_def) ultimately show ?case apply safe subgoal for i apply (safe dest!: spec[of _ i]) subgoal premises prems using prems(1,3,4) by eventually_elim auto done done next case (Ln fa) then have "\i\<^sub>F x in at x. isDERIV i fa (list_of_eucl x @ params)" by auto moreover have iD: "i < DIM('a) \ isDERIV i fa (list_of_eucl x @ params)" "interpret_floatarith fa (list_of_eucl x @ params) > 0" for i using Ln.prems(1)[OF ] by force+ from Ln have m: "max_Var_floatarith fa \ DIM('a) + length params" by simp from has_derivative_continuous[OF interpret_floatarith_FDERIV_floatarith_append, OF iD(1) m] have "isCont (\x. interpret_floatarith fa (list_of_eucl x @ params)) x" by simp then have "\\<^sub>F x in at x. interpret_floatarith fa (list_of_eucl x @ params) > 0" using iD(2) order_tendstoD by (auto simp: isCont_def) ultimately show ?case by (auto elim: eventually_elim2) next case (Power fa m) then show ?case by (cases m) auto next case (Floor fa) then have "\i\<^sub>F x in at x. isDERIV i fa (list_of_eucl x @ params)" by auto moreover have iD: "i < DIM('a) \ isDERIV i fa (list_of_eucl x @ params)" "interpret_floatarith fa (list_of_eucl x @ params) \ \" for i using Floor.prems(1)[OF ] by force+ from Floor have m: "max_Var_floatarith fa \ DIM('a) + length params" by simp from has_derivative_continuous[OF interpret_floatarith_FDERIV_floatarith_append, OF iD(1) m] have cont: "isCont (\x. interpret_floatarith fa (list_of_eucl x @ params)) x" by simp let ?i = "\x. interpret_floatarith fa (list_of_eucl x @ params)" have "\\<^sub>F y in at x. ?i y > floor (?i x)" "\\<^sub>F y in at x. ?i y < ceiling (?i x)" using cont by (auto simp: isCont_def eventually_floor_less eventually_less_ceiling iD(2)) then have "\\<^sub>F x in at x. ?i x \ \" apply eventually_elim apply (auto simp: Ints_def) by linarith ultimately show ?case by (auto elim: eventually_elim2) qed (fastforce intro: DIM_positive elim: eventually_elim2)+ lemma eventually_isFDERIV: assumes iD: "isFDERIV DIM('a) [0.. DIM('a::executable_euclidean_space) + length params" shows "\\<^sub>F (x::'a) in at x. isFDERIV DIM('a) [0.. DIM('a::executable_euclidean_space) + length params" shows "\\<^sub>F (x::'a) in at x. isFDERIV DIM('a) [0..auto simp: isFDERIV_def\) lemma interpret_floatarith_FDERIV_floatariths_eucl_of_env: assumes iD: "isFDERIV DIM('a) xs fas vs" assumes fresh: "freshs_floatariths (fas) ds" assumes [simp]: "length ds = DIM ('a)" "\i. i \ set xs \ i < length vs" "distinct xs" "\i. i \ set ds \ i < length vs" "distinct ds" shows "((\x::'a::executable_euclidean_space. eucl_of_list (interpret_floatariths fas (list_updates xs (list_of_eucl x) vs))::'a) has_derivative (\d. eucl_of_list (interpret_floatariths (FDERIV_floatariths fas xs (map Var ds)) (list_updates ds (list_of_eucl d) vs)))) (at (eucl_of_env xs vs))" by (subst has_derivative_componentwise_within) (auto simp add: eucl_of_list_inner isFDERIV_lengthD[OF iD] intro!: interpret_floatarith_FDERIV_floatarith_eucl_of_env iD[THEN isFDERIV_isDERIV_D] fresh_floatariths_fresh_floatarithI fresh) lemma interpret_floatarith_FDERIV_floatariths_append: assumes iD: "isFDERIV DIM('a) [0.. DIM('a) + length ramsch" assumes [simp]: "length fas = DIM('a)" shows "((\x::'a::executable_euclidean_space. eucl_of_list (interpret_floatariths fas (list_of_eucl x@ramsch))::'a) has_derivative (\d. eucl_of_list (interpret_floatariths (FDERIV_floatariths fas [0.. ia < DIM('a) + length ramsch" for ia using m by simp have m_nth': "ia < max_Var_floatarith (fas ! j) \ ia < DIM('a) + length ramsch" if "j < DIM('a)" for j ia using m_nth max_Var_floatariths_lessI that by auto have "((\xa::'a. eucl_of_list (interpret_floatariths fas (list_updates [0..d. eucl_of_list (interpret_floatariths (FDERIV_floatariths fas [0.. DIM('a)" assumes [simp]: "length fas = DIM('a)" shows "((\x::'a::executable_euclidean_space. eucl_of_list (interpret_floatariths fas (list_of_eucl x))::'a) has_derivative (\d. eucl_of_list (interpret_floatariths (FDERIV_floatariths fas [0.. 'b::linorder_topology" shows "continuous_on A f \ continuous_on A g \ continuous_on A (\x. min (f x) (g x))" by (auto simp: continuous_on_def intro!: tendsto_min) lemmas [continuous_intros] = continuous_on_max lemma continuous_on_if_const[continuous_intros]: "continuous_on s f \ continuous_on s g \ continuous_on s (\x. if p then f x else g x)" by (cases p) auto lemma continuous_on_floatarith: assumes "continuous_on_floatarith fa" "length xs = DIM('a)" "distinct xs" shows "continuous_on UNIV (\x. interpret_floatarith fa (list_updates xs (list_of_eucl (x::'a::executable_euclidean_space)) vs))" using assms by (induction fa) (auto intro!: continuous_intros split: if_splits simp: list_updates_nth list_of_eucl_nth_if) fun open_form :: "form \ bool" where "open_form (Bound x a b f) = False" | "open_form (Assign x a f) = False" | "open_form (Less a b) \ continuous_on_floatarith a \ continuous_on_floatarith b" | "open_form (LessEqual a b) = False" | "open_form (AtLeastAtMost x a b) = False" | "open_form (Conj f g) \ open_form f \ open_form g" | "open_form (Disj f g) \ open_form f \ open_form g" lemma open_form: assumes "open_form f" "length xs = DIM('a::executable_euclidean_space)" "distinct xs" shows "open (Collect (\x::'a. interpret_form f (list_updates xs (list_of_eucl x) vs)))" using assms by (induction f) (auto intro!: open_Collect_less continuous_on_floatarith open_Collect_conj open_Collect_disj) primrec isnFDERIV where "isnFDERIV N fas xs ds vs 0 = True" | "isnFDERIV N fas xs ds vs (Suc n) \ isFDERIV N xs (FDERIV_n_floatariths fas xs (map Var ds) n) vs \ isnFDERIV N fas xs ds vs n" lemma one_add_square_eq_0: "1 + (x)\<^sup>2 \ (0::real)" by (sos "((R<1 + (([~1] * A=0) + (R<1 * (R<1 * [x]^2)))))") lemma isDERIV_fold_const_fa[intro]: assumes "isDERIV x fa vs" shows "isDERIV x (fold_const_fa fa) vs" using assms apply (induction fa) subgoal by (auto simp: fold_const_fa.simps split: floatarith.splits option.splits) subgoal by (auto simp: fold_const_fa.simps split: floatarith.splits) subgoal by (auto simp: fold_const_fa.simps split: floatarith.splits option.splits) subgoal by (auto simp: fold_const_fa.simps split: floatarith.splits) subgoal by (auto simp: fold_const_fa.simps split: floatarith.splits) subgoal by (auto simp: fold_const_fa.simps split: floatarith.splits) subgoal by (auto simp: fold_const_fa.simps split: floatarith.splits) subgoal by (auto simp: fold_const_fa.simps split: floatarith.splits) subgoal by (auto simp: fold_const_fa.simps split: floatarith.splits) subgoal by (auto simp: fold_const_fa.simps split: floatarith.splits) subgoal by (auto simp: fold_const_fa.simps split: floatarith.splits) subgoal by (auto simp: fold_const_fa.simps split: floatarith.splits) subgoal by (auto simp: fold_const_fa.simps split: floatarith.splits) subgoal by (auto simp: fold_const_fa.simps split: floatarith.splits) subgoal for fa n by (cases n) (auto simp: fold_const_fa.simps split: floatarith.splits nat.splits) subgoal by (auto simp: fold_const_fa.simps split: floatarith.splits) (subst (asm) fold_const_fa[symmetric], force)+ subgoal by (auto simp: fold_const_fa.simps split: floatarith.splits) subgoal by (auto simp: fold_const_fa.simps split: floatarith.splits) done lemma isDERIV_fold_const_fa_minus[intro!]: assumes "isDERIV x (fold_const_fa fa) vs" shows "isDERIV x (fold_const_fa (Minus fa)) vs" using assms by (induction fa) (auto simp: fold_const_fa.simps split: floatarith.splits) lemma isDERIV_fold_const_fa_plus[intro!]: assumes "isDERIV x (fold_const_fa fa) vs" assumes "isDERIV x (fold_const_fa fb) vs" shows "isDERIV x (fold_const_fa (Add fa fb)) vs" using assms by (induction fa) (auto simp: fold_const_fa.simps split: floatarith.splits option.splits) lemma isDERIV_fold_const_fa_mult[intro!]: assumes "isDERIV x (fold_const_fa fa) vs" assumes "isDERIV x (fold_const_fa fb) vs" shows "isDERIV x (fold_const_fa (Mult fa fb)) vs" using assms by (induction fa) (auto simp: fold_const_fa.simps split: floatarith.splits option.splits) lemma isDERIV_fold_const_fa_power[intro!]: assumes "isDERIV x (fold_const_fa fa) vs" shows "isDERIV x (fold_const_fa (fa ^\<^sub>e n)) vs" apply (cases n, simp add: fold_const_fa.simps split: floatarith.splits) using assms by (induction fa) (auto simp: fold_const_fa.simps split: floatarith.splits option.splits) lemma isDERIV_fold_const_fa_inverse[intro!]: assumes "isDERIV x (fold_const_fa fa) vs" assumes "interpret_floatarith fa vs \ 0" shows "isDERIV x (fold_const_fa (Inverse fa)) vs" using assms by (simp add: fold_const_fa.simps) lemma add_square_ne_zero[simp]: "(y::'a::linordered_idom) > 0 \ y + x\<^sup>2 \ 0" by auto (metis less_add_same_cancel2 power2_less_0) lemma isDERIV_FDERIV_floatarith: assumes "isDERIV x fa vs" "\i. i < length ds \ isDERIV x (ds ! i) vs" assumes [simp]: "length xs = length ds" shows "isDERIV x (FDERIV_floatarith fa xs ds) vs" using assms apply (induction fa) subgoal by (auto simp: FDERIV_floatarith_def isDERIV_inner_iff) subgoal by (auto simp: FDERIV_floatarith_def isDERIV_inner_iff) subgoal by (auto simp: FDERIV_floatarith_def isDERIV_inner_iff) subgoal by (auto simp: FDERIV_floatarith_def isDERIV_inner_iff) subgoal by (auto simp: FDERIV_floatarith_def isDERIV_inner_iff) subgoal by (auto simp: FDERIV_floatarith_def isDERIV_inner_iff) subgoal by (auto simp: FDERIV_floatarith_def isDERIV_inner_iff) subgoal by (auto simp: FDERIV_floatarith_def isDERIV_inner_iff) subgoal by (auto simp: FDERIV_floatarith_def isDERIV_inner_iff) subgoal by (auto simp: FDERIV_floatarith_def isDERIV_inner_iff) subgoal by (auto simp: FDERIV_floatarith_def isDERIV_inner_iff) subgoal by (auto simp: FDERIV_floatarith_def isDERIV_inner_iff) subgoal by (auto simp: FDERIV_floatarith_def isDERIV_inner_iff) subgoal by (auto simp: FDERIV_floatarith_def isDERIV_inner_iff) subgoal for fa n by (cases n) (auto simp: FDERIV_floatarith_def isDERIV_inner_iff) subgoal by (auto simp: FDERIV_floatarith_def isDERIV_inner_iff) subgoal by (auto simp: FDERIV_floatarith_def isDERIV_inner_iff) subgoal by (auto simp: FDERIV_floatarith_def isDERIV_inner_iff) done lemma isDERIV_FDERIV_floatariths: assumes "isFDERIV N xs fas vs" "isFDERIV N xs ds vs" and [simp]: "length fas = length ds" shows "isFDERIV N xs (FDERIV_floatariths fas xs ds) vs" using assms by (auto simp: isFDERIV_def FDERIV_floatariths_def intro!: isDERIV_FDERIV_floatarith) lemma isFDERIV_imp_isFDERIV_FDERIV_n: assumes "length fas = length ds" shows "isFDERIV N xs fas vs \ isFDERIV N xs ds vs \ isFDERIV N xs (FDERIV_n_floatariths fas xs ds n) vs" using assms by (induction n) (auto intro!: isDERIV_FDERIV_floatariths) lemma isFDERIV_map_Var: assumes [simp]: "length ds = N" "length xs = N" shows "isFDERIV N xs (map Var ds) vs" by (auto simp: isFDERIV_def) theorem isFDERIV_imp_isnFDERIV: assumes "isFDERIV N xs fas vs" and [simp]: "length fas = N" "length xs = N" "length ds = N" shows "isnFDERIV N fas xs ds vs n" using assms by (induction n) (auto intro!: isFDERIV_imp_isFDERIV_FDERIV_n isFDERIV_map_Var) lemma eventually_isnFDERIV: assumes iD: "isnFDERIV DIM('a) fas [0.. 2 * DIM('a::executable_euclidean_space)" shows "\\<^sub>F (x::'a) in at x. isnFDERIV DIM('a) fas [0..\<^sub>F x in at x. isnFDERIV DIM('a) fas [0.. DIM('a) + length (list_of_eucl d)" by (auto intro!: max_Var_floatarith_FDERIV_n_floatariths[THEN order_trans] m[THEN order_trans]) from eventually_isFDERIV[OF 2 this] 1 show ?case by eventually_elim simp qed simp lemma isFDERIV_open: assumes "max_Var_floatariths fas \ DIM('a)" shows "open {x::'a. isFDERIV DIM('a::executable_euclidean_space) [0..x. DERIV_floatarith x fa) xs) vs \ (einterpret ds vs::'a)" by (auto simp: FDERIV_floatarith_def interpret_floatarith_inner_floatariths) lemma interpret_floatariths_FDERIV_floatariths_cong: assumes [simp]: "length d1s = DIM('a::executable_euclidean_space)" "length d2s = DIM('a)" "length fas1 = length fas2" assumes fresh1: "freshs_floatariths fas1 d1s" assumes fresh2: "freshs_floatariths fas2 d2s" assumes eq1: "\i. i < length fas1 \ interpret_floatariths (map (\x. DERIV_floatarith x (fas1 ! i)) [0..x. DERIV_floatarith x (fas2 ! i)) [0..i. i < DIM('a) \ xs1 ! (d1s ! i) = xs2 ! (d2s ! i)" shows "interpret_floatariths (FDERIV_floatariths fas1 [0..x. x = n \ s x = n" shows "subst_floatarith (\x. Var (s x)) (DERIV_floatarith n fa) = DERIV_floatarith n (subst_floatarith (\x. Var (s x)) fa)" using assms proof (induction fa) case (Power fa n) then show ?case by (cases n) auto qed force+ lemma subst_floatarith_inner_floatariths[simp]: assumes "length fs = length gs" shows "subst_floatarith s (inner_floatariths fs gs) = inner_floatariths (map (subst_floatarith s) fs) (map (subst_floatarith s) gs)" using assms by (induction rule: list_induct2) auto fun_cases subst_floatarith_Num: "subst_floatarith s fa = Num y" and subst_floatarith_Add: "subst_floatarith s fa = Add x y" and subst_floatarith_Minus: "subst_floatarith s fa = Minus y" lemma Num_eq_subst_Var[simp]: "Num x = subst_floatarith (\x. Var (s x)) fa \ fa = Num x" by (cases fa) auto lemma Add_eq_subst_VarE: assumes "Add fa1 fa2 = subst_floatarith (\x. Var (s x)) fa" obtains a1 a2 where "fa = Add a1 a2" "fa1 = subst_floatarith (\x. Var (s x)) a1" "fa2 = subst_floatarith (\x. Var (s x)) a2" using assms by (cases fa) auto lemma subst_floatarith_eq_self[simp]: "subst_floatarith s f = f" if "max_Var_floatarith f = 0" using that by (induction f) auto lemma fold_const_fa_unique: "False" if "(\x. f = Num x)" using that[of 0] that[of 1] by auto lemma zero_unique: False if "(\x::float. x = 0)" using that[of 0] that[of 1] by auto lemma fold_const_fa_Mult_eq_NumE: assumes "fold_const_fa (Mult a b) = Num x" obtains y z where "fold_const_fa a = Num y" "fold_const_fa b = Num z" "x = y * z" | y where "fold_const_fa a = Num 0" "x = 0" | y where "fold_const_fa b = Num 0" "x = 0" using assms by atomize_elim (auto simp: fold_const_fa.simps split!: option.splits if_splits elim!: dest_Num_fa_Some dest_Num_fa_None) lemma fold_const_fa_Add_eq_NumE: assumes "fold_const_fa (Add a b) = Num x" obtains y z where "fold_const_fa a = Num y" "fold_const_fa b = Num z" "x = y + z" using assms by atomize_elim (auto simp: fold_const_fa.simps split!: option.splits if_splits elim!: dest_Num_fa_Some dest_Num_fa_None) lemma subst_floatarith_Var_fold_const_fa[symmetric]: "fold_const_fa (subst_floatarith (\x. Var (s x)) fa) = subst_floatarith (\x. Var (s x)) (fold_const_fa fa)" proof (induction fa) case (Add fa1 fa2) then show ?case apply (auto simp: fold_const_fa.simps split!: floatarith.splits option.splits if_splits elim!: dest_Num_fa_Some) apply (metis Num_eq_subst_Var dest_Num_fa.simps(1) option.simps(3)) apply (metis Num_eq_subst_Var dest_Num_fa.simps(1) option.simps(3)) apply (metis Num_eq_subst_Var dest_Num_fa.simps(1) option.simps(3)) apply (metis Num_eq_subst_Var dest_Num_fa.simps(1) option.simps(3)) apply (metis Num_eq_subst_Var dest_Num_fa.simps(1) option.simps(3)) apply (metis Num_eq_subst_Var dest_Num_fa.simps(1) option.simps(3)) apply (metis Num_eq_subst_Var dest_Num_fa.simps(1) option.simps(3)) apply (metis Num_eq_subst_Var dest_Num_fa.simps(1) option.simps(3)) apply (metis Num_eq_subst_Var dest_Num_fa.simps(1) option.simps(3)) apply (metis Num_eq_subst_Var dest_Num_fa.simps(1) option.simps(3)) done next case (Mult fa1 fa2) then show ?case apply (auto simp: fold_const_fa.simps split!: floatarith.splits option.splits if_splits elim!: dest_Num_fa_Some) apply (metis Num_eq_subst_Var dest_Num_fa.simps(1) option.simps(3)) apply (metis Num_eq_subst_Var dest_Num_fa.simps(1) option.simps(3)) apply (metis Num_eq_subst_Var dest_Num_fa.simps(1) option.simps(3)) apply (metis Num_eq_subst_Var dest_Num_fa.simps(1) option.simps(3)) apply (metis Num_eq_subst_Var dest_Num_fa.simps(1) option.simps(3)) apply (metis Num_eq_subst_Var dest_Num_fa.simps(1) option.simps(3)) apply (metis Num_eq_subst_Var dest_Num_fa.simps(1) option.simps(3)) apply (metis Num_eq_subst_Var dest_Num_fa.simps(1) option.simps(3)) apply (metis Num_eq_subst_Var dest_Num_fa.simps(1) option.simps(3)) apply (metis Num_eq_subst_Var dest_Num_fa.simps(1) option.simps(3)) apply (metis Num_eq_subst_Var dest_Num_fa.simps(1) option.simps(3)) apply (metis Num_eq_subst_Var dest_Num_fa.simps(1) option.simps(3)) apply (metis Num_eq_subst_Var dest_Num_fa.simps(1) option.simps(3)) apply (metis Num_eq_subst_Var dest_Num_fa.simps(1) option.simps(3)) apply (metis Num_eq_subst_Var dest_Num_fa.simps(1) option.simps(3)) done next case (Min) then show ?case by (auto simp: fold_const_fa.simps split: floatarith.splits) next case (Max) then show ?case by (auto simp: fold_const_fa.simps split: floatarith.splits) qed (auto simp: fold_const_fa.simps split!: floatarith.splits option.splits if_splits elim!: dest_Num_fa_Some) lemma subst_floatarith_eq_Num[simp]: "(subst_floatarith (\x. Var (s x)) fa = Num x) \ fa = Num x" by (induction fa) (auto simp: ) lemma fold_const_fa_subst_eq_Num0_iff[simp]: "fold_const_fa (subst_floatarith (\x. Var (s x)) fa) = Num x \ fold_const_fa fa = Num x" unfolding subst_floatarith_Var_fold_const_fa[symmetric] by simp lemma subst_floatarith_Var_FDERIV_floatarith: assumes len: "length xs = DIM('a::executable_euclidean_space)" and [simp]: "length ds = DIM('a)" assumes eq: "\x y. x \ set xs \ (y = x) = (s y = x)" shows "subst_floatarith (\x. Var (s x)) (FDERIV_floatarith fa xs ds) = (FDERIV_floatarith (subst_floatarith (\x. Var (s x)) fa) xs (map (subst_floatarith (\x. Var (s x))) ds))" proof - have [simp]: "\x. x \ set xs \ subst_floatarith (\x. Var (s x)) (DERIV_floatarith x fa1) = (DERIV_floatarith x (subst_floatarith (\x. Var (s x)) fa1))" for fa1 by (rule subst_floatarith_Var_DERIV_floatarith) (rule eq) have map_eq: "(map (\xa. if xa = s x then Num 1 else Num 0) xs) = (map (\xa. if xa = x then Num 1 else Num 0) xs)" for x apply (subst map_eq_conv) using eq[of x x] eq[of "s x"] by (auto simp: ) show ?thesis using len by (induction fa) (auto simp: FDERIV_floatarith_def o_def if_distrib subst_floatarith_Var_fold_const_fa fold_const_fa.simps(18) map_eq cong: map_cong if_cong) qed lemma subst_floatarith_Var_FDERIV_n_nth: assumes len: "length xs = DIM('a::executable_euclidean_space)" and [simp]: "length ds = DIM('a)" assumes eq: "\x y. x \ set xs \ (y = x) = (s y = x)" assumes [simp]: "i < length fas" shows "subst_floatarith (\x. Var (s x)) (FDERIV_n_floatariths fas xs ds n ! i) = (FDERIV_n_floatariths (map (subst_floatarith (\x. Var (s x))) fas) xs (map (subst_floatarith (\x. Var (s x))) ds) n ! i)" proof (induction n) case (Suc n) show ?case by (simp add: subst_floatarith_Var_FDERIV_floatarith[OF len _ eq] Suc.IH[symmetric]) qed simp lemma subst_floatarith_Var_max_Var_floatarith: assumes "\i. i < max_Var_floatarith fa \ s i = i" shows "subst_floatarith (\i. Var (s i)) fa = fa" using assms by (induction fa) auto lemma interpret_floatarith_subst_floatarith_idem: assumes mv: "max_Var_floatarith fa \ length vs" assumes idem: "\j. j < max_Var_floatarith fa \ vs ! s j = vs ! j" shows "interpret_floatarith (subst_floatarith (\i. Var (s i)) fa) vs = interpret_floatarith fa vs" using assms by (induction fa) auto lemma isDERIV_subst_Var_floatarith: assumes mv: "max_Var_floatarith fa \ length vs" assumes idem: "\j. j < max_Var_floatarith fa \ vs ! s j = vs ! j" assumes "\j. s j = i \ j = i" shows "isDERIV i (subst_floatarith (\i. Var (s i)) fa) vs = isDERIV i fa vs" using mv idem proof (induction fa) case (Power fa n) then show ?case by (cases n) auto qed (auto simp: interpret_floatarith_subst_floatarith_idem) lemma isFDERIV_subst_Var_floatarith: assumes mv: "max_Var_floatariths fas \ length vs" assumes idem: "\j. j < max_Var_floatariths fas \ vs ! (s j) = vs ! j" assumes "\i j. i \ set xs \ s j = i \ j = i" shows "isFDERIV n xs (map (subst_floatarith (\i. Var (s i))) fas) vs = isFDERIV n xs fas vs" proof - have mv: "\i. i < length fas \ max_Var_floatarith (fas ! i) \ length vs" apply (rule order_trans[OF _ mv]) by (intro max_Var_floatarith_le_max_Var_floatariths_nth) have idem: "\i j. i < length fas \ j < max_Var_floatarith (fas ! i) \ vs ! s j = vs ! j" using idem by (auto simp: dest!: max_Var_floatariths_lessI) show ?thesis unfolding isFDERIV_def using mv idem assms(3) by (auto simp: isDERIV_subst_Var_floatarith) qed lemma interpret_floatariths_append[simp]: "interpret_floatariths (xs @ ys) vs = interpret_floatariths xs vs @ interpret_floatariths ys vs" by (induction xs) auto lemma not_fresh_inner_floatariths: assumes "length xs = length ys" shows "\ fresh_floatarith (inner_floatariths xs ys) i \ \fresh_floatariths xs i \ \fresh_floatariths ys i" using assms by (induction xs ys rule: list_induct2) auto lemma fresh_inner_floatariths: assumes "length xs = length ys" shows "fresh_floatarith (inner_floatariths xs ys) i \ fresh_floatariths xs i \ fresh_floatariths ys i" using not_fresh_inner_floatariths assms by auto lemma not_fresh_floatariths_map: " \ fresh_floatariths (map f xs) i \ (\x \ set xs. \fresh_floatarith (f x) i)" by (induction xs) auto lemma fresh_floatariths_map: " fresh_floatariths (map f xs) i \ (\x \ set xs. fresh_floatarith (f x) i)" by (induction xs) auto lemma fresh_floatarith_fold_const_fa: "fresh_floatarith fa i \ fresh_floatarith (fold_const_fa fa) i" by (induction fa) (auto simp: fold_const_fa.simps split: floatarith.splits option.splits) lemma fresh_floatarith_fold_const_fa_Add[intro!]: assumes "fresh_floatarith (fold_const_fa a) i" "fresh_floatarith (fold_const_fa b) i" shows "fresh_floatarith (fold_const_fa (Add a b)) i" using assms by (auto simp: fold_const_fa.simps split!: floatarith.splits option.splits) lemma fresh_floatarith_fold_const_fa_Mult[intro!]: assumes "fresh_floatarith (fold_const_fa a) i" "fresh_floatarith (fold_const_fa b) i" shows "fresh_floatarith (fold_const_fa (Mult a b)) i" using assms by (auto simp: fold_const_fa.simps split!: floatarith.splits option.splits) lemma fresh_floatarith_fold_const_fa_Minus[intro!]: assumes "fresh_floatarith (fold_const_fa b) i" shows "fresh_floatarith (fold_const_fa (Minus b)) i" using assms by (auto simp: fold_const_fa.simps split!: floatarith.splits) lemma fresh_FDERIV_floatarith: "fresh_floatarith ode_e i \ fresh_floatariths ds i \ length ds = DIM('a) \ fresh_floatarith (FDERIV_floatarith ode_e [0.. fresh_floatarith (FDERIV_floatarith ode_e [0.. length ds = DIM('a) \ \fresh_floatarith ode_e i \ \fresh_floatariths ds i" using fresh_FDERIV_floatarith by auto lemma not_fresh_FDERIV_floatariths: "\ fresh_floatariths (FDERIV_floatariths ode_e [0.. length ds = DIM('a) \ \fresh_floatariths ode_e i \ \fresh_floatariths ds i" by (induction ode_e) (auto simp: FDERIV_floatariths_def dest!: not_fresh_FDERIV_floatarith) lemma isDERIV_FDERIV_floatarith_linear: fixes x h::"'a::executable_euclidean_space" assumes "\k. k < DIM('a) \ isDERIV i (DERIV_floatarith k fa) xs" assumes "max_Var_floatarith fa \ DIM('a)" assumes [simp]: "length xs = DIM('a)" "length hs = DIM('a)" shows "isDERIV i (FDERIV_floatarith fa [0..i j k. i < DIM('a) \ j < DIM('a) \ k < DIM('a) \ isDERIV i (DERIV_floatarith k (fas ! j)) (xs)" assumes [simp]: "length fas = DIM('a::executable_euclidean_space)" assumes [simp]: "length xs = DIM('a)" "length hs = DIM('a)" assumes "max_Var_floatariths fas \ DIM('a)" shows "isFDERIV DIM('a) [0..ij length fas = n \ length xs = n)" lemma isFDERIV_approx: "bounded_by vs VS \ isFDERIV_approx prec n xs fas VS \ isFDERIV n xs fas vs" by (auto simp: isFDERIV_approx_def isFDERIV_def intro!: isDERIV_approx) primrec isnFDERIV_approx where "isnFDERIV_approx p N fas xs ds vs 0 = True" | "isnFDERIV_approx p N fas xs ds vs (Suc n) \ isFDERIV_approx p N xs (FDERIV_n_floatariths fas xs (map Var ds) n) vs \ isnFDERIV_approx p N fas xs ds vs n" lemma isnFDERIV_approx: "bounded_by vs VS \ isnFDERIV_approx prec N fas xs ds VS n \ isnFDERIV N fas xs ds vs n" by (induction n) (auto intro!: isFDERIV_approx) fun plain_floatarith::"nat \ floatarith \ bool" where "plain_floatarith N (floatarith.Add a b) \ plain_floatarith N a \ plain_floatarith N b" | "plain_floatarith N (floatarith.Mult a b) \ plain_floatarith N a \ plain_floatarith N b" | "plain_floatarith N (floatarith.Minus a) \ plain_floatarith N a" | "plain_floatarith N (floatarith.Pi) \ True" | "plain_floatarith N (floatarith.Num n) \ True" | "plain_floatarith N (floatarith.Var i) \ i < N" | "plain_floatarith N (floatarith.Max a b) \ plain_floatarith N a \ plain_floatarith N b" | "plain_floatarith N (floatarith.Min a b) \ plain_floatarith N a \ plain_floatarith N b" | "plain_floatarith N (floatarith.Power a n) \ plain_floatarith N a" | "plain_floatarith N (floatarith.Cos a) \ False" \ \TODO: should be plain!\ | "plain_floatarith N (floatarith.Arctan a) \ False" \ \TODO: should be plain!\ | "plain_floatarith N (floatarith.Abs a) \ plain_floatarith N a" | "plain_floatarith N (floatarith.Exp a) \ False" \ \TODO: should be plain!\ | "plain_floatarith N (floatarith.Sqrt a) \ False" \ \TODO: should be plain!\ | "plain_floatarith N (floatarith.Floor a) \ plain_floatarith N a" | "plain_floatarith N (floatarith.Powr a b) \ False" | "plain_floatarith N (floatarith.Inverse a) \ False" | "plain_floatarith N (floatarith.Ln a) \ False" lemma plain_floatarith_approx_not_None: assumes "plain_floatarith N fa" "N \ length XS" "\i. i < N \ XS ! i \ None" shows "approx p fa XS \ None" using assms by (induction fa) (auto simp: Let_def split_beta' prod_eq_iff approx.simps) definition "Rad_of w = w * (Pi / Num 180)" lemma interpret_Rad_of[simp]: "interpret_floatarith (Rad_of w) xs = rad_of (interpret_floatarith w xs)" by (auto simp: Rad_of_def rad_of_def) definition "Deg_of w = Num 180 * w / Pi" lemma interpret_Deg_of[simp]: "interpret_floatarith (Deg_of w) xs = deg_of (interpret_floatarith w xs)" by (auto simp: Deg_of_def deg_of_def inverse_eq_divide) unbundle no_floatarith_notation end diff --git a/thys/Applicative_Lifting/applicative.ML b/thys/Applicative_Lifting/applicative.ML --- a/thys/Applicative_Lifting/applicative.ML +++ b/thys/Applicative_Lifting/applicative.ML @@ -1,1369 +1,1369 @@ (* Author: Joshua Schneider, ETH Zurich *) signature APPLICATIVE = sig type afun val intern: Context.generic -> xstring -> string val extern: Context.generic -> string -> xstring val afun_of_generic: Context.generic -> string -> afun val afun_of: Proof.context -> string -> afun val afuns_of_term_generic: Context.generic -> term -> afun list val afuns_of_term: Proof.context -> term -> afun list val afuns_of_typ_generic: Context.generic -> typ -> afun list val afuns_of_typ: Proof.context -> typ -> afun list val name_of_afun: afun -> binding val unfolds_of_afun: afun -> thm list type afun_inst val match_afun_inst: Proof.context -> afun -> term * int -> afun_inst val import_afun_inst: afun -> Proof.context -> afun_inst * Proof.context val inner_sort_of: afun_inst -> sort val mk_type: afun_inst -> typ -> typ val mk_pure: afun_inst -> typ -> term val lift_term: afun_inst -> term -> term val mk_ap: afun_inst -> typ * typ -> term val mk_comb: afun_inst -> typ -> term * term -> term val mk_set: afun_inst -> typ -> term val dest_type: Proof.context -> afun_inst -> typ -> typ option val dest_type': Proof.context -> afun_inst -> typ -> typ val dest_pure: Proof.context -> afun_inst -> term -> term val dest_comb: Proof.context -> afun_inst -> term -> term * term val infer_comb: Proof.context -> afun_inst -> term * term -> term val subst_lift_term: afun_inst -> (term * term) list -> term -> term val generalize_lift_terms: afun_inst -> term list -> Proof.context -> term list * Proof.context val afun_unfold_tac: Proof.context -> afun -> int -> tactic val afun_fold_tac: Proof.context -> afun -> int -> tactic val unfold_all_tac: Proof.context -> int -> tactic val normalform_conv: Proof.context -> afun -> conv val normalize_rel_tac: Proof.context -> afun -> int -> tactic val general_normalform_conv: Proof.context -> afun -> cterm * cterm -> thm * thm val general_normalize_rel_tac: Proof.context -> afun -> int -> tactic val forward_lift_rule: Proof.context -> afun -> thm -> thm val unfold_wrapper_tac: Proof.context -> afun option -> int -> tactic val fold_wrapper_tac: Proof.context -> afun option -> int -> tactic val normalize_wrapper_tac: Proof.context -> afun option -> int -> tactic val lifting_wrapper_tac: Proof.context -> afun option -> int -> tactic val setup_combinators: (string * thm) list -> local_theory -> local_theory val combinator_rule_attrib: string list option -> attribute val parse_opt_afun: afun option context_parser val applicative_cmd: (((((binding * string list) * string) * string) * string option) * string option) -> local_theory -> Proof.state val print_afuns: Proof.context -> unit val add_unfold_attrib: xstring option -> attribute val forward_lift_attrib: xstring -> attribute end; structure Applicative : APPLICATIVE = struct open Ctr_Sugar_Util (** General utilities **) fun fold_options xs = fold (fn x => (case x of SOME x' => cons x' | NONE => I)) xs []; fun the_pair [x, y] = (x, y) | the_pair _ = raise General.Size; fun strip_comb2 (f $ x $ y) = (f, (x, y)) | strip_comb2 t = raise TERM ("strip_comb2", [t]); fun mk_comb_pattern (t, n) = let val Ts = take n (binder_types (fastype_of t)); val maxidx = maxidx_of_term t; val vars = map (fn (T, i) => ((Name.uu, maxidx + i), T)) (Ts ~~ (1 upto n)); in (vars, Term.betapplys (t, map Var vars)) end; fun match_comb_pattern ctxt tn u = let val thy = Proof_Context.theory_of ctxt; val (vars, pat) = mk_comb_pattern tn; val envs = Pattern.match thy (pat, u) (Vartab.empty, Vartab.empty) handle Pattern.MATCH => raise TERM ("match_comb_pattern", [u, pat]); in (vars, envs) end; fun dest_comb_pattern ctxt tn u = let val (vars, (_, env)) = match_comb_pattern ctxt tn u; in map (the o Envir.lookup1 env) vars end; fun norm_term_types tyenv t = Term_Subst.map_types_same (Envir.norm_type_same tyenv) t handle Same.SAME => t; val mk_TFrees_of = mk_TFrees' oo replicate; fun mk_Free name typ ctxt = yield_singleton Variable.variant_fixes name ctxt |>> (fn name' => Free (name', typ)); (*tuples with explicit sentinel*) fun mk_tuple' ts = fold_rev (curry HOLogic.mk_prod) ts HOLogic.unit; fun strip_tuple' (Const (@{const_name Unity}, _)) = [] | strip_tuple' (Const (@{const_name Pair}, _) $ t1 $ t2) = t1 :: strip_tuple' t2 | strip_tuple' t = raise TERM ("strip_tuple'", [t]); fun mk_eq_on S = let val (SA, ST) = `HOLogic.dest_setT (fastype_of S); in Const (@{const_name eq_on}, ST --> BNF_Util.mk_pred2T SA SA) $ S end; (* Polymorphic terms and term groups *) type poly_type = typ list * typ; type poly_term = typ list * term; fun instantiate_poly_type (tvars, T) insts = typ_subst_atomic (tvars ~~ insts) T; fun instantiate_poly_term (tvars, t) insts = subst_atomic_types (tvars ~~ insts) t; fun dest_poly_type ctxt (tvars, T) U = let val thy = Proof_Context.theory_of ctxt; val tyenv = Sign.typ_match thy (T, U) Vartab.empty handle Type.TYPE_MATCH => raise TYPE ("dest_poly_type", [U, T], []); in map (Type.lookup tyenv o dest_TVar) tvars end; fun poly_type_to_term (tvars, T) = (tvars, Logic.mk_type T); fun poly_type_of_term (tvars, t) = (tvars, Logic.dest_type t); (* Schematic variables are treated uniformly in packed terms, thus forming an ad hoc context of type variables. Otherwise, morphisms are allowed to rename schematic variables non-consistently in separate terms, and occasionally will do so. *) fun pack_poly_term (tvars, t) = HOLogic.mk_prod (mk_tuple' (map Logic.mk_type tvars), t); fun unpack_poly_term t = let val (tvars, t') = HOLogic.dest_prod t; in (map Logic.dest_type (strip_tuple' tvars), t') end; val pack_poly_terms = mk_tuple' o map pack_poly_term; val unpack_poly_terms = map unpack_poly_term o strip_tuple'; (*match and instantiate schematic type variables which are not "quantified" in the packed term*) fun match_poly_terms_type ctxt (pt, i) (U, maxidx) = let val thy = Proof_Context.theory_of ctxt; val pt' = Logic.incr_indexes ([], [], maxidx + 1) pt; val (tvars, T) = poly_type_of_term (nth (unpack_poly_terms pt') i); val tyenv = Sign.typ_match thy (T, U) Vartab.empty handle Type.TYPE_MATCH => raise TYPE ("match_poly_terms", [U, T], []); val tyenv' = fold Vartab.delete_safe (map (#1 o dest_TVar) tvars) tyenv; val pt'' = Envir.subst_term_types tyenv' pt'; in unpack_poly_terms pt'' end; fun match_poly_terms ctxt (pt, i) (t, maxidx) = match_poly_terms_type ctxt (pt, i) (fastype_of t, maxidx); (*fix schematic type variables which are not "quantified", as well as schematic term variables*) fun import_poly_terms pt ctxt = let fun insert_paramTs (tvars, t) = fold_types (fold_atyps (fn TVar v => if member (op =) tvars (TVar v) then I else insert (op =) v | _ => I)) t; val paramTs = rev (fold insert_paramTs (unpack_poly_terms pt) []); val (tfrees, ctxt') = Variable.invent_types (map #2 paramTs) ctxt; val instT = TVars.make (paramTs ~~ map TFree tfrees); val params = map (apsnd (Term_Subst.instantiateT instT)) (rev (Term.add_vars pt [])); val (frees, ctxt'') = Variable.variant_fixes (map (Name.clean o #1 o #1) params) ctxt'; val inst = Vars.make (params ~~ map Free (frees ~~ map #2 params)); val pt' = Term_Subst.instantiate (instT, inst) pt; in (unpack_poly_terms pt', ctxt'') end; (** Internal representation **) (* Applicative functors *) type rel_thms = { pure_transfer: thm, ap_rel_fun: thm }; fun map_rel_thms f {pure_transfer, ap_rel_fun} = {pure_transfer = f pure_transfer, ap_rel_fun = f ap_rel_fun}; type afun_thms = { hom: thm, ichng: thm, reds: thm Symtab.table, rel_thms: rel_thms option, rel_intros: thm list, pure_comp_conv: thm }; fun map_afun_thms f {hom, ichng, reds, rel_thms, rel_intros, pure_comp_conv} = {hom = f hom, ichng = f ichng, reds = Symtab.map (K f) reds, rel_thms = Option.map (map_rel_thms f) rel_thms, rel_intros = map f rel_intros, pure_comp_conv = f pure_comp_conv}; datatype afun = AFun of { name: binding, terms: term, rel: term option, thms: afun_thms, unfolds: thm Item_Net.T }; fun rep_afun (AFun af) = af; val name_of_afun = #name o rep_afun; val terms_of_afun = #terms o rep_afun; val rel_of_afun = #rel o rep_afun; val thms_of_afun = #thms o rep_afun; val unfolds_of_afun = Item_Net.content o #unfolds o rep_afun; val red_of_afun = Symtab.lookup o #reds o thms_of_afun; val has_red_afun = is_some oo red_of_afun; fun mk_afun name terms rel thms = AFun {name = name, terms = terms, rel = rel, thms = thms, unfolds = Thm.item_net}; fun map_afun f1 f2 f3 f4 f5 (AFun {name, terms, rel, thms, unfolds}) = AFun {name = f1 name, terms = f2 terms, rel = f3 rel, thms = f4 thms, unfolds = f5 unfolds}; fun map_unfolds f thms = fold Item_Net.update (map f (Item_Net.content thms)) Thm.item_net; fun morph_afun phi = let val binding = Morphism.binding phi; val term = Morphism.term phi; val thm = Morphism.thm phi; in map_afun binding term (Option.map term) (map_afun_thms thm) (map_unfolds thm) end; val transfer_afun = morph_afun o Morphism.transfer_morphism; fun add_unfolds_afun thms = map_afun I I I I (fold Item_Net.update thms); fun patterns_of_afun af = let val [Tt, (_, pure), (_, ap), _] = unpack_poly_terms (terms_of_afun af); val (_, T) = poly_type_of_term Tt; in [#2 (mk_comb_pattern (pure, 1)), #2 (mk_comb_pattern (ap, 2)), Net.encode_type T] end; (* Combinator rules *) datatype combinator_rule = Combinator_Rule of { strong_premises: string Ord_List.T, weak_premises: bool, conclusion: string, eq_thm: thm }; fun rep_combinator_rule (Combinator_Rule rule) = rule; val conclusion_of_rule = #conclusion o rep_combinator_rule; val thm_of_rule = #eq_thm o rep_combinator_rule; fun eq_combinator_rule (rule1, rule2) = pointer_eq (rule1, rule2) orelse Thm.eq_thm (thm_of_rule rule1, thm_of_rule rule2); fun is_applicable_rule rule have_weak have_premises = let val {strong_premises, weak_premises, ...} = rep_combinator_rule rule; in (have_weak orelse not weak_premises) andalso have_premises strong_premises end; fun map_combinator_rule f1 f2 f3 f4 (Combinator_Rule {strong_premises, weak_premises, conclusion, eq_thm}) = Combinator_Rule {strong_premises = f1 strong_premises, weak_premises = f2 weak_premises, conclusion = f3 conclusion, eq_thm = f4 eq_thm}; fun transfer_combinator_rule thy = map_combinator_rule I I I (Thm.transfer thy); fun mk_combinator_rule comb_names weak_premises thm = let val (lhs, rhs) = Logic.dest_equals (Thm.prop_of thm); val conclusion = the (Symtab.lookup comb_names (#1 (dest_Const lhs))); val premises = Ord_List.make fast_string_ord (fold_options (map (Symtab.lookup comb_names o #1) (Term.add_consts rhs []))); val weak_premises' = Ord_List.make fast_string_ord (these weak_premises); val strong_premises = Ord_List.subtract fast_string_ord weak_premises' premises; in Combinator_Rule {strong_premises = strong_premises, weak_premises = is_some weak_premises, conclusion = conclusion, eq_thm = thm} end; (* Generic data *) (*FIXME: needs tests, especially around theory merging*) fun merge_afuns _ (af1, af2) = if pointer_eq (af1, af2) then raise Change_Table.SAME else map_afun I I I I (fn thms1 => Item_Net.merge (thms1, #unfolds (rep_afun af2))) af1; structure Data = Generic_Data ( type T = { combinators: thm Symtab.table * combinator_rule list, afuns: afun Name_Space.table, patterns: (string * term list) Item_Net.T }; val empty = { combinators = (Symtab.empty, []), afuns = Name_Space.empty_table "applicative functor", patterns = Item_Net.init (op = o apply2 #1) #2 }; val extend = I; fun merge ({combinators = (cd1, cr1), afuns = a1, patterns = p1}, {combinators = (cd2, cr2), afuns = a2, patterns = p2}) = {combinators = (Symtab.merge (K true) (cd1, cd2), Library.merge eq_combinator_rule (cr1, cr2)), afuns = Name_Space.join_tables merge_afuns (a1, a2), patterns = Item_Net.merge (p1, p2)}; ); fun get_combinators context = let val thy = Context.theory_of context; val {combinators = (defs, rules), ...} = Data.get context; in (Symtab.map (K (Thm.transfer thy)) defs, map (transfer_combinator_rule thy) rules) end; val get_afun_table = #afuns o Data.get; val get_afun_space = Name_Space.space_of_table o get_afun_table; val get_patterns = #patterns o Data.get; fun map_data f1 f2 f3 {combinators, afuns, patterns} = {combinators = f1 combinators, afuns = f2 afuns, patterns = f3 patterns}; val intern = Name_Space.intern o get_afun_space; fun extern context = Name_Space.extern (Context.proof_of context) (get_afun_space context); local fun undeclared name = error ("Undeclared applicative functor " ^ quote name); in fun afun_of_generic context name = case Name_Space.lookup (get_afun_table context) name of SOME af => transfer_afun (Context.theory_of context) af | NONE => undeclared name; val afun_of = afun_of_generic o Context.Proof; fun update_afun name f context = if Name_Space.defined (get_afun_table context) name then Data.map (map_data I (Name_Space.map_table_entry name f) I) context else undeclared name; end; fun match_term context = map #1 o Item_Net.retrieve_matching (get_patterns context); fun match_typ context = match_term context o Net.encode_type; (*works only with terms which are combinations of pure and ap*) fun afuns_of_term_generic context = map (afun_of_generic context) o match_term context; val afuns_of_term = afuns_of_term_generic o Context.Proof; fun afuns_of_typ_generic context = map (afun_of_generic context) o match_typ context; val afuns_of_typ = afuns_of_typ_generic o Context.Proof; fun all_unfolds_of_generic context = let val unfolds_of = map (Thm.transfer'' context) o unfolds_of_afun; in Name_Space.fold_table (fn (_, af) => append (unfolds_of af)) (get_afun_table context) [] end; val all_unfolds_of = all_unfolds_of_generic o Context.Proof; (** Term construction and destruction **) type afun_inst = { T: poly_type, pure: poly_term, ap: poly_term, set: poly_term }; fun mk_afun_inst [T, pure, ap, set] = {T = poly_type_of_term T, pure = pure, ap = ap, set = set}; fun pack_afun_inst {T, pure, ap, set} = pack_poly_terms [poly_type_to_term T, pure, ap, set]; fun match_afun_inst ctxt af = match_poly_terms ctxt (terms_of_afun af, 0) #> mk_afun_inst; fun import_afun_inst_raw terms = import_poly_terms terms #>> mk_afun_inst; val import_afun_inst = import_afun_inst_raw o terms_of_afun; fun inner_sort_of {T = (tvars, _), ...} = Type.sort_of_atyp (the_single tvars); fun mk_type {T, ...} = instantiate_poly_type T o single; fun mk_pure {pure, ...} = instantiate_poly_term pure o single; fun mk_ap {ap, ...} (T1, T2) = instantiate_poly_term ap [T1, T2]; fun mk_set {set, ...} = instantiate_poly_term set o single; fun lift_term af_inst t = Term.betapply (mk_pure af_inst (Term.fastype_of t), t); fun mk_comb af_inst funT (t1, t2) = Term.betapplys (mk_ap af_inst (dest_funT funT), [t1, t2]); fun dest_type ctxt {T, ...} = the_single o dest_poly_type ctxt T; val dest_type' = the_default HOLogic.unitT ooo dest_type; fun dest_pure ctxt {pure = (_, pure), ...} = the_single o dest_comb_pattern ctxt (pure, 1); fun dest_comb ctxt {ap = (_, ap), ...} = the_pair o dest_comb_pattern ctxt (ap, 2); fun infer_comb ctxt af_inst (t1, t2) = let val funT = the_default (dummyT --> dummyT) (dest_type ctxt af_inst (fastype_of t1)); in mk_comb af_inst funT (t1, t2) end; (*lift a term, except for non-combination subterms mapped by subst*) fun subst_lift_term af_inst subst tm = let fun subst_lift (s $ t) = (case (subst_lift s, subst_lift t) of (NONE, NONE) => NONE | (SOME s', NONE) => SOME (mk_comb af_inst (fastype_of s) (s', lift_term af_inst t)) | (NONE, SOME t') => SOME (mk_comb af_inst (fastype_of s) (lift_term af_inst s, t')) | (SOME s', SOME t') => SOME (mk_comb af_inst (fastype_of s) (s', t'))) | subst_lift t = AList.lookup (op aconv) subst t; in (case subst_lift tm of NONE => lift_term af_inst tm | SOME tm' => tm') end; fun add_lifted_vars (s $ t) = add_lifted_vars s #> add_lifted_vars t | add_lifted_vars (Abs (_, _, t)) = Term.add_vars t | add_lifted_vars _ = I; (*lift terms, where schematic variables are generalized to the functor and then fixed*) fun generalize_lift_terms af_inst ts ctxt = let val vars = subtract (op =) (fold add_lifted_vars ts []) (fold Term.add_vars ts []); val (var_names, Ts) = split_list vars; val (free_names, ctxt') = Variable.variant_fixes (map #1 var_names) ctxt; val Ts' = map (mk_type af_inst) Ts; val subst = map Var vars ~~ map Free (free_names ~~ Ts'); in (map (subst_lift_term af_inst subst) ts, ctxt') end; (** Reasoning with applicative functors **) (* Utilities *) val clean_name = perhaps (perhaps_apply [try Name.dest_skolem, try Name.dest_internal]); (*based on term_name from Pure/term.ML*) fun term_to_vname (Const (x, _)) = Long_Name.base_name x | term_to_vname (Free (x, _)) = clean_name x | term_to_vname (Var ((x, _), _)) = clean_name x | term_to_vname _ = "x"; fun afuns_of_rel precise ctxt t = let val (_, (lhs, rhs)) = Variable.focus NONE t ctxt |> #1 |> #2 |> Logic.strip_imp_concl |> Envir.beta_eta_contract |> HOLogic.dest_Trueprop |> strip_comb2; in if precise then (case afuns_of_term ctxt lhs of [] => afuns_of_term ctxt rhs | afs => afs) else afuns_of_typ ctxt (fastype_of lhs) end; fun AUTO_AFUNS precise tac ctxt opt_af = case opt_af of SOME af => tac [af] | NONE => SUBGOAL (fn (goal, i) => (case afuns_of_rel precise ctxt goal of [] => no_tac | afs => tac afs i) handle TERM _ => no_tac); fun AUTO_AFUN precise tac = AUTO_AFUNS precise (tac o hd); fun binop_par_conv cv ct = let val ((binop, arg1), arg2) = Thm.dest_comb ct |>> Thm.dest_comb; val (th1, th2) = cv (arg1, arg2); in Drule.binop_cong_rule binop th1 th2 end; fun binop_par_conv_tac cv = CONVERSION (HOLogic.Trueprop_conv (binop_par_conv cv)); val fold_goal_tac = SELECT_GOAL oo Raw_Simplifier.fold_goals_tac; (* Unfolding of lifted constants *) fun afun_unfold_tac ctxt af = Raw_Simplifier.rewrite_goal_tac ctxt (unfolds_of_afun af); fun afun_fold_tac ctxt af = fold_goal_tac ctxt (unfolds_of_afun af); fun unfold_all_tac ctxt = Raw_Simplifier.rewrite_goal_tac ctxt (all_unfolds_of ctxt); (* Basic conversions *) fun pure_conv ctxt {pure = (_, pure), ...} cv ct = let val ([var], (tyenv, env)) = match_comb_pattern ctxt (pure, 1) (Thm.term_of ct); val arg = the (Envir.lookup1 env var); val thm = cv (Thm.cterm_of ctxt arg); in if Thm.is_reflexive thm then Conv.all_conv ct else let val pure_inst = Envir.subst_term_types tyenv pure; in Drule.arg_cong_rule (Thm.cterm_of ctxt pure_inst) thm end end; fun ap_conv ctxt {ap = (_, ap), ...} cv1 cv2 ct = let val ([var1, var2], (tyenv, env)) = match_comb_pattern ctxt (ap, 2) (Thm.term_of ct); val (arg1, arg2) = apply2 (the o Envir.lookup1 env) (var1, var2); val thm1 = cv1 (Thm.cterm_of ctxt arg1); val thm2 = cv2 (Thm.cterm_of ctxt arg2); in if Thm.is_reflexive thm1 andalso Thm.is_reflexive thm2 then Conv.all_conv ct else let val ap_inst = Envir.subst_term_types tyenv ap; in Drule.binop_cong_rule (Thm.cterm_of ctxt ap_inst) thm1 thm2 end end; (* Normal form conversion *) (*convert a term into applicative normal form*) fun normalform_conv ctxt af ct = let val {hom, ichng, pure_comp_conv, ...} = thms_of_afun af; val the_red = the o red_of_afun af; val leaf_conv = Conv.rewr_conv (mk_meta_eq (the_red "I") |> Thm.symmetric); val merge_conv = Conv.rewr_conv (mk_meta_eq hom); val swap_conv = Conv.rewr_conv (mk_meta_eq ichng); val rotate_conv = Conv.rewr_conv (mk_meta_eq (the_red "B") |> Thm.symmetric); val pure_rotate_conv = Conv.rewr_conv (mk_meta_eq pure_comp_conv); val af_inst = match_afun_inst ctxt af (Thm.term_of ct, Thm.maxidx_of_cterm ct); fun left_conv cv = ap_conv ctxt af_inst cv Conv.all_conv; fun norm_pure_nf ct = ((pure_rotate_conv then_conv left_conv norm_pure_nf) else_conv merge_conv) ct; val norm_nf_pure = swap_conv then_conv norm_pure_nf; fun norm_nf_nf ct = ((rotate_conv then_conv left_conv (left_conv norm_pure_nf then_conv norm_nf_nf)) else_conv norm_nf_pure) ct; fun normalize ct = ((ap_conv ctxt af_inst normalize normalize then_conv norm_nf_nf) else_conv pure_conv ctxt af_inst Conv.all_conv else_conv leaf_conv) ct; in normalize ct end; val normalize_rel_tac = binop_par_conv_tac o apply2 oo normalform_conv; (* Bracket abstraction and generalized unlifting *) (*TODO: use proper conversions*) datatype apterm = Pure of term (*includes pure application*) | ApVar of int * term (*unique index, instantiated term*) | Ap of apterm * apterm; fun apterm_vars (Pure _) = I | apterm_vars (ApVar v) = cons v | apterm_vars (Ap (t1, t2)) = apterm_vars t1 #> apterm_vars t2; fun occurs_any _ (Pure _) = false | occurs_any vs (ApVar (i, _)) = exists (fn j => i = j) vs | occurs_any vs (Ap (t1, t2)) = occurs_any vs t1 orelse occurs_any vs t2; fun term_of_apterm ctxt af_inst t = let fun tm_of (Pure t) = t | tm_of (ApVar (_, t)) = t | tm_of (Ap (t1, t2)) = infer_comb ctxt af_inst (tm_of t1, tm_of t2); in tm_of t end; fun apterm_of_term ctxt af_inst t = let fun aptm_of t i = case try (dest_comb ctxt af_inst) t of SOME (t1, t2) => i |> aptm_of t1 ||>> aptm_of t2 |>> Ap | NONE => if can (dest_pure ctxt af_inst) t then (Pure t, i) else (ApVar (i, t), i + 1); in aptm_of t end; (*find a common variable sequence for two applicative terms, depending on available combinators*) fun consolidate ctxt af (t1, t2) = let fun common_inst (i, t) (j, insts) = case Termtab.lookup insts t of SOME k => (((i, t), k), (j, insts)) | NONE => (((i, t), j), (j + 1, Termtab.update (t, j) insts)); val (vars, _) = (0, Termtab.empty) |> fold_map common_inst (apterm_vars t1 []) ||>> fold_map common_inst (apterm_vars t2 []); fun merge_adjacent (([], _), _) [] = [] | merge_adjacent ((is, t), d) [] = [((is, t), d)] | merge_adjacent (([], _), _) (((i, t), d)::xs) = merge_adjacent (([i], t), d) xs | merge_adjacent ((is, t), d) (((i', t'), d')::xs) = if d = d' then merge_adjacent ((i'::is, t), d) xs else ((is, t), d) :: merge_adjacent (([i'], t'), d') xs; fun align _ [] = NONE | align ((i, t), d) (((i', t'), d')::xs) = if d = d' then SOME ([((i @ i', t), d)], xs) else Option.map (apfst (cons ((i', t'), d'))) (align ((i, t), d) xs); fun merge ([], ys) = ys | merge (xs, []) = xs | merge ((xs as ((is1, t1), d1)::xs'), ys as (((is2, t2), d2)::ys')) = if d1 = d2 then ((is1 @ is2, t1), d1) :: merge (xs', ys') else case (align ((is2, t2), d2) xs, align ((is1, t1), d1) ys) of (SOME (zs, xs''), NONE) => zs @ merge (xs'', ys') | (NONE, SOME (zs, ys'')) => zs @ merge (xs', ys'') | _ => ((is1, t1), d1) :: ((is2, t2), d2) :: merge (xs', ys'); fun unbalanced vs = error ("Unbalanced opaque terms " ^ commas_quote (map (Syntax.string_of_term ctxt o #2 o #1) vs)); fun mismatch (t1, t2) = error ("Mismatched opaque terms " ^ quote (Syntax.string_of_term ctxt t1) ^ " and " ^ quote (Syntax.string_of_term ctxt t2)); fun same ([], []) = [] | same ([], ys) = unbalanced ys | same (xs, []) = unbalanced xs | same ((((i1, t1), d1)::xs), (((i2, t2), d2)::ys)) = if d1 = d2 then ((i1 @ i2, t1), d1) :: same (xs, ys) else mismatch (t1, t2); in vars |> has_red_afun af "C" ? apply2 (sort (int_ord o apply2 #2)) |> apply2 (if has_red_afun af "W" then merge_adjacent (([], Term.dummy), 0) else map (apfst (apfst single))) |> (if has_red_afun af "K" then merge else same) |> map #1 end; fun ap_cong ctxt af_inst thm1 thm2 = let val funT = the_default (dummyT --> dummyT) (dest_type ctxt af_inst (Thm.typ_of_cterm (Thm.lhs_of thm1))); val ap_inst = Thm.cterm_of ctxt (mk_ap af_inst (dest_funT funT)); in Drule.binop_cong_rule ap_inst thm1 thm2 end; fun rewr_subst_ap ctxt af_inst rewr thm1 thm2 = let val rule1 = ap_cong ctxt af_inst thm1 thm2; val rule2 = Conv.rewr_conv rewr (Thm.rhs_of rule1); in Thm.transitive rule1 rule2 end; fun merge_pures ctxt af_inst merge_thm tt = let fun merge (Pure t) = SOME (Thm.reflexive (Thm.cterm_of ctxt t)) | merge (ApVar _) = NONE | merge (Ap (tt1, tt2)) = case merge tt1 of NONE => NONE | SOME thm1 => case merge tt2 of NONE => NONE | SOME thm2 => SOME (rewr_subst_ap ctxt af_inst merge_thm thm1 thm2); in merge tt end; exception ASSERT of string; (*abstract over a variable (opaque subterm)*) fun eliminate ctxt (af, af_inst) tt (v, v_tm) = let val {hom, ichng, ...} = thms_of_afun af; val the_red = the o red_of_afun af; val hom_conv = mk_meta_eq hom; val ichng_conv = mk_meta_eq ichng; val mk_combI = Thm.symmetric o mk_meta_eq; val id_conv = mk_combI (the_red "I"); val comp_conv = mk_combI (the_red "B"); val flip_conv = Option.map mk_combI (red_of_afun af "C"); val const_conv = Option.map mk_combI (red_of_afun af "K"); val dup_conv = Option.map mk_combI (red_of_afun af "W"); val rewr_subst_ap = rewr_subst_ap ctxt af_inst; fun extract_comb n thm = Pure (thm |> Thm.rhs_of |> funpow n Thm.dest_arg1 |> Thm.term_of); fun refl_step tt = (tt, Thm.reflexive (Thm.cterm_of ctxt (term_of_apterm ctxt af_inst tt))); fun comb2_step def (tt1, thm1) (tt2, thm2) = let val thm = rewr_subst_ap def thm1 thm2; in (Ap (Ap (extract_comb 3 thm, tt1), tt2), thm) end; val B_step = comb2_step comp_conv; fun swap_B_step (tt1, thm1) thm2 = let val thm3 = rewr_subst_ap ichng_conv thm1 thm2; val thm4 = Thm.transitive thm3 (Conv.rewr_conv comp_conv (Thm.rhs_of thm3)); in (Ap (Ap (extract_comb 3 thm4, extract_comb 1 thm3), tt1), thm4) end; fun I_step tm = let val thm = Conv.rewr_conv id_conv (Thm.cterm_of ctxt tm) in (extract_comb 1 thm, thm) end; fun W_step s1 s2 = let val (Ap (Ap (tt1, tt2), tt3), thm1) = B_step s1 s2; val thm2 = Conv.rewr_conv comp_conv (Thm.rhs_of thm1 |> funpow 2 Thm.dest_arg1); val thm3 = merge_pures ctxt af_inst hom_conv tt3 |> the; val (tt4, thm4) = swap_B_step (Ap (Ap (extract_comb 3 thm2, tt1), tt2), thm2) thm3; val var = Thm.rhs_of thm1 |> Thm.dest_arg; val thm5 = rewr_subst_ap (the dup_conv) thm4 (Thm.reflexive var); val thm6 = Thm.transitive thm1 thm5; in (Ap (extract_comb 2 thm6, tt4), thm6) end; fun S_step s1 s2 = let val (Ap (Ap (tt1, tt2), tt3), thm1) = comb2_step (the flip_conv) s1 s2; val thm2 = Conv.rewr_conv comp_conv (Thm.rhs_of thm1 |> Thm.dest_arg1); val var = Thm.rhs_of thm1 |> Thm.dest_arg; val thm3 = rewr_subst_ap (the dup_conv) thm2 (Thm.reflexive var); val thm4 = Thm.transitive thm1 thm3; val tt = Ap (extract_comb 2 thm4, Ap (Ap (extract_comb 3 thm2, Ap (tt1, tt2)), tt3)); in (tt, thm4) end; fun K_step tt tm = let val ct = Thm.cterm_of ctxt tm; val T_opt = Term.fastype_of tm |> dest_type ctxt af_inst |> Option.map (Thm.ctyp_of ctxt); val thm = Thm.instantiate' [T_opt] [SOME ct] (Conv.rewr_conv (the const_conv) (term_of_apterm ctxt af_inst tt |> Thm.cterm_of ctxt)) in (Ap (extract_comb 2 thm, tt), thm) end; fun unreachable _ = raise ASSERT "eliminate: assertion failed"; fun elim (Pure _) = unreachable () | elim (ApVar (i, t)) = if exists (fn x => x = i) v then I_step t else unreachable () | elim (Ap (t1, t2)) = (case (occurs_any v t1, occurs_any v t2) of (false, false) => unreachable () | (false, true) => B_step (refl_step t1) (elim t2) | (true, false) => (case merge_pures ctxt af_inst hom_conv t2 of SOME thm => swap_B_step (elim t1) thm | NONE => comb2_step (the flip_conv) (elim t1) (refl_step t2)) | (true, true) => if is_some flip_conv then S_step (elim t1) (elim t2) else W_step (elim t1) (elim t2)); in if occurs_any v tt then elim tt else K_step tt v_tm end; (*convert a pair of terms into equal canonical forms, modulo pure terms*) fun general_normalform_conv ctxt af cts = let val (t1, t2) = apply2 (Thm.term_of) cts; val maxidx = Int.max (apply2 Thm.maxidx_of_cterm cts); (* TODO: is there a better strategy for finding the instantiated functor? *) val af_inst = match_afun_inst ctxt af (t1, maxidx); val ((apt1, apt2), _) = 0 |> apterm_of_term ctxt af_inst t1 ||>> apterm_of_term ctxt af_inst t2; val vs = consolidate ctxt af (apt1, apt2); val merge_thm = mk_meta_eq (#hom (thms_of_afun af)); fun elim_all tt [] = the (merge_pures ctxt af_inst merge_thm tt) | elim_all tt (v::vs) = let val (tt', rule1) = eliminate ctxt (af, af_inst) tt v; val rule2 = elim_all tt' vs; val (_, vartm) = dest_comb ctxt af_inst (Thm.term_of (Thm.rhs_of rule1)); val rule3 = ap_cong ctxt af_inst rule2 (Thm.reflexive (Thm.cterm_of ctxt vartm)); in Thm.transitive rule1 rule3 end; in (elim_all apt1 vs, elim_all apt2 vs) end; val general_normalize_rel_tac = binop_par_conv_tac oo general_normalform_conv; (* Reduce canonical forms to base relation *) fun rename_params names i st = let val (_, Bs, Bi, C) = Thm.dest_state (st, i); val Bi' = Logic.list_rename_params names Bi; in Thm.renamed_prop (Logic.list_implies (Bs @ [Bi'], C)) st end; (* R' (pure f <> x1 <> ... <> xn) (pure g <> x1 <> ... <> xn) ===> !!y1 ... yn. [| yi : setF xi ... |] ==> R (f y1 ... yn) (g y1 ... yn), where either both R and R' are equality, or R' = relF R for relator relF of the functor. The premises yi : setF xi are added only in the latter case and if the set operator is available. Succeeds if partial progress can be made. The names of the new parameters yi are derived from the arguments xi. *) fun head_cong_tac ctxt af renames = let val {rel_intros, ...} = thms_of_afun af; fun term_name tm = case AList.lookup (op aconv) renames tm of SOME n => n | NONE => term_to_vname tm; fun gather_vars' af_inst tm = case try (dest_comb ctxt af_inst) tm of SOME (t1, t2) => term_name t2 :: gather_vars' af_inst t1 | NONE => []; fun gather_vars prop = case prop of Const (@{const_name Trueprop}, _) $ (_ $ rhs) => rev (gather_vars' (match_afun_inst ctxt af (rhs, maxidx_of_term prop)) rhs) | _ => []; in SUBGOAL (fn (subgoal, i) => (REPEAT_DETERM (resolve_tac ctxt rel_intros i) THEN REPEAT_DETERM (resolve_tac ctxt [ext, @{thm rel_fun_eq_onI}] i ORELSE eresolve_tac ctxt [@{thm UNIV_E}] i) THEN PRIMITIVE (rename_params (gather_vars subgoal) i))) end; (* Forward lifting *) (* TODO: add limited support for premises, where used variables are not generalized in the conclusion *) fun forward_lift_rule ctxt af thm = let val thm = Object_Logic.rulify ctxt thm; val (af_inst, ctxt_inst) = import_afun_inst af ctxt; val (prop, ctxt_Ts) = yield_singleton Variable.importT_terms (Thm.prop_of thm) ctxt_inst; val (lhs, rhs) = prop |> HOLogic.dest_Trueprop |> HOLogic.dest_eq; val ([lhs', rhs'], ctxt_lifted) = generalize_lift_terms af_inst [lhs, rhs] ctxt_Ts; val lifted = HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs', rhs')); val (lifted', ctxt') = yield_singleton (Variable.import_terms true) lifted ctxt_lifted; fun tac {prems, context} = HEADGOAL (general_normalize_rel_tac context af THEN' head_cong_tac context af [] THEN' resolve_tac context [prems MRS thm]); val thm' = singleton (Variable.export ctxt' ctxt) (Goal.prove ctxt' [] [] lifted' tac); val thm'' = Raw_Simplifier.fold_rule ctxt (unfolds_of_afun af) thm'; in thm'' end; fun forward_lift_attrib name = Thm.rule_attribute [] (fn context => fn thm => let val af = afun_of_generic context (intern context name) (* FIXME !?!? *) in forward_lift_rule (Context.proof_of context) af thm end); (* High-level tactics *) fun unfold_wrapper_tac ctxt = AUTO_AFUNS false (fn afs => Simplifier.safe_asm_full_simp_tac (ctxt addsimps flat (map unfolds_of_afun afs))) ctxt; fun fold_wrapper_tac ctxt = AUTO_AFUN true (fold_goal_tac ctxt o unfolds_of_afun) ctxt; fun WRAPPER tac ctxt opt_af = REPEAT_DETERM o resolve_tac ctxt [@{thm allI}] THEN' Subgoal.FOCUS (fn {context = ctxt, params, ...} => let val renames = map (swap o apsnd Thm.term_of) params in AUTO_AFUNS false (EVERY' o map (afun_unfold_tac ctxt)) ctxt opt_af 1 THEN AUTO_AFUN true (fn af => afun_unfold_tac ctxt af THEN' CONVERSION Drule.beta_eta_conversion THEN' tac ctxt af THEN' head_cong_tac ctxt af renames) ctxt opt_af 1 end) ctxt THEN' Raw_Simplifier.rewrite_goal_tac ctxt [Drule.triv_forall_equality]; val normalize_wrapper_tac = WRAPPER normalize_rel_tac; val lifting_wrapper_tac = WRAPPER general_normalize_rel_tac; val parse_opt_afun = Scan.peek (fn context => Scan.option Parse.name >> Option.map (intern context #> afun_of_generic context)); (** Declaration **) (* Combinator setup *) fun declare_combinators combs phi = let val (names, thms) = split_list combs; val thms' = map (Morphism.thm phi) thms; fun add_combs (defs, rules) = (fold (Symtab.insert (K false)) (names ~~ thms') defs, rules); in Data.map (map_data add_combs I I) end; val setup_combinators = Local_Theory.declaration {syntax = false, pervasive = false} o declare_combinators; fun combinator_of_red thm = let val (lhs, _) = Logic.dest_equals (Thm.prop_of thm); val (head, _) = strip_comb lhs; in #1 (dest_Const head) end; fun register_combinator_rule weak_premises thm context = let val (lhs, rhs) = Logic.dest_equals (Thm.prop_of thm); val ltvars = Term.add_tvars lhs []; val rtvars = Term.add_tvars rhs []; val _ = if exists (not o member op = ltvars) rtvars then Pretty.breaks [Pretty.str "Combinator equation", Pretty.quote (Syntax.pretty_term (Context.proof_of context) (Thm.prop_of thm)), Pretty.str "has additional type variables on right-hand side."] |> Pretty.block |> Pretty.string_of |> error else (); val (defs, _) = #combinators (Data.get context); val comb_names = Symtab.make (map (fn (name, thm) => (combinator_of_red thm, name)) (Symtab.dest defs)); val rule = mk_combinator_rule comb_names weak_premises thm; fun add_rule (defs, rules) = (defs, insert eq_combinator_rule rule rules); in Data.map (map_data add_rule I I) context end; val combinator_rule_attrib = Thm.declaration_attribute o register_combinator_rule; (* Derivation of combinator reductions *) fun combinator_closure rules have_weak combs = let fun apply rule (cs, changed) = if not (Ord_List.member fast_string_ord cs (conclusion_of_rule rule)) andalso is_applicable_rule rule have_weak (fn prems => Ord_List.subset fast_string_ord (prems, cs)) then (Ord_List.insert fast_string_ord (conclusion_of_rule rule) cs, true) else (cs, changed); fun loop cs = (case fold apply rules (cs, false) of (cs', true) => loop cs' | (_, false) => cs); in loop combs end; fun derive_combinator_red ctxt af_inst red_thms (base_thm, eq_thm) = let val base_prop = Thm.prop_of base_thm; val tvars = Term.add_tvars base_prop []; val (Ts, ctxt_Ts) = mk_TFrees_of (length tvars) (inner_sort_of af_inst) ctxt; val base_prop' = base_prop |> Term_Subst.instantiate (TVars.make (tvars ~~ Ts), Vars.empty); val (lhs, rhs) = Logic.dest_equals base_prop'; val ([lhs', rhs'], ctxt') = generalize_lift_terms af_inst [lhs, rhs] ctxt_Ts; val lifted_prop = (lhs', rhs') |> HOLogic.mk_eq |> HOLogic.mk_Trueprop; val unfold_comb_conv = HOLogic.Trueprop_conv (HOLogic.eq_conv (Conv.top_sweep_conv (K (Conv.rewr_conv eq_thm)) ctxt') Conv.all_conv); val tac = HEADGOAL (CONVERSION unfold_comb_conv THEN' Raw_Simplifier.rewrite_goal_tac ctxt' red_thms THEN' resolve_tac ctxt' [@{thm refl}]); in singleton (Variable.export ctxt' ctxt) (Goal.prove ctxt' [] [] lifted_prop (K tac)) end; (*derive all instantiations with pure terms which can be simplified by homomorphism*) (*FIXME: more of a workaround than a sensible solution*) fun weak_red_closure ctxt (af_inst, merge_thm) strong_red = let val (lhs, _) = Thm.prop_of strong_red |> Logic.dest_equals; val vars = rev (Term.add_vars lhs []); fun closure [] prev thms = (prev::thms) | closure ((v, af_T)::vs) prev thms = (case try (dest_type ctxt af_inst) af_T of NONE => closure vs prev thms | SOME T_opt => let val (T, ctxt') = (case T_opt of NONE => yield_singleton Variable.invent_types (inner_sort_of af_inst) ctxt |>> TFree | SOME T => (T, ctxt)); val (v', ctxt'') = mk_Free (#1 v) T ctxt'; val pure_v = Thm.cterm_of ctxt'' (lift_term af_inst v'); - val next = Drule.instantiate_normalize ([], [((v, af_T), pure_v)]) prev; + val next = Drule.instantiate_normalize (TVars.empty, Vars.make [((v, af_T), pure_v)]) prev; val next' = Raw_Simplifier.rewrite_rule ctxt'' [merge_thm] next; val next'' = singleton (Variable.export ctxt'' ctxt) next'; in closure vs next'' (prev::thms) end); in closure vars strong_red [] end; fun combinator_red_closure ctxt (comb_defs, rules) (af_inst, merge_thm) weak_reds combs = let val have_weak = not (null weak_reds); val red_thms0 = Symtab.fold (fn (_, thm) => cons (mk_meta_eq thm)) combs weak_reds; val red_thms = flat (map (weak_red_closure ctxt (af_inst, merge_thm)) red_thms0); fun apply rule ((cs, rs), changed) = if not (Symtab.defined cs (conclusion_of_rule rule)) andalso is_applicable_rule rule have_weak (forall (Symtab.defined cs)) then let val conclusion = conclusion_of_rule rule; val def = the (Symtab.lookup comb_defs conclusion); val new_red_thm = derive_combinator_red ctxt af_inst rs (def, thm_of_rule rule); val new_red_thms = weak_red_closure ctxt (af_inst, merge_thm) (mk_meta_eq new_red_thm); in ((Symtab.update (conclusion, new_red_thm) cs, new_red_thms @ rs), true) end else ((cs, rs), changed); fun loop xs = (case fold apply rules (xs, false) of (xs', true) => loop xs' | (_, false) => xs); in #1 (loop (combs, red_thms)) end; (* Preparation of AFun data *) fun mk_terms ctxt (raw_pure, raw_ap, raw_rel, raw_set) = let val thy = Proof_Context.theory_of ctxt; val show_typ = quote o Syntax.string_of_typ ctxt; val show_term = quote o Syntax.string_of_term ctxt; fun closed_poly_term t = let val poly_t = singleton (Variable.polymorphic ctxt) t; in case Term.add_vars (singleton (Variable.export_terms (Proof_Context.augment t ctxt) ctxt) t) [] of [] => (case (Term.hidden_polymorphism poly_t) of [] => poly_t | _ => error ("Hidden type variables in term " ^ show_term t)) | _ => error ("Locally free variables in term " ^ show_term t) end; val pure = closed_poly_term raw_pure; val (tvar, T1) = fastype_of pure |> dest_funT |>> dest_TVar handle TYPE _ => error ("Bad type for pure: " ^ show_typ (fastype_of pure)); val maxidx_pure = maxidx_of_term pure; val ap = Logic.incr_indexes ([], [], maxidx_pure + 1) (closed_poly_term raw_ap); fun bad_ap _ = error ("Bad type for ap: " ^ show_typ (fastype_of ap)); val (T23, (T2, T3)) = fastype_of ap |> dest_funT ||> dest_funT handle TYPE _ => bad_ap (); val maxidx_common = Term.maxidx_term ap maxidx_pure; (*unify type variables, while keeping the live variables separate*) fun no_unifier (T, U) = error ("Unable to infer common functor type from " ^ commas (map show_typ [T, U])); fun unify_ap_type T (tyenv, maxidx) = let val argT = TVar ((Name.aT, maxidx + 1), []); val T1' = Term_Subst.instantiateT (TVars.make [(tvar, argT)]) T1; val (tyenv', maxidx') = Sign.typ_unify thy (T1', T) (tyenv, maxidx + 1) handle Type.TUNIFY => no_unifier (T1', T); in (argT, (tyenv', maxidx')) end; val (ap_args, (ap_env, maxidx_env)) = fold_map unify_ap_type [T2, T3, T23] (Vartab.empty, maxidx_common); val [T2_arg, T3_arg, T23_arg] = map (Envir.norm_type ap_env) ap_args; val (tvar2, tvar3) = (dest_TVar T2_arg, dest_TVar T3_arg) handle TYPE _ => bad_ap (); val _ = if T23_arg = T2_arg --> T3_arg then () else bad_ap (); val sort = foldl1 (Sign.inter_sort thy) (map #2 [tvar, tvar2, tvar3]); val _ = Sign.of_sort thy (Term.aT sort --> Term.aT sort, sort) orelse error ("Sort constraint " ^ quote (Syntax.string_of_sort ctxt sort) ^ " not closed under function types"); fun update_sort (v, S) (tyenv, maxidx) = (Vartab.update_new (v, (S, TVar ((Name.aT, maxidx + 1), sort))) tyenv, maxidx + 1); val (common_env, _) = fold update_sort [tvar, tvar2, tvar3] (ap_env, maxidx_env); val tvar' = Envir.norm_type common_env (TVar tvar); val pure' = norm_term_types common_env pure; val (tvar2', tvar3') = apply2 (Envir.norm_type common_env) (T2_arg, T3_arg); val ap' = norm_term_types common_env ap; fun bad_set set = error ("Bad type for set: " ^ show_typ (fastype_of set)); fun mk_set set = let val tyenv = Sign.typ_match thy (domain_type (fastype_of set), range_type (fastype_of pure')) Vartab.empty handle Type.TYPE_MATCH => bad_set set; val set' = Envir.subst_term_types tyenv set; val set_tvar = fastype_of set' |> range_type |> HOLogic.dest_setT |> dest_TVar handle TYPE _ => bad_set set; val _ = if Term.eq_tvar (dest_TVar tvar', set_tvar) then () else bad_set set; in ([tvar'], set') end val set = (case raw_set of NONE => ([tvar'], Abs ("x", tvar', HOLogic.mk_UNIV tvar')) | SOME t => mk_set (closed_poly_term t)); val terms = Term_Subst.zero_var_indexes (pack_poly_terms [poly_type_to_term ([tvar'], range_type (fastype_of pure')), ([tvar'], pure'), ([tvar2', tvar3'], ap'), set]); (*TODO: also infer the relator type?*) fun bad_rel rel = error ("Bad type for rel: " ^ show_typ (fastype_of rel)); fun mk_rel rel = let val ((T1, T2), (T1_af, T2_af)) = fastype_of rel |> dest_funT |>> BNF_Util.dest_pred2T ||> BNF_Util.dest_pred2T; val _ = (dest_TVar T1; dest_TVar T2); val _ = if T1 = T2 then bad_rel rel else (); val af_inst = mk_afun_inst (match_poly_terms_type ctxt (terms, 0) (T1_af, maxidx_of_term rel)); val (T1', T2') = apply2 (dest_type ctxt af_inst) (T1_af, T2_af); val _ = if (is_none T1' andalso is_none T2') orelse (T1' = SOME T1 andalso T2' = SOME T2) then () else bad_rel rel; in Term_Subst.zero_var_indexes (pack_poly_terms [([T1, T2], rel)]) end handle TYPE _ => bad_rel rel; val rel = Option.map (mk_rel o closed_poly_term) raw_rel; in (terms, rel) end; fun mk_rel_intros {pure_transfer, ap_rel_fun} = let val pure_rel_intro = pure_transfer RS @{thm rel_funD}; in [pure_rel_intro, ap_rel_fun] end; fun mk_afun_thms ctxt af_inst (hom_thm, ichng_thm, reds, rel_axioms) = let val pure_comp_conv = let val ([T1, T2, T3], ctxt_Ts) = mk_TFrees_of 3 (inner_sort_of af_inst) ctxt; val (((g, f), x), ctxt') = ctxt_Ts |> mk_Free "g" (T2 --> T3) ||>> mk_Free "f" (mk_type af_inst (T1 --> T2)) ||>> mk_Free "x" (mk_type af_inst T1); val comb = mk_comb af_inst; val lhs = comb (T2 --> T3) (lift_term af_inst g, comb (T1 --> T2) (f, x)); val B_g = Abs ("f", T1 --> T2, Abs ("x", T1, Term.betapply (g, Bound 1 $ Bound 0))); val rhs = comb (T1 --> T3) (comb ((T1 --> T2) --> T1 --> T3) (lift_term af_inst B_g, f), x); val prop = HOLogic.mk_eq (lhs, rhs) |> HOLogic.mk_Trueprop; val merge_rule = mk_meta_eq hom_thm; val B_intro = the (Symtab.lookup reds "B") |> mk_meta_eq |> Thm.symmetric; val tac = HEADGOAL (Raw_Simplifier.rewrite_goal_tac ctxt' [B_intro, merge_rule] THEN' resolve_tac ctxt' [@{thm refl}]); in singleton (Variable.export ctxt' ctxt) (Goal.prove ctxt' [] [] prop (K tac)) end; val eq_intros = let val ([T1, T2], ctxt_Ts) = mk_TFrees_of 2 (inner_sort_of af_inst) ctxt; val T12 = mk_type af_inst (T1 --> T2); val (((((x, y), x'), f), g), ctxt') = ctxt_Ts |> mk_Free "x" T1 ||>> mk_Free "y" T1 ||>> mk_Free "x" (mk_type af_inst T1) ||>> mk_Free "f" T12 ||>> mk_Free "g" T12; val pure_fun = mk_pure af_inst T1; val pure_cong = Drule.infer_instantiate' ctxt' (map (SOME o Thm.cterm_of ctxt') [x, y, pure_fun]) @{thm arg_cong}; val ap_fun = mk_ap af_inst (T1, T2); val ap_cong1 = Drule.infer_instantiate' ctxt' (map (SOME o Thm.cterm_of ctxt') [f, g, ap_fun, x']) @{thm arg1_cong}; in Variable.export ctxt' ctxt [pure_cong, ap_cong1] end; val rel_intros = case rel_axioms of NONE => [] | SOME axioms => mk_rel_intros axioms; in {hom = hom_thm, ichng = ichng_thm, reds = reds, rel_thms = rel_axioms, rel_intros = eq_intros @ rel_intros, pure_comp_conv = pure_comp_conv} end; fun reuse_TFrees n S (ctxt, Ts) = let val have_n = Int.min (n, length Ts); val (more_Ts, ctxt') = mk_TFrees_of (n - have_n) S ctxt; in (take have_n Ts @ more_Ts, (ctxt', Ts @ more_Ts)) end; fun mk_comb_prop lift_pos thm af_inst ctxt_Ts = let val base = Thm.prop_of thm; val tvars = Term.add_tvars base []; val (Ts, (ctxt', Ts')) = reuse_TFrees (length tvars) (inner_sort_of af_inst) ctxt_Ts; val base' = base |> Term_Subst.instantiate (TVars.make (tvars ~~ Ts), Vars.empty); val (lhs, rhs) = Logic.dest_equals base'; val (_, lhs_args) = strip_comb lhs; val lift_var = Var o apsnd (mk_type af_inst) o dest_Var; val (lhs_args', subst) = fold_index (fn (i, v) => if member (op =) lift_pos i then apfst (cons v) else map_prod (cons (lift_var v)) (cons (v, lift_var v))) lhs_args ([], []); val (lhs', rhs') = apply2 (subst_lift_term af_inst subst) (lhs, rhs); val lifted = (lhs', rhs') |> HOLogic.mk_eq |> HOLogic.mk_Trueprop; in (fold Logic.all lhs_args' lifted, (ctxt', Ts')) end; fun mk_homomorphism_prop af_inst ctxt_Ts = let val ([T1, T2], (ctxt', Ts')) = reuse_TFrees 2 (inner_sort_of af_inst) ctxt_Ts; val ((f, x), _) = ctxt' |> mk_Free "f" (T1 --> T2) ||>> mk_Free "x" T1; val lhs = mk_comb af_inst (T1 --> T2) (lift_term af_inst f, lift_term af_inst x); val rhs = lift_term af_inst (f $ x); val prop = HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs)); in (Logic.all f (Logic.all x prop), (ctxt', Ts')) end; fun mk_interchange_prop af_inst ctxt_Ts = let val ([T1, T2], (ctxt', Ts')) = reuse_TFrees 2 (inner_sort_of af_inst) ctxt_Ts; val ((f, x), _) = ctxt' |> mk_Free "f" (mk_type af_inst (T1 --> T2)) ||>> mk_Free "x" T1; val lhs = mk_comb af_inst (T1 --> T2) (f, lift_term af_inst x); val T_x = Abs ("f", T1 --> T2, Bound 0 $ x); val rhs = mk_comb af_inst ((T1 --> T2) --> T2) (lift_term af_inst T_x, f); val prop = HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs)); in (Logic.all f (Logic.all x prop), (ctxt', Ts')) end; fun mk_rel_props (af_inst, rel_inst) ctxt_Ts = let fun mk_af_rel tm = let val (T1, T2) = BNF_Util.dest_pred2T (fastype_of tm); in betapply (instantiate_poly_term rel_inst [T1, T2], tm) end; val ([T1, T2, T3], (ctxt', Ts')) = reuse_TFrees 3 (inner_sort_of af_inst) ctxt_Ts; val (pure_R, _) = mk_Free "R" (T1 --> T2 --> @{typ bool}) ctxt'; val rel_pure = BNF_Util.mk_rel_fun pure_R (mk_af_rel pure_R) $ mk_pure af_inst T1 $ mk_pure af_inst T2; val pure_prop = Logic.all pure_R (HOLogic.mk_Trueprop rel_pure); val ((((f, g), x), ap_R), _) = ctxt' |> mk_Free "f" (mk_type af_inst (T1 --> T2)) ||>> mk_Free "g" (mk_type af_inst (T1 --> T3)) ||>> mk_Free "x" (mk_type af_inst T1) ||>> mk_Free "R" (T2 --> T3 --> @{typ bool}); val fun_rel = BNF_Util.mk_rel_fun (mk_eq_on (mk_set af_inst T1 $ x)) ap_R; val rel_ap = Logic.mk_implies (HOLogic.mk_Trueprop (mk_af_rel fun_rel $ f $ g), HOLogic.mk_Trueprop (mk_af_rel ap_R $ mk_comb af_inst (T1 --> T2) (f, x) $ mk_comb af_inst (T1 --> T3) (g, x))); val ap_prop = fold_rev Logic.all [ap_R, f, g, x] rel_ap; in ([pure_prop, ap_prop], (ctxt', Ts')) end; fun mk_interchange ctxt ((comb_defs, _), comb_unfolds) (af_inst, merge_thm) reds = let val T_def = the (Symtab.lookup comb_defs "T"); val T_red = the (Symtab.lookup reds "T"); val (weak_prop, (ctxt', _)) = mk_comb_prop [0] T_def af_inst (ctxt, []); fun tac {context, ...} = HEADGOAL (Raw_Simplifier.rewrite_goal_tac context [Thm.symmetric merge_thm] THEN' resolve_tac context [T_red]); val weak_red = singleton (Variable.export ctxt' ctxt) (Goal.prove ctxt' [] [] weak_prop tac); in Raw_Simplifier.rewrite_rule ctxt (comb_unfolds) weak_red RS sym end; fun mk_weak_reds ctxt ((comb_defs, _), comb_unfolds) af_inst (hom_thm, ichng_thm, reds) = let val unfolded_reds = Symtab.map (K (Raw_Simplifier.rewrite_rule ctxt comb_unfolds)) reds; val af_thms = mk_afun_thms ctxt af_inst (hom_thm, ichng_thm, unfolded_reds, NONE); val af = mk_afun Binding.empty (pack_afun_inst af_inst) NONE af_thms; fun tac {context, ...} = HEADGOAL (normalize_wrapper_tac context (SOME af) THEN' Raw_Simplifier.rewrite_goal_tac context comb_unfolds THEN' resolve_tac context [refl]); fun mk comb lift_pos = let val def = the (Symtab.lookup comb_defs comb); val (prop, (ctxt', _)) = mk_comb_prop lift_pos def af_inst (ctxt, []); val hol_thm = singleton (Variable.export ctxt' ctxt) (Goal.prove ctxt' [] [] prop tac); in mk_meta_eq hol_thm end; val uncurry_thm = mk_meta_eq (forward_lift_rule ctxt af @{thm uncurry_pair}); in [mk "C" [1], mk "C" [2], uncurry_thm] end; fun mk_comb_reds ctxt combss af_inst user_combs (hom_thm, user_thms, ichng_thms) = let val ((comb_defs, comb_rules), comb_unfolds) = combss; val merge_thm = mk_meta_eq hom_thm; val user_reds = Symtab.make (user_combs ~~ user_thms); val reds0 = combinator_red_closure ctxt (comb_defs, comb_rules) (af_inst, merge_thm) [] user_reds; val ichng_thm = case ichng_thms of [] => singleton (Variable.export ctxt ctxt) (mk_interchange ctxt combss (af_inst, merge_thm) reds0) | [thm] => thm; val weak_reds = mk_weak_reds ctxt combss af_inst (hom_thm, ichng_thm, reds0); val reds1 = combinator_red_closure ctxt (comb_defs, comb_rules) (af_inst, merge_thm) weak_reds reds0; val unfold = Raw_Simplifier.rewrite_rule ctxt comb_unfolds; in (Symtab.map (K unfold) reds1, ichng_thm) end; fun note_afun_thms af = let val thms = thms_of_afun af; val named_thms = [("homomorphism", [#hom thms]), ("interchange", [#ichng thms]), ("afun_rel_intros", #rel_intros thms)] @ map (fn (name, thm) => ("pure_" ^ name ^ "_conv", [thm])) (Symtab.dest (#reds thms)) @ (case #rel_thms thms of NONE => [] | SOME rel_thms' => [("pure_transfer", [#pure_transfer rel_thms']), ("ap_rel_fun_cong", [#ap_rel_fun rel_thms'])]); val base_name = Binding.name_of (name_of_afun af); fun mk_note (name, thms) = ((Binding.qualify true base_name (Binding.name name), []), [(thms, [])]); in Local_Theory.notes (map mk_note named_thms) #> #2 end; fun register_afun af = let fun decl phi context = Data.map (fn {combinators, afuns, patterns} => let val af' = morph_afun phi af; val (name, afuns') = Name_Space.define context true (name_of_afun af', af') afuns; val patterns' = Item_Net.update (name, patterns_of_afun af') patterns; in {combinators = combinators, afuns = afuns', patterns = patterns'} end) context; in Local_Theory.declaration {syntax = false, pervasive = false} decl end; fun applicative_cmd (((((name, flags), raw_pure), raw_ap), raw_rel), raw_set) lthy = let val comb_unfolds = Named_Theorems.get lthy @{named_theorems combinator_unfold}; val comb_reprs = Named_Theorems.get lthy @{named_theorems combinator_repr}; val (comb_defs, comb_rules) = get_combinators (Context.Proof lthy); val _ = fold (fn name => if Symtab.defined comb_defs name then I else error ("Unknown combinator " ^ quote name)) flags (); val _ = if has_duplicates op = flags then warning "Ignoring duplicate combinators" else (); val user_combs0 = Ord_List.make fast_string_ord flags; val raw_pure' = Syntax.read_term lthy raw_pure; val raw_ap' = Syntax.read_term lthy raw_ap; val raw_rel' = Option.map (Syntax.read_term lthy) raw_rel; val raw_set' = Option.map (Syntax.read_term lthy) raw_set; val (terms, rel) = mk_terms lthy (raw_pure', raw_ap', raw_rel', raw_set'); val derived_combs0 = combinator_closure comb_rules false user_combs0; val required_combs = Ord_List.make fast_string_ord ["B", "I"]; val user_combs = Ord_List.union fast_string_ord user_combs0 (Ord_List.subtract fast_string_ord derived_combs0 required_combs); val derived_combs1 = combinator_closure comb_rules false user_combs; val derived_combs2 = combinator_closure comb_rules true derived_combs1; fun is_redundant comb = eq_list (op =) (derived_combs2, (combinator_closure comb_rules true (Ord_List.remove fast_string_ord comb user_combs))); val redundant_combs = filter is_redundant user_combs; val _ = if null redundant_combs then () else warning ("Redundant combinators: " ^ commas redundant_combs); val prove_interchange = not (Ord_List.member fast_string_ord derived_combs1 "T"); val (af_inst, ctxt_af) = import_afun_inst_raw terms lthy; (* TODO: reuse TFrees from above *) val (rel_insts, ctxt_inst) = (case rel of NONE => (NONE, ctxt_af) | SOME r => let val (rel_inst, ctxt') = import_poly_terms r ctxt_af |>> the_single; val T = fastype_of (#2 rel_inst) |> range_type |> domain_type; val af_inst = match_poly_terms_type ctxt' (terms, 0) (T, ~1) |> mk_afun_inst; in (SOME (af_inst, rel_inst), ctxt') end); val mk_propss = [apfst single o mk_homomorphism_prop af_inst, fold_map (fn comb => mk_comb_prop [] (the (Symtab.lookup comb_defs comb)) af_inst) user_combs, if prove_interchange then apfst single o mk_interchange_prop af_inst else pair [], if is_some rel then mk_rel_props (the rel_insts) else pair []]; val (propss, (ctxt_Ts, _)) = fold_map I mk_propss (ctxt_inst, []); fun repr_tac ctxt = Raw_Simplifier.rewrite_goals_tac ctxt comb_reprs; fun after_qed thmss lthy' = let val [[hom_thm], user_thms, ichng_thms, rel_thms] = map (Variable.export lthy' ctxt_inst) thmss; val (reds, ichng_thm) = mk_comb_reds ctxt_inst ((comb_defs, comb_rules), comb_unfolds) af_inst user_combs (hom_thm, user_thms, ichng_thms); val rel_axioms = case rel_thms of [] => NONE | [thm1, thm2] => SOME {pure_transfer = thm1, ap_rel_fun = thm2}; val af_thms = mk_afun_thms ctxt_inst af_inst (hom_thm, ichng_thm, reds, rel_axioms); val af_thms = map_afun_thms (singleton (Variable.export ctxt_inst lthy)) af_thms; val af = mk_afun name terms rel af_thms; in lthy |> register_afun af |> note_afun_thms af end; in Proof.theorem NONE after_qed ((map o map) (rpair []) propss) ctxt_Ts |> Proof.refine (Method.Basic (SIMPLE_METHOD o repr_tac)) |> Seq.the_result "" end; fun print_afuns ctxt = let fun pretty_afun (name, af) = let val [pT, (_, pure), (_, ap), (_, set)] = unpack_poly_terms (terms_of_afun af); val ([tvar], T) = poly_type_of_term pT; val rel = Option.map (#2 o the_single o unpack_poly_terms) (rel_of_afun af); val combinators = Symtab.keys (#reds (thms_of_afun af)); in Pretty.block (Pretty.fbreaks ([Pretty.block [Pretty.str name, Pretty.str ":", Pretty.brk 1, Pretty.quote (Syntax.pretty_typ ctxt T), Pretty.brk 1, Pretty.str "of", Pretty.brk 1, Syntax.pretty_typ ctxt tvar], Pretty.block [Pretty.str "pure:", Pretty.brk 1, Pretty.quote (Syntax.pretty_term ctxt pure)], Pretty.block [Pretty.str "ap:", Pretty.brk 1, Pretty.quote (Syntax.pretty_term ctxt ap)], Pretty.block [Pretty.str "set:", Pretty.brk 1, Pretty.quote (Syntax.pretty_term ctxt set)]] @ (case rel of NONE => [] | SOME rel' => [Pretty.block [Pretty.str "rel:", Pretty.brk 1, Pretty.quote (Syntax.pretty_term ctxt rel')]]) @ [Pretty.block ([Pretty.str "combinators:", Pretty.brk 1] @ Pretty.commas (map Pretty.str combinators))])) end; val afuns = sort_by #1 (Name_Space.fold_table cons (get_afun_table (Context.Proof ctxt)) []); in Pretty.writeln (Pretty.big_list "Registered applicative functors:" (map pretty_afun afuns)) end; (* Unfolding *) fun add_unfold_thm name thm context = let val (lhs, _) = Thm.prop_of thm |> HOLogic.dest_Trueprop |> HOLogic.dest_eq handle TERM _ => error "Not an equation"; val names = case name of SOME n => [intern context n] | NONE => case match_typ context (Term.fastype_of lhs) of ns as (_::_) => ns | [] => error "Unable to determine applicative functor instance"; val _ = map (afun_of_generic context) names; (*TODO: check equation*) val thm' = mk_meta_eq thm; in fold (fn n => update_afun n (add_unfolds_afun [thm'])) names context end; fun add_unfold_attrib name = Thm.declaration_attribute (add_unfold_thm name); (*TODO: attribute to delete unfolds*) end; diff --git a/thys/Auto2_HOL/logic_steps.ML b/thys/Auto2_HOL/logic_steps.ML --- a/thys/Auto2_HOL/logic_steps.ML +++ b/thys/Auto2_HOL/logic_steps.ML @@ -1,1081 +1,1081 @@ (* File: logic_steps.ML Author: Bohua Zhan Core (logic) proofsteps. *) signature LOGIC_PROOFSTEPS = sig (* General logic *) val shadow_prop_item: proofstep val shadow_term_item: proofstep val exists_elim_prfstep: proofstep val add_logic_proofsteps: theory -> theory val mk_all_disj: term list * term list -> term val strip_all_disj: term -> term list * term list val norm_all_disj: Proof.context -> conv val replace_disj_vars: Proof.context -> term list * term list -> term list * term list val disj_prop_match: Proof.context -> id_inst -> term * (term list * term list) * ((indexname * typ) list * cterm list) -> id_inst_th list val norm_conj: conv (* DISJ items. *) val TY_DISJ: string val disj_to_ritems: bool -> term -> thm -> raw_item list val disj_to_update: bool -> term -> box_id * int option * thm -> raw_update val dest_tname_of_disj: cterm list -> term * cterm list val is_match_prem_only: box_item -> bool val analyze_disj_th: Proof.context -> thm -> term * thm val disj_rewr_terms: term list -> term list val output_disj_fn: item_output val disj_prop_matcher: item_matcher val reduce_disj_True: conv val match_update_prfstep: proofstep val match_one_sch_prfstep: proofstep val disj_match_iff_prfstep: proofstep val disj_create_case_prfstep: proofstep val disj_shadow_prfstep: proofstep val add_disj_proofsteps: theory -> theory (* Normalizers *) val split_not_imp_th: thm -> thm list val split_conj_gen_th: Proof.context -> thm -> thm list val eq_normalizer: normalizer val property_normalizer: normalizer val disj_normalizer: normalizer val logic_thm_update: Proof.context -> box_id * thm -> raw_update val add_disj_normalizers: theory -> theory end; structure Logic_ProofSteps : LOGIC_PROOFSTEPS = struct fun boolVar s = Var ((s, 0), boolT) (* Shadowing based on equivalence. For both PROP and TERM items, shadowing is based on subterm equivalence, skipping any Not (~) at head. *) fun shadow_item_fn ctxt item1 item2 = if #sc item1 = 0 andalso #sc item2 = 0 then [] else let val id = BoxItem.merged_id ctxt [item1, item2] val _ = assert (forall (BoxItem.match_ty_strs [TY_TERM, TY_PROP]) [item1, item2]) "shadow_item_fn" val (tname1, tname2) = (the_single (#tname item1), the_single (#tname item2)) handle List.Empty => raise Fail "shadow_item_fn" val (t1, t2) = (Thm.term_of tname1, Thm.term_of tname2) val (lhs, rhs) = if fastype_of t1 = boolT andalso is_neg t1 andalso fastype_of t2 = boolT andalso is_neg t2 then (UtilLogic.get_cneg tname1, UtilLogic.get_cneg tname2) else (tname1, tname2) val equiv_ids = (RewriteTable.subequiv_info ctxt id (lhs, rhs)) |> map fst |> filter BoxID.has_incr_id |> Util.max_partial (BoxID.is_eq_ancestor ctxt) val item_to_shadow = if #uid item1 > #uid item2 then item1 else item2 fun process_id id' = ShadowItem {id = id', item = item_to_shadow} in map process_id equiv_ids end val shadow_prop_item = {name = "shadow_prop", args = [TypedUniv TY_PROP, TypedUniv TY_PROP], func = TwoStep shadow_item_fn} val shadow_term_item = {name = "shadow_term", args = [TypedUniv TY_TERM, TypedUniv TY_TERM], func = TwoStep shadow_item_fn} fun eq_abs_fn ctxt item1 item2 = let val {id = id1, tname = tname1, ...} = item1 val {id = id2, tname = tname2, ...} = item2 val id = BoxID.merge_boxes ctxt (id1, id2) val (ct1, ct2) = apply2 the_single (tname1, tname2) val (t1, t2) = apply2 Thm.term_of (ct1, ct2) in if not (Util.is_abs t1) orelse not (Util.is_abs t2) then [] else if RewriteTable.is_equiv id ctxt (ct1, ct2) then [] else if Term_Ord.term_ord (t2, t1) = LESS then [] else let fun process_equiv (id', eq_th) = let val (lhs, rhs) = (Util.lhs_of eq_th, Util.rhs_of eq_th) in AddItems {id = id', sc = SOME 1, raw_items = [Fact (TY_EQ, [lhs, rhs], eq_th)]} end in (Matcher.rewrite_match ctxt (t1, ct2) (id, fo_init)) |> map (fn ((id, _), eq_th) => (id, eq_th)) |> filter (BoxID.has_incr_id o fst) |> Util.max_partial (BoxID.id_is_eq_ancestor ctxt) |> map process_equiv end end val eq_abs_prfstep = {name = "eq_abs", args = [TypedUniv TY_TERM, TypedUniv TY_TERM], func = TwoStep eq_abs_fn} (* Given an assumption of the form EX x. A, we produce an assumption A with x in A replaced by a free variable. To avoid name collisions, when the update is produced x is replaced by an "internal" free variable, with suffix '_'. When the update is applied, that internal free variable is replaced by a fresh variable as determined by the context. We produce at most two variables at a time. *) fun exists_elim_fn ctxt {id, prop, ...} = if not (BoxID.has_incr_id id) then [] else let val t = prop_of' prop in if is_ex t orelse is_bex t orelse (is_neg t andalso is_obj_all (dest_not t)) orelse (is_neg t andalso is_ball (dest_not t)) then let val (ritems, th') = prop |> apply_to_thm' (UtilLogic.normalize_exists ctxt) |> Update.apply_exists_ritems ctxt in [AddItems {id = id, sc = NONE, raw_items = ritems @ [Update.thm_to_ritem th']}] end else [] end val exists_elim_prfstep = {name = "exists_elim", args = [TypedUniv TY_PROP], func = OneStep exists_elim_fn} val add_logic_proofsteps = fold add_prfstep [ shadow_prop_item, shadow_term_item, exists_elim_prfstep, eq_abs_prfstep ] (* Given (x_1 ... x_n, A_1 ... A_n), create the corresponding forall-disj term. *) fun mk_all_disj (vars, terms) = case vars of [] => list_disj terms | var :: vars' => mk_obj_all var (mk_all_disj (vars', terms)) (* Normalize t into a disjunction of terms. *) fun strip_disj t = if is_disj t then maps strip_disj [dest_arg1 t, dest_arg t] else if is_imp t then maps strip_disj [get_neg (dest_arg1 t), dest_arg t] else if is_neg t then let val t' = dest_not t in if is_neg t' then strip_disj (dest_not t') else if is_conj t' then maps strip_disj [get_neg (dest_arg1 t'), get_neg (dest_arg t')] else [t] end else [t] (* Normalize a term into the form !x_1 ... x_n. A_1 | ... | A_n *) fun strip_all_disj t = if is_obj_all t then case t of _ $ Abs (abs as (_, T, _)) => let val (x, body) = Term.dest_abs abs val var = Free (x, T) val (vars, disjs) = strip_all_disj body in if exists (Util.occurs_free var) disjs then (var :: vars, disjs) else (vars, disjs) end | f $ arg => strip_all_disj (f $ UtilLogic.force_abs_form arg) | _ => raise Fail "strip_all_disj" else if is_ball t then case t of _ $ S $ Abs (abs as (_, T, _)) => let val (x, body) = Term.dest_abs abs val var = Free (x, T) val mem = mk_mem (var, S) val (vars, disjs) = strip_all_disj body in (var :: vars, get_neg mem :: disjs) end | f $ S $ arg => strip_all_disj (f $ S $ UtilLogic.force_abs_form arg) | _ => raise Fail "strip_all_disj" else if is_neg t andalso is_ex (dest_not t) then case dest_not t of _ $ Abs (abs as (_, T, _)) => let val (x, body) = Term.dest_abs abs val var = Free (x, T) val (vars, disjs) = strip_all_disj (get_neg body) in if exists (Util.occurs_free var) disjs then (var :: vars, disjs) else (vars, disjs) end | f $ arg => strip_all_disj (get_neg (f $ UtilLogic.force_abs_form arg)) | _ => raise Fail "strip_all_disj" else if is_neg t andalso is_bex (dest_not t) then case dest_not t of _ $ S $ Abs (abs as (_, T, _)) => let val (x, body) = Term.dest_abs abs val var = Free (x, T) val mem = mk_mem (var, S) val (vars, disjs) = strip_all_disj (get_neg body) in (var :: vars, get_neg mem :: disjs) end | f $ S $ arg => strip_all_disj (get_neg (f $ S $ UtilLogic.force_abs_form arg)) | _ => raise Fail "strip_all_disj" else if is_disj t then let val (v1, ts1) = strip_all_disj (dest_arg1 t) val (v2, ts2) = strip_all_disj (dest_arg t) in (v1 @ v2, ts1 @ ts2) end else if is_imp t then let val (v1, ts1) = strip_all_disj (get_neg (dest_arg1 t)) val (v2, ts2) = strip_all_disj (dest_arg t) in (v1 @ v2, ts1 @ ts2) end else if is_neg t then let val t' = dest_not t in if is_neg t' then strip_all_disj (dest_not t') else if is_conj t' then let val (v1, ts1) = strip_all_disj (get_neg (dest_arg1 t')) val (v2, ts2) = strip_all_disj (get_neg (dest_arg t')) in (v1 @ v2, ts1 @ ts2) end else ([], [t]) end else ([], [t]) (* Normalize (A_1 | A_2 | ... | A_m) | (B_1 | B_2 | ... | B_n) *) fun norm_disj_clauses ct = let val (arg1, _) = Util.dest_binop_args (Thm.term_of ct) in if is_disj arg1 then Conv.every_conv [rewr_obj_eq UtilBase.disj_assoc_th, Conv.arg_conv norm_disj_clauses] ct else Conv.all_conv ct end (* Normalize ct. *) fun norm_disj ct = let val t = Thm.term_of ct val _ = assert (fastype_of t = boolT) "norm_disj: wrong type" in if is_disj t then Conv.every_conv [Conv.binop_conv norm_disj, norm_disj_clauses] ct else if is_imp t then Conv.every_conv [rewr_obj_eq UtilBase.imp_conv_disj_th, norm_disj] ct else if is_neg t andalso is_neg (dest_not t) then Conv.every_conv [rewr_obj_eq UtilBase.nn_cancel_th, norm_disj] ct else if is_neg t andalso is_conj (dest_not t) then Conv.every_conv [rewr_obj_eq UtilBase.de_Morgan_conj_th, norm_disj] ct else Conv.all_conv ct end (* Normalize to forall at the top-level *) fun norm_all_disj ctxt ct = let val t = Thm.term_of ct in if is_obj_all t then if not (Util.is_subterm (Bound 0) (dest_arg t)) then Conv.every_conv [rewr_obj_eq UtilBase.all_trivial_th, norm_all_disj ctxt] ct else Conv.every_conv [Conv.binder_conv (norm_all_disj o snd) ctxt] ct else if is_ball t then Conv.every_conv [rewr_obj_eq UtilBase.Ball_def_th, norm_all_disj ctxt] ct else if is_neg t andalso is_ex (dest_not t) then Conv.every_conv [rewr_obj_eq UtilBase.not_ex_th, norm_all_disj ctxt] ct else if is_neg t andalso is_bex (dest_not t) then Conv.every_conv [ Conv.arg_conv (rewr_obj_eq UtilBase.Bex_def_th), norm_all_disj ctxt] ct else if is_disj t then let val eq1 = Conv.binop_conv (norm_all_disj ctxt) ct val rhs = Util.rhs_of eq1 in if is_obj_all (dest_arg1 rhs) then Conv.every_conv [Conv.rewr_conv eq1, rewr_obj_eq UtilBase.disj_commute_th, rewr_obj_eq UtilBase.swap_all_disj_th, norm_all_disj ctxt] ct else if is_obj_all (dest_arg rhs) then Conv.every_conv [Conv.rewr_conv eq1, rewr_obj_eq UtilBase.swap_all_disj_th, norm_all_disj ctxt] ct else Conv.every_conv [Conv.rewr_conv eq1, norm_disj_clauses] ct end else if is_imp t then Conv.every_conv [rewr_obj_eq UtilBase.imp_conv_disj_th, norm_all_disj ctxt] ct else if is_neg t andalso is_neg (dest_not t) then Conv.every_conv [rewr_obj_eq UtilBase.nn_cancel_th, norm_all_disj ctxt] ct else if is_neg t andalso is_conj (dest_not t) then Conv.every_conv [rewr_obj_eq UtilBase.de_Morgan_conj_th, norm_all_disj ctxt] ct else Conv.all_conv ct end fun mk_disj_eq eq_ths = case eq_ths of [] => raise Fail "mk_disj_eq" | [eq] => eq | eq :: eqs' => Drule.binop_cong_rule UtilBase.cDisj eq (mk_disj_eq eqs') (* Sort A | (B_1 | B_2 | ... | B_n), assuming the right side is sorted. *) fun sort_disj_clause_aux ct = if not (is_disj (Thm.term_of ct)) then Conv.all_conv ct else let val (arg1, arg2) = Util.dest_binop_args (Thm.term_of ct) in if is_disj arg2 then if Term_Ord.term_ord (dest_arg1 arg2, arg1) = LESS then Conv.every_conv [rewr_obj_eq (obj_sym UtilBase.disj_assoc_th), Conv.arg1_conv (rewr_obj_eq UtilBase.disj_commute_th), rewr_obj_eq UtilBase.disj_assoc_th, Conv.arg_conv sort_disj_clause_aux] ct else Conv.all_conv ct else if Term_Ord.term_ord (arg2, arg1) = LESS then rewr_obj_eq (obj_sym UtilBase.disj_commute_th) ct else Conv.all_conv ct end (* Sort A_1 | ... | A_n. *) fun sort_disj_clause ct = if not (is_disj (Thm.term_of ct)) then Conv.all_conv ct else Conv.every_conv [Conv.arg_conv sort_disj_clause, sort_disj_clause_aux] ct (* Apply cv on the body of !x y z. ... .*) fun forall_body_conv cv ctxt ct = if is_obj_all (Thm.term_of ct) then Conv.binder_conv (fn (_, ctxt) => forall_body_conv cv ctxt) ctxt ct else cv ct fun norm_all_disj_sorted ctxt ct = Conv.every_conv [norm_all_disj ctxt, forall_body_conv sort_disj_clause ctxt] ct fun abstract_eq ctxt var eq = let val (x, T) = Term.dest_Free var val all_const = Const (UtilBase.All_name, (T --> boolT) --> boolT) in Drule.arg_cong_rule (Thm.cterm_of ctxt all_const) (Thm.abstract_rule x (Thm.cterm_of ctxt var) eq) end fun replace_disj_vars ctxt (vars, disjs) = let val vars' = vars |> map Term.dest_Free |> Variable.variant_frees ctxt [] |> map Free val subst = vars ~~ vars' in (vars', map (subst_atomic subst) disjs) end (* Matching for all-disj propositions *) fun disj_prop_match ctxt (id, (tyinst, inst)) (t, (var_t, ts), (var_u, cus)) = let val thy = Proof_Context.theory_of ctxt val us = map Thm.term_of cus in if length var_t <> length var_u orelse length ts <> length us then [] else let (* First match the types (return [] if no match). *) val tys_t = map fastype_of var_t val tys_u = map snd var_u val tyinst' = fold (Sign.typ_match thy) (tys_t ~~ tys_u) tyinst val var_t' = map (Envir.subst_term_types tyinst') var_t val ts' = ts |> map (Envir.subst_term_types tyinst') val var_ct' = map (Thm.cterm_of ctxt) var_t' - val cus' = cus |> map (Thm.instantiate_cterm ([], var_u ~~ var_ct')) + val cus' = cus |> map (Thm.instantiate_cterm (TVars.empty, Vars.make (var_u ~~ var_ct'))) (* Match the type-instantiated pattern with term. *) val insts = Matcher.rewrite_match_subset ctxt var_t' (ts', cus') (id, (tyinst', inst)) fun process_inst ((id', instsp'), ths) = let (* Equality between normalized t and u *) val eq_th = ths |> mk_disj_eq |> fold (abstract_eq ctxt) (rev var_t') |> apply_to_lhs (norm_all_disj_sorted ctxt) |> apply_to_rhs (norm_all_disj_sorted ctxt) (* Equality between un-normalized t and u *) val t' = Util.subst_term_norm instsp' t val norm1 = norm_all_disj_sorted ctxt (Thm.cterm_of ctxt t') val cu = Thm.cterm_of ctxt (mk_all_disj (var_t', map Thm.term_of cus')) val norm2 = norm_all_disj_sorted ctxt cu val eq_th' = Util.transitive_list [norm1, eq_th, meta_sym norm2] in ((id', instsp'), eq_th') end in map process_inst insts end handle Type.TYPE_MATCH => [] end fun norm_conj_de_Morgan ct = let val t = Thm.term_of ct in if is_neg t andalso is_disj (dest_not t) then Conv.every_conv [rewr_obj_eq UtilBase.de_Morgan_disj_th, Conv.arg_conv norm_conj_de_Morgan] ct else Conv.all_conv ct end fun norm_conj_not_imp ct = let val t = Thm.term_of ct in if is_neg t andalso is_imp (dest_not t) then Conv.every_conv [rewr_obj_eq UtilBase.not_imp_th, Conv.arg_conv norm_conj_not_imp] ct else Conv.all_conv ct end fun norm_nn_cancel ct = let val t = Thm.term_of ct in if is_neg t andalso is_neg (dest_not t) then Conv.every_conv [rewr_obj_eq UtilBase.nn_cancel_th, norm_nn_cancel] ct else Conv.all_conv ct end fun norm_conj ct = let val t = Thm.term_of ct val _ = assert (fastype_of t = boolT) "norm_conj: wrong type" in if is_neg t andalso is_neg (dest_not t) then norm_nn_cancel ct else if is_neg t andalso is_imp (dest_not t) then norm_conj_not_imp ct else if is_neg t andalso is_disj (dest_not t) then norm_conj_de_Morgan ct else Conv.all_conv ct end (* DISJ items. *) val TY_DISJ = "DISJ" (* Given a theorem in the form of a disjunction, possibly containing schematic variables, return the corresponding DISJ item. *) fun disj_to_ritems prem_only disj_head th = let val subs = strip_disj (prop_of' th) in if length subs = 1 then if Util.has_vars (Thm.prop_of th) then let fun th_to_ritem th = Fact (TY_DISJ, UtilLogic.term_of_bool prem_only :: disj :: strip_disj (prop_of' th), th) in map th_to_ritem (th |> apply_to_thm' norm_conj |> UtilLogic.split_conj_th) end else [Fact (TY_PROP, [prop_of' th], th)] else let val tname = UtilLogic.term_of_bool prem_only :: disj_head :: subs in [Fact (TY_DISJ, tname, th)] end end fun disj_to_update prem_only disj_head (id, sc, th) = if Thm.prop_of th aconv pFalse then ResolveBox {id = id, th = th} else AddItems {id = id, sc = sc, raw_items = disj_to_ritems prem_only disj_head th} fun get_disj_head t = if is_disj t then disj else if is_imp t then get_disj_head (dest_arg t) else if is_neg t andalso is_neg (dest_not t) then get_disj_head (dest_not (dest_not t)) else if is_neg t andalso is_conj (dest_not t) then conj else if is_obj_all t orelse is_ball t then case dest_arg t of Abs (x, T, body) => get_disj_head (snd (Term.dest_abs (x, T, body))) | _ => imp else if is_neg t andalso is_ex (dest_not t) then conj else if is_neg t andalso is_bex (dest_not t) then conj else imp (* Given a theorem th, return equivalent theorem in disjunctive form, with possible schematic variables. Also return whether th is "active", that is, whether it is originally a conjunctive goal or disjunctive fact, as opposed to implications. *) fun analyze_disj_th ctxt th = let val head = get_disj_head (prop_of' th) val th' = th |> apply_to_thm' (norm_all_disj ctxt) |> apply_to_thm (UtilLogic.to_meta_conv ctxt) |> Util.forall_elim_sch in (head, th') end (* Deconstruct the tname of a DISJ item. *) fun dest_tname_of_disj tname = case tname of _ :: disj_head :: rest => (Thm.term_of disj_head, rest) | _ => raise Fail "dest_tname_of_disj: too few terms in tname." (* Determine whether the item is for matching premises only (from the first entry in tname. *) fun is_match_prem_only {tname, ...} = UtilLogic.bool_of_term (Thm.term_of (hd tname)) fun is_active_head disj_head = (not (disj_head aconv imp)) fun disj_rewr_terms ts = if UtilLogic.bool_of_term (hd ts) then [] else drop 2 ts fun output_disj_fn ctxt (ts, _) = let val (match_prem, disj_head, subs) = (hd ts, hd (tl ts), tl (tl ts)) val prefix = if UtilLogic.bool_of_term match_prem then "(match_prem) " else "" in if disj_head aconv disj then prefix ^ ((foldr1 mk_disj subs) |> Syntax.string_of_term ctxt) else if disj_head aconv conj then prefix ^ ((foldr1 mk_conj (map get_neg subs)) |> get_neg |> Syntax.string_of_term ctxt) else if disj_head aconv imp then prefix ^ ((foldr1 mk_imp (subs |> split_last |> apfst (map get_neg) |> apsnd single |> (op @))) |> Syntax.string_of_term ctxt) else raise Fail "output_disj_fn: unexpected disj_head." end val disj_prop_matcher = let fun pre_match pat {tname, ...} ctxt = let val (var_t, ts) = (strip_all_disj pat) |> replace_disj_vars ctxt val (_, cus) = dest_tname_of_disj tname val us = map Thm.term_of cus val var_u = fold Term.add_vars us [] in length ts = length us andalso length var_t = length var_u end fun match pat item ctxt (id, instsp) = let val {tname, prop = th, ...} = item val (_, cus) = dest_tname_of_disj tname val us = map Thm.term_of cus val var_u = fold Term.add_vars us [] val (var_t, ts) = (strip_all_disj pat) |> replace_disj_vars ctxt fun process_perm perm = map (pair (map Var perm)) (disj_prop_match ctxt (id, instsp) (pat, (var_t, ts), (perm, cus))) fun process_inst (var_u, ((id', instsp'), eq_th)) = let val eq_th' = make_trueprop_eq (meta_sym eq_th) val forall_th = th |> fold Thm.forall_intr (rev (map (Thm.cterm_of ctxt) var_u)) |> apply_to_thm (UtilLogic.to_obj_conv ctxt) in ((id', instsp'), Thm.equal_elim eq_th' forall_th) end in if length var_t <> length var_u orelse length ts <> length us then [] else var_u |> Util.all_permutes |> maps process_perm |> map process_inst end in {pre_match = pre_match, match = match} end (* Given ct in the form p_1 | ... | p_n, apply cv to each of p_i. *) fun ac_disj_conv cv ct = if is_disj (Thm.term_of ct) then Conv.every_conv [Conv.arg1_conv cv, Conv.arg_conv (ac_disj_conv cv)] ct else cv ct (* Assume ct is a disjunction, associating to the right. *) fun reduce_disj_True ct = if is_disj (Thm.term_of ct) then ((rewr_obj_eq UtilBase.disj_True1_th) else_conv ((Conv.arg_conv reduce_disj_True) then_conv (rewr_obj_eq UtilBase.disj_True2_th))) ct else Conv.all_conv ct (* Handles also the case where pat is in not-conj or imp form. *) fun match_prop ctxt (id, item2) pat = let val disj_pats = strip_disj pat (* th is pat'(inst), where pat' is one of the disjunctive terms of pat. *) fun process_inst ((id, inst), th) = let (* Construct the theorem pat'(inst) == True. *) val to_eqT_cv = (th RS UtilBase.eq_True_th) |> rewr_obj_eq |> Conv.try_conv (* Rewrite pat(inst) using the above, then rewrite to True. *) val pat_eqT = pat |> Thm.cterm_of ctxt |> norm_disj |> Util.subst_thm ctxt inst |> apply_to_rhs (ac_disj_conv to_eqT_cv) |> apply_to_rhs reduce_disj_True |> to_obj_eq val patT = pat_eqT RS UtilBase.eq_True_inv_th in ((id, inst), patT) end val insts1 = (ItemIO.match_arg ctxt (PropMatch pat) item2 (id, fo_init)) |> filter (BoxID.has_incr_id o fst o fst) val insts2 = if length disj_pats > 1 then map process_inst (maps (match_prop ctxt (id, item2)) disj_pats) else [] in insts1 @ insts2 end (* Given theorem ~P, cancel any disjunct that is aconv to P. It is possible to leave one disjunct P un-cancelled. *) fun disj_cancel_cv ctxt th ct = if is_disj (Thm.term_of ct) then Conv.every_conv [Conv.arg_conv (disj_cancel_cv ctxt th), Conv.try_conv (rewr_obj_eq (th RS UtilBase.or_cancel1_th)), Conv.try_conv (rewr_obj_eq (th RS UtilBase.or_cancel2_th))] ct else Conv.all_conv ct (* Given theorem P and a disjuncion theorem, return new disjunction theorem with ~P cancelled. If all disjuncts can be cancelled, return False. *) fun disj_cancel_prop ctxt th prop = let val th' = if is_neg (prop_of' th) then th else th RS UtilBase.nn_create_th val prop' = prop |> apply_to_thm' (disj_cancel_cv ctxt th') val P = th' |> prop_of' |> dest_not in if prop_of' prop' aconv P then [th', prop'] MRS UtilBase.contra_triv_th else prop' end (* Reduce a disjunction p_1 | ... | t | ... | p_n by matching ~t with the second item. If the disjunction contains schematic variables, t must have either zero or the largest number of schematic variables. *) fun match_update_fn ctxt item1 item2 = if is_match_prem_only item1 andalso Util.has_vars (Thm.prop_of (#prop item1)) then [] else let val {id, prop, tname, ...} = item1 val thy = Proof_Context.theory_of ctxt val (disj_head, csubs) = dest_tname_of_disj tname val subs = map Thm.term_of csubs fun count_var t = length (Term.add_vars t []) val max_nvar = fold (curry Int.max) (map count_var subs) 0 fun is_priority_term t = if is_neg t then exists (is_ex orf is_bex) (strip_conj (dest_not t)) else exists (is_obj_all orf is_ball) (strip_disj t) val has_priority_term = exists is_priority_term (map get_neg subs) val (NO_MATCH, SLOW_MATCH, YES_MATCH) = (0, 1, 2) (* Test whether to perform matching on pattern. *) fun test_do_match t = let val nvar = count_var t val neg_t = get_neg t in if not (Util.is_pattern t) then NO_MATCH else if length subs > 1 andalso Property.is_property_prem thy neg_t then NO_MATCH else if has_priority_term andalso not (is_priority_term neg_t) then SLOW_MATCH else if is_mem neg_t andalso Term.is_Var (dest_arg1 neg_t) andalso null (Term.add_frees (dest_arg neg_t) []) then SLOW_MATCH else if nvar = 0 orelse nvar = max_nvar then YES_MATCH else NO_MATCH end (* Match the negation of subs[i] with th2. For each match, instantiate in prop all schematic variables in t, so that t becomes ~th2. Then remove t from prop in the instantiated version. *) fun get_matches i = let val t = nth subs i val do_match = test_do_match t fun process_inst ((id', inst), th) = let val prop' = prop |> Util.subst_thm_thy thy inst |> disj_cancel_prop ctxt th val sc = if do_match = SLOW_MATCH then 200 else 10 in disj_to_update false disj_head (id', SOME sc, prop') :: (if count_var t > 0 then [] else [ShadowItem {id = id', item = item1}]) end in if do_match = NO_MATCH then [] else t |> get_neg |> match_prop ctxt (id, item2) |> maps process_inst end fun get_matches_no_var () = let fun process_inst (id', ths) = let val prop' = prop |> fold (disj_cancel_prop ctxt) ths in disj_to_update false disj_head (id', SOME 1, prop') :: (if is_match_prem_only item1 then [] else [ShadowItem {id = id', item = item1}]) end fun get_match_at_id id' insts = insts |> filter (fn ((id, _), _) => BoxID.is_eq_ancestor ctxt id id') |> map snd |> take 1 fun get_matches_at_id all_insts id' = (id', maps (get_match_at_id id') all_insts) fun merge_matches all_insts = let val ids = distinct (op =) (maps (map (fst o fst)) all_insts) in map (get_matches_at_id all_insts) ids end val _ = assert (length subs >= 2) val ts = [hd subs, List.last subs] in ts |> map get_neg |> map (match_prop ctxt (id, item2)) |> merge_matches |> maps process_inst end in if max_nvar = 0 then get_matches_no_var () else maps get_matches (0 upto (length subs - 1)) end val match_update_prfstep = {name = "disj_match_update", args = [TypedUniv TY_DISJ, PropMatch (boolVar "A")], func = TwoStep match_update_fn} (* For DISJ items with a single term, of form f p1 ... pn, match t against each of p_i. *) fun match_one_sch_fn ctxt item1 item2 = if is_match_prem_only item1 then [] else let val {id, tname, prop = th1, ...} = item1 val thy = Proof_Context.theory_of ctxt val subs = (dest_tname_of_disj tname) |> snd |> map Thm.term_of in if length subs > 1 then [] else let val t = the_single subs val args = Util.dest_args t fun count_var t = length (Term.add_vars t []) val nvar = count_var t fun get_matches i = if count_var (nth args i) < nvar then [] else let val arg = nth args i val targ = TypedMatch (TY_TERM, arg) val insts = (ItemIO.match_arg ctxt targ item2 (id, fo_init)) |> filter (BoxID.has_incr_id o fst o fst) fun inst_to_updt ((id', inst), _) = let val th1' = Util.subst_thm_thy thy inst th1 val prop' = prop_of' th1' in if is_eq_term prop' andalso RewriteTable.is_equiv_t id' ctxt (dest_eq prop') then [] else [Update.thm_update (id', th1')] end in maps inst_to_updt insts end in maps get_matches (0 upto (length args - 1)) end end val match_one_sch_prfstep = {name = "disj_match_one_sch", args = [TypedUniv TY_DISJ, TypedUniv TY_TERM], func = TwoStep match_one_sch_fn} fun disj_match_iff_fn ctxt {id, tname, prop, ...} = if not (BoxID.has_incr_id id) then [] else let val (_, csubs) = dest_tname_of_disj tname val subs = map Thm.term_of csubs in if length subs > 1 then [] else if not (is_eq_term (the_single subs) andalso fastype_of (dest_arg (the_single subs)) = boolT) then [] else let val cv = (UtilLogic.to_obj_conv ctxt) then_conv (Trueprop_conv norm_disj) val forward = prop |> UtilLogic.equiv_forward_th |> apply_to_thm cv val backward = prop |> UtilLogic.equiv_backward_th |> apply_to_thm cv in [disj_to_update false imp (id, NONE, forward), disj_to_update false imp (id, NONE, backward)] end end val disj_match_iff_prfstep = {name = "disj_match_iff", args = [TypedUniv TY_DISJ], func = OneStep disj_match_iff_fn} (* For active case, create box checking the next case. *) fun disj_create_case_fn _ {id, tname, ...} = if not (BoxID.has_incr_id id) then [] else if exists Util.has_vars (map Thm.term_of tname) then [] else let val (disj_head, csubs) = dest_tname_of_disj tname in if not (is_active_head disj_head) then [] else if length csubs = 1 then [] else let val subs = map Thm.term_of csubs in [AddBoxes {id = id, sc = NONE, init_assum = mk_Trueprop (hd subs)}] end end val disj_create_case_prfstep = {name = "disj_create_case", args = [TypedUniv TY_DISJ], func = OneStep disj_create_case_fn} (* item1 dominates item2 if the disjunctive terms in item1 is a subset of that for item2. *) fun disj_shadow_fn ctxt (item1 as {tname = tname1, ...}) (item2 as {tname = tname2, ...}) = let val id = BoxItem.merged_id ctxt [item1, item2] val (disj_head1, subs1) = dest_tname_of_disj tname1 val (disj_head2, subs2) = dest_tname_of_disj tname2 in if not (BoxID.has_incr_id id) then [] else if not (is_active_head disj_head1) andalso is_active_head disj_head2 then [] else if is_match_prem_only item1 andalso not (is_match_prem_only item2) then [] else if subset (op aconvc) (subs1, subs2) then [ShadowItem {id = id, item = item2}] else [] end val disj_shadow_prfstep = {name = "disj_shadow", args = [TypedUniv TY_DISJ, TypedUniv TY_DISJ], func = TwoStep disj_shadow_fn} val add_disj_proofsteps = fold ItemIO.add_item_type [ (TY_DISJ, SOME disj_rewr_terms, SOME output_disj_fn, NONE) ] #> fold ItemIO.add_prop_matcher [ (TY_DISJ, disj_prop_matcher) ] #> fold add_prfstep [ match_update_prfstep, match_one_sch_prfstep, disj_match_iff_prfstep, disj_create_case_prfstep, disj_shadow_prfstep ] (* Normalizers *) fun split_not_imp_th th = th |> apply_to_thm' norm_conj_not_imp |> UtilLogic.split_conj_th (* Generalized form of splitting A & B. Also deal with cases ~(A | B) and ~(A --> B). *) fun split_conj_gen_th _ th = th |> apply_to_thm' norm_conj |> UtilLogic.split_conj_th fun eq_normalizer _ ritem = case ritem of Handler _ => [ritem] | Fact (ty_str, _, th) => if ty_str <> TY_PROP then [ritem] else if is_eq_term (prop_of' th) then let val (lhs, rhs) = dest_eq (prop_of' th) in if fastype_of lhs = boolT then map Update.thm_to_ritem (UtilLogic.split_conj_th (th RS UtilBase.iffD_th)) else [Fact (TY_EQ, [lhs, rhs], th)] end else [ritem] fun property_normalizer _ ritem = case ritem of Handler _ => [ritem] | Fact (ty_str, _, th) => if ty_str <> TY_PROP then [ritem] else if Property.is_property (prop_of' th) then [Fact (TY_PROPERTY, [prop_of' th], th)] else [ritem] fun disj_normalizer ctxt ritem = case ritem of Handler _ => [ritem] | Fact (ty_str, _, th) => if ty_str <> TY_PROP then [ritem] else let val t = prop_of' th in if is_neg t andalso is_conj (dest_not t) orelse is_disj t orelse is_imp t orelse is_obj_all t orelse is_ball t orelse is_neg t andalso is_ex (dest_not t) orelse is_neg t andalso is_bex (dest_not t) then let val (disj_head, disj_th) = analyze_disj_th ctxt th val prem_only = Auto2_State.lookup_prem_only ctxt t val disj_th = if prem_only then disj_th else snd (Normalizer.meta_use_vardefs disj_th) in disj_to_ritems prem_only disj_head disj_th end else [ritem] end fun logic_thm_update ctxt (id, th) = let val t = prop_of' th in if is_obj_all t orelse is_ball t orelse is_neg t andalso is_ex (dest_not t) orelse is_neg t andalso is_bex (dest_not t) orelse is_disj t orelse is_imp t orelse is_neg t andalso is_conj (dest_not t) then let val (disj_head, disj_th) = analyze_disj_th ctxt th val raw_items = disj_to_ritems true disj_head disj_th in AddItems {id = id, sc = NONE, raw_items = raw_items} end else Update.thm_update (id, th) end val add_disj_normalizers = Normalizer.add_th_normalizer ( "split_conj_gen", split_conj_gen_th ) #> fold Normalizer.add_normalizer [ ("eq", eq_normalizer), ("property", property_normalizer), ("disj", disj_normalizer) ] end (* structure Logic_ProofSteps. *) val _ = Theory.setup Logic_ProofSteps.add_logic_proofsteps val _ = Theory.setup Logic_ProofSteps.add_disj_proofsteps val _ = Theory.setup Logic_ProofSteps.add_disj_normalizers val TY_DISJ = Logic_ProofSteps.TY_DISJ diff --git a/thys/Auto2_HOL/util.ML b/thys/Auto2_HOL/util.ML --- a/thys/Auto2_HOL/util.ML +++ b/thys/Auto2_HOL/util.ML @@ -1,926 +1,926 @@ (* File: util.ML Author: Bohua Zhan Utility functions. *) signature BASIC_UTIL = sig (* Exceptions *) val assert: bool -> string -> unit (* Types *) val propT: typ (* Lists *) val the_pair: 'a list -> 'a * 'a val the_triple: 'a list -> 'a * 'a * 'a val filter_split: ('a -> bool) -> 'a list -> 'a list * 'a list (* Managing matching environments. *) val fo_init: Type.tyenv * Envir.tenv val lookup_instn: Type.tyenv * Envir.tenv -> string * int -> term val lookup_inst: Type.tyenv * Envir.tenv -> string -> term (* Tracing functions. *) val trace_t: Proof.context -> string -> term -> unit val trace_tlist: Proof.context -> string -> term list -> unit val trace_thm: Proof.context -> string -> thm -> unit val trace_fullthm: Proof.context -> string -> thm -> unit val trace_thm_global: string -> thm -> unit val trace_fullthm_global: string -> thm -> unit (* Terms *) val dest_arg: term -> term val dest_arg1: term -> term (* Theorems *) val apply_to_thm: conv -> thm -> thm val meta_sym: thm -> thm val apply_to_lhs: conv -> thm -> thm val apply_to_rhs: conv -> thm -> thm end; signature UTIL = sig include BASIC_UTIL (* Lists *) val max: ('a * 'a -> order) -> 'a list -> 'a val max_partial: ('a -> 'a -> bool) -> 'a list -> 'a list val subsets: 'a list -> 'a list list val all_permutes: 'a list -> 'a list list val all_pairs: 'a list * 'b list -> ('a * 'b) list val remove_dup_lists: ('a * 'a -> order) -> 'a list * 'a list -> 'a list * 'a list val is_subseq: ('a * 'a -> bool) -> 'a list * 'a list -> bool (* Strings. *) val is_prefix_str: string -> string -> bool val is_just_internal: string -> bool (* Managing matching environments. *) val defined_instn: Type.tyenv * Envir.tenv -> string * int -> bool val lookup_tyinst: Type.tyenv * Envir.tenv -> string -> typ val update_env: indexname * term -> Type.tyenv * Envir.tenv -> Type.tyenv * Envir.tenv val eq_env: (Type.tyenv * Envir.tenv) * (Type.tyenv * Envir.tenv) -> bool (* Matching. *) val first_order_match_list: theory -> (term * term) list -> Type.tyenv * Envir.tenv -> Type.tyenv * Envir.tenv (* Printing functions, mostly from Isabelle Cookbook. *) val string_of_terms: Proof.context -> term list -> string val string_of_terms_global: theory -> term list -> string val string_of_tyenv: Proof.context -> Type.tyenv -> string val string_of_env: Proof.context -> Envir.tenv -> string val string_of_list: ('a -> string) -> 'a list -> string val string_of_list': ('a -> string) -> 'a list -> string val string_of_bool: bool -> string (* Managing context. *) val declare_free_term: term -> Proof.context -> Proof.context (* Terms. *) val is_abs: term -> bool val is_implies: term -> bool val dest_binop: term -> term * (term * term) val dest_binop_head: term -> term val dest_binop_args: term -> term * term val dest_args: term -> term list val dest_argn: int -> term -> term val get_head_name: term -> string val is_meta_eq: term -> bool val occurs_frees: term list -> term -> bool val occurs_free: term -> term -> bool val has_vars: term -> bool val is_subterm: term -> term -> bool val has_subterm: term list -> term -> bool val is_head: term -> term -> bool val lambda_abstract: term -> term -> term val is_pattern_list: term list -> bool val is_pattern: term -> bool val normalize_meta_all_imp: Proof.context -> conv val swap_meta_imp_alls: Proof.context -> conv val normalize_meta_horn: Proof.context -> conv val strip_meta_horn: term -> term list * (term list * term) val list_meta_horn: term list * (term list * term) -> term val to_internal_vars: Proof.context -> term list * term -> term list * term val rename_abs_term: term list -> term -> term val print_term_detail: Proof.context -> term -> string (* cterms. *) val dest_cargs: cterm -> cterm list val dest_binop_cargs: cterm -> cterm * cterm (* Theorems. *) val arg_backn_conv: int -> conv -> conv val argn_conv: int -> conv -> conv val comb_equiv: cterm * thm list -> thm val name_of_thm: thm -> string val update_name_of_thm: thm -> string -> thm -> thm val lhs_of: thm -> term val rhs_of: thm -> term val assume_meta_eq: theory -> term * term -> thm val assume_thm: Proof.context -> term -> thm val subst_thm_thy: theory -> Type.tyenv * Envir.tenv -> thm -> thm val subst_thm: Proof.context -> Type.tyenv * Envir.tenv -> thm -> thm val send_first_to_hyps: thm -> thm val send_all_to_hyps: thm -> thm val subst_thm_atomic: (cterm * cterm) list -> thm -> thm val subst_term_norm: Type.tyenv * Envir.tenv -> term -> term val concl_conv_n: int -> conv -> conv val concl_conv: conv -> conv val transitive_list: thm list -> thm val skip_n_conv: int -> conv -> conv val pattern_rewr_conv: term -> (term * thm) list -> conv val eq_cong_th: int -> term -> term -> Proof.context -> thm val forall_elim_sch: thm -> thm (* Conversions *) val reverse_eta_conv: Proof.context -> conv val repeat_n_conv: int -> conv -> conv (* Miscellaneous. *) val test_conv: Proof.context -> conv -> string -> string * string -> unit val term_pat_setup: theory -> theory val cterm_pat_setup: theory -> theory val type_pat_setup: theory -> theory val timer: string * (unit -> 'a) -> 'a val exn_trace: (unit -> 'a) -> 'a end; structure Util : UTIL = struct fun assert b exn_str = if b then () else raise Fail exn_str val propT = @{typ prop} fun max comp lst = let fun max2 t1 t2 = if comp (t1, t2) = LESS then t2 else t1 in case lst of [] => raise Fail "max: empty list" | l :: ls => fold max2 ls l end (* Given a function comp, remove y for each pair (x, y) such that comp x y = true (if x dominates y). *) fun max_partial comp lst = let fun helper taken remains = case remains of [] => taken | x :: xs => if exists (fn y => comp y x) taken then helper taken xs else helper (x :: filter_out (fn y => comp x y) taken) xs in helper [] lst end (* Return all subsets of lst. *) fun subsets [] = [[]] | subsets (l::ls) = let val prev = subsets ls in prev @ map (cons l) prev end (* List of all permutations of xs *) fun all_permutes xs = case xs of [] => [[]] | [x] => [[x]] | [x, y] => [[x, y], [y, x]] | _ => maps (fn i => map (cons (nth xs i)) (all_permutes (nth_drop i xs))) (0 upto (length xs - 1)) (* Convert list to pair. List must consist of exactly two items. *) fun the_pair lst = case lst of [i1, i2] => (i1, i2) | _ => raise Fail "the_pair" (* Convert list to triple. List must consist of exactly three items. *) fun the_triple lst = case lst of [i1, i2, i3] => (i1, i2, i3) | _ => raise Fail "the_triple" (* Split list into (ins, outs), where ins satisfy f, and outs don't. *) fun filter_split _ [] = ([], []) | filter_split f (x :: xs) = let val (ins, outs) = filter_split f xs in if f x then (x :: ins, outs) else (ins, x :: outs) end (* Form the Cartesian product of two lists. *) fun all_pairs (l1, l2) = maps (fn y => (map (fn x => (x, y)) l1)) l2 (* Given two sorted lists, remove all pairs of terms that appear in both lists, counting multiplicity. *) fun remove_dup_lists ord (xs, ys) = case xs of [] => ([], ys) | x :: xs' => case ys of [] => (xs, []) | y :: ys' => case ord (x, y) of LESS => apfst (cons x) (remove_dup_lists ord (xs', ys)) | EQUAL => remove_dup_lists ord (xs', ys') | GREATER => apsnd (cons y) (remove_dup_lists ord (xs, ys')) (* Whether l1 is a subsequence of l2. *) fun is_subseq eq (l1, l2) = case l1 of [] => true | x :: l1' => case l2 of [] => false | y :: l2' => if eq (x, y) then is_subseq eq (l1', l2') else is_subseq eq (l1, l2') (* Whether pre is a prefix of str. *) fun is_prefix_str pre str = is_prefix (op =) (String.explode pre) (String.explode str) (* Test whether x is followed by exactly one _. *) fun is_just_internal x = Name.is_internal x andalso not (Name.is_skolem x) val fo_init = (Vartab.empty, Vartab.empty) (* Lookup a Vartab inst with string and integer specifying indexname. *) fun defined_instn (_, inst) (str, n) = Vartab.defined inst (str, n) fun lookup_instn (_, inst) (str, n) = case Vartab.lookup inst (str, n) of NONE => raise Fail ("lookup_inst: not found " ^ str ^ (if n = 0 then "" else string_of_int n)) | SOME (_, u) => u fun lookup_inst (tyinst, inst) str = lookup_instn (tyinst, inst) (str, 0) fun lookup_tyinst (tyinst, _) str = case Vartab.lookup tyinst (str, 0) of NONE => raise Fail ("lookup_tyinst: not found " ^ str) | SOME (_, T) => T fun update_env (idx, t) (tyenv, tenv) = (tyenv, tenv |> Vartab.update_new (idx, (type_of t, t))) (* A rough comparison, simply compare the corresponding terms. *) fun eq_env ((_, inst1), (_, inst2)) = let val data1 = Vartab.dest inst1 val data2 = Vartab.dest inst2 fun compare_data (((x, i), (ty, t)), ((x', i'), (ty', t'))) = x = x' andalso i = i' andalso ty = ty' andalso t aconv t' in eq_set compare_data (data1, data2) end fun first_order_match_list thy pairs inst = case pairs of [] => inst | (t, u) :: rest => let val inst' = Pattern.first_order_match thy (t, u) inst in first_order_match_list thy rest inst' end fun string_of_terms ctxt ts = ts |> map (Syntax.pretty_term ctxt) |> Pretty.commas |> Pretty.block |> Pretty.string_of fun string_of_terms_global thy ts = ts |> map (Syntax.pretty_term_global thy) |> Pretty.commas |> Pretty.block |> Pretty.string_of fun pretty_helper aux env = env |> Vartab.dest |> map aux |> map (fn (s1, s2) => Pretty.block [s1, Pretty.str " := ", s2]) |> Pretty.enum "," "[" "]" |> Pretty.string_of fun string_of_tyenv ctxt tyenv = let fun get_typs (v, (s, T)) = (TVar (v, s), T) val print = apply2 (Syntax.pretty_typ ctxt) in pretty_helper (print o get_typs) tyenv end fun string_of_env ctxt env = let fun get_ts (v, (T, t)) = (Var (v, T), t) val print = apply2 (Syntax.pretty_term ctxt) in pretty_helper (print o get_ts) env end fun string_of_list func lst = Pretty.str_list "[" "]" (map func lst) |> Pretty.string_of fun string_of_list' func lst = if length lst = 1 then func (the_single lst) else string_of_list func lst fun string_of_bool b = if b then "true" else "false" fun trace_t ctxt s t = tracing (s ^ " " ^ (Syntax.string_of_term ctxt t)) fun trace_tlist ctxt s ts = tracing (s ^ " " ^ (string_of_terms ctxt ts)) fun trace_thm ctxt s th = tracing (s ^ " " ^ (th |> Thm.prop_of |> Syntax.string_of_term ctxt)) fun trace_fullthm ctxt s th = tracing (s ^ " [" ^ (Thm.hyps_of th |> string_of_terms ctxt) ^ "] ==> " ^ (Thm.prop_of th |> Syntax.string_of_term ctxt)) fun trace_thm_global s th = let val thy = Thm.theory_of_thm th in tracing (s ^ " " ^ (th |> Thm.prop_of |> Syntax.string_of_term_global thy)) end fun trace_fullthm_global s th = let val thy = Thm.theory_of_thm th in tracing (s ^ " [" ^ (Thm.hyps_of th |> string_of_terms_global thy) ^ "] ==> " ^ (Thm.prop_of th |> Syntax.string_of_term_global thy)) end fun declare_free_term t ctxt = if not (is_Free t) then raise Fail "declare_free_term: t not free." else ctxt |> Variable.add_fixes_direct [t |> Term.dest_Free |> fst] |> Variable.declare_term t fun is_abs t = case t of Abs _ => true | _ => false (* Whether a given term is of the form A ==> B. *) fun is_implies t = let val _ = assert (fastype_of t = propT) "is_implies: wrong type" in case t of @{const Pure.imp} $ _ $ _ => true | _ => false end (* Extract the last two arguments on t, collecting the rest into f. *) fun dest_binop t = case t of f $ a $ b => (f, (a, b)) | _ => raise Fail "dest_binop" fun dest_binop_head t = fst (dest_binop t) fun dest_binop_args t = snd (dest_binop t) (* Return the argument of t. If t is f applied to multiple arguments, return the last argument. *) fun dest_arg t = case t of _ $ arg => arg | _ => raise Fail "dest_arg" (* Return the first of two arguments of t. *) fun dest_arg1 t = case t of _ $ arg1 $ _ => arg1 | _ => raise Fail "dest_arg1" (* Return the list of all arguments of t. *) fun dest_args t = t |> Term.strip_comb |> snd (* Return the nth argument of t, counting from left and starting at zero. *) fun dest_argn n t = nth (dest_args t) n (* Return the name of the head function. *) fun get_head_name t = case Term.head_of t of Const (nm, _) => nm | _ => raise Fail "get_head_name" (* Whether the term is of the form A == B. *) fun is_meta_eq t = let val _ = assert (fastype_of t = propT) "is_meta_eq_term: wrong type" in case t of Const (@{const_name Pure.eq}, _) $ _ $ _ => true | _ => false end (* Given free variable freevar (or a list of free variables freevars), determine whether any of the inputs appears in t. *) fun occurs_frees freevars t = inter (op aconv) (map Free (Term.add_frees t [])) freevars <> [] fun occurs_free freevar t = occurs_frees [freevar] t (* Whether the given term contains schematic variables. *) fun has_vars t = length (Term.add_vars t []) > 0 (* Whether subt is a subterm of t. *) fun is_subterm subt t = exists_subterm (fn t' => t' aconv subt) t (* Whether any of subts is a subterm of t. *) fun has_subterm subts t = exists_subterm (fn t' => member (op aconv) subts t') t (* Whether s is a head term of t. *) fun is_head s t = let val (sf, sargs) = Term.strip_comb s val (tf, targs) = Term.strip_comb t in sf aconv tf andalso is_prefix (op aconv) sargs targs end (* If stmt is P(t), return %t. P(t). *) fun lambda_abstract t stmt = Term.lambda t (Term.abstract_over (t, stmt)) (* A more general criterion for patterns. In combinations, any argument that is a pattern (in the more general sense) frees up checking for any functional schematic variables in that argument. *) fun is_pattern_list ts = let fun is_funT T = case T of Type ("fun", _) => true | _ => false fun get_fun_vars t = (Term.add_vars t []) |> filter (is_funT o snd) |> map Var fun test_list exclude_vars ts = case ts of [] => true | [t] => test_term exclude_vars t | t :: ts' => if test_term exclude_vars t then test_list (merge (op aconv) (exclude_vars, get_fun_vars t)) ts' else if test_list exclude_vars ts' then test_term (distinct (op aconv) (exclude_vars @ maps get_fun_vars ts')) t else false and test_term exclude_vars t = case t of Abs (_, _, t') => test_term exclude_vars t' | _ => let val (head, args) = strip_comb t in if is_Var head then if member (op aconv) exclude_vars head then test_list exclude_vars args else forall is_Bound args andalso not (has_duplicates (op aconv) args) else test_list exclude_vars args end in test_list [] ts end fun is_pattern t = is_pattern_list [t] (* Push !!x to the right as much as possible. *) fun normalize_meta_all_imp_once ct = Conv.try_conv ( Conv.every_conv [Conv.rewr_conv (Thm.symmetric @{thm Pure.norm_hhf_eq}), Conv.arg_conv normalize_meta_all_imp_once]) ct fun normalize_meta_all_imp ctxt ct = let val t = Thm.term_of ct in if Logic.is_all t then Conv.every_conv [Conv.binder_conv (normalize_meta_all_imp o snd) ctxt, normalize_meta_all_imp_once] ct else if is_implies t then Conv.arg_conv (normalize_meta_all_imp ctxt) ct else Conv.all_conv ct end (* Rewrite A ==> !!v_i. B to !!v_i. A ==> B. *) fun swap_meta_imp_alls ctxt ct = let val t = Thm.term_of ct in if is_implies t andalso Logic.is_all (dest_arg t) then Conv.every_conv [Conv.rewr_conv @{thm Pure.norm_hhf_eq}, Conv.binder_conv (swap_meta_imp_alls o snd) ctxt] ct else Conv.all_conv ct end (* Normalize a horn clause into standard form !!v_i. A_i ==> B. *) fun normalize_meta_horn ctxt ct = let val t = Thm.term_of ct in if Logic.is_all t then Conv.binder_conv (normalize_meta_horn o snd) ctxt ct else if is_implies t then Conv.every_conv [Conv.arg_conv (normalize_meta_horn ctxt), swap_meta_imp_alls ctxt] ct else Conv.all_conv ct end (* Deconstruct a horn clause !!v_i. A_i ==> B into (v_i, (A_i, B)). *) fun strip_meta_horn t = case t of Const (@{const_name Pure.all}, _) $ Abs (abs as (_, T, _)) => let val (x, body) = Term.dest_abs abs val var = Free (x, T) val (vars, (assums, concl)) = strip_meta_horn body in (var :: vars, (assums, concl)) end | @{const Pure.imp} $ P $ Q => let val (vars, (assums, concl)) = strip_meta_horn Q in (vars, (P :: assums, concl)) end | _ => ([], ([], t)) fun list_meta_horn (vars, (As, B)) = (Logic.list_implies (As, B)) |> fold Logic.all (rev vars) fun to_internal_vars ctxt (vars, body) = let val vars' = vars |> map Term.dest_Free |> Variable.variant_frees ctxt [] |> map (apfst Name.internal) |> map Free val subst = vars ~~ vars' in (vars', subst_atomic subst body) end fun rename_abs_term vars t = case vars of [] => t | var :: rest => let val (x, _) = Term.dest_Free var in case t of A $ Abs (_, T1, body) => A $ Abs (x, T1, rename_abs_term rest body) | _ => error "rename_abs_term" end fun print_term_detail ctxt t = case t of Const (s, ty) => "Const (" ^ s ^ ", " ^ (Syntax.string_of_typ ctxt ty) | Free (s, ty) => "Free (" ^ s ^ ", " ^ (Syntax.string_of_typ ctxt ty) ^ ")" | Var ((x, i), ty) => "Var ((" ^ x ^ ", " ^ (string_of_int i) ^ "), " ^ (Syntax.string_of_typ ctxt ty) ^ ")" | Bound n => "Bound " ^ (string_of_int n) | Abs (s, ty, b) => "Abs (" ^ s ^ ", " ^ (Syntax.string_of_typ ctxt ty) ^ ", " ^ (print_term_detail ctxt b) ^ ")" | u $ v => "(" ^ print_term_detail ctxt u ^ ") $ (" ^ print_term_detail ctxt v ^ ")" (* Version for ct. *) fun dest_cargs ct = ct |> Drule.strip_comb |> snd fun dest_binop_cargs ct = (Thm.dest_arg1 ct, Thm.dest_arg ct) (* Apply cv to nth argument of t, counting from right and starting at 0. *) fun arg_backn_conv n cv ct = if n = 0 then Conv.arg_conv cv ct else Conv.fun_conv (arg_backn_conv (n-1) cv) ct (* Apply cv to nth argument of t, counting from left and starting at 0. *) fun argn_conv n cv ct = let val args_count = ct |> Thm.term_of |> dest_args |> length val _ = assert (n >= 0 andalso n < args_count) in arg_backn_conv (args_count - n - 1) cv ct end (* Given a head cterm f (function to be applied), and a list of equivalence theorems of arguments, produce an equivalent theorem for the overall term. *) fun comb_equiv (cf, arg_equivs) = Library.foldl (uncurry Thm.combination) (Thm.reflexive cf, arg_equivs) (* Retrive name of theorem. *) fun name_of_thm th = if Thm.has_name_hint th then Thm.get_name_hint th else raise Fail "name_of_thm: not found" (* Set the name of th to the name of ori_th, followed by suffix. *) fun update_name_of_thm ori_th suffix th = if Thm.has_name_hint ori_th then th |> Thm.put_name_hint (Thm.get_name_hint ori_th ^ suffix) else th val lhs_of = Thm.term_of o Thm.lhs_of val rhs_of = Thm.term_of o Thm.rhs_of fun assume_meta_eq thy (t1, t2) = Thm.assume (Thm.global_cterm_of thy (Logic.mk_equals (t1, t2))) fun assume_thm ctxt t = if type_of t <> propT then raise Fail "assume_thm: t is not of type prop" else Thm.assume (Thm.cterm_of ctxt t) (* Similar to Envir.subst_term. Apply an instantiation to a theorem. *) fun subst_thm_thy thy (tyinsts, insts) th = let fun process_tyenv (v, (S, T)) = ((v, S), Thm.global_ctyp_of thy T) val tys = map process_tyenv (Vartab.dest tyinsts) fun process_tenv (v, (T, u)) = ((v, Envir.subst_type tyinsts T), Thm.global_cterm_of thy u) val ts = map process_tenv (Vartab.dest insts) in - th |> Drule.instantiate_normalize (tys, ts) + th |> Drule.instantiate_normalize (TVars.make tys, Vars.make ts) end fun subst_thm ctxt (tyinsts, insts) th = subst_thm_thy (Proof_Context.theory_of ctxt) (tyinsts, insts) th fun send_first_to_hyps th = let val cprem = Thm.cprem_of th 1 in Thm.implies_elim th (Thm.assume cprem) end fun send_all_to_hyps th = let val _ = assert (forall (not o has_vars) (Thm.prems_of th)) "send_all_to_hyps: schematic variables in hyps." in funpow (Thm.nprems_of th) send_first_to_hyps th end (* Replace using subst the internal variables in th. This proceeds in several steps: first, pull any hypotheses of the theorem involving the replaced variables into statement of the theorem, perform the replacement (using forall_intr then forall_elim), finally return the hypotheses to their original place. *) fun subst_thm_atomic subst th = let val old_cts = map fst subst val old_ts = map Thm.term_of old_cts val new_cts = map snd subst val chyps = filter (fn ct => has_subterm old_ts (Thm.term_of ct)) (Thm.chyps_of th) in th |> fold Thm.implies_intr chyps |> fold Thm.forall_intr old_cts |> fold Thm.forall_elim (rev new_cts) |> funpow (length chyps) send_first_to_hyps end (* Substitution into terms used in auto2. Substitute types first and instantiate the types in the table of term instantiations. Also perform beta_norm at the end. *) fun subst_term_norm (tyinsts, insts) t = let fun inst_tenv tenv = tenv |> Vartab.dest |> map (fn (ixn, (T, v)) => (ixn, (Envir.subst_type tyinsts T, v))) |> Vartab.make in t |> Envir.subst_term_types tyinsts |> Envir.subst_term (tyinsts, inst_tenv insts) |> Envir.beta_norm end (* Apply the conversion cv to the statement of th, yielding the equivalent theorem. *) fun apply_to_thm cv th = let val eq = cv (Thm.cprop_of th) in if Thm.is_reflexive eq then th else Thm.equal_elim eq th end (* Given th of form A == B, get th' of form B == A. *) val meta_sym = Thm.symmetric (* Apply conv to rewrite the left hand side of th. *) fun apply_to_lhs cv th = let val eq = cv (Thm.lhs_of th) in if Thm.is_reflexive eq then th else Thm.transitive (meta_sym eq) th end (* Apply conv to rewrite the right hand side of th. *) fun apply_to_rhs cv th = let val eq = cv (Thm.rhs_of th) in if Thm.is_reflexive eq then th else Thm.transitive th eq end (* Using cv, rewrite the part of ct after stripping i premises. *) fun concl_conv_n i cv ct = if i = 0 then cv ct else (Conv.arg_conv (concl_conv_n (i-1) cv)) ct (* Rewrite part of ct after stripping all premises. *) fun concl_conv cv ct = case Thm.term_of ct of @{const Pure.imp} $ _ $ _ => Conv.arg_conv (concl_conv cv) ct | _ => cv ct (* Given a list of theorems A = B, B = C, etc., apply Thm.transitive to get equality between start and end. *) fun transitive_list ths = let fun rev_transitive btoc atob = let val (b, c) = btoc |> Thm.cprop_of |> Thm.dest_equals val (a, b') = atob |> Thm.cprop_of |> Thm.dest_equals in if b aconvc b' then if a aconvc b then btoc else if b aconvc c then atob else Thm.transitive atob btoc else let val _ = map (trace_thm_global "ths:") ths in raise Fail "transitive_list: intermediate does not agree" end end in fold rev_transitive (tl ths) (hd ths) end (* Skip to argument n times. For example, if applied to rewrite a proposition in implication form (==> or -->), it will skip the first n assumptions. *) fun skip_n_conv n cv = if n <= 0 then cv else Conv.arg_conv (skip_n_conv (n-1) cv) (* Given a term t, and pairs (a_i, eq_i), where a_i's are atomic subterms of t. Suppose the input is obtained by replacing each a_i by the left side of eq_i in t, obtain equality from t to the term obtained by replacing each a_i by the right side of eq_i in t. *) fun pattern_rewr_conv t eq_ths ct = case t of Bound _ => raise Fail "pattern_rewr_conv: bound variable." | Abs _ => raise Fail "pattern_rewr_conv: abs not supported." | _ $ _ => let val (f, arg) = Term.strip_comb t val (f', arg') = Drule.strip_comb ct val _ = assert (f aconv Thm.term_of f') "pattern_rewr_conv: input does not match pattern." val ths = map (fn (t, ct) => pattern_rewr_conv t eq_ths ct) (arg ~~ arg') in comb_equiv (f', ths) end | Const _ => let val _ = assert (t aconv Thm.term_of ct) "pattern_rewr_conv: input does not match pattern." in Conv.all_conv ct end | _ => (* Free and Var cases *) (case AList.lookup (op aconv) eq_ths t of NONE => Conv.all_conv ct | SOME eq_th => let val _ = assert (lhs_of eq_th aconv (Thm.term_of ct)) "pattern_rewr_conv: wrong substitution." in eq_th end) (* Given integer i, term b_i, and a term A = f(a_1, ..., a_n), produce the theorem a_i = b_i ==> A = f(a_1, ..., b_i, ..., a_n). *) fun eq_cong_th i bi A ctxt = let val thy = Proof_Context.theory_of ctxt val (cf, cargs) = Drule.strip_comb (Thm.cterm_of ctxt A) val _ = assert (i < length cargs) "eq_cong_th: i out of bounds." val ai = Thm.term_of (nth cargs i) val eq_i = assume_meta_eq thy (ai, bi) val eqs = map Thm.reflexive (take i cargs) @ [eq_i] @ map Thm.reflexive (drop (i + 1) cargs) val eq_A = comb_equiv (cf, eqs) in Thm.implies_intr (Thm.cprop_of eq_i) eq_A end (* Convert theorems of form !!x y. P x y into P ?x ?y (arbitrary number of quantifiers). *) fun forall_elim_sch th = case Thm.prop_of th of Const (@{const_name Pure.all}, _) $ Abs (x, T, _) => let val var_xs = map fst (Term.add_var_names (Thm.prop_of th) []) val x' = if member (op =) var_xs x then singleton (Name.variant_list var_xs) x else x val thy = Thm.theory_of_thm th in th |> Thm.forall_elim (Thm.global_cterm_of thy (Var ((x', 0), T))) |> forall_elim_sch end | _ => th (* Given P of function type, produce P == %x. P x. *) fun reverse_eta_conv ctxt ct = let val t = Thm.term_of ct val argT = Term.domain_type (fastype_of t) handle Match => raise CTERM ("reverse_eta_conv", [ct]) val rhs = Abs ("x", argT, t $ Bound 0) val eq_th = Thm.eta_conversion (Thm.cterm_of ctxt rhs) in meta_sym eq_th end (* Repeat cv exactly n times. *) fun repeat_n_conv n cv t = if n = 0 then Conv.all_conv t else (cv then_conv (repeat_n_conv (n-1) cv)) t (* Generic function for testing a conv. *) fun test_conv ctxt cv err_str (str1, str2) = let val (t1, t2) = (Proof_Context.read_term_pattern ctxt str1, Proof_Context.read_term_pattern ctxt str2) val th = cv (Thm.cterm_of ctxt t1) in if t1 aconv (lhs_of th) andalso t2 aconv (rhs_of th) then () else let val _ = trace_t ctxt "Input:" t1 val _ = trace_t ctxt "Expected:" t2 val _ = trace_t ctxt "Actual:" (Thm.prop_of th) in raise Fail err_str end end (* term_pat and typ_pat, from Isabelle Cookbook. *) val term_pat_setup = let val parser = Args.context -- Scan.lift Args.embedded_inner_syntax fun term_pat (ctxt, str) = str |> Proof_Context.read_term_pattern ctxt |> ML_Syntax.print_term |> ML_Syntax.atomic in ML_Antiquotation.inline @{binding "term_pat"} (parser >> term_pat) end val cterm_pat_setup = let val parser = Args.context -- Scan.lift Args.embedded_inner_syntax fun cterm_pat (ctxt, str) = str |> Proof_Context.read_term_pattern ctxt |> ML_Syntax.print_term |> ML_Syntax.atomic |> prefix "Thm.cterm_of ML_context" in ML_Antiquotation.value @{binding "cterm_pat"} (parser >> cterm_pat) end val type_pat_setup = let val parser = Args.context -- Scan.lift Args.embedded_inner_syntax fun typ_pat (ctxt, str) = let val ctxt' = Proof_Context.set_mode Proof_Context.mode_schematic ctxt in str |> Syntax.read_typ ctxt' |> ML_Syntax.print_typ |> ML_Syntax.atomic end in ML_Antiquotation.inline @{binding "typ_pat"} (parser >> typ_pat) end (* Time the given function f : unit -> 'a. *) fun timer (msg, f) = let val t_start = Timing.start () val res = f () val t_end = Timing.result t_start in (writeln (msg ^ (Timing.message t_end)); res) end (* When exception is shown when running function f : unit -> 'a, print stack trace. *) fun exn_trace f = Runtime.exn_trace f end (* structure Util. *) structure Basic_Util: BASIC_UTIL = Util open Basic_Util val _ = Theory.setup (Util.term_pat_setup) val _ = Theory.setup (Util.cterm_pat_setup) val _ = Theory.setup (Util.type_pat_setup) diff --git a/thys/Automatic_Refinement/Lib/Indep_Vars.thy b/thys/Automatic_Refinement/Lib/Indep_Vars.thy --- a/thys/Automatic_Refinement/Lib/Indep_Vars.thy +++ b/thys/Automatic_Refinement/Lib/Indep_Vars.thy @@ -1,59 +1,59 @@ theory Indep_Vars imports Main Refine_Util Mpat_Antiquot begin definition [simp]: "INDEP v \ True" lemma INDEPI: "INDEP v" by simp ML \ signature INDEP_VARS = sig val indep_tac: Proof.context -> tactic' end structure Indep_Vars: INDEP_VARS = struct local fun vsubterms (Abs (_,_,t)) = vsubterms t | vsubterms (t as (_$_)) = let val (f,args) = strip_comb t val args_vsts = map vsubterms args |> flat in case f of (Var (name,vT)) => [(name,vT,fastype_of t,args)]@args_vsts | _ => vsubterms f @ args_vsts end | vsubterms _ = [] fun indep_vars ctxt t st = let val cert = Thm.cterm_of ctxt fun inst_of (name,vT,T,args) = let val Ts = map fastype_of args |> rev val t' = fold absdummy Ts (Var (name,T)) val inst = ((name, vT), cert t') in inst end val inst = vsubterms t |> distinct ((op =) o apply2 #1) |> map inst_of - val st' = Drule.instantiate_normalize ([],inst) st + val st' = Drule.instantiate_normalize (TVars.empty, Vars.make inst) st |> Conv.fconv_rule (Thm.beta_conversion true) in Seq.single st' end fun indep_tac_aux ctxt i st = case Logic.concl_of_goal (Thm.prop_of st) i of @{mpat "Trueprop (INDEP ?v)"} => (indep_vars ctxt v THEN resolve_tac ctxt @{thms INDEPI} i) st | _ => Seq.empty in (* Remove explicit parameters from schematic variable. *) fun indep_tac ctxt = IF_EXGOAL (CONVERSION Thm.eta_conversion THEN' indep_tac_aux ctxt) end end \ end diff --git a/thys/Automatic_Refinement/Lib/Refine_Util.thy b/thys/Automatic_Refinement/Lib/Refine_Util.thy --- a/thys/Automatic_Refinement/Lib/Refine_Util.thy +++ b/thys/Automatic_Refinement/Lib/Refine_Util.thy @@ -1,964 +1,964 @@ section "General Utilities" theory Refine_Util imports Refine_Util_Bootstrap1 Mpat_Antiquot Mk_Term_Antiquot begin definition conv_tag where "conv_tag n x == x" \ \Used internally for @{text "pat_conv"}-conversion\ lemma shift_lambda_left: "(f \ \x. g x) \ (\x. f x \ g x)" by simp ML \ infix 0 THEN_ELSE' THEN_ELSE_COMB' infix 1 THEN_ALL_NEW_FWD THEN_INTERVAL infix 2 ORELSE_INTERVAL infix 3 ->> signature BASIC_REFINE_UTIL = sig include BASIC_REFINE_UTIL (* Resolution with matching *) val RSm: Proof.context -> thm -> thm -> thm val is_Abs: term -> bool val is_Comb: term -> bool val has_Var: term -> bool val is_TFree: typ -> bool val is_def_thm: thm -> bool (* Tacticals *) type tactic' = int -> tactic type itactic = int -> int -> tactic val IF_EXGOAL: (int -> tactic) -> tactic' val COND': (term -> bool) -> tactic' val CONCL_COND': (term -> bool) -> tactic' val THEN_ELSE': tactic' * (tactic' * tactic') -> tactic' val THEN_ELSE_COMB': tactic' * ((tactic'*tactic'->tactic') * tactic' * tactic') -> tactic' val INTERVAL_FWD: tactic' -> int -> int -> tactic val THEN_ALL_NEW_FWD: tactic' * tactic' -> tactic' val REPEAT_ALL_NEW_FWD: tactic' -> tactic' val REPEAT_DETERM': tactic' -> tactic' val REPEAT': tactic' -> tactic' val ALL_GOALS_FWD': tactic' -> tactic' val ALL_GOALS_FWD: tactic' -> tactic val APPEND_LIST': tactic' list -> tactic' val SINGLE_INTERVAL: itactic -> tactic' val THEN_INTERVAL: itactic * itactic -> itactic val ORELSE_INTERVAL: itactic * itactic -> itactic val CAN': tactic' -> tactic' val NTIMES': tactic' -> int -> tactic' (* Only consider results that solve subgoal. If none, return all results unchanged. *) val TRY_SOLVED': tactic' -> tactic' (* Case distinction with tactics. Generalization of THEN_ELSE to lists. *) val CASES': (tactic' * tactic) list -> tactic' (* Tactic that depends on subgoal term structure *) val WITH_subgoal: (term -> tactic') -> tactic' (* Tactic that depends on subgoal's conclusion term structure *) val WITH_concl: (term -> tactic') -> tactic' (* Tactic version of Variable.trade. Import, apply tactic, and export results. One effect is that schematic variables in the goal are fixed, and thus cannot be instantiated by the tactic. *) val TRADE: (Proof.context -> tactic') -> Proof.context -> tactic' (* Tactics *) val fo_rtac: thm -> Proof.context -> tactic' val fo_resolve_tac: thm list -> Proof.context -> tactic' val rprems_tac: Proof.context -> tactic' val rprem_tac: int -> Proof.context -> tactic' val elim_all_tac: Proof.context -> thm list -> tactic val prefer_tac: int -> tactic val insert_subgoal_tac: cterm -> tactic' val insert_subgoals_tac: cterm list -> tactic' val eqsubst_inst_tac: Proof.context -> bool -> int list -> ((indexname * Position.T) * string) list -> thm -> int -> tactic val eqsubst_inst_meth: (Proof.context -> Proof.method) context_parser (* Parsing *) val ->> : 'a context_parser *('a * Context.generic -> 'b * Context.generic) -> 'b context_parser end signature REFINE_UTIL = sig include BASIC_REFINE_UTIL val order_by: ('a * 'a -> order) -> ('b -> 'a) -> 'b list -> 'b list val build_res_net: thm list -> (int * thm) Net.net (* Terms *) val fo_matchp: theory -> cterm -> term -> term list option val fo_matches: theory -> cterm -> term -> bool val anorm_typ: typ -> typ val anorm_term: term -> term val import_cterms: bool -> cterm list -> Proof.context -> cterm list * Proof.context val subsume_sort: ('a -> term) -> theory -> 'a list -> 'a list val subsume_sort_gen: ('a -> term) -> Context.generic -> 'a list -> 'a list val mk_compN1: typ list -> int -> term -> term -> term val mk_compN: int -> term -> term -> term val dest_itselfT: typ -> typ val dummify_tvars: term -> term (* a\\x. f x \ a ?x \ f ?x *) val shift_lambda_left: thm -> thm val shift_lambda_leftN: int -> thm -> thm (* Left-Bracketed Structures *) (* Map [] to z, and [x1,...,xN] to f(...f(f(x1,x2),x3)...) *) val list_binop_left: 'a -> ('a * 'a -> 'a) -> 'a list -> 'a (* Map [] to z, [x] to i x, [x1,...,xN] to f(...f(f(x1,x2),x3)...), thread state *) val fold_binop_left: ('c -> 'b * 'c) -> ('a -> 'c -> 'b * 'c) -> ('b * 'b -> 'b) -> 'a list -> 'c -> 'b * 'c (* Tuples, handling () as empty tuple *) val strip_prodT_left: typ -> typ list val list_prodT_left: typ list -> typ val mk_ltuple: term list -> term (* Fix a tuple of new frees *) val fix_left_tuple_from_Ts: string -> typ list -> Proof.context -> term * Proof.context (* HO-Patterns with tuples *) (* Lambda-abstraction over list of terms, recognizing tuples *) val lambda_tuple: term list -> term -> term (* Instantiate tuple-types in specified variables *) val instantiate_tuples: Proof.context -> (indexname*typ) list -> thm -> thm (* Instantiate tuple-types in variables from given term *) val instantiate_tuples_from_term_tac: Proof.context -> term -> tactic (* Instantiate tuple types in variables of subgoal *) val instantiate_tuples_subgoal_tac: Proof.context -> tactic' (* Rules *) val abs_def: Proof.context -> thm -> thm (* Rule combinators *) (* Iterate rule on theorem until it fails *) val repeat_rule: (thm -> thm) -> thm -> thm (* Apply rule on theorem and assert that theorem was changed *) val changed_rule: (thm -> thm) -> thm -> thm (* Try rule on theorem, return theorem unchanged if rule fails *) val try_rule: (thm -> thm) -> thm -> thm (* Singleton version of Variable.trade *) val trade_rule: (Proof.context -> thm -> thm) -> Proof.context -> thm -> thm (* Combine with first matching theorem *) val RS_fst: thm -> thm list -> thm (* Instantiate first matching theorem *) val OF_fst: thm list -> thm list -> thm (* Conversion *) val trace_conv: conv val monitor_conv: string -> conv -> conv val monitor_conv': string -> (Proof.context -> conv) -> Proof.context -> conv val fixup_vars: cterm -> thm -> thm val fixup_vars_conv: conv -> conv val fixup_vars_conv': (Proof.context -> conv) -> Proof.context -> conv val pat_conv': cterm -> (string -> Proof.context -> conv) -> Proof.context -> conv val pat_conv: cterm -> (Proof.context -> conv) -> Proof.context -> conv val HOL_concl_conv: (Proof.context -> conv) -> Proof.context -> conv val import_conv: (Proof.context -> conv) -> Proof.context -> conv val fix_conv: Proof.context -> conv -> conv val ite_conv: conv -> conv -> conv -> conv val cfg_trace_f_tac_conv: bool Config.T val f_tac_conv: Proof.context -> (term -> term) -> tactic -> conv (* Conversion combinators to choose first matching position *) (* Try argument, then function *) val fcomb_conv: conv -> conv (* Descend over function or abstraction *) val fsub_conv: (Proof.context -> conv) -> Proof.context -> conv (* Apply to topmost matching position *) val ftop_conv: (Proof.context -> conv) -> Proof.context -> conv (* Parsing *) val parse_bool_config: string -> bool Config.T -> bool context_parser val parse_paren_list: 'a context_parser -> 'a list context_parser val parse_paren_lists: 'a context_parser -> 'a list list context_parser (* 2-step configuration parser *) (* Parse boolean config, name or no_name. *) val parse_bool_config': string -> bool Config.T -> Token.T list -> (bool Config.T * bool) * Token.T list (* Parse optional (p1,...,pn). Empty list if nothing parsed. *) val parse_paren_list': 'a parser -> Token.T list -> 'a list * Token.T list (* Apply list of (config,value) pairs *) val apply_configs: ('a Config.T * 'a) list -> Proof.context -> Proof.context end structure Refine_Util: REFINE_UTIL = struct open Basic_Refine_Util fun RSm ctxt thA thB = let val (thA, ctxt') = ctxt |> Variable.declare_thm thA |> Variable.declare_thm thB |> yield_singleton (apfst snd oo Variable.import true) thA val thm = thA RS thB val thm = singleton (Variable.export ctxt' ctxt) thm |> Drule.zero_var_indexes in thm end fun is_Abs (Abs _) = true | is_Abs _ = false fun is_Comb (_$_) = true | is_Comb _ = false fun has_Var (Var _) = true | has_Var (Abs (_,_,t)) = has_Var t | has_Var (t1$t2) = has_Var t1 orelse has_Var t2 | has_Var _ = false fun is_TFree (TFree _) = true | is_TFree _ = false fun is_def_thm thm = case thm |> Thm.prop_of of Const (@{const_name "Pure.eq"},_)$_$_ => true | _ => false type tactic' = int -> tactic type itactic = int -> int -> tactic (* Fail if subgoal does not exist *) fun IF_EXGOAL tac i st = if i <= Thm.nprems_of st then tac i st else no_tac st; fun COND' P = IF_EXGOAL (fn i => fn st => (if P (Thm.prop_of st |> curry Logic.nth_prem i) then all_tac st else no_tac st) handle TERM _ => no_tac st | Pattern.MATCH => no_tac st ) (* FIXME: Subtle difference between Logic.concl_of_goal and this: concl_of_goal converts loose bounds to frees! *) fun CONCL_COND' P = COND' (strip_all_body #> Logic.strip_imp_concl #> P) fun (tac1 THEN_ELSE' (tac2,tac3)) x = tac1 x THEN_ELSE (tac2 x,tac3 x); (* If first tactic succeeds, combine its effect with "comb tac2", otherwise use tac_else. Example: tac1 THEN_ELSE_COMB ((THEN_ALL_NEW),tac2,tac_else) *) fun (tac1 THEN_ELSE_COMB' (comb,tac2,tac_else)) i st = let val rseq = tac1 i st in case seq_is_empty rseq of (true,_) => tac_else i st | (false,rseq) => comb (K(K( rseq )), tac2) i st end (* Apply tactic to subgoals in interval, in a forward manner, skipping over emerging subgoals *) fun INTERVAL_FWD tac l u st = if l>u then all_tac st else (tac l THEN (fn st' => let val ofs = Thm.nprems_of st' - Thm.nprems_of st; in if ofs < ~1 then raise THM ( "INTERVAL_FWD: Tac solved more than one goal",~1,[st,st']) else INTERVAL_FWD tac (l+1+ofs) (u+ofs) st' end)) st; (* Apply tac2 to all subgoals emerged from tac1, in forward manner. *) fun (tac1 THEN_ALL_NEW_FWD tac2) i st = (tac1 i THEN (fn st' => INTERVAL_FWD tac2 i (i + Thm.nprems_of st' - Thm.nprems_of st) st') ) st; fun REPEAT_ALL_NEW_FWD tac = tac THEN_ALL_NEW_FWD (TRY o (fn i => REPEAT_ALL_NEW_FWD tac i)); (* Repeat tac on subgoal. Determinize each step. Stop if tac fails or subgoal is solved. *) fun REPEAT_DETERM' tac i st = let val n = Thm.nprems_of st in REPEAT_DETERM (COND (has_fewer_prems n) no_tac (tac i)) st end fun REPEAT' tac i st = let val n = Thm.nprems_of st in REPEAT (COND (has_fewer_prems n) no_tac (tac i)) st end (* Apply tactic to all goals in a forward manner. Newly generated goals are ignored. *) fun ALL_GOALS_FWD' tac i st = (tac i THEN (fn st' => let val i' = i + Thm.nprems_of st' + 1 - Thm.nprems_of st; in if i' <= Thm.nprems_of st' then ALL_GOALS_FWD' tac i' st' else all_tac st' end )) st; fun ALL_GOALS_FWD tac = ALL_GOALS_FWD' tac 1; fun APPEND_LIST' tacs = fold_rev (curry (op APPEND')) tacs (K no_tac); fun SINGLE_INTERVAL tac i = tac i i fun ((tac1:itactic) THEN_INTERVAL (tac2:itactic)) = (fn i => fn j => fn st => ( tac1 i j THEN (fn st' => tac2 i (j + Thm.nprems_of st' - Thm.nprems_of st) st') ) st ):itactic fun tac1 ORELSE_INTERVAL tac2 = (fn i => fn j => tac1 i j ORELSE tac2 i j) (* Fail if tac fails, otherwise do nothing *) fun CAN' tac i st = case tac i st |> Seq.pull of NONE => Seq.empty | SOME _ => Seq.single st (* Repeat tactic n times *) fun NTIMES' _ 0 _ st = Seq.single st | NTIMES' tac n i st = (tac THEN' NTIMES' tac (n-1)) i st (* Resolve with rule. Use first-order unification. From cookbook, added exception handling *) fun fo_rtac thm = Subgoal.FOCUS (fn {context = ctxt, concl, ...} => let val concl_pat = Drule.strip_imp_concl (Thm.cprop_of thm) val insts = Thm.first_order_match (concl_pat, concl) in resolve_tac ctxt [Drule.instantiate_normalize insts thm] 1 end handle Pattern.MATCH => no_tac ) fun fo_resolve_tac thms ctxt = FIRST' (map (fn thm => fo_rtac thm ctxt) thms); (* Resolve with premises. Copied and adjusted from Goal.assume_rule_tac. *) fun rprems_tac ctxt = Goal.norm_hhf_tac ctxt THEN' CSUBGOAL (fn (goal, i) => let fun non_atomic (Const (@{const_name Pure.imp}, _) $ _ $ _) = true | non_atomic (Const (@{const_name Pure.all}, _) $ _) = true | non_atomic _ = false; val ((_, goal'), ctxt') = Variable.focus_cterm NONE goal ctxt; val goal'' = Drule.cterm_rule (singleton (Variable.export ctxt' ctxt)) goal'; val Rs = filter (non_atomic o Thm.term_of) (Drule.strip_imp_prems goal''); val ethms = Rs |> map (fn R => (Simplifier.norm_hhf ctxt (Thm.trivial R))); in eresolve_tac ctxt ethms i end ); (* Resolve with premise. Copied and adjusted from Goal.assume_rule_tac. *) fun rprem_tac n ctxt = Goal.norm_hhf_tac ctxt THEN' CSUBGOAL (fn (goal, i) => let val ((_, goal'), ctxt') = Variable.focus_cterm NONE goal ctxt; val goal'' = Drule.cterm_rule (singleton (Variable.export ctxt' ctxt)) goal'; val R = nth (Drule.strip_imp_prems goal'') (n - 1) val rl = Simplifier.norm_hhf ctxt (Thm.trivial R) in eresolve_tac ctxt [rl] i end ); fun elim_all_tac ctxt thms = ALLGOALS (REPEAT_ALL_NEW (ematch_tac ctxt thms)) fun prefer_tac i = defer_tac i THEN PRIMITIVE (Thm.permute_prems 0 ~1) fun order_by ord f = sort (ord o apply2 f) (* CLONE from tactic.ML *) local (*insert one tagged rl into the net*) fun insert_krl (krl as (_,th)) = Net.insert_term (K false) (Thm.concl_of th, krl); in (*build a net of rules for resolution*) fun build_res_net rls = fold_rev insert_krl (tag_list 1 rls) Net.empty; end fun insert_subgoals_tac cts i = PRIMITIVE ( Thm.permute_prems 0 (i - 1) #> fold_rev Thm.implies_intr cts #> Thm.permute_prems 0 (~i + 1) ) fun insert_subgoal_tac ct i = insert_subgoals_tac [ct] i local (* Substitution with dynamic instantiation of parameters. By Lars Noschinski. *) fun eqsubst_tac' ctxt asm = if asm then EqSubst.eqsubst_asm_tac ctxt else EqSubst.eqsubst_tac ctxt fun subst_method inst_tac tac = Args.goal_spec -- Scan.lift (Args.mode "asm" -- Scan.optional (Args.parens (Scan.repeat Parse.nat)) [0]) -- Scan.optional (Scan.lift (Parse.and_list1 (Parse.position Args.var -- (Args.$$$ "=" |-- Parse.!!! Args.embedded_inner_syntax)) --| Args.$$$ "in")) [] -- Attrib.thms >> (fn (((quant, (asm, occL)), insts), thms) => fn ctxt => METHOD (fn facts => if null insts then quant (Method.insert_tac ctxt facts THEN' tac ctxt asm occL thms) else (case thms of [thm] => quant ( Method.insert_tac ctxt facts THEN' inst_tac ctxt asm occL insts thm) | _ => error "Cannot have instantiations with multiple rules"))); in fun eqsubst_inst_tac ctxt asm occL insts thm = Subgoal.FOCUS ( fn {context=ctxt,...} => let val ctxt' = ctxt |> Proof_Context.set_mode Proof_Context.mode_schematic (* FIXME !? *) val thm' = thm |> Rule_Insts.read_instantiate ctxt' insts [] in eqsubst_tac' ctxt asm occL [thm'] 1 end ) ctxt val eqsubst_inst_meth = subst_method eqsubst_inst_tac eqsubst_tac' end (* Match pattern against term, and return list of values for non-dummy variables. A variable is considered dummy if its name starts with an underscore (_)*) fun fo_matchp thy cpat t = let fun ignore (Var ((name,_),_)) = String.isPrefix "_" name | ignore _ = true; val pat = Thm.term_of cpat; val pvars = fold_aterms ( fn t => fn l => if is_Var t andalso not (ignore t) then t::l else l ) pat [] |> rev val inst = Pattern.first_order_match thy (pat,t) (Vartab.empty,Vartab.empty); in SOME (map (Envir.subst_term inst) pvars) end handle Pattern.MATCH => NONE; val fo_matches = is_some ooo fo_matchp fun anorm_typ ty = let val instT = Term.add_tvarsT ty [] |> map_index (fn (i,(n,s)) => (n,TVar (("t"^string_of_int i,0),s))) val ty = Term.typ_subst_TVars instT ty; in ty end; fun anorm_term t = let val instT = Term.add_tvars t [] |> map_index (fn (i,(n,s)) => (n,TVar (("t"^string_of_int i,0),s))) val t = Term.subst_TVars instT t; val inst = Term.add_vars t [] |> map_index (fn (i,(n,s)) => (n,Var (("v"^string_of_int i,0),s))) val t = Term.subst_Vars inst t; in t end; fun import_cterms is_open cts ctxt = let val ts = map Thm.term_of cts val (ts', ctxt') = Variable.import_terms is_open ts ctxt val cts' = map (Thm.cterm_of ctxt) ts' in (cts', ctxt') end (* Order a list of items such that more specific items come before less specific ones. The term to be compared is extracted by a function that is given as parameter. *) fun subsume_sort f thy items = let val rhss = map (Envir.beta_eta_contract o f) items fun freqf thy net rhs = Net.match_term net rhs |> filter (fn p => Pattern.matches thy (p,rhs)) |> length val net = fold (fn rhs => Net.insert_term_safe (op =) (rhs,rhs)) rhss Net.empty val freqs = map (freqf thy net) rhss val res = freqs ~~ items |> sort (rev_order o int_ord o apply2 fst) |> map snd in res end fun subsume_sort_gen f = subsume_sort f o Context.theory_of fun mk_comp1 env (f, g) = let val fT = fastype_of1 (env, f); val gT = fastype_of1 (env, g); val compT = fT --> gT --> domain_type gT --> range_type fT; in Const ("Fun.comp", compT) $ f $ g end; fun mk_compN1 _ 0 f g = f$g | mk_compN1 env 1 f g = mk_comp1 env (f, g) | mk_compN1 env n f g = let val T = fastype_of1 (env, g) |> domain_type val g = incr_boundvars 1 g $ Bound 0 val env = T::env in Abs ("x"^string_of_int n,T,mk_compN1 env (n-1) f g) end val mk_compN = mk_compN1 [] fun abs_def ctxt = Local_Defs.meta_rewrite_rule ctxt #> Drule.abs_def fun trace_conv ct = (tracing (@{make_string} ct); Conv.all_conv ct); fun monitor_conv msg conv ct = let val _ = tracing (msg ^ " (gets): " ^ @{make_string} ct); val res = conv ct handle exc => (if Exn.is_interrupt exc then Exn.reraise exc else tracing (msg ^ " (raises): " ^ @{make_string} exc); Exn.reraise exc) val _ = tracing (msg ^ " (yields): " ^ @{make_string} res); in res end fun monitor_conv' msg conv ctxt ct = monitor_conv msg (conv ctxt) ct fun fixup_vars ct thm = let val lhs = Thm.lhs_of thm val inst = Thm.first_order_match (lhs,ct) val thm' = Thm.instantiate inst thm in thm' end fun fixup_vars_conv conv ct = fixup_vars ct (conv ct) fun fixup_vars_conv' conv ctxt = fixup_vars_conv (conv ctxt) local fun tag_ct ctxt name ct = let val t = Thm.term_of ct; val ty = fastype_of t; val t' = Const (@{const_name conv_tag},@{typ unit}-->ty-->ty) $Free (name,@{typ unit})$t; val ct' = Thm.cterm_of ctxt t'; in ct' end fun mpat_conv pat ctxt ct = let val (tym,tm) = Thm.first_order_match (pat,ct); - val tm' = map (fn (pt as ((name, _), _),ot) => (pt, tag_ct ctxt name ot)) tm; - val ct' = Thm.instantiate_cterm (tym,tm') pat; + val tm' = Vars.map (fn ((name, _), _) => tag_ct ctxt name) tm; + val ct' = Thm.instantiate_cterm (tym, tm') pat; val rthm = Goal.prove_internal ctxt [] (Thm.cterm_of ctxt (Logic.mk_equals (apply2 Thm.term_of (ct, ct')))) (K (simp_tac (put_simpset HOL_basic_ss ctxt addsimps @{thms conv_tag_def}) 1)) |> Goal.norm_result ctxt in fixup_vars ct rthm end handle Pattern.MATCH => raise (CTERM ("mpat_conv: No match",[pat,ct])); fun tag_conv cnv ctxt ct = case Thm.term_of ct of Const (@{const_name conv_tag},_)$Free(name,_)$_ => ( (Conv.rewr_conv (@{thm conv_tag_def}) then_conv (cnv name) ctxt) ct) | _ => Conv.all_conv ct; fun all_tag_conv cnv = Conv.bottom_conv (tag_conv cnv); in (* Match against pattern, and apply parameter conversion to matches of variables prefixed by hole_prefix. *) fun pat_conv' cpat cnv ctxt = mpat_conv cpat ctxt then_conv (all_tag_conv cnv ctxt); fun pat_conv cpat conv = pat_conv' cpat (fn name => case name of "HOLE" => conv | _ => K Conv.all_conv); end fun HOL_concl_conv cnv = Conv.params_conv ~1 (fn ctxt => Conv.concl_conv ~1 ( HOLogic.Trueprop_conv (cnv ctxt))); fun import_conv conv ctxt ct = let val (ct',ctxt') = yield_singleton (import_cterms true) ct ctxt val res = conv ctxt' ct' val res' = singleton (Variable.export ctxt' ctxt) res |> fixup_vars ct in res' end fun fix_conv ctxt conv ct = let val thm = conv ct val eq = Logic.mk_equals (Thm.term_of ct, Thm.term_of ct) |> head_of in if (Thm.term_of (Thm.lhs_of thm) aconv Thm.term_of ct) then thm else thm RS Thm.trivial (Thm.mk_binop (Thm.cterm_of ctxt eq) ct (Thm.rhs_of thm)) end fun ite_conv cv cv1 cv2 ct = let val eq1 = SOME (cv ct) handle THM _ => NONE | CTERM _ => NONE | TERM _ => NONE | TYPE _ => NONE; val res = case eq1 of NONE => cv2 ct | SOME eq1 => let val eq2 = cv1 (Thm.rhs_of eq1) in if Thm.is_reflexive eq1 then eq2 else if Thm.is_reflexive eq2 then eq1 else Thm.transitive eq1 eq2 end in res end val cfg_trace_f_tac_conv = Attrib.setup_config_bool @{binding trace_f_tac_conv} (K false) (* Transform term and prove equality to original by tactic *) fun f_tac_conv ctxt f tac ct = let val t = Thm.term_of ct val t' = f t val goal = Logic.mk_equals (t,t') val _ = if Config.get ctxt cfg_trace_f_tac_conv then tracing (Syntax.string_of_term ctxt goal) else () val goal = Thm.cterm_of ctxt goal val thm = Goal.prove_internal ctxt [] goal (K tac) in thm end (* Apply function to result and context *) fun (p->>f) ctks = let val (res,(context,tks)) = p ctks val (res,context) = f (res, context) in (res,(context,tks)) end fun parse_bool_config name cfg = ( Scan.lift (Args.$$$ name) ->> (apsnd (Config.put_generic cfg true) #>> K true) || Scan.lift (Args.$$$ ("no_"^name)) ->> (apsnd (Config.put_generic cfg false) #>> K false) ) fun parse_paren_list p = Scan.lift ( Args.$$$ "(") |-- Parse.enum1' "," p --| Scan.lift (Args.$$$ ")" ) fun parse_paren_lists p = Scan.repeat (parse_paren_list p) val _ = Theory.setup (Method.setup @{binding fo_rule} (Attrib.thms >> (fn thms => fn ctxt => SIMPLE_METHOD' ( fo_resolve_tac thms ctxt))) "resolve using first-order matching" #> Method.setup @{binding rprems} (Scan.lift (Scan.option Parse.nat) >> (fn i => fn ctxt => SIMPLE_METHOD' ( case i of NONE => rprems_tac ctxt | SOME i => rprem_tac i ctxt )) ) "resolve with premises" #> Method.setup @{binding elim_all} (Attrib.thms >> (fn thms => fn ctxt => SIMPLE_METHOD (elim_all_tac ctxt thms))) "repeteadly apply elimination rules to all subgoals" #> Method.setup @{binding subst_tac} eqsubst_inst_meth "single-step substitution (dynamic instantiation)" #> Method.setup @{binding clarsimp_all} ( Method.sections Clasimp.clasimp_modifiers >> K (fn ctxt => SIMPLE_METHOD ( CHANGED_PROP (ALLGOALS (Clasimp.clarsimp_tac ctxt)))) ) "simplify and clarify all subgoals") (* Filter alternatives that solve a subgoal. If no alternative solves goal, return result sequence unchanged *) fun TRY_SOLVED' tac i st = let val res = tac i st val solved = Seq.filter (fn st' => Thm.nprems_of st' < Thm.nprems_of st) res in case Seq.pull solved of SOME _ => solved | NONE => res end local fun CASES_aux [] = no_tac | CASES_aux ((tac1, tac2)::cs) = tac1 1 THEN_ELSE (tac2, CASES_aux cs) in (* Accepts a list of pairs of (pattern_tac', worker_tac), and applies worker_tac to results of first successful pattern_tac'. *) val CASES' = SELECT_GOAL o CASES_aux end (* TODO/FIXME: There seem to be no guarantees when eta-long forms are introduced by unification. So, we have to expect eta-long forms everywhere, which may be a problem when matching terms syntactically. *) fun WITH_subgoal tac = CONVERSION Thm.eta_conversion THEN' IF_EXGOAL (fn i => fn st => tac (nth (Thm.prems_of st) (i - 1)) i st) fun WITH_concl tac = CONVERSION Thm.eta_conversion THEN' IF_EXGOAL (fn i => fn st => tac (Logic.concl_of_goal (Thm.prop_of st) i) i st ) fun TRADE tac ctxt i st = let val orig_ctxt = ctxt val (st,ctxt) = yield_singleton (apfst snd oo Variable.import true) st ctxt val seq = tac ctxt i st |> Seq.map (singleton (Variable.export ctxt orig_ctxt)) in seq end (* Try argument, then function *) fun fcomb_conv conv = let open Conv in arg_conv conv else_conv fun_conv conv end (* Descend over function or abstraction *) fun fsub_conv conv ctxt = let open Conv in fcomb_conv (conv ctxt) else_conv abs_conv (conv o snd) ctxt else_conv no_conv end (* Apply to topmost matching position *) fun ftop_conv conv ctxt ct = (conv ctxt else_conv fsub_conv (ftop_conv conv) ctxt) ct (* Iterate rule on theorem until it fails *) fun repeat_rule n thm = case try n thm of SOME thm => repeat_rule n thm | NONE => thm (* Apply rule on theorem and assert that theorem was changed *) fun changed_rule n thm = let val thm' = n thm in if Thm.eq_thm_prop (thm, thm') then raise THM ("Same",~1,[thm,thm']) else thm' end (* Try rule on theorem *) fun try_rule n thm = case try n thm of SOME thm => thm | NONE => thm fun trade_rule f ctxt thm = singleton (Variable.trade (map o f) ctxt) thm fun RS_fst thm thms = let fun r [] = raise THM ("RS_fst, no matches",~1,thm::thms) | r (thm'::thms) = case try (op RS) (thm,thm') of NONE => r thms | SOME thm => thm in r thms end fun OF_fst thms insts = let fun r [] = raise THM ("OF_fst, no matches",length thms,thms@insts) | r (thm::thms) = case try (op OF) (thm,insts) of NONE => r thms | SOME thm => thm in r thms end (* Map [] to z, and [x1,...,xN] to f(...f(f(x1,x2),x3)...) *) fun list_binop_left z f = let fun r [] = z | r [T] = T | r (T::Ts) = f (r Ts,T) in fn l => r (rev l) end (* Map [] to z, [x] to i x, [x1,...,xN] to f(...f(f(x1,x2),x3)...), thread state *) fun fold_binop_left z i f = let fun r [] ctxt = z ctxt | r [T] ctxt = i T ctxt | r (T::Ts) ctxt = let val (Ti,ctxt) = i T ctxt val (Tsi,ctxt) = r Ts ctxt in (f (Tsi,Ti),ctxt) end in fn l => fn ctxt => r (rev l) ctxt end fun strip_prodT_left (Type (@{type_name Product_Type.prod},[A,B])) = strip_prodT_left A @ [B] | strip_prodT_left (Type (@{type_name Product_Type.unit},[])) = [] | strip_prodT_left T = [T] val list_prodT_left = list_binop_left HOLogic.unitT HOLogic.mk_prodT (* Make tuple with left-bracket structure *) val mk_ltuple = list_binop_left HOLogic.unit HOLogic.mk_prod (* Fix a tuple of new frees *) fun fix_left_tuple_from_Ts name = fold_binop_left (fn ctxt => (@{term "()"},ctxt)) (fn T => fn ctxt => let val (x,ctxt) = yield_singleton Variable.variant_fixes name ctxt val x = Free (x,T) in (x,ctxt) end) HOLogic.mk_prod (* Replace all type-vars by dummyT *) val dummify_tvars = map_types (map_type_tvar (K dummyT)) fun dest_itselfT (Type (@{type_name itself},[A])) = A | dest_itselfT T = raise TYPE("dest_itselfT",[T],[]) fun shift_lambda_left thm = thm RS @{thm shift_lambda_left} fun shift_lambda_leftN i = funpow i shift_lambda_left (* TODO: Naming should be without ' for basic parse, and with ' for context_parser! *) fun parse_bool_config' name cfg = (Args.$$$ name #>> K (cfg,true)) || (Args.$$$ ("no_"^name) #>> K (cfg,false)) fun parse_paren_list' p = Scan.optional (Args.parens (Parse.enum1 "," p)) [] fun apply_configs l ctxt = fold (fn (cfg,v) => fn ctxt => Config.put cfg v ctxt) l ctxt fun lambda_tuple [] t = t | lambda_tuple (@{mpat "(?a,?b)"}::l) t = let val body = lambda_tuple (a::b::l) t in @{mk_term "case_prod ?body"} end | lambda_tuple (x::l) t = lambda x (lambda_tuple l t) fun get_tuple_inst ctxt (iname,T) = let val (argTs,T) = strip_type T fun cr (Type (@{type_name prod},[T1,T2])) ctxt = let val (x1,ctxt) = cr T1 ctxt val (x2,ctxt) = cr T2 ctxt in (HOLogic.mk_prod (x1,x2), ctxt) end | cr T ctxt = let val (name, ctxt) = yield_singleton Variable.variant_fixes "x" ctxt in (Free (name,T),ctxt) end val ctxt = Variable.set_body false ctxt (* Prevent generation of skolem-names *) val (args,ctxt) = fold_map cr argTs ctxt fun fl (@{mpat "(?x,?y)"}) = fl x @ fl y | fl t = [t] val fargs = flat (map fl args) val fTs = map fastype_of fargs val v = Var (iname,fTs ---> T) val v = list_comb (v,fargs) val v = lambda_tuple args v in Thm.cterm_of ctxt v end fun instantiate_tuples ctxt inTs = let val inst = inTs ~~ map (get_tuple_inst ctxt) inTs in - Thm.instantiate ([],inst) + Thm.instantiate (TVars.empty, Vars.make inst) end val _ = COND' fun instantiate_tuples_from_term_tac ctxt t st = let val vars = Term.add_vars t [] in PRIMITIVE (instantiate_tuples ctxt vars) st end fun instantiate_tuples_subgoal_tac ctxt = WITH_subgoal (fn t => K (instantiate_tuples_from_term_tac ctxt t)) end structure Basic_Refine_Util: BASIC_REFINE_UTIL = Refine_Util open Basic_Refine_Util \ attribute_setup zero_var_indexes = \ Scan.succeed (Thm.rule_attribute [] (K Drule.zero_var_indexes)) \ "Set variable indexes to zero, renaming to avoid clashes" end diff --git a/thys/Case_Labeling/casify.ML b/thys/Case_Labeling/casify.ML --- a/thys/Case_Labeling/casify.ML +++ b/thys/Case_Labeling/casify.ML @@ -1,399 +1,400 @@ signature CASIFY = sig datatype options = Options of { simp_all_cases: bool, split_right_only: bool, protect_subgoals: bool } val hyp_subst_tac: Proof.context -> int -> tactic val SPLIT_subst_tac: Proof.context -> int -> tactic val extract_cases_tac: context_tactic val prepare_labels_tac: Proof.context -> int -> tactic val split_bind_all_tac: {right_only: bool, simp_all_cases: bool} -> Proof.context -> int -> tactic val casify_tac: options -> context_tactic val casify_options: options -> options parser val casify_method_setup: options -> (Proof.context -> Method.method) context_parser end structure Casify : CASIFY = struct val bind_unnamedN = "case" val case_premsN = "prems" val case_unnamedN = "unnamed" datatype 'a prg_ctxt = Block of (string * 'a list) datatype 'a prg_concl = Prg_Concl of ((string * 'a) option * term) fun dest_var (Const (@{const_name Case_Labeling.VAR}, _) $ t) = Util.dest_tuple t | dest_var t = raise TERM ("dest_var", [t]) val dest_vars = HOLogic.dest_list #> maps dest_var val dest_ct = let fun mk_block na ic vs = let val (_, idx) = HOLogic.dest_number ic in (idx, Block (HOLogic.dest_string na, dest_vars vs)) end fun dest_block (Const (@{const_name Pair}, _) $ na $ (Const (@{const_name Pair}, _) $ ic $ vs)) = mk_block na ic vs | dest_block t = raise TERM ("dest_block", [t]) in HOLogic.dest_list #> rev #> map dest_block end fun dest_VC (Const (@{const_name Case_Labeling.VC}, _) $ ct $ t) = (dest_ct ct, t) | dest_VC t = ([], t) fun try_dest_Trueprop t = case try HOLogic.dest_Trueprop t of NONE => t | SOME t' => t' fun dest_BIND (Const (@{const_name Case_Labeling.BIND}, _) $ na $ ic $ t) = let val s = HOLogic.dest_string na val (_, n) = HOLogic.dest_number ic in (SOME (s,n), t) end | dest_BIND t = (NONE, t) fun dest_SPLIT (Const (@{const_name Case_Labeling.SPLIT}, _) $ t $ u) = (t,u) | dest_SPLIT t = raise TERM ("dest_SPLIT", [t]) fun dest_Bound (Bound i) = [i] | dest_Bound t = raise TERM ("dest_bound", [t]) fun dest_HIER (Const (@{const_name Case_Labeling.HIER},_) $ ct $ t) = (SOME (length (HOLogic.dest_list ct)), t) | dest_HIER t = (NONE, t) fun parse_label prop = let val vars = Term.strip_all_vars prop val ((prems, label), _) = prop |> (Logic.strip_horn o Term.strip_all_body) ||> try_dest_Trueprop ||>> dest_VC in { vars=vars, label=label, prems=prems} end fun parse_prem prop = let val vars = Term.strip_all_vars prop val (prems, (hier, concl)) = prop |> (Logic.strip_horn o Term.strip_all_body) ||> try_dest_Trueprop ||> dest_HIER val prop' = Logic.list_all (vars, Logic.list_implies (prems, HOLogic.mk_Trueprop concl)) in (hier, prop') end fun strip_prg_ctxt n ((params, t), _) = let fun lookup_delete eq y = let fun aux _ [] = error ("Term is not a parameter: " ^ @{make_string} y) | aux acc ((k,v) :: xs) = if eq (k, y) then (v, rev acc @ xs) else aux ((k,v) :: acc) xs in aux [] end fun upd_label_vars ps [] = [(~n, Block (case_unnamedN, map snd ps))] | upd_label_vars ps [(n, Block (s,vs))] = let val (vs', ps') = fold_map (lookup_delete (op aconv)) vs ps in [(n, Block (s, vs' @ map snd ps'))] end | upd_label_vars ps ((n, Block (s,vs)) :: ct) = let val (vs', ps') = fold_map (lookup_delete (op aconv)) vs ps in (n, Block (s,vs')) :: upd_label_vars ps' ct end val build_param_map = map (fn (s,(s', T)) => (Free (s',T), {fix= (s,T), abs=(s',T)})) val (asms, (label, concl)) = t |> Logic.strip_horn ||> try_dest_Trueprop ||> dest_VC val prg_ctxt = upd_label_vars (build_param_map params) label val ctxt_len = length prg_ctxt val asms' = map (parse_prem #> apfst (the_default ctxt_len)) asms in (prg_ctxt, (asms', concl)) end fun strip_prg_ctxts xs = let val n = length xs in map_index (fn (i,x) => strip_prg_ctxt (n - i) x) xs end datatype ('a,'b) precase = Precase of {fixes: (binding * typ) list, assumes: ('b * term list) list, binds: (indexname * term option) list, cases: ('a * ('a,'b) precase) list} fun bindings args = map (apfst Binding.name) args fun coalesce_order ord = sort (Util.fst_ord ord) #> AList.coalesce (is_equal o ord) fun unique_names xs = fst (fold_map (Util.infst Name.variant) xs Name.context) fun build_precase (prg_ctxt, (prems, t))= let val sorted_prems = prems |> map (apsnd (fn t => (fst (dest_BIND (try_dest_Trueprop t)), t))) |> sort (prod_ord int_ord (prod_ord (option_ord (prod_ord fast_string_ord int_ord)) (K EQUAL))) fun drop_labels (Const (@{const_name "Case_Labeling.BIND"}, _) $ _ $ _ $ t ) = drop_labels t | drop_labels (t1 $ t2) = drop_labels t1 $ drop_labels t2 | drop_labels (Abs (x,T,t)) = Abs (x,T, drop_labels t) | drop_labels t = t (* XXX integrate with drop_labels, to save a second pass *) fun find_binds (t as Const (@{const_name "Case_Labeling.BIND"}, _) $ _ $ _ $ _) = ( case dest_BIND t of (NONE, t) => find_binds t | (SOME (s,n), t) => (n,(s,t)) :: find_binds t) | find_binds (t1 $ t2) = find_binds t1 @ find_binds t2 | find_binds (Abs (_,_,t)) = find_binds t | find_binds _ = [] fun has_loose_bounds t = case loose_bnos t of [] => false | _ :: _ => (warning "loose bounds in term"; true) (*XXX*) fun unique_binds _ [] = [] | unique_binds acc ((s,x) :: xs) = (case AList.lookup (op=) acc s of NONE => ((s,0), x) :: unique_binds ((s,0) :: acc) xs | SOME n => ((s,n+1), x) :: unique_binds (AList.update (op=) (s,n+1) acc) xs ) fun mk_precase _ _ [] [] = [] | mk_precase _ _ [] (_::_) = error "premise with a too long HIERarchy label" | mk_precase n abs_ofixes (bl :: prg_ctxt) prems = let val (m, Block (s, vars)) = bl val fixes = map #fix vars val params = map #abs vars val (prems1, prems2) = chop_prefix (fn (m,_) => m = n) prems val abs_fixes = abs_ofixes o fold_rev Term.absfree params val prems' = prems1 |> map (fn (_, (x,t)) => (swap (the_default (case_premsN, ~1) x), t)) |> coalesce_order (prod_ord int_ord string_ord) |> map (apfst snd) |> unique_names val assumes = map (apsnd (map abs_fixes)) prems' val fixes' = bindings fixes val binds = let val prem_bs = maps find_binds (maps snd prems') val (concl_cases, concl_bs) = case prg_ctxt of [] => ([(bind_unnamedN, abs_fixes t)], find_binds t) | (_::_) => ([],[]) val bs = concl_bs @ prem_bs |> sort (prod_ord int_ord (prod_ord string_ord Term_Ord.fast_term_ord)) |> map snd |> map (apsnd abs_fixes) |> filter_out (has_loose_bounds o snd) in concl_cases @ bs |> unique_binds [] |> map (apsnd SOME) end val precase = Precase { fixes = fixes', assumes = assumes, binds = binds, cases = mk_precase (n+1) abs_fixes prg_ctxt prems2 } in [((m,s), precase)] end in mk_precase 1 drop_labels prg_ctxt sorted_prems end fun merge_precases precases = let val _ = precases : ((int * string) * (int * string, string) precase) list (*XXX*) fun merge_p (_, []) = error "empty case" (*XXX*) | merge_p (s, Precase {fixes, assumes, binds, cases} :: pcs) = let fun sel_cases (Precase {cases, ...}) = cases val cases' = merge_precases (cases @ maps sel_cases pcs) in (s, Rule_Cases.Case {fixes = fixes, assumes = assumes, binds = binds, cases = cases'}) end val cases = precases |> coalesce_order (prod_ord int_ord string_ord) |> map (apfst snd) |> unique_names |> map merge_p in cases end fun mk_cases' ctxt = Thm.prems_of #> map (fn t => Variable.focus NONE t ctxt) #> strip_prg_ctxts #> maps build_precase #> merge_precases #> map (apsnd SOME) fun normalize_conv cv ctxt ct = Conv.bottom_conv (fn ctxt => Conv.try_conv ( cv then_conv normalize_conv cv ctxt )) ctxt ct fun prepare_labels_tac ctxt = let val suc_numeral_simps = @{thms Suc_numeral_simps[THEN eq_reflection] append.simps[THEN eq_reflection]} val app_simps = @{thms append.simps[THEN eq_reflection]} val suc_to_num_conv = normalize_conv (Conv.rewrs_conv (suc_numeral_simps @ app_simps)) val label_fun_conv = Conv.fun_conv o suc_to_num_conv fun label_conv ctxt ct = (case Thm.term_of ct of Const (@{const_name Case_Labeling.BIND}, _) $ _ $ _ $ _ => label_fun_conv ctxt ct | Const (@{const_name Case_Labeling.HIER}, _) $ _ $ _ => label_fun_conv ctxt ct | Const (@{const_name Case_Labeling.VC}, _) $ _ $ _ => label_fun_conv ctxt ct | _ => Conv.no_conv ct) fun norm_labels_conv ctxt ct = Conv.bottom_conv (Conv.try_conv o label_conv) ctxt ct in CONVERSION (norm_labels_conv ctxt) end fun extract_cases_tac (ctxt, st) = let val tac = unfold_tac ctxt @{thms LABEL_simps} in CONTEXT_CASES (mk_cases' ctxt st) tac (ctxt, st) end (* Splits the nth \ into its components and simplifies any case_prod over this variable*) structure Splitsubst = Hypsubst ( val dest_Trueprop = HOLogic.dest_Trueprop val dest_eq = dest_SPLIT val dest_imp = HOLogic.dest_imp val eq_reflection = @{thm SPLIT_reflection} val rev_eq_reflection = @{thm rev_SPLIT_reflection} val imp_intr = @{thm impI} val rev_mp = @{thm rev_mp} val subst = @{thm SPLIT_subst} val sym = @{thm SPLIT_sym} val thin_refl = @{thm SPLIT_thin_refl} ) val hyp_subst_tac = Splitsubst.hyp_subst_tac fun SPLIT_subst_tac ctxt = REPEAT_ALL_NEW (REPEAT_ALL_NEW (ematch_tac ctxt @{thms SPLIT_prodE}) THEN' Splitsubst.hyp_subst_tac ctxt) local val case_prod_th = @{thm split_conv[THEN eq_reflection]} fun case_prod_conv ctxt = Conv.forall_conv (fn (x, ctxt) => Conv.forall_conv (fn (y, ctxt) => fn ct => let val insts = [NONE, SOME x, SOME y] val typ_insts = map (Option.map Thm.ctyp_of_cterm) insts val th = Thm.instantiate' typ_insts insts case_prod_th in Conv.bottom_conv (fn _ => Conv.try_conv (Conv.rewr_conv th)) ctxt ct end ) ctxt) ctxt in fun split_nth_all_conv {right_only:bool} = let val desc_conv = if right_only then Conv.try_conv else Conv.repeat_conv fun conv ctxt 0 ct = desc_conv ( Conv.rewr_conv @{thm split_paired_all} then_conv case_prod_conv ctxt then_conv Conv.try_conv (conv ctxt 1) ) ct | conv ctxt n ct = Conv.forall_conv (fn (_, ctxt) => conv ctxt (n - 1)) ctxt ct in conv end fun split_bind_all_conv right_only ctxt ct = let val {vars, label, prems} = parse_label (Thm.term_of ct) val nvars = length vars val bv_ts = label |> map (fn x => case x of (_, Block (_, vs)) => vs) |> flat val sv_ts = prems |> map_filter (try dest_SPLIT) |> map snd |> maps Util.dest_tuple val lvars = bv_ts @ sv_ts |> maps dest_Bound |> sort_distinct int_ord |> map (fn n => nvars - n - 1) in Conv.every_conv (map (split_nth_all_conv right_only ctxt) lvars) ct end fun RAWCONV cv i st = Seq.single (Conv.gconv_rule cv i st) fun split_bind_all_tac {right_only: bool, simp_all_cases: bool} ctxt = let val rewr_tac = if simp_all_cases then Raw_Simplifier.rewrite_goal_tac ctxt [case_prod_th] else K all_tac in RAWCONV (split_bind_all_conv {right_only = right_only} ctxt) THEN' rewr_tac end end datatype options = Options of { simp_all_cases: bool, split_right_only: bool, protect_subgoals: bool} fun casify_tac (Options { simp_all_cases, split_right_only, protect_subgoals }) (ctxt, st) = let - fun inst_disambig ct = Thm.instantiate ([], [((("n", 0), @{typ nat}), ct)]) @{thm DISAMBIG_I} + fun inst_disambig ct = + Thm.instantiate (TVars.empty, Vars.make [((("n", 0), @{typ nat}), ct)]) @{thm DISAMBIG_I} fun disambig_tac i = if protect_subgoals then match_tac ctxt [inst_disambig (Thm.cterm_of ctxt (HOLogic.mk_number @{typ nat} i))] else K all_tac val prep_tac = DETERM (ALLGOALS (fn i => EVERY' [ prepare_labels_tac ctxt, split_bind_all_tac { simp_all_cases = simp_all_cases, right_only = split_right_only } ctxt, TRY o SPLIT_subst_tac ctxt, disambig_tac i ] i )) in (ctxt, st) |> (prep_tac THEN_CONTEXT extract_cases_tac) end local fun set_simp_all_cases simp_all_cases (Options { simp_all_cases = _, split_right_only, protect_subgoals}) = Options { simp_all_cases = simp_all_cases, split_right_only = split_right_only, protect_subgoals = protect_subgoals } fun set_protect_subgoals protect_subgoals (Options { simp_all_cases, split_right_only, protect_subgoals = _}) = Options { simp_all_cases = simp_all_cases, split_right_only = split_right_only, protect_subgoals = protect_subgoals } fun set_split_right_only split_right_only (Options { simp_all_cases, split_right_only = _, protect_subgoals}) = Options { simp_all_cases = simp_all_cases, split_right_only = split_right_only, protect_subgoals = protect_subgoals } val options = map (fn (s,f) => Args.parens (Args.$$$ s) >> K f) [ ("simp", set_simp_all_cases true), ("no_simp", set_simp_all_cases false), ("disambig_subgoals", set_protect_subgoals true), ("no_disambig_subgoals", set_protect_subgoals false), ("split_right_only", set_split_right_only true), ("no_split_right_only", set_split_right_only false) ] fun scan_alt scans = fold (fn scan1 => fn scan2 => scan1 || scan2) scans Scan.fail in fun casify_options def = Scan.repeat (scan_alt options) >> (fn xs => fold (fn f => fn x => f x) xs def) end fun casify_method_setup def = Scan.lift (casify_options def) >> (fn opt => fn _ => Util.SIMPLE_METHOD_CASES (casify_tac opt)) end diff --git a/thys/Collections/ICF/tools/ICF_Tools.thy b/thys/Collections/ICF/tools/ICF_Tools.thy --- a/thys/Collections/ICF/tools/ICF_Tools.thy +++ b/thys/Collections/ICF/tools/ICF_Tools.thy @@ -1,312 +1,311 @@ section \General ML-level tools\ theory ICF_Tools imports Main begin lemma meta_same_imp_rule: "(\PROP P; PROP P\ \ PROP Q) \ (PROP P \ PROP Q)" by rule (* TODO: Replace by distinct_prems_rl *) ML \ infix 0 ##; fun (f ## g) (a,b) = (f a, g b) signature ICF_TOOLS = sig (* Generic *) val gen_variant: (string -> bool) -> string -> string val map_option: ('a -> 'b) -> 'a option -> 'b option val parse_cpat: cterm context_parser - val rename_cterm: (cterm * cterm) -> - ((indexname * sort) * ctyp) list * ((indexname * typ) * cterm) list + val rename_cterm: (cterm * cterm) -> ctyp TVars.table * cterm Vars.table val renames_cterm: (cterm * cterm) -> bool val import_cterm: cterm -> Proof.context -> cterm * Proof.context val inst_meta_cong: Proof.context -> cterm -> thm (* val thms_from_main: string -> thm list val thm_from_main: string -> thm *) val sss_add: thm list -> Proof.context -> Proof.context val changed_conv: conv -> conv val repeat_top_sweep_conv: (Proof.context -> conv) -> Proof.context -> conv val rem_dup_prems: Proof.context -> thm -> thm (* Definition Theorems *) val dest_def_eq: term -> term * term val norm_def_thm: thm -> thm val dthm_lhs: thm -> term val dthm_rhs: thm -> term val dthm_params: thm -> term list val dthm_head: thm -> term val dt_lhs: term -> term val dt_rhs: term -> term val dt_params: term -> term list val dt_head: term -> term val chead_of: cterm -> cterm val chead_of_thm: thm -> cterm (* Simple definition: name\term, fixes variables *) val define_simple: string -> term -> local_theory -> ((term * thm) * local_theory) (* Wrapping stuff inside local theory context *) val wrap_lthy_result_global: (local_theory -> 'a * local_theory) -> (morphism -> 'a -> 'b) -> theory -> 'b * theory val wrap_lthy_global: (local_theory -> local_theory) -> theory -> theory val wrap_lthy_result_local: (local_theory -> 'a * local_theory) -> (morphism -> 'a -> 'b) -> local_theory -> 'b * local_theory val wrap_lthy_local: (local_theory -> local_theory) -> local_theory -> local_theory (* Wrapped versions of simple definition *) val define_simple_global: string -> term -> theory -> ((term * thm) * theory) val define_simple_local: string -> term -> local_theory -> ((term * thm) * local_theory) (* Revert abbreviations matching pattern (TODO/FIXME: HACK) *) val revert_abbrevs: string -> theory -> theory end; structure ICF_Tools: ICF_TOOLS = struct fun gen_variant decl s = let fun search s = if not (decl s) then s else search (Symbol.bump_string s); in if not (decl s) then s else search (Symbol.bump_init s) end; val parse_cpat = Args.context -- Scan.lift Args.embedded_inner_syntax >> (fn (ctxt, str) => Proof_Context.read_term_pattern ctxt str |> Thm.cterm_of ctxt ); (* Renaming first-order match *) fun rename_cterm (ct1,ct2) = ( Thm.first_order_match (ct2,ct1); Thm.first_order_match (ct1,ct2)); val renames_cterm = can rename_cterm; fun import_cterm ct ctxt = let val (t', ctxt') = yield_singleton (Variable.import_terms true) (Thm.term_of ct) ctxt; val ct' = Thm.cterm_of ctxt' t'; in (ct', ctxt') end (* Get theorem by name, that is visible in HOL.Main. Moreover, the theory of this theorem will be HOL.Main, which is required to avoid non-trivial theory merges as they may occur when using thm-antiquotation. (cf. post https://lists.cam.ac.uk/pipermail/cl-isabelle-users/2012-August/msg00175.html on Isabelle mailing list) *)(* fun thms_from_main name = let val xthmref = Facts.named name; val thy = @{theory Main}; val name = Facts.ref_name xthmref |> Global_Theory.intern_fact thy; val name = case name of "_" => "Pure.asm_rl" | name => name; val fs = Global_Theory.facts_of thy; val thms = Facts.lookup (Context.Theory thy) fs name |> the |> #2 |> map (Thm.transfer thy); in thms end fun thm_from_main name = thms_from_main name |> Facts.the_single (name, Position.none) *) (* Unfold with simpset fun unfold_ss ss = let val simple_prover = SINGLE o (fn ss => ALLGOALS (resolve_tac (Raw_Simplifier.prems_of ss))); in Raw_Simplifier.rewrite_thm (true,false,false) simple_prover ss end; *) local fun sss_add_single thm ss = let val simps = Raw_Simplifier.dest_ss (simpset_of ss) |> #simps |> map #2; val ess = ss delsimps simps; val thm' = simplify ss thm; val new_simps = simps |> map (simplify (ess addsimps [thm'])); val ss' = ess addsimps (thm'::new_simps) in ss' end in val sss_add = fold sss_add_single end local open Conv; in fun changed_conv cnv = (fn (ct:cterm) => let val thm = cnv ct in if Thm.is_reflexive thm then raise THM ("changed_conv: Not changed",~1,[thm]) else thm end) fun repeat_top_sweep_conv cnv ctxt = repeat_conv (changed_conv (top_sweep_conv cnv ctxt)); end (* Remove duplicate premises (stable) *) fun rem_dup_prems ctxt thm = let val prems = Thm.prems_of thm; val perm = prems |> tag_list 0 |> map swap |> Termtab.make_list |> Termtab.dest |> map snd |> sort (int_ord o apply2 hd) |> flat; val thm' = Drule.rearrange_prems perm thm |> Conv.fconv_rule (Raw_Simplifier.rewrite ctxt true @{thms meta_same_imp_rule}); in thm' end; fun dest_def_eq (Const (@{const_name Pure.eq},_)$l$r) = (l,r) | dest_def_eq (Const (@{const_name HOL.Trueprop},_) $(Const (@{const_name HOL.eq},_)$l$r)) = (l,r) | dest_def_eq t = raise TERM ("No definitional equation",[t]); fun norm_def_thm thm = case Thm.concl_of thm of (Const (@{const_name Pure.eq},_)$_$_) => thm | _ => thm RS eq_reflection; val dt_lhs = dest_def_eq #> fst; val dt_rhs = dest_def_eq #> snd; val dt_params = dt_lhs #> strip_comb #> snd; val dt_head = dt_lhs #> head_of; val dthm_lhs = Thm.concl_of #> dt_lhs; val dthm_rhs = Thm.concl_of #> dt_rhs; val dthm_params = Thm.concl_of #> dt_params; val dthm_head = Thm.concl_of #> dt_head; (* Head of function application (cterm) *) fun chead_of ct = case Thm.term_of ct of (_$_) => chead_of (Thm.dest_fun ct) | _ => ct; val chead_of_thm = norm_def_thm #> Thm.lhs_of #> chead_of; val meta_cong_rl = @{thm "eq_reflection"} OF @{thms "arg_cong"} OF @{thms "meta_eq_to_obj_eq"} fun inst_meta_cong ctxt ct = let val (ct, ctxt') = import_cterm ct ctxt; val mc_thm = meta_cong_rl; val fpat = mc_thm |> Thm.cprop_of |> Drule.strip_imp_concl |> Thm.dest_arg1 |> chead_of; val inst = infer_instantiate ctxt [(#1 (dest_Var (Thm.term_of fpat)), ct)] mc_thm; val inst' = singleton (Variable.export ctxt' ctxt) inst; in inst' end (* Define name\rhs, generate _def theorem. *) fun define_simple name rhs lthy = let (* Import type variables *) val (rhs,lthy) = yield_singleton Variable.importT_terms rhs lthy; val ((ft,(_,def_thm)),lthy) = Local_Theory.define ((Binding.name name,NoSyn), ((Binding.name (Thm.def_name name),[]),rhs)) lthy; in ((ft,def_thm),lthy) end; fun wrap_lthy_result_global f rmap thy = let val lthy = Named_Target.theory_init thy; val (r,lthy) = f lthy; val (r,thy) = Local_Theory.exit_result_global rmap (r,lthy); in (r,thy) end fun wrap_lthy_global f = wrap_lthy_result_global (pair () o f) (K I) #> #2; fun wrap_lthy_result_local f rmap lthy = let val lthy = (snd o Local_Theory.begin_nested) lthy; val (r,lthy) = f lthy; val m = Local_Theory.target_morphism lthy; val lthy = Local_Theory.end_nested lthy; val r = rmap m r; in (r,lthy) end fun wrap_lthy_local f = wrap_lthy_result_local (pair () o f) (K I) #> #2; (* Define name\rhs, yielding constant *) fun define_simple_global name rhs thy = let val lthy = Named_Target.theory_init thy; val (r,lthy) = define_simple name rhs lthy; fun map_res m (t,thm) = (Morphism.term m t,Morphism.thm m thm); val (r,thy) = Local_Theory.exit_result_global (map_res) (r,lthy); in (r,thy) end; (* Define name\rhs, yielding constant *) fun define_simple_local name rhs lthy = let val lthy = (snd o Local_Theory.begin_nested) lthy; val (r,lthy) = define_simple name rhs lthy; val m = Local_Theory.target_morphism lthy; val lthy = Local_Theory.end_nested lthy; fun map_res m (t,thm) = (Morphism.term m t,Morphism.thm m thm); val (r,lthy) = (map_res m r,lthy); in (r,lthy) end; fun map_option _ NONE = NONE | map_option f (SOME a) = SOME (f a); fun revert_abbrevs mpat thy = let val ctxt = Proof_Context.init_global thy; val match_prefix = if Long_Name.is_qualified mpat then mpat else Long_Name.qualify (Context.theory_name thy) mpat; val {const_space, constants, ...} = Sign.consts_of thy |> Consts.dest; val names = Name_Space.extern_entries true ctxt const_space constants |> map_filter (fn ((name, _), (_, SOME _)) => if Long_Name.qualifier name = match_prefix then SOME name else NONE | _ => NONE) val _ = if null names then warning ("ICF_Tools.revert_abbrevs: No names with prefix: " ^ match_prefix) else (); in fold (Sign.revert_abbrev "") names thy end end; \ attribute_setup rem_dup_prems = \ Scan.succeed (Thm.rule_attribute [] (ICF_Tools.rem_dup_prems o Context.proof_of)) \ "Remove duplicate premises" method_setup dup_subgoals = \ Scan.succeed (fn ctxt => SIMPLE_METHOD (PRIMITIVE (ICF_Tools.rem_dup_prems ctxt))) \ "Remove duplicate subgoals" end diff --git a/thys/Collections/ICF/tools/Locale_Code.thy b/thys/Collections/ICF/tools/Locale_Code.thy --- a/thys/Collections/ICF/tools/Locale_Code.thy +++ b/thys/Collections/ICF/tools/Locale_Code.thy @@ -1,386 +1,386 @@ section \Code Generation from Locales\ theory Locale_Code imports ICF_Tools Ord_Code_Preproc begin text \ Provides a simple mechanism to prepare code equations for constants stemming from locale interpretations. The usage pattern is as follows: \setup Locale_Code.checkpoint\ is called before a series of interpretations, and afterwards, \setup Locale_Code.prepare\ is called. Afterwards, the code generator will correctly recognize expressions involving terms from the locale interpretation. \ text \Tag to indicate pattern deletion\ definition LC_DEL :: "'a \ unit" where "LC_DEL a \ ()" ML \ signature LOCALE_CODE = sig type pat_eq = cterm * thm list val open_block: theory -> theory val close_block: theory -> theory val del_pat: cterm -> theory -> theory val add_pat_eq: cterm -> thm list -> theory -> theory val lc_decl_eq: thm list -> local_theory -> local_theory val lc_decl_del: term -> local_theory -> local_theory val setup: theory -> theory val get_unf_ss: theory -> simpset val tracing_enabled: bool Unsynchronized.ref end structure Locale_Code :LOCALE_CODE = struct open ICF_Tools val tracing_enabled = Unsynchronized.ref false; type pat_eq = cterm * thm list type block_data = {idx:int, del_pats: cterm list, add_pateqs: pat_eq list} val closed_block = {idx = ~1, del_pats=[], add_pateqs=[]}; fun init_block idx = {idx = idx, del_pats=[], add_pateqs=[]}; fun is_open ({idx,...}:block_data) = idx <> ~1; fun assert_open bd = if is_open bd then () else error "Locale_Code: No open block"; fun assert_closed bd = if is_open bd then error "Locale_Code: Block already open" else (); fun merge_bd (bd1,bd2) = ( if is_open bd1 orelse is_open bd2 then error "Locale_Code: Merge with open block" else (); closed_block ); fun bd_add_del_pats ps {idx,del_pats,add_pateqs} = {idx = idx, del_pats = ps@del_pats, add_pateqs = add_pateqs}; fun bd_add_add_pateqs pes {idx,del_pats,add_pateqs} = {idx = idx, del_pats = del_pats, add_pateqs = pes@add_pateqs}; structure BlockData = Theory_Data ( type T = block_data val empty = (closed_block) val extend = I val merge = merge_bd ); structure FoldSSData = Oc_Simpset ( val prio = 5; val name = "Locale_Code"; ); fun add_unf_thms thms thy = let val ctxt = Proof_Context.init_global thy val thms = map Thm.symmetric thms in FoldSSData.map (fn ss => put_simpset ss ctxt |> sss_add thms |> simpset_of ) thy end val get_unf_ss = FoldSSData.get; (* First order match with fixed head *) fun match_fixed_head (pat,obj) = let (* Match heads *) val inst = Thm.first_order_match (chead_of pat, chead_of obj); val pat = Thm.instantiate_cterm inst pat; (* Now match whole pattern *) val inst = Thm.first_order_match (pat, obj); in inst end; val matches_fixed_head = can match_fixed_head; (* First order match of heads only *) fun match_heads (pat,obj) = Thm.first_order_match (chead_of pat, chead_of obj); val matches_heads = can match_heads; val pat_nargs = Thm.term_of #> strip_comb #> #2 #> length; (* Adjust a theorem to exactly match pattern *) fun norm_thm_pat (thm,pat) = let val thm = norm_def_thm thm; val na_pat = pat_nargs pat; val lhs = Thm.lhs_of thm; val na_lhs = pat_nargs lhs; val lhs' = if na_lhs > na_pat then funpow (na_lhs - na_pat) Thm.dest_fun lhs else lhs; val inst = Thm.first_order_match (lhs',pat); in Thm.instantiate inst thm end; fun del_pat_matches cpat (epat,_) = if pat_nargs cpat = 0 then matches_heads (cpat,epat) else matches_fixed_head (cpat,epat); (* Pattern-Eqs from specification *) local datatype action = ADD of (cterm * thm list) | DEL of cterm fun filter_pat_eq thy thms pat = let val cpat = Thm.global_cterm_of thy pat; in if (pat_nargs cpat = 0) then NONE else let val thms' = fold (fn thm => fn acc => case try norm_thm_pat (thm, cpat) of NONE => acc | SOME thm => thm::acc ) thms []; in case thms' of [] => NONE | _ => SOME (ADD (cpat,thms')) end end; fun process_actions acc [] = acc | process_actions acc (ADD peq::acts) = process_actions (peq::acc) acts | process_actions acc (DEL cpat::acts) = let val acc' = filter (not o curry renames_cterm cpat o fst) acc; val _ = if length acc = length acc' then warning ("Locale_Code: LC_DEL without effect: " ^ @{make_string} cpat) else (); in process_actions acc' acts end; fun pat_eqs_of_spec thy {rough_classification = Spec_Rules.Equational _, terms = pats, rules = thms, ...} = map_filter (filter_pat_eq thy thms) pats | pat_eqs_of_spec thy {rough_classification = Spec_Rules.Unknown, terms = [Const (@{const_name LC_DEL},_)$pat], ...} = [(DEL (Thm.global_cterm_of thy pat))] | pat_eqs_of_spec _ _ = []; in fun pat_eqs_of_specs thy specs = map (pat_eqs_of_spec thy) specs |> flat |> rev |> process_actions []; end; fun is_proper_pat cpat = let val pat = Thm.term_of cpat; val (f,args) = strip_comb pat; in is_Const f andalso args <> [] andalso not (is_Var (hd (rev args))) end; (* Instantiating pattern-eq *) local (* Get constant name for instantiation pattern *) fun inst_name lthy pat = let val (fname,params) = case strip_comb pat of ((Const (fname,_)),params) => (fname,params) | _ => raise TERM ("inst_name: Expected pattern",[pat]); fun pname (Const (n,_)) = Long_Name.base_name n | pname (s$t) = pname s ^ "_" ^ pname t | pname _ = Name.uu; in space_implode "_" (Long_Name.base_name fname::map pname params) |> gen_variant (can (Proof_Context.read_const {proper = true, strict = false} lthy)) end; in fun inst_pat_eq (cpat,thms) = wrap_lthy_result_global (fn lthy => let val (((instT,inst),thms),lthy) = Variable.import true thms lthy; - val cpat = Thm.instantiate_cterm (TVars.dest instT, Vars.dest inst) cpat; + val cpat = Thm.instantiate_cterm (instT, inst) cpat; val pat = Thm.term_of cpat; val name = inst_name lthy pat; val ((_,(_,def_thm)),lthy) = Local_Theory.define ((Binding.name name,NoSyn), ((Binding.name (Thm.def_name name),[]),pat)) lthy; val thms' = map (Local_Defs.fold lthy [def_thm]) thms; in ((def_thm,thms'),lthy) end) (fn m => fn (def_thm,thms') => (Morphism.thm m def_thm, map (Morphism.thm m) thms')) #> (fn ((def_thm,thms'),thy) => let val thy = thy |> add_unf_thms [def_thm] |> Code.declare_default_eqns_global (map (rpair true) thms'); in thy end) end (* Bookkeeping *) fun new_specs thy = let val bd = BlockData.get thy; val _ = assert_open bd; val ctxt = Proof_Context.init_global thy; val srules = Spec_Rules.get ctxt; val res = take (length srules - #idx bd) srules; in res end fun open_block thy = let val bd = BlockData.get thy; val _ = assert_closed bd; val ctxt = Proof_Context.init_global thy; val idx = length (Spec_Rules.get ctxt); val thy = BlockData.map (K (init_block idx)) thy; in thy end; fun process_block bd thy = let fun filter_del_pats cpat peqs = let val peqs' = filter (not o del_pat_matches cpat) peqs val _ = if length peqs = length peqs' then warning ("Locale_Code: No pattern-eqs matching filter: " ^ @{make_string} cpat) else (); in peqs' end; fun filter_add_pats (orig_pat,_) = forall (fn (add_pat,_) => not (renames_cterm (orig_pat,add_pat))) (#add_pateqs bd); val specs = new_specs thy; val peqs = pat_eqs_of_specs thy specs |> fold filter_del_pats (#del_pats bd) |> filter filter_add_pats; val peqs = peqs @ #add_pateqs bd; val peqs = rev peqs; (* Important: Process equations in the order in that they have been added! *) val _ = if !tracing_enabled then map (fn peq => (tracing (@{make_string} peq); ())) peqs else []; val thy = thy |> fold inst_pat_eq peqs; in thy end; fun close_block thy = let val bd = BlockData.get thy; val _ = assert_open bd; val thy = process_block bd thy |> BlockData.map (K closed_block); in thy end; fun del_pat cpat thy = let val bd = BlockData.get thy; val _ = assert_open bd; val bd = bd_add_del_pats [cpat] bd; val thy = BlockData.map (K bd) thy; in thy end; fun add_pat_eq cpat thms thy = let val _ = is_proper_pat cpat orelse raise CTERM ("add_pat_eq: Not a proper pattern",[cpat]); fun ntp thm = case try norm_thm_pat (thm,cpat) of NONE => raise THM ("add_pat_eq: Theorem does not match pattern",~1,[thm]) | SOME thm => thm; val thms = map ntp thms; val thy = BlockData.map (bd_add_add_pateqs [(cpat,thms)]) thy; in thy end; local fun cpat_of_thm thm = let fun strip ct = case Thm.term_of ct of (_$Var _) => strip (Thm.dest_fun ct) | _ => ct; in strip (Thm.lhs_of thm) end; fun adjust_length (cpat1,cpat2) = let val n1 = cpat1 |> Thm.term_of |> strip_comb |> #2 |> length; val n2 = cpat2 |> Thm.term_of |> strip_comb |> #2 |> length; in if n1>n2 then (funpow (n1-n2) Thm.dest_fun cpat1, cpat2) else (cpat1, funpow (n2-n1) Thm.dest_fun cpat2) end fun find_match cpat cpat' = SOME (cpat,rename_cterm (cpat',cpat)) handle Pattern.MATCH => (case Thm.term_of cpat' of _$_ => find_match (Thm.dest_fun cpat) (Thm.dest_fun cpat') | _ => NONE ); (* Common head of definitional theorems *) fun comp_head thms = case map norm_def_thm thms of [] => NONE | thm::thms => let fun ch [] r = SOME r | ch (thm::thms) (cpat,acc) = let val cpat' = cpat_of_thm thm; val (cpat,cpat') = adjust_length (cpat,cpat') in case find_match cpat cpat' of NONE => NONE | SOME (cpat,inst) => ch thms (cpat, Drule.instantiate_normalize inst thm :: acc) end; in ch thms (cpat_of_thm thm,[thm]) end; in fun lc_decl_eq thms lthy = case comp_head thms of SOME (cpat,thms) => let val _ = if !tracing_enabled then tracing ("decl_eq: " ^ @{make_string} cpat ^ ": " ^ @{make_string} thms) else (); fun decl m = let val cpat'::thms' = Morphism.fact m (Drule.mk_term cpat :: thms); val cpat' = Drule.dest_term cpat'; in Context.mapping (BlockData.map (bd_add_add_pateqs [(cpat',thms')])) I end in lthy |> Local_Theory.declaration {syntax = false, pervasive = false} decl end | NONE => raise THM ("Locale_Code.lc_decl_eq: No common pattern",~1,thms); end; fun lc_decl_del pat = let val ty = fastype_of pat; val dpat = Const (@{const_name LC_DEL},ty --> @{typ unit})$pat; in Spec_Rules.add Binding.empty Spec_Rules.Unknown [dpat] [] end (* Package setup *) val setup = FoldSSData.setup; end \ setup Locale_Code.setup attribute_setup lc_delete = \ Parse.and_list1' ICF_Tools.parse_cpat >> (fn cpats => Thm.declaration_attribute (K (Context.mapping (fold Locale_Code.del_pat cpats) I))) \ "Locale_Code: Delete patterns for current block" attribute_setup lc_add = \ Parse.and_list1' (ICF_Tools.parse_cpat -- Attrib.thms) >> (fn peqs => Thm.declaration_attribute (K (Context.mapping (fold (uncurry Locale_Code.add_pat_eq) peqs) I))) \ "Locale_Code: Add pattern-eqs for current block" end diff --git a/thys/Combinatorics_Words/Reverse_Symmetry.thy b/thys/Combinatorics_Words/Reverse_Symmetry.thy --- a/thys/Combinatorics_Words/Reverse_Symmetry.thy +++ b/thys/Combinatorics_Words/Reverse_Symmetry.thy @@ -1,424 +1,424 @@ (* Title: CoW/Reverse_Symmetry.thy Author: Martin Raška, Charles University *) theory Reverse_Symmetry imports Main begin chapter "Reverse symmetry" text \This theory deals with a mechanism which produces new facts on lists from already known facts by the reverse symmetry of lists, induced by the mapping @{term rev}. It constructs the rule attribute ``reversed'' which produces the symmetrical fact using so-called reversal rules, which are rewriting rules that may be applied to obtain the symmetrical fact. An example of such a reversal rule is the already existing @{thm rev_append[symmetric, no_vars]}. Some additional reversal rules are given in this theory. The symmetrical fact 'A[reversed]' is constructed from the fact 'A' in the following manner: 1. each schematic variable @{term "xs::'a list"} of type @{typ "'a list"} is instantiated by @{term "rev xs"}; 2. constant Nil is substituted by @{term "rev Nil"}; 3. each quantification of @{typ "'a list"} type variable @{term "\(xs::'a list). P xs"} is substituted by (logically equivalent) quantification @{term "\xs. P (rev xs)"}, similarly for $\forall$, $\exists$ and $\exists!$ quantifiers; each bounded quantification of @{typ "'a list"} type variable @{term "\(xs::'a list) \ A. P xs"} is substituted by (logically equivalent) quantification @{term "\xs\rev ` A. P (rev xs)"}, similarly for bounded $\exists$ quantifier; 4. simultaneous rewrites according to a the current list of reversal rules are performed; 5. final correctional rewrites related to reversion of @{const "Cons"} are performed. List of reversal rules is maintained by declaration attribute ``reversal\_rule'' with standard ``add'' and ``del'' options. See examples at the end of the file. \ section \Quantifications and maps\ lemma all_surj_conv: assumes "surj f" shows "(\x. PROP P (f x)) \ (\y. PROP P y)" proof fix y assume "\x. PROP P (f x)" then have "PROP P (f (inv f y))". then show "PROP P y" unfolding surj_f_inv_f[OF assms]. qed lemma All_surj_conv: assumes "surj f" shows "(\x. P (f x)) \ (\y. P y)" proof (intro iffI allI) fix y assume "\x. P (f x)" then have "P (f (inv f y))".. then show "P y" unfolding surj_f_inv_f[OF assms]. qed simp lemma Ex_surj_conv: assumes "surj f" shows "(\x. P (f x)) \ (\y. P y)" proof assume "\x. P (f x)" then obtain x where "P (f x)".. then show "\x. P x".. next assume "\y. P y" then obtain y where "P y".. then have "P (f (inv f y))" unfolding surj_f_inv_f[OF assms]. then show "\x. P (f x)".. qed lemma Ex1_bij_conv: assumes "bij f" shows "(\!x. P (f x)) \ (\!y. P y)" proof have imp: "\!y. Q y" if bij: "bij g" and ex1: "\!x. Q (g x)" for g Q proof - from ex1E[OF ex1, rule_format] obtain x where ex: "Q (g x)" and uniq: "\x'. Q (g x') \ x' = x" by blast { fix y assume "Q y" then have "Q (g (inv g y))" unfolding surj_f_inv_f[OF bij_is_surj[OF bij]]. from uniq[OF this] have "x = inv g y".. then have "y = g x" unfolding bij_inv_eq_iff[OF bij].. } with ex show "\!y. Q y".. qed show "\!x. P (f x) \ \!y. P y" using imp[OF assms]. assume "\!y. P y" then have "\!y. P (f (inv f y))" unfolding surj_f_inv_f[OF bij_is_surj[OF assms]]. from imp[OF bij_imp_bij_inv[OF assms] this] show "\!x. P (f x)". qed lemma Ball_inj_conv: assumes "inj f" shows "(\y\f ` A. P (inv f y)) \ (\x\A. P x)" using ball_simps(9)[of f A "\y. P (inv f y)"] unfolding inv_f_f[OF assms]. lemma Bex_inj_conv: assumes "inj f" shows "(\y\f ` A. P (inv f y)) \ (\x\A. P x)" using bex_simps(7)[of f A "\y. P (inv f y)"] unfolding inv_f_f[OF assms]. subsection \Quantifications and reverse\ lemma rev_involution: "rev \ rev = id" by auto lemma rev_bij: "bij rev" using o_bij[OF rev_involution rev_involution]. lemma rev_inv: "inv rev = rev" using inv_unique_comp[OF rev_involution rev_involution]. lemmas all_rev_conv = all_surj_conv[OF bij_is_surj[OF rev_bij]] and All_rev_conv = All_surj_conv[OF bij_is_surj[OF rev_bij]] and Ex_rev_conv = Ex_surj_conv[OF bij_is_surj[OF rev_bij]] and Ex1_rev_conv = Ex1_bij_conv[OF rev_bij] and Ball_rev_conv = Ball_inj_conv[OF bij_is_inj[OF rev_bij], unfolded rev_inv] and Bex_rev_conv = Bex_inj_conv[OF bij_is_inj[OF rev_bij], unfolded rev_inv] section \Attributes\ context begin subsection \Definitons of reverse wrapers\ private definition rev_Nil_wrap :: "'a list" where "rev_Nil_wrap = rev Nil" private definition all_rev_wrap :: "('a list \ prop) \ prop" where "all_rev_wrap P \ (\x. PROP P (rev x))" private definition All_rev_wrap :: "('a list \ bool) \ bool" where "All_rev_wrap P = (\x. P (rev x))" private definition Ex_rev_wrap :: "('a list \ bool) \ bool" where "Ex_rev_wrap P = (\x. P (rev x))" private definition Ex1_rev_wrap :: "('a list \ bool) \ bool" where "Ex1_rev_wrap P = (\!x. P (rev x))" private definition Ball_rev_wrap :: "'a list set \ ('a list \ bool) \ bool" where "Ball_rev_wrap A P = (\x \ rev ` A. P (rev x))" private definition Bex_rev_wrap :: "'a list set \ ('a list \ bool) \ bool" where "Bex_rev_wrap A P = (\x \ rev ` A. P (rev x))" subsection \Initial reversal rules\ private lemma rev_Nil: "rev Nil = Nil" by simp private lemmas init_rev_wrap = eq_reflection[OF trans[OF rev_Nil[symmetric] rev_Nil_wrap_def[symmetric]]] transitive[OF all_rev_conv[symmetric] all_rev_wrap_def[symmetric], of P P for P] eq_reflection[OF trans[OF All_rev_conv[symmetric] All_rev_wrap_def[symmetric], of P P for P]] eq_reflection[OF trans[OF Ex_rev_conv[symmetric] Ex_rev_wrap_def[symmetric], of P P for P]] (* Ex_rev_wrapI *) eq_reflection[OF trans[OF Ex1_rev_conv[symmetric] Ex1_rev_wrap_def[symmetric], of P P for P]] eq_reflection[OF trans[OF Ball_rev_conv[symmetric] Ball_rev_wrap_def[symmetric], of P P for P]] eq_reflection[OF trans[OF Bex_rev_conv[symmetric] Bex_rev_wrap_def[symmetric], of P P for P]] private lemma all_rev_unwrap: "all_rev_wrap (\x. PROP P x) \ (\x. PROP P (rev x))" using all_rev_wrap_def. private lemma All_rev_unwrap: "All_rev_wrap (\x. P x) = (\x. P (rev x))" using All_rev_wrap_def. private lemma Ex_rev_unwrap: "Ex_rev_wrap (\x. P x) = (\x. P (rev x))" using Ex_rev_wrap_def. private lemma Ex1_rev_unwrap: "Ex1_rev_wrap (\x. P x) = (\!x. P (rev x))" using Ex1_rev_wrap_def. private lemma Ball_rev_unwrap: "Ball_rev_wrap A (\x. P x) = (\x \ rev ` A. P (rev x))" using Ball_rev_wrap_def. private lemma Bex_rev_unwrap: "Bex_rev_wrap A (\x. P x) = (\x \ rev ` A. P (rev x))" using Bex_rev_wrap_def. private lemmas init_rev_unwrap = eq_reflection[OF rev_Nil_wrap_def] all_rev_unwrap eq_reflection[OF All_rev_unwrap] eq_reflection[OF Ex_rev_unwrap] eq_reflection[OF Ex1_rev_unwrap] eq_reflection[OF Ball_rev_unwrap] eq_reflection[OF Bex_rev_unwrap] subsection \Cons reversion\ definition snocs :: "'a list \ 'a list \ 'a list" where "snocs xs ys = xs @ ys" lemma Cons_rev: "a # rev u = rev (snocs u [a])" unfolding snocs_def by simp lemma snocs_snocs: "snocs (snocs xs (y # ys)) zs = snocs xs (y # snocs ys zs)" unfolding snocs_def by simp subsection \Final corrections\ lemma snocs_Nil: "snocs [] xs = xs" unfolding snocs_def by simp lemma snocs_is_append: "snocs xs ys = xs @ ys" unfolding snocs_def.. private lemmas final_correct1 = snocs_Nil private lemmas final_correct2 = snocs_is_append subsection \Declaration attribute \reversal_rule\\ ML \ structure Reversal_Rules = Named_Thms( val name = @{binding "reversal_rule"} val description = "Rules performing reverse-symmetry transformation on theorems on lists" ) \ attribute_setup reversal_rule = \Attrib.add_del (Thm.declaration_attribute Reversal_Rules.add_thm) (Thm.declaration_attribute Reversal_Rules.del_thm)\ "maintaining a list of reversal rules" subsection \Rule attribute \reversed\\ ML \ val eq_refl = @{thm eq_reflection} fun pure_eq_of th = case Thm.prop_of th of Const (\<^const_name>\Pure.eq\, _) $ _ $ _ => SOME (th) | Const (\<^const_name>\Trueprop\, _) $ (Const (\<^const_name>\HOL.eq\, _) $ _ $ _ ) => SOME (th RS eq_refl) | _ => NONE val init_rev_wrap = @{thms init_rev_wrap} val init_unwrap = @{thms init_rev_unwrap} val final_correct1 = map_filter pure_eq_of @{thms final_correct1} val final_correct2 = map_filter pure_eq_of @{thms final_correct2} fun reverse ths context th = let val ctxt = Context.proof_of context val rules = map_filter pure_eq_of (ths @ Reversal_Rules.get ctxt) val vars = Term.add_vars (Thm.full_prop_of th) [] fun add_inst_vars [] inst_vars = inst_vars | add_inst_vars ( ((v, i), T) :: vars ) inst_vars = case T of Type(\<^type_name>\list\, _) => add_inst_vars vars ( ( ((v, i), T), Thm.cterm_of ctxt ( Const(\<^const_name>\rev\, Type(\<^type_name>\fun\, [T, T])) $ Var((v, i), T) ) ) :: inst_vars ) | _ => add_inst_vars vars inst_vars in th |> Drule.instantiate_normalize - ([], add_inst_vars vars []) + (TVars.empty, Vars.make (add_inst_vars vars [])) |> Simplifier.rewrite_rule ctxt init_rev_wrap |> Simplifier.rewrite_rule ctxt init_unwrap |> Simplifier.rewrite_rule ctxt rules |> Simplifier.rewrite_rule ctxt final_correct1 |> Simplifier.rewrite_rule ctxt final_correct2 end val reversed = Scan.optional (Scan.lift (Args.add -- Args.colon) |-- Attrib.thms) [] >> (fn ths => Thm.rule_attribute [] (reverse ths)) \ attribute_setup reversed = \reversed\ "Transforms the theorem on lists to reverse-symmetric version" end section \Declaration of basic reversal rules\ lemma hd_last_Nil: "hd [] = last []" unfolding hd_def last_def by simp lemma last_rev_hd: "last(rev xs) = hd xs" by (induct xs, simp add: hd_last_Nil, simp) lemma hd_rev_last: "hd(rev xs) = last xs" by (induct xs, simp add: hd_last_Nil, simp) lemma tl_rev: "tl (rev xs) = rev (butlast xs)" unfolding rev_swap[symmetric] butlast_rev[of "rev xs", symmetric] rev_rev_ident.. lemma if_rev: "(if P then rev u else rev v) = rev (if P then u else v)" by simp lemma rev_in_conv: "rev u \ A \ u \ rev ` A" using image_iff by fastforce lemma in_lists_rev: "u \ lists A \ rev u \ lists A" by auto lemma rev_in_lists: "rev u \ lists A \ u \ lists A" by auto lemma rev_lists_conv: "rev ` lists A = lists A" proof (intro equalityI subsetI) fix x show "x \ rev ` lists A \ x \ lists A" unfolding rev_in_conv[symmetric] using rev_in_lists. show "x \ lists A \ x \ rev ` lists A" unfolding rev_in_conv[symmetric] using in_lists_rev. qed lemma coset_rev: "List.coset (rev xs) = List.coset xs" by simp lemmas [reversal_rule] = Cons_rev snocs_snocs rev_append[symmetric] last_rev_hd hd_rev_last tl_rev butlast_rev rev_is_rev_conv length_rev take_rev drop_rev rotate_rev if_rev rev_in_conv rev_lists_conv set_rev coset_rev section \Examples\ context begin subsection \Cons and append\ private lemma example_Cons_append: assumes "xs = [a, b]" and "ys = [b, a, b]" shows "xs @ xs @ xs = a # b # a # ys" using assms by simp thm example_Cons_append example_Cons_append[reversed] example_Cons_append[reversed, reversed] thm not_Cons_self not_Cons_self[reversed] thm neq_Nil_conv neq_Nil_conv[reversed] subsection \Induction rules\ thm list_nonempty_induct list_nonempty_induct[reversed] (* needs work *) list_nonempty_induct[reversed, where P="\x. P (rev x)" for P, unfolded rev_rev_ident] thm list_induct2 list_induct2[reversed] (* needs work *) list_induct2[reversed, where P="\x y. P (rev x) (rev y)" for P, unfolded rev_rev_ident] subsection \hd, tl, last, butlast\ thm hd_append hd_append[reversed] last_append thm length_tl length_tl[reversed] length_butlast thm hd_Cons_tl hd_Cons_tl[reversed] append_butlast_last_id append_butlast_last_id[reversed] subsection \set\ thm hd_in_set hd_in_set[reversed] last_in_set thm set_ConsD set_ConsD[reversed] thm split_list_first split_list_first[reversed] thm split_list_first_prop split_list_first_prop[reversed] split_list_first_prop[reversed, unfolded append_assoc append_Cons append_Nil] split_list_last_prop thm split_list_first_propE split_list_first_propE[reversed] split_list_first_propE[reversed, unfolded append_assoc append_Cons append_Nil] split_list_last_propE subsection \rotate\ private lemma rotate1_hd_tl': "xs \ [] \ rotate 1 xs = tl xs @ [hd xs]" by (cases xs) simp_all thm rotate1_hd_tl' rotate1_hd_tl'[reversed] end end \ No newline at end of file diff --git a/thys/Deriving/Comparator_Generator/compare_code.ML b/thys/Deriving/Comparator_Generator/compare_code.ML --- a/thys/Deriving/Comparator_Generator/compare_code.ML +++ b/thys/Deriving/Comparator_Generator/compare_code.ML @@ -1,67 +1,67 @@ signature COMPARE_CODE = sig (* changes the code equations of some constant such that two consecutive comparisons via <=, <, or = are turned into one invocation of the comparator. The difference to a standard code_unfold is that here we change the code-equations where an additional sort-constraint on compare_order can be added. Otherwise, there would be no compare-function. *) val change_compare_code : term (* the constant *) -> string list (* the list of type parameters which should be constraint to @{sort compare_order} *) -> local_theory -> local_theory end structure Compare_Code : COMPARE_CODE = struct fun drop_leading_qmark s = if String.isPrefix "?" s then String.substring (s,1,String.size s - 1) else s fun change_compare_code const inst_names lthy = let val inst_names = map drop_leading_qmark inst_names val const_string = quote (Pretty.string_of (Syntax.pretty_term lthy const)) val cname = (case const of Const (cname,_) => cname | _ => error ("expected constant as input, but got " ^ const_string)) val cert = Code.get_cert lthy [] cname val code_eqs = cert |> Code.equations_of_cert (Proof_Context.theory_of lthy) |> snd |> these |> map (fst o snd) |> map_filter I val _ = if null code_eqs then error "could not find code equations" else () (* adding sort-constraint compare_order within code equations*) val const' = hd code_eqs |> Thm.concl_of |> Logic.dest_equals |> fst |> strip_comb |> fst val types = Term.add_tvars const' [] val ctyp_of = TVar #> Thm.ctyp_of lthy fun filt s = not (member (op =) (@{sort ord} @ @{sort linorder} @ @{sort order}) s) val map_types = maps (fn ty => if List.exists (fn tn => tn = (fst o fst) ty) inst_names then [(ty, ctyp_of (apsnd (fn ss => filter filt ss @ @{sort compare_order}) ty))] else []) types - val code_eqs = map (Thm.instantiate (map_types, [])) code_eqs + val code_eqs = map (Thm.instantiate (TVars.make map_types, Vars.empty)) code_eqs (* replace comparisons and register code eqns *) val new_code_eqs = map (Local_Defs.unfold lthy @{thms two_comparisons_into_compare}) code_eqs val _ = if map Thm.prop_of new_code_eqs = map Thm.prop_of code_eqs then warning ("Code equations for " ^ const_string ^ " did not change\n" ^ "Perhaps you have to provide some type variables which should be restricted to compare_order\n" ^ (@{make_string} (map TVar types ~~ map snd types, const', code_eqs))) else () val lthy = Local_Theory.note ( (Binding.name (Long_Name.base_name cname ^ "_compare_code"), @{attributes [code]}), new_code_eqs) lthy |> snd in lthy end fun change_compare_code_cmd const tnames_option lthy = change_compare_code (Syntax.read_term lthy const) tnames_option lthy val _ = Outer_Syntax.local_theory @{command_keyword compare_code} "turn comparisons via <= and < into compare within code-equations" (Scan.optional (@{keyword "("} |-- (Parse.list Parse.string) --| @{keyword ")"}) [] -- Parse.term >> (fn (inst, c) => change_compare_code_cmd c inst)) end diff --git a/thys/IMP2/lib/subgoal_focus_some.ML b/thys/IMP2/lib/subgoal_focus_some.ML --- a/thys/IMP2/lib/subgoal_focus_some.ML +++ b/thys/IMP2/lib/subgoal_focus_some.ML @@ -1,181 +1,181 @@ (* Generalized version of Subgoal.FOCUS, where the premises to be be focused on can be selected by a filter function. This generalizes the do_prems from bool to Proof.context -> cterm -> bool Author: Peter Lammich. Derived from subgoal_focus.ML. *) signature SUBGOAL_FOCUS_SOME = sig type focus = {context: Proof.context, params: (string * cterm) list, prems: thm list, asms: (bool * cterm) list, - concl: cterm, schematics: ((indexname * sort) * ctyp) list * ((indexname * typ) * cterm) list} + concl: cterm, schematics: ctyp TVars.table * cterm Vars.table} type prem_filter = Proof.context -> cterm -> bool val focus_params: Proof.context -> int -> binding list option -> thm -> focus * thm val focus_params_fixed: Proof.context -> int -> binding list option -> thm -> focus * thm val focus_prems: Proof.context -> int -> binding list option -> thm -> focus * thm val focus: Proof.context -> int -> binding list option -> thm -> focus * thm val focus_some_prems: prem_filter -> Proof.context -> int -> binding list option -> thm -> focus * thm val retrofit: Proof.context -> Proof.context -> (string * cterm) list -> (bool * cterm) list -> int -> thm -> thm -> thm Seq.seq val FOCUS_PARAMS: (focus -> tactic) -> Proof.context -> int -> tactic val FOCUS_PARAMS_FIXED: (focus -> tactic) -> Proof.context -> int -> tactic val FOCUS_PREMS: (focus -> tactic) -> Proof.context -> int -> tactic val FOCUS: (focus -> tactic) -> Proof.context -> int -> tactic val FOCUS_SOME_PREMS: prem_filter -> (focus -> tactic) -> Proof.context -> int -> tactic end structure Subgoal_Focus_Some : SUBGOAL_FOCUS_SOME = struct type focus = {context: Proof.context, params: (string * cterm) list, prems: thm list, asms: (bool * cterm) list, - concl: cterm, schematics: ((indexname * sort) * ctyp) list * ((indexname * typ) * cterm) list} + concl: cterm, schematics: ctyp TVars.table * cterm Vars.table} type prem_filter = Proof.context -> cterm -> bool fun partition P l = (filter P l, filter_out P l) fun invert_perm l = tag_list 0 l |> map swap |> order_list fun gen_focus (do_prems, do_concl) ctxt i bindings raw_st = let val st = raw_st |> Thm.transfer (Proof_Context.theory_of ctxt) |> Raw_Simplifier.norm_hhf_protect ctxt; val ((schematic_types, [st']), ctxt1) = Variable.importT [st] ctxt; val ((params, goal), ctxt2) = Variable.focus_cterm bindings (Thm.cprem_of st' i) ctxt1; val (asms, concl) = (Drule.strip_imp_prems goal, Drule.strip_imp_concl goal) val asms = map (`(do_prems ctxt2)) asms val fasms = filter fst asms |> map snd val nasms = filter_out fst asms |> map snd val concl = Drule.list_implies (nasms, concl) val text = fasms @ (if do_concl then [concl] else []); val ((_, inst), ctxt3) = Variable.import_inst true (map Thm.term_of text) ctxt2; val schematic_terms = Vars.fold (fn (v, t) => cons (v, Thm.cterm_of ctxt3 t)) inst []; - val schematics = (TVars.dest schematic_types, schematic_terms); + val schematics = (schematic_types, Vars.make schematic_terms); val asms' = map (apsnd (Thm.instantiate_cterm schematics)) asms; val fasms' = filter fst asms' |> map snd val concl' = Thm.instantiate_cterm schematics concl; val (prems, context) = Assumption.add_assumes fasms' ctxt3; in ({context = context, params = params, prems = prems, asms = asms', concl = concl', schematics = schematics}, Goal.init concl') end; val focus_params = gen_focus (K (K false), false); val focus_params_fixed = gen_focus (K (K false), true); val focus_prems = gen_focus (K (K true), false); val focus = gen_focus (K (K true), true); fun focus_some_prems flt = gen_focus (flt,false) (* B [?'b, ?y] ---------------- B ['b, y params] *) fun lift_import idx params th ctxt = let val ((_, [th']), ctxt') = Variable.importT [th] ctxt; val Ts = map Thm.typ_of_cterm params; val ts = map Thm.term_of params; val prop = Thm.full_prop_of th'; val concl_vars = Term.add_vars (Logic.strip_imp_concl prop) []; val vars = rev (Term.add_vars prop []); val (ys, ctxt'') = Variable.variant_fixes (map (Name.clean o #1 o #1) vars) ctxt'; fun var_inst v y = let val ((x, i), T) = v; val (U, args) = if member (op =) concl_vars v then (T, []) else (Ts ---> T, ts); val u = Free (y, U); in ((Var v, list_comb (u, args)), (u, Var ((x, i + idx), U))) end; val (inst1, inst2) = split_list (map (apply2 (apply2 (Thm.cterm_of ctxt))) (map2 var_inst vars ys)); - - val th'' = Thm.instantiate ([], map (apfst (Term.dest_Var o Thm.term_of)) inst1) th'; + val inst = Vars.build (fold (Vars.add o apfst (Term.dest_Var o Thm.term_of)) inst1); + val th'' = Thm.instantiate (TVars.empty, inst) th'; in ((inst2, th''), ctxt'') end; (* [x, A x] : B x ==> C ------------------ [!!x. A x ==> B x] : C *) fun lift_subgoals ctxt params asms th = let fun lift ct = fold_rev (Thm.all_name ctxt) params (Drule.list_implies (asms, ct)); val unlift = fold (Thm.elim_implies o Thm.assume) asms o Drule.forall_elim_list (map #2 params) o Thm.assume; val subgoals = map lift (Drule.strip_imp_prems (Thm.cprop_of th)); val th' = fold (Thm.elim_implies o unlift) subgoals th; in (subgoals, th') end; fun retrofit ctxt1 ctxt0 params all_asms i st1 st0 = let val asms = filter fst all_asms |> map snd val perm = tag_list 0 all_asms |> partition (fst o snd) |> op @ |> map fst val perm = invert_perm perm val idx = Thm.maxidx_of st0 + 1; val ps = map #2 params; val ((subgoal_inst, st2), ctxt2) = lift_import idx ps st1 ctxt1; val (subgoals, st3) = lift_subgoals ctxt2 params asms st2; val result = st3 |> Goal.conclude |> Drule.implies_intr_list asms |> Drule.rearrange_prems perm |> Drule.forall_intr_list ps |> Drule.implies_intr_list subgoals |> fold_rev (Thm.forall_intr o #1) subgoal_inst |> fold (Thm.forall_elim o #2) subgoal_inst |> Thm.adjust_maxidx_thm idx |> singleton (Variable.export ctxt2 ctxt0); in Thm.bicompose (SOME ctxt0) {flatten = true, match = false, incremented = false} (false, result, Thm.nprems_of st1) i st0 end; fun GEN_FOCUS flags tac ctxt i st = if Thm.nprems_of st < i then Seq.empty else let val (args as {context = ctxt', params, asms, ...}, st') = gen_focus flags (ctxt |> Variable.set_bound_focus true) i NONE st; in Seq.lifts (retrofit ctxt' ctxt params asms i) (tac args st') st end; val FOCUS_PARAMS = GEN_FOCUS (K (K false), false); val FOCUS_PARAMS_FIXED = GEN_FOCUS (K (K false), true); val FOCUS_PREMS = GEN_FOCUS (K (K true), false); val FOCUS = GEN_FOCUS (K (K true), true); fun FOCUS_SOME_PREMS flt = GEN_FOCUS (flt, true); end diff --git a/thys/Nominal2/nominal_dt_quot.ML b/thys/Nominal2/nominal_dt_quot.ML --- a/thys/Nominal2/nominal_dt_quot.ML +++ b/thys/Nominal2/nominal_dt_quot.ML @@ -1,751 +1,751 @@ (* Title: nominal_dt_alpha.ML Author: Christian Urban Author: Cezary Kaliszyk Performing quotient constructions, lifting theorems and deriving support properties for the quotient types. *) signature NOMINAL_DT_QUOT = sig val define_qtypes: (string list * binding * mixfix) list -> typ list -> term list -> thm list -> local_theory -> Quotient_Info.quotients list * local_theory val define_qconsts: typ list -> (string * term * mixfix * thm) list -> local_theory -> Quotient_Info.quotconsts list * local_theory val define_qperms: typ list -> string list -> (string * sort) list -> (string * term * mixfix * thm) list -> thm list -> local_theory -> local_theory val define_qsizes: typ list -> string list -> (string * sort) list -> (string * term * mixfix * thm) list -> local_theory -> local_theory val lift_thms: typ list -> thm list -> thm list -> Proof.context -> thm list * Proof.context val prove_supports: Proof.context -> thm list -> term list -> thm list val prove_fsupp: Proof.context -> typ list -> thm -> thm list -> thm list val fs_instance: typ list -> string list -> (string * sort) list -> thm list -> local_theory -> local_theory val prove_fv_supp: typ list -> term list -> term list -> term list -> term list -> thm list -> thm list -> thm list -> thm list -> thm -> bclause list list -> Proof.context -> thm list val prove_bns_finite: typ list -> term list -> thm -> thm list -> Proof.context -> thm list val prove_perm_bn_alpha_thms: typ list -> term list -> term list -> thm -> thm list -> thm list -> thm list -> Proof.context -> thm list val prove_permute_bn_thms: typ list -> term list -> term list -> thm -> thm list -> thm list -> thm list -> Proof.context -> thm list val prove_strong_exhausts: Proof.context -> thm list -> bclause list list list -> thm list -> thm list -> thm list -> thm list -> thm list -> thm list val prove_strong_induct: Proof.context -> thm -> thm list -> thm list -> bclause list list list -> thm list end structure Nominal_Dt_Quot: NOMINAL_DT_QUOT = struct open Nominal_Permeq fun lookup xs x = the (AList.lookup (op=) xs x) (* defines the quotient types *) fun define_qtypes qtys_descr alpha_tys alpha_trms alpha_equivp_thms lthy = let val qty_args1 = map2 (fn ty => fn trm => (ty, trm, false)) alpha_tys alpha_trms val qty_args2 = map2 (fn descr => fn args1 => (descr, args1, (NONE, NONE))) qtys_descr qty_args1 val qty_args3 = qty_args2 ~~ alpha_equivp_thms in fold_map (Quotient_Type.add_quotient_type {overloaded = false}) qty_args3 lthy end (* a wrapper for lifting a raw constant *) fun lift_raw_const qtys (qconst_name, rconst, mx', rsp_thm) lthy = let val rty = fastype_of rconst val qty = Quotient_Term.derive_qtyp lthy qtys rty val lhs_raw = Free (qconst_name, qty) val rhs_raw = rconst val raw_var = (Binding.name qconst_name, NONE, mx') val ((binding, _, mx), ctxt) = Proof_Context.cert_var raw_var lthy val lhs = Syntax.check_term ctxt lhs_raw val rhs = Syntax.check_term ctxt rhs_raw val (lhs_str, lhs_ty) = dest_Free lhs handle TERM _ => error "Constant already defined." val _ = if null (strip_abs_vars rhs) then () else error "The definiens cannot be an abstraction" val _ = if is_Const rhs then () else warning "The definiens is not a constant" in Quotient_Def.add_quotient_def (((binding, mx), Binding.empty_atts), (lhs, rhs)) rsp_thm ctxt end (* defines quotient constants *) fun define_qconsts qtys consts_specs lthy = let val (qconst_infos, lthy') = fold_map (lift_raw_const qtys) consts_specs lthy val phi = Proof_Context.export_morphism lthy' (Proof_Context.transfer (Proof_Context.theory_of lthy') lthy) in (map (Quotient_Info.transform_quotconsts phi) qconst_infos, lthy') end (* defines the quotient permutations and proves pt-class *) fun define_qperms qtys qfull_ty_names tvs perm_specs raw_perm_laws lthy = let val lthy1 = lthy |> Local_Theory.exit_global |> Class.instantiation (qfull_ty_names, tvs, @{sort pt}) val (_, lthy2) = define_qconsts qtys perm_specs lthy1 val ((_, raw_perm_laws'), lthy3) = Variable.importT raw_perm_laws lthy2 val lifted_perm_laws = map (Quotient_Tacs.lifted lthy3 qtys []) raw_perm_laws' |> Variable.exportT lthy3 lthy2 fun tac ctxt = Class.intro_classes_tac ctxt [] THEN (ALLGOALS (resolve_tac ctxt lifted_perm_laws)) in lthy2 |> Class.prove_instantiation_exit tac |> Named_Target.theory_init end (* defines the size functions and proves size-class *) fun define_qsizes qtys qfull_ty_names tvs size_specs lthy = let fun tac ctxt = Class.intro_classes_tac ctxt [] in lthy |> Local_Theory.exit_global |> Class.instantiation (qfull_ty_names, tvs, @{sort size}) |> snd o (define_qconsts qtys size_specs) |> Class.prove_instantiation_exit tac |> Named_Target.theory_init end (* lifts a theorem and cleans all "_raw" parts from variable names *) local val any = Scan.one (Symbol.not_eof) val raw = Scan.this_string "_raw" val exclude = Scan.repeat (Scan.unless raw any) --| raw >> implode val parser = Scan.repeat (exclude || any) in fun unraw_str s = s |> raw_explode |> Scan.finite Symbol.stopper parser >> implode |> fst end fun unraw_vars_thm ctxt thm = let fun unraw_var_str ((s, i), T) = ((unraw_str s, i), T) val vars = Term.add_vars (Thm.prop_of thm) [] val vars' = map (Thm.cterm_of ctxt o Var o unraw_var_str) vars in - Thm.instantiate ([], (vars ~~ vars')) thm + Thm.instantiate (TVars.empty, Vars.make (vars ~~ vars')) thm end fun unraw_bounds_thm th = let val trm = Thm.prop_of th val trm' = Term.map_abs_vars unraw_str trm in Thm.rename_boundvars trm trm' th end fun lift_thms qtys simps thms ctxt = (map (Quotient_Tacs.lifted ctxt qtys simps #> unraw_bounds_thm #> unraw_vars_thm ctxt #> Drule.zero_var_indexes) thms, ctxt) fun mk_supports_goal ctxt qtrm = let val vs = fresh_args ctxt qtrm val rhs = list_comb (qtrm, vs) val lhs = fold (curry HOLogic.mk_prod) vs @{term "()"} |> mk_supp in mk_supports lhs rhs |> HOLogic.mk_Trueprop end fun supports_tac ctxt perm_simps = let val simpset1 = put_simpset HOL_basic_ss ctxt addsimps @{thms supports_def fresh_def[symmetric]} val simpset2 = put_simpset HOL_ss ctxt addsimps @{thms swap_fresh_fresh fresh_Pair} in EVERY' [ simp_tac simpset1, eqvt_tac ctxt (eqvt_strict_config addpres perm_simps), simp_tac simpset2 ] end fun prove_supports_single ctxt perm_simps qtrm = let val goal = mk_supports_goal ctxt qtrm val ctxt' = Proof_Context.augment goal ctxt in Goal.prove ctxt' [] [] goal (K (HEADGOAL (supports_tac ctxt perm_simps))) |> singleton (Proof_Context.export ctxt' ctxt) end fun prove_supports ctxt perm_simps qtrms = map (prove_supports_single ctxt perm_simps) qtrms (* finite supp lemmas for qtypes *) fun prove_fsupp ctxt qtys qinduct qsupports_thms = let val (vs, ctxt') = Variable.variant_fixes (replicate (length qtys) "x") ctxt val goals = vs ~~ qtys |> map Free |> map (mk_finite o mk_supp) |> foldr1 (HOLogic.mk_conj) |> HOLogic.mk_Trueprop val tac = EVERY' [ resolve_tac ctxt' @{thms supports_finite}, resolve_tac ctxt' qsupports_thms, asm_simp_tac (put_simpset HOL_ss ctxt' addsimps @{thms finite_supp supp_Pair finite_Un}) ] in Goal.prove ctxt' [] [] goals (K (HEADGOAL (resolve_tac ctxt' [qinduct] THEN_ALL_NEW tac))) |> singleton (Proof_Context.export ctxt' ctxt) |> Old_Datatype_Aux.split_conj_thm |> map zero_var_indexes end (* finite supp instances *) fun fs_instance qtys qfull_ty_names tvs qfsupp_thms lthy = let val lthy1 = lthy |> Local_Theory.exit_global |> Class.instantiation (qfull_ty_names, tvs, @{sort fs}) fun tac ctxt = Class.intro_classes_tac ctxt [] THEN (ALLGOALS (resolve_tac ctxt qfsupp_thms)) in lthy1 |> Class.prove_instantiation_exit tac |> Named_Target.theory_init end (* proves that fv and fv_bn equals supp *) fun gen_mk_goals fv supp = let val arg_ty = fastype_of fv |> domain_type in (arg_ty, fn x => HOLogic.mk_eq (fv $ x, supp x)) end fun mk_fvs_goals fv = gen_mk_goals fv mk_supp fun mk_fv_bns_goals fv_bn alpha_bn = gen_mk_goals fv_bn (mk_supp_rel alpha_bn) fun add_simpset ctxt thms = put_simpset HOL_basic_ss ctxt addsimps thms fun symmetric thms = map (fn thm => thm RS @{thm sym}) thms val supp_Abs_set = @{thms supp_Abs(1)[symmetric]} val supp_Abs_res = @{thms supp_Abs(2)[symmetric]} val supp_Abs_lst = @{thms supp_Abs(3)[symmetric]} fun mk_supp_abs ctxt (BC (Set, _, _)) = EqSubst.eqsubst_tac ctxt [1] supp_Abs_set | mk_supp_abs ctxt (BC (Res, _, _)) = EqSubst.eqsubst_tac ctxt [1] supp_Abs_res | mk_supp_abs ctxt (BC (Lst, _, _)) = EqSubst.eqsubst_tac ctxt [1] supp_Abs_lst fun mk_supp_abs_tac ctxt [] = [] | mk_supp_abs_tac ctxt (BC (_, [], _)::xs) = mk_supp_abs_tac ctxt xs | mk_supp_abs_tac ctxt (bc::xs) = (DETERM o mk_supp_abs ctxt bc)::mk_supp_abs_tac ctxt xs fun mk_bn_supp_abs_tac ctxt trm = trm |> fastype_of |> body_type |> (fn ty => case ty of @{typ "atom set"} => simp_tac (add_simpset ctxt supp_Abs_set) | @{typ "atom list"} => simp_tac (add_simpset ctxt supp_Abs_lst) | _ => raise TERM ("mk_bn_supp_abs_tac", [trm])) val thms1 = @{thms supp_Pair supp_eqvt[symmetric] Un_assoc conj_assoc} val thms2 = @{thms de_Morgan_conj Collect_disj_eq finite_Un} val thms3 = @{thms alphas prod_alpha_def prod_fv.simps rel_prod_conv permute_prod_def prod.rec prod.case prod.inject not_True_eq_False empty_def[symmetric] finite.emptyI} fun prove_fv_supp qtys qtrms fvs fv_bns alpha_bns fv_simps eq_iffs perm_simps fv_bn_eqvts qinduct bclausess ctxt = let val goals1 = map mk_fvs_goals fvs val goals2 = map2 mk_fv_bns_goals fv_bns alpha_bns fun tac ctxt = SUBGOAL (fn (goal, i) => let val (fv_fun, arg) = goal |> Envir.eta_contract |> Logic.strip_assums_concl |> HOLogic.dest_Trueprop |> fst o HOLogic.dest_eq |> dest_comb val supp_abs_tac = case (AList.lookup (op=) (qtrms ~~ bclausess) (head_of arg)) of SOME bclauses => EVERY' (mk_supp_abs_tac ctxt bclauses) | NONE => mk_bn_supp_abs_tac ctxt fv_fun val eqvt_rconfig = eqvt_relaxed_config addpres (perm_simps @ fv_bn_eqvts) in EVERY' [ TRY o asm_full_simp_tac (add_simpset ctxt (@{thm supp_Pair[symmetric]}::fv_simps)), TRY o supp_abs_tac, TRY o simp_tac (add_simpset ctxt @{thms supp_def supp_rel_def}), TRY o eqvt_tac ctxt eqvt_rconfig, TRY o simp_tac (add_simpset ctxt (@{thms Abs_eq_iff} @ eq_iffs)), TRY o asm_full_simp_tac (add_simpset ctxt thms3), TRY o simp_tac (add_simpset ctxt thms2), TRY o asm_full_simp_tac (add_simpset ctxt (thms1 @ (symmetric fv_bn_eqvts)))] i end) in induct_prove qtys (goals1 @ goals2) qinduct tac ctxt |> map (atomize ctxt) |> map (simplify (put_simpset HOL_basic_ss ctxt addsimps @{thms fun_eq_iff[symmetric]})) end fun prove_bns_finite qtys qbns qinduct qbn_simps ctxt = let fun mk_goal qbn = let val arg_ty = domain_type (fastype_of qbn) val finite = @{term "finite :: atom set => bool"} in (arg_ty, fn x => finite $ (to_set (qbn $ x))) end val props = map mk_goal qbns val ss_tac = asm_full_simp_tac (put_simpset HOL_basic_ss ctxt addsimps (qbn_simps @ @{thms set_simps set_append finite_insert finite.emptyI finite_Un})) in induct_prove qtys props qinduct (K ss_tac) ctxt end fun prove_perm_bn_alpha_thms qtys qperm_bns alpha_bns qinduct qperm_bn_simps qeq_iffs qalpha_refls ctxt = let val ([p], ctxt') = Variable.variant_fixes ["p"] ctxt val p = Free (p, @{typ perm}) fun mk_goal qperm_bn alpha_bn = let val arg_ty = domain_type (fastype_of alpha_bn) in (arg_ty, fn x => (mk_id (Abs ("", arg_ty, alpha_bn $ Bound 0 $ (qperm_bn $ p $ Bound 0)))) $ x) end val props = map2 mk_goal qperm_bns alpha_bns val ss = @{thm id_def}::qperm_bn_simps @ qeq_iffs @ qalpha_refls val ss_tac = asm_full_simp_tac (put_simpset HOL_ss ctxt addsimps ss) in induct_prove qtys props qinduct (K ss_tac) ctxt' |> Proof_Context.export ctxt' ctxt |> map (simplify (put_simpset HOL_basic_ss ctxt addsimps @{thms id_def})) end fun prove_permute_bn_thms qtys qbns qperm_bns qinduct qperm_bn_simps qbn_defs qbn_eqvts ctxt = let val ([p], ctxt') = Variable.variant_fixes ["p"] ctxt val p = Free (p, @{typ perm}) fun mk_goal qbn qperm_bn = let val arg_ty = domain_type (fastype_of qbn) in (arg_ty, fn x => (mk_id (Abs ("", arg_ty, HOLogic.mk_eq (mk_perm p (qbn $ Bound 0), qbn $ (qperm_bn $ p $ Bound 0)))) $ x)) end val props = map2 mk_goal qbns qperm_bns val ss = @{thm id_def}::qperm_bn_simps @ qbn_defs val ss_tac = EVERY' [asm_full_simp_tac (put_simpset HOL_basic_ss ctxt' addsimps ss), TRY o eqvt_tac ctxt' (eqvt_strict_config addpres qbn_eqvts), TRY o asm_full_simp_tac (put_simpset HOL_basic_ss ctxt')] in induct_prove qtys props qinduct (K ss_tac) ctxt' |> Proof_Context.export ctxt' ctxt |> map (simplify (put_simpset HOL_basic_ss ctxt addsimps @{thms id_def})) end (*** proves strong exhauts theorems ***) (* fixme: move into nominal_library *) fun abs_const bmode ty = let val (const_name, binder_ty, abs_ty) = case bmode of Lst => (@{const_name "Abs_lst"}, @{typ "atom list"}, @{type_name abs_lst}) | Set => (@{const_name "Abs_set"}, @{typ "atom set"}, @{type_name abs_set}) | Res => (@{const_name "Abs_res"}, @{typ "atom set"}, @{type_name abs_res}) in Const (const_name, [binder_ty, ty] ---> Type (abs_ty, [ty])) end fun mk_abs bmode trm1 trm2 = abs_const bmode (fastype_of trm2) $ trm1 $ trm2 fun is_abs_eq thm = let fun is_abs trm = case (head_of trm) of Const (@{const_name "Abs_set"}, _) => true | Const (@{const_name "Abs_lst"}, _) => true | Const (@{const_name "Abs_res"}, _) => true | _ => false in thm |> Thm.prop_of |> HOLogic.dest_Trueprop |> HOLogic.dest_eq |> fst |> is_abs end (* adds a freshness condition to the assumptions *) fun mk_ecase_prems lthy c (params, prems, concl) bclauses = let val tys = map snd params val binders = get_all_binders bclauses fun prep_binder (opt, i) = let val t = Bound (length tys - i - 1) in case opt of NONE => setify_ty lthy (nth tys i) t | SOME bn => to_set_ty (fastype_of1 (tys, bn $ t)) (bn $ t) end val prems' = case binders of [] => prems (* case: no binders *) | _ => binders (* case: binders *) |> map prep_binder |> fold_union_env tys |> (fn t => mk_fresh_star t c) |> (fn t => HOLogic.mk_Trueprop t :: prems) in mk_full_horn params prems' concl end (* derives the freshness theorem that there exists a p, such that (p o as) #* (c, t1,..., tn) *) fun fresh_thm ctxt c parms binders bn_finite_thms = let fun prep_binder (opt, i) = case opt of NONE => setify ctxt (nth parms i) | SOME bn => to_set (bn $ (nth parms i)) fun prep_binder2 (opt, i) = case opt of NONE => atomify ctxt (nth parms i) | SOME bn => bn $ (nth parms i) val rhs = HOLogic.mk_tuple ([c] @ parms @ (map prep_binder2 binders)) val lhs = binders |> map prep_binder |> fold_union |> mk_perm (Bound 0) val goal = mk_fresh_star lhs rhs |> mk_exists ("p", @{typ perm}) |> HOLogic.mk_Trueprop val ss = bn_finite_thms @ @{thms supp_Pair finite_supp finite_sets_supp} @ @{thms finite.intros finite_Un finite_set finite_fset} in Goal.prove ctxt [] [] goal (K (HEADGOAL (resolve_tac ctxt @{thms at_set_avoiding1} THEN_ALL_NEW (simp_tac (put_simpset HOL_ss ctxt addsimps ss))))) end (* derives an abs_eq theorem of the form Exists q. [as].x = [p o as].(q o x) for non-recursive binders Exists q. [as].x = [q o as].(q o x) for recursive binders *) fun abs_eq_thm ctxt fprops p parms bn_eqvt permute_bns (bclause as (BC (bmode, binders, bodies))) = case binders of [] => [] | _ => let val rec_flag = is_recursive_binder bclause val binder_trm = comb_binders ctxt bmode parms binders val body_trm = foldl1 HOLogic.mk_prod (map (nth parms) bodies) val abs_lhs = mk_abs bmode binder_trm body_trm val abs_rhs = if rec_flag then mk_abs bmode (mk_perm (Bound 0) binder_trm) (mk_perm (Bound 0) body_trm) else mk_abs bmode (mk_perm p binder_trm) (mk_perm (Bound 0) body_trm) val abs_eq = HOLogic.mk_eq (abs_lhs, abs_rhs) val peq = HOLogic.mk_eq (mk_perm (Bound 0) binder_trm, mk_perm p binder_trm) val goal = HOLogic.mk_conj (abs_eq, peq) |> (fn t => HOLogic.mk_exists ("q", @{typ "perm"}, t)) |> HOLogic.mk_Trueprop val ss = fprops @ @{thms set_simps set_append union_eqvt} @ @{thms fresh_star_Un fresh_star_Pair fresh_star_list fresh_star_singleton fresh_star_fset fresh_star_set} val tac1 = if rec_flag then resolve_tac ctxt @{thms Abs_rename_set' Abs_rename_res' Abs_rename_lst'} else resolve_tac ctxt @{thms Abs_rename_set Abs_rename_res Abs_rename_lst} val tac2 = EVERY' [simp_tac (put_simpset HOL_basic_ss ctxt addsimps ss), TRY o simp_tac (put_simpset HOL_ss ctxt)] in [ Goal.prove ctxt [] [] goal (K (HEADGOAL (tac1 THEN_ALL_NEW tac2))) |> (if rec_flag then eqvt_rule ctxt (eqvt_strict_config addpres bn_eqvt) else eqvt_rule ctxt (eqvt_strict_config addpres permute_bns)) ] end val setify = @{lemma "xs = ys ==> set xs = set ys" by simp} fun case_tac ctxt c bn_finite_thms eq_iff_thms bn_eqvt permute_bns perm_bn_alphas prems bclausess qexhaust_thm = let fun aux_tac prem bclauses = case (get_all_binders bclauses) of [] => EVERY' [resolve_tac ctxt [prem], assume_tac ctxt] | binders => Subgoal.SUBPROOF (fn {params, prems, concl, context = ctxt, ...} => let val parms = map (Thm.term_of o snd) params val fthm = fresh_thm ctxt c parms binders bn_finite_thms val ss = @{thms fresh_star_Pair union_eqvt fresh_star_Un} val (([(_, fperm)], fprops), ctxt') = Obtain.result (fn ctxt' => EVERY1 [eresolve_tac ctxt [exE], full_simp_tac (put_simpset HOL_basic_ss ctxt' addsimps ss), REPEAT o (eresolve_tac ctxt @{thms conjE})]) [fthm] ctxt val abs_eq_thms = flat (map (abs_eq_thm ctxt' fprops (Thm.term_of fperm) parms bn_eqvt permute_bns) bclauses) val ((_, eqs), ctxt'') = Obtain.result (fn ctxt'' => EVERY1 [ REPEAT o (eresolve_tac ctxt @{thms exE}), REPEAT o (eresolve_tac ctxt @{thms conjE}), REPEAT o (dresolve_tac ctxt [setify]), full_simp_tac (put_simpset HOL_basic_ss ctxt'' addsimps @{thms set_append set_simps})]) abs_eq_thms ctxt' val (abs_eqs, peqs) = split_filter is_abs_eq eqs val fprops' = map (eqvt_rule ctxt'' (eqvt_strict_config addpres permute_bns)) fprops @ map (eqvt_rule ctxt'' (eqvt_strict_config addpres bn_eqvt)) fprops (* for freshness conditions *) val tac1 = SOLVED' (EVERY' [ simp_tac (put_simpset HOL_basic_ss ctxt'' addsimps peqs), rewrite_goal_tac ctxt'' (@{thms fresh_star_Un[THEN eq_reflection]}), conj_tac ctxt'' (DETERM o resolve_tac ctxt'' fprops') ]) (* for equalities between constructors *) val tac2 = SOLVED' (EVERY' [ resolve_tac ctxt [@{thm ssubst} OF prems], rewrite_goal_tac ctxt'' (map safe_mk_equiv eq_iff_thms), rewrite_goal_tac ctxt'' (map safe_mk_equiv abs_eqs), conj_tac ctxt'' (DETERM o resolve_tac ctxt'' (@{thms refl} @ perm_bn_alphas)) ]) (* proves goal "P" *) val side_thm = Goal.prove ctxt'' [] [] (Thm.term_of concl) (K (EVERY1 [ resolve_tac ctxt'' [prem], RANGE [tac1, tac2] ])) |> singleton (Proof_Context.export ctxt'' ctxt) in resolve_tac ctxt [side_thm] 1 end) ctxt in EVERY1 [resolve_tac ctxt [qexhaust_thm], RANGE (map2 aux_tac prems bclausess)] end fun prove_strong_exhausts lthy exhausts bclausesss bn_finite_thms eq_iff_thms bn_eqvt permute_bns perm_bn_alphas = let val ((_, exhausts'), lthy') = Variable.import true exhausts lthy val ([c, a], lthy'') = Variable.variant_fixes ["c", "'a"] lthy' val c = Free (c, TFree (a, @{sort fs})) val (ecases, main_concls) = exhausts' (* ecases are of the form (params, prems, concl) *) |> map Thm.prop_of |> map Logic.strip_horn |> split_list val ecases' = (map o map) strip_full_horn ecases val premss = (map2 o map2) (mk_ecase_prems lthy'' c) ecases' bclausesss fun tac bclausess exhaust {prems, context} = case_tac context c bn_finite_thms eq_iff_thms bn_eqvt permute_bns perm_bn_alphas prems bclausess exhaust fun prove prems bclausess exhaust concl = Goal.prove lthy'' [] prems concl (tac bclausess exhaust) in map4 prove premss bclausesss exhausts' main_concls |> Proof_Context.export lthy'' lthy end (** strong induction theorems **) fun add_c_prop c c_ty trm = let val (P, arg) = dest_comb trm val (P_name, P_ty) = dest_Free P val (ty_args, bool) = strip_type P_ty in Free (P_name, (c_ty :: ty_args) ---> bool) $ c $ arg end fun add_qnt_c_prop c_name c_ty trm = trm |> HOLogic.dest_Trueprop |> incr_boundvars 1 |> add_c_prop (Bound 0) c_ty |> HOLogic.mk_Trueprop |> mk_all (c_name, c_ty) fun prep_prem lthy c c_name c_ty bclauses (params, prems, concl) = let val tys = map snd params val binders = get_all_binders bclauses fun prep_binder (opt, i) = let val t = Bound (length tys - i - 1) in case opt of NONE => setify_ty lthy (nth tys i) t | SOME bn => to_set_ty (fastype_of1 (tys, bn $ t)) (bn $ t) end val prems' = prems |> map (incr_boundvars 1) |> map (add_qnt_c_prop c_name c_ty) val prems'' = case binders of [] => prems' (* case: no binders *) | _ => binders (* case: binders *) |> map prep_binder |> fold_union_env tys |> incr_boundvars 1 |> (fn t => mk_fresh_star_ty c_ty t (Bound 0)) |> (fn t => HOLogic.mk_Trueprop t :: prems') val concl' = concl |> HOLogic.dest_Trueprop |> incr_boundvars 1 |> add_c_prop (Bound 0) c_ty |> HOLogic.mk_Trueprop in mk_full_horn (params @ [(c_name, c_ty)]) prems'' concl' end fun prove_strong_induct lthy induct exhausts size_thms bclausesss = let val ((_, [induct']), lthy') = Variable.import true [induct] lthy val ([c_name, a], lthy'') = Variable.variant_fixes ["c", "'a"] lthy' val c_ty = TFree (a, @{sort fs}) val c = Free (c_name, c_ty) val (prems, concl) = induct' |> Thm.prop_of |> Logic.strip_horn val concls = concl |> HOLogic.dest_Trueprop |> HOLogic.dest_conj |> map (add_c_prop c c_ty) |> map HOLogic.mk_Trueprop val prems' = prems |> map strip_full_horn |> map2 (prep_prem lthy'' c c_name c_ty) (flat bclausesss) fun pat_tac ctxt thm = Subgoal.FOCUS (fn {params, context = ctxt', ...} => let val ty_parms = map (fn (_, ct) => (fastype_of (Thm.term_of ct), ct)) params val vs = Term.add_vars (Thm.prop_of thm) [] val vs_tys = map (Type.legacy_freeze_type o snd) vs val assigns = map (lookup ty_parms) vs_tys val thm' = infer_instantiate ctxt' (map #1 vs ~~ assigns) thm in resolve_tac ctxt' [thm'] 1 end) ctxt THEN_ALL_NEW asm_full_simp_tac (put_simpset HOL_basic_ss ctxt) fun size_simp_tac ctxt = simp_tac (put_simpset size_ss ctxt addsimps (@{thms comp_def snd_conv} @ size_thms)) in Goal.prove_common lthy'' NONE [] prems' concls (fn {prems, context = ctxt} => Induction_Schema.induction_schema_tac ctxt prems THEN RANGE (map (pat_tac ctxt) exhausts) 1 THEN prove_termination_ind ctxt 1 THEN ALLGOALS (size_simp_tac ctxt)) |> Proof_Context.export lthy'' lthy end end (* structure *) diff --git a/thys/Ordinary_Differential_Equations/Ex/Lorenz/Lorenz_Approximation.thy b/thys/Ordinary_Differential_Equations/Ex/Lorenz/Lorenz_Approximation.thy --- a/thys/Ordinary_Differential_Equations/Ex/Lorenz/Lorenz_Approximation.thy +++ b/thys/Ordinary_Differential_Equations/Ex/Lorenz/Lorenz_Approximation.thy @@ -1,3050 +1,3051 @@ section"Example: Lorenz attractor" theory Lorenz_Approximation imports "HOL-ODE-Numerics.ODE_Numerics" Result_File_Coarse begin lemma replicate_numeral [simp]: "replicate (numeral k) x = x # replicate (pred_numeral k) x" by (simp add: numeral_eq_Suc) text \\label{sec:lorenz}\ text \TODO: move to isabelle? \ lifting_update blinfun.lifting lifting_forget blinfun.lifting lemma eventually_uniformly_on: "(\\<^sub>F x in uniformly_on T l. P x) = (\e>0. \f. (\x\T. dist (f x) (l x) < e) \ P f)" unfolding uniformly_on_def apply (subst eventually_INF) apply safe subgoal for E apply (cases "E = {}") subgoal by (auto intro!: exI[where x=1]) subgoal apply (auto simp: INF_principal_finite eventually_principal elim!: ) proof goal_cases case (1 x) from 1 have "0 < Min E" apply (subst Min_gr_iff) apply force apply force apply force done have *: "(\e\E. {f. \t\T. dist (f t) (l t) < e}) = {f. \t\T. dist (f t) (l t) < Min E}" using 1 apply (auto simp: ) apply (subst Min_gr_iff) apply force apply force apply force apply (drule bspec, assumption) apply (rule less_le_trans, assumption) apply auto done from 1 have "\f. (\x\T. dist (f x) (l x) < Min E) \ P f" unfolding * by simp then show ?case using 1(4)[rule_format, OF \0 < Min E\] by auto qed done subgoal for e apply (rule exI[where x="{e}"]) by (auto simp: eventually_principal) done lemma op_cast_image_impl[autoref_rules]: "(\x. x, op_cast_image::'a::executable_euclidean_space set \ 'b::executable_euclidean_space set) \ aform.appr_rel \ aform.appr_rel" if "DIM('a) = DIM('b)" using that apply (auto simp: aform.appr_rel_def intro!: relcompI) unfolding lv_rel_def set_rel_br by (force simp: intro!: brI dest!: brD) lemma cast_bl_blinfun_of_list[simp]: "cast_bl (blinfun_of_list xs::'a \\<^sub>L 'a) = (blinfun_of_list xs::'b\\<^sub>L'b)" if "DIM('a::executable_euclidean_space) = DIM('b::executable_euclidean_space)" using that apply (auto simp: cast_bl_rep intro!: blinfun_eqI) by (auto simp: blinfun_of_list_def blinfun_of_matrix_apply linear_sum linear.scaleR sum_Basis_sum_nth_Basis_list linear_cast) lemma cast_idem[simp]: "cast x = x" by (auto simp: cast_def) lemma cast_bl_idem[simp]: "cast_bl x = x" by (auto simp: cast_bl_rep intro!: blinfun_eqI) lemma op_cast_eucl1_image_impl[autoref_rules]: "(\x. x, op_cast_eucl1_image::'a::executable_euclidean_space c1_info set \ 'b::executable_euclidean_space c1_info set) \ aform.appr1_rel \ aform.appr1_rel" if "DIM_precond TYPE('a) D" "DIM_precond TYPE('b) D" using that proof (auto, goal_cases) case (1 a b a') then show ?case apply (auto simp: aform.appr1_rel_br set_rel_br br_def) subgoal for w x y z apply (auto simp: aform.c1_info_of_appr_def cast_eucl1_def aform.c1_info_invar_def split: option.splits) apply (rule image_eqI) apply (rule cast_eucl_of_list, force, force simp: Joints_imp_length_eq, force) subgoal for s t apply (rule image_eqI[where x="t"]) supply [simp del] = eucl_of_list_take_DIM apply (auto simp: flow1_of_list_def) apply (subst cast_eucl_of_list) subgoal by simp subgoal by (auto dest!: Joints_imp_length_eq simp: power2_eq_square flow1_of_list_def[abs_def]) subgoal by simp done done subgoal for w x apply (rule image_eqI[where x="cast_eucl1 (w, x)"]) apply (auto simp: aform.c1_info_of_appr_def cast_eucl1_def aform.c1_info_invar_def split: option.splits) apply (rule image_eqI) apply (rule cast_eucl_of_list, force, force simp: Joints_imp_length_eq, force) subgoal for s t apply (rule image_eqI[where x="t"]) supply [simp del] = eucl_of_list_take_DIM apply (auto simp: flow1_of_list_def) apply (subst cast_eucl_of_list) subgoal by simp subgoal by (auto dest!: Joints_imp_length_eq simp: power2_eq_square flow1_of_list_def[abs_def]) subgoal by simp done done done qed lemma less_3_iff: "i < (3::nat) \ i = 0 \ i = 1 \ i = 2" by arith definition mat3_of_vec::"R3 \ real^3^3" where "mat3_of_vec x = (let xs = list_of_eucl x in eucl_of_list [xs!0,0,0, xs!1,0,0, xs!2,0,0])" lemma ll3: "{..<3} = {0,1,2::nat}" by (auto simp: less_3_iff) lemma mat3_of_vec: "cast (mat3_of_vec x *v eucl_of_list [1, 0, 0]) = x" by (auto simp: mat3_of_vec_def eucl_of_list_matrix_vector_mult_eq_sum_nth_Basis_list linear_sum linear_cast linear.scaleR ll3 linear_add Basis_list_R3 inner_prod_def prod_eq_iff) primrec bisect_form where "bisect_form p f xs l u 0 = (l, u)" | "bisect_form p f xs l u (Suc n) = (let m = (l + u)/2 in if approx_form_aform p (f m) xs then bisect_form p f xs l m n else bisect_form p f xs m u n)" text \This should prove that the expansion estimates are sufficient.\ lemma expansion_main: "expansion_main (coarse_results) = Some True" by eval context includes floatarith_notation begin definition "matrix_of_degrees2\<^sub>e = (let e = Var 2; ur = Rad_of (Var 0); vr = Rad_of (Var 1); x1 = Cos ur; x2 = Cos vr; y1 = Sin ur; y2 = Sin vr in [x1 + (e * (x2 - x1)), 0, 0, y1 + (e * (y2 - y1)), 0, 0, 0, 0, 0])" definition "matrix_of_degrees2 u v = approx_floatariths 30 matrix_of_degrees2\<^sub>e (aforms_of_ivls [u, v, 0] [u, v, 1])" text \following \vector_field.h\ / \vector_field.cc\\ abbreviation "S \ 10::real" abbreviation "B \ 8/3::real" abbreviation "TEMP \ sqrt((S + 1) * (S + 1) + 4 * S * (28 - 1))" abbreviation "K1 \ S / TEMP" abbreviation "K2 \ (S - 1 + TEMP) / (2 * S)" abbreviation "K3 \ (S - 1 - TEMP) / (2 * S)" abbreviation "E1 \ (- (S + 1) + TEMP) / 2" abbreviation "E2 \ (- (S + 1) - TEMP) / 2" abbreviation "E3 \ - B" abbreviation "C1 \ \X. X ! 0 + X ! 1" abbreviation "C2 \ \X. K1 * C1 X * X ! 2" schematic_goal lorenz_fas: "[E1 * X!0 - C2 X, E2 * X!1 + C2 X, E3 * X!2 + C1 X * (K2 * X!0 + K3 * X!1)] = interpret_floatariths ?fas X" by (reify_floatariths) concrete_definition lorenz_fas uses lorenz_fas end interpretation lorenz: ode_interpretation true_form UNIV lorenz_fas "\(X0, X1, X2). (E1 * X0 - K1 * (X0 + X1) * X2, E2 * X1 + K1 * (X0 + X1) * X2, E3 * X2 + (X0 + X1) * (K2 * X0 + K3 * X1))::real*real*real" "d::3" for d by standard (auto simp: lorenz_fas_def less_Suc_eq_0_disj nth_Basis_list_prod Basis_list_real_def mk_ode_ops_def eucl_of_list_prod power2_eq_square inverse_eq_divide intro!: isFDERIV_I) value [code] "length (slp_of_fas lorenz_fas)" definition "mig_aform p x = mig_componentwise (Inf_aform' p x) (Sup_aform' p x)" context includes floatarith_notation begin definition "mig_aforms p x = real_of_float ((lower o the) ((approx p (Norm (map (Num o float_of o (mig_aform p)) x))) []))" definition "column_of_c1_info x N j = (map (\i. the (snd x) ! i) (map (\i. i * N + j) [0..e b))) (rotate_z_fa (Rad_of (R\<^sub>e a)))) xs)" definition "perspective_projection_aforms xs = the (approx_slp_outer 30 3 (rotate_zx_slp (-30) (-60) (map Var [0..<3])) xs)" definition "print_lorenz_aform print_fun cx cy cz ci cd1 cd2 = (\a b. let (s1, n) = ((-6), False); _ = print_fun (String.implode (''# gen(''@ show a@''): ''@ shows_aforms_hr (b) '''' @ ''\'')); _ = print_fun (String.implode (''# box(''@ show a@''): ''@ shows_box_of_aforms_hr (b) '''' @ ''\'')); ((x0, y0, z0), (x1, y1, z1)) = case (map (Inf_aform' 30) (take 3 b), map (Sup_aform' 30) (take 3 b)) of ([x0, y0, z0], [x1, y1, z1]) \ ((x0, y0, z0), (x1, y1, z1)); _ = print_fun (String.implode (shows_segments_of_aform 0 1 b ((shows cx o shows_space o shows z0 o shows_space o shows z1)'''') ''\'')); _ = print_fun (String.implode (shows_segments_of_aform 0 2 b ((shows cy o shows_space o shows y0 o shows_space o shows y1)'''') ''\'')); _ = print_fun (String.implode (shows_segments_of_aform 1 2 b ((shows cz o shows_space o shows x0 o shows_space o shows x1)'''') ''\'')); PS = perspective_projection_aforms b; _ = print_fun (String.implode (shows_segments_of_aform 0 1 PS ((shows ci o shows_space o shows (fst (PS ! 2)) o shows_space o shows (fst (b ! 2))) '''') ''\'')) in if \ a \ length b > 10 then print_fun (String.implode (shows_aforms_vareq 3 [(0, 1), (0, 2), (1, 2)] [0..<3] cd1 cd2 (FloatR 1 s1 * (if n then real_divl 30 1 (max (mig_aforms 30 (map (\i. b ! i) [3,6,9])) (mig_aforms 30 (map (\i. b ! i) [4,7,10]))) else 1)) \ \always length \2^s!\\ ''# no C1 info'' b '''')) else ())" definition "print_lorenz_aform_std print_fun = print_lorenz_aform print_fun ''0x000001'' ''0x000002'' ''0x000012'' ''0x000003'' [''0xa66f00'', ''0x06266f'', ''0xc60000''] [''0xffaa00'', ''0x1240ab'', ''0xc60000'']" definition "lorenz_optns print_funo = (let pf = the_default (\_ _. ()) (map_option print_lorenz_aform_std print_funo); tf = the_default (\_ _. ()) (map_option (\print_fun a b. let _ = print_fun (String.implode (''# '' @ a @ ''\'')) in case b of Some b \ (print_fun (String.implode (''# '' @ shows_box_of_aforms_hr (b) '''' @ ''\''))) | None \ ()) print_funo) in \ precision = 30, adaptive_atol = FloatR 1 (-30), adaptive_rtol = FloatR 1 (-30), method_id = 2, start_stepsize = FloatR 1 (- 8), iterations = 40, halve_stepsizes = 10, widening_mod = 40, rk2_param = FloatR 1 0, default_reduce = correct_girard 30 50 25, printing_fun = pf, tracing_fun = tf \)" definition lorenz_optns' where "lorenz_optns' pf m N rk2p a = lorenz_optns pf \ default_reduce := correct_girard 30 m N, rk2_param := rk2p, adaptive_atol := a, adaptive_rtol := a \" definition mirror_irects where "mirror_irects = map (\irect. case irect of [i, j, k] \ [if j < 0 then - i else i , abs j, k] | irect \ irect)" definition "print_irects irects = (let _ = map (\is. let _ = map (\j. let _ = print (String.implode (show j)) in print (STR '' '')) is in print (STR ''\'')) irects in ())" abbreviation "aforms_of_ivl \ \x. aforms_of_ivls (fst x) (snd x)" definition "conefield_propagation\<^sub>e = ([Deg_of (Arctan (Var (6) / Var (3))), Deg_of (Arctan (Var (7) / Var (4))), Min (Norm [Var(3), Var (6), Var (9)]) (Norm [Var(4), Var (7), Var (10)])])" definition "conefield_propagation DX = approx_floatariths 30 conefield_propagation\<^sub>e DX" definition "conefield_propagation_slp = slp_of_fas conefield_propagation\<^sub>e" lemma [simp]: "length conefield_propagation_slp = 51" by eval definition op_with_unit_matrix :: "'a::real_normed_vector \ 'a \ 'a \\<^sub>L 'a" where "op_with_unit_matrix X = (X, 1\<^sub>L)" context includes blinfun.lifting begin lemma matrix_vector_mult_blinfun_works[simp]: "matrix e *v g = e g" for e::"(real^'n) \\<^sub>L (real^'m)" by transfer (simp add: bounded_linear_def matrix_works) end lemma length_conefield_propagation\<^sub>e[simp]: "length conefield_propagation\<^sub>e = 3" by (simp add: conefield_propagation\<^sub>e_def) lemma interpret_floatariths_conefield_propagation: "interpret_floatariths conefield_propagation\<^sub>e (list_of_eucl (vec1_of_flow1 (xDX::(real^3) \ ((real^3)\\<^sub>L(real^3))))) = (let DX = snd xDX; DXu = DX (eucl_of_list [1, 0, 0]); DXv = DX (eucl_of_list [0, 1, 0]) in [deg_of (arctan (vec_nth DXu 1 / vec_nth DXu 0)), deg_of (arctan (vec_nth DXv 1 / vec_nth DXv 0)), min (norm DXu) (norm DXv)] )" apply (auto simp: conefield_propagation\<^sub>e_def Let_def interpret_mvmult_nth[where 'n=3 and 'm=3] inverse_eq_divide vec1_of_flow1_def nth_append) apply (auto simp: matrix_inner_Basis_list ) apply (auto simp: interpret_floatarith_norm[where 'a="real ^ 3"] einterpret_mvmult_fa[where 'n=3 and 'm=3] matrix_inner_Basis_list nth_append) by (auto simp: matrix_def axis_eq_eucl_of_list eucl_of_list_012) definition "conefield_bounds_form l u = (fold Conj [ Less (-90) (N\<^sub>r l), LessEqual (N\<^sub>r l) (N\<^sub>r u), LessEqual (Var 9) (0), LessEqual 0 (Var 9), Less (N\<^sub>r u) (90), Less 0 (Var 3), AtLeastAtMost (Var 6) (Tan (Rad_of (N\<^sub>r l)) * Var 3) (Tan (Rad_of (N\<^sub>r u)) * Var 3)] true_form)" definition "contract_angles X i = (snd (bisect_form 30 (\x. conefield_bounds_form x (89)) X 89 (-89) i), snd (bisect_form 30 (conefield_bounds_form (-89)) X (-89) 89 i))" definition "approx_conefield_bounds (DX::(R3 \ (R3 \\<^sub>L R3)) set) l u = do { let DX = (cast_eucl1 ` DX::3 eucl1 set); DXo \ aform.vec1rep DX; DX \ (case DXo of None \ do { let _ = aform.print_msg (''# approx_conefield_bounds failed DXo...''); SUCCEED } | Some DX \ RETURN DX); let _ = aform.trace_set (''# approx_conefield_bounds DX: '') (Some DX); approx_form_spec (conefield_bounds_form l u) (list_of_eucl ` DX) }" lemma [autoref_rules]: includes autoref_syntax shows "(conefield_bounds_form, conefield_bounds_form) \ Id \ Id \ Id " by auto lemma [autoref_rules_raw]: "DIM_precond TYPE((real, 3) vec \ ((real, 3) vec, 3) vec) 12" "DIM_precond TYPE(R3) 3" "DIM_precond TYPE((real, 3) vec) 3" by auto schematic_goal approx_conefield_bounds_impl: includes autoref_syntax fixes optns::"real aform numeric_options" assumes [autoref_rules]: "(DXi, DX) \ aform.appr1_rel" assumes [autoref_rules]: "(li, l) \ Id" assumes [autoref_rules]: "(ui, u) \ Id" notes [autoref_rules] = aform.print_msg_impl[where optns = optns] aform.ivl_rep_of_set_autoref[where optns = optns] aform.transfer_operations(12)[where optns = optns] aform.approx_euclarithform[where optns=optns] aform.trace_set_impl[of optns] shows "(nres_of ?r, approx_conefield_bounds $ DX $ l $ u) \ \bool_rel\nres_rel" unfolding autoref_tag_defs unfolding approx_conefield_bounds_def including art by autoref_monadic concrete_definition approx_conefield_bounds_impl for optns li ui DXi uses approx_conefield_bounds_impl lemmas [autoref_rules] = approx_conefield_bounds_impl.refine context includes autoref_syntax begin lemma [autoref_rules]: "(real_of_ereal, real_of_ereal) \ ereal_rel \ rnv_rel" "(\, \) \ ereal_rel" by auto end lemma interpret_form_true_form[simp]: "interpret_form true_form \ \_. True" by (force simp: true_form_def intro!: eq_reflection) lemma interpret_form_conefield_bounds_form_list: "interpret_form (conefield_bounds_form L U) [x, y, z, ux, vx, wx, uy, vy, wy, uz, vz, wz] \ (0 < ux \ -90 < L \ L \ U \ U < 90 \ uz = 0 \ uy \ tan (rad_of U) * ux \ tan (rad_of L) * ux \ uy)" if "U \ float" "L \ float" "e \ float" "em \ float" using that by (auto simp: conefield_bounds_form_def L2_set_def) lemma list_of_eucl_eucl1_3: includes vec_syntax shows "(list_of_eucl (vec1_of_flow1 (xDX::(real^3) \ ((real^3)\\<^sub>L(real^3))))) = (let (x, DX) = xDX; DXu = DX (eucl_of_list [1, 0, 0]); DXv = DX (eucl_of_list [0, 1, 0]); DXw = DX (eucl_of_list [0, 0, 1]) in [x $ 0, x $ 1, x $ 2, vec_nth DXu 0, vec_nth DXv 0, vec_nth DXw 0, vec_nth DXu 1, vec_nth DXv 1, vec_nth DXw 1, vec_nth DXu 2, vec_nth DXv 2, vec_nth DXw 2])" apply (auto simp: matrix_inner_Basis_list Let_def vec1_of_flow1_def concat_map_map_index less_Suc_eq_0_disj list_of_eucl_matrix eval_nat_numeral aform.inner_Basis_eq_vec_nth intro!: nth_equalityI split: prod.splits) by (auto simp: matrix_def axis_eq_eucl_of_list eucl_of_list_012) lemma interpret_form_conefield_bounds_form: "interpret_form (conefield_bounds_form L U) (list_of_eucl (vec1_of_flow1 (xDX::(real^3) \ ((real^3)\\<^sub>L(real^3))))) = (let DX = snd xDX; DXu = DX (eucl_of_list [1, 0, 0]); DXv = DX (eucl_of_list [0, 1, 0]); uz = vec_nth DXu 2; uy = vec_nth DXu 1; ux = vec_nth DXu 0; vy = vec_nth DXv 1; vx = vec_nth DXv 0 in ux > 0 \ -90 < L \ L \ U \ U < 90 \ uz = 0 \ (uy / ux) \ {tan (rad_of L) .. tan (rad_of U)} )" if "L \ float" "U \ float" using that unfolding list_of_eucl_eucl1_3 Let_def by (auto split: prod.splits simp: interpret_form_conefield_bounds_form_list divide_simps) lemma approx_conefield_bounds_cast: "approx_conefield_bounds DX L U \ SPEC (\b. b \ (\(x, dx) \ cast_eucl1 ` DX::3 eucl1 set. let u' = dx (eucl_of_list [1, 0, 0]) in vec_nth u' 1 / vec_nth u' 0 \ {tan (rad_of L) .. tan (rad_of U)} \ vec_nth u' 2 = 0 \ vec_nth u' 0 > 0 \ -90 < L \ L \ U \ U < 90) )" if "L \ float" "U \ float" unfolding approx_conefield_bounds_def apply refine_vcg apply (auto simp: env_len_def ) subgoal for a b c apply (drule bspec, assumption) unfolding interpret_form_conefield_bounds_form[OF that] by (auto simp: Let_def divide_simps) done lemma approx_conefield_bounds[le, refine_vcg]: "approx_conefield_bounds DX l u \ SPEC (\b. b \ (\(x, dx) \ DX::R3 c1_info set. let (u'1, u'2, u'3) = dx ((1, 0, 0)) in u'2 / u'1 \ {tan (rad_of l) .. tan (rad_of u)} \ u'3 = 0 \ u'1 > 0 \ -90 < l \ l \ u \ u < 90) )" if "l \ float" "u \ float" apply (rule approx_conefield_bounds_cast[le, OF that]) apply (auto dest!: bspec simp: Let_def split: prod.splits) by (auto simp: cast_eucl1_def cast_def) schematic_goal MU\<^sub>e: "-E2 / E1 = interpret_floatarith ?fas []" by (reify_floatariths) concrete_definition MU\<^sub>e uses MU\<^sub>e schematic_goal NU\<^sub>e: "-E3 / E1 = interpret_floatarith ?fas []" by (reify_floatariths) concrete_definition NU\<^sub>e uses NU\<^sub>e definition "approx_ivls p fas xs = do { let xs = ivls_of_aforms p xs; res \ those (map (\f. approx p f xs) fas); Some (map (real_of_float o lower) res, map (real_of_float o upper) res) }" definition "deform p t exit XDX = (case XDX of (lu, (X, DX)) \ let d = ldec 0.1; d = if exit then (1 - real_of_float (lb_sqrt 30 (1 - 2 * float_of (d)))) else d; sd = real_of_float ((float_of (d*d))); C0Deform = aforms_of_ivls [-sd,-sd,-sd] [sd, sd, sd]; result = msum_aforms' X C0Deform in (lu, case DX of None \ (result, None) | Some DX \ let C1_norm = 2 * d; C1_norm = if exit then real_divr 30 C1_norm (1 - C1_norm) else C1_norm; l = -C1_norm; u = C1_norm; D_M = aforms_of_ivls [1 + l,0,0, 0,1 + l,0, 0,0,1 + l] [1 + u,0,0, 0,1 + u,0, 0,0,1 + u]; (ri, ru) = the (approx_ivls p (mmult_fa 3 3 3 (map Var [0..<9]) (map Var [9..<18])) (D_M @ DX)); Dresult = aforms_of_ivls ri ru; resultDresult = product_aforms result Dresult in (take 3 resultDresult, Some (drop 3 resultDresult))))" definition "ivls_of_aforms' p r = (map (Inf_aform' p) r, map (Sup_aform' p) r)" definition "compute_half_exit p t XDX = (case XDX of ((l, u::ereal), (X, DX)) \ let \ \ASSERTING that \Y\ straddles zero\ (x0, y0, _) = case map (Inf_aform' p) X of [x,y,z] \ (x, y, z); (x1, y1, _) = case map (Sup_aform' p) X of [x,y,z] \ (x, y, z); splitting = x0 = 0 \ x1 = 0; sign_x = if (x0 + x1) / 2 > 0 then 1 else -1; mag_x = max (abs x0) (abs x1); sign_x\<^sub>e = N\<^sub>r sign_x; exit_rad\<^sub>e = N\<^sub>r (ldec 0.1); X\<^sub>e = Var (0); Y\<^sub>e = Var (1); Z\<^sub>e = Var (2); max_x_over_r\<^sub>e = N\<^sub>r mag_x / exit_rad\<^sub>e; abs_x_over_r\<^sub>e = (Abs X\<^sub>e) / exit_rad\<^sub>e; result = (if splitting then let result = (the (approx_floatariths p [sign_x\<^sub>e * exit_rad\<^sub>e, Y\<^sub>e * Powr (max_x_over_r\<^sub>e) MU\<^sub>e, Z\<^sub>e * Powr (max_x_over_r\<^sub>e) NU\<^sub>e] X)); (ir, sr) = ivls_of_aforms' p result in aforms_of_ivls (ir[1:=min 0 (ir!1), 2:=min 0 (ir!2)]) (sr[1:=max 0 (sr!1), 2:=max 0 (sr!2)]) else the (approx_floatariths p [sign_x\<^sub>e * exit_rad\<^sub>e, Y\<^sub>e * Powr (abs_x_over_r\<^sub>e) MU\<^sub>e, Z\<^sub>e * Powr (abs_x_over_r\<^sub>e) NU\<^sub>e] X)); _ = () in ((l::ereal, \::ereal), (case DX of None \ (result, None) | Some DX \ let ux\<^sub>e = Var (3); uy\<^sub>e = Var (6); P21\<^sub>e = if splitting then (MU\<^sub>e / exit_rad\<^sub>e) * Y\<^sub>e * sign_x\<^sub>e * Powr (max_x_over_r\<^sub>e) (MU\<^sub>e - 1) \ \No need for \Hull(0)\ because \y\ straddles zero\ else (MU\<^sub>e / exit_rad\<^sub>e) * Y\<^sub>e * sign_x\<^sub>e * Powr (abs_x_over_r\<^sub>e) (MU\<^sub>e - 1); P22\<^sub>e = if splitting then Powr (max_x_over_r\<^sub>e) MU\<^sub>e else Powr (abs_x_over_r\<^sub>e) MU\<^sub>e; P31\<^sub>e = if splitting then sign_x\<^sub>e * (NU\<^sub>e / exit_rad\<^sub>e) * Z\<^sub>e * Powr (max_x_over_r\<^sub>e) (NU\<^sub>e - 1) \ \No need for \Hull(\)\ because scaling afterwards\ else sign_x\<^sub>e * (NU\<^sub>e / exit_rad\<^sub>e) * Z\<^sub>e * Powr (abs_x_over_r\<^sub>e) (NU\<^sub>e - 1); ry = (P21\<^sub>e * ux\<^sub>e) + (P22\<^sub>e * uy\<^sub>e); rz = P31\<^sub>e * ux\<^sub>e; (iDr, sDr) = the (approx_ivls p ([0, 0, 0, ry, 0, 0, rz, 0, 0]) (X @ DX)); Dresult = aforms_of_ivls (iDr[3:=min 0 (iDr!3)]) (sDr[3:=max 0 (sDr!3)]); resultDresult = product_aforms result Dresult in (take 3 resultDresult, Some (drop 3 resultDresult)) )))" fun list3 where "list3 [a,b,c] = (a, b, c)" definition "split_x n x0 y0 z0 x1 y1 z1 = (let elem = (\(x0, x1). aforms_of_ivls [x0, y0, z0] [x1, y1, z1]); coord = (\x0 n i. i * x0 * FloatR 1 (-int n)); us = map (coord x0 n) (rev [0..<(2^n)]) @ map (coord x1 n) [Suc 0..X'. ((l, u), (X', DX'))) X's; Xes = map (compute_half_exit p t) XDX's; Xlumpies = map (deform p t True) Xes in Xlumpies)" definition "cube_enteri = (map ldec [-0.1, -0.00015, 0.1, 0.8,0,0, 0.0005,0,0, 0,0,0], map udec [ 0.1, 0.00015, 0.1, 1.7,0,0, 0.002,0,0, 0,0,0])" definition "cube_enter = set_of_ivl (pairself eucl_of_list cube_enteri)" value [code] "println ((show) (map (ivls_of_aforms' 100 o list_of_appr1e_aform) (compute_cube_exit 30 (FloatR 1 (-10)) ((ereal 1, ereal 1), (aforms_of_ivls (take 3 (fst cube_enteri)) (take 3 (snd cube_enteri)), Some (aforms_of_ivls (drop 3 (fst cube_enteri)) (drop 3 (snd cube_enteri))))))))" definition "cube_exiti = [(aforms_of_ivls (map ldec [-0.12, -0.024, -0.012]) (map udec [-0.088, 0.024, 0.13]), Some (aforms_of_ivls (map ldec [0,0,0, -0.56,0,0, -0.6,0,0]) (map udec [0,0,0, 0.56,0,0, -0.08,0,0]))), (aforms_of_ivls (map ldec [ 0.088, -0.024, -0.012]) (map udec [ 0.12, 0.024, 0.13]), Some (aforms_of_ivls (map ldec [0,0,0, -0.53,0,0, 0.08,0,0]) (map udec [0,0,0, 0.56,0,0, 0.6,0,0])))]" definition "cube_exitv = aform.c1_info_of_apprs cube_exiti" lemma cube_enteri[autoref_rules]: "(cube_enteri, cube_enter::'a set) \ lvivl_rel" if "DIM_precond TYPE('a::executable_euclidean_space) 12" using that by (auto simp: cube_enteri_def cube_enter_def set_of_ivl_def intro!: brI lv_relivl_relI) lemma cube_exiti[autoref_rules]: "(cube_exiti, cube_exitv::'n eucl1 set) \ clw_rel aform.appr1_rel" if "DIM_precond TYPE('n::enum rvec) 3" unfolding cube_exitv_def cube_exiti_def apply (rule aform.clw_rel_appr1_relI) using that by (auto simp: aform.c1_info_invar_def power2_eq_square) definition "lorenz_interrupt (optns::real aform numeric_options) b (eX::3 eucl1 set) = do { ((el, eu), X) \ scaleR2_rep eX; let fX = fst ` X; fentry \ op_image_fst_ivl (cube_enter::3 vec1 set); interrupt \ aform.op_subset (fX:::aform.appr_rel) fentry; (ol, ou) \ ivl_rep fentry; aform.CHECKs (ST ''asdf'') (0 < el \ ol \ ou); let _ = (if b then aform.trace_set (ST ''Potential Interrupt: '') (Some fX) else ()); let _ = (if b then aform.trace_set (ST ''With: '') (Some ({ol .. ou::3 rvec}:::aform.appr_rel)) else ()); if \b \ \interrupt then RETURN (op_empty_coll, mk_coll eX) else do { vX \ aform.vec1rep X; let _ = (if b then aform.trace_set1e (ST ''Actual Interrupt: '') (Some eX) else ()); let l = (eucl_of_list [-1/2/2,-1/2/2,-1/2/2]::3 rvec); let u = eucl_of_list [1/2/2, 1/2/2, 1/2/2]; ASSERT (l \ u); let CX = mk_coll ({l .. u}:::aform.appr_rel); (C0::3 eucl1 set) \ scaleRe_ivl_coll_spec el eu (fst ` cube_exitv \ UNIV); (C1::3 eucl1 set) \ scaleRe_ivl_coll_spec el eu (cube_exitv); case vX of None \ RETURN (CX, C0) | Some vX \ do { b \ aform.op_subset vX cube_enter; aform.CHECKs (ST ''FAILED: TANGENT VECTORs are not contained'') b; RETURN (CX, C1) } } }" lemma [autoref_rules]: includes autoref_syntax shows "(real_of_int, real_of_int) \ int_rel \ rnv_rel" "(ldec, ldec) \ Id \ rnv_rel" "(udec, udec) \ Id \ rnv_rel" by auto schematic_goal lorenz_interrupti: includes autoref_syntax assumes[autoref_rules]: "(bi, b) \ bool_rel" "(Xi, X::3 eucl1 set) \ aform.appr1e_rel" "(optnsi, optns) \ Id" shows "(nres_of ?r, lorenz_interrupt optns b X) \ \clw_rel aform.appr_rel \\<^sub>r clw_rel aform.appr1e_rel\nres_rel" unfolding lorenz_interrupt_def including art by autoref_monadic concrete_definition lorenz_interrupti for optnsi1 bi Xi uses lorenz_interrupti[where optnsi = "optnsi" and optnsa = "\_ _ _ _ _ _ _ _. optnsi" and optnsb = "\_ _ _ _ _ _ _ _ _. optnsi" and optnsc = "\_ _ _ _ _ _ _ _ _ _ _. optnsi" and optnsd = "\_ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _. optnsi" and optnse = "\_ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _. optnsi" and optnsf = "\_ _ _ _ _ _ _ _ _. optnsi" and optns = "\_ _ _ _ _. optnsi" for optnsi] lemma lorenz_interrupti_refine[autoref_rules]: includes autoref_syntax shows "(\optnsi bi Xi. (lorenz_interrupti optnsi bi Xi), lorenz_interrupt) \ num_optns_rel \ bool_rel \ aform.appr1e_rel \ \clw_rel aform.appr_rel \\<^sub>r clw_rel aform.appr1e_rel\dres_nres_rel" using lorenz_interrupti.refine by (auto simp: nres_rel_def dres_nres_rel_def) definition "(large_cube::R3 set) = {-1/4 .. 1/4} \ {-1/4 .. 1/4} \ {-1/4 .. 1/4}" definition "cube_entry = (cast_eucl1 ` (flow1_of_vec1 ` cube_enter::3 eucl1 set)::R3 c1_info set)" definition "cube_exit = (cast_eucl1 ` (cube_exitv::3 eucl1 set)::R3 c1_info set)" text \protect locale parameters\ lemma flow0_cong[cong]: "auto_ll_on_open.flow0 ode X = auto_ll_on_open.flow0 ode X" by (auto simp:) lemma existence_ivl0_cong[cong]: "auto_ll_on_open.existence_ivl0 ode X = auto_ll_on_open.existence_ivl0 ode X" by (auto simp:) lemma Dflow_cong[cong]: "c1_on_open_euclidean.Dflow ode ode_d X = c1_on_open_euclidean.Dflow ode ode_d X" by (auto simp:) lemma flowsto_cong[cong]: "c1_on_open_euclidean.flowsto ode ode_d D = c1_on_open_euclidean.flowsto ode ode_d D" by (auto simp:) lemma poincare_mapsto_cong[cong]: "c1_on_open_euclidean.poincare_mapsto ode X = c1_on_open_euclidean.poincare_mapsto ode X" by (auto simp:) lemma returns_to_cong[cong]: "auto_ll_on_open.returns_to ode X = auto_ll_on_open.returns_to ode X" by (auto simp:) lemma return_time_cong[cong]: "auto_ll_on_open.return_time ode X = auto_ll_on_open.return_time ode X" by (auto simp: ) lemma poincare_map_cong[cong]: "auto_ll_on_open.poincare_map ode X = auto_ll_on_open.poincare_map ode X" by (auto simp: ) lemma eq_nth_iff_index: "distinct xs \ n < length xs \ i = xs ! n \ index xs i = n" using index_nth_id by fastforce lemma cast_in_BasisI: "(cast i::'a) \ Basis" if "(i::'c) \ Basis""DIM('c::executable_euclidean_space) = DIM('a::executable_euclidean_space)" using that by (auto simp: cast_def eucl_of_list_nth inner_Basis if_distrib if_distribR eq_nth_iff_index cong: if_cong) lemma cast_le_iff: "(cast (x::'a)::'c) \ y \ x \ cast y" if "DIM('c::executable_euclidean_space) = DIM('a::executable_euclidean_space)" apply (auto simp: eucl_le[where 'a='a] eucl_le[where 'a='c] dest!: bspec intro!: ) apply (rule cast_in_BasisI, assumption) apply (auto simp: that) apply (metis cast_eqI2 cast_inner that) apply (rule cast_in_BasisI, assumption) apply (auto simp: that) apply (metis cast_eqI2 cast_inner that) done lemma cast_le_cast_iff: "(cast (x::'a)::'c) \ cast y \ x \ y" if "DIM('c::executable_euclidean_space) = DIM('a::executable_euclidean_space)" apply (auto simp: eucl_le[where 'a='a] eucl_le[where 'a='c] dest!: bspec intro!: ) apply (rule cast_in_BasisI, assumption) apply (auto simp: that) apply (rule cast_in_BasisI, assumption) apply (auto simp: that) by (metis cast_eqI2 cast_inner that) lemma cast_image_Icc[simp]: "cast ` {a .. b::'c} = {cast a .. cast b::'a}" if "DIM('c::executable_euclidean_space) = DIM('a::executable_euclidean_space)" using that apply (auto simp: cast_le_iff dest!:) subgoal for x apply (rule image_eqI[where x="cast x"]) by (auto simp: cast_le_iff) done lemma cast_eucl1_image_scaleR2: "cast_eucl1 ` scaleR2 l u X = scaleR2 l u (cast_eucl1 ` (X::'b c1_info set)::'a c1_info set)" if "DIM('a::executable_euclidean_space) = DIM('b::executable_euclidean_space)" using that apply (auto simp: scaleR2_def image_def vimage_def cast_eucl1_def linear.scaleR linear_cast_bl) apply force+ apply (rule exI conjI)+ apply assumption apply (rule exI conjI)+ apply assumption apply (rule bexI) prefer 2 apply assumption apply force by (auto simp: linear.scaleR linear_cast_bl) lemma scaleR2_diff_prod2: "scaleR2 d e (X) - Y \ UNIV = scaleR2 d e (X - Y \ UNIV)" by (force simp: scaleR2_def vimage_def image_def) end lemma (in c1_on_open_euclidean) flowsto_scaleR2I: "flowsto (scaleR2 d e X0) T (CX \ UNIV) (scaleR2 d e Y)" if "flowsto (X0) T (CX \ UNIV) (Y)" using that apply (auto simp: flowsto_def scaleR2_def) apply (drule bspec, assumption) apply auto apply (rule bexI) prefer 2 apply assumption apply auto subgoal for x a b h by (auto intro!: image_eqI[where x="(x, (flow0 a h, Dflow a h o\<^sub>L b))"] blinfun_eqI simp: blinfun.bilinear_simps) done definition "aforms_of_resultrect x0 x1 y0 y1 = aforms_of_ivl (ivl_of_resultrect x0 x1 y0 y1)" definition "flatten_varveq x = fst x @ the_default [] (snd x)" (* LessEqual (N\<^sub>r e) (N\<^sub>r em *\<^sub>e floatarith.Min (norm\<^sub>e [(3)\<^sub>e, (6)\<^sub>e, (9)\<^sub>e]) (norm\<^sub>e [(4)\<^sub>e, (7)\<^sub>e, (10)\<^sub>e])) *) derive "show" ereal definition \::"(real*real*real) set" where "\ = {(-6, -6, 27) .. (6, 6, 27)}" definition \\<^sub>l\<^sub>e::"(real*real*real) set" where "\\<^sub>l\<^sub>e = {(x, y, z). z \ 27}" definition "results = symmetrize coarse_results" definition "results_at x = {res \ set results. x \ source_of_res res}" text \a part of the stable manifold (up to the (backward) first intersection with \\\)\ definition \::"(real*real*real) set" where "\ = {x. {0..} \ lorenz.existence_ivl0 x \ (\t>0. lorenz.flow0 x t \ \) \ (lorenz.flow0 x \ 0) at_top}" definition "\\<^sub>i intr = (if intr then \ else {})" definition "\\<^sub>i\<^sub>v intr = cast ` (\\<^sub>i intr)" definition "sourcei_of_res res = source_of_res res - (\\<^sub>i (invoke_nf res))" definition "resultsi_at x = {res \ set results. x \ sourcei_of_res res}" definition "N = \(source_of_res ` (set results))" definition "\ x = \(conefield_of_res ` (results_at x))" definition "R = lorenz.poincare_map \" definition "DR x = frechet_derivative (lorenz.poincare_map \) (at x within \\<^sub>l\<^sub>e)" definition "\ x = Min (expansion ` results_at x)" definition "\\<^sub>p x = Min (preexpansion ` results_at x)" abbreviation returns_to (infixl "returns'_to" 50) where "(x returns_to P) \ lorenz.returns_to P x" lemma closed_\[intro, simp]: "closed \" by (auto simp: \_def) lemma \_stable: "lorenz.stable_on (- \) \" unfolding lorenz.stable_on_def proof (intro allI impI) fix t x0 assume outside: "\s\{0<..t}. lorenz.flow0 x0 s \ - \" assume assms: "lorenz.flow0 x0 t \ \" "t \ lorenz.existence_ivl0 x0" "0 < t" from assms have *: "{0..} \ lorenz.existence_ivl0 (lorenz.flow0 x0 t)" "(lorenz.flow0 (lorenz.flow0 x0 t) \ 0) at_top" by (auto simp: \_def) have nonneg_exivl: "s \ lorenz.existence_ivl0 x0" if "s \ 0" for s proof (cases "s \ t") case True then show ?thesis using \0 \ s\ assms(2) lorenz.ivl_subset_existence_ivl[of t x0] by auto next case False define u where "u = s - t" with False have "u > 0" "s = t + u" by auto note this(2) also have "t + u \ lorenz.existence_ivl0 x0" apply (rule lorenz.existence_ivl_trans) apply fact using * \u > 0\ by auto finally show ?thesis . qed show "x0 \ \" unfolding \_def proof (safe intro!: nonneg_exivl) have "\\<^sub>F s in at_top. (s::real) \ 0" using eventually_ge_at_top by blast then have "\\<^sub>F s in at_top. lorenz.flow0 (lorenz.flow0 x0 t) s = lorenz.flow0 x0 (s + t)" proof eventually_elim case (elim s) then have "s \ lorenz.existence_ivl0 x0" using nonneg_exivl[OF \0 \ s\] by simp then have "lorenz.flow0 (lorenz.flow0 x0 t) s = lorenz.flow0 x0 (t + s)" apply (subst lorenz.flow_trans) using assms * elim by auto then show ?case by (simp add: ac_simps) qed then have "((\s. (lorenz.flow0 x0 (s + t))) \ 0) at_top" by (blast intro: * Lim_transform_eventually) then show "(lorenz.flow0 x0 \ 0) at_top" unfolding aform.tendsto_at_top_translate_iff . next fix s::real assume s: "0 < s" "lorenz.flow0 x0 s \ \" show False proof (cases "s \ t") case True then show ?thesis using outside s by (auto simp: \_def) next case False then obtain u where u: "u = s - t" "u > 0" by auto then have "lorenz.flow0 x0 (s) = lorenz.flow0 x0 (t + u)" by (simp add: algebra_simps) also have "\ = lorenz.flow0 (lorenz.flow0 x0 t) u" apply (subst lorenz.flow_trans) subgoal by fact subgoal unfolding u apply (rule lorenz.diff_existence_ivl_trans) apply fact+ apply (rule nonneg_exivl) using \0 < s\ by simp subgoal by simp done also from assms(1) \u > 0\ have "\ \ \" by (auto simp: \_def) finally show ?thesis using s by auto qed qed qed lemma (in auto_ll_on_open) stable_on_empty[intro, simp]: "stable_on asdf {}" by (auto simp: stable_on_def) lemma \\<^sub>i_stable: "lorenz.stable_on (- \) (\\<^sub>i b)" using \_stable unfolding \\<^sub>i_def apply (cases b) subgoal by auto subgoal using lorenz.stable_on_empty by (auto simp: \\<^sub>i_def) done definition "\\<^sub>v = (cast ` \)" definition "NF = lorenz.flowsto (cube_entry - \ \ UNIV) {0..} (large_cube \ UNIV) cube_exit" lemma NF0: "lorenz.flowsto ((fst ` cube_entry - \) \ UNIV) {0..} (large_cube \ UNIV) (fst ` cube_exit \ UNIV)" if NF using that unfolding NF_def lorenz.flowsto_def apply (auto simp: NF_def) apply (drule bspec, force) by auto lemma [autoref_rules]: includes autoref_syntax shows "(\_. (), \\<^sub>i\<^sub>v) \ bool_rel \ ghost_rel" by (auto simp: ghost_rel_def) lemma lorenz_interrupt[le, refine_vcg]: "lorenz_interrupt optns b X \ SPEC (\(CX, R). lorenz.flowsto ((cast_eucl1 ` X::R3 c1_info set) - (\\<^sub>i b \ UNIV)) {0..} (cast ` CX \ UNIV) (cast_eucl1 ` R))" if NF unfolding lorenz_interrupt_def apply refine_vcg subgoal by (rule lorenz.flowsto_self) auto subgoal by (auto simp: eucl_le[where 'a="3 rvec"] eucl_of_list_inner Basis_list_vec_def Basis_list_real_def) subgoal for a b c d e f g h i j k l m n p apply (auto) apply (simp add: cast_eucl1_image_scaleR2 scaleR2_diff_prod2) apply (erule make_neg_goal) apply (rule lorenz.flowsto_scaleR2I) using NF0[OF that] apply (rule lorenz.flowsto_subset) subgoal for q apply (auto simp: scaleR2_def cast_eucl1_def) apply (auto simp: linear_cast_bl linear.scaleR cube_entry_def cast_eucl1_def image_image) subgoal premises prems for r s t u v proof - from prems have "fst ` c \ fst ` (cube_enter::3 vec1 set)" by auto with \(u, v) \ c\ obtain w where "((u, w)::3 vec1) \ cube_enter" by auto from _ this have "cast u \ (\x. cast (fst (x::3 vec1))) ` cube_enter" by (rule image_eqI) auto then show ?thesis using prems by blast qed subgoal by (auto simp: \\<^sub>i_def) done subgoal by simp subgoal premises _ by (auto simp: eucl_le[where 'a="3 rvec"] eucl_of_list_inner Basis_list_vec_def Basis_list_real_def eucl_of_list_def large_cube_def) subgoal premises prems for x supply [[simproc del: defined_all]] apply (auto simp:) subgoal for A B C D E apply (rule image_eqI[where x="cast_eucl1 (((B, C, D), A))"]) apply (auto simp: cast_eucl1_def) subgoal premises prems using prems(1) apply (auto simp: cube_exit_def aform.c1_info_of_apprs_def cube_exiti_def cast_eucl1_def aform.c1_info_of_appr_def) done done done done subgoal for a b c d e f g h i j k l m n p apply (auto) apply (erule make_neg_goal, thin_tac "\ _") apply (simp add: cast_eucl1_image_scaleR2 scaleR2_diff_prod2) apply (rule lorenz.flowsto_scaleR2I) using that[unfolded NF_def] apply (rule lorenz.flowsto_subset) subgoal for q apply (auto simp: scaleR2_def cast_eucl1_def ) apply (auto simp: linear_cast_bl linear.scaleR cube_entry_def cast_eucl1_def image_image) subgoal premises prems for r s t u v proof - from prems have \vec1_of_flow1 (u, v) \ cube_enter\ by auto from _ this have "(cast u, cast_bl v) \ (\x. (cast (fst (x::3 vec1)), cast_bl (snd (flow1_of_vec1 x)))) ` cube_enter" by (rule image_eqI) (auto simp: ) then show ?thesis using prems by blast qed subgoal by (auto simp: \\<^sub>i_def) done subgoal by simp subgoal premises _ by (auto simp: eucl_le[where 'a="3 rvec"] eucl_of_list_inner Basis_list_vec_def Basis_list_real_def eucl_of_list_def large_cube_def) subgoal by (simp add: cube_exit_def) done done definition "lorenz_S X = (case X of (x, y, z) \ (-x, -y, z))" lemma lorenz_symI: "((\t. lorenz_S (f t)) has_vderiv_on lf') T" if "(f has_vderiv_on f') T" "\t. t \ T \ lf' t = lorenz_S (f' t)" using that by (auto simp: has_vderiv_on_def lorenz_S_def split_beta' has_vector_derivative_def intro!: derivative_eq_intros) lemma lorenz_S: "t \ lorenz.existence_ivl0 (lorenz_S X)" (is ?th1) "lorenz.flow0 (lorenz_S X) t = lorenz_S (lorenz.flow0 X t)" (is ?th2) if "t \ lorenz.existence_ivl0 X" proof - have 1: "((\t. lorenz_S (lorenz.flow0 X t)) solves_ode (\_ (X0, X1, X2). (E1 * X0 - K1 * (X0 + X1) * X2, E2 * X1 + K1 * (X0 + X1) * X2, E3 * X2 + (X0 + X1) * (K2 * X0 + K3 * X1)))) {0--t} UNIV" apply (rule solves_odeI) apply (rule lorenz_symI) apply (rule lorenz.flow_has_vderiv_on_compose) apply simp apply simp apply (rule derivative_intros) apply (rule refl) using that apply (rule lorenz.in_existence_between_zeroI) apply assumption apply (rule refl) unfolding lorenz_S_def apply (split prod.splits)+ apply (simp add: field_simps) apply simp done have "lorenz.flow0 X 0 = X" unfolding lorenz.flow_initial_time_if by simp then have 2: "lorenz_S (lorenz.flow0 X 0) = lorenz_S X" "is_interval {0--t}" "0 \ {0--t}" "{0--t} \ UNIV" by auto from lorenz.maximal_existence_flow[OF 1 2] show ?th1 ?th2 by fast+ qed lemma \\<^sub>l\<^sub>e_impl[autoref_rules]: "(Sctn [0, 0, 1] 27, \\<^sub>l\<^sub>e) \ \lv_rel\below_rel" apply (auto simp: below_rel_def \\<^sub>l\<^sub>e_def below_halfspace_def sctn_rel_def intro!: relcompI[where b="Sctn (0, 0, 1) 27"] brI lv_relI) subgoal unfolding lv_rel_def by (auto intro!: brI) unfolding le_halfspace_def by (auto intro!: brI) lemma [autoref_rules]: "((), \\<^sub>v) \ ghost_rel" by (auto intro!: ghost_relI) no_notation vec_nth (infixl "$" 90) and vec_lambda (binder "\" 10) abbreviation "guards_rel \ \clw_rel (\\lv_rel\ivl_rel, \lv_rel\plane_rel\inter_rel) \\<^sub>r aform.reach_optns_rel\list_rel" definition "aform_poincare_onto_from optns = aform.poincare_onto_from" lemma aform_poincare_onto_from[autoref_rules]: includes autoref_syntax shows "DIM_precond TYPE('b rvec) E \ (XSi, XS::'b::enum eucl1 set) \ clw_rel aform.appr1e_rel \ (sctni, sctn) \ \lv_rel\sctn_rel \ (ivli, ivl) \ \lv_rel\ivl_rel \ (Si, Sa) \ \lv_rel\halfspaces_rel \ (guardsi, guards) \ guards_rel \ (symstartd, symstart) \ aform.appr1e_rel \ \clw_rel aform.appr_rel \\<^sub>r clw_rel aform.appr1e_rel\dres_nres_rel \ ((), trap) \ ghost_rel \ (roi, roptn) \ aform.reach_optns_rel \ (odoi, odo) \ ode_ops_rel \ (optnsi, optns) \ num_optns_rel \ (nres_of (solve_poincare_map_aform optnsi E odoi symstartd Si guardsi ivli sctni roi XSi), aform_poincare_onto_from $ optns $ odo $ symstart $ trap $ Sa $ guards $ ivl $ sctn $ roptn $ XS) \ \clw_rel aform.appr1e_rel\nres_rel" unfolding autoref_tag_defs aform_poincare_onto_from_def using aform.poincare_onto_from_impl.refine[OF _ aform_ncc aform_ncc, where 'a='b, of E odoi odo XSi XS Si Sa guardsi guards ivli ivl sctni sctn roi roptn "(\x. nres_of (symstartd x))" symstart symstartd trap optnsi, unfolded autoref_tag_defs, OF _ _ _ _ _ _ _ _ _ order_refl] by (auto simp: dest!: aform.dres_nres_rel_nres_relD) definition "lorenz_odo_impl = init_ode_ops True True lorenz.odo" interpretation autoref_op_pat_def lorenz.odo . lemma lorenz_odo_impl[autoref_rules]: "(lorenz_odo_impl, lorenz.odo) \ ode_ops_rel" by (auto simp: ode_ops_rel_def lorenz_odo_impl_def) definition lorenz_poincare where "lorenz_poincare optns interrupt guards roptn XS0 = aform_poincare_onto_from optns lorenz.odo (lorenz_interrupt optns interrupt) (\\<^sub>i\<^sub>v interrupt:::ghost_rel) ((below_halfspaces {Sctn (eucl_of_list [0, 0, 1]) 27}::(real^3) set):::\lv_rel\halfspaces_rel) guards (op_atLeastAtMost_ivl (eucl_of_list [-6, -6, 27]:::lv_rel) (eucl_of_list [6, 6, 27]:::lv_rel):::lvivl_rel::(real^3) set) (Sctn (eucl_of_list [0, 0, -1]) (- 27)::(real^3) sctn) roptn XS0" lemma [autoref_rules_raw]: includes autoref_syntax shows "((), (OP \\<^sub>i\<^sub>v ::: bool_rel \ ghost_rel) $ (OP intr ::: bool_rel)) \ ghost_rel" by (auto simp: ghost_rel_def) schematic_goal lorenz_poincare_impl[autoref_rules]: includes autoref_syntax assumes [autoref_rules]: "(XSi, XS) \ clw_rel aform.appr1e_rel" "(intri, intr) \ bool_rel" "(guardsi, guards) \ guards_rel" "(roi, roptn) \ aform.reach_optns_rel" "(optnsi, optns) \ num_optns_rel" shows "(nres_of ?r, lorenz_poincare $ optns $ intr $ guards $ roptn $ XS) \ \clw_rel aform.appr1e_rel\nres_rel" unfolding autoref_tag_defs unfolding lorenz_poincare_def including art supply [autoref_rules_raw] = ghost_relI by autoref_monadic lemma cast_image_eqI: "cast ` X = Y" if "DIM('a) = DIM('b)" "(X::'a::executable_euclidean_space set) = cast ` (Y::'b::executable_euclidean_space set)" using that by (auto simp: image_image) lemma transfer_\[transfer_rule]: "(rel_set lorenz.rel_ve) \\<^sub>v \" unfolding \\<^sub>v_def by (auto simp: lorenz.rel_ve_cast' intro!: rel_setI) lemma transfer_\\<^sub>l\<^sub>e[transfer_rule]: "(rel_set lorenz.rel_ve) (cast ` \\<^sub>l\<^sub>e) \\<^sub>l\<^sub>e" by (auto simp: lorenz.rel_ve_cast' intro!: rel_setI) lemma transfer_\\<^sub>i[transfer_rule]: "(rel_fun (=) (rel_set lorenz.rel_ve)) \\<^sub>i\<^sub>v \\<^sub>i" unfolding \\<^sub>i\<^sub>v_def by (auto simp: lorenz.rel_ve_cast' intro!: rel_setI) lemma transfer_\[transfer_rule]: "(rel_set lorenz.rel_ve) (cast ` \) \" by (auto simp: lorenz.rel_ve_cast' intro!: rel_setI) lemma len_fas: "length lorenz_fas = 3" by (auto simp: lorenz_fas_def) lemma lorenz_poincare[le, refine_vcg]: "lorenz_poincare optns intr guards roptn XS \ SPEC (\R. aform.poincare_mapsto lorenz.odo (cast ` \) (XS - (\\<^sub>i\<^sub>v intr \ UNIV)) (cast ` \\<^sub>l\<^sub>e) UNIV R)" if [refine_vcg]: NF unfolding lorenz_poincare_def aform_poincare_onto_from_def apply (refine_vcg) subgoal by (simp add: aform.wd_def aform.ode_e_def len_fas) subgoal for a b c d apply (auto simp: lorenz.flowsto_eq[symmetric]) proof goal_cases case 1 from 1(2)[unfolded lorenz.flowsto_eq[symmetric]] show ?case by transfer (simp add: lorenz.avflowsto_eq) qed subgoal unfolding aform.stable_on_def unfolding autoref_tag_defs proof (intro allI impI, goal_cases) case (1 t x0) from 1 have t: "t \ lorenz.v.existence_ivl0 x0" using lorenz.vex_ivl_eq by simp from 1 have f: "lorenz.v.flow0 x0 t \ \\<^sub>i\<^sub>v intr" using lorenz.vflow_eq[OF t] by simp from 1 have "lorenz.v.flow0 x0 s \ - cast ` \" if "s\{0<..t}" for s proof - from that t have s: "s \ lorenz.v.existence_ivl0 x0" by (auto dest!: lorenz.a.v.closed_segment_subset_existence_ivl simp: closed_segment_eq_real_ivl) have "lorenz.v.flow0 x0 s \ aform.Csafe lorenz.odo - op_atLeastAtMost_ivl (eucl_of_list [- 6, - 6, 27]) (eucl_of_list [6, 6, 27]) \ plane_of (Sctn (eucl_of_list [0, 0, - 1]) (- 27))" using 1(4)[rule_format, OF that] unfolding lorenz.vflow_eq[OF s] by auto also have "\ \ - cast ` \" by (auto simp: eucl_le[where 'a="real^3"] eucl_of_list_inner axis_eq_axis cast_def Basis_list_real_def Basis_list_vec3 \_def plane_of_def eucl_of_list_inner_eq inner_lv_rel_def) finally show "lorenz.v.flow0 x0 s \ - cast ` \" . qed show "x0 \ \\<^sub>i\<^sub>v intr" by (rule \\<^sub>i_stable[Transfer.untransferred, unfolded lorenz.v.stable_on_def, rule_format]) fact+ qed subgoal for R proof (clarsimp, goal_cases) case 1 note 1(2) also have "({eucl_of_list [- 6, - 6, 27]..eucl_of_list [6, 6, 27]::3 rvec} \ plane_of (Sctn (eucl_of_list [0, 0, - 1]) (- 27))) = cast ` \" apply auto apply (auto simp: \_def o_def plane_of_def eucl_of_list_def Basis_list_R3 Basis_list_vec3) subgoal by (auto simp: cast_def eucl_of_list_def Basis_list_R3 Basis_list_vec3) subgoal by (auto simp: cast_def eucl_of_list_def Basis_list_R3 Basis_list_vec3) subgoal apply (auto simp: cast_le_iff[symmetric]) by (auto simp: cast_def eucl_of_list_def Basis_list_R3 Basis_list_vec3 less_eq_prod_def list_of_eucl_def inner_simps inner_axis_axis) subgoal apply (auto simp: cast_le_iff) by (auto simp: cast_def eucl_of_list_def Basis_list_R3 Basis_list_vec3 less_eq_prod_def list_of_eucl_def inner_simps inner_axis_axis) subgoal by (auto simp: cast_def eucl_of_list_def Basis_list_R3 Basis_list_vec3 less_eq_prod_def list_of_eucl_def inner_simps inner_axis_axis) done also have "(below_halfspaces {Sctn (eucl_of_list [0, 0, 1]::3 rvec) 27}) = (cast ` \\<^sub>l\<^sub>e)" apply (auto simp: \\<^sub>l\<^sub>e_def below_halfspaces_def below_halfspace_def le_halfspace_def eucl_of_list_def Basis_list_R3 Basis_list_vec3) apply (rule image_eqI) apply (rule cast_cast[symmetric]) by (auto simp: cast_def list_of_eucl_def o_def plane_of_def inner_simps inner_axis_axis eucl_of_list_def Basis_list_R3 Basis_list_vec3) finally show ?case by (rule aform.poincare_mapsto_subset) (force simp: lorenz.vX_def intro: cast_cast[symmetric])+ qed done context includes floatarith_notation begin definition "mat1\<^sub>e = [Var 0, Var 1, Var 2, Var 3, 0, 0, Var 4, 0, 0, Var 5, 0, 0]" definition mat1_nres::"3 rvec set \ 3 rvec set \ 3 eucl1 set nres" where "mat1_nres X v = do { Xv \ aform.approx_slp_appr mat1\<^sub>e (slp_of_fas mat1\<^sub>e) (concat ` listset [list_of_eucl ` X, list_of_eucl ` v]); RETURN (flow1_of_vec1 ` Xv) }" lemma [simp]: "(x, x') \ aform.appr_rel \ aform.ncc x'" using aform_ncc[where 'a='a] by (auto simp: aform.ncc_precond_def) lemma mat1e_autoref[autoref_rules]: "(mat1\<^sub>e, mat1\<^sub>e) \ \Id\list_rel" by auto schematic_goal mat1_impl: includes autoref_syntax assumes [autoref_rules]: "(Xi, X) \ aform.appr_rel" "(vi, v) \ aform.appr_rel" shows "(nres_of ?r, mat1_nres $ X $ v) \ \aform.appr1_rel\nres_rel" unfolding mat1_nres_def including art by autoref_monadic concrete_definition mat1_impl for Xi vi uses mat1_impl lemmas [autoref_rules] = mat1_impl.refine lemma mat_nres[le, refine_vcg]: "mat1_nres X v \ SPEC (\M. X \ v \ (\x. (fst x, blinfun_apply (snd x) (eucl_of_list [1, 0, 0]))) ` M)" unfolding mat1_nres_def apply refine_vcg apply (auto simp: dest!: bspec) apply (auto simp: mat1\<^sub>e_def image_image ) subgoal for x a b apply (rule image_eqI[where x="eucl_of_list [(list_of_eucl a @ list_of_eucl b) ! 0, (list_of_eucl a @ list_of_eucl b) ! Suc 0, (list_of_eucl a @ list_of_eucl b) ! 2, (list_of_eucl a @ list_of_eucl b) ! 3, 0, 0, (list_of_eucl a @ list_of_eucl b) ! 4, 0, 0, (list_of_eucl a @ list_of_eucl b) ! 5, 0, 0]"]) apply (auto simp: eucl_of_list_prod eucl_of_list_inner nth_append Basis_list_vec3 intro!: euclidean_eqI[where 'a="3 rvec"]) unfolding Basis_list[symmetric] Basis_list_vec3 by (auto simp: flow1_of_vec1_def blinfun_of_vmatrix.rep_eq Basis_list_vec3 inner_simps matrix_vector_mult_eq_list_of_eucl_nth ll3 inner_axis_axis) done definition [simp]: "op_image_cast_eucl1e = (`) cast_eucl1" definition [simp]: "op_image_cast_eucl1e_coll = (`) cast_eucl1" lemma prod_relI'': "\(fst ab, a')\R1; (snd ab, b')\R2\ \ (ab,(a', b'))\\R1,R2\prod_rel" by (auto simp: prod_rel_def) lemma strange_aux_lemma: "(b, b') \ A \ (b, snd (a'a, b')) \ A" by auto lemma [autoref_rules]: includes autoref_syntax assumes "DIM_precond TYPE('a::executable_euclidean_space) D" "DIM_precond TYPE('b::executable_euclidean_space) D" shows "(\x. x, (op_image_cast_eucl1e::('a::executable_euclidean_space c1_info set \ 'b::executable_euclidean_space c1_info set))) \ aform.appr1e_rel \ aform.appr1e_rel" (is ?th1) and "(\x. x, op_image_cast_eucl1e_coll::'a::executable_euclidean_space c1_info set \ 'b::executable_euclidean_space c1_info set) \ clw_rel aform.appr1e_rel \ clw_rel aform.appr1e_rel" (is ?th2) proof - show 1: ?th1 unfolding scaleR2_rel_def apply (rule subsetD) apply (rule fun_rel_comp_dist) apply (rule relcompI) apply (rule fun_relI) apply (erule prod_relE) apply simp apply (rule prod_relI) apply simp apply (rule fst_conv[symmetric]) apply (rule op_cast_eucl1_image_impl[OF assms, param_fo]) apply (rule strange_aux_lemma) apply (auto simp: br_def scaleR2_def image_def vimage_def cast_eucl1_def) subgoal for a b c d e f g apply (rule exI[where x=e] conjI)+ apply assumption apply (rule exI conjI)+ apply assumption apply (rule exI conjI)+ apply (rule bexI) prefer 2 apply assumption apply (rule conjI) apply force apply (rule refl) using assms by (auto simp: blinfun.bilinear_simps linear_cast linear.scaleR intro!: blinfun_eqI) subgoal for a b c d e f g apply (rule exI[where x=e] exI conjI)+ apply assumption apply (rule exI conjI)+ apply assumption apply (rule bexI) prefer 2 apply assumption apply force using assms by (auto simp: blinfun.bilinear_simps linear_cast linear.scaleR intro!: blinfun_eqI) done have id_map: "(\x. x) = (map (\x. x))" by simp show ?th2 apply (subst id_map) apply (rule lift_clw_rel_map) apply (rule relator_props)+ subgoal using 1 by auto subgoal by auto done qed definition "lorenz_poincare_tangents optns interrupt guards roptn c1 (X0::R3 set) (tangents::R3 set) = do { X0tanmat \ (if c1 then do { R \ mat1_nres (cast ` X0) (cast ` tangents); RETURN (R::3 eucl1 set) } else RETURN (cast ` X0 \ UNIV)); XDX0 \ scaleRe_ivl_spec 1 1 (X0tanmat); let _ = aform.trace_set1e ''START'' (Some XDX0); let _ = aform.print_set1e False (XDX0); P \ lorenz_poincare optns interrupt guards roptn ((mk_coll XDX0:::clw_rel aform.appr1e_rel)); RETURN (op_image_cast_eucl1e_coll P::R3 c1_info set) }" lemma [autoref_rules_raw]: "DIM(real \ real \ real) = DIM((real, 3) vec)" by auto schematic_goal lorenz_poincare_tangents_impl: includes autoref_syntax assumes [autoref_rules]: "(optnsi, optns) \ Id" "(intrri, intr) \ bool_rel" "(guardsi, guards) \ guards_rel" "(roi, roptn) \ aform.reach_optns_rel" "(c1i, c1) \ bool_rel" "(X0i, X0) \ aform.appr_rel" "(tangentsi, tangents) \ aform.appr_rel" shows "(nres_of ?r, lorenz_poincare_tangents $ optns $ intr $ guards $ roptn $ c1 $ (X0::R3 set) $ tangents) \ \clw_rel aform.appr1e_rel\nres_rel" unfolding lorenz_poincare_tangents_def including art by (autoref_monadic) concrete_definition lorenz_poincare_tangents_impl uses lorenz_poincare_tangents_impl[where optnsa = "\_ _ _ _ _ _ _ _. optns" and optnsb = "\_ _ _ _ _ _ _ _ _. optns" and optnsi = "optns" and optnsc = "optns" and optns = "\_ _ _ _ _ _ _. optns" for optns optnsc] lemma lorenz_poincare_tangents_impl_refine[autoref_rules]: includes autoref_syntax shows "(\optnsi intrri guardsi roi c1i X0i tangentsi. nres_of (lorenz_poincare_tangents_impl optnsi intrri guardsi roi c1i X0i tangentsi), lorenz_poincare_tangents) \ num_optns_rel \ bool_rel \ guards_rel \ aform.reach_optns_rel \ bool_rel \ aform.appr_rel \ aform.appr_rel \ \clw_rel aform.appr1e_rel\nres_rel" using lorenz_poincare_tangents_impl.refine by force lemma transfer_UNIV_rel_blinfun[transfer_rule]: "(rel_set (rel_blinfun lorenz.rel_ve lorenz.rel_ve)) UNIV UNIV" apply (auto intro!: rel_setI simp: rel_blinfun_def) subgoal for x apply (rule exI[where x="cast_bl x"]) by (auto intro!: rel_funI simp: lorenz.rel_ve_cast) subgoal for x apply (rule exI[where x="cast_bl x"]) by (auto intro!: rel_funI simp: lorenz.rel_ve_cast) done lemma lorenz_vX[simp]: "lorenz.vX = (UNIV::3 rvec set)" by (force simp: lorenz.vX_def intro!: cast_cast[symmetric]) lemma closed_cast_\[intro, simp]: "closed (cast ` \::3 rvec set)" by (auto simp: \_def ) lemma blinfun_apply_transfer[transfer_rule]: "(rel_fun (rel_blinfun lorenz.rel_ve lorenz.rel_ve) (rel_fun (rel_prod (=) (rel_prod (=) (=))) lorenz.rel_ve)) (blinfun_apply o cast_bl) blinfun_apply" by (auto intro!: rel_funI simp: rel_blinfun_def lorenz.rel_ve_cast dest!: rel_funD) lemma lorenz_poincare_tangents[le, refine_vcg]: "lorenz_poincare_tangents optns intr guards roptn c1 (X0::R3 set) tangents \ SPEC (\x. (if c1 then \tans. X0 \ tangents \ (\(x, y). (x, blinfun_apply y (1, 0, 0))) ` tans \ lorenz.poincare_mapsto \ (tans - \\<^sub>i intr \ UNIV) (\\<^sub>l\<^sub>e) UNIV x else lorenz.poincare_mapsto \ ((X0 - \\<^sub>i intr) \ UNIV) (\\<^sub>l\<^sub>e) UNIV x))" if [refine_vcg]: NF unfolding lorenz_poincare_tangents_def apply refine_vcg apply auto apply (subst lorenz.poincare_mapsto_eq[symmetric]) apply simp proof goal_cases case (2 R) then show ?case apply transfer apply (subst lorenz.vpoincare_mapsto_eq[symmetric]) apply (auto simp: ) apply (rule aform.poincare_mapsto_subset, assumption) by (force simp: scaleR2_def )+ next case (1 tans R) show ?case apply (rule exI[where x="cast_eucl1 ` tans"]) apply (rule conjI) subgoal including lifting_syntax using 1(2) by transfer (force simp: cast_def o_def) subgoal using 1(3) apply transfer apply (subst lorenz.avpoincare_mapsto_eq[symmetric]) by (auto simp: ) done qed definition of_mat1_image::"R3 c1_info set \ R3 set nres" where [refine_vcg_def]: "of_mat1_image X = SPEC (\R. R = (\x. blinfun_apply (snd x) (1, 0, 0)) ` X)" lemma of_mat1_image_impl[autoref_rules]: "(\x. (case x of (_, Some xs) \ RETURN [xs ! 0, xs ! 3, xs ! 6] | (_, None) \ SUCCEED), of_mat1_image) \ aform.appr1_rel \ \aform.appr_rel\nres_rel" apply (auto simp: of_mat1_image_def RETURN_RES_refine_iff nres_rel_def aform.appr1_rel_internal aform.appr_rel_def intro!: relcompI split: option.splits) unfolding aforms_rel_def apply (rule brI) apply (auto simp: ) unfolding lv_rel_def set_rel_br apply (rule brI) prefer 2 apply (force simp: Joints_imp_length_eq) apply (auto elim!: mem_Joints_appendE simp: flow1_of_list_def Joints_imp_length_eq) subgoal for a b c d e f g h i j apply (rule image_eqI[where x="[j ! 0, j ! 3, j! 6]"]) apply (auto simp: blinfun_of_list_def blinfun_of_matrix_apply Basis_prod_def Basis_list_R3 Basis_list_vec3 eval_nat_numeral zero_prod_def) apply (rule map_nth_Joints'[of _ _ "[0, Suc (Suc (Suc 0)), Suc (Suc (Suc (Suc (Suc (Suc 0)))))]", simplified]) apply auto done subgoal for a b c d e f unfolding image_image apply (auto simp: Joints_def valuate_def) subgoal for g apply (rule image_eqI) prefer 2 apply (rule image_eqI[where x=g]) apply (rule refl) apply (auto simp: ) apply (auto simp: blinfun_of_list_def blinfun_of_matrix_apply flow1_of_list_def Basis_prod_def Basis_list_R3 Basis_list_vec3 eval_nat_numeral zero_prod_def) done done done definition [refine_vcg_def]: "floatdegs res = SPEC (\_::unit. min_deg res \ float \ max_deg res \ float)" definition [simp]: "isinfloat x \ x \ float" lemma [code]: "isinfloat (real_of_float x) = True" by (auto) lemma floatdegs_impl[autoref_rules]: includes autoref_syntax shows "(\res. (if isinfloat (min_deg res) \ isinfloat (max_deg res) then RETURN () else SUCCEED), floatdegs) \ Id \ \unit_rel\nres_rel" by (auto simp: nres_rel_def floatdegs_def) definition "check_c1_entry optns em P (res0::result) (res::result) = do { uv_ret \ of_mat1_image P; nuv \ aform.mig_set 3 uv_ret; floatdegs res0; floatdegs res; let e' = em * ereal nuv; b1 \ approx_conefield_bounds P (min_deg res) (max_deg res); let b2 = e' \ preexpansion res; let b3 = e' \ expansion res0; let _ = aform.print_msg ((shows ''# check_c1_entry: '' o shows_list [b1, b2, b3] o shows_space o shows_list [e', preexpansion res, expansion res0]) ''''); RETURN (em \ 0 \ b1 \ b2 \ b3) }" lemma [autoref_itype]: "shows_prec ::\<^sub>i i_nat \\<^sub>i A \\<^sub>i i_string \\<^sub>i i_string" by auto lemma [autoref_rules]: includes autoref_syntax shows "PREFER_id A \ (shows_list, shows_list) \ \A\list_rel \ string_rel \ string_rel" "(shows_prec, shows_prec) \ nat_rel \ string_rel \ string_rel \ string_rel" "(shows_prec, shows_prec) \ nat_rel \ ereal_rel \ string_rel \ string_rel" "(shows_prec, shows_prec::_\result \_) \ nat_rel \ Id \ string_rel \ string_rel" "(shows_space, shows_space) \ string_rel \ string_rel" by (auto simp: string_rel_def) lemma [autoref_rules]: includes autoref_syntax shows "(expansion, expansion) \ Id \ rnv_rel" "(preexpansion, preexpansion) \ Id \ rnv_rel" "(min_deg, min_deg) \ Id \ rnv_rel" "(max_deg, max_deg) \ Id \ rnv_rel" by auto interpretation autoref_op_pat_def aform.mig_set . lemma [autoref_rules_raw]: "DIM_precond TYPE(real \ real \ real) (OP 3 ::: nat_rel)" by simp schematic_goal check_c1_entry_impl: includes autoref_syntax assumes [autoref_rules]: "(optnsi, optns) \ Id" "(res0i, res0) \ Id" "(resi, res) \ Id" "(emi, em) \ ereal_rel" "(Pei, P) \ aform.appr1_rel" shows "(nres_of ?r, check_c1_entry optns em P res0 res) \ \bool_rel\nres_rel" unfolding check_c1_entry_def including art by autoref_monadic concrete_definition check_c1_entry_impl uses check_c1_entry_impl[ where optns = "\_ . optnsi" and optnsi="optnsi" and optnsc=optns and optnsa="\_ _ _ _ _. optnsi" and optnsb="\_ _ _ _ _ _ _ _ . optnsi" and optns="\_. optnsi" for optns optnsi] lemmas check_c1_entry_impl_refine[autoref_rules] = check_c1_entry_impl.refine[autoref_higher_order_rule] definition "c1_entry_correct (em::ereal) (P::R3 c1_info set) res0 res = (\(_, d)\P. case (d (1, 0, 0)) of (dx, dy, dz) \ dz = 0 \ dx > 0 \ -90 < min_deg res \ min_deg res \ max_deg res \ max_deg res < 90 \ ereal (preexpansion res) \ em * (norm (dx, dy, dz)) \ ereal (expansion res0) \ em * (norm (dx, dy, dz)) \ dy / dx \ {tan (rad_of (min_deg res)) .. tan (rad_of (max_deg res))})" lemma check_c1_entry[le, refine_vcg]: "check_c1_entry optns em P res0 res \ SPEC (\b. b \ c1_entry_correct em P res0 res)" unfolding check_c1_entry_def c1_entry_correct_def apply refine_vcg apply (auto dest!: bspec simp:) apply (rule order_trans, assumption, rule ereal_mult_left_mono, force, force) apply (rule order_trans, assumption, rule ereal_mult_left_mono, force, force) done subsection \options for the lorenz system\ definition aform_numeric_optns::"_ \ _ \ _ \ _ \ _ \ _ \ _ \ _ \ _ \ _ \ _ \ _ \ real aform numeric_options" where "aform_numeric_optns = numeric_options.fields" fun zbucket::"real \ real \ real \ real \ real \ real \ real \ ((real list \ real list) \ real list sctn) list" where "zbucket d (x0,x1) (y0, y1) (z0, z1) = [zsec' (x0 - d, x0 + d) (y0 - d, y1 + d) z0, \ \bottom\ xsec' x0 (y0 - d, y1 + d) (z0 - d, z1), \ \left\ xsec x1 (y0 - d, y1 + d) (z0 - d, z1), \ \right\ ysec' (x0 - d, x1 + d) y0 (z0 - d, z1), \ \backno\ ysec (x0 - d, x1 + d) y1 (z0 - d, z1)] \ \front\" subsubsection \Hybridizations\ definition "reduce_weak_params c1 = (if c1 then (12::nat, 0::nat) else (3, 0))" definition "reduce_hard_params c1 = (if c1 then (0::nat, 100::nat) else (0, 100))" definition "ro_split_weak c1 w = (case reduce_weak_params c1 of (m, n) \ ro (w + 1) m n w w (-5))" definition "ro_split_weak' c1 w = (case reduce_weak_params c1 of (m, n) \ ro w m n w w (-5))" definition "ro_split_weak'' c1 w = (case reduce_weak_params c1 of (m, n) \ ro (w + 2) m n w w (-5))" definition "ro_split_weak4' c1 w = (case reduce_weak_params c1 of (m, n) \ ro (w + 4) m n w w (-5))" definition "ro_split_weak2 c1 w w2 = (case reduce_weak_params c1 of (m, n) \ ro (w + 1) m n w w2 (-5))" definition "ro_split_weak2' c1 w w2 = (case reduce_weak_params c1 of (m, n) \ ro (w) m n w w2 (-5))" definition "ro_split_hard c1 w0 w1 = (case reduce_hard_params c1 of (m, n) \ ro (w0 + 1) m n w0 w1 (-5))" definition "ro_split_hard'' c1 w0 w1 = (case reduce_hard_params c1 of (m, n) \ ro (w0 + 2) m n w0 w1 (-5))" definition "ro_split_not c1 w = (case reduce_weak_params c1 of (m, n) \ ro 0 m n w w (-5))" definition "ro_split_not2 c1 w w2 = (case reduce_weak_params c1 of (m, n) \ ro 0 m n w w2 (-5))" definition "xsecs x y z = [xsec' (-x) (-y, y) (0, z), xsec x (-y, y) (0, z)]" type_synonym run_options = "(nat \ nat) \ int \ (((real list \ real list) \ real list sctn) list \ real aform reach_options) list \ real aform reach_options" abbreviation "p1 \ ldec 0.1" definition mode_middle::"_ \ run_options" where "mode_middle c1 = (reduce_weak_params c1, -14, [([zsec' (-2, 2) (-1, 1) 10], ro_split_weak' c1 (-3)), (xsecs (5 * p1) 10 10 @ xsecs p1 10 (6) @ [zsec' (-p1, p1) (-p1, p1) p1], ro_split_hard c1 (-5) (-2)), (xsecs (3/2/2) 10 (10), ro_split_not2 c1 0 (-2)), \ \To collect after interrupt\ ([zsec (-30, -6) (-10, 10) 30, zsec (6, 30) (-10, 10) 30], ro_split_not2 c1 (-1) (-3)) ], ro_split_weak4' c1 (-5))" definition mode_inner3::"bool\bool\run_options" where "mode_inner3 c1 very_inner = (reduce_weak_params c1, -15, (if very_inner then [([zsec' (-2, 2) (-1, 1) 10], ro_split_weak' c1 (-2))] else [])@ [(xsecs (3/2) 15 27@xsecs 1 (10) (11/2), ro_split_weak2 c1 (-2) (-1)), ([ zsec (-30, -6) (-10, 10) 30, zsec (6, 30) (-10, 10) 30], ro_split_not2 c1 (-1) (-3)) ], if very_inner then ro_split_weak4' c1 (-5) else ro_split_weak'' c1 (-5))" definition mode_inner2::"bool \ real \ run_options" where "mode_inner2 c1 x = (reduce_weak_params c1, -14, [(xsecs x 10 27, ro_split_weak2' c1 (-2) (-1)), ([zsec ( -30, -6) (-10, 10) 30, zsec (6, 30) (-10, 10) 30], ro_split_not2 c1 (-3) (-3))], ro_split_not c1 (-6))" definition "ro_outer c1 w = (case reduce_weak_params c1 of (m, n) \ ro w m n (-6) (-6) (-5))" definition mode_outer::"bool\_\_\run_options" where "mode_outer c1 w i = (reduce_weak_params c1, (-i), [([zsec (-30, -6) (-10, 10) 27, zsec (6, 30) (-10, 10) 27], ro_split_not2 c1 (-3) (-4))], ro_outer c1 w)" definition lookup_mode::"bool \ result \ _" where "lookup_mode c1 i = (if gridx0 i \ - 1024 then mode_outer c1 (-3) 16 else if gridx0 i \ - 120 then mode_outer c1 (-3) 14 else if gridx0 i \ 107 then mode_inner2 c1 (4) else if gridx0 i \ 169 then mode_inner3 c1 False else if gridx0 i \ 196 then mode_inner3 c1 True else if gridx0 i \ 201 then mode_middle c1 else if gridx0 i \ 235 then mode_inner3 c1 True else if gridx0 i \ 290 then mode_inner3 c1 False else if gridx0 i \ 450 then mode_inner2 c1 4 else mode_outer c1 (-3) 14)" definition mode_ro_spec::"bool \ result \ ((nat \ nat) \ int \ ((real, 3) vec set \ unit) list \ unit) nres" where [refine_vcg_def]: "mode_ro_spec c1 res = SPEC (\_. True)" lemma reach_options_rel_br: "reach_options_rel TYPE('ty) = br (\_. ()) (\_. True)" by (auto simp: reach_options_rel_def br_def) lemma mode_ro_spec_impl[autoref_rules]: includes autoref_syntax shows "(\b x. RETURN (lookup_mode b x), mode_ro_spec) \ bool_rel \ Id \ \(nat_rel \\<^sub>r nat_rel) \\<^sub>r int_rel \\<^sub>r guards_rel \\<^sub>r aform.reach_optns_rel\nres_rel" supply [simp del] = prod_rel_id_simp apply (rule fun_relI) apply (rule fun_relI) apply (rule nres_relI) unfolding mode_ro_spec_def apply (rule RETURN_SPEC_refine) apply (auto simp: mode_ro_spec_def nres_rel_def RETURN_RES_refine_iff) apply (rule exI)+ apply (rule prod_relI'' IdI)+ unfolding lv_rel_def ivl_rel_def br_rel_prod br_chain plane_rel_br inter_rel_br clw_rel_br br_list_rel Id_br prod_eq_iff reach_options_rel_br apply (rule brI refl)+ defer apply (rule brI) apply (rule refl) apply auto apply (auto simp: lookup_mode_def mode_outer_def mode_inner2_def mode_inner3_def xsecs_def mode_middle_def) done lemma [autoref_rules]: includes autoref_syntax shows "(ivl_of_res, ivl_of_res) \ Id \ \rnv_rel\list_rel \\<^sub>r \rnv_rel\list_rel" by auto lemma [autoref_rules]: includes autoref_syntax shows "(Polygon.pairself, Polygon.pairself) \ (A \ C) \ (A \\<^sub>r A) \ (C \\<^sub>r C)" by (auto dest: fun_relD) lemma set_of_ivl_impl[autoref_rules]: includes autoref_syntax shows "(\x. x, set_of_ivl) \ (A \\<^sub>r A) \ \A\ivl_rel" by (auto simp: ivl_rel_def br_def) lemma eucl_of_list_pad: includes autoref_syntax shows "DIM_precond TYPE('a::executable_euclidean_space) D \ (\xs. take D xs @ replicate (D - length xs) 0, eucl_of_list::_\'a) \ rl_rel \ lv_rel" unfolding lv_rel_def by (auto simp: intro!: brI) concrete_definition eucl_of_list_pad uses eucl_of_list_pad lemmas [autoref_rules] = eucl_of_list_pad.refine lemma source_of_res_impl[autoref_rules]: includes autoref_syntax shows "(ivl_of_res, source_of_res) \ Id \ \lv_rel\ivl_rel" unfolding source_of_res_def apply (auto simp: ivl_rel_def intro!: relcompI brI) subgoal for a apply (auto simp: ivl_of_res_def ivl_of_resultrect_def intro!: lv_relI) unfolding lv_rel_def apply (auto intro!: brI) done done definition tangent_seg_of_res :: "real aform numeric_options \ result \ R3 set nres" where "tangent_seg_of_res optns res0 = do { let fas = map (OP (nth matrix_of_degrees2\<^sub>e)) [0, 3, 6]; let u = min_deg res0; let v = max_deg res0; aform.approx_slp_appr fas (slp_of_fas fas) (lv_ivl [u, v, 0] [u, v, 1]) }" lemmas [refine_vcg_def] = tangent_seg_of_res_spec_def lemma tangent_seg_of_res[le, refine_vcg]: "tangent_seg_of_res optns res \ tangent_seg_of_res_spec res" unfolding tangent_seg_of_res_def tangent_seg_of_res_spec_def apply refine_vcg apply (auto simp: matrix_of_degrees2\<^sub>e_def Let_def in_segment) subgoal for x a b c u by (drule bspec[where x="[min_deg res, max_deg res, u]"]) (auto simp: tangent_of_deg_def lv_ivl_def algebra_simps intro!:) done lemma [autoref_rules]: includes autoref_syntax shows "(nth matrix_of_degrees2\<^sub>e, nth matrix_of_degrees2\<^sub>e) \ nat_rel \ Id" by auto schematic_goal tangent_seg_of_res_impl: includes autoref_syntax assumes [autoref_rules]: "(resi, res) \ Id" "(optnsi, optns) \ num_optns_rel" shows "(nres_of ?r, tangent_seg_of_res optns res) \ \aform.appr_rel\nres_rel" unfolding tangent_seg_of_res_def including art by autoref_monadic concrete_definition tangent_seg_of_res_impl uses tangent_seg_of_res_impl lemmas [autoref_rules] = tangent_seg_of_res_impl.refine[where optnsi = optnsi and optnsa=optns and optns="\_ _ _. optnsi" for optns optnsi, autoref_higher_order_rule] lemma return_of_res_impl: includes autoref_syntax shows "(\results res. (get_results (inf_retx res) (inf_rety res) (sup_retx res) (sup_rety res) results), return_of_res) \ \Id\list_rel \ Id \ \Id\list_wset_rel" by (auto simp: return_of_res_def list_wset_rel_def intro!: brI) concrete_definition return_of_res_impl uses return_of_res_impl lemmas [autoref_rules] = return_of_res_impl.refine lemma lorenz_optns'_impl[autoref_rules]: includes autoref_syntax shows "(lorenz_optns', lorenz_optns') \ \Id \ unit_rel\option_rel \ nat_rel \ nat_rel \ rnv_rel \ rnv_rel \ num_optns_rel" by auto lemma [autoref_rules]: includes autoref_syntax shows "(results, results) \ \Id\list_rel" "(invoke_nf, invoke_nf) \ Id \ bool_rel" by auto definition "check_line_nres print_fun m0 n0 c1 res0 = do { let X0 = source_of_res res0; (X0l, X0u) \ ivl_rep X0; ((m::nat, n::nat), a::int, modes, ro) \ mode_ro_spec c1 res0; let interrupt = invoke_nf res0; let optns = lorenz_optns' print_fun (the_default m m0) (the_default n n0) 1 (FloatR 1 a); tangents \ tangent_seg_of_res optns res0; aform.CHECKs (ST ''check_line_nres le'') (X0l \ X0u); sp \ aform.subset_spec_plane X0 (Sctn (eucl_of_list [0, 0, 1]) 27); aform.CHECKs (ST ''check_line_nres le'') sp; ASSERT (X0l \ X0u); Pe \ lorenz_poincare_tangents optns interrupt modes ro c1 ({X0l .. X0u}) tangents; PeS \ sets_of_coll Pe; let RETs = (return_of_res results res0); let RET = \((mk_coll ` (source_of_res ` RETs:::\lvivl_rel\list_wset_rel):::\clw_rel lvivl_rel\list_wset_rel)); every \ WEAK_ALL\<^bsup>\Pe. \P em eM Rivls. em > 0 \ Pe = scaleR2 em eM P \ fst ` P \ \Rivls \ (\Rivl \ Rivls. (\res\RETs. Rivl \ source_of_res res \ (c1 \ c1_entry_correct em P res0 res)))\<^esup> PeS (\Pe. do { let _ = aform.trace_set1e (ST ''# Return Element: '') (Some Pe); ((em, eM), P) \ scaleR2_rep Pe; aform.CHECKs (ST ''check_line_nres pos'') (0 < em); let R = (fst ` P:::aform.appr_rel); (Ri, Rs) \ op_ivl_rep_of_set R; let Rivl = (op_atLeastAtMost_ivl Ri Rs); Rivls \ aform.split_along_ivls2 3 (mk_coll Rivl) RET; Rivlss \ sets_of_coll Rivls; WEAK_ALL\<^bsup>\Rivl. \res\RETs. Rivl \ source_of_res res \ (c1 \ c1_entry_correct em P res0 res)\<^esup> Rivlss (\Rivl. do { b \ WEAK_EX\<^bsup>\res. Rivl \ source_of_res res \ (c1 \ c1_entry_correct em P res0 res)\<^esup> RETs (\res. do { let src = (source_of_res res:::lvivl_rel); let subs = Rivl \ src; cones \ if \(c1 \ subs) then RETURN True else check_c1_entry optns em P res0 res; RETURN (subs \ cones) }); let _ = aform.print_msg ((shows (ST ''# return of '') o shows res0 o shows (if b then ST '' OK'' else ST '' FAILED''))''''); RETURN b }) }); RETURN (every, Pe, RET) }" definition "aform_subset_spec_plane optns = aform.subset_spec_plane" lemma aform_subset_spec_plane_impl[autoref_rules]: includes autoref_syntax shows "DIM_precond TYPE('a::executable_euclidean_space) D \ (Xi, X::'a set) \ \lv_rel\ivl_rel \ (sctni, sctn) \ \lv_rel\sctn_rel \ (optnsi, optns) \ num_optns_rel \ (nres_of (subset_spec_plane_impl_aform optnsi D Xi sctni), aform_subset_spec_plane $ optns $ X $ sctn) \ \bool_rel\nres_rel" using aform.subset_spec_plane_impl.refine[where 'a='a, of D Xi X sctni sctn optnsi] by (force simp: aform_subset_spec_plane_def) schematic_goal check_line_impl: includes autoref_syntax assumes [autoref_rules]: "(pfi, pf) \ \Id \ unit_rel\option_rel" "(c1i, c1) \ bool_rel" "(res0i, res0) \ Id" "(m0i, m0) \ \nat_rel\option_rel" "(n0i, n0) \ \nat_rel\option_rel" shows "(nres_of ?r, check_line_nres $ pf $ m0 $ n0 $ c1 $ res0) \ \bool_rel \\<^sub>r clw_rel aform.appr1e_rel \\<^sub>r clw_rel lvivl_rel\nres_rel" unfolding check_line_nres_def including art by autoref_monadic concrete_definition check_line_impl uses check_line_impl[where optns = "\_ _ _ _ _ _ _ _ _ _ _ _ _ _ _ . lorenz_optns pfi" and optnsa = "\_ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _. lorenz_optns pfi" and optnsb = "\_ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _. lorenz_optns pfi" and optnsc = "\_ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _. lorenz_optns pfi" and optnsd = "\_ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _. lorenz_optns pfi" and optnse = "\_ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _. lorenz_optns pfi" and optnsf = "\ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _. lorenz_optns pfi" and pfi = pfi for pfi] lemmas [autoref_rules] = check_line_impl.refine lemma check_line_nres: "check_line_nres pf m0 n0 c1 res0 \ SPEC (\(every, Pe, RET). \P. Pe = \P \ (if c1 then \tans. source_of_res res0 \ {tangent_of_deg (min_deg res0)--tangent_of_deg (max_deg res0)} \ (\(x, y). (x, blinfun_apply y (1, 0, 0))) ` tans \ lorenz.poincare_mapsto \ (tans - \\<^sub>i (invoke_nf res0) \ UNIV) \\<^sub>l\<^sub>e UNIV (Pe) else lorenz.poincare_mapsto \ ((sourcei_of_res res0) \ UNIV) \\<^sub>l\<^sub>e UNIV (\P)) \ source_of_res res0 \ plane_of (Sctn (0, 0, 1) 27) \ (every \ (\x\P. \P em. em > 0 \ (\eM. x = scaleR2 em eM P) \ (\Rivls. fst ` P \ \Rivls \ (\Rivl\Rivls. \res\return_of_res results res0. Rivl \ source_of_res res \ (c1 \ c1_entry_correct em P res0 res))))))" if [refine_vcg]: NF unfolding check_line_nres_def sourcei_of_res_def apply (refine_vcg, clarsimp_all) using [[goals_limit=1]] subgoal for a0 a b c d e f g h i j k l m n p q r s t apply (rule exI[where x=i]) apply (rule exI[where x=j]) apply (rule conjI) apply force apply (rule conjI) apply (rule exI[where x=k]) apply force apply (rule exI[where x=s]) apply (rule conjI) defer apply force apply blast done subgoal for x0 y0 z0 x1 y1 z1 tangents s R every apply (rule exI[where x=R]) apply auto subgoal for tans apply (rule exI[where x=tans]) by auto subgoal for tans apply (rule exI[where x=tans]) by auto done done definition "print_sets_color (print_fun::String.literal \ unit) (c::string) (X::'a::executable_euclidean_space set) = ()" definition "print_lorenz_color print_fun cx cy cz ci cd1 cd2 P = ()" definition "print_aforms print_fun c aforms = fold (\a _. print_fun (String.implode (shows_segments_of_aform 0 1 a c ''\''))) aforms ()" lemma print_sets_color_impl[autoref_rules]: includes autoref_syntax shows "(\print_fun c X. print_aforms print_fun c X, print_sets_color) \ (Id \ unit_rel) \ string_rel \ clw_rel aform.appr_rel \ unit_rel" by auto lemma print_lorenz_color_impl[autoref_rules]: includes autoref_syntax shows "(\print_fun cx cy cz ci cd1 cd2 P. fold (\(_, x) b. print_lorenz_aform print_fun cx cy cz ci cd1 cd2 False (fst x @ the_default [] (snd x)) ) P (), print_lorenz_color) \ (Id \ unit_rel) \ string_rel \ string_rel \ string_rel \ string_rel \ \\string_rel\list_rel, \\string_rel\list_rel, (\clw_rel aform.appr1e_rel, unit_rel\fun_rel)\fun_rel\fun_rel" by auto definition check_line_core where "check_line_core print_funo m0 n0 c1 i = do { let print_fun = the_default (\_. ()) print_funo; CHECK (\_. print_fun (STR ''Hey, out of bounds!'')) (i < length results); let res = ((results:::\Id\list_rel) ! (i:::nat_rel)); (r, P, B) \ check_line_nres print_funo m0 n0 c1 res; let _ = print_sets_color print_fun (ST ''0x007f00'') (aform.sets_of_ivls B); (_, Pu) \ scaleR2_rep_coll P; let _ = print_sets_color print_fun (ST ''0x7f0000'') (aform.op_image_fst_coll (Pu:::clw_rel aform.appr1_rel):::clw_rel aform.appr_rel); let _ = print_lorenz_color print_fun (ST ''0x7f0000'') (ST ''0x7f0001'') (ST ''0x7f0002'') (ST ''0x7f0003'') [(ST ''0xc60000''), (ST ''0xc60000''), (ST ''0xc60000'')] [(ST ''0xf60000''), (ST ''0xf60000''), (ST ''0xf60000'')] P; let _ = (if r then print_fun (String.implode ((show (ST ''# VERIFIED '') @ show i @ show (ST ''\'')))) else print_fun (String.implode ((show (ST ''# Failed to verify '') @ show i @ show (ST ''\'')) ))); RETURN r }" lemma [autoref_rules]: includes autoref_syntax shows "(shows_prec, shows_prec) \ nat_rel \ nat_rel \ string_rel \ string_rel" "(shows_prec, shows_prec) \ nat_rel \ string_rel \ string_rel \ string_rel" "(String.implode, String.implode) \ string_rel \ Id" by (auto simp: string_rel_def) schematic_goal check_line_core_impl: includes autoref_syntax assumes [autoref_rules]: "(pfi, pf) \ \Id \ unit_rel\option_rel" "(c1i, c1) \ bool_rel" "(ii, i) \ nat_rel" "(m0i, m0) \ \nat_rel\option_rel" "(n0i, n0) \ \nat_rel\option_rel" shows "(nres_of ?f, check_line_core $ pf $ m0 $ n0 $ c1 $ i) \ \bool_rel\nres_rel" unfolding check_line_core_def including art by autoref_monadic concrete_definition check_line_core_impl for pfi m0i n0i c1i ii uses check_line_core_impl lemmas [autoref_rules] = check_line_core_impl.refine definition "c1i_of_res res = sourcei_of_res res \ conefield_of_res res" definition "correct_res res = ((\(x, dx) \ c1i_of_res res. x \ plane_of (Sctn (0, 0, 1) 27) \ dx \ plane_of (Sctn (0, 0, 1) 0) \ ((lorenz.returns_to \ x \ lorenz.return_time \ differentiable at x within \\<^sub>l\<^sub>e \ (\D. (lorenz.poincare_map \ has_derivative D) (at x within \\<^sub>l\<^sub>e) \ norm (D dx) \ expansion res * norm dx \ (\res2 \ return_of_res results res. (lorenz.poincare_map \ x, D dx) \ c1_of_res res2 \ norm (D dx) \ preexpansion res2 * norm dx))))))" lemma check_line_nres_c0_correct: "check_line_nres pf m0 n0 c1 res \ SPEC (\(every, Pe, RET). every \ (\x \ sourcei_of_res res. lorenz.poincare_map \ x \ \(source_of_res ` return_of_res results res)))" if NF apply (rule check_line_nres[OF \NF\, le]) apply (auto simp: c1i_of_res_def lorenz.poincare_mapsto_def) subgoal apply (drule bspec, force) apply (auto dest!: spec[where x="1\<^sub>L"]) apply (drule bspec, force) apply (force simp: scaleR2_def image_def vimage_def) done subgoal premises prems for a b c d e tans proof - obtain t where "((c, d, e), t) \ tans - \\<^sub>i (invoke_nf res) \ UNIV" "((c, d, e), tangent_of_deg (min_deg res)) = (\(x, y). (x, blinfun_apply y (1, 0, 0))) ((c, d, e), t)" using prems by (auto simp: sourcei_of_res_def) with prems(6)[rule_format, of "((c, d, e), t)"] prems(3) obtain D x where "tangent_of_deg (min_deg res) = blinfun_apply t (1, 0, 0)" "(c, d, e) returns_to \" "fst ` (tans - \\<^sub>i (invoke_nf res) \ UNIV) \ \\<^sub>l\<^sub>e" "lorenz.return_time \ differentiable at (c, d, e) within \\<^sub>l\<^sub>e" "(lorenz.poincare_map \ has_derivative blinfun_apply D) (at (c, d, e) within \\<^sub>l\<^sub>e)" "x \ b" "(lorenz.poincare_map \ (c, d, e), D o\<^sub>L t) \ x" by auto with prems show ?thesis subgoal apply (auto dest!: bspec[OF _ \x \ b\]) apply (auto simp: scaleR2_def image_def vimage_def) apply (auto simp: subset_iff) by fastforce \\slow\ done qed subgoal for a b c d e f apply (drule bspec[where A="sourcei_of_res res"]) apply force apply (auto dest!: spec[where x="1\<^sub>L"]) apply (drule bspec, force) apply auto apply (auto simp: scaleR2_def image_def vimage_def) apply (auto simp: subset_iff) by fastforce\ \slow\ done lemma cone_conefield[intro, simp]: "cone (conefield a b)" unfolding conefield_alt_def by (rule cone_cone_hull) lemma in_segment_norm_bound: "c \ {a -- b} \ norm c \ max (norm a) (norm b)" apply (auto simp: in_segment max_def intro!: norm_triangle_le) apply (auto simp: algebra_simps intro: add_mono mult_left_mono mult_right_mono) using affine_ineq by blast lemma norm_tangent_of_deg[simp]: "norm (tangent_of_deg d) = 1" by (auto simp: tangent_of_deg_def norm_prod_def) lemma check_line_nres_c1_correct: "check_line_nres pf m0 n0 True res \ SPEC (\(correct, Pe, RET). correct \ correct_res res)" if NF proof (rule check_line_nres[OF \NF\, le], clarsimp, goal_cases) case P: (1 a P tans) let ?tans = "{tangent_of_deg (min_deg res)--tangent_of_deg (max_deg res)}" have tans_plane: "?tans \ UNIV \ UNIV \ {0}" by (auto simp: in_segment tangent_of_deg_def) from P have *: "x \ plane_of (Sctn (0, 0, 1) 27)" if "x \ sourcei_of_res res" for x using that by (auto simp: that sourcei_of_res_def) from tans_plane P have **: "dx \ plane_of (Sctn (0, 0, 1) 0)" if "x \ sourcei_of_res res" "dx \ conefield_of_res res" for x dx proof - from tans_plane that obtain c dx' dy' where c: "dx = c *\<^sub>R (dx', dy', 0)" "(dx', dy', 0) \ ?tans" "c \ 0" unfolding conefield_of_res_def conefield_alt_def cone_hull_expl by auto then show ?thesis by (auto simp: plane_of_def) qed show ?case unfolding correct_res_def proof (intro ballI conjI, clarsimp_all simp add: * ** c1i_of_res_def c1_of_res_def sourcei_of_res_def, goal_cases) case source: (1 x y z dx dy dz) from tans_plane source obtain c dx' dy' where c: "(dx, dy, dz) = c *\<^sub>R (dx', dy', 0)" "(dx', dy', 0) \ ?tans" "c \ 0" unfolding conefield_of_res_def conefield_alt_def cone_hull_expl by auto from c source P obtain t where tans: "((x, y, z), t) \ tans" "blinfun_apply t (1, 0, 0) = (dx', dy', 0)" by auto from P(3) tans(1) source(3) obtain Re D where Re: "(x, y, z) returns_to \" "fst ` (tans - \\<^sub>i (invoke_nf res) \ UNIV) \ \\<^sub>l\<^sub>e" "lorenz.return_time \ differentiable at (x, y, z) within \\<^sub>l\<^sub>e" "(lorenz.poincare_map \ has_derivative blinfun_apply D) (at (x, y, z) within \\<^sub>l\<^sub>e)" "Re \ P" "(lorenz.poincare_map \ (x, y, z), D o\<^sub>L t) \ Re" by (auto simp: lorenz.poincare_mapsto_def dest!: bspec) from P(5)[rule_format, OF \Re \ P\] obtain R em eM Rivls where R: "Re = scaleR2 em eM R" "em > 0" "fst ` R \ \Rivls" "\Rivl. Rivl\Rivls \ \resa\return_of_res results res. Rivl \ source_of_res resa \ c1_entry_correct em R res resa" by auto have "lorenz.poincare_map \ (x, y, z) \ fst ` R" and s2: "(lorenz.poincare_map \ (x, y, z), D o\<^sub>L t) \ scaleR2 em eM R" using Re R by (auto simp: scaleR2_def) then obtain Rivl res' where Rivl: "lorenz.poincare_map \ (x, y, z) \ Rivl" "Rivl \ Rivls" "res' \ return_of_res results res" "Rivl \ source_of_res res'" and c1: "c1_entry_correct em R res res'" using R by force from s2 obtain ed Dt where Dt: "em \ ereal ed" "ereal ed \ eM" "D o\<^sub>L t = ed *\<^sub>R Dt" "(lorenz.poincare_map \ (x, y, z), Dt) \ R" by (force simp: scaleR2_def) then have Dt_simp[simp]: "Dt = inverse ed *\<^sub>R (D o\<^sub>L t)" using \0 < em\ by (cases em) (auto simp: intro!: simp: blinfun.bilinear_simps inverse_eq_divide) from c1[unfolded c1_entry_correct_def, rule_format, OF Dt(4)] obtain dxr dyr where dxrdyr: "blinfun_apply D (dx', dy', 0) /\<^sub>R ed = (dxr, dyr, 0)" "ereal (preexpansion res') \ em * ereal (norm (dxr, dyr, 0::real))" "ereal (expansion res) \ em * ereal (norm (dxr, dyr, 0::real))" "-90 < min_deg res'" "min_deg res' \ max_deg res'" "tan (rad_of (min_deg res')) \ (dyr / dxr)" "(dyr / dxr) \ tan (rad_of (max_deg res'))" "max_deg res' < 90" "0 < dxr" by (auto simp: blinfun.bilinear_simps tans) then obtain emr where emr: "em = ereal emr" "0 < emr" "emr \ ed" "(preexpansion res') \ emr * (norm (dxr, dyr, 0::real))" "(expansion res) \ emr * (norm (dxr, dyr, 0::real))" using \0 < em\ Dt by (cases em) (auto simp: simp: blinfun.bilinear_simps divide_simps prod_eq_iff) from dxrdyr have Ddx'dy': "D (dx', dy', 0) = ed *\<^sub>R (dxr, dyr, 0)" using \0 < em\ Dt by (cases em) (auto simp: simp: blinfun.bilinear_simps divide_simps prod_eq_iff) note \(x, y, z) returns_to \\ moreover note \lorenz.return_time \ differentiable at (x, y, z) within \\<^sub>l\<^sub>e\ moreover note \(lorenz.poincare_map \ has_derivative D) (at (x, y, z) within \\<^sub>l\<^sub>e)\ moreover note \res' \ return_of_res results res\ moreover have "lorenz.poincare_map \ (x, y, z) \ source_of_res res'" using Rivl by force moreover have \0 \ ed\ using Dt \0 < em\ by (cases em) auto have \D (dx, dy, dz) \ conefield_of_res res'\ unfolding c blinfun.bilinear_simps conefield_of_res_def Ddx'dy' apply (intro mem_cone, simp_all add: \0 \ ed\ \0 \ c\ tangent_of_deg_def) apply (rule conefield_prod3I) unfolding fun_cong[OF tan_def, symmetric] subgoal by fact subgoal using dxrdyr apply (intro cos_gt_zero_pi) unfolding rad_of_lt_iff rad_of_gt_iff by (auto simp: deg_of_def) subgoal using dxrdyr apply (intro cos_gt_zero_pi) unfolding rad_of_lt_iff rad_of_gt_iff by (auto simp: deg_of_def) subgoal by fact subgoal by fact done moreover have norms_le: "emr * norm (dx', dy', 0::real) * (\c\ * norm (dxr, dyr, 0::real)) \ \ed\ * (\c\ * norm (dxr, dyr, 0::real))" proof - from c(2)[THEN in_segment_norm_bound] have "norm (dx', dy', 0::real) \ 1" by auto also have "\ \ ed / emr" using dxrdyr emr unfolding Ddx'dy' by auto finally show ?thesis using emr by (intro mult_right_mono) (auto simp: divide_simps ac_simps) qed then have "expansion res * norm (dx, dy, dz) \ norm (D (dx, dy, dz))" unfolding c blinfun.bilinear_simps conefield_of_res_def Ddx'dy' norm_scaleR apply - apply (rule order_trans) apply (rule mult_right_mono) apply (rule emr) by (auto simp: ac_simps) moreover have "preexpansion res' * norm (dx, dy, dz) \ norm (D (dx, dy, dz))" using norms_le unfolding c blinfun.bilinear_simps conefield_of_res_def Ddx'dy' norm_scaleR apply - apply (rule order_trans) apply (rule mult_right_mono) apply (rule emr) by (auto simp: ac_simps) ultimately show ?case by blast qed qed lemma conefield_ne_empyt[simp]: "conefield a b \ {}" by (auto simp: conefield_def conesegment_def cone_hull_empty_iff[symmetric]) lemma in_return_of_resD: "res' \ return_of_res results res \ res' \ set results" by (auto simp: return_of_res_def get_results_def) lemma finite_results_at[intro, simp]: "finite (results_at x)" by (auto simp: results_at_def) lemma lorenz_bounds_lemma: "x returns_to \" "R x \ N" "(R has_derivative DR x) (at x within \\<^sub>l\<^sub>e)" "\c. c \ \ x \ DR x c \ \ (R x)" "\c. c \ \ x \ norm (DR x c) \ \ x * norm c" "\c. c \ \ x \ norm (DR x c) \ \\<^sub>p (R x) * norm c" if "x \ N - \" NF "\res. res \ set results \ correct_res res" proof - from \x \ N - \\ obtain res where res: "res \ set results" "x \ sourcei_of_res res" by (auto simp: N_def sourcei_of_res_def \\<^sub>i_def) then have ne: "c1i_of_res res \ {}" by (auto simp: c1i_of_res_def conefield_of_res_def) from res this obtain dx where dx: "(x, dx) \ c1i_of_res res" by (auto simp: c1i_of_res_def) from that(3)[OF \res \ set _\] have "correct_res res" by simp from this[unfolded correct_res_def, rule_format, OF dx] res obtain res' D where res': "x returns_to \" "lorenz.return_time \ differentiable at x within \\<^sub>l\<^sub>e" "(lorenz.poincare_map \ has_derivative D) (at x within \\<^sub>l\<^sub>e)" "expansion res * norm dx \ norm (D dx)" "res' \ return_of_res results res" "(lorenz.poincare_map \ x, D dx) \ c1_of_res res'" "preexpansion res' * norm dx \ norm (D dx)" by auto show "x returns_to \" by fact show "R x \ N" using res' by (auto simp: R_def N_def N_def c1i_of_res_def c1_of_res_def in_return_of_resD sourcei_of_res_def) show "(R has_derivative DR x) (at x within \\<^sub>l\<^sub>e)" apply (auto simp: R_def DR_def N_def c1_of_res_def in_return_of_resD) apply (subst frechet_derivative_works[symmetric]) apply (rule differentiableI) by fact next fix dx assume "dx \ \ x" then obtain res where res: "res \ set results" and dx: "(x, dx) \ c1_of_res res" by (auto simp: \_def results_at_def c1_of_res_def ) then have dx: "(x, dx) \ c1i_of_res res" using \x \ N - _\ by (auto simp: c1i_of_res_def sourcei_of_res_def c1_of_res_def \\<^sub>i_def) from res dx have ne: "c1i_of_res res \ {}" by (auto simp: c1_of_res_def conefield_of_res_def) from that(3)[OF \res \ set _\] have "correct_res res" by simp from that this[unfolded correct_res_def, rule_format, OF dx] res obtain res' D where res': "x returns_to \" "x \ plane_of (Sctn (0, 0, 1) 27)" "lorenz.return_time \ differentiable at x within \\<^sub>l\<^sub>e" "(lorenz.poincare_map \ has_derivative D) (at x within \\<^sub>l\<^sub>e)" "expansion res * norm dx \ norm (D dx)" "res' \ return_of_res results res" "(lorenz.poincare_map \ x, D dx) \ c1_of_res res'" "preexpansion res' * norm dx \ norm (D dx)" by auto have DRD: "DR x = D" unfolding DR_def apply (rule frechet_derivative_unique_within) apply (subst frechet_derivative_works[symmetric]) apply (rule differentiableI) apply fact apply fact using \x \ plane_of _\ apply safe subgoal for _ _ _ e by (auto simp: \\<^sub>l\<^sub>e_def Basis_prod_def prod_eq_iff plane_of_def prod_eq_iff inner_prod_def intro!: exI[where x="-e/2"]) done have [intro, simp]: "res' \ results_at (lorenz.poincare_map \ x)" using res' by (auto simp: c1_of_res_def results_at_def in_return_of_resD R_def intro!: exI[where x=res']) have [intro, simp]: "res \ results_at x" using res dx by (auto simp: c1i_of_res_def results_at_def sourcei_of_res_def) show "DR x dx \ \ (R x)" unfolding DRD \_def using res' by (auto simp: c1_of_res_def R_def) have "\ x * norm dx \ expansion res * norm dx" by (rule mult_right_mono) (auto simp: \_def) also have "\ \ norm (DR x dx)" unfolding DRD by fact finally show "\ x * norm dx \ norm (DR x dx)" . have "\\<^sub>p (R x) * norm dx \ preexpansion res' * norm dx" by (rule mult_right_mono) (auto simp: \\<^sub>p_def R_def) also have "\ \ norm (DR x dx)" unfolding DRD by fact finally show "\\<^sub>p (R x) * norm dx \ norm (DR x dx)" . qed lemma check_line_core_correct: "check_line_core pf m0 n0 True i \ SPEC (\correct. correct \ correct_res (results ! i))" if [refine_vcg]: NF unfolding check_line_core_def supply [refine_vcg] = check_line_nres_c1_correct[le] by refine_vcg text \The symmetric reduction\ lemma source_of_res_mirror: "(x, y, z) \ source_of_res (mirror_result res) \ (-x, -y, z) \ source_of_res res" by (cases res) (auto simp: source_of_res_def ivl_of_res_def ivl_of_resultrect_def set_of_ivl_def) lemma conefield_of_res_mirror[simp]: "(x, y, z) \ conefield_of_res (mirror_result res) \ (x, y, z) \ conefield_of_res res" by (cases res) (auto simp: conefield_of_res_def ivl_of_res_def) lemma c1_of_res_mirror: "((x, y, z), dx, dy, dz) \ c1_of_res (mirror_result res) \ ((-x, -y, z), dx, dy, dz) \ c1_of_res res" by (auto simp: c1_of_res_def source_of_res_mirror) lemmas [simp] = lorenz_S(2) lemma lorenz_S_idem[simp]: "lorenz_S (lorenz_S x) = (x::R3)" by (auto simp: lorenz_S_def split_beta') lemma lorenz_S_exivl[simp]: "lorenz.existence_ivl0 (lorenz_S X) = lorenz.existence_ivl0 X" using lorenz_S(1)[of _ X] using lorenz_S(1)[of _ "lorenz_S X"] by auto lemma lorenz_S_zero[simp]: "lorenz_S x = 0 \ (x::R3) = 0" by (auto simp: lorenz_S_def split_beta' prod_eq_iff) lemma lorenz_S_returns_toI[simp]: "x returns_to (lorenz_S ` P) \ lorenz_S x returns_to P" apply (auto simp: lorenz.returns_to_def) subgoal premises prems for t proof - have " \\<^sub>F s in at_right 0. s < t" using tendsto_ident_at \0 < t\ by (rule order_tendstoD) then have " \\<^sub>F s in at_right 0. s \ lorenz.existence_ivl0 x" unfolding eventually_at_filter apply eventually_elim using \0 < t\ lorenz.closed_segment_subset_existence_ivl[OF prems(3)] by (auto simp: closed_segment_eq_real_ivl subset_iff) then show ?thesis using prems(1) by eventually_elim force qed done lemma lorenz_S_returns_to[simp]: "lorenz_S x returns_to P \ x returns_to (lorenz_S ` P)" using lorenz_S_returns_toI[of P x] lorenz_S_returns_toI[of "lorenz_S ` P" "lorenz_S x"] by (auto simp: image_image) lemma lorenz_S_image_Sigma[simp]: "lorenz_S ` \ = \" apply (auto simp: \_def lorenz_S_def) apply (rule image_eqI) apply (rule lorenz_S_idem[symmetric]) apply (auto simp: \_def lorenz_S_def) done lemma linear_lorenz_S: "linear lorenz_S" by unfold_locales (auto simp: lorenz_S_def) lemma inj_lorenz_S: "inj_on (lorenz_S::R3 \ _) G" by (rule inj_onI) (auto simp: lorenz_S_def prod_eq_iff) lemma lorenz_S_return_time: "lorenz.return_time P (lorenz_S x) = lorenz.return_time (lorenz_S ` P) x" if "x returns_to (lorenz_S ` P)" "closed P" proof - from lorenz.returns_toE[OF that(1)] obtain t0 t1 where f: "0 < t0" "t0 \ t1" " t1 \ lorenz.existence_ivl0 x" "lorenz.flow0 x t1 \ lorenz_S ` P" "\t. 0 < t \ t < t0 \ lorenz.flow0 x t \ lorenz_S ` P" by auto have [simp]: "lorenz.return_time (lorenz_S ` P) x \ lorenz.existence_ivl0 x" by (auto intro!: that closed_injective_linear_image linear_lorenz_S lorenz.return_time_exivl inj_lorenz_S) have c': "closed (lorenz_S ` P)" by (auto intro!: that closed_injective_linear_image linear_lorenz_S lorenz.return_time_exivl lorenz.return_time_pos inj_lorenz_S) show ?thesis using f(1-4) using lorenz.return_time_returns[OF that(1) c'] apply (intro lorenz.return_time_eqI) apply (auto intro!: that closed_injective_linear_image linear_lorenz_S lorenz.return_time_exivl lorenz.return_time_pos c' inj_lorenz_S) subgoal premises prems for a b c d e f g proof - have [simp]: "a \ lorenz.existence_ivl0 x" using _ that(1) apply (rule lorenz.less_return_time_imp_exivl) using prems that(2) c' by auto have "lorenz.return_time (lorenz_S ` P) x \ a" apply (rule lorenz.return_time_le) using prems apply (auto intro!: that closed_injective_linear_image linear_lorenz_S lorenz.return_time_exivl lorenz.return_time_pos c' inj_lorenz_S) apply (rule image_eqI) apply (rule lorenz_S_idem[symmetric]) by auto then show ?thesis using prems by simp qed done qed lemma lorenz_S_poincare_map: "lorenz.poincare_map P (lorenz_S x) = lorenz_S (lorenz.poincare_map (lorenz_S ` P) x)" if "x returns_to (lorenz_S ` P)" "closed P" using that unfolding lorenz.poincare_map_def apply (auto simp: lorenz_S_return_time) apply (subst lorenz_S) by (auto intro!: lorenz.return_time_exivl that closed_injective_linear_image linear_lorenz_S inj_lorenz_S) lemma [continuous_intros]: "isCont (lorenz_S::_\R3) x" "continuous_on (G::R3 set) lorenz_S" by (auto simp:lorenz_S_def[abs_def] split_beta' continuous_intros) lemma filtermap_lorenz_S_le: "filtermap lorenz_S (at x within lorenz_S ` P) \(at (lorenz_S x::R3) within P)"\ \TODO: generalize!\ unfolding at_within_def apply (auto simp: intro!: antisym filtermap_inf[le] filtermap_inf[ge]) apply (rule inf.coboundedI1) apply (subst filtermap_nhds_open_map) apply (auto simp: intro!: invariance_of_domain inj_lorenz_S continuous_intros) apply (rule inf.coboundedI2) apply (auto simp: image_image ) apply (auto simp: lorenz_S_def split_beta')[] done lemma filtermap_lorenz_S_eq: "filtermap lorenz_S (at (x::R3) within lorenz_S ` P) = (at (lorenz_S x::R3) within P)" apply (rule antisym) using filtermap_lorenz_S_le[of "x" P] apply simp subgoal proof - have "filtermap lorenz_S (at (lorenz_S x) within P) \ filtermap lorenz_S (filtermap lorenz_S (at x within lorenz_S ` P))" using filtermap_lorenz_S_le[of "lorenz_S x" "lorenz_S ` P"] by (auto simp: image_image filtermap_filtermap) then show ?thesis apply (subst (asm) filtermap_mono_strong) by (auto simp: inj_lorenz_S) qed done lemma norm_lorenz_S[simp]: "norm (lorenz_S x) = norm x" by (auto simp: lorenz_S_def norm_prod_def split_beta') lemma bl_lorenz_S: "bounded_linear (lorenz_S)" by unfold_locales (auto simp: lorenz_S_def norm_prod_def intro!: exI[where x=1]) lemma filtermap_lorenz_S_eq_bot[simp]: "filtermap (lorenz_S::R3\_) F = bot \ F = bot" apply (auto simp: ) apply (subst (asm) filtermap_bot[symmetric]) apply (subst (asm) filtermap_eq_strong) by (auto simp: inj_lorenz_S) lemma netlimit_filtermap[simp]: "at x within X \ bot \ netlimit (filtermap lorenz_S (at x within X)) = lorenz_S (x::R3)" apply (rule tendsto_Lim) unfolding filterlim_filtermap apply simp by (auto intro!: tendsto_eq_intros simp: split_beta' lorenz_S_def[abs_def]) lemma lorenz_S_halfspace [simp]: "lorenz_S ` \\<^sub>l\<^sub>e = \\<^sub>l\<^sub>e" apply (auto simp: \\<^sub>l\<^sub>e_def lorenz_S_def[abs_def]) apply (rule image_eqI) apply auto apply (rule sym) apply (rule minus_minus) apply (rule minus_minus[symmetric]) done lemma closure_Sigma_le_eq: "closure \\<^sub>l\<^sub>e = \\<^sub>l\<^sub>e" proof (rule closure_closed) have "\\<^sub>l\<^sub>e = {x. x \ (0, 0, 1) \ 27}" by (auto simp: \\<^sub>l\<^sub>e_def ) also have "closed \" by (rule closed_halfspace_component_le) finally show "closed \\<^sub>l\<^sub>e" . qed lemma closure_Sigma_le[simp]: "closure (\\<^sub>l\<^sub>e - {x}) = \\<^sub>l\<^sub>e" proof (cases "x \ \\<^sub>l\<^sub>e") case that: True have "closure \\<^sub>l\<^sub>e \ closure (insert x (\\<^sub>l\<^sub>e - {x}))" by (rule closure_mono) auto also have "\ = insert x (closure (\\<^sub>l\<^sub>e - {x}))" apply (subst closure_insert) by simp also have "x \ closure (\\<^sub>l\<^sub>e - {x})" apply (rule closed_sequentially[where f="\n. x - (0, 0, inverse (Suc n))"]) apply (rule closed_closure) subgoal apply (auto simp: ) apply (rule subsetD) apply (rule closure_subset) using that apply (auto simp: \\<^sub>l\<^sub>e_def prod_eq_iff) apply (rule order_trans) apply (rule diff_right_mono) apply (assumption) apply simp done subgoal apply (rule tendsto_eq_intros) apply (rule tendsto_intros) apply (rule tendsto_intros) apply (rule tendsto_intros) apply (rule tendsto_intros) apply (rule tendsto_intros) apply (rule LIMSEQ_inverse_real_of_nat) by (auto simp: prod_eq_iff) done then have "insert x (closure (\\<^sub>l\<^sub>e - {x})) \ closure (\\<^sub>l\<^sub>e - {x})" by auto finally have "closure \\<^sub>l\<^sub>e \ closure (\\<^sub>l\<^sub>e - {x})" . moreover have "closure (\\<^sub>l\<^sub>e - {x}) \ closure (\\<^sub>l\<^sub>e)" by (rule closure_mono) auto ultimately have "closure (\\<^sub>l\<^sub>e - {x}) = closure (\\<^sub>l\<^sub>e)" by simp also have "\ = \\<^sub>l\<^sub>e" by (rule closure_Sigma_le_eq) finally show ?thesis . next case False then show ?thesis apply simp apply (rule closure_Sigma_le_eq) done qed lemma lorenz_S_return_time_has_derivative: assumes "(lorenz.return_time \ has_derivative D) (at x within \\<^sub>l\<^sub>e)" and "lorenz.returns_to \ x" and "x \ \\<^sub>l\<^sub>e" shows "(lorenz.return_time \ has_derivative D o lorenz_S) (at (lorenz_S x) within \\<^sub>l\<^sub>e)" proof - have [simp]: "\trivial_limit (at x within \\<^sub>l\<^sub>e)" unfolding at_within_eq_bot_iff using assms by simp interpret bounded_linear "lorenz_S::R3\_" by (rule bl_lorenz_S) have "\\<^sub>F x in at x within \\<^sub>l\<^sub>e. (x::R3) returns_to \" by (blast intro: lorenz.eventually_returns_to_continuousI has_derivative_continuous assms) then have "\\<^sub>F y in at x within \\<^sub>l\<^sub>e. inverse (norm (y - x)) * (lorenz.return_time \ y - lorenz.return_time \ x - D (y - x)) = inverse (norm (lorenz_S y - lorenz_S x)) * (lorenz.return_time \ (lorenz_S y) - lorenz.return_time \ (lorenz_S x) - D (y - x))" by eventually_elim (auto simp: lorenz_S_return_time assms diff[symmetric]) then show ?thesis using assms apply (subst filtermap_lorenz_S_eq[symmetric]) apply (auto simp: has_derivative_def filterlim_filtermap) unfolding o_def apply (rule bounded_linear_compose, assumption, rule bl_lorenz_S) unfolding diff lorenz_S_idem apply (auto simp: Lim_ident_at) apply (blast intro: Lim_transform_eventually) done qed lemma lorenz_S_return_time_differentiable: "lorenz.return_time \ differentiable at (lorenz_S x) within \\<^sub>l\<^sub>e" if "lorenz.return_time \ differentiable at x within \\<^sub>l\<^sub>e" "lorenz.returns_to \ x" "x \ \\<^sub>l\<^sub>e" proof - from that obtain D where "(lorenz.return_time \ has_derivative D) (at x within \\<^sub>l\<^sub>e)" by (auto simp: differentiable_def) then have "(lorenz.return_time \ has_derivative D o lorenz_S) (at (lorenz_S x) within \\<^sub>l\<^sub>e)" by (rule lorenz_S_return_time_has_derivative) fact+ then show ?thesis by (auto simp: differentiable_def) qed lemma lorenz_S_has_derivative: "(lorenz_S has_derivative lorenz_S) (at (x::R3) within X)" by (auto simp: lorenz_S_def[abs_def] split_beta' intro!: derivative_eq_intros) lemma lorenz_S_poincare_map_has_derivative: assumes "(lorenz.poincare_map \ has_derivative D) (at x within \\<^sub>l\<^sub>e)" "(lorenz.return_time \ has_derivative Dr) (at x within \\<^sub>l\<^sub>e)" "lorenz.returns_to \ x" "x \ \\<^sub>l\<^sub>e" shows "(lorenz.poincare_map \ has_derivative lorenz_S o D o lorenz_S) (at (lorenz_S x) within \\<^sub>l\<^sub>e)" proof - have [simp]: "\trivial_limit (at x within \\<^sub>l\<^sub>e)" unfolding at_within_eq_bot_iff using assms by simp interpret bounded_linear "lorenz_S::R3\_" by (rule bl_lorenz_S) have "\\<^sub>F x in at x within \\<^sub>l\<^sub>e. (x::R3) returns_to \" by (blast intro: lorenz.eventually_returns_to_continuousI has_derivative_continuous assms) then have "\\<^sub>F y in at x within \\<^sub>l\<^sub>e. (lorenz_S (lorenz.poincare_map \ y) - lorenz_S (lorenz.poincare_map \ x) - lorenz_S (D (y - x))) /\<^sub>R norm (y - x) = (lorenz.poincare_map \ (lorenz_S y) - lorenz.poincare_map \ (lorenz_S x) - lorenz_S (D (y - x))) /\<^sub>R norm (lorenz_S y - lorenz_S x)" by eventually_elim (auto simp: lorenz_S_return_time lorenz_S_poincare_map assms diff[symmetric]) then show ?thesis using has_derivative_compose[OF assms(1) lorenz_S_has_derivative] assms apply (subst filtermap_lorenz_S_eq[symmetric]) apply (auto simp: has_derivative_def filterlim_filtermap) unfolding o_def apply (rule bounded_linear_compose, rule bl_lorenz_S) apply (rule bounded_linear_compose, assumption, rule bl_lorenz_S) unfolding diff lorenz_S_idem apply (auto simp: Lim_ident_at) apply (blast intro: Lim_transform_eventually) done qed lemma [simp]: "expansion (mirror_result res) = expansion res" by (cases res) auto lemma lorenz_S_on_plane: "lorenz_S (dx, dy, 0::real) = - (dx, dy, 0)" by (auto simp: lorenz_S_def ) lemma mirror_result_idem[simp]: "mirror_result (mirror_result x) = x" by (cases x) (auto simp: mirror_result_def) lemma mirror_in_set: "x \ set results \ mirror_result x \ set results" by (auto simp: results_def symmetrize_def) lemma mirror_result_in: "mirror_result res2 \ return_of_res results (mirror_result res)" if "res2 \ return_of_res results res" proof - from that have "res2 \ set results" by (rule in_return_of_resD) from mirror_in_set[OF this] have "mirror_result res2 \ set results" . then show ?thesis apply (cases res2; cases res) using that by (auto simp: return_of_res_def get_results_def) qed lemma in_source_of_res_mirrorI: "(x::R3) \ source_of_res (mirror_result (r))" if "lorenz_S x \ source_of_res r" using that apply (cases r; cases x) by (auto simp: source_of_res_def set_of_ivl_def ivl_of_res_def lorenz_S_def less_eq_prod_def ivl_of_resultrect_def) lemma conefield_of_res_mirror_simp[simp]: "conefield_of_res (mirror_result res2) = conefield_of_res res2" by (cases res2) (auto simp: conefield_of_res_def) lemma lorenz_minus_planeI: "lorenz_S (- x) = x" if "snd (snd (x::R3)) = 0" using that by (auto simp: lorenz_S_def split_beta' prod_eq_iff) lemma preexpansion_mirror_result[simp]: "preexpansion (mirror_result res2) = preexpansion res2" by (cases res2) (auto simp: ) lemma lorenz_S_tendsto_0I: "(lorenz.flow0 (lorenz_S x) \ 0) at_top" if "{0..} \ lorenz.existence_ivl0 x" "(lorenz.flow0 x \ 0) at_top" proof (rule Lim_transform_eventually) have "\\<^sub>F s in at_top. (s::real) \ 0" using eventually_ge_at_top by blast then show "\\<^sub>F s in at_top. lorenz_S (lorenz.flow0 x s) = lorenz.flow0 (lorenz_S x) s" by eventually_elim (use that in auto) show "((\s. lorenz_S (lorenz.flow0 x s)) \ 0) at_top" unfolding Zfun_def[symmetric] by (rule bounded_linear.tendsto_zero[OF bl_lorenz_S that(2)]) qed lemma lorenz_S_tendsto_0_iff: "(lorenz.flow0 (lorenz_S x) \ 0) at_top \ (lorenz.flow0 x \ 0) at_top" if "{0..} \ lorenz.existence_ivl0 x" using lorenz_S_tendsto_0I[of x, OF that] lorenz_S_tendsto_0I[of "lorenz_S x"] that by auto lemma lorenz_S_eq_iff[simp]: "lorenz_S y = lorenz_S x \ y = x" for x y::"real*real*real" by (auto simp: lorenz_S_def split: prod.splits) lemma lorenz_S_\: "lorenz_S x \ \ \ x \ \" apply (auto simp: \_def lorenz_S_tendsto_0_iff ) subgoal for t apply (auto simp: dest!: spec[where x=t]) apply (subst (asm) lorenz_S) apply auto apply (subst (asm) (2) lorenz_S_image_Sigma[symmetric]) by (simp del: lorenz_S_image_Sigma) subgoal for t apply (auto simp: dest!: spec[where x=t]) apply (subst (asm) lorenz_S) apply auto apply (subst (asm) lorenz_S_image_Sigma[symmetric]) apply (auto simp del: lorenz_S_image_Sigma) done done lemma sourcei_of_res_mirror: "(x, y, z) \ sourcei_of_res (mirror_result res) \ (-x, -y, z) \ sourcei_of_res res" using lorenz_S_\[of "(x, y, z)"] by (cases res) (auto simp: source_of_res_def sourcei_of_res_def ivl_of_res_def ivl_of_resultrect_def set_of_ivl_def \\<^sub>i_def lorenz_S_def) lemma c1i_of_res_mirror: "((x, y, z), dx, dy, dz) \ c1i_of_res (mirror_result res) \ ((-x, -y, z), dx, dy, dz) \ c1i_of_res res" by (auto simp: c1i_of_res_def sourcei_of_res_mirror) lemma correct_res_mirror_result: "correct_res (mirror_result res)" if "correct_res res" unfolding correct_res_def proof (clarsimp simp add: c1i_of_res_mirror, goal_cases) case (1 x y z dx dy dz) then have 1: "(lorenz_S (x, y, z), dx, dy, dz) \ c1i_of_res res" by (auto simp: lorenz_S_def) from that[unfolded correct_res_def, rule_format, OF 1, simplified] have "(lorenz_S (x, y, z)) \ plane_of (Sctn (0, 0, 1) 27)" "(dx, dy, dz) \ plane_of (Sctn (0, 0, 1) 0)" by auto then have plane: "(x, y, z) \ plane_of (Sctn (0, 0, 1) 27)" "(dx, dy, dz) \ plane_of (Sctn (0, 0, 1) 0)" by (auto simp: plane_of_def lorenz_S_def) then show ?case proof (clarsimp, goal_cases) case mem: 1 with that[unfolded correct_res_def, rule_format, OF 1, simplified] obtain D res2 where D: "lorenz_S (x, y, z) returns_to \" "lorenz.return_time \ differentiable at (lorenz_S (x, y, z)) within \\<^sub>l\<^sub>e" "(lorenz.poincare_map \ has_derivative D) (at (lorenz_S (x, y, z)) within \\<^sub>l\<^sub>e)" "expansion res * norm (dx, dy, dz) \ norm (D (dx, dy, dz))" "res2 \ return_of_res results res" "(lorenz.poincare_map \ (lorenz_S (x, y, z)), D (dx, dy, dz)) \ c1_of_res res2" "preexpansion res2 * norm (dx, dy, dz) \ norm (D (dx, dy, dz))" by auto from plane have S_le: "lorenz_S (x, y, z) \ \\<^sub>l\<^sub>e" by (auto simp: \\<^sub>l\<^sub>e_def plane_of_def lorenz_S_def) interpret linear D by (rule has_derivative_linear; fact) have ret: "(x, y, z) returns_to \" using D(1) lorenz_S_returns_to by simp moreover have "lorenz.return_time \ differentiable at (x, y, z) within \\<^sub>l\<^sub>e" using lorenz_S_return_time_differentiable[OF D(2) D(1) S_le] by auto moreover from D obtain Dr where Dr: "(lorenz.return_time \ has_derivative Dr) (at (lorenz_S (x, y, z)) within \\<^sub>l\<^sub>e)" by (auto simp: differentiable_def) let ?D = "lorenz_S \ D \ lorenz_S" have "(lorenz.poincare_map \ has_derivative ?D) (at (x, y, z) within \\<^sub>l\<^sub>e)" using lorenz_S_poincare_map_has_derivative[OF D(3) Dr D(1) S_le] by auto moreover from plane have [simp]: "dz = 0" by (auto simp: plane_of_def) have "expansion (mirror_result res) * norm (dx, dy, dz) \ norm (?D (dx, dy, dz))" using D(4) apply (auto simp: ) unfolding lorenz_S_on_plane neg by simp moreover have \mirror_result res2 \ return_of_res results (mirror_result res)\ using D(5) by (rule mirror_result_in) moreover have "(lorenz.poincare_map \ (x, y, z), ?D (dx, dy, dz)) \ c1_of_res (mirror_result res2)" using D(6) apply (subst (asm) lorenz_S_poincare_map) apply auto apply fact apply (auto simp: c1_of_res_def in_source_of_res_mirrorI) unfolding lorenz_S_on_plane neg apply (subst lorenz_minus_planeI) apply (auto simp: conefield_of_res_def conefield_alt_def cone_hull_expl in_segment tangent_of_deg_def) done moreover have "preexpansion (mirror_result res2) * norm (dx, dy, dz) \ norm (?D (dx, dy, dz))" using D(7) apply (auto simp: ) unfolding lorenz_S_on_plane neg by simp ultimately show ?case by (force intro!: exI[where x = ?D] bexI[where x="mirror_result res2"]) qed qed lemma reduce_lorenz_symmetry: "Ball (set results) correct_res" if "Ball (set coarse_results) correct_res" using that by (auto simp: results_def symmetrize_def intro!: correct_res_mirror_result) end subsection \Code Generation\ definition [code_abbrev]: "my_divide_integer (i::integer) (j::integer) = i div j" code_printing constant my_divide_integer \ (SML) "IntInf.div/ (_,/ _)" subsection \Tuning code equations\ definition mult_twopow_int::"int \ int \ int" where "mult_twopow_int x n = x * (power_int 2 n)" definition div_twopow_int :: "int \ int \ int" where "div_twopow_int x n = x div (power_int 2 n)" context includes integer.lifting begin lift_definition mult_twopow_integer :: "integer \ integer \ integer" is mult_twopow_int . lift_definition div_twopow_integer :: "integer \ integer \ integer" is div_twopow_int . end lemma compute_float_round_down[code]: "float_round_down prec (Float m e) = (let d = bitlen \m\ - int prec - 1 in if 0 < d then Float (div_twopow_int m d) (e + d) else Float m e)" including float.lifting using Float.compute_float_down[of "Suc prec - bitlen \m\ - e" m e, symmetric] by transfer (auto simp add: field_simps abs_mult log_mult bitlen_alt_def truncate_down_def div_twopow_int_def power_int_def cong del: if_weak_cong) lemma compute_float_plus_down[code]: fixes p::nat and m1 e1 m2 e2::int shows "float_plus_down p (Float m1 e1) (Float m2 e2) = (if m1 = 0 then float_round_down p (Float m2 e2) else if m2 = 0 then float_round_down p (Float m1 e1) else (if e1 \ e2 then (let k1 = Suc p - nat (bitlen \m1\) in if bitlen \m2\ > e1 - e2 - k1 - 2 then float_round_down p ((Float m1 e1) + (Float m2 e2)) else float_round_down p (Float (mult_twopow_int m1 (int k1 + 2) + sgn m2) (e1 - int k1 - 2))) else float_plus_down p (Float m2 e2) (Float m1 e1)))" using Float.compute_float_plus_down[of p m1 e1 m2 e2] by (auto simp: mult_twopow_int_def Let_def power_int_def nat_add_distrib) subsection \Codegen\ definition "is_dRETURN_True x = (case x of dRETURN b \ b | _ \ False)" definition "file_output_option s f = (case s of None \ f None | Some s \ file_output (String.implode s) (\pf. f (Some pf)))" definition "check_line_lookup_out s m0 n0 c1 i = is_dRETURN_True (file_output_option s (\pf. check_line_core_impl pf m0 n0 c1 i))" fun alternating where "alternating [] xs = xs" | "alternating xs [] = xs" | "alternating (x#xs) (y#ys) = x#y#alternating xs ys" definition "ordered_lines = alternating (rev [0..<222]) ([222..<400])" \ \the hard ones ``first'', potentially useless due to nondeterministic \Parallel.map\\ definition "parallel_check filenameo m n c1 ns = Parallel.forall (\i. let _ = print (String.implode (''# Starting '' @ show i @ ''\'')); b = check_line_lookup_out (map_option (\f. f @ show i) filenameo) (Some m) (Some n) c1 i; _ = if b then print (String.implode (''# Success: '' @ show i @ ''\'')) else print (String.implode (''# Failed: '' @ show i @ ''\'')) in b ) ns" ML \val check_line = @{computation_check terms: Trueprop parallel_check ordered_lines check_line_core_impl check_line_lookup_out (* bool *) True False (* num *) Num.One Num.Bit0 Num.Bit1 (* nat *) Suc "0::nat" "1::nat" "numeral::num\nat" (* int / integer*) "numeral::num\int" "numeral::num\integer" "uminus::_\int" "uminus::_\integer" int_of_integer integer_of_int "0::int" "1::int" (* Pairs *) "Pair::_ \ _\ (real list \ real list)" "Pair::_\_\(real list \ real list) \ real list sctn" "Pair::_\_\((real list \ real list) \ real list sctn) list \ real aform reach_options" (* Option *) "None::nat option" "Some::_\nat option" "None::string option" "Some::_\string option" (* Lists *) "Nil::real list" "Cons::_\_\real list" "Nil::nat list" "Cons::_\_\nat list" "Nil::real aform list" "Cons::_\_\real aform list" "Nil::((real list \ real list) \ real list sctn) list" "Cons::_\_\((real list \ real list) \ real list sctn) 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" (* String *) String.Char String.implode "Cons::char \ char list \ char list" "Nil::char list" (* float *) Float float_of_int float_of_nat (* real *) "numeral::num\real" "real_of_float" "(/)::real\real\real" "uminus::real\_" real_divl real_divr real_of_int (* section *) "Sctn::_\_\real list sctn" (* aform *) "aforms_of_ivls::_\_\real aform list" (* input *) coarse_results (* modes *) xsec xsec' ysec ysec' zsec zsec' zbucket lookup_mode ro ro_outer mode_outer (* unit *) "()" }\ lemma is_dRETURN_True_iff[simp]: "is_dRETURN_True x \ (x = dRETURN True)" by (auto simp: is_dRETURN_True_def split: dres.splits) lemma check_line_core_impl_True: "check_line_core_impl pfo m n True i = dRETURN True \ NF \ correct_res (results ! i)" apply (cases "check_line_core_impl pfo m n True i") using check_line_core_correct[of pfo m n i] check_line_core_impl.refine[of pfo pfo True True i i m m n n] apply (auto simp: nres_rel_def) apply (drule order_trans[where y="check_line_core pfo m n True i"]) apply assumption by auto lemma check_line_lookup_out: "correct_res (results ! i)" if "\s m n. check_line_lookup_out s m n True i" NF using that by (auto simp: check_line_lookup_out_def file_output_iff check_line_core_impl_True file_output_option_def split: dres.splits option.splits) definition "check_lines c1 ns = list_all (\i. \s m n. check_line_lookup_out s m n c1 i) ns" lemma check_linesI: "check_lines c1 ns" if "parallel_check s m n c1 ns" using that by (auto simp: parallel_check_def check_lines_def list_all_iff) subsection \Automate generation of lemmas\ lemma length_coarse_results[simp]: "length coarse_results = 400" by (simp add: coarse_results_def) lemma correct_res_coarse_resultsI: "correct_res (results ! i) \ i < 400 \ correct_res (coarse_results ! i)" by (auto simp: results_def symmetrize_def nth_append) lemma Ball_coarseI: "Ball (set coarse_results) correct_res" if NF "check_lines True xs" "set xs = {..<400}" using that by (force simp: check_lines_def list_all_iff in_set_conv_nth intro!: correct_res_coarse_resultsI check_line_lookup_out) ML \map_option (using_master_directory_term @{context}) (SOME "a")\ ML \ fun mk_optionT ty = Type (@{type_name "option"}, [ty]) fun mk_None ty = Const (@{const_name "None"}, mk_optionT ty) fun mk_Some ty x = Const (@{const_name "Some"}, ty --> mk_optionT ty) $ x fun mk_option ty _ NONE = mk_None ty | mk_option ty f (SOME x) = mk_Some ty (f x) fun check_lines_tac' s m n ctxt = resolve_tac ctxt - [Thm.instantiate ([], - [("s", @{typ "string option"}, mk_option @{typ string} (using_master_directory_term ctxt) s), + [Thm.instantiate (TVars.empty, + Vars.make ([ + ("s", @{typ "string option"}, mk_option @{typ string} (using_master_directory_term ctxt) s), ("m", @{typ nat}, HOLogic.mk_nat m), ("n", @{typ nat}, HOLogic.mk_nat n)] - |> map (fn (s, ty, t) => (((s, 0), ty), Thm.cterm_of ctxt t))) + |> map (fn (s, ty, t) => (((s, 0), ty), Thm.cterm_of ctxt t)))) @{thm check_linesI}] THEN' CONVERSION (check_line ctxt) THEN' resolve_tac ctxt @{thms TrueI} \ method_setup parallel_check = \ Scan.lift (Parse.maybe Parse.string) -- Scan.lift Parse.nat -- Scan.lift Parse.nat >> (fn ((s, m), n) => fn ctxt => SIMPLE_METHOD' (check_lines_tac' s m n ctxt)) \ lemma lorenz_bounds_lemma_asym: "\x \ N - \. x returns_to \" "R ` (N - \) \ N" "\x \ N - \. (R has_derivative DR x) (at x within \\<^sub>l\<^sub>e)" "\x \ N - \. DR x ` \ x \ \ (R x)" "\x \ N - \. \c \ \ x. norm (DR x c) \ \ x * norm c" "\x \ N - \. \c \ \ x. norm (DR x c) \ \\<^sub>p (R x) * norm c" if NF "Ball (set results) correct_res" using that by (auto intro!: lorenz_bounds_lemma) 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')]); + val inst = Thm.instantiate (TVars.empty, Vars.make [(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)]) + (TVars.empty, Vars.make [((("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)]) + (TVars.empty, Vars.make [((("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)]) + val inst_approx = (TVars.empty, Vars.make [((("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)"}), + (TVars.make [((("'i", 0), @{sort "{enum}"}), mk_numeralT (d + 1) |> Thm.ctyp_of ctxt)], + Vars.make [((("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)"}), + Thm.instantiate (TVars.empty, + Vars.make [((("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)"}), + Thm.instantiate (TVars.empty, + Vars.make [((("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_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/Refine_Imperative_HOL/Lib/User_Smashing.thy b/thys/Refine_Imperative_HOL/Lib/User_Smashing.thy --- a/thys/Refine_Imperative_HOL/Lib/User_Smashing.thy +++ b/thys/Refine_Imperative_HOL/Lib/User_Smashing.thy @@ -1,67 +1,67 @@ theory User_Smashing imports Pure begin (* Alternative flex-flex smasher by Simon Wimmer *) ML \ fun enumerate xs = fold (fn x => fn (i, xs) => (i +1, (x, i) :: xs)) xs (0, []) |> snd \ ML \ fun dummy_abs _ [] t = t | dummy_abs n (T :: Ts) t = Abs ("x" ^ Int.toString n, T, dummy_abs (n + 1) Ts t) \ ML \ fun common_prefix Ts (t1 as Abs (_, T, t)) (u1 as Abs (_, U, u)) = if U = T then common_prefix (T :: Ts) t u else ([], t1, u1) | common_prefix Ts t u = (Ts, t, u); fun dest_app acc (t $ u) = dest_app (u :: acc) t | dest_app acc t = (t, acc); fun add_bound (Bound i, n) bs = (i, n) :: bs | add_bound _ bs = bs; \ ML \ fun smash_pair ctxt thm (t, u) = let val idx = Thm.maxidx_of thm + 1; val ctxt' = ctxt; val (Ts, t1, _) = common_prefix [] t u; val (tas, t2) = Term.strip_abs t; val (uas, u2) = Term.strip_abs u; val (tx as Var (_, T1), ts) = Term.strip_comb t2; val (ux as Var (_, U1), us) = Term.strip_comb u2; val Ts1 = Term.binder_types T1; val Us1 = Term.binder_types U1; val T = Term.fastype_of1 (Ts, t1); val tshift = length tas - length Ts; val ushift = length uas - length Ts; val tbs = fold add_bound (enumerate (rev ts)) [] |> map (apfst (fn i => i - tshift)); val ubs = fold add_bound (enumerate (rev us)) [] |> map (apfst (fn i => i - ushift)); val bounds = inter (op =) (map fst tbs) (map fst ubs) |> distinct (=); val T' = map (nth Ts) bounds ---> T; val v = Var (("simon", idx), T'); val tbs' = map (fn i => find_first (fn (j, _) => i = j) tbs |> the |> snd |> Bound) bounds; val t' = list_comb (v, tbs') |> dummy_abs 0 Ts1; (* Need to add bounds for superfluous abstractions here *) val ubs' = map (fn i => find_first (fn (j, _) => i = j) ubs |> the |> snd |> Bound) bounds; val u' = list_comb (v, ubs') |> dummy_abs 0 Us1; val subst = [(Term.dest_Var tx, Thm.cterm_of ctxt' t'), (Term.dest_Var ux, Thm.cterm_of ctxt' u')]; in - instantiate_normalize ([], subst) thm + instantiate_normalize (TVars.empty, Vars.make subst) thm end; fun smash ctxt thm = case (Thm.tpairs_of thm) of [] => thm | (p :: _) => smash_pair ctxt thm p; fun smashed_attrib ctxt thm = (NONE, SOME (smash ctxt thm)); \ ML \ val smash_new_rule = Seq.single oo smash; \ end diff --git a/thys/Refine_Imperative_HOL/Sepref_Rules.thy b/thys/Refine_Imperative_HOL/Sepref_Rules.thy --- a/thys/Refine_Imperative_HOL/Sepref_Rules.thy +++ b/thys/Refine_Imperative_HOL/Sepref_Rules.thy @@ -1,1745 +1,1745 @@ section \Refinement Rule Management\ theory Sepref_Rules imports Sepref_Basic Sepref_Constraints begin text \This theory contains tools for managing the refinement rules used by Sepref\ text \The theories are based on uncurried functions, i.e., every function has type @{typ "'a\'b"}, where @{typ 'a} is the tuple of parameters, or unit if there are none. \ subsection \Assertion Interface Binding\ text \Binding of interface types to refinement assertions\ definition intf_of_assn :: "('a \ _ \ assn) \ 'b itself \ bool" where [simp]: "intf_of_assn a b = True" lemma intf_of_assnI: "intf_of_assn R TYPE('a)" by simp named_theorems_rev intf_of_assn \Links between refinement assertions and interface types\ lemma intf_of_assn_fallback: "intf_of_assn (R :: 'a \ _ \ assn) TYPE('a)" by simp subsection \Function Refinement with Precondition\ definition fref :: "('c \ bool) \ ('a \ 'c) set \ ('b \ 'd) set \ (('a \ 'b) \ ('c \ 'd)) set" ("[_]\<^sub>f _ \ _" [0,60,60] 60) where "[P]\<^sub>f R \ S \ {(f,g). \x y. P y \ (x,y)\R \ (f x, g y)\S}" abbreviation freft ("_ \\<^sub>f _" [60,60] 60) where "R \\<^sub>f S \ ([\_. True]\<^sub>f R \ S)" lemma rel2p_fref[rel2p]: "rel2p (fref P R S) = (\f g. (\x y. P y \ rel2p R x y \ rel2p S (f x) (g y)))" by (auto simp: fref_def rel2p_def[abs_def]) lemma fref_cons: assumes "(f,g) \ [P]\<^sub>f R \ S" assumes "\c a. (c,a)\R' \ Q a \ P a" assumes "R' \ R" assumes "S \ S'" shows "(f,g) \ [Q]\<^sub>f R' \ S'" using assms unfolding fref_def by fastforce lemmas fref_cons' = fref_cons[OF _ _ order_refl order_refl] lemma frefI[intro?]: assumes "\x y. \P y; (x,y)\R\ \ (f x, g y)\S" shows "(f,g)\fref P R S" using assms unfolding fref_def by auto lemma fref_ncI: "(f,g)\R\S \ (f,g)\R\\<^sub>fS" apply (rule frefI) apply parametricity done lemma frefD: assumes "(f,g)\fref P R S" shows "\P y; (x,y)\R\ \ (f x, g y)\S" using assms unfolding fref_def by auto lemma fref_ncD: "(f,g)\R\\<^sub>fS \ (f,g)\R\S" apply (rule fun_relI) apply (drule frefD) apply simp apply assumption+ done lemma fref_compI: "fref P R1 R2 O fref Q S1 S2 \ fref (\x. Q x \ (\y. (y,x)\S1 \ P y)) (R1 O S1) (R2 O S2)" unfolding fref_def apply (auto) apply blast done lemma fref_compI': "\ (f,g)\fref P R1 R2; (g,h)\fref Q S1 S2 \ \ (f,h) \ fref (\x. Q x \ (\y. (y,x)\S1 \ P y)) (R1 O S1) (R2 O S2)" using fref_compI[of P R1 R2 Q S1 S2] by auto lemma fref_unit_conv: "(\_. c, \_. a) \ fref P unit_rel S \ (P () \ (c,a)\S)" by (auto simp: fref_def) lemma fref_uncurry_conv: "(uncurry c, uncurry a) \ fref P (R1\\<^sub>rR2) S \ (\x1 y1 x2 y2. P (y1,y2) \ (x1,y1)\R1 \ (x2,y2)\R2 \ (c x1 x2, a y1 y2) \ S)" by (auto simp: fref_def) lemma fref_mono: "\ \x. P' x \ P x; R' \ R; S \ S' \ \ fref P R S \ fref P' R' S'" unfolding fref_def by auto blast lemma fref_composeI: assumes FR1: "(f,g)\fref P R1 R2" assumes FR2: "(g,h)\fref Q S1 S2" assumes C1: "\x. P' x \ Q x" assumes C2: "\x y. \P' x; (y,x)\S1\ \ P y" assumes R1: "R' \ R1 O S1" assumes R2: "R2 O S2 \ S'" assumes FH: "f'=f" "h'=h" shows "(f',h') \ fref P' R' S'" unfolding FH apply (rule subsetD[OF fref_mono fref_compI'[OF FR1 FR2]]) using C1 C2 apply blast using R1 apply blast using R2 apply blast done lemma fref_triv: "A\Id \ (f,f)\[P]\<^sub>f A \ Id" by (auto simp: fref_def) subsection \Heap-Function Refinement\ text \ The following relates a heap-function with a pure function. It contains a precondition, a refinement assertion for the arguments before and after execution, and a refinement relation for the result. \ (* TODO: We only use this with keep/destroy information, so we could model the parameter relations as such (('a\'ai \ assn) \ bool) *) definition hfref :: " ('a \ bool) \ (('a \ 'ai \ assn) \ ('a \ 'ai \ assn)) \ ('b \ 'bi \ assn) \ (('ai \ 'bi Heap) \ ('a\'b nres)) set" ("[_]\<^sub>a _ \ _" [0,60,60] 60) where "[P]\<^sub>a RS \ T \ { (f,g) . \c a. P a \ hn_refine (fst RS a c) (f c) (snd RS a c) T (g a)}" abbreviation hfreft ("_ \\<^sub>a _" [60,60] 60) where "RS \\<^sub>a T \ ([\_. True]\<^sub>a RS \ T)" lemma hfrefI[intro?]: assumes "\c a. P a \ hn_refine (fst RS a c) (f c) (snd RS a c) T (g a)" shows "(f,g)\hfref P RS T" using assms unfolding hfref_def by blast lemma hfrefD: assumes "(f,g)\hfref P RS T" shows "\c a. P a \ hn_refine (fst RS a c) (f c) (snd RS a c) T (g a)" using assms unfolding hfref_def by blast lemma hfref_to_ASSERT_conv: "NO_MATCH (\_. True) P \ (a,b)\[P]\<^sub>a R \ S \ (a,\x. ASSERT (P x) \ b x) \ R \\<^sub>a S" unfolding hfref_def apply (clarsimp; safe; clarsimp?) apply (rule hn_refine_nofailI) apply (simp add: refine_pw_simps) subgoal for xc xa apply (drule spec[of _ xc]) apply (drule spec[of _ xa]) by simp done text \ A pair of argument refinement assertions can be created by the input assertion and the information whether the parameter is kept or destroyed by the function. \ primrec hf_pres :: "('a \ 'b \ assn) \ bool \ ('a \ 'b \ assn)\('a \ 'b \ assn)" where "hf_pres R True = (R,R)" | "hf_pres R False = (R,invalid_assn R)" abbreviation hfkeep :: "('a \ 'b \ assn) \ ('a \ 'b \ assn)\('a \ 'b \ assn)" ("(_\<^sup>k)" [1000] 999) where "R\<^sup>k \ hf_pres R True" abbreviation hfdrop :: "('a \ 'b \ assn) \ ('a \ 'b \ assn)\('a \ 'b \ assn)" ("(_\<^sup>d)" [1000] 999) where "R\<^sup>d \ hf_pres R False" abbreviation "hn_kede R kd \ hn_ctxt (snd (hf_pres R kd))" abbreviation "hn_keep R \ hn_kede R True" abbreviation "hn_dest R \ hn_kede R False" lemma keep_drop_sels[simp]: "fst (R\<^sup>k) = R" "snd (R\<^sup>k) = R" "fst (R\<^sup>d) = R" "snd (R\<^sup>d) = invalid_assn R" by auto lemma hf_pres_fst[simp]: "fst (hf_pres R k) = R" by (cases k) auto text \ The following operator combines multiple argument assertion-pairs to argument assertion-pairs for the product. It is required to state argument assertion-pairs for uncurried functions. \ definition hfprod :: " (('a \ 'b \ assn)\('a \ 'b \ assn)) \ (('c \ 'd \ assn)\('c \ 'd \ assn)) \ ((('a\'c) \ ('b \ 'd) \ assn) \ (('a\'c) \ ('b \ 'd) \ assn))" (infixl "*\<^sub>a" 65) where "RR *\<^sub>a SS \ (prod_assn (fst RR) (fst SS), prod_assn (snd RR) (snd SS))" lemma hfprod_fst_snd[simp]: "fst (A *\<^sub>a B) = prod_assn (fst A) (fst B)" "snd (A *\<^sub>a B) = prod_assn (snd A) (snd B)" unfolding hfprod_def by auto subsubsection \Conversion from fref to hfref\ (* TODO: Variant of import-param! Automate this! *) lemma fref_to_pure_hfref': assumes "(f,g) \ [P]\<^sub>f R\\S\nres_rel" assumes "\x. x\Domain R \ R\``Collect P \ f x = RETURN (f' x)" shows "(return o f', g) \ [P]\<^sub>a (pure R)\<^sup>k\pure S" apply (rule hfrefI) apply (rule hn_refineI) using assms apply ((sep_auto simp: fref_def pure_def pw_le_iff pw_nres_rel_iff refine_pw_simps eintros del: exI)) apply force done subsubsection \Conversion from hfref to hnr\ text \This section contains the lemmas. The ML code is further down. \ lemma hf2hnr: assumes "(f,g) \ [P]\<^sub>a R \ S" shows "\x xi. P x \ hn_refine (emp * hn_ctxt (fst R) x xi) (f$xi) (emp * hn_ctxt (snd R) x xi) S (g$x)" using assms unfolding hfref_def by (auto simp: hn_ctxt_def) (*lemma hf2hnr_new: assumes "(f,g) \ [P]\<^sub>a R \ S" shows "\x xi. (\h. h\fst R x xi \ P x) \ hn_refine (emp * hn_ctxt (fst R) x xi) (f xi) (emp * hn_ctxt (snd R) x xi) S (g$x)" using assms unfolding hfref_def by (auto simp: hn_ctxt_def intro: hn_refine_preI) *) (* Products that stem from currying are tagged by a special refinement relation *) definition [simp]: "to_hnr_prod \ prod_assn" lemma to_hnr_prod_fst_snd: "fst (A *\<^sub>a B) = to_hnr_prod (fst A) (fst B)" "snd (A *\<^sub>a B) = to_hnr_prod (snd A) (snd B)" unfolding hfprod_def by auto (* Warning: This lemma is carefully set up to be applicable as an unfold rule, for more than one level of uncurrying*) lemma hnr_uncurry_unfold: " (\x xi. P x \ hn_refine (\ * hn_ctxt (to_hnr_prod A B) x xi) (fi xi) (\' * hn_ctxt (to_hnr_prod A' B') x xi) R (f x)) \ (\b bi a ai. P (a,b) \ hn_refine (\ * hn_ctxt B b bi * hn_ctxt A a ai) (fi (ai,bi)) (\' * hn_ctxt B' b bi * hn_ctxt A' a ai) R (f (a,b)) )" by (auto simp: hn_ctxt_def prod_assn_def star_aci) lemma hnr_intro_dummy: "\x xi. P x \ hn_refine (\ x xi) (c xi) (\' x xi) R (a x) \ \x xi. P x \ hn_refine (emp*\ x xi) (c xi) (emp*\' x xi) R (a x)" by simp lemma hn_ctxt_ctxt_fix_conv: "hn_ctxt (hn_ctxt R) = hn_ctxt R" by (simp add: hn_ctxt_def[abs_def]) lemma uncurry_APP: "uncurry f$(a,b) = f$a$b" by auto (* TODO: Replace by more general rule. *) lemma norm_RETURN_o: "\f. (RETURN o f)$x = (RETURN$(f$x))" "\f. (RETURN oo f)$x$y = (RETURN$(f$x$y))" "\f. (RETURN ooo f)$x$y$z = (RETURN$(f$x$y$z))" "\f. (\x. RETURN ooo f x)$x$y$z$a = (RETURN$(f$x$y$z$a))" "\f. (\x y. RETURN ooo f x y)$x$y$z$a$b = (RETURN$(f$x$y$z$a$b))" by auto lemma norm_return_o: "\f. (return o f)$x = (return$(f$x))" "\f. (return oo f)$x$y = (return$(f$x$y))" "\f. (return ooo f)$x$y$z = (return$(f$x$y$z))" "\f. (\x. return ooo f x)$x$y$z$a = (return$(f$x$y$z$a))" "\f. (\x y. return ooo f x y)$x$y$z$a$b = (return$(f$x$y$z$a$b))" by auto lemma hn_val_unit_conv_emp[simp]: "hn_val unit_rel x y = emp" by (auto simp: hn_ctxt_def pure_def) subsubsection \Conversion from hnr to hfref\ text \This section contains the lemmas. The ML code is further down. \ abbreviation "id_assn \ pure Id" abbreviation "unit_assn \ id_assn :: unit \ _" lemma pure_unit_rel_eq_empty: "unit_assn x y = emp" by (auto simp: pure_def) lemma uc_hfprod_sel: "fst (A *\<^sub>a B) a c = (case (a,c) of ((a1,a2),(c1,c2)) \ fst A a1 c1 * fst B a2 c2)" "snd (A *\<^sub>a B) a c = (case (a,c) of ((a1,a2),(c1,c2)) \ snd A a1 c1 * snd B a2 c2)" unfolding hfprod_def prod_assn_def[abs_def] by auto subsubsection \Conversion from relation to fref\ text \This section contains the lemmas. The ML code is further down. \ definition "CURRY R \ { (f,g). (uncurry f, uncurry g) \ R }" lemma fref_param1: "R\S = fref (\_. True) R S" by (auto simp: fref_def fun_relD) lemma fref_nest: "fref P1 R1 (fref P2 R2 S) \ CURRY (fref (\(a,b). P1 a \ P2 b) (R1\\<^sub>rR2) S)" apply (rule eq_reflection) by (auto simp: fref_def CURRY_def) lemma in_CURRY_conv: "(f,g) \ CURRY R \ (uncurry f, uncurry g) \ R" unfolding CURRY_def by auto lemma uncurry0_APP[simp]: "uncurry0 c $ x = c" by auto lemma fref_param0I: "(c,a)\R \ (uncurry0 c, uncurry0 a) \ fref (\_. True) unit_rel R" by (auto simp: fref_def) subsubsection \Composition\ definition hr_comp :: "('b \ 'c \ assn) \ ('b \ 'a) set \ 'a \ 'c \ assn" \ \Compose refinement assertion with refinement relation\ where "hr_comp R1 R2 a c \ \\<^sub>Ab. R1 b c * \((b,a)\R2)" definition hrp_comp :: "('d \ 'b \ assn) \ ('d \ 'c \ assn) \ ('d \ 'a) set \ ('a \ 'b \ assn) \ ('a \ 'c \ assn)" \ \Compose argument assertion-pair with refinement relation\ where "hrp_comp RR' S \ (hr_comp (fst RR') S, hr_comp (snd RR') S) " lemma hr_compI: "(b,a)\R2 \ R1 b c \\<^sub>A hr_comp R1 R2 a c" unfolding hr_comp_def by sep_auto lemma hr_comp_Id1[simp]: "hr_comp (pure Id) R = pure R" unfolding hr_comp_def[abs_def] pure_def apply (intro ext ent_iffI) by sep_auto+ lemma hr_comp_Id2[simp]: "hr_comp R Id = R" unfolding hr_comp_def[abs_def] apply (intro ext ent_iffI) by sep_auto+ (*lemma hr_comp_invalid[simp]: "hr_comp (\a c. true) R a c = true * \(\b. (b,a)\R)" unfolding hr_comp_def[abs_def] apply (intro ext ent_iffI) apply sep_auto+ done*) lemma hr_comp_emp[simp]: "hr_comp (\a c. emp) R a c = \(\b. (b,a)\R)" unfolding hr_comp_def[abs_def] apply (intro ext ent_iffI) apply sep_auto+ done lemma hr_comp_prod_conv[simp]: "hr_comp (prod_assn Ra Rb) (Ra' \\<^sub>r Rb') = prod_assn (hr_comp Ra Ra') (hr_comp Rb Rb')" unfolding hr_comp_def[abs_def] prod_assn_def[abs_def] apply (intro ext ent_iffI) apply solve_entails apply clarsimp apply sep_auto apply clarsimp apply (intro ent_ex_preI) apply (rule ent_ex_postI) apply (sep_auto split: prod.splits) done lemma hr_comp_pure: "hr_comp (pure R) S = pure (R O S)" apply (intro ext) apply (rule ent_iffI) unfolding hr_comp_def[abs_def] apply (sep_auto simp: pure_def)+ done lemma hr_comp_is_pure[safe_constraint_rules]: "is_pure A \ is_pure (hr_comp A B)" by (auto simp: hr_comp_pure is_pure_conv) lemma hr_comp_the_pure: "is_pure A \ the_pure (hr_comp A B) = the_pure A O B" unfolding is_pure_conv by (clarsimp simp: hr_comp_pure) lemma rdomp_hrcomp_conv: "rdomp (hr_comp A R) x \ (\y. rdomp A y \ (y,x)\R)" by (auto simp: rdomp_def hr_comp_def) lemma hn_rel_compI: "\nofail a; (b,a)\\R2\nres_rel\ \ hn_rel R1 b c \\<^sub>A hn_rel (hr_comp R1 R2) a c" unfolding hr_comp_def hn_rel_def nres_rel_def apply (clarsimp intro!: ent_ex_preI) apply (drule (1) order_trans) apply (simp add: ret_le_down_conv) by sep_auto lemma hr_comp_precise[constraint_rules]: assumes [safe_constraint_rules]: "precise R" assumes SV: "single_valued S" shows "precise (hr_comp R S)" apply (rule preciseI) unfolding hr_comp_def apply clarsimp by (metis SV assms(1) preciseD single_valuedD) lemma hr_comp_assoc: "hr_comp (hr_comp R S) T = hr_comp R (S O T)" apply (intro ext) unfolding hr_comp_def apply (rule ent_iffI; clarsimp) apply sep_auto apply (rule ent_ex_preI; clarsimp) (* TODO: sep_auto/solve_entails is too eager splitting the subgoal here! *) apply sep_auto done lemma hnr_comp: assumes R: "\b1 c1. P b1 \ hn_refine (R1 b1 c1 * \) (c c1) (R1p b1 c1 * \') R (b b1)" assumes S: "\a1 b1. \Q a1; (b1,a1)\R1'\ \ (b b1,a a1)\\R'\nres_rel" assumes PQ: "\a1 b1. \Q a1; (b1,a1)\R1'\ \ P b1" assumes Q: "Q a1" shows "hn_refine (hr_comp R1 R1' a1 c1 * \) (c c1) (hr_comp R1p R1' a1 c1 * \') (hr_comp R R') (a a1)" unfolding hn_refine_alt proof clarsimp assume NF: "nofail (a a1)" show " > c c1 <\r. hn_rel (hr_comp R R') (a a1) r * (hr_comp R1p R1' a1 c1 * \')>\<^sub>t" apply (subst hr_comp_def) apply (clarsimp intro!: norm_pre_ex_rule) proof - fix b1 assume R1: "(b1, a1) \ R1'" from S R1 Q have R': "(b b1, a a1) \ \R'\nres_rel" by blast with NF have NFB: "nofail (b b1)" by (simp add: nres_rel_def pw_le_iff refine_pw_simps) from PQ R1 Q have P: "P b1" by blast with NFB R have "> c c1 <\r. hn_rel R (b b1) r * (R1p b1 c1 * \')>\<^sub>t" unfolding hn_refine_alt by auto thus "> c c1 <\r. hn_rel (hr_comp R R') (a a1) r * (hr_comp R1p R1' a1 c1 * \')>\<^sub>t" apply (rule cons_post_rule) apply (solve_entails) by (intro ent_star_mono hn_rel_compI[OF NF R'] hr_compI[OF R1] ent_refl) qed qed lemma hnr_comp1_aux: assumes R: "\b1 c1. P b1 \ hn_refine (hn_ctxt R1 b1 c1) (c c1) (hn_ctxt R1p b1 c1) R (b$b1)" assumes S: "\a1 b1. \Q a1; (b1,a1)\R1'\ \ (b$b1,a$a1)\\R'\nres_rel" assumes PQ: "\a1 b1. \Q a1; (b1,a1)\R1'\ \ P b1" assumes Q: "Q a1" shows "hn_refine (hr_comp R1 R1' a1 c1) (c c1) (hr_comp R1p R1' a1 c1) (hr_comp R R') (a a1)" using assms hnr_comp[where \=emp and \'=emp and a=a and b=b and c=c and P=P and Q=Q] unfolding hn_ctxt_def by auto lemma hfcomp: assumes A: "(f,g) \ [P]\<^sub>a RR' \ S" assumes B: "(g,h) \ [Q]\<^sub>f T \ \U\nres_rel" shows "(f,h) \ [\a. Q a \ (\a'. (a',a)\T \ P a')]\<^sub>a hrp_comp RR' T \ hr_comp S U" using assms unfolding fref_def hfref_def hrp_comp_def apply clarsimp apply (rule hnr_comp1_aux[of P "fst RR'" f "snd RR'" S g "\a. Q a \ (\a'. (a',a)\T \ P a')" T h U]) apply (auto simp: hn_ctxt_def) done lemma hfref_weaken_pre_nofail: assumes "(f,g) \ [P]\<^sub>a R \ S" shows "(f,g) \ [\x. nofail (g x) \ P x]\<^sub>a R \ S" using assms unfolding hfref_def hn_refine_def by auto lemma hfref_cons: assumes "(f,g) \ [P]\<^sub>a R \ S" assumes "\x. P' x \ P x" assumes "\x y. fst R' x y \\<^sub>t fst R x y" assumes "\x y. snd R x y \\<^sub>t snd R' x y" assumes "\x y. S x y \\<^sub>t S' x y" shows "(f,g) \ [P']\<^sub>a R' \ S'" unfolding hfref_def apply clarsimp apply (rule hn_refine_cons) apply (rule assms(3)) defer apply (rule entt_trans[OF assms(4)]; sep_auto) apply (rule assms(5)) apply (frule assms(2)) using assms(1) unfolding hfref_def apply auto done subsubsection \Composition Automation\ text \This section contains the lemmas. The ML code is further down. \ lemma prod_hrp_comp: "hrp_comp (A *\<^sub>a B) (C \\<^sub>r D) = hrp_comp A C *\<^sub>a hrp_comp B D" unfolding hrp_comp_def hfprod_def by simp lemma hrp_comp_keep: "hrp_comp (A\<^sup>k) B = (hr_comp A B)\<^sup>k" by (auto simp: hrp_comp_def) lemma hr_comp_invalid: "hr_comp (invalid_assn R1) R2 = invalid_assn (hr_comp R1 R2)" apply (intro ent_iffI entailsI ext) unfolding invalid_assn_def hr_comp_def by auto lemma hrp_comp_dest: "hrp_comp (A\<^sup>d) B = (hr_comp A B)\<^sup>d" by (auto simp: hrp_comp_def hr_comp_invalid) definition "hrp_imp RR RR' \ \a b. (fst RR' a b \\<^sub>t fst RR a b) \ (snd RR a b \\<^sub>t snd RR' a b)" lemma hfref_imp: "hrp_imp RR RR' \ [P]\<^sub>a RR \ S \ [P]\<^sub>a RR' \ S" apply clarsimp apply (erule hfref_cons) apply (simp_all add: hrp_imp_def) done lemma hrp_imp_refl: "hrp_imp RR RR" unfolding hrp_imp_def by auto lemma hrp_imp_reflI: "RR = RR' \ hrp_imp RR RR'" unfolding hrp_imp_def by auto lemma hrp_comp_cong: "hrp_imp A A' \ B=B' \ hrp_imp (hrp_comp A B) (hrp_comp A' B')" by (sep_auto simp: hrp_imp_def hrp_comp_def hr_comp_def entailst_def) lemma hrp_prod_cong: "hrp_imp A A' \ hrp_imp B B' \ hrp_imp (A*\<^sub>aB) (A'*\<^sub>aB')" by (sep_auto simp: hrp_imp_def prod_assn_def intro: entt_star_mono) lemma hrp_imp_trans: "hrp_imp A B \ hrp_imp B C \ hrp_imp A C" unfolding hrp_imp_def by (fastforce intro: entt_trans) lemma fcomp_norm_dflt_init: "x\[P]\<^sub>a R \ T \ hrp_imp R S \ x\[P]\<^sub>a S \ T" apply (erule rev_subsetD) by (rule hfref_imp) definition "comp_PRE R P Q S \ \x. S x \ (P x \ (\y. (y,x)\R \ Q x y))" lemma comp_PRE_cong[cong]: assumes "R\R'" assumes "\x. P x \ P' x" assumes "\x. S x \ S' x" assumes "\x y. \P x; (y,x)\R; y\Domain R; S' x \ \ Q x y \ Q' x y" shows "comp_PRE R P Q S \ comp_PRE R' P' Q' S'" using assms by (fastforce simp: comp_PRE_def intro!: eq_reflection ext) lemma fref_compI_PRE: "\ (f,g)\fref P R1 R2; (g,h)\fref Q S1 S2 \ \ (f,h) \ fref (comp_PRE S1 Q (\_. P) (\_. True)) (R1 O S1) (R2 O S2)" using fref_compI[of P R1 R2 Q S1 S2] unfolding comp_PRE_def by auto lemma PRE_D1: "(Q x \ P x) \ comp_PRE S1 Q (\x _. P x) S x" by (auto simp: comp_PRE_def) lemma PRE_D2: "(Q x \ (\y. (y,x)\S1 \ S x \ P x y)) \ comp_PRE S1 Q P S x" by (auto simp: comp_PRE_def) lemma fref_weaken_pre: assumes "\x. P x \ P' x" assumes "(f,h) \ fref P' R S" shows "(f,h) \ fref P R S" apply (rule rev_subsetD[OF assms(2) fref_mono]) using assms(1) by auto lemma fref_PRE_D1: assumes "(f,h) \ fref (comp_PRE S1 Q (\x _. P x) X) R S" shows "(f,h) \ fref (\x. Q x \ P x) R S" by (rule fref_weaken_pre[OF PRE_D1 assms]) lemma fref_PRE_D2: assumes "(f,h) \ fref (comp_PRE S1 Q P X) R S" shows "(f,h) \ fref (\x. Q x \ (\y. (y,x)\S1 \ X x \ P x y)) R S" by (rule fref_weaken_pre[OF PRE_D2 assms]) lemmas fref_PRE_D = fref_PRE_D1 fref_PRE_D2 lemma hfref_weaken_pre: assumes "\x. P x \ P' x" assumes "(f,h) \ hfref P' R S" shows "(f,h) \ hfref P R S" using assms by (auto simp: hfref_def) lemma hfref_weaken_pre': assumes "\x. \P x; rdomp (fst R) x\ \ P' x" assumes "(f,h) \ hfref P' R S" shows "(f,h) \ hfref P R S" apply (rule hfrefI) apply (rule hn_refine_preI) using assms by (auto simp: hfref_def rdomp_def) lemma hfref_weaken_pre_nofail': assumes "(f,g) \ [P]\<^sub>a R \ S" assumes "\x. \nofail (g x); Q x\ \ P x" shows "(f,g) \ [Q]\<^sub>a R \ S" apply (rule hfref_weaken_pre[OF _ assms(1)[THEN hfref_weaken_pre_nofail]]) using assms(2) by blast lemma hfref_compI_PRE_aux: assumes A: "(f,g) \ [P]\<^sub>a RR' \ S" assumes B: "(g,h) \ [Q]\<^sub>f T \ \U\nres_rel" shows "(f,h) \ [comp_PRE T Q (\_. P) (\_. True)]\<^sub>a hrp_comp RR' T \ hr_comp S U" apply (rule hfref_weaken_pre[OF _ hfcomp[OF A B]]) by (auto simp: comp_PRE_def) lemma hfref_compI_PRE: assumes A: "(f,g) \ [P]\<^sub>a RR' \ S" assumes B: "(g,h) \ [Q]\<^sub>f T \ \U\nres_rel" shows "(f,h) \ [comp_PRE T Q (\x y. P y) (\x. nofail (h x))]\<^sub>a hrp_comp RR' T \ hr_comp S U" using hfref_compI_PRE_aux[OF A B, THEN hfref_weaken_pre_nofail] apply (rule hfref_weaken_pre[rotated]) apply (auto simp: comp_PRE_def) done lemma hfref_PRE_D1: assumes "(f,h) \ hfref (comp_PRE S1 Q (\x _. P x) X) R S" shows "(f,h) \ hfref (\x. Q x \ P x) R S" by (rule hfref_weaken_pre[OF PRE_D1 assms]) lemma hfref_PRE_D2: assumes "(f,h) \ hfref (comp_PRE S1 Q P X) R S" shows "(f,h) \ hfref (\x. Q x \ (\y. (y,x)\S1 \ X x \ P x y)) R S" by (rule hfref_weaken_pre[OF PRE_D2 assms]) lemma hfref_PRE_D3: assumes "(f,h) \ hfref (comp_PRE S1 Q P X) R S" shows "(f,h) \ hfref (comp_PRE S1 Q P X) R S" using assms . lemmas hfref_PRE_D = hfref_PRE_D1 hfref_PRE_D3 subsection \Automation\ text \Purity configuration for constraint solver\ lemmas [safe_constraint_rules] = pure_pure text \Configuration for hfref to hnr conversion\ named_theorems to_hnr_post \to_hnr converter: Postprocessing unfold rules\ lemma uncurry0_add_app_tag: "uncurry0 (RETURN c) = uncurry0 (RETURN$c)" by simp lemmas [to_hnr_post] = norm_RETURN_o norm_return_o uncurry0_add_app_tag uncurry0_apply uncurry0_APP hn_val_unit_conv_emp mult_1[of "x::assn" for x] mult_1_right[of "x::assn" for x] named_theorems to_hfref_post \to_hfref converter: Postprocessing unfold rules\ lemma prod_casesK[to_hfref_post]: "case_prod (\_ _. k) = (\_. k)" by auto lemma uncurry0_hfref_post[to_hfref_post]: "hfref (uncurry0 True) R S = hfref (\_. True) R S" apply (fo_rule arg_cong fun_cong)+ by auto (* Currently not used, we keep it in here anyway. *) text \Configuration for relation normalization after composition\ named_theorems fcomp_norm_unfold \fcomp-normalizer: Unfold theorems\ named_theorems fcomp_norm_simps \fcomp-normalizer: Simplification theorems\ named_theorems fcomp_norm_init "fcomp-normalizer: Initialization rules" named_theorems fcomp_norm_trans "fcomp-normalizer: Transitivity rules" named_theorems fcomp_norm_cong "fcomp-normalizer: Congruence rules" named_theorems fcomp_norm_norm "fcomp-normalizer: Normalization rules" named_theorems fcomp_norm_refl "fcomp-normalizer: Reflexivity rules" text \Default Setup\ lemmas [fcomp_norm_unfold] = prod_rel_comp nres_rel_comp Id_O_R R_O_Id lemmas [fcomp_norm_unfold] = hr_comp_Id1 hr_comp_Id2 lemmas [fcomp_norm_unfold] = hr_comp_prod_conv lemmas [fcomp_norm_unfold] = prod_hrp_comp hrp_comp_keep hrp_comp_dest hr_comp_pure (*lemmas [fcomp_norm_unfold] = prod_casesK uncurry0_hfref_post*) lemma [fcomp_norm_simps]: "CONSTRAINT is_pure P \ pure (the_pure P) = P" by simp lemmas [fcomp_norm_simps] = True_implies_equals lemmas [fcomp_norm_init] = fcomp_norm_dflt_init lemmas [fcomp_norm_trans] = hrp_imp_trans lemmas [fcomp_norm_cong] = hrp_comp_cong hrp_prod_cong (*lemmas [fcomp_norm_norm] = hrp_comp_dest*) lemmas [fcomp_norm_refl] = refl hrp_imp_refl lemma ensure_fref_nresI: "(f,g)\[P]\<^sub>f R\S \ (RETURN o f, RETURN o g)\[P]\<^sub>f R\\S\nres_rel" by (auto intro: nres_relI simp: fref_def) lemma ensure_fref_nres_unfold: "\f. RETURN o (uncurry0 f) = uncurry0 (RETURN f)" "\f. RETURN o (uncurry f) = uncurry (RETURN oo f)" "\f. (RETURN ooo uncurry) f = uncurry (RETURN ooo f)" by auto text \Composed precondition normalizer\ named_theorems fcomp_prenorm_simps \fcomp precondition-normalizer: Simplification theorems\ text \Support for preconditions of the form \_\Domain R\, where \R\ is the relation of the next more abstract level.\ declare DomainI[fcomp_prenorm_simps] lemma auto_weaken_pre_init_hf: assumes "\x. PROTECT P x \ P' x" assumes "(f,h) \ hfref P' R S" shows "(f,h) \ hfref P R S" using assms by (auto simp: hfref_def) lemma auto_weaken_pre_init_f: assumes "\x. PROTECT P x \ P' x" assumes "(f,h) \ fref P' R S" shows "(f,h) \ fref P R S" using assms by (auto simp: fref_def) lemmas auto_weaken_pre_init = auto_weaken_pre_init_hf auto_weaken_pre_init_f lemma auto_weaken_pre_uncurry_step: assumes "PROTECT f a \ f'" shows "PROTECT (\(x,y). f x y) (a,b) \ f' b" using assms by (auto simp: curry_def dest!: meta_eq_to_obj_eq intro!: eq_reflection) lemma auto_weaken_pre_uncurry_finish: "PROTECT f x \ f x" by (auto) lemma auto_weaken_pre_uncurry_start: assumes "P \ P'" assumes "P'\Q" shows "P\Q" using assms by (auto) lemma auto_weaken_pre_comp_PRE_I: assumes "S x \ P x" assumes "\y. \(y,x)\R; P x; S x\ \ Q x y" shows "comp_PRE R P Q S x" using assms by (auto simp: comp_PRE_def) lemma auto_weaken_pre_to_imp_nf: "(A\B\C) = (A\B \ C)" "((A\B)\C) = (A\B\C)" by auto lemma auto_weaken_pre_add_dummy_imp: "P \ True \ P" by simp text \Synthesis for hfref statements\ definition hfsynth_ID_R :: "('a \ _ \ assn) \ 'a \ bool" where [simp]: "hfsynth_ID_R _ _ \ True" lemma hfsynth_ID_R_D: fixes I :: "'a itself" assumes "hfsynth_ID_R R a" assumes "intf_of_assn R I" shows "a ::\<^sub>i I" by simp lemma hfsynth_hnr_from_hfI: assumes "\x xi. P x \ hfsynth_ID_R (fst R) x \ hn_refine (emp * hn_ctxt (fst R) x xi) (f$xi) (emp * hn_ctxt (snd R) x xi) S (g$x)" shows "(f,g) \ [P]\<^sub>a R \ S" using assms unfolding hfref_def by (auto simp: hn_ctxt_def) lemma hfsynth_ID_R_uncurry_unfold: "hfsynth_ID_R (to_hnr_prod R S) (a,b) \ hfsynth_ID_R R a \ hfsynth_ID_R S b" "hfsynth_ID_R (fst (hf_pres R k)) \ hfsynth_ID_R R" by (auto intro!: eq_reflection) ML \ signature SEPREF_RULES = sig (* Analysis of relations, both fref and fun_rel *) (* "R1\...\Rn\_" / "[_]\<^sub>f ((R1\\<^sub>rR2)...\\<^sub>rRn)" \ "[R1,...,Rn]" *) val binder_rels: term -> term list (* "_\...\_\S" / "[_]\<^sub>f _ \ S" \ "S" *) val body_rel: term -> term (* Map \/fref to (precond,args,res). NONE if no/trivial precond. *) val analyze_rel: term -> term option * term list * term (* Make trivial ("\_. True") precond *) val mk_triv_precond: term list -> term (* Make "[P]\<^sub>f ((R1\\<^sub>rR2)...\\<^sub>rRn) \ S". Insert trivial precond if NONE. *) val mk_rel: term option * term list * term -> term (* Map relation to (args,res) *) val strip_rel: term -> term list * term (* Make hfprod (op *\<^sub>a) *) val mk_hfprod : term * term -> term val mk_hfprods : term list -> term (* Determine interface type of refinement assertion, using default fallback if necessary. Use named_thms intf_of_assn for configuration. *) val intf_of_assn : Proof.context -> term -> typ (* Convert a parametricity theorem in higher-order form to uncurried fref-form. For functions without arguments, a unit-argument is added. TODO/FIXME: Currently this only works for higher-order theorems, i.e., theorems of the form (f,g)\R1\\\Rn. First-order theorems are silently treated as refinement theorems for functions with zero arguments, i.e., a unit-argument is added. *) val to_fref : Proof.context -> thm -> thm (* Convert a parametricity or fref theorem to first order form *) val to_foparam : Proof.context -> thm -> thm (* Convert schematic hfref goal to hnr-goal *) val prepare_hfref_synth_tac : Proof.context -> tactic' (* Convert theorem in hfref-form to hnr-form *) val to_hnr : Proof.context -> thm -> thm (* Convert theorem in hnr-form to hfref-form *) val to_hfref: Proof.context -> thm -> thm (* Convert theorem to given form, if not yet in this form *) val ensure_fref : Proof.context -> thm -> thm val ensure_fref_nres : Proof.context -> thm -> thm val ensure_hfref : Proof.context -> thm -> thm val ensure_hnr : Proof.context -> thm -> thm type hnr_analysis = { thm: thm, (* Original theorem, may be normalized *) precond: term, (* Precondition, abstracted over abs-arguments *) prems : term list, (* Premises not depending on arguments *) ahead: term * bool, (* Abstract function, has leading RETURN *) chead: term * bool, (* Concrete function, has leading return *) argrels: (term * bool) list, (* Argument relations, preserved (keep-flag) *) result_rel: term (* Result relation *) } val analyze_hnr: Proof.context -> thm -> hnr_analysis val pretty_hnr_analysis: Proof.context -> hnr_analysis -> Pretty.T val mk_hfref_thm: Proof.context -> hnr_analysis -> thm (* Simplify precondition of fref/hfref-theorem *) val simplify_precond: Proof.context -> thm -> thm (* Normalize hfref-theorem after composition *) val norm_fcomp_rule: Proof.context -> thm -> thm (* Replace "pure ?A" by "?A'" and is_pure constraint, then normalize *) val add_pure_constraints_rule: Proof.context -> thm -> thm (* Compose fref/hfref and fref theorem, to produce hfref theorem. The input theorems may also be in ho-param or hnr form, and are converted accordingly. *) val gen_compose : Proof.context -> thm -> thm -> thm (* FCOMP-attribute *) val fcomp_attrib: attribute context_parser end structure Sepref_Rules: SEPREF_RULES = struct local open Refine_Util Relators in fun binder_rels @{mpat "?F \ ?G"} = F::binder_rels G | binder_rels @{mpat "fref _ ?F _"} = strip_prodrel_left F | binder_rels _ = [] local fun br_aux @{mpat "_ \ ?G"} = br_aux G | br_aux R = R in fun body_rel @{mpat "fref _ _ ?G"} = G | body_rel R = br_aux R end fun strip_rel R = (binder_rels R, body_rel R) fun analyze_rel @{mpat "fref (\_. True) ?R ?S"} = (NONE,strip_prodrel_left R,S) | analyze_rel @{mpat "fref ?P ?R ?S"} = (SOME P,strip_prodrel_left R,S) | analyze_rel R = let val (args,res) = strip_rel R in (NONE,args,res) end fun mk_triv_precond Rs = absdummy (map rel_absT Rs |> list_prodT_left) @{term True} fun mk_rel (P,Rs,S) = let val R = list_prodrel_left Rs val P = case P of SOME P => P | NONE => mk_triv_precond Rs in @{mk_term "fref ?P ?R ?S"} end end fun mk_hfprod (a, b) = @{mk_term "?a*\<^sub>a?b"} local fun mk_hfprods_rev [] = @{mk_term "unit_assn\<^sup>k"} | mk_hfprods_rev [Rk] = Rk | mk_hfprods_rev (Rkn::Rks) = mk_hfprod (mk_hfprods_rev Rks, Rkn) in val mk_hfprods = mk_hfprods_rev o rev end fun intf_of_assn ctxt t = let val orig_ctxt = ctxt val (t,ctxt) = yield_singleton (Variable.import_terms false) t ctxt val v = TVar (("T",0),Proof_Context.default_sort ctxt ("T",0)) |> Logic.mk_type val goal = @{mk_term "Trueprop (intf_of_assn ?t ?v)"} val i_of_assn_rls = Named_Theorems_Rev.get ctxt @{named_theorems_rev intf_of_assn} @ @{thms intf_of_assn_fallback} fun tac ctxt = REPEAT_ALL_NEW (resolve_tac ctxt i_of_assn_rls) val thm = Goal.prove ctxt [] [] goal (fn {context,...} => ALLGOALS (tac context)) val intf = case Thm.concl_of thm of @{mpat "Trueprop (intf_of_assn _ (?v AS\<^sub>p TYPE (_)))"} => v | _ => raise THM("Intf_of_assn: Proved a different theorem?",~1,[thm]) val intf = singleton (Variable.export_terms ctxt orig_ctxt) intf |> Logic.dest_type in intf end datatype rthm_type = RT_HOPARAM (* (_,_) \ _ \ \ \ _ *) | RT_FREF (* (_,_) \ [_]\<^sub>f _ \ _ *) | RT_HNR (* hn_refine _ _ _ _ _ *) | RT_HFREF (* (_,_) \ [_]\<^sub>a _ \ _ *) | RT_OTHER fun rthm_type thm = case Thm.concl_of thm |> HOLogic.dest_Trueprop of @{mpat "(_,_) \ fref _ _ _"} => RT_FREF | @{mpat "(_,_) \ hfref _ _ _"} => RT_HFREF | @{mpat "hn_refine _ _ _ _ _"} => RT_HNR | @{mpat "(_,_) \ _"} => RT_HOPARAM (* TODO: Distinction between ho-param and fo-param *) | _ => RT_OTHER fun to_fref ctxt thm = let open Conv in case Thm.concl_of thm |> HOLogic.dest_Trueprop of @{mpat "(_,_)\_\_"} => Local_Defs.unfold0 ctxt @{thms fref_param1} thm |> fconv_rule (repeat_conv (Refine_Util.ftop_conv (K (rewr_conv @{thm fref_nest})) ctxt)) |> Local_Defs.unfold0 ctxt @{thms in_CURRY_conv} | @{mpat "(_,_)\_"} => thm RS @{thm fref_param0I} | _ => raise THM ("to_fref: Expected theorem of form (_,_)\_",~1,[thm]) end fun to_foparam ctxt thm = let val unf_thms = @{thms split_tupled_all prod_rel_simp uncurry_apply cnv_conj_to_meta Product_Type.split} in case Thm.concl_of thm of @{mpat "Trueprop ((_,_) \ fref _ _ _)"} => (@{thm frefD} OF [thm]) |> Thm.forall_intr_vars |> Local_Defs.unfold0 ctxt unf_thms |> Variable.gen_all ctxt | @{mpat "Trueprop ((_,_) \ _)"} => Parametricity.fo_rule thm | _ => raise THM("Expected parametricity or fref theorem",~1,[thm]) end fun to_hnr ctxt thm = (thm RS @{thm hf2hnr}) |> Local_Defs.unfold0 ctxt @{thms to_hnr_prod_fst_snd keep_drop_sels} (* Resolve fst and snd over *\<^sub>a and R\<^sup>k, R\<^sup>d *) |> Local_Defs.unfold0 ctxt @{thms hnr_uncurry_unfold} (* Resolve products for uncurried parameters *) |> Local_Defs.unfold0 ctxt @{thms uncurry_apply uncurry_APP assn_one_left split} (* Remove the uncurry modifiers, the emp-dummy, and unfold product cases *) |> Local_Defs.unfold0 ctxt @{thms hn_ctxt_ctxt_fix_conv} (* Remove duplicate hn_ctxt tagging *) |> Local_Defs.unfold0 ctxt @{thms all_to_meta imp_to_meta HOL.True_implies_equals HOL.implies_True_equals Pure.triv_forall_equality cnv_conj_to_meta} (* Convert to meta-level, remove vacuous condition *) |> Local_Defs.unfold0 ctxt (Named_Theorems.get ctxt @{named_theorems to_hnr_post}) (* Post-Processing *) |> Goal.norm_result ctxt |> Conv.fconv_rule Thm.eta_conversion (* Convert schematic hfref-goal to hn_refine goal *) fun prepare_hfref_synth_tac ctxt = let val i_of_assn_rls = Named_Theorems_Rev.get ctxt @{named_theorems_rev intf_of_assn} @ @{thms intf_of_assn_fallback} val to_hnr_post_rls = Named_Theorems.get ctxt @{named_theorems to_hnr_post} val i_of_assn_tac = ( REPEAT' ( DETERM o dresolve_tac ctxt @{thms hfsynth_ID_R_D} THEN' DETERM o SOLVED' (REPEAT_ALL_NEW (resolve_tac ctxt i_of_assn_rls)) ) ) in (* Note: To re-use the to_hnr infrastructure, we first work with $-tags on the abstract function, which are finally removed. *) resolve_tac ctxt @{thms hfsynth_hnr_from_hfI} THEN_ELSE' ( SELECT_GOAL ( unfold_tac ctxt @{thms to_hnr_prod_fst_snd keep_drop_sels hf_pres_fst} (* Distribute fst,snd over product and hf_pres *) THEN unfold_tac ctxt @{thms hnr_uncurry_unfold hfsynth_ID_R_uncurry_unfold} (* Curry parameters *) THEN unfold_tac ctxt @{thms uncurry_apply uncurry_APP assn_one_left split} (* Curry parameters (II) and remove emp assertion *) (*THEN unfold_tac ctxt @{thms hn_ctxt_ctxt_fix_conv} (* Remove duplicate hn_ctxt (Should not be necessary) *)*) THEN unfold_tac ctxt @{thms all_to_meta imp_to_meta HOL.True_implies_equals HOL.implies_True_equals Pure.triv_forall_equality cnv_conj_to_meta} (* Convert precondition to meta-level *) THEN ALLGOALS i_of_assn_tac (* Generate _::\<^sub>i_ premises*) THEN unfold_tac ctxt to_hnr_post_rls (* Postprocessing *) THEN unfold_tac ctxt @{thms APP_def} (* Get rid of $ - tags *) ) , K all_tac ) end (************************************) (* Analyze hnr *) structure Termtab2 = Table( type key = term * term val ord = prod_ord Term_Ord.fast_term_ord Term_Ord.fast_term_ord); type hnr_analysis = { thm: thm, precond: term, prems : term list, ahead: term * bool, chead: term * bool, argrels: (term * bool) list, result_rel: term } fun analyze_hnr (ctxt:Proof.context) thm = let (* Debug information: Stores string*term pairs, which are pretty-printed on error *) val dbg = Unsynchronized.ref [] fun add_dbg msg ts = ( dbg := (msg,ts) :: !dbg; () ) fun pretty_dbg (msg,ts) = Pretty.block [ Pretty.str msg, Pretty.str ":", Pretty.brk 1, Pretty.list "[" "]" (map (Syntax.pretty_term ctxt) ts) ] fun pretty_dbgs l = map pretty_dbg l |> Pretty.fbreaks |> Pretty.block fun trace_dbg msg = Pretty.block [Pretty.str msg, Pretty.fbrk, pretty_dbgs (rev (!dbg))] |> Pretty.string_of |> tracing fun fail msg = (trace_dbg msg; raise THM(msg,~1,[thm])) fun assert cond msg = cond orelse fail msg; (* Heads may have a leading return/RETURN. The following code strips off the leading return, unless it has the form "return x" for an argument x *) fun check_strip_leading args t f = (* Handle the case RETURN x, where x is an argument *) if Termtab.defined args f then (t,false) else (f,true) fun strip_leading_RETURN args (t as @{mpat "RETURN$(?f)"}) = check_strip_leading args t f | strip_leading_RETURN args (t as @{mpat "RETURN ?f"}) = check_strip_leading args t f | strip_leading_RETURN _ t = (t,false) fun strip_leading_return args (t as @{mpat "return$(?f)"}) = check_strip_leading args t f | strip_leading_return args (t as @{mpat "return ?f"}) = check_strip_leading args t f | strip_leading_return _ t = (t,false) (* The following code strips the arguments of the concrete or abstract function. It knows how to handle APP-tags ($), and stops at PR_CONST-tags. Moreover, it only strips actual arguments that occur in the precondition-section of the hn_refine-statement. This ensures that non-arguments, like maxsize, are treated correctly. *) fun strip_fun _ (t as @{mpat "PR_CONST _"}) = (t,[]) | strip_fun s (t as @{mpat "?f$?x"}) = check_arg s t f x | strip_fun s (t as @{mpat "?f ?x"}) = check_arg s t f x | strip_fun _ f = (f,[]) and check_arg s t f x = if Termtab.defined s x then strip_fun s f |> apsnd (curry op :: x) else (t,[]) (* Arguments in the pre/postcondition are wrapped into hn_ctxt tags. This function strips them off. *) fun dest_hn_ctxt @{mpat "hn_ctxt ?R ?a ?c"} = ((a,c),R) | dest_hn_ctxt _ = fail "Invalid hn_ctxt parameter in pre or postcondition" fun dest_hn_refine @{mpat "(hn_refine ?G ?c ?G' ?R ?a)"} = (G,c,G',R,a) | dest_hn_refine _ = fail "Conclusion is not a hn_refine statement" (* Strip separation conjunctions. Special case for "emp", which is ignored. *) fun is_emp @{mpat emp} = true | is_emp _ = false val strip_star' = Sepref_Basic.strip_star #> filter (not o is_emp) (* Compare Termtab2s for equality of keys *) fun pairs_eq pairs1 pairs2 = Termtab2.forall (Termtab2.defined pairs1 o fst) pairs2 andalso Termtab2.forall (Termtab2.defined pairs2 o fst) pairs1 fun atomize_prem @{mpat "Trueprop ?p"} = p | atomize_prem _ = fail "Non-atomic premises" (* Make HOL conjunction list *) fun mk_conjs [] = @{const True} | mk_conjs [p] = p | mk_conjs (p::ps) = HOLogic.mk_binop @{const_name "HOL.conj"} (p,mk_conjs ps) (***********************) (* Start actual analysis *) val _ = add_dbg "thm" [Thm.prop_of thm] val prems = Thm.prems_of thm val concl = Thm.concl_of thm |> HOLogic.dest_Trueprop val (G,c,G',R,a) = dest_hn_refine concl val pre_pairs = G |> strip_star' |> tap (add_dbg "precondition") |> map dest_hn_ctxt |> Termtab2.make val post_pairs = G' |> strip_star' |> tap (add_dbg "postcondition") |> map dest_hn_ctxt |> Termtab2.make val _ = assert (pairs_eq pre_pairs post_pairs) "Parameters in precondition do not match postcondition" val aa_set = pre_pairs |> Termtab2.keys |> map fst |> Termtab.make_set val ca_set = pre_pairs |> Termtab2.keys |> map snd |> Termtab.make_set val (a,leading_RETURN) = strip_leading_RETURN aa_set a val (c,leading_return) = strip_leading_return ca_set c val _ = add_dbg "stripped abstract term" [a] val _ = add_dbg "stripped concrete term" [c] val (ahead,aargs) = strip_fun aa_set a; val (chead,cargs) = strip_fun ca_set c; val _ = add_dbg "abstract head" [ahead] val _ = add_dbg "abstract args" aargs val _ = add_dbg "concrete head" [chead] val _ = add_dbg "concrete args" cargs val _ = assert (length cargs = length aargs) "Different number of abstract and concrete arguments"; val _ = assert (not (has_duplicates op aconv aargs)) "Duplicate abstract arguments" val _ = assert (not (has_duplicates op aconv cargs)) "Duplicate concrete arguments" val argpairs = aargs ~~ cargs val ap_set = Termtab2.make_set argpairs val _ = assert (pairs_eq pre_pairs ap_set) "Arguments from pre/postcondition do not match operation's arguments" val pre_rels = map (the o (Termtab2.lookup pre_pairs)) argpairs val post_rels = map (the o (Termtab2.lookup post_pairs)) argpairs val _ = add_dbg "pre-rels" pre_rels val _ = add_dbg "post-rels" post_rels fun adjust_hf_pres @{mpat "snd (?R\<^sup>k)"} = R | adjust_hf_pres t = t val post_rels = map adjust_hf_pres post_rels fun is_invalid R @{mpat "invalid_assn ?R'"} = R aconv R' | is_invalid _ @{mpat "snd (_\<^sup>d)"} = true | is_invalid _ _ = false fun is_keep (R,R') = if R aconv R' then true else if is_invalid R R' then false else fail "Mismatch between pre and post relation for argument" val keep = map is_keep (pre_rels ~~ post_rels) val argrels = pre_rels ~~ keep val aa_set = Termtab.make_set aargs val ca_set = Termtab.make_set cargs fun is_precond t = (exists_subterm (Termtab.defined ca_set) t andalso fail "Premise contains concrete argument") orelse exists_subterm (Termtab.defined aa_set) t val (preconds, prems) = split is_precond prems val precond = map atomize_prem preconds |> mk_conjs |> fold lambda aargs val _ = add_dbg "precond" [precond] val _ = add_dbg "prems" prems in { thm = thm, precond = precond, prems = prems, ahead = (ahead,leading_RETURN), chead = (chead,leading_return), argrels = argrels, result_rel = R } end fun pretty_hnr_analysis ctxt ({thm,precond,ahead,chead,argrels,result_rel,...}) : Pretty.T = let val _ = thm (* Suppress unused warning for thm *) fun pretty_argrel (R,k) = Pretty.block [ Syntax.pretty_term ctxt R, if k then Pretty.str "\<^sup>k" else Pretty.str "\<^sup>d" ] val pretty_chead = case chead of (t,false) => Syntax.pretty_term ctxt t | (t,true) => Pretty.block [Pretty.str "return ", Syntax.pretty_term ctxt t] val pretty_ahead = case ahead of (t,false) => Syntax.pretty_term ctxt t | (t,true) => Pretty.block [Pretty.str "RETURN ", Syntax.pretty_term ctxt t] in Pretty.fbreaks [ (*Display.pretty_thm ctxt thm,*) Pretty.block [ Pretty.enclose "[" "]" [pretty_chead, pretty_ahead], Pretty.enclose "[" "]" [Syntax.pretty_term ctxt precond], Pretty.brk 1, Pretty.block (Pretty.separate " \" (map pretty_argrel argrels @ [Syntax.pretty_term ctxt result_rel])) ] ] |> Pretty.block end fun mk_hfref_thm ctxt ({thm,precond,prems,ahead,chead,argrels,result_rel}) = let fun mk_keep (R,true) = @{mk_term "?R\<^sup>k"} | mk_keep (R,false) = @{mk_term "?R\<^sup>d"} (* TODO: Move, this is of general use! *) fun mk_uncurry f = @{mk_term "uncurry ?f"} (* Uncurry function for the given number of arguments. For zero arguments, add a unit-parameter. *) fun rpt_uncurry n t = if n=0 then @{mk_term "uncurry0 ?t"} else if n=1 then t else funpow (n-1) mk_uncurry t (* Rewrite uncurried lambda's to \(_,_). _ form. Use top-down rewriting to correctly handle nesting to the left. TODO: Combine with abstraction and uncurry-procedure, and mark the deviation about uncurry as redundant intermediate step to be eliminated. *) fun rew_uncurry_lambda t = let val rr = map (Logic.dest_equals o Thm.prop_of) @{thms uncurry_def uncurry0_def} val thy = Proof_Context.theory_of ctxt in Pattern.rewrite_term_top thy rr [] t end (* Shortcuts for simplification tactics *) fun gsimp_only ctxt sec = let val ss = put_simpset HOL_basic_ss ctxt |> sec in asm_full_simp_tac ss end fun simp_only ctxt thms = gsimp_only ctxt (fn ctxt => ctxt addsimps thms) (********************************) (* Build theorem statement *) (* \prems\ \ (chead,ahead) \ [precond] rels \ R *) (* Uncurry precondition *) val num_args = length argrels val precond = precond |> rpt_uncurry num_args |> rew_uncurry_lambda (* Convert to nicer \((...,_),_) - form*) (* Re-attach leading RETURN/return *) fun mk_RETURN (t,r) = if r then let val T = funpow num_args range_type (fastype_of (fst ahead)) val tRETURN = Const (@{const_name RETURN}, T --> Type(@{type_name nres},[T])) in Refine_Util.mk_compN num_args tRETURN t end else t fun mk_return (t,r) = if r then let val T = funpow num_args range_type (fastype_of (fst chead)) val tRETURN = Const (@{const_name return}, T --> Type(@{type_name Heap},[T])) in Refine_Util.mk_compN num_args tRETURN t end else t (* Hrmpf!: Gone for good from 2015\2016. Inserting ctxt-based substitute here. *) fun certify_inst ctxt (instT, inst) = - (map (apsnd (Thm.ctyp_of ctxt)) (TVars.dest instT), - map (apsnd (Thm.cterm_of ctxt)) (Vars.dest inst)); + (TVars.map (K (Thm.ctyp_of ctxt)) instT, + Vars.map (K (Thm.cterm_of ctxt)) inst); (* fun mk_RETURN (t,r) = if r then @{mk_term "RETURN o ?t"} else t fun mk_return (t,r) = if r then @{mk_term "return o ?t"} else t *) (* Uncurry abstract and concrete function, append leading return *) val ahead = ahead |> mk_RETURN |> rpt_uncurry num_args val chead = chead |> mk_return |> rpt_uncurry num_args (* Add keep-flags and summarize argument relations to product *) val argrel = map mk_keep argrels |> rev (* TODO: Why this rev? *) |> mk_hfprods (* Produce final result statement *) val result = @{mk_term "Trueprop ((?chead,?ahead) \ [?precond]\<^sub>a ?argrel \ ?result_rel)"} val result = Logic.list_implies (prems,result) (********************************) (* Prove theorem *) (* Create context and import result statement and original theorem *) val orig_ctxt = ctxt (*val thy = Proof_Context.theory_of ctxt*) val (insts, ctxt) = Variable.import_inst true [result] ctxt val insts' = certify_inst ctxt insts val result = Term_Subst.instantiate insts result val thm = Thm.instantiate insts' thm (* Unfold APP tags. This is required as some APP-tags have also been unfolded by analysis *) val thm = Local_Defs.unfold0 ctxt @{thms APP_def} thm (* Tactic to prove the theorem. A first step uses hfrefI to get a hnr-goal. This is then normalized in several consecutive steps, which get rid of uncurrying. Finally, the original theorem is used for resolution, where the pre- and postcondition, and result relation are connected with a consequence rule, to handle unfolded hn_ctxt-tags, re-ordered relations, and introduced unit-parameters (TODO: Mark artificially introduced unit-parameter specially, it may get confused with intentional unit-parameter, e.g., functional empty_set ()!) *) fun tac ctxt = resolve_tac ctxt @{thms hfrefI} THEN' gsimp_only ctxt (fn c => c addsimps @{thms uncurry_def hn_ctxt_def uncurry0_def keep_drop_sels uc_hfprod_sel o_apply APP_def} |> Splitter.add_split @{thm prod.split} ) THEN' TRY o ( REPEAT_ALL_NEW (match_tac ctxt @{thms allI impI}) THEN' simp_only ctxt @{thms Product_Type.split prod.inject}) THEN' TRY o REPEAT_ALL_NEW (ematch_tac ctxt @{thms conjE}) THEN' TRY o hyp_subst_tac ctxt THEN' simp_only ctxt @{thms triv_forall_equality} THEN' ( resolve_tac ctxt @{thms hn_refine_cons[rotated]} THEN' (resolve_tac ctxt [thm] THEN_ALL_NEW assume_tac ctxt)) THEN_ALL_NEW simp_only ctxt @{thms hn_ctxt_def entt_refl pure_unit_rel_eq_empty mult_ac mult_1 mult_1_right keep_drop_sels} (* Prove theorem *) val result = Thm.cterm_of ctxt result val rthm = Goal.prove_internal ctxt [] result (fn _ => ALLGOALS (tac ctxt)) (* Export statement to original context *) val rthm = singleton (Variable.export ctxt orig_ctxt) rthm (* Post-processing *) val rthm = Local_Defs.unfold0 ctxt (Named_Theorems.get ctxt @{named_theorems to_hfref_post}) rthm in rthm end fun to_hfref ctxt = analyze_hnr ctxt #> mk_hfref_thm ctxt (***********************************) (* Composition *) local fun norm_set_of ctxt = { trans_rules = Named_Theorems.get ctxt @{named_theorems fcomp_norm_trans}, cong_rules = Named_Theorems.get ctxt @{named_theorems fcomp_norm_cong}, norm_rules = Named_Theorems.get ctxt @{named_theorems fcomp_norm_norm}, refl_rules = Named_Theorems.get ctxt @{named_theorems fcomp_norm_refl} } fun init_rules_of ctxt = Named_Theorems.get ctxt @{named_theorems fcomp_norm_init} fun unfold_rules_of ctxt = Named_Theorems.get ctxt @{named_theorems fcomp_norm_unfold} fun simp_rules_of ctxt = Named_Theorems.get ctxt @{named_theorems fcomp_norm_simps} in fun norm_fcomp_rule ctxt = let open PO_Normalizer Refine_Util val norm1 = gen_norm_rule (init_rules_of ctxt) (norm_set_of ctxt) ctxt val norm2 = Local_Defs.unfold0 ctxt (unfold_rules_of ctxt) val norm3 = Conv.fconv_rule ( Simplifier.asm_full_rewrite (put_simpset HOL_basic_ss ctxt addsimps simp_rules_of ctxt)) val norm = changed_rule (try_rule norm1 o try_rule norm2 o try_rule norm3) in repeat_rule norm end end fun add_pure_constraints_rule ctxt thm = let val orig_ctxt = ctxt val t = Thm.prop_of thm fun cnv (@{mpat (typs) "pure (mpaq_STRUCT (mpaq_Var ?x _) :: (?'v_c\?'v_a) set)"}) = let val T = a --> c --> @{typ assn} val t = Var (x,T) val t = @{mk_term "(the_pure ?t)"} in [(x,T,t)] end | cnv (t$u) = union op= (cnv t) (cnv u) | cnv (Abs (_,_,t)) = cnv t | cnv _ = [] val pvars = cnv t val _ = (pvars |> map #1 |> has_duplicates op=) andalso raise TERM ("Duplicate indexname with different type",[t]) (* This should not happen *) val substs = map (fn (x,_,t) => (x,t)) pvars val t' = subst_Vars substs t fun mk_asm (x,T,_) = let val t = Var (x,T) val t = @{mk_term "Trueprop (CONSTRAINT is_pure ?t)"} in t end val assms = map mk_asm pvars fun add_prems prems t = let val prems' = Logic.strip_imp_prems t val concl = Logic.strip_imp_concl t in Logic.list_implies (prems@prems', concl) end val t' = add_prems assms t' val (t',ctxt) = yield_singleton (Variable.import_terms true) t' ctxt val thm' = Goal.prove_internal ctxt [] (Thm.cterm_of ctxt t') (fn _ => ALLGOALS (resolve_tac ctxt [thm] THEN_ALL_NEW assume_tac ctxt)) val thm' = norm_fcomp_rule ctxt thm' val thm' = singleton (Variable.export ctxt orig_ctxt) thm' in thm' end val cfg_simp_precond = Attrib.setup_config_bool @{binding fcomp_simp_precond} (K true) local fun mk_simp_thm ctxt t = let val st = t |> HOLogic.mk_Trueprop |> Thm.cterm_of ctxt |> Goal.init val ctxt = Context_Position.set_visible false ctxt val ctxt = ctxt addsimps ( refine_pw_simps.get ctxt @ Named_Theorems.get ctxt @{named_theorems fcomp_prenorm_simps} @ @{thms split_tupled_all cnv_conj_to_meta} ) val trace_incomplete_transfer_tac = COND (Thm.prems_of #> exists (strip_all_body #> Logic.strip_imp_concl #> Term.is_open)) (print_tac ctxt "Failed transfer from intermediate level:") all_tac val tac = ALLGOALS (resolve_tac ctxt @{thms auto_weaken_pre_comp_PRE_I} ) THEN ALLGOALS (Simplifier.asm_full_simp_tac ctxt) THEN trace_incomplete_transfer_tac THEN ALLGOALS (TRY o filter_prems_tac ctxt (K false)) THEN Local_Defs.unfold0_tac ctxt [Drule.triv_forall_equality] val st' = tac st |> Seq.take 1 |> Seq.list_of val thm = case st' of [st'] => Goal.conclude st' | _ => raise THM("Simp_Precond: Simp-Tactic failed",~1,[st]) (* Check generated premises for leftover intermediate stuff *) val _ = exists (Logic.is_all) (Thm.prems_of thm) andalso raise THM("Simp_Precond: Transfer from intermediate level failed",~1,[thm]) val thm = thm (*|> map (Simplifier.asm_full_simplify ctxt)*) |> Conv.fconv_rule (Object_Logic.atomize ctxt) |> Local_Defs.unfold0 ctxt @{thms auto_weaken_pre_to_imp_nf} val thm = case Thm.concl_of thm of @{mpat "Trueprop (_ \ _)"} => thm | @{mpat "Trueprop _"} => thm RS @{thm auto_weaken_pre_add_dummy_imp} | _ => raise THM("Simp_Precond: Generated odd theorem, expected form 'P\Q'",~1,[thm]) in thm end in fun simplify_precond ctxt thm = let val orig_ctxt = ctxt val thm = Refine_Util.OF_fst @{thms auto_weaken_pre_init} [asm_rl,thm] val thm = Local_Defs.unfold0 ctxt @{thms split_tupled_all} thm OF @{thms auto_weaken_pre_uncurry_start} fun rec_uncurry thm = case try (fn () => thm OF @{thms auto_weaken_pre_uncurry_step}) () of NONE => thm OF @{thms auto_weaken_pre_uncurry_finish} | SOME thm => rec_uncurry thm val thm = rec_uncurry thm |> Conv.fconv_rule Thm.eta_conversion val t = case Thm.prems_of thm of t::_ => t | _ => raise THM("Simp-Precond: Expected at least one premise",~1,[thm]) val (t,ctxt) = yield_singleton (Variable.import_terms false) t ctxt val ((_,t),ctxt) = Variable.focus NONE t ctxt val t = case t of @{mpat "Trueprop (_ \ ?t)"} => t | _ => raise TERM("Simp_Precond: Expected implication",[t]) val simpthm = mk_simp_thm ctxt t |> singleton (Variable.export ctxt orig_ctxt) val thm = thm OF [simpthm] val thm = Local_Defs.unfold0 ctxt @{thms prod_casesK} thm in thm end fun simplify_precond_if_cfg ctxt = if Config.get ctxt cfg_simp_precond then simplify_precond ctxt else I end (* fref O fref *) fun compose_ff ctxt A B = (@{thm fref_compI_PRE} OF [A,B]) |> norm_fcomp_rule ctxt |> simplify_precond_if_cfg ctxt |> Conv.fconv_rule Thm.eta_conversion (* hfref O fref *) fun compose_hf ctxt A B = (@{thm hfref_compI_PRE} OF [A,B]) |> norm_fcomp_rule ctxt |> simplify_precond_if_cfg ctxt |> Conv.fconv_rule Thm.eta_conversion |> add_pure_constraints_rule ctxt |> Conv.fconv_rule Thm.eta_conversion fun ensure_fref ctxt thm = case rthm_type thm of RT_HOPARAM => to_fref ctxt thm | RT_FREF => thm | _ => raise THM("Expected parametricity or fref theorem",~1,[thm]) fun ensure_fref_nres ctxt thm = let val thm = ensure_fref ctxt thm in case Thm.concl_of thm of @{mpat (typs) "Trueprop (_\fref _ _ (_::(_ nres\_)set))"} => thm | @{mpat "Trueprop ((_,_)\fref _ _ _)"} => (thm RS @{thm ensure_fref_nresI}) |> Local_Defs.unfold0 ctxt @{thms ensure_fref_nres_unfold} | _ => raise THM("Expected fref-theorem",~1,[thm]) end fun ensure_hfref ctxt thm = case rthm_type thm of RT_HNR => to_hfref ctxt thm | RT_HFREF => thm | _ => raise THM("Expected hnr or hfref theorem",~1,[thm]) fun ensure_hnr ctxt thm = case rthm_type thm of RT_HNR => thm | RT_HFREF => to_hnr ctxt thm | _ => raise THM("Expected hnr or hfref theorem",~1,[thm]) fun gen_compose ctxt A B = let val rtA = rthm_type A in if rtA = RT_HOPARAM orelse rtA = RT_FREF then compose_ff ctxt (ensure_fref ctxt A) (ensure_fref ctxt B) else compose_hf ctxt (ensure_hfref ctxt A) ((ensure_fref_nres ctxt B)) end val parse_fcomp_flags = Refine_Util.parse_paren_lists (Refine_Util.parse_bool_config "prenorm" cfg_simp_precond) val fcomp_attrib = parse_fcomp_flags |-- Attrib.thm >> (fn B => Thm.rule_attribute [] (fn context => fn A => let val ctxt = Context.proof_of context in gen_compose ctxt A B end)) end \ attribute_setup to_fref = \ Scan.succeed (Thm.rule_attribute [] (Sepref_Rules.to_fref o Context.proof_of)) \ "Convert parametricity theorem to uncurried fref-form" attribute_setup to_foparam = \ Scan.succeed (Thm.rule_attribute [] (Sepref_Rules.to_foparam o Context.proof_of)) \ \Convert param or fref rule to first order rule\ (* Overloading existing param_fo - attribute from Parametricity.thy *) attribute_setup param_fo = \ Scan.succeed (Thm.rule_attribute [] (Sepref_Rules.to_foparam o Context.proof_of)) \ \Convert param or fref rule to first order rule\ attribute_setup to_hnr = \ Scan.succeed (Thm.rule_attribute [] (Sepref_Rules.to_hnr o Context.proof_of)) \ "Convert hfref-rule to hnr-rule" attribute_setup to_hfref = \Scan.succeed ( Thm.rule_attribute [] (Context.proof_of #> Sepref_Rules.to_hfref) )\ \Convert hnr to hfref theorem\ attribute_setup ensure_fref_nres = \Scan.succeed ( Thm.rule_attribute [] (Context.proof_of #> Sepref_Rules.ensure_fref_nres) )\ attribute_setup sepref_dbg_norm_fcomp_rule = \Scan.succeed ( Thm.rule_attribute [] (Context.proof_of #> Sepref_Rules.norm_fcomp_rule) )\ attribute_setup sepref_simplify_precond = \Scan.succeed ( Thm.rule_attribute [] (Context.proof_of #> Sepref_Rules.simplify_precond) )\ \Simplify precondition of fref/hfref-theorem\ attribute_setup FCOMP = Sepref_Rules.fcomp_attrib "Composition of refinement rules" end diff --git a/thys/Simpl/generalise_state.ML b/thys/Simpl/generalise_state.ML --- a/thys/Simpl/generalise_state.ML +++ b/thys/Simpl/generalise_state.ML @@ -1,301 +1,301 @@ (* Title: generalise_state.ML Author: Norbert Schirmer, TU Muenchen Copyright (C) 2005-2007 Norbert Schirmer This library is free software; you can redistribute it and/or modify it under the terms of the GNU Lesser General Public License as published by the Free Software Foundation; either version 2.1 of the License, or (at your option) any later version. This library is distributed in the hope that it will be useful, but WITHOUT ANY WARRANTY; without even the implied warranty of MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU Lesser General Public License for more details. You should have received a copy of the GNU Lesser General Public License along with this library; if not, write to the Free Software Foundation, Inc., 59 Temple Place, Suite 330, Boston, MA 02111-1307 USA *) signature SPLIT_STATE = sig val isState: term -> bool val abs_state: term -> term option val abs_var: Proof.context -> term -> (string * typ) val split_state: Proof.context -> string -> typ -> term -> (term * term list) val ex_tac: Proof.context -> term list -> tactic (* the term-list is the list of selectors as returned by split_state. They may be used to construct the instatiation of the existentially quantified state. *) end; functor GeneraliseFun (structure SplitState: SPLIT_STATE) = struct val genConj = @{thm generaliseConj}; val genImp = @{thm generaliseImp}; val genImpl = @{thm generaliseImpl}; val genAll = @{thm generaliseAll}; val gen_all = @{thm generalise_all}; val genEx = @{thm generaliseEx}; val genRefl = @{thm generaliseRefl}; val genRefl' = @{thm generaliseRefl'}; val genTrans = @{thm generaliseTrans}; val genAllShift = @{thm generaliseAllShift}; val gen_allShift = @{thm generalise_allShift}; val meta_spec = @{thm meta_spec}; val protectRefl = @{thm protectRefl}; val protectImp = @{thm protectImp}; fun gen_thm decomp (t,ct) = let val (ts,cts,recomb) = decomp (t,ct) in recomb (map (gen_thm decomp) (ts~~cts)) end; fun dest_prop (Const (@{const_name Pure.prop}, _) $ P) = P | dest_prop t = raise TERM ("dest_prop", [t]); fun prem_of thm = #1 (Logic.dest_implies (dest_prop (Thm.prop_of thm))); fun conc_of thm = #2 (Logic.dest_implies (dest_prop (Thm.prop_of thm))); fun dest_All (Const (@{const_name "All"},_)$t) = t | dest_All t = raise TERM ("dest_All",[t]); fun SIMPLE_OF ctxt rule prems = let val mx = fold (fn thm => fn i => Int.max (Thm.maxidx_of thm,i)) prems 0; in DistinctTreeProver.discharge ctxt prems (Thm.incr_indexes (mx + 1) rule) end; infix 0 OF_RAW fun tha OF_RAW thb = thb COMP (Drule.incr_indexes thb tha); fun SIMPLE_OF_RAW ctxt tha thb = SIMPLE_OF ctxt tha [thb]; datatype qantifier = Meta_all | Hol_all | Hol_ex fun list_exists (vs, x) = fold_rev (fn (x, T) => fn P => HOLogic.exists_const T $ Abs (x, T, P)) vs x; fun spec' cv thm = let (* thm = Pure.prop ((all x. P x) ==> Q), where "all" is meta or HOL *) val (ct1,ct2) = thm |> Thm.cprop_of |> Thm.dest_comb |> #2 |> Thm.dest_comb |> #2 |> Thm.dest_comb; in (case Thm.term_of ct1 of Const (@{const_name "Trueprop"},_) => let val (Var (sP,_)$Var (sV,sVT)) = HOLogic.dest_Trueprop (Thm.concl_of spec); val cvT = Thm.ctyp_of_cterm cv; val vT = Thm.typ_of cvT; in Thm.instantiate - ([(dest_TVar sVT, cvT)], - [((sP, vT --> HOLogic.boolT), #2 (Thm.dest_comb ct2)), + (TVars.make [(dest_TVar sVT, cvT)], + Vars.make [((sP, vT --> HOLogic.boolT), #2 (Thm.dest_comb ct2)), ((sV, vT), cv)]) spec end | Const (@{const_name Pure.all},_) => let val (Var (sP,_)$Var (sV,sVT)) = Thm.concl_of meta_spec; val cvT = Thm.ctyp_of_cterm cv; val vT = Thm.typ_of cvT; in Thm.instantiate - ([(dest_TVar sVT, cvT)], - [((sP, vT --> propT), ct2), + (TVars.make [(dest_TVar sVT, cvT)], + Vars.make [((sP, vT --> propT), ct2), ((sV, vT),cv)]) meta_spec end | _ => raise THM ("spec'",0,[thm])) end; fun split_thm qnt ctxt s T t = let val (t',vars) = SplitState.split_state ctxt s T t; val vs = map (SplitState.abs_var ctxt) vars; val prop = (case qnt of Meta_all => Logic.list_all (vs,t') | Hol_all => HOLogic.mk_Trueprop (HOLogic.list_all (vs, t')) | Hol_ex => Logic.mk_implies (HOLogic.mk_Trueprop (list_exists (vs, t')), HOLogic.mk_Trueprop (HOLogic.mk_exists (s,T,t)))) in (case qnt of Hol_ex => Goal.prove ctxt [] [] prop (fn _ => SplitState.ex_tac ctxt vars) | _ => let val rP = conc_of genRefl'; - val thm0 = Thm.instantiate ([], [(dest_Var rP, Thm.cterm_of ctxt prop)]) genRefl'; + val thm0 = Thm.instantiate (TVars.empty, Vars.make [(dest_Var rP, Thm.cterm_of ctxt prop)]) genRefl'; fun elim_all v thm = let val cv = Thm.cterm_of ctxt v; val spc = Goal.protect 0 (spec' cv thm); in SIMPLE_OF ctxt genTrans [thm,spc] end; val thm = fold elim_all vars thm0; in thm end) end; fun eta_expand ctxt ct = let val mi = Thm.maxidx_of_cterm ct; val T = domain_type (Thm.typ_of_cterm ct); val x = Thm.cterm_of ctxt (Var (("s",mi+1),T)); in Thm.lambda x (Thm.apply ct x) end; fun split_abs ctxt ct = (case Thm.term_of ct of Abs x => (x, Thm.dest_abs NONE ct) | _ => split_abs ctxt (eta_expand ctxt ct)) fun decomp ctxt (Const (@{const_name HOL.conj}, _) $ t $ t', ct) = ([t,t'],snd (Drule.strip_comb ct), fn [thm,thm'] => SIMPLE_OF ctxt genConj [thm,thm']) | decomp ctxt ((allc as Const (@{const_name "All"},aT)) $ f, ct) = let val cf = snd (Thm.dest_comb ct); val (abst as (x,T,_),(cx',cb)) = split_abs ctxt cf; val Free (x',_) = Thm.term_of cx'; val (Const (@{const_name Pure.all},_)$Abs (s,_,_)) = genAll |> Thm.prems_of |> hd |> dest_prop; val genAll' = Drule.rename_bvars [(s,x)] genAll; val (Const (@{const_name Pure.all},_)$Abs (s',_,_)) = genAllShift |> Thm.prems_of |> hd |> dest_prop; val genAllShift' = Drule.rename_bvars [(s',x)] genAllShift; in if SplitState.isState (allc$Abs abst) then ([Thm.term_of cb],[cb], fn [thm] => let val P = HOLogic.dest_Trueprop (dest_prop (prem_of thm)); val thm' = split_thm Hol_all ctxt x' T P; val thm1 = genAllShift' OF_RAW Goal.protect 0 (Thm.forall_intr cx' (Goal.conclude thm')); val thm2 = genAll' OF_RAW Goal.protect 0 (Thm.forall_intr cx' (Goal.conclude thm)); in SIMPLE_OF ctxt genTrans [thm1,thm2] end) else ([Thm.term_of cb],[cb], fn [thm] => genAll' OF_RAW Goal.protect 0 (Thm.forall_intr cx' (Goal.conclude thm))) end | decomp ctxt ((exc as Const (@{const_name "Ex"},_)) $ f, ct) = let val cf = snd (Thm.dest_comb ct); val (abst as (x,T,_),(cx',cb)) = split_abs ctxt cf; val Free (x',_) = Thm.term_of cx'; val (Const (@{const_name Pure.all},_)$Abs (s,_,_)) = genEx |> Thm.prems_of |> hd |> dest_prop; val genEx' = Drule.rename_bvars [(s,x)] genEx; in if SplitState.isState (exc$Abs abst) then ([Thm.term_of cb],[cb], fn [thm] => let val P = HOLogic.dest_Trueprop (dest_prop (prem_of thm)); val thm' = split_thm Hol_ex ctxt x' T P; in SIMPLE_OF_RAW ctxt protectImp (Goal.protect 0 thm') end ) else ([Thm.term_of cb],[cb], fn [thm] => genEx' OF_RAW Goal.protect 0 (Thm.forall_intr cx' (Goal.conclude thm))) end | decomp ctxt (Const (@{const_name HOL.implies},_)$P$Q, ct) = let val [cP,cQ] = (snd (Drule.strip_comb ct)); in ([Q],[cQ],fn [thm] => let val X = genImp |> Thm.concl_of |> dest_prop |> Logic.dest_implies |> #1 |> dest_prop |> HOLogic.dest_Trueprop |> HOLogic.dest_imp |> #1 |> dest_Var; - val genImp' = Thm.instantiate ([],[(X,cP)]) genImp; + val genImp' = Thm.instantiate (TVars.empty, Vars.make [(X,cP)]) genImp; in SIMPLE_OF ctxt genImp' [thm] end) end | decomp ctxt (Const (@{const_name Pure.imp},_)$P$Q, ct) = let val [cP,cQ] = (snd (Drule.strip_comb ct)); in ([Q],[cQ],fn [thm] => let val X = genImpl |> Thm.concl_of |> dest_prop |> Logic.dest_implies |> #1 |> dest_prop |> Logic.dest_implies |> #1 |> dest_Var; - val genImpl' = Thm.instantiate ([],[(X,cP)]) genImpl; + val genImpl' = Thm.instantiate (TVars.empty, Vars.make [(X,cP)]) genImpl; in SIMPLE_OF ctxt genImpl' [thm] end) end | decomp ctxt ((allc as Const (@{const_name Pure.all},_)) $ f, ct) = let val cf = snd (Thm.dest_comb ct); val (abst as (x,T,_),(cx',cb)) = split_abs ctxt cf; val Free (x',_) = Thm.term_of cx'; val (Const (@{const_name Pure.all},_)$Abs (s,_,_)) = gen_all |> Thm.prems_of |> hd |> dest_prop; val gen_all' = Drule.rename_bvars [(s,x)] gen_all; val (Const (@{const_name Pure.all},_)$Abs (s',_,_)) = gen_allShift |> Thm.prems_of |> hd |> dest_prop; val gen_allShift' = Drule.rename_bvars [(s',x)] gen_allShift; in if SplitState.isState (allc$Abs abst) then ([Thm.term_of cb],[cb], fn [thm] => let val P = dest_prop (prem_of thm); val thm' = split_thm Meta_all ctxt x' T P; val thm1 = gen_allShift' OF_RAW Goal.protect 0 (Thm.forall_intr cx' (Goal.conclude thm')); val thm2 = gen_all' OF_RAW Goal.protect 0 (Thm.forall_intr cx' (Goal.conclude thm)); in SIMPLE_OF ctxt genTrans [thm1,thm2] end) else ([Thm.term_of cb],[cb], fn [thm] => gen_all' OF_RAW Goal.protect 0 (Thm.forall_intr cx' (Goal.conclude thm))) end | decomp ctxt (Const (@{const_name "Trueprop"},_)$P, ct) = ([P],snd (Drule.strip_comb ct),fn [thm] => thm) | decomp ctxt (t, ct) = ([],[], fn [] => let val rP = HOLogic.dest_Trueprop (dest_prop (conc_of genRefl)); - in Thm.instantiate ([],[(dest_Var rP, ct)]) genRefl end) + in Thm.instantiate (TVars.empty, Vars.make [(dest_Var rP, ct)]) genRefl end) fun generalise ctxt ct = gen_thm (decomp ctxt) (Thm.term_of ct,ct); (* -------- (init) #C ==> #C *) fun init ct = Thm.instantiate' [] [SOME ct] protectRefl; fun generalise_over_tac ctxt P i st = let val t = List.nth (Thm.prems_of st, i - 1); (* FIXME !? *) in (case P t of SOME t' => let val ct = Thm.cterm_of ctxt t'; val meta_spec_protect' = infer_instantiate ctxt [(("x", 0), ct)] @{thm meta_spec_protect}; in (init (Thm.adjust_maxidx_cterm 0 (List.nth (Drule.cprems_of st, i - 1))) |> resolve_tac ctxt [meta_spec_protect'] 1 |> Seq.maps (fn st' => Thm.bicompose NONE {flatten = true, match = false, incremented = false} (false, Goal.conclude st', Thm.nprems_of st') i st)) end | NONE => no_tac st) end; fun generalise_over_all_states_tac ctxt i = REPEAT (generalise_over_tac ctxt SplitState.abs_state i); fun generalise_tac ctxt i st = let val ct = List.nth (Drule.cprems_of st, i - 1); val ct' = Thm.dest_equals_rhs (Thm.cprop_of (Thm.eta_conversion ct)); val r = Goal.conclude (generalise ctxt ct'); in (init (Thm.adjust_maxidx_cterm 0 (List.nth (Drule.cprems_of st, i - 1))) |> (resolve_tac ctxt [r] 1 THEN resolve_tac ctxt [Drule.protectI] 1) |> Seq.maps (fn st' => Thm.bicompose NONE {flatten = true, match = false, incremented = false} (false, Goal.conclude st', Thm.nprems_of st') i st)) end fun GENERALISE ctxt i = generalise_over_all_states_tac ctxt i THEN generalise_tac ctxt i end; diff --git a/thys/Simpl/hoare.ML b/thys/Simpl/hoare.ML --- a/thys/Simpl/hoare.ML +++ b/thys/Simpl/hoare.ML @@ -1,3410 +1,3410 @@ (* Title: hoare.ML Author: Norbert Schirmer, TU Muenchen Copyright (C) 2004-2007 Norbert Schirmer This library is free software; you can redistribute it and/or modify it under the terms of the GNU Lesser General Public License as published by the Free Software Foundation; either version 2.1 of the License, or (at your option) any later version. This library is distributed in the hope that it will be useful, but WITHOUT ANY WARRANTY; without even the implied warranty of MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU Lesser General Public License for more details. You should have received a copy of the GNU Lesser General Public License along with this library; if not, write to the Free Software Foundation, Inc., 59 Temple Place, Suite 330, Boston, MA 02111-1307 USA *) signature HOARE = sig datatype hoareMode = Partial | Total val gen_proc_rec: Proof.context -> hoareMode -> int -> thm datatype state_kind = Record | Function datatype par_kind = In | Out val deco: string val proc_deco: string val par_deco: string -> string val chopsfx: string -> string -> string val is_state_var: string -> bool val extern: Proof.context -> string -> string val remdeco: Proof.context -> string -> string val remdeco': string -> string val undeco: Proof.context -> term -> term val varname: string -> string val resuffix: string -> string -> string -> string type proc_info = {params: ((par_kind * string) list), recursive: bool, state_kind: state_kind} type hoare_tac = (bool -> int -> tactic) -> Proof.context -> hoareMode -> int -> tactic type hoare_data = {proc_info: proc_info Symtab.table, active_procs: string list list, default_state_kind: state_kind, generate_guard: (stamp * (Proof.context -> term -> term option)), wp_tacs: (string * hoare_tac) list, hoare_tacs: (string * hoare_tac) list, vcg_simps: thm list} val get_data: Proof.context -> hoare_data val get_params: string -> Proof.context -> (par_kind * string) list option val get_default_state_kind: Proof.context -> state_kind val get_state_kind: string -> Proof.context -> state_kind option val clique_name: string list -> string val install_generate_guard: (Proof.context -> term -> term option) -> Context.generic -> Context.generic val generate_guard: Proof.context -> term -> term option val BasicSimpTac: Proof.context -> state_kind -> bool -> thm list -> (int -> tactic) -> int -> tactic val hoare: (Proof.context -> Proof.method) context_parser val hoare_raw: (Proof.context -> Proof.method) context_parser val vcg: (Proof.context -> Proof.method) context_parser val vcg_step: (Proof.context -> Proof.method) context_parser val hoare_rule: (Proof.context -> Proof.method) context_parser val add_foldcongsimps: thm list -> theory -> theory val get_foldcong_ss : theory -> simpset val add_foldcongs : thm list -> theory -> theory val modeqN : string val modexN : string val implementationN : string val specL : string val vcg_tac : string -> string -> string list -> Proof.context -> int -> tactic val hoare_rule_tac : Proof.context -> thm list -> int -> tactic datatype 'a bodykind = BodyTyp of 'a | BodyTerm of 'a val proc_specs : (bstring * string) list parser val add_params : morphism -> string -> (par_kind * string) list -> Context.generic -> Context.generic val set_default_state_kind : state_kind -> Context.generic -> Context.generic val add_state_kind : morphism -> string -> state_kind -> Context.generic -> Context.generic val add_recursive : morphism -> string -> Context.generic -> Context.generic end; structure Hoare: HOARE = struct (* Misc *) val record_vanish = Attrib.setup_config_bool @{binding hoare_record_vanish} (K true); val use_generalise = Attrib.setup_config_bool @{binding hoare_use_generalise} (K false); val sort_variables = Attrib.setup_config_bool @{binding hoare_sort_variables} (K true); val use_cond_inv_modifies = Attrib.setup_config_bool @{binding hoare_use_cond_inv_modifies} (K true); val hoare_trace = Attrib.setup_config_bool @{binding hoare_trace} (K false); val body_def_sfx = "_body"; val programN = "\"; val hoare_ctxtL = "hoare"; val specL = "_spec"; val procL = "_proc"; val bodyP = "_impl"; val modifysfx = "_modifies"; val modexN = "Hoare.mex"; val modeqN = "Hoare.meq"; val KNF = @{const_name StateFun.K_statefun}; (* Some abstract syntax operations *) val Trueprop = HOLogic.mk_Trueprop; infix 0 ===; val (op ===) = Trueprop o HOLogic.mk_eq; fun is_empty_set (Const (@{const_name Orderings.bot}, _)) = true | is_empty_set _ = false; fun mk_Int Ts A B = let val T = fastype_of1 (Ts, A) in Const (@{const_name Lattices.inf}, T --> T --> T) $ A $ B end; fun mk_Un T (A, B) = Const (@{const_name Lattices.sup}, T --> T --> T) $ A $ B; fun dest_Un (Const (@{const_name Lattices.sup}, _) $ t1 $ t2) = dest_Un t1 @ dest_Un t2 | dest_Un t = [t] fun mk_UN' dT rT t = let val dTs = HOLogic.mk_setT dT; val rTs = HOLogic.mk_setT rT; in Const (@{const_name Complete_Lattices.Sup}, rTs --> rT) $ (Const (@{const_name image}, (dT --> rT) --> dTs --> rTs) $ t $ Const (@{const_name Orderings.top}, dTs)) end; fun mk_UN ((x, T), P) = mk_UN' T (fastype_of P) (absfree (x, T) P); fun dest_UN (Const (@{const_name Complete_Lattices.Sup}, _) $ (Const (@{const_name Set.image}, _) $ Abs (x, T, t) $ Const (@{const_name Orderings.top}, _))) = let val (vars, body) = dest_UN t in ((x, T) :: vars, body) end | dest_UN t = ([], t); fun tap_UN (Const (@{const_name Complete_Lattices.Sup}, _) $ (Const (@{const_name Set.image}, _) $ t $ Const (@{const_name Orderings.top}, _))) = SOME t | tap_UN _ = NONE; (* Fetching the rules *) datatype hoareMode = Partial | Total fun get_rule p t Partial = p | get_rule p t Total = t val Guard = get_rule @{thm HoarePartial.Guard} @{thm HoareTotal.Guard}; val GuardStrip = get_rule @{thm HoarePartial.GuardStrip} @{thm HoareTotal.GuardStrip}; val GuaranteeAsGuard = get_rule @{thm HoarePartial.GuaranteeAsGuard} @{thm HoareTotal.GuaranteeAsGuard}; val Guarantee = get_rule @{thm HoarePartial.Guarantee} @{thm HoareTotal.Guarantee}; val GuaranteeStrip = get_rule @{thm HoarePartial.GuaranteeStrip} @{thm HoareTotal.GuaranteeStrip}; val GuardsNil = get_rule @{thm HoarePartial.GuardsNil} @{thm HoareTotal.GuardsNil}; val GuardsCons = get_rule @{thm HoarePartial.GuardsCons} @{thm HoareTotal.GuardsCons}; val GuardsConsGuaranteeStrip = get_rule @{thm HoarePartial.GuardsConsGuaranteeStrip} @{thm HoareTotal.GuardsConsGuaranteeStrip}; val Skip = get_rule @{thm HoarePartial.Skip} @{thm HoareTotal.Skip}; val Basic = get_rule @{thm HoarePartial.Basic} @{thm HoareTotal.Basic}; val BasicCond = get_rule @{thm HoarePartial.BasicCond} @{thm HoareTotal.BasicCond}; val Spec = get_rule @{thm HoarePartial.Spec} @{thm HoareTotal.Spec}; val SpecIf = get_rule @{thm HoarePartial.SpecIf} @{thm HoareTotal.SpecIf}; val Throw = get_rule @{thm HoarePartial.Throw} @{thm HoareTotal.Throw}; val Raise = get_rule @{thm HoarePartial.raise} @{thm HoareTotal.raise}; val Catch = get_rule @{thm HoarePartial.Catch} @{thm HoareTotal.Catch}; val CondCatch = get_rule @{thm HoarePartial.condCatch} @{thm HoareTotal.condCatch}; val CatchSwap = get_rule @{thm HoarePartial.CatchSwap} @{thm HoareTotal.CatchSwap}; val CondCatchSwap = get_rule @{thm HoarePartial.condCatchSwap} @{thm HoareTotal.condCatchSwap}; val Seq = get_rule @{thm HoarePartial.Seq} @{thm HoareTotal.Seq}; val SeqSwap = get_rule @{thm HoarePartial.SeqSwap} @{thm HoareTotal.SeqSwap}; val BSeq = get_rule @{thm HoarePartial.BSeq} @{thm HoareTotal.BSeq}; val Cond = get_rule @{thm HoarePartial.Cond} @{thm HoareTotal.Cond}; val CondInv'Partial = @{thm HoarePartial.CondInv'}; val CondInv'Total = @{thm HoareTotal.CondInv'}; val CondInv' = get_rule CondInv'Partial CondInv'Total; val SwitchNil = get_rule @{thm HoarePartial.switchNil} @{thm HoareTotal.switchNil}; val SwitchCons = get_rule @{thm HoarePartial.switchCons} @{thm HoareTotal.switchCons}; val CondSwap = get_rule @{thm HoarePartial.CondSwap} @{thm HoareTotal.CondSwap}; val While = get_rule @{thm HoarePartial.While} @{thm HoareTotal.While}; val WhileAnnoG = get_rule @{thm HoarePartial.WhileAnnoG} @{thm HoareTotal.WhileAnnoG}; val WhileAnnoFix = get_rule @{thm HoarePartial.WhileAnnoFix'} @{thm HoareTotal.WhileAnnoFix'}; val WhileAnnoGFix = get_rule @{thm HoarePartial.WhileAnnoGFix} @{thm HoareTotal.WhileAnnoGFix}; val BindR = get_rule @{thm HoarePartial.Bind} @{thm HoareTotal.Bind}; val Block = get_rule @{thm HoarePartial.Block} @{thm HoareTotal.Block}; val BlockSwap = get_rule @{thm HoarePartial.BlockSwap} @{thm HoareTotal.BlockSwap}; val Proc = get_rule @{thm HoarePartial.ProcSpec} @{thm HoareTotal.ProcSpec}; val ProcNoAbr = get_rule @{thm HoarePartial.ProcSpecNoAbrupt} @{thm HoareTotal.ProcSpecNoAbrupt}; val ProcBody = get_rule @{thm HoarePartial.ProcBody} @{thm HoareTotal.ProcBody}; val CallBody = get_rule @{thm HoarePartial.CallBody} @{thm HoareTotal.CallBody}; val FCall = get_rule @{thm HoarePartial.FCall} @{thm HoareTotal.FCall}; val ProcRecSpecs = get_rule @{thm HoarePartial.ProcRecSpecs} @{thm HoareTotal.ProcRecSpecs}; val ProcModifyReturnSameFaults = get_rule @{thm HoarePartial.ProcModifyReturnSameFaults} @{thm HoareTotal.ProcModifyReturnSameFaults}; val ProcModifyReturn = get_rule @{thm HoarePartial.ProcModifyReturn} @{thm HoareTotal.ProcModifyReturn}; val ProcModifyReturnNoAbr = get_rule @{thm HoarePartial.ProcModifyReturnNoAbr} @{thm HoareTotal.ProcModifyReturnNoAbr}; val ProcModifyReturnNoAbrSameFaultsPartial = @{thm HoarePartial.ProcModifyReturnNoAbrSameFaults}; val ProcModifyReturnNoAbrSameFaultsTotal = @{thm HoareTotal.ProcModifyReturnNoAbrSameFaults}; val ProcModifyReturnNoAbrSameFaults = get_rule ProcModifyReturnNoAbrSameFaultsPartial ProcModifyReturnNoAbrSameFaultsTotal; val TrivPost = get_rule @{thm HoarePartial.TrivPost} @{thm HoareTotal.TrivPost}; val TrivPostNoAbr = get_rule @{thm HoarePartial.TrivPostNoAbr} @{thm HoareTotal.TrivPostNoAbr}; val DynProcProcPar = get_rule @{thm HoarePartial.DynProcProcPar} @{thm HoareTotal.DynProcProcPar}; val DynProcProcParNoAbr = get_rule @{thm HoarePartial.DynProcProcParNoAbrupt} @{thm HoareTotal.DynProcProcParNoAbrupt}; val ProcProcParModifyReturn = get_rule @{thm HoarePartial.ProcProcParModifyReturn} @{thm HoareTotal.ProcProcParModifyReturn}; val ProcProcParModifyReturnSameFaultsPartial = @{thm HoarePartial.ProcProcParModifyReturnSameFaults}; val ProcProcParModifyReturnSameFaultsTotal = @{thm HoareTotal.ProcProcParModifyReturnSameFaults}; val ProcProcParModifyReturnSameFaults = get_rule ProcProcParModifyReturnSameFaultsPartial ProcProcParModifyReturnSameFaultsTotal; val ProcProcParModifyReturnNoAbr = get_rule @{thm HoarePartial.ProcProcParModifyReturnNoAbr} @{thm HoareTotal.ProcProcParModifyReturnNoAbr}; val ProcProcParModifyReturnNoAbrSameFaultsPartial = @{thm HoarePartial.ProcProcParModifyReturnNoAbrSameFaults}; val ProcProcParModifyReturnNoAbrSameFaultsTotal = @{thm HoareTotal.ProcProcParModifyReturnNoAbrSameFaults}; val ProcProcParModifyReturnNoAbrSameFaults = get_rule ProcProcParModifyReturnNoAbrSameFaultsPartial ProcProcParModifyReturnNoAbrSameFaultsTotal; val DynCom = get_rule @{thm HoarePartial.DynComConseq} @{thm HoareTotal.DynComConseq}; val AugmentContext = get_rule @{thm HoarePartial.augment_context'} @{thm HoareTotal.augment_context'}; val AugmentEmptyFaults = get_rule @{thm HoarePartial.augment_emptyFaults} @{thm HoareTotal.augment_emptyFaults}; val AsmUN = get_rule @{thm HoarePartial.AsmUN} @{thm HoareTotal.AsmUN}; val SpecAnno = get_rule @{thm HoarePartial.SpecAnno'} @{thm HoareTotal.SpecAnno'}; val SpecAnnoNoAbrupt = get_rule @{thm HoarePartial.SpecAnnoNoAbrupt} @{thm HoareTotal.SpecAnnoNoAbrupt}; val LemAnno = get_rule @{thm HoarePartial.LemAnno} @{thm HoareTotal.LemAnno}; val LemAnnoNoAbrupt = get_rule @{thm HoarePartial.LemAnnoNoAbrupt} @{thm HoareTotal.LemAnnoNoAbrupt}; val singleton_conv_sym = @{thm Set.singleton_conv2} RS sym; val anno_defs = [@{thm Language.whileAnno_def},@{thm Language.whileAnnoG_def},@{thm Language.specAnno_def}, @{thm Language.whileAnnoGFix_def},@{thm Language.whileAnnoFix_def},@{thm Language.lem_def}]; val strip_simps = @{thm Language.strip_simp} :: @{thm Option.option.map(2)} :: @{thms Language.strip_guards_simps}; val normalize_simps = [@{thm Language.while_def}, @{thm Language.bseq_def}, @{thm List.append_Nil}, @{thm List.append_Cons}] @ @{thms List.list.cases} @ @{thms Language.flatten_simps} @ @{thms Language.sequence.simps} @ @{thms Language.normalize_simps} @ @{thms Language.guards.simps} @ [@{thm fst_conv}, @{thm snd_conv}]; val K_rec_convs = []; val K_fun_convs = [@{thm StateFun.K_statefun_apply}, @{thm StateFun.K_statefun_comp}]; val K_convs = K_rec_convs @ K_fun_convs; val K_rec_congs = []; val K_fun_congs = [@{thm StateFun.K_statefun_cong}]; val K_congs = K_rec_congs @ K_fun_congs; (* misc. aux. functions *) (* first_subterm * yields result x of P for first subterm for which P is (SOME x), and all bound variables on the path * to that term *) fun first_subterm_dest P = let fun first abs_vars t = (case P t of SOME x => SOME (abs_vars,x) |_=> (case t of u $ v => (case first abs_vars u of NONE => first abs_vars v | SOME x => SOME x) | Abs (c,T,u) => first (abs_vars @ [(c,T)]) u | _ => NONE)) in first [] end; (* first_subterm * yields first subterm for which P holds, and all bound variables on the path * to that term *) fun first_subterm P = let fun P' t = if P t then SOME t else NONE; in first_subterm_dest P' end; (* max_subterm_dest * yields results of P for all maximal subterms for which P is (SOME x), * and all bound variables on the path to that subterm *) fun max_subterms_dest P = let fun collect abs_vars t = (case P t of SOME x => [(abs_vars,x)] |_=> (case t of u $ v => collect abs_vars u @ collect abs_vars v | Abs (c,T,u) => collect (abs_vars @ [(c,T)]) u | _ => [])) in collect [] end; fun last [] = raise Empty | last [x] = x | last (_::xs) = last xs; fun dest_splits (Const (@{const_name case_prod},_)$Abs (n,T,t)) = (n,T)::dest_splits t | dest_splits (Const (@{const_name case_prod},_)$Abs (n,T,t)$_) = (n,T)::dest_splits t | dest_splits (Abs (n,T,_)) = [(n,T)] | dest_splits _ = []; fun idx eq [] x = ~1 | idx eq (x::rs) y = if eq x y then 0 else let val i = idx eq rs y in if i < 0 then i else i+1 end; fun resuffix sfx1 sfx2 s = suffix sfx2 (unsuffix sfx1 s) handle Fail _ => s; (* state space representation dependent functions *) datatype state_kind = Record | Function fun state_simprocs Record = [Record.simproc] | state_simprocs Function = [Record.simproc, StateFun.lookup_simproc]; fun state_upd_simproc Record = Record.upd_simproc | state_upd_simproc Function = StateFun.update_simproc; fun state_ex_sel_eq_simproc Record = Record.ex_sel_eq_simproc | state_ex_sel_eq_simproc Function = StateFun.ex_lookup_eq_simproc; val state_split_simp_tac = Record.split_simp_tac val state_hierarchy = Record.dest_recTs fun stateT_id T = case (state_hierarchy T) of [] => NONE | Ts => SOME (last Ts); fun globalsT (Type (_, T :: _)) = SOME T | globalsT _ = NONE; fun stateT_ids T = (case stateT_id T of NONE => NONE | SOME sT => (case globalsT T of NONE => SOME [sT] | SOME gT => (case stateT_id gT of NONE => SOME [sT] | SOME gT' => SOME [sT,gT']))); datatype par_kind = In | Out (*** utilities ***) (* utils for variable name decorations *) val deco = "_'"; val proc_deco = "_'proc"; fun par_deco name = deco ^ name ^ deco; fun chopsfx sfx str = (case try (unsuffix sfx) str of SOME s => s | NONE => str) val is_state_var = can (unsuffix deco); (* removes the suffix of the string beginning with deco. * "xys_'a" --> "xys"; * The a is also chopped, since sometimes the bound variables * are renamed, I think SELECT_GOAL in rename_goal is to blame *) fun remdeco' str = let fun chop (p::ps) (x::xs) = chop ps xs | chop [] xs = [] | chop (p::ps) [] = error "remdeco: code should never be reached"; fun remove prf (s as (x::xs)) = if is_prefix (op =) prf s then chop prf s else (x::remove prf xs) | remove prf [] = []; in String.implode (remove (String.explode deco) (String.explode str)) end; fun extern ctxt s = (case try (Proof_Context.extern_const ctxt o Lexicon.unmark_const) s of NONE => s | SOME s' => s'); fun remdeco ctxt s = remdeco' (extern ctxt s); fun undeco ctxt (Const (c, T)) = Const (remdeco ctxt c, T) | undeco ctxt ((f as Const (@{syntax_const "_free"},_)) $ Free (x, T)) = (*f$*)Const (remdeco' x, T) | undeco ctxt (Const _ $ _ $ ((Const (@{syntax_const "_free"},_)) $ Free (x, T))) = (*f$*)Const (remdeco' x, T) | undeco ctxt (Free (c, T)) = Const (remdeco' c, T) | undeco ctxt x = x fun varname x = x ^ deco val dest_string = map (chr o HOLogic.dest_char) o HOLogic.dest_list; fun dest_string' t = (case try dest_string t of SOME s => implode s | NONE => (case t of Free (s,_) => s | Const (s,_) => Long_Name.base_name s | _ => raise TERM ("dest_string'",[t]))) fun is_state_space_var Tids t = let fun is_stateT T = (case stateT_id T of NONE => 0 | SOME id => if member (op =) Tids id then ~1 else 0); in (case t of Const _ $ Abs (_,T,_) => is_stateT T | Free (_,T) => is_stateT T | _ => 0) end; datatype callMode = Static | Parameter fun proc_name Static (Const (p,_)$_) = resuffix deco proc_deco (Long_Name.base_name p) | proc_name Static (Const (@{const_name StateFun.lookup},_)$_$Free (p,_)$_) = suffix proc_deco (remdeco' (Long_Name.base_name p)) | proc_name Static p = dest_string' p | proc_name Parameter (Const (p,_)) = resuffix deco proc_deco (Long_Name.base_name p) | proc_name Parameter (Abs (_,_,Const (p,_)$Bound 0)) = resuffix deco proc_deco (Long_Name.base_name p) | proc_name Parameter (Abs (_,_,Const (@{const_name StateFun.lookup},_)$_$Free (p,_)$_)) = suffix proc_deco (remdeco' (Long_Name.base_name p)) | proc_name _ t = raise TERM ("proc_name",[t]); fun dest_call (Const (@{const_name Language.call},_)$init$pname$return$c) = (init,pname,return,c,Static,true) | dest_call (Const (@{const_name Language.fcall},_)$init$pname$return$_$c) = (init,pname,return,c,Static,true) | dest_call (Const (@{const_name Language.com.Call},_)$pname) = (Bound 0,pname,Bound 0,Bound 0,Static,false) | dest_call (Const (@{const_name Language.dynCall},_)$init$pname$return$c) = (init,pname,return,c,Parameter,true) | dest_call t = raise TERM ("Hoare.dest_call: unexpected term",[t]); fun dest_whileAnno (Const (@{const_name Language.whileAnnoG},_) $gs$b$I$V$c) = (SOME gs,b,I,V,c,false) | dest_whileAnno (Const (@{const_name Language.whileAnno},_) $b$I$V$c) = (NONE,b,I,V,c,false) | dest_whileAnno (Const (@{const_name Language.whileAnnoGFix},_)$gs$b$I$V$c) = (SOME gs,b,I,V,c,true) | dest_whileAnno (Const (@{const_name Language.whileAnnoFix},_) $b$I$V$c) = (NONE,b,I,V,c,true) | dest_whileAnno t = raise TERM ("Hoare.dest_while: unexpected term",[t]); fun dest_Guard (Const (@{const_name Language.com.Guard},_)$f$g$c) = (f,g,c,false) | dest_Guard (Const (@{const_name Language.guaranteeStrip},_)$f$g$c) = (f,g,c,true) | dest_Guard t = raise TERM ("Hoare.dest_guard: unexpected term",[t]); (*** extend theory by procedure definition ***) fun add_declaration name decl thy = thy |> Named_Target.init [] name |> Local_Theory.declaration {syntax = false, pervasive = false} decl |> Local_Theory.exit |> Proof_Context.theory_of; (* data kind 'HOL/hoare' *) type proc_info = {params: ((par_kind * string) list), recursive: bool, state_kind: state_kind} type hoare_tac = (bool -> int -> tactic) -> Proof.context -> hoareMode -> int -> tactic; type hoare_data = {proc_info: proc_info Symtab.table, active_procs: string list list, default_state_kind: state_kind, generate_guard: (stamp * (Proof.context -> term -> term option)), wp_tacs: (string * hoare_tac) list, hoare_tacs: (string * hoare_tac) list, vcg_simps: thm list}; fun make_hoare_data proc_info active_procs default_state_kind generate_guard wp_tacs hoare_tacs vcg_simps = {proc_info = proc_info, active_procs = active_procs, default_state_kind = default_state_kind, generate_guard = generate_guard, wp_tacs = wp_tacs, hoare_tacs = hoare_tacs, vcg_simps = vcg_simps}; structure Hoare_Data = Generic_Data ( type T = hoare_data; val empty = make_hoare_data (Symtab.empty: proc_info Symtab.table) ([]:string list list) (Function) (stamp (),(K (K NONE)): Proof.context -> term -> term option) ([]:(string * hoare_tac) list) ([]:(string * hoare_tac) list) ([]:thm list); val extend = I; (* FIXME exponential blowup due to append !? *) fun merge ({proc_info = proc_info1, active_procs = active_procs1, default_state_kind = _, generate_guard = (stmp1,generate_gaurd1), wp_tacs = wp_tacs1, hoare_tacs = hoare_tacs1, vcg_simps = vcg_simps1}, {proc_info = proc_info2, active_procs = active_procs2, default_state_kind = default_state_kind2, generate_guard = (stmp2, _), wp_tacs = wp_tacs2, hoare_tacs = hoare_tacs2, vcg_simps=vcg_simps2}) : T = if stmp1=stmp2 then make_hoare_data (Symtab.merge (K true) (proc_info1,proc_info2)) (active_procs1 @ active_procs2) (default_state_kind2) (stmp1,generate_gaurd1) (wp_tacs1 @ wp_tacs2) (hoare_tacs1 @ hoare_tacs2) (Thm.merge_thms (vcg_simps1,vcg_simps2)) else error ("Theories have different aux. functions to generate guards") ); val get_data = Hoare_Data.get o Context.Proof; (* access 'params' *) fun mk_free ctxt name = let val ctxt' = Context.proof_of ctxt; val n' = Variable.intern_fixed ctxt' name |> perhaps Long_Name.dest_hidden; val T' = Proof_Context.infer_type ctxt' (n', dummyT) handle ERROR _ => dummyT in (Free (n',T')) end; fun morph_name ctxt phi name = (case Morphism.term phi (mk_free ctxt name) of Free (x,_) => x | _ => name); datatype 'a bodykind = BodyTyp of 'a | BodyTerm of 'a fun set_default_state_kind sk ctxt = let val {proc_info,active_procs,default_state_kind,generate_guard,wp_tacs,hoare_tacs, vcg_simps,...} = Hoare_Data.get ctxt; val data = make_hoare_data proc_info active_procs sk generate_guard wp_tacs hoare_tacs vcg_simps; in Hoare_Data.put data ctxt end; val get_default_state_kind = #default_state_kind o get_data; fun add_active_procs phi ps ctxt = let val {proc_info,active_procs,default_state_kind,generate_guard,wp_tacs,hoare_tacs, vcg_simps,...} = Hoare_Data.get ctxt; val data = make_hoare_data proc_info ((map (morph_name ctxt phi) ps)::active_procs) default_state_kind generate_guard wp_tacs hoare_tacs vcg_simps; in Hoare_Data.put data ctxt end; fun add_hoare_tacs tacs ctxt = let val {proc_info,active_procs, default_state_kind, generate_guard,wp_tacs,hoare_tacs, vcg_simps,...} = Hoare_Data.get ctxt; val data = make_hoare_data proc_info active_procs default_state_kind generate_guard wp_tacs (hoare_tacs@tacs) vcg_simps; in Hoare_Data.put data ctxt end; fun map_vcg_simps f ctxt = let val {proc_info,active_procs,default_state_kind,generate_guard,wp_tacs,hoare_tacs, vcg_simps,...} = Hoare_Data.get ctxt; val data = make_hoare_data proc_info active_procs default_state_kind generate_guard wp_tacs hoare_tacs (f vcg_simps); in Hoare_Data.put data ctxt end; fun thy_attrib f = Thm.declaration_attribute (fn thm => map_vcg_simps (f thm)); val vcg_simpadd = Thm.add_thm val vcg_simpdel = Thm.del_thm val vcg_simp_add = thy_attrib vcg_simpadd; val vcg_simp_del = thy_attrib vcg_simpdel; (* add 'procedure' *) fun mk_proc_info params recursive state_kind = {params=params,recursive=recursive,state_kind=state_kind}; val empty_proc_info = {params=[],recursive=false,state_kind=Record}; fun map_proc_info_params f {params,recursive,state_kind} = mk_proc_info (f params) recursive state_kind; fun map_proc_info_recursive f {params,recursive,state_kind} = mk_proc_info params (f recursive) state_kind; fun map_proc_info_state_kind f {params,recursive,state_kind} = mk_proc_info params recursive (f state_kind); fun add_params phi name frmls ctxt = let val {proc_info,active_procs,default_state_kind,generate_guard,wp_tacs,hoare_tacs, vcg_simps,...} = Hoare_Data.get ctxt; val params = map (apsnd (morph_name ctxt phi)) frmls; val f = map_proc_info_params (K params); val default = f empty_proc_info; val proc_info' = Symtab.map_default (morph_name ctxt phi name,default) f proc_info; val data = make_hoare_data proc_info' active_procs default_state_kind generate_guard wp_tacs hoare_tacs vcg_simps; in Hoare_Data.put data ctxt end; fun get_params name ctxt = Option.map #params (Symtab.lookup (#proc_info (get_data ctxt)) name); fun add_recursive phi name ctxt = let val {proc_info,active_procs,default_state_kind,generate_guard,wp_tacs,hoare_tacs, vcg_simps,...} = Hoare_Data.get ctxt; val f = map_proc_info_recursive (K true); val default = f empty_proc_info; val proc_info'= Symtab.map_default (morph_name ctxt phi name,default) f proc_info; val data = make_hoare_data proc_info' active_procs default_state_kind generate_guard wp_tacs hoare_tacs vcg_simps; in Hoare_Data.put data ctxt end; fun get_recursive name ctxt = Option.map #recursive (Symtab.lookup (#proc_info (get_data ctxt)) name); fun add_state_kind phi name sk ctxt = let val {proc_info,active_procs,default_state_kind,generate_guard,wp_tacs,hoare_tacs, vcg_simps,...} = Hoare_Data.get ctxt; val f = map_proc_info_state_kind (K sk); val default = f empty_proc_info; val proc_info'= Symtab.map_default (morph_name ctxt phi name,default) f proc_info; val data = make_hoare_data proc_info' active_procs default_state_kind generate_guard wp_tacs hoare_tacs vcg_simps; in Hoare_Data.put data ctxt end; fun get_state_kind name ctxt = Option.map #state_kind (Symtab.lookup (#proc_info (get_data ctxt)) name); fun install_generate_guard f ctxt = let val {proc_info,active_procs, default_state_kind, generate_guard,wp_tacs,hoare_tacs, vcg_simps,...} = Hoare_Data.get ctxt; val data = make_hoare_data proc_info active_procs default_state_kind (stamp (), f) wp_tacs hoare_tacs vcg_simps in Hoare_Data.put data ctxt end; fun generate_guard ctxt = snd (#generate_guard (get_data ctxt)) ctxt; fun check_procedures_definition procs thy = let val ctxt = Proof_Context.init_global thy; fun already_defined name = if is_some (get_params name ctxt) then ["procedure " ^ quote name ^ " already defined"] else [] val err_already_defined = maps (already_defined o #1) procs; fun duplicate_procs names = (case duplicates (op =) names of [] => [] | dups => ["Duplicate procedures " ^ commas_quote dups]); val err_duplicate_procs = duplicate_procs (map #1 procs); fun duplicate_pars name pars = (case duplicates (op =) (map fst pars) of [] => [] | dups => ["Duplicate parameters in procedure " ^ quote name ^ ": " ^ commas_quote dups]); val err_duplicate_pars = maps (fn (name,inpars,outpars,locals,_,_,_) => duplicate_pars name (inpars @ locals) @ duplicate_pars name (outpars @ locals)) procs; (* FIXME: Check that no global variables are used as result parameters *) val errs = err_already_defined @ err_duplicate_procs @ err_duplicate_pars; in if null errs then () else error (cat_lines errs) end; fun add_parameter_info phi cname (name,(inpars,outpars,state_kind)) ctxt = let fun par_deco' T = if T = "" then deco else par_deco (cname name); val pars = map (fn (par,T) => (In,suffix (par_deco' T) par)) inpars@ map (fn (par,T) => (Out,suffix (par_deco' T) par)) outpars; val ctxt_decl = ctxt |> add_params phi name pars |> add_state_kind phi name state_kind in ctxt_decl end; fun mk_loc_exp xs = let fun mk_expr s = (s,(("",false),(Expression.Named [],[]))) in (map mk_expr xs,[]) end; val parametersN = "_parameters"; val variablesN = "_variables"; val signatureN = "_signature"; val bodyN = "_body"; val implementationN = "_impl"; val cliqueN = "_clique"; val clique_namesN = "_clique_names"; val NoBodyN = @{const_name Vcg.NoBody}; val statetypeN = "StateType"; val proc_nameT = HOLogic.stringT; fun expression_no_pos (expr, fixes) : Expression.expression = (map (fn (name, inst) => ((name, Position.none), inst)) expr, fixes); fun add_locale name expr elems thy = thy |> Expression.add_locale (Binding.name name) (Binding.name name) [] expr elems |> snd |> Local_Theory.exit; fun add_locale' name expr elems thy = thy |> Expression.add_locale (Binding.name name) (Binding.name name) [] expr elems ||> Local_Theory.exit; fun add_locale_cmd name expr elems thy = thy |> Expression.add_locale_cmd (Binding.name name) (Binding.name name) [] (expression_no_pos expr) elems |> snd |> Local_Theory.exit; fun read_typ thy raw_T env = let val ctxt' = Proof_Context.init_global thy |> fold (Variable.declare_typ o TFree) env; val T = Syntax.read_typ ctxt' raw_T; val env' = Term.add_tfreesT T env; in (T, env') end; fun add_variable_statespaces (cname, (inpars, outpars, locvars)) thy = let val inpars' = if forall (fn (_,T) => T = "") inpars then [] else inpars; val outpars' = if forall (fn (_,T) => T = "") outpars then [] else outpars; fun prep_comp (n, T) env = let val (T', env') = read_typ thy T env handle ERROR msg => cat_error msg ("The error(s) above occurred in component " ^ quote n) in ((n, T'), env') end; val (in_outs,in_out_env) = fold_map prep_comp (distinct (op =) (inpars'@outpars')) []; val (locs,var_env) = fold_map prep_comp locvars in_out_env; val parSP = cname ^ parametersN; val in_outs' = map (apfst (suffix (par_deco cname))) in_outs; val in_out_args = map fst in_out_env; val varSP = cname ^ variablesN; val locs' = map (apfst (suffix (par_deco cname))) locs; val var_args = map fst var_env; in if null inpars' andalso null outpars' andalso null locvars then thy |> add_locale_cmd parSP ([],[]) [] |> Proof_Context.theory_of |> add_locale_cmd varSP ([],[]) [] |> Proof_Context.theory_of else thy |> StateSpace.define_statespace_i (SOME statetypeN) in_out_args parSP [] in_outs' |> StateSpace.define_statespace_i (SOME statetypeN) var_args varSP [((cname, false), ((map TFree in_out_env),parSP,[]))] locs' end; fun intern_locale thy = Locale.intern thy #> perhaps Long_Name.dest_hidden; fun apply_in_context thy lexp f t = let fun name_variant lname = if intern_locale thy lname = lname then lname else name_variant (lname ^ "'"); in thy (* Create a dummy locale in dummy theory just to read the term *) |> add_locale_cmd (name_variant "foo") lexp [] |> (fn ctxt => f ctxt t) end; fun add_abbrev loc mode name spec thy = thy |> Named_Target.init [] loc |> (fn lthy => let val t = Syntax.read_term (Local_Theory.target_of lthy) spec; in Local_Theory.abbrev mode ((Binding.name name, NoSyn), t) lthy end) |> #2 |> Local_Theory.exit |> Proof_Context.theory_of; exception TOPSORT of string fun topsort less [] = [] | topsort less xs = let fun list_all P xs = fold (fn x => fn b => b andalso P x) xs true; fun split_min n (x::xs) = if n=0 then raise TOPSORT "no minimum in list" else if list_all (less x) xs then (x,xs) else split_min (n-1) (xs@[x]); fun tsort [] = [] | tsort xs = let val (x,xs') = split_min (length xs) xs; in x::tsort xs' end; in tsort xs end; fun clique_name clique = (foldr1 (fn (a,b) => a ^ "_" ^ b) (map (unsuffix proc_deco) clique)); fun error_to_warning msg f thy = f thy handle ERROR msg' => (warning (msg' ^ "\n" ^ msg); thy); fun procedures_definition locname procs thy = let val procs' = map (fn (name,a,b,c,d,e,f) => (suffix proc_deco name,a,b,c,d,e,f)) procs; val _ = check_procedures_definition procs' thy; val name_pars = map (fn (name,inpars,outpars,_,_,_,sk) => (name,(inpars,outpars,sk))) procs'; val name_vars = map (fn (name,inpars,outpars,locals,_,_,_) => (name,(inpars,outpars,locals))) procs'; val name_body = map (fn (name,_,_,_,body,_,_) => (name,body)) procs'; val name_pars_specs = map (fn (name,inpars,outpars,_,_,specs,sk) => (name,(inpars,outpars,sk),specs)) procs'; val names = map #1 procs'; val sk = #7 (hd procs'); val thy = thy |> Context.theory_map (set_default_state_kind sk); val (all_callss,cliques,is_recursive,has_body) = let val ctxt = Context.Theory thy |> fold (add_parameter_info Morphism.identity (unsuffix proc_deco)) name_pars |> StateSpace.set_silent true fun read_body (_, body) = Syntax.read_term (Context.proof_of ctxt) body; val bodies = map read_body name_body; fun dcall t = (case try dest_call t of SOME (_,p,_,_,m,_) => SOME (proc_name m p) | _ => NONE); fun in_names x = if member (op =) names x then SOME x else NONE; fun add_edges n = fold (fn x => Graph.add_edge (n, x)); val all_callss = map (map snd o max_subterms_dest dcall) bodies; val callss = map (map_filter in_names) all_callss; val graph = fold (fn n => Graph.new_node (n, ())) names Graph.empty; val graph' = fold2 add_edges names callss graph; fun idx x = find_index (fn y => x=y) names; fun name_ord (a,b) = int_ord (idx a, idx b); val cliques = Graph.strong_conn graph'; val cliques' = map (sort name_ord) cliques; val my_calls = these o AList.lookup (op =) (names ~~ map (distinct (op =)) callss); val my_body = AList.lookup (op =) (names ~~ bodies); fun is_recursive n = exists (fn [_] => false | ns => member (op =) ns n) (Graph.strong_conn graph'); fun has_body n = (case my_body n of SOME (Const (c,_)) => c <> NoBodyN | _ => true) fun clique_less c1 c2 = null (inter (op =) (distinct (op =) (maps my_calls c1)) c2); val cliques'' = topsort clique_less cliques'; in (all_callss,cliques'',is_recursive,has_body) end; (* cliques may only depend on ones to the left, so it is safe to * add the locales from the left to the right. *) fun my_clique n = Library.find_first (fn c => member (op =) c n) cliques; fun lname sfx clique = suffix sfx (clique_name clique); fun cname n = clique_name (the (my_clique n)); fun parameter_info_decl phi = fold (add_parameter_info phi cname) name_pars; fun get_loc sfx clique n = if member (op =) clique n then NONE else SOME (resuffix proc_deco sfx n); fun parent_locales thy sfx clique = let val calls = distinct (op =) (flat (map_filter (AList.lookup (op =) (names ~~ all_callss)) clique)); in map (intern_locale thy) (distinct (op =) (map_filter (get_loc sfx clique) calls)) end; val names_all_callss = names ~~ map (distinct (op =)) all_callss; val get_calls = the o AList.lookup (op =) names_all_callss; fun clique_vars clique = let fun add name (ins,outs,locs) = let val (nins,nouts,nlocs) = the (AList.lookup (op =) name_vars name) in (ins@nins,outs@nouts,locs@nlocs) end; val (is,os,ls) = fold add clique ([],[],[]); in (lname "" clique, (distinct (op =) is, distinct (op =) os, distinct (op =) ls)) end; fun add_signature_locale (cname, name) thy = let val name' = unsuffix proc_deco name; val fixes = [Element.Fixes [(Binding.name name, SOME proc_nameT, NoSyn)]]; (* FIXME: may use HOLogic.typeT as soon as locale type-inference works properly *) val pE = mk_loc_exp [intern_locale thy (suffix parametersN cname)]; val sN = suffix signatureN name'; in thy |> add_locale sN pE fixes |> Proof_Context.theory_of |> (fn thy => add_declaration (intern_locale thy sN) parameter_info_decl thy) end; fun mk_bdy_def read_term name = let val name' = unsuffix proc_deco name; val bdy = read_term (the (AList.lookup (op =) name_body name)); val bdy_defN = suffix body_def_sfx name'; val b = Binding.name bdy_defN; in ((b, NoSyn), ((Thm.def_binding b, []), bdy)) end; fun add_body_locale (name, _) thy = let val name' = unsuffix proc_deco name; val callees = filter_out (fn n => n = name) (get_calls name) val fixes = [Element.Fixes [(Binding.name name, SOME proc_nameT, NoSyn)]]; (* FIXME: may use HOLogic.typeT as soon as locale type-inference works properly *) val pE = mk_loc_exp (map (intern_locale thy) ([lname variablesN (the (my_clique name))]@ the_list locname@ map (resuffix proc_deco signatureN) callees)); fun def lthy = let val read = Syntax.read_term (Context.proof_map (add_active_procs Morphism.identity (the (my_clique name))) (Local_Theory.target_of lthy)) in mk_bdy_def read name end; fun add_decl_and_def lname ctxt = ctxt |> Proof_Context.theory_of |> Named_Target.init [] lname |> Local_Theory.declaration {syntax = false, pervasive = false} parameter_info_decl |> (fn lthy => if has_body name then snd (Local_Theory.define (def lthy) lthy) else lthy) |> Local_Theory.exit |> Proof_Context.theory_of; in thy |> add_locale' (suffix bodyN name') pE fixes |-> add_decl_and_def end; fun mk_def_eq thy read_term name = if has_body name then let (* FIXME: All the read_term stuff is just because type-inference/abbrevs for * new locale elements does not work right now; * We read the term to expand the abbreviations, then we print it again * (without folding the abbreviation) and reread as string *) val name' = unsuffix proc_deco name; val bdy_defN = suffix body_def_sfx name'; val rhs = read_term ("Some " ^ bdy_defN); val nt = read_term name; val Free (gamma,_) = read_term programN; val eq = HOLogic.Trueprop$ HOLogic.mk_eq (Free (gamma,fastype_of nt --> fastype_of rhs)$nt,rhs) val consts = Sign.consts_of thy; val eqs = YXML.string_of_body (Term_XML.Encode.term consts (Consts.dummy_types consts eq)); val assms = Element.Assumes [((Binding.name (suffix bodyP name'), []),[(eqs,[])])] in [assms] end else []; fun add_impl_locales clique thy = let val cliqN = lname cliqueN clique; val cnamesN = lname clique_namesN clique; val multiple_procs = length clique > 1; val add_distinct_procs_namespace = if multiple_procs then StateSpace.namespace_definition cnamesN proc_nameT ([],[]) [] clique else I; val bodyLs = map (suffix bodyN o unsuffix proc_deco) clique; fun pE thy = mk_loc_exp (map (intern_locale thy) (hoare_ctxtL::bodyLs) @ (parent_locales thy implementationN clique) @ (if multiple_procs then [intern_locale thy cnamesN] else [])); fun read_term thy = apply_in_context thy (pE thy) Syntax.read_term; fun elems thy = maps (mk_def_eq thy (read_term thy)) clique; fun add_recursive_info phi name = if is_recursive name then (add_recursive phi name) else I; fun proc_declaration phi = add_active_procs phi clique; fun recursive_declaration phi ctxt = ctxt |> fold (add_recursive_info phi) clique; fun add_impl_locale name thy = let val implN = suffix implementationN (unsuffix proc_deco name); val parentN = intern_locale thy cliqN val parent = mk_loc_exp [parentN]; in thy |> add_locale_cmd implN parent [] |> Proof_Context.theory_of |> (fn thy => Interpretation.global_sublocale parentN (mk_loc_exp [intern_locale thy implN]) [] thy) |> Proof.global_terminal_proof ((Method.Basic (fn ctxt => Method.SIMPLE_METHOD (Locale.intro_locales_tac {strict = true, eager = false} ctxt [])), Position.no_range), NONE) |> Proof_Context.theory_of end; in thy |> add_distinct_procs_namespace |> (fn thy => add_locale_cmd cliqN (pE thy) (elems thy) thy) |> Proof_Context.theory_of |> fold add_impl_locale clique |> (fn thy => add_declaration (intern_locale thy cliqN) proc_declaration thy) |> (fn thy => add_declaration (intern_locale thy cliqN) recursive_declaration thy) end; fun add_spec_locales (name, _, specs) thy = let val name' = unsuffix proc_deco name; val ps = (suffix signatureN name' :: the_list locname); val ps' = hoare_ctxtL :: ps ; val pE = mk_loc_exp (map (intern_locale thy) ps) val pE' = mk_loc_exp (map (intern_locale thy) ps') fun read thy = apply_in_context thy (mk_loc_exp [intern_locale thy (suffix cliqueN (cname name))]) (Syntax.read_prop); fun proc_declaration phi = (*parameter_info_decl phi o already in signature *) add_active_procs phi (the (my_clique name)); fun add_locale'' (thm_name,spec) thy = let val spec' = read thy spec; val elem = Element.Assumes [((Binding.name thm_name, []), [(spec', [])])]; in thy |> add_locale thm_name pE' [elem] |> Proof_Context.theory_of |> (fn thy => add_declaration (intern_locale thy thm_name) proc_declaration thy) |> error_to_warning ("abbreviation: '" ^ thm_name ^ "' not added") (add_abbrev (intern_locale thy (suffix cliqueN (cname name))) Syntax.mode_input thm_name spec) end; in thy |> fold add_locale'' specs end; in thy |> fold (add_variable_statespaces o clique_vars) cliques |> fold (fn c => fold (fn n => add_signature_locale (lname "" c, n)) c) cliques |> fold add_body_locale name_pars |> fold add_impl_locales cliques |> fold add_spec_locales name_pars_specs end; (********************* theory extender interface ********************************) (** package setup **) (* outer syntax *) val var_declP = Parse.name -- (@{keyword "::"} |-- Parse.!!! Parse.embedded); val var_declP' = Parse.name >> (fn n => (n,"")); val localsP = Scan.repeat var_declP; val argP = var_declP; val argP' = var_declP'; val not_eqP = Scan.ahead (Scan.unless @{keyword "="} (Scan.one (K true))) val proc_decl_statespace = (Parse.short_ident --| @{keyword "("}) -- ((Parse.list argP) -- (Scan.optional (@{keyword "|"} |-- Parse.list argP) []) --| @{keyword ")"}) --| not_eqP val proc_decl_record = (Parse.short_ident --| @{keyword "("}) -- ((Parse.list argP') -- (Scan.optional (@{keyword "|"} |-- Parse.list argP') []) --| @{keyword ")"}) --| Scan.option @{keyword "="} val proc_decl = proc_decl_statespace >> pair Function || proc_decl_record >> pair Record; val loc_decl = Scan.optional (@{keyword "where"} |-- localsP --| @{keyword "in"}) [] val proc_body = Parse.embedded (*>> BodyTerm*) fun proc_specs x = (Parse.!!! (Scan.repeat (Parse_Spec.thm_name ":" -- Parse.embedded)) >> map (fn ((thm_name, _), prop) => (Binding.name_of thm_name, prop))) x val par_loc = Scan.option (@{keyword "("} |-- @{keyword "imports"} |-- Parse.name --| @{keyword ")"}); val _ = Outer_Syntax.command @{command_keyword procedures} "define procedures" (par_loc -- (Parse.and_list1 (proc_decl -- loc_decl -- proc_body -- proc_specs)) >> (fn (loc,decls) => let val decls' = map (fn ((((state_kind,(name,(ins,outs))),ls),bdy),specs) => (name,ins,outs,ls,bdy,specs,state_kind)) decls in Toplevel.theory (procedures_definition loc decls') end)); val _ = Outer_Syntax.command @{command_keyword hoarestate} "define state space for hoare logic" (StateSpace.statespace_decl >> (fn ((args,name),(parents,comps)) => Toplevel.theory (StateSpace.define_statespace args name parents (map (apfst (suffix deco)) comps)))); (*************************** Auxiliary Functions for integration of ********************) (*************************** automatic program analysers ********************) fun dest_conjs t = (case HOLogic.dest_conj t of [t1,t2] => dest_conjs t1 @ dest_conjs t2 | ts => ts); fun split_guard (Const (@{const_name Collect},CT)$(Abs (s,T,t))) = let fun mkCollect t = Const (@{const_name Collect},CT)$(Abs (s,T,t)); in map mkCollect (dest_conjs t) end | split_guard t = [t]; fun split_guards gs = let fun norm c f g = map (fn g => c$f$g) (split_guard g); fun norm_guard ((c as Const (@{const_name Pair},_))$f$g) = norm c f g | norm_guard ((c as Const (@{const_name Language.guaranteeStripPair},_))$f$g) = norm c f g | norm_guard t = [t]; in maps norm_guard (HOLogic.dest_list gs) end fun fold_com f t = let (* traverse does not descend into abstractions, like in DynCom, call, etc. *) fun traverse cnt (c as Const (@{const_name Language.com.Skip},_)) = (cnt,f cnt c [] []) | traverse cnt ((c as Const (@{const_name Language.com.Basic},_))$g) = (cnt, f cnt c [g] []) | traverse cnt ((c as Const (@{const_name Language.com.Spec},_))$r) = (cnt, f cnt c [r] []) | traverse cnt ((c as Const (@{const_name Language.com.Seq},_))$c1$c2) = let val (cnt1,v1) = traverse cnt c1; val (cnt2,v2) = traverse cnt1 c2; in (cnt2, f cnt c [] [v1,v2]) end | traverse cnt ((c as Const (@{const_name Language.com.Cond},_))$b$c1$c2) = let val (cnt1,v1) = traverse cnt c1; val (cnt2,v2) = traverse cnt1 c2; in (cnt2, f cnt c [b] [v1,v2]) end | traverse cnt ((c as Const (@{const_name Language.com.While},_))$b$c1) = let val (cnt1,v1) = traverse cnt c1 in (cnt1, f cnt c [b] [v1]) end | traverse cnt ((c as Const (@{const_name Language.com.Call},_))$p) = (cnt, f cnt c [p] []) | traverse cnt ((c as Const (@{const_name Language.com.DynCom},_))$c1) = (cnt, f cnt c [c1] []) | traverse cnt ((c as Const (@{const_name Language.com.Guard},_))$flt$g$c1) = let val (cnt1,v1) = traverse (cnt + length (split_guard g)) c1 in (cnt1, f cnt c [flt,g] [v1]) end | traverse cnt (c as Const (@{const_name Language.com.Throw},_)) = (cnt,f cnt c [] []) | traverse cnt ((c as Const (@{const_name Language.com.Catch},_))$c1$c2) = let val (cnt1,v1) = traverse cnt c1; val (cnt2,v2) = traverse cnt1 c2; in (cnt2, f cnt c [] [v1,v2]) end | traverse cnt ((c as Const (@{const_name Language.guards},_))$gs$c1) = let val (cnt1,v1) = traverse (cnt + length (split_guards gs)) c1; in (cnt1, f cnt c [gs] [v1]) end | traverse cnt ((c as Const (@{const_name Language.block},_))$init$c1$return$c2) = let val (cnt1,v1) = traverse cnt c1 in (cnt1, f cnt c [init,return,c2] [v1]) end | traverse cnt ((c as Const (@{const_name Language.call},_))$init$p$return$c1) = (cnt, f cnt c [init,p,return,c1] []) | traverse cnt ((c as Const (@{const_name Language.whileAnno},_))$b$I$V$c1) = let val (cnt1,v1) = traverse cnt c1 in (cnt1, f cnt c [b,I,V] [v1]) end | traverse cnt ((c as Const (@{const_name Language.whileAnnoG},_))$gs$b$I$V$c1) = let val (cnt1,v1) = traverse (cnt + length (split_guards gs)) c1 in (cnt1, f cnt c [gs,b,I,V] [v1]) end | traverse _ t = raise TERM ("fold_com: unknown command",[t]); in snd (traverse 0 t) end; (*************************** Tactics ****************************************************) (*** Aux. tactics ***) fun cond_rename_bvars cond name thm = let fun rename (tm as (Abs (x, T, t))) = if cond tm then Abs (name x, T, rename t) else Abs (x, T, rename t) | rename (t $ u) = rename t $ rename u | rename a = a; in Thm.renamed_prop (rename (Thm.prop_of thm)) thm end; val rename_bvars = cond_rename_bvars (K true); fun trace_tac ctxt str st = (if Config.get ctxt hoare_trace then tracing str else (); all_tac st); fun error_tac str st = (error str;no_tac st); fun rename_goal ctxt name = EVERY' [K (trace_tac ctxt "rename_goal -- START"), SELECT_GOAL (PRIMITIVE (rename_bvars name)), K (trace_tac ctxt "rename_goal -- STOP")]; (* splits applications of tupled arguments to a schematic Variables, e.g. * ALL a b. ?P (a,b) --> ?Q (a,b) gets * ALL a b. ?P a b --> ?Q a b * only tuples nested to the right are splitted. *) fun split_pair_apps ctxt thm = let val t = Thm.prop_of thm; fun mk_subst subst (Abs (x,T,t)) = mk_subst subst t | mk_subst subst (t as (t1$t2)) = (case strip_comb t of (var as Var (v,vT),args) => (if not (AList.defined (op =) subst var) then let val len = length args; val (argTs,bdyT) = strip_type vT; val (z, _) = Name.variant "z" (fold Term.declare_term_frees args Name.context); val frees = map (apfst (fn i => z^string_of_int i)) (0 upto (len - 1) ~~ argTs); fun splitT (Type (@{type_name Product_Type.prod}, [T1, T2])) = T1::splitT T2 | splitT T = [T]; fun pair_depth (Const (@{const_name Pair},aT)$t1$t2) = 1 + (pair_depth t2) | pair_depth _ = 0; fun mk_sel max free i = let val snds = funpow i HOLogic.mk_snd (Free free) in if i=max then snds else HOLogic.mk_fst snds end; fun split (free,arg) = let val depth = (pair_depth arg); in if depth = 0 then [Free free] else map (mk_sel depth free) (0 upto depth) end; val args' = maps split (frees ~~ args); val argTs' = maps splitT argTs; val inst = fold_rev absfree frees (list_comb (Var (v,argTs' ---> bdyT), args')) in subst@[(var,inst)] end else subst) | _ => mk_subst (mk_subst subst t1) t2) | mk_subst subst t = subst; val subst = map (fn (v,t) => (dest_Var v, Thm.cterm_of ctxt t)) (mk_subst [] t); in full_simplify (put_simpset HOL_basic_ss ctxt addsimps [@{thm fst_conv}, @{thm snd_conv}]) - (Drule.instantiate_normalize ([],subst) thm) + (Drule.instantiate_normalize (TVars.empty, Vars.make subst) thm) end; (* Generates split theorems, for !!,!,? quantifiers and for UN, e.g. * ALL x. P x = ALL a b. P a b *) fun mk_split_thms ctxt (vars as _::_) = let val thy = Proof_Context.theory_of ctxt; val names = map fst vars; val types = map snd vars; val free_vars = map Free vars; val pT = foldr1 HOLogic.mk_prodT types; val x = (singleton (Name.variant_list names) "x", pT); val xp = foldr1 HOLogic.mk_prod free_vars; val tfree_names = fold Term.add_tfree_names free_vars []; val zeta = TFree (singleton (Name.variant_list tfree_names) "z", Sign.defaultS thy); val split_meta_prop = let val P = Free (singleton (Name.variant_list names) "P", pT --> Term.propT) in Logic.mk_equals (Logic.all (Free x) (P $ Free x), fold_rev Logic.all free_vars (P $ xp)) end; val P = Free (singleton (Name.variant_list names) "P", pT --> HOLogic.boolT); val split_object_prop = let fun ALL vs t = Library.foldr (fn ((v,T),t) => HOLogic.mk_all (v,T,t)) (vs,t) in (ALL [x] (P $ Free x)) === (ALL vars (P $ xp)) end; val split_ex_prop = let fun EX vs t = Library.foldr (fn ((v,T),t) => HOLogic.mk_exists (v,T,t)) (vs,t) in (EX [x] (P $ Free x)) === (EX vars (P $ xp)) end; val split_UN_prop = let val P = Free (singleton (Name.variant_list names) "P", pT --> HOLogic.mk_setT zeta); fun UN vs t = Library.foldr mk_UN (vs, t) in (UN [x] (P $ Free x)) === (UN vars (P $ xp)) end; fun prove_simp simps prop = let val ([prop'], _) = Variable.importT_terms [prop] ctxt (* FIXME continue context!? *) in Goal.prove_global thy [] [] prop' (fn {context = ctxt, ...} => ALLGOALS (simp_tac (put_simpset HOL_basic_ss ctxt addsimps simps))) end; val split_meta = prove_simp [@{thm split_paired_all}] split_meta_prop; val split_object = prove_simp [@{thm split_paired_All}] split_object_prop; val split_ex = prove_simp [@{thm split_paired_Ex}] split_ex_prop; val split_UN = prove_simp [@{thm Hoare.split_paired_UN}] split_UN_prop; in [split_meta,split_object,split_ex,split_UN] end | mk_split_thms _ _ = raise Match; fun rename_aux_var name rule = let fun is_aux_var (Abs ("Z",TVar(_,_),_)) = true | is_aux_var _ = false; in cond_rename_bvars is_aux_var (K name) rule end; (* adapts single auxiliary variable in a rule to potentialy multiple auxiliary * variables in actual specification, e.g. if vars are a b, * split_app=false: ALL Z. ?P Z gets to ALL a b. ?P (a,b) * split_app=true: ALL Z. ?P Z gets to ALL a b. ?P a b * If only one auxiliary variable is given, the variables are just renamed, * If no auxiliary is given, unit is inserted for Z: * ALL Z. ?P Z gets P () *) fun adapt_aux_var ctxt split_app (vars as (_::_::_)) tvar_rules = let val thy = Proof_Context.theory_of ctxt; val max_idx = fold Integer.max (map (Thm.maxidx_of o snd) tvar_rules) 0; val types = map (fn i => TVar (("z",i),Sign.defaultS thy)) (max_idx + 1 upto (max_idx + length vars)); fun tvar n = (n, Sign.defaultS thy); val pT = Thm.ctyp_of ctxt (foldr1 HOLogic.mk_prodT types); - val rules' = map (fn (z,r) => (Drule.instantiate_normalize ([(tvar z,pT)],[]) r)) tvar_rules; + val rules' = map (fn (z,r) => (Drule.instantiate_normalize (TVars.make [(tvar z,pT)], Vars.empty) r)) tvar_rules; val splits = mk_split_thms ctxt (vars ~~ types); val rules'' = map (full_simplify (put_simpset HOL_basic_ss ctxt addsimps splits)) rules'; in if split_app then (map (split_pair_apps ctxt) rules'') else rules'' end | adapt_aux_var _ _ ([name]) tvar_rules = map (rename_aux_var name o snd) tvar_rules | adapt_aux_var ctxt _ ([]) tvar_rules = let val thy = Proof_Context.theory_of ctxt; fun tvar n = (n, Sign.defaultS thy); val uT = Thm.ctyp_of ctxt HOLogic.unitT; - val rules' = map (fn (z,r) => (Drule.instantiate_normalize ([(tvar z,uT)],[]) r)) tvar_rules; + val rules' = map (fn (z,r) => (Drule.instantiate_normalize (TVars.make [(tvar z,uT)], Vars.empty) r)) tvar_rules; val splits = [@{thm Hoare.unit_meta},@{thm Hoare.unit_object},@{thm Hoare.unit_ex},@{thm Hoare.unit_UN}]; val rules'' = map (full_simplify (put_simpset HOL_basic_ss ctxt addsimps splits)) rules'; in rules'' end (* Generates a rule for recursion for n procedures out of general recursion rule *) fun gen_call_rec_rule ctxt specs_name n rule = let val thy = Proof_Context.theory_of ctxt; val maxidx = Thm.maxidx_of rule; val vars = Term.add_vars (Thm.prop_of rule) []; fun get_type n = the (AList.lookup (op =) vars (n, 0)); val (Type (_, [Type (_, [assT, Type (_, [pT,_])])])) = get_type specs_name; val zT = TVar (("z",maxidx+1),Sign.defaultS thy) fun mk_var i n T = Var ((n ^ string_of_int i,0),T); val quadT = HOLogic.mk_prodT (assT, HOLogic.mk_prodT (pT, HOLogic.mk_prodT (assT,assT))); val quadT_set = HOLogic.mk_setT quadT; fun mk_spec i = let val quadruple = HOLogic.mk_ptuple (HOLogic.flat_tupleT_paths quadT) quadT [mk_var i "P" (zT --> assT)$Bound 0, mk_var i "p" pT, mk_var i "Q" (zT --> assT)$Bound 0, mk_var i "A" (zT --> assT)$Bound 0]; val single = HOLogic.mk_set quadT [quadruple]; in mk_UN' zT quadT_set (Abs ("Z", zT, single)) end; val Specs = foldr1 (mk_Un quadT_set) (map mk_spec (1 upto n)); val rule' = Thm.instantiate' [] [SOME (Thm.cterm_of ctxt Specs)] rule |> full_simplify (put_simpset (simpset_of @{theory_context Main}) ctxt addsimps [@{thm Hoare.conjE_simp},@{thm Hoare.in_Specs_simp},@{thm Hoare.in_set_Un_simp},@{thm split_all_conj}, @{thm image_Un},@{thm image_Un_single_simp}] ) |> rename_bvars (fn s => if member (op =) ["s","\"] s then s else "Z") in rule' end; fun gen_proc_rec ctxt mode n = gen_call_rec_rule ctxt "Specs" n (ProcRecSpecs mode); (*** verification condition generator ***) (* simplifications on "Collect" sets, like {s. P s} Int {s. Q s} = {s. P s & Q s} *) fun merge_assertion_simp_tac ctxt thms = simp_tac (put_simpset HOL_basic_ss ctxt addsimps ([@{thm Hoare.CollectInt_iff},@{thm HOL.conj_assoc},@{thm Hoare.Compl_Collect},singleton_conv_sym, @{thm Set.Int_empty_right},@{thm Set.Int_empty_left},@{thm Un_empty_right},@{thm Un_empty_left}]@thms)) ; (* The following probably shouldn't live here, but refactoring so that Hoare could depend on recursive_records does not look feasible. The upshot is that there's a duplicate foldcong_ss set here. *) structure FoldCongData = Theory_Data ( type T = simpset; val empty = HOL_basic_ss; val copy = I; val extend = I; val merge = merge_ss; ) val get_foldcong_ss = FoldCongData.get fun add_foldcongs congs thy = FoldCongData.map (fn ss => Proof_Context.init_global thy |> put_simpset ss |> fold Simplifier.add_cong congs |> simpset_of) thy fun add_foldcongsimps simps thy = FoldCongData.map (fn ss => Proof_Context.init_global thy |> put_simpset ss |> (fn ctxt => ctxt addsimps simps) |> simpset_of) thy (* propagates state into "Collect" sets and simplifies selections updates like: * s:{s. P s} = P s *) fun in_assertion_simp_tac ctxt state_kind thms i = let val vcg_simps = #vcg_simps (get_data ctxt); val fold_simps = get_foldcong_ss (Proof_Context.theory_of ctxt) in EVERY [simp_tac (put_simpset HOL_basic_ss ctxt addsimps ([mem_Collect_eq,@{thm Set.Un_iff},@{thm Set.Int_iff}, @{thm Set.empty_subsetI}, @{thm Set.empty_iff}, UNIV_I, @{thm Hoare.Collect_False}]@thms@K_convs@vcg_simps) addsimprocs (state_simprocs state_kind) |> fold Simplifier.add_cong K_congs) i THEN_MAYBE (simp_tac (put_simpset fold_simps ctxt addsimprocs [state_upd_simproc state_kind]) i) ] end; fun assertion_simp_tac ctxt state_kind thms i = merge_assertion_simp_tac ctxt [] i THEN_MAYBE in_assertion_simp_tac ctxt state_kind thms i (* simplify equality test on strings (and datatype-constructors) and propagate result*) fun string_eq_simp_tac ctxt = simp_tac (put_simpset HOL_basic_ss ctxt addsimps @{thms list.inject list.distinct char.inject cong_exp_iff_simps simp_thms}); fun assertion_string_eq_simp_tac ctxt state_kind thms i = assertion_simp_tac ctxt state_kind thms i THEN_MAYBE string_eq_simp_tac ctxt i; fun before_set2pred_simp_tac ctxt = (simp_tac (put_simpset HOL_basic_ss ctxt addsimps [singleton_conv_sym, @{thm Hoare.CollectInt_iff}, @{thm Hoare.Compl_Collect}])); (*****************************************************************************) (** set2pred transforms sets inclusion into predicates implication, **) (** maintaining the original variable names. **) (** Ex. "{x. x=0} <= {x. x <= 1}" -set2pred-> "x=0 --> x <= 1" **) (** Subgoals containing intersections (A Int B) or complement sets (-A) **) (** are first simplified by "before_set2pred_simp_tac", that returns only **) (** subgoals of the form "{x. P x} <= {x. Q x}", which are easily **) (** transformed. **) (** This transformation may solve very easy subgoals due to a ligth **) (** simplification done by full_simp_tac **) (*****************************************************************************) fun set2pred_tac ctxt i thm = ((before_set2pred_simp_tac ctxt i) THEN_MAYBE (EVERY [trace_tac ctxt "set2pred", resolve_tac ctxt [subsetI] i, resolve_tac ctxt [CollectI] i, dresolve_tac ctxt [CollectD] i, full_simp_tac (put_simpset HOL_basic_ss ctxt) i ])) thm (*****************************************************************************) (** BasicSimpTac is called to simplify all verification conditions. It does **) (** a light simplification by applying "mem_Collect_eq" **) (** then it tries to solve subgoals of the form "A <= A" and then if **) (** set2pred is true it **) (** transforms any other into predicates, applying then **) (** the tactic chosen by the user, which may solve the subgoal completely **) (** (MaxSimpTac). **) (*****************************************************************************) fun MaxSimpTac ctxt tac i = TRY (FIRST[resolve_tac ctxt [subset_refl] i, set2pred_tac ctxt i THEN_MAYBE tac i, trace_tac ctxt "final_tac failed" ]); fun BasicSimpTac ctxt state_kind set2pred thms tac i = EVERY [(trace_tac ctxt "BasicSimpTac -- START --"), assertion_simp_tac ctxt state_kind thms i THEN_MAYBE (if set2pred then MaxSimpTac ctxt tac i else TRY (resolve_tac ctxt [subset_refl] i)), (trace_tac ctxt "BasicSimpTac -- STOP --")]; (* EVERY [(trace_tac ctxt "BasicSimpTac -- START --"), simp_tac (HOL_basic_ss addsimps [mem_Collect_eq,@{thm Hoare.CollectInt_iff}, @{thm Set.empty_subsetI}, @{thm Set.empty_iff}, UNIV_I] addsimprocs [state_simproc sk]) i THEN_MAYBE simp_tac (HOL_basic_ss addsimprocs [state_upd_simproc sk]) i THEN_MAYBE (if set2pred then MaxSimpTac ctxt tac i else TRY (rtac subset_refl i)), (trace_tac ctxt "BasicSimpTac -- STOP --")]; *) (* fun simp_state_eq_tac Record state_space = full_simp_tac (HOL_basic_ss addsimprocs (state_simprocs Record)) THEN_MAYBE' full_simp_tac (HOL_basic_ss addsimprocs [state_upd_simproc Record]) THEN_MAYBE' (state_split_simp_tac [] state_space) | simp_state_eq_tac StateFun state_space = *) fun post_conforms_tac ctxt state_kind i = EVERY [REPEAT1 (resolve_tac ctxt [allI,impI] i), ((fn i => TRY (resolve_tac ctxt [conjI] i)) THEN_ALL_NEW (fn i => (REPEAT (resolve_tac ctxt [allI,impI] i)) THEN (full_simp_tac (put_simpset HOL_basic_ss ctxt addsimps [mem_Collect_eq,@{thm Set.singleton_iff},@{thm Set.empty_iff},UNIV_I] addsimprocs (state_simprocs state_kind)) i))) i]; fun dest_hoare_raw (Const(@{const_name HoarePartialDef.hoarep},_)$G$T$F$P$C$Q$A) = (P,C,Q,A,Partial,G,T,F) | dest_hoare_raw (Const(@{const_name HoareTotalDef.hoaret},_)$G$T$F$P$C$Q$A) = (P,C,Q,A,Total,G,T,F) | dest_hoare_raw t = raise TERM ("Hoare.dest_hoare_raw: unexpected term",[t]) fun mk_hoare_abs Ts (P,C,Q,A,mode,G,T,F) = let val hoareT = map (curry fastype_of1 Ts) [G,T,F,P,C,Q,A] ---> HOLogic.boolT; val hoareC = (case mode of Partial => Const (@{const_name HoarePartialDef.hoarep},hoareT) | Total => Const (@{const_name HoareTotalDef.hoaret},hoareT)); in hoareC$G$T$F$P$C$Q$A end; val is_hoare = can dest_hoare_raw fun dest_hoare t = let val triple = (strip_qnt_body @{const_name "All"} o HOLogic.dest_Trueprop o strip_qnt_body @{const_name Pure.all}) t; in dest_hoare_raw triple end; fun get_aux_tvar rule = let fun aux_hoare (Abs ("Z",TVar (z,_),t)) = if is_hoare (strip_qnt_body @{const_name All} t) then SOME z else NONE | aux_hoare _ = NONE; in (case first_subterm_dest (aux_hoare) (Thm.prop_of rule) of SOME (_,z) => (z,rule) | NONE => raise TERM ("get_aux_tvar: No auxiliary variable of hoare-rule found", [Thm.prop_of rule])) end; fun strip_vars t = let val bdy = (HOLogic.dest_Trueprop o Logic.strip_assums_concl) t; in strip_qnt_vars @{const_name Pure.all} t @ strip_qnt_vars @{const_name All} bdy end; local (* ex_simps are necessary in case of multiple logical variables. The state will usually be the first variable. EX s a b. s=s' ... . We have to transport EX s to s=s' to perform the substitution *) val conseq1_ss_base = simpset_of (put_simpset HOL_basic_ss @{context} addsimps ([mem_Collect_eq,@{thm Set.singleton_iff},@{thm Set.Int_iff}, @{thm Set.empty_iff},UNIV_I, @{thm HOL.conj_assoc}, @{thm disj_assoc}] @ @{thms Hoare.all_imp_eq_triv} @K_convs @ @{thms simp_thms} @ @{thms ex_simps} @ @{thms all_simps}) delsimps [@{thm Hoare.all_imp_to_ex}] |> fold Simplifier.add_cong K_congs) val conseq1_ss_record = simpset_of (put_simpset conseq1_ss_base @{context} addsimprocs (state_simprocs Record)); val conseq1_ss_fun = simpset_of (put_simpset conseq1_ss_base @{context} addsimprocs (state_simprocs Function)); fun conseq1_ss Record = conseq1_ss_record | conseq1_ss Function = conseq1_ss_fun; val conseq2_ss_base = simpset_of (put_simpset HOL_basic_ss @{context} addsimps (@{thms Hoare.all_imp_eq_triv} @ @{thms simp_thms} @ @{thms ex_simps} @ @{thms all_simps}) delsimps [@{thm Hoare.all_imp_to_ex}] |> Simplifier.add_cong @{thm imp_cong}); val conseq2_ss_record = simpset_of (put_simpset conseq2_ss_base @{context} addsimprocs [state_upd_simproc Record, state_ex_sel_eq_simproc Record]); val conseq2_ss_fun = simpset_of (put_simpset conseq2_ss_base @{context} addsimprocs [state_upd_simproc Function, state_ex_sel_eq_simproc Function]); fun conseq2_ss Record = conseq2_ss_record | conseq2_ss Function = conseq2_ss_fun; in fun raw_conseq_simp_tac ctxt state_kind thms i = let val ctxt' = Config.put simp_depth_limit 0 ctxt; in simp_tac (put_simpset (conseq1_ss state_kind) ctxt' addsimps thms) i THEN_MAYBE simp_tac (put_simpset (conseq2_ss state_kind) ctxt') i end end val conseq_simp_tac = raw_conseq_simp_tac; (* Generates the hoare-quadruples that can be derived out of the hoare-context T *) fun gen_context_thms ctxt mode params G T F = let val Type (_,[comT]) = range_type (fastype_of G); fun destQuadruple (Const (@{const_name Set.insert},_) $ PpQA $ Const (@{const_name Orderings.bot}, _)) = PpQA | destQuadruple t = raise Match; fun mkCallQuadruple (Const (@{const_name Pair}, _) $ P $ (Const (@{const_name Pair}, _) $ p $ (Const(@{const_name Pair}, _) $ Q $ A))) = let val Call_p = Const (@{const_name Language.com.Call}, fastype_of p --> comT) $ p; in (P, Call_p, Q, A) end; fun mkHoare mode G T F (vars,PpQA) = let val hoare = (case mode of Partial => @{const_name HoarePartialDef.hoarep} | Total => @{const_name HoareTotalDef.hoaret}); (* FIXME: Use future Proof_Context.rename_vars or make closed term and remove by hand *) (* fun free_params ps t = foldr (fn ((x,xT),t) => snd (variant_abs (x,xT,t))) (ps,t); val PpQA' = mkCallQuadruple (strip_qnt_body @{const_name Pure.all} (free_params params (Term.list_all (vars,PpQA)))); *) val params' = (Variable.variant_frees ctxt [PpQA] params); val bnds = map Bound (0 upto (length vars - 1)); fun free_params_vars t = subst_bounds (bnds @ rev (map Free params' ), t) fun free_params t = subst_bounds (rev (map Free params' ), t) val (P',p',Q',A') = mkCallQuadruple (free_params_vars PpQA); val T' = free_params T; val G' = free_params G; val F' = free_params F; val bdy = mk_hoare_abs [] (P',p',Q',A',mode,G',T',F'); in (HOLogic.mk_Trueprop (HOLogic.list_all (vars,bdy)), map fst params') end; fun hoare_context_specs mode G T F = let fun mk t = try (mkHoare mode G T F o apsnd destQuadruple o dest_UN) t; in map_filter mk (dest_Un T) end; fun mk_prove mode (prop,params) = let val vars = map fst (strip_qnt_vars @{const_name All} (HOLogic.dest_Trueprop (Logic.strip_assums_concl prop))); val [asmUN'] = adapt_aux_var ctxt true vars [get_aux_tvar (AsmUN mode)] in Goal.prove ctxt params [] prop (fn thms => EVERY[trace_tac ctxt "extracting specifications from hoare context", resolve_tac ctxt [asmUN'] 1, DEPTH_SOLVE_1 (resolve_tac ctxt [subset_refl,refl] 1 ORELSE ((resolve_tac ctxt [@{thm Hoare.subset_unI1}] 1 APPEND resolve_tac ctxt [@{thm Hoare.subset_unI2}] 1) ORELSE (resolve_tac ctxt [@{thm Hoare.subset_singleton_insert1}] 1 APPEND resolve_tac ctxt [@{thm Hoare.subset_singleton_insert2}] 1))) ORELSE error_tac ("vcg: cannot extract specifications from context") ]) end; val specs = hoare_context_specs mode G T F; in map (mk_prove mode) specs end; fun is_modifies_clause t = exists_subterm (fn (Const (@{const_name Hoare.meq},_)) => true| _ => false) (#3 (dest_hoare (Logic.strip_assums_concl t))) handle (TERM _) => false; val is_spec_clause = not o is_modifies_clause; (* e.g: Intg => the_Intg lift Intg => lift the_Intg map Ingt => map the_Intg Hp o lift Intg => lift the_Intg o the_Hp *) fun swap_constr_destr f (t as (Const (@{const_name Fun.id},_))) = t | swap_constr_destr f (t as (Const (c,Type ("fun",[T,valT])))) = (Const (f c, Type ("fun",[valT,T])) handle Empty => raise TERM ("Hoare.swap_constr_destr",[t])) | swap_constr_destr f (Const ("StateFun.map_fun",Type ("fun", (* FIXME unknown "StateFun.map_fun" !? *) [Type ("fun",[T,valT]), Type ("fun",[Type ("fun",[xT,T']), Type ("fun",[xT',valT'])])]))$g) = Const ("StateFun.map_fun",Type("fun",[Type ("fun",[valT,T]), Type ("fun",[Type ("fun",[xT,valT']), Type ("fun",[xT',T'])])]))$ swap_constr_destr f g | swap_constr_destr f (Const (@{const_name Fun.comp},Type ("fun", [Type ("fun",[bT',cT]), Type ("fun",[Type ("fun",[aT ,bT]), Type ("fun",[aT',cT'])])]))$h$g) = let val h'=swap_constr_destr f h; val g'=swap_constr_destr f g; in Const (@{const_name Fun.comp},Type ("fun", [Type ("fun",[bT,aT]), Type ("fun",[Type ("fun",[cT,bT']), Type ("fun",[cT',aT'])])]))$g'$h' end | swap_constr_destr f (Const (@{const_name List.map},Type ("fun", [Type ("fun",[aT,bT]), Type ("fun",[asT,bsT])]))$g) = (Const (@{const_name List.map},Type ("fun", [Type ("fun",[bT,aT]), Type ("fun",[bsT,asT])]))$swap_constr_destr f g) | swap_constr_destr f t = raise TERM ("Hoare.swap_constr_destr",[t]); (* FIXME: unused? *) val destr_to_constr = let fun convert c = let val (path,base) = split_last (Long_Name.explode c); in Long_Name.implode (path @ ["val",unprefix "the_" base]) end; in swap_constr_destr convert end; fun gen_call_tac cont_tac mode cmode state_kind state_space ctxt asms spec_sfx pname return has_args _ = let val thy = Proof_Context.theory_of ctxt; val pname' = unsuffix proc_deco pname; val spec = (case AList.lookup (op =) asms pname of SOME s => SOME s | NONE => try (Proof_Context.get_thm ctxt) (suffix spec_sfx pname')); fun auxvars_for p t = (case first_subterm_dest (try dest_call) t of SOME (vars,(_,p',_,_,m,_)) => (if m=Static andalso p=(dest_string' p') then SOME vars else NONE) | NONE => NONE); fun get_auxvars_for p t = (case (map_filter ((auxvars_for p) o snd) (max_subterms_dest tap_UN t)) of (vars::_) => vars | _ => []); fun spec_tac augment_rule augment_emptyFaults _ spec i = let val spec' = augment_emptyFaults OF [spec] handle THM _ => spec; in EVERY [resolve_tac ctxt [augment_rule] i, resolve_tac ctxt [spec'] (i+1), TRY (resolve_tac ctxt [subset_refl, @{thm Set.empty_subsetI}, @{thm Set.Un_upper1}, @{thm Set.Un_upper2}] i)] end; fun check_spec name P thm = (case try dest_hoare (Thm.concl_of thm) of SOME spc => (case try dest_call (#2 spc) of SOME (_,p,_,_,m,_) => if proc_name m p = name andalso P (Thm.concl_of thm) then SOME (#5 spc,thm) else NONE | _ => NONE) | _ => NONE) fun find_dyn_specs name P thms = map_filter (check_spec name P) thms; fun get_spec name P thms = case find_dyn_specs name P thms of (spec_mode,spec)::_ => SOME (spec_mode,spec) | _ => NONE; fun solve_spec augment_rule _ augment_emptyFaults mode _ (SOME spec_mode) (SOME spec) i= if mode=Partial andalso spec_mode=Total then resolve_tac ctxt [@{thm HoareTotal.hoaret_to_hoarep'}] i THEN spec_tac augment_rule augment_emptyFaults mode spec i else if mode=spec_mode then spec_tac augment_rule augment_emptyFaults mode spec i else error("vcg: cannot use a partial correctness specification of " ^ pname' ^ " for a total correctness proof") | solve_spec _ asmUN_rule _ mode Static _ _ i =(* try to infer spec out of context *) EVERY[trace_tac ctxt "inferring specification from hoare context1", resolve_tac ctxt [asmUN_rule] i, DEPTH_SOLVE_1 (resolve_tac ctxt [subset_refl,refl] i ORELSE ((resolve_tac ctxt [@{thm Hoare.subset_unI1}] i APPEND resolve_tac ctxt [@{thm Hoare.subset_unI2}] i) ORELSE (resolve_tac ctxt [@{thm Hoare.subset_singleton_insert1}] i APPEND resolve_tac ctxt [@{thm Hoare.subset_singleton_insert2}] i))) ORELSE error_tac ("vcg: cannot infer specification of " ^ pname' ^ " from context") (* if tactic for DEPTH_SOLVE_1 would create new subgoals, use SELECT_GOAL and DEPTH_SOLVE *) ] | solve_spec augment_rule asmUN_rule augment_emptyFaults mode Parameter _ _ i = (* try to infer spec out of assumptions *) let fun tac thms = (case (find_dyn_specs pname is_spec_clause thms) of (spec_mode,spec)::_ => solve_spec augment_rule asmUN_rule augment_emptyFaults mode Parameter (SOME spec_mode) (SOME spec) 1 | _ => all_tac) in Subgoal.FOCUS (tac o #prems) ctxt i end val strip_spec_vars = strip_qnt_vars @{const_name All} o HOLogic.dest_Trueprop; fun apply_call_tac ctxt pname mode cmode spec_mode spec_goal is_abr spec (subgoal,i) = let val spec_vars = map fst (case spec of SOME sp => (strip_spec_vars (Thm.concl_of sp)) | NONE => (case try (dest_hoare) subgoal of SOME (_,_,_,_,_,_,Theta,_) => get_auxvars_for pname Theta | _ => [])); fun get_call_rule Static mode is_abr = if is_abr then Proc mode else ProcNoAbr mode | get_call_rule Parameter mode is_abr = if is_abr then DynProcProcPar mode else DynProcProcParNoAbr mode; val [call_rule,augment_ctxt_rule,asmUN_rule, augment_emptyFaults] = adapt_aux_var ctxt true spec_vars (map get_aux_tvar [get_call_rule cmode mode is_abr, AugmentContext mode, AsmUN mode, AugmentEmptyFaults mode]); in EVERY [resolve_tac ctxt [call_rule] i, trace_tac ctxt "call_tac -- basic_tac -- solving spec", solve_spec augment_ctxt_rule asmUN_rule augment_emptyFaults mode cmode spec_mode spec spec_goal] end; fun basic_tac spec i = let val msg ="Theorem " ^pname'^spec_sfx ^ " is no proper specification for procedure " ^pname'^ "; trying to infer specification from hoare context"; fun spec' s mode abr = if is_modifies_clause (Thm.concl_of s) then if abr then (TrivPost mode) OF [s] else (TrivPostNoAbr mode) OF [s] else s; val (is_abr,spec_mode,spec,spec_has_args) = (* is there a proper specification fact? *) case spec of NONE => (true,NONE,NONE,false) | SOME s => case try dest_hoare (Thm.concl_of s) of NONE => (warning msg;(true,NONE,NONE,false)) | SOME (_,c,Q,spec_abr,spec_mode,_,_,_) => case try dest_call c of NONE => (warning msg;(true,NONE,NONE,false)) | SOME (_,p,_,_,m,spec_has_args) => if proc_name m p = pname then if (mode=Total andalso spec_mode=Partial) then (warning msg;(true,NONE,NONE,false)) else if is_empty_set spec_abr then (false,SOME spec_mode, SOME (spec' s spec_mode false),spec_has_args) else (true,SOME spec_mode, SOME (spec' s spec_mode true),spec_has_args) else (warning msg;(true,NONE,NONE,false)); val () = if spec_has_args then error "procedure call in specification must be parameterless!" else (); val spec_goal = i+2; in EVERY[trace_tac ctxt "call_tac -- basic_tac -- START --", SUBGOAL (apply_call_tac ctxt pname mode cmode spec_mode spec_goal is_abr spec) i, resolve_tac ctxt [allI] (i+1), resolve_tac ctxt [allI] (i+1), cont_tac (i+1), trace_tac ctxt "call_tac -- basic_tac -- simplify", conseq_simp_tac ctxt state_kind [@{thm StateSpace.upd_globals_def}] i, trace_tac ctxt "call_tac -- basic_tac -- STOP --"] end; fun get_modifies (Const (@{const_name Collect},_) $ Abs (_,_,m)) = m | get_modifies t = raise TERM ("gen_call_tac.get_modifies",[t]); fun get_updates (Abs (_,_,t)) = get_updates t | get_updates (Const (@{const_name Hoare.mex},_) $ t) = get_updates t | get_updates (Const (@{const_name Hoare.meq},T) $ _ $ upd) = (T,upd) | get_updates t = raise TERM ("gen_call_tac.get_updates",[t]); (* return has the form: %s t. s(|globals:=globals t,...|) * should be translated to %s t. s(|globals := globals s(|m := m (globals t),...|),...|) * for all m in the modifies list. *) fun mk_subst gT meqT = fst (Sign.typ_unify thy (gT,domain_type meqT) (Vartab.empty,0)); fun mk_selR subst gT (upd,uT) = let val vT = range_type (hd (binder_types uT)); in Const (unsuffix Record.updateN upd,gT --> (Envir.norm_type subst vT)) end; (* lookup:: "('v => 'a) => 'n => ('n => 'v) => 'a" * update:: "('v => 'a) => ('a => 'v) => 'n => ('a => 'a) => ('n => 'v) => ('n => 'v)" *) fun mk_selF subst uT d n = let val vT_a::a_vT::nT::aT_aT::n_vT::_ = binder_types uT; val lT = (Envir.norm_type subst (vT_a --> nT --> n_vT --> (domain_type aT_aT))); val d' = map_types (Envir.norm_type subst) d; in Const (@{const_name StateFun.lookup},lT)$d'$n end; fun mk_rupdR subst gT (upd,uT) = let val [vT,_] = binder_types uT in Const (upd,(Envir.norm_type subst vT) --> gT --> gT) end; fun K_fun kn uT = let val T=range_type (hd (binder_types uT)) in Const (kn,T --> T --> T) end; fun K_rec uT t = let val T=range_type (hd (binder_types uT)) in Abs ("_", T, incr_boundvars 1 t) end; fun mk_supdF subst uT d c n = let val uT' = Envir.norm_type subst uT; val c' = map_types (Envir.norm_type subst) c; val d' = map_types (Envir.norm_type subst) d; in Const (@{const_name StateFun.update},uT')$d'$c'$n end; fun modify_updatesR subst gT glob ((Const (upd,uT))$_$(Const _$Z)) = mk_rupdR subst gT (upd,uT)$ (K_rec uT (mk_selR subst gT (upd,uT)$(glob$Bound 0)))$(glob$Bound 1) | modify_updatesR subst gT glob ((Const (upd,uT))$_$s) = mk_rupdR subst gT (upd,uT)$ (K_rec uT (mk_selR subst gT (upd,uT)$(glob$Bound 0)))$ modify_updatesR subst gT glob s | modify_updatesR subst gT glob ((_$Z)) = (glob$Bound 1) (* may_not_modify *) | modify_updatesR _ _ _ t = raise TERM ("gen_call_tac.modify_updatesR",[t]); fun modify_updatesF subst _ glob (Const (@{const_name StateFun.update},uT)$d$c$n$_$(Const globs$Z)) = mk_supdF subst uT d c n$ (K_fun KNF uT$(mk_selF subst uT d n$(glob$Bound 0)))$(glob$Bound 1) | modify_updatesF subst gT glob (Const (@{const_name StateFun.update},uT)$d$c$n$_$s) = mk_supdF subst uT d c n$ (K_fun KNF uT$(mk_selF subst uT d n$(glob$Bound 0)))$modify_updatesF subst gT glob s | modify_updatesF subst _ glob ((globs$Z)) = (glob$Bound 1) (* may_not_modify *) | modify_updatesF _ _ _ t = raise TERM ("gen_call_tac.modify_updatesF",[t]); fun modify_updates Record = modify_updatesR | modify_updates _ = modify_updatesF fun globalsT (Const (gupd,T)) = domain_type (hd (binder_types T)) | globalsT t = raise TERM ("gen_call_tac.globalsT",[t]); fun mk_upd meqT mods (gupd$(Abs (dmy,dmyT,(glob$Bound 1)))$Bound 1) = let val gT = (globalsT gupd); val subst = mk_subst gT meqT; in (gupd$(Abs (dmy,dmyT,incr_boundvars 1 (modify_updates state_kind subst gT glob mods)))$Bound 1) end | mk_upd meqT mods (upd$v$s) = upd$v$mk_upd meqT mods s | mk_upd _ _ t = raise TERM ("gen_call_tac.mk_upd",[t]); fun modify_return (meqT,mods) (Abs (s,T,Abs (t,U,upd))) = (Abs (s,T,Abs (t,U,mk_upd meqT mods upd))) | modify_return _ t = raise TERM ("get_call_tac.modify_return",[t]); fun modify_tac spec modifies_thm i = try (fn () => let val (_,call,modif_spec_nrm,modif_spec_abr,modif_spec_mode,G,Theta,_) = dest_hoare (Thm.concl_of modifies_thm); val is_abr = not (is_empty_set modif_spec_abr); val emptyTheta = is_empty_set Theta; (*val emptyFaults = is_empty_set F;*) val spec_has_args = #6 (dest_call call); val () = if spec_has_args then error "procedure call in modifies-specification must be parameterless!" else (); val (mxprem,ModRet) = (case cmode of Static => (8,if is_abr then if emptyTheta then (ProcModifyReturn mode) else (ProcModifyReturnSameFaults mode) else if emptyTheta then (ProcModifyReturnNoAbr mode) else (ProcModifyReturnNoAbrSameFaults mode)) | Parameter => (9,if is_abr then if emptyTheta then (ProcProcParModifyReturn mode) else (ProcProcParModifyReturnSameFaults mode) else if emptyTheta then (ProcProcParModifyReturnNoAbr mode) else (ProcProcParModifyReturnNoAbrSameFaults mode))); val to_prove_prem = (case cmode of Static => 0 | Parameter => 1); val spec_goal = if is_abr then i + mxprem - 5 else i + mxprem - 6 val mods_nrm = modif_spec_nrm |> get_modifies |> get_updates; val return' = modify_return mods_nrm return; (* val return' = if is_abr then let val mods_abr = modif_spec_abr |> get_modifies |> get_updates; in modify_return mods_abr return end else return;*) val cret = Thm.cterm_of ctxt return'; val (_,_,return'_var,_,_,_) = nth (Thm.prems_of ModRet) to_prove_prem |> dest_hoare |> #2 |> dest_call; val ModRet' = infer_instantiate ctxt [(#1 (dest_Var return'_var), cret)] ModRet; val modifies_thm_partial = if modif_spec_mode = Total then @{thm HoareTotal.hoaret_to_hoarep'} OF [modifies_thm] else modifies_thm; fun solve_modifies_tac i = (clarsimp_tac ((ctxt |> put_claset (claset_of @{theory_context Set}) |> put_simpset (simpset_of @{theory_context Set})) addsimps ([@{thm Hoare.mex_def},@{thm Hoare.meq_def},@{thm StateSpace.upd_globals_def}]@K_convs) addsimprocs (state_upd_simproc Record::(state_simprocs state_kind)) |> fold Simplifier.add_cong K_congs) i) THEN_MAYBE EVERY [trace_tac ctxt "modify_tac: splitting record", state_split_simp_tac ctxt [] state_space i]; val cnt = i + mxprem; in EVERY[trace_tac ctxt "call_tac -- modifies_tac --", resolve_tac ctxt [ModRet'] i, solve_spec (AugmentContext Partial) (AsmUN Partial) (AugmentEmptyFaults Partial) Partial Static (SOME Partial) (SOME modifies_thm_partial) spec_goal, if is_abr then EVERY [trace_tac ctxt "call_tac -- Solving abrupt modifies --", solve_modifies_tac (cnt - 6)] else all_tac, trace_tac ctxt "call_tac -- Solving Modifies --", solve_modifies_tac (cnt - 7), basic_tac spec (cnt - 8), if cmode = Parameter then EVERY [resolve_tac ctxt [subset_refl] (cnt - 8), simp_tac (put_simpset HOL_basic_ss ctxt addsimps (@{thm Hoare.CollectInt_iff} :: @{thms simp_thms})) 1] else all_tac] end) () |> (fn SOME res => res | NONE => raise TERM ("get_call_tac.modify_tac: no proper modifies spec", [])); fun specs_of_assms_tac thms = (case get_spec pname is_spec_clause thms of SOME (_,spec) => (case get_spec pname is_modifies_clause thms of SOME (_,modifies_thm) => modify_tac (SOME spec) modifies_thm 1 | NONE => basic_tac (SOME spec) 1) | NONE => (warning ("no proper specification for procedure " ^pname^ " in assumptions"); all_tac)); val test_modify_in_ctxt_tac = let val mname = (suffix modifysfx pname'); val mspec = (case try (Proof_Context.get_thm ctxt) mname of SOME s => SOME s | NONE => (case AList.lookup (op =) asms pname of SOME s => if is_modifies_clause (Thm.concl_of s) then SOME s else NONE | NONE => NONE)); in (case mspec of NONE => basic_tac spec | SOME modifies_thm => (case check_spec pname is_modifies_clause modifies_thm of SOME _ => modify_tac spec modifies_thm | NONE => (warning ("ignoring theorem " ^ (suffix modifysfx pname') ^ "; no proper modifies specification for procedure "^pname'); basic_tac spec))) end; fun inline_bdy_tac has_args i = (case try (Proof_Context.get_thm ctxt) (suffix bodyP pname') of NONE => no_tac | SOME impl => (case try (Proof_Context.get_thm ctxt) (suffix (body_def_sfx^"_def") pname') of NONE => no_tac | SOME bdy => (tracing ("No specification found for procedure \"" ^ pname' ^ "\". Inlining procedure!"); if has_args then EVERY [trace_tac ctxt "inline_bdy_tac args", resolve_tac ctxt [CallBody mode] i, resolve_tac ctxt [impl] (i+3), resolve_tac ctxt [allI] (i+2), resolve_tac ctxt [allI] (i+2), in_assertion_simp_tac ctxt state_kind [] (i+2), cont_tac (i+2), resolve_tac ctxt [allI] (i+1),in_assertion_simp_tac ctxt state_kind [bdy] (i+1), cont_tac (i+1), in_assertion_simp_tac ctxt state_kind [@{thm StateSpace.upd_globals_def}] i] else EVERY [trace_tac ctxt "inline_bdy_tac no args", resolve_tac ctxt [ProcBody mode] i, resolve_tac ctxt [impl] (i+2), simp_tac (put_simpset HOL_basic_ss ctxt addsimps [bdy]) (i+1), cont_tac (i+1)]))); in (case cmode of Static => if get_recursive pname ctxt = SOME false andalso is_none spec then inline_bdy_tac has_args else test_modify_in_ctxt_tac | Parameter => (case spec of NONE => (tracing "no spec found!"; Subgoal.FOCUS (specs_of_assms_tac o #prems) ctxt) | SOME spec => (tracing "found spec!"; case check_spec pname is_spec_clause spec of SOME _ => test_modify_in_ctxt_tac | NONE => (warning ("ignoring theorem " ^ (suffix spec_sfx pname') ^ "; no proper specification for procedure " ^pname'); Subgoal.FOCUS (specs_of_assms_tac o #prems) ctxt)))) end; fun call_tac cont_tac mode state_kind state_space ctxt asms spec_sfx t = let val (_,c,_,_,_,_,_,F) = dest_hoare (Logic.strip_assums_concl t); fun gen_tac (_,pname,return,c,cmode,has_args) = gen_call_tac cont_tac mode cmode state_kind state_space ctxt asms spec_sfx (proc_name cmode pname) return has_args F; in gen_tac (dest_call c) end handle TERM _ => K no_tac; fun solve_in_Faults_tac ctxt i = resolve_tac ctxt [UNIV_I, @{thm in_insert_hd}] i ORELSE SELECT_GOAL (SOLVE (simp_tac (put_simpset (simpset_of @{theory_context Set}) ctxt) 1)) i; local fun triv_simp ctxt = merge_assertion_simp_tac ctxt [mem_Collect_eq] THEN' simp_tac (put_simpset HOL_basic_ss ctxt addsimps @{thms simp_thms} |> fold Simplifier.add_cong [@{thm conj_cong}, @{thm imp_cong}]); (* a guarded while produces stupid things, since the guards are put at the end of the body and in the invariant (rule WhileAnnoG): - guard: g /\ g - guarantee: g --> g *) in fun guard_tac ctxt strip cont_tac mode (t,i) = let val (_,c,_,_,_,_,_,F) = dest_hoare (Logic.strip_assums_concl t); val (_,_,_,doStrip) = dest_Guard c; val guarantees = if strip orelse doStrip then [GuardStrip mode, GuaranteeStrip mode] else [Guarantee mode] fun basic_tac i = EVERY [resolve_tac ctxt [Guard mode, GuaranteeAsGuard mode] i, trace_tac ctxt "Guard", cont_tac (i+1), triv_simp ctxt i] fun guarantee_tac i = EVERY [resolve_tac ctxt guarantees i, solve_in_Faults_tac ctxt (i+2), cont_tac (i+1), triv_simp ctxt i] in if is_empty_set F then EVERY [trace_tac ctxt "Guard: basic_tac", basic_tac i] else EVERY [trace_tac ctxt "Guard: trying guarantee_tac", guarantee_tac i ORELSE basic_tac i] end handle TERM _ => no_tac end; fun wf_tac ctxt = simp_tac (put_simpset HOL_basic_ss ctxt addsimps [@{thm Wellfounded.wf_measure},@{thm Wellfounded.wf_lex_prod},@{thm Wfrec.wf_same_fst}, @{thm Hoare.wf_measure_lex_prod},@{thm Wellfounded.wf_inv_image}]); fun in_rel_simp ctxt = simp_tac (put_simpset HOL_basic_ss ctxt addsimps [@{thm Hoare.in_measure_iff},@{thm Hoare.in_lex_iff},@{thm Hoare.in_mlex_iff},@{thm Hoare.in_inv_image_iff}, @{thm split_conv}]); fun while_annotate_tac ctxt inv i st = let val annotateWhile = Thm.lift_rule (Thm.cprem_of st i) @{thm HoarePartial.reannotateWhileNoGuard}; val lifted_inv = fold_rev Term.abs (Logic.strip_params (Logic.get_goal (Thm.prop_of st) i)) inv; val invVar = (#1 o strip_comb o #3 o dest_whileAnno o #2 o dest_hoare) (List.last (Thm.prems_of annotateWhile)); val annotate = infer_instantiate ctxt [(#1 (dest_Var invVar), Thm.cterm_of ctxt lifted_inv)] annotateWhile; in ((trace_tac ctxt ("annotating While with: " ^ Syntax.string_of_term ctxt lifted_inv )) THEN compose_tac ctxt (false,annotate,1) i) st end; fun cond_annotate_tac ctxt inv mode (_,i) st = let val annotateCond = Thm.lift_rule (Thm.cprem_of st i) (CondInv' mode); val lifted_inv = fold_rev Term.abs (Logic.strip_params (Logic.get_goal (Thm.prop_of st) i)) inv; val invVar = List.last (Thm.prems_of annotateCond) |> dest_hoare |> #3 |> strip_comb |> #1; val annotate = infer_instantiate ctxt [(#1 (dest_Var invVar), Thm.cterm_of ctxt lifted_inv)] annotateCond; in ((trace_tac ctxt ("annotating Cond with: "^ Syntax.string_of_term ctxt lifted_inv)) THEN compose_tac ctxt (false,annotate,5) i) st end; fun basic_while_tac ctxt state_kind cont_tac tac mode i = let fun common_tac i = EVERY[if mode=Total then wf_tac ctxt (i+3) else all_tac, BasicSimpTac ctxt state_kind true [] tac (i+2), if mode=Total then in_rel_simp ctxt (i+1) THEN (resolve_tac ctxt [allI] (i+1)) else all_tac, cont_tac (i+1) ] fun basic_tac i = EVERY [resolve_tac ctxt [While mode] i, common_tac i] in EVERY [trace_tac ctxt "basic_while_tac: basic_tac", basic_tac i] end; fun while_tac ctxt state_kind inv cont_tac tac mode t i= let val basic_tac = basic_while_tac ctxt state_kind cont_tac tac mode; in (case inv of NONE => basic_tac i | SOME I => EVERY [while_annotate_tac ctxt I i, basic_tac i]) end handle TERM _ => no_tac fun dest_split (Abs (x,T,t)) = let val (vs,recomb,bdy) = dest_split t; in ((x,T)::vs,fn t' => Abs (x,T,recomb t'),bdy) end | dest_split (c as Const (@{const_name case_prod},_)$Abs(x,T,t)) = let val (vs,recomb,bdy) = dest_split t; in ((x,T)::vs,fn t' => c$Abs (x,T,recomb t'),bdy) end | dest_split t = ([],I,t); fun whileAnnoG_tac ctxt strip_guards mode t i st = let val (_,c,_,_,_,_,_,F) = dest_hoare (Logic.strip_assums_concl t); val (SOME grds,_,I,_,_,fix) = dest_whileAnno c; val Rule = if fix then WhileAnnoGFix mode else WhileAnnoG mode; fun extract_faults (Const (@{const_name Set.insert}, _) $ t $ _) = [t] | extract_faults _ = []; fun leave_grd fs (Const (@{const_name Pair}, _) $ f $ g) = if member (op =) fs f andalso strip_guards then NONE else SOME g | leave_grd fs (Const (@{const_name Language.guaranteeStripPair}, _) $ f $ g) = if member (op =) fs f then NONE else SOME g | leave_grd fs _ = NONE; val (I_vs,I_recomb,I') = dest_split I; val grds' = map_filter (leave_grd (extract_faults F)) (HOLogic.dest_list grds); val pars = (Logic.strip_params (Logic.get_goal (Thm.prop_of st) i)); val J = fold_rev Term.abs pars (I_recomb (fold_rev (mk_Int (map snd (pars@I_vs))) grds' I')); val WhileG = Thm.lift_rule (Thm.cprem_of st i) Rule; val invVar = (fst o strip_comb o #3 o dest_whileAnno o (fn xs => nth xs 1) o snd o strip_comb o #2 o dest_hoare) (List.last (Thm.prems_of WhileG)); val WhileGInst = infer_instantiate ctxt [(#1 (dest_Var invVar), Thm.cterm_of ctxt J)] WhileG; in ((trace_tac ctxt ("WhileAnnoG, adding guards to invariant: " ^ Syntax.string_of_term ctxt J)) THEN compose_tac ctxt (false,WhileGInst,1) i) st end handle TERM _ => no_tac st | Bind => no_tac st (* renames bound state variable according to name given in goal, * before rule specAnno is applied, and solves sidecondition *) fun gen_Anno_tac dest rules tac cont_tac ctxt state_kind (t,i) st = let val vars = (dest o #2 o dest_hoare) (Logic.strip_assums_concl t); val rules' = (case (List.filter (not o null) (map dest_splits vars)) of [] => rules |(xs::_) => adapt_aux_var ctxt false (map fst xs) (map get_aux_tvar rules)); in EVERY [resolve_tac ctxt rules' i, tac, simp_tac (put_simpset HOL_basic_ss ctxt addsimps ([@{thm split_conv}, refl, @{thm Hoare.triv_All_eq}]@anno_defs) addsimprocs [@{simproc case_prod_beta}]) (i+2), simp_tac (put_simpset HOL_basic_ss ctxt addsimps [@{thm split_conv}]) (i+1), simp_tac (put_simpset HOL_basic_ss ctxt addsimps [@{thm split_conv}]) i, REPEAT (resolve_tac ctxt [allI] (i+1)), cont_tac (i+1), conseq_simp_tac ctxt state_kind [] i] st end handle TERM _ => no_tac st; fun specAnno_tac ctxt state_kind cont_tac mode = let fun dest_specAnno (Const (@{const_name Language.specAnno},_)$P$c$Q$A) = [P,c,Q,A] | dest_specAnno t = raise TERM ("dest_specAnno",[t]); val rules = [SpecAnnoNoAbrupt mode,SpecAnno mode]; in gen_Anno_tac dest_specAnno rules all_tac cont_tac ctxt state_kind end; fun whileAnnoFix_tac ctxt state_kind cont_tac mode (t,i) = let fun dest (Const (@{const_name Language.whileAnnoFix},_)$b$I$V$c) = [I,V,c] | dest t = raise TERM ("dest_whileAnnoFix",[t]); val rules = [WhileAnnoFix mode]; fun wf_tac' i = EVERY [REPEAT (resolve_tac ctxt [allI] i), simp_tac (put_simpset HOL_basic_ss ctxt addsimps [@{thm split_conv}]) i, wf_tac ctxt i]; val tac = if mode=Total then EVERY [wf_tac' (i+3),in_rel_simp ctxt (i+1)] else all_tac in gen_Anno_tac dest rules tac cont_tac ctxt state_kind (t,i) end; fun lemAnno_tac ctxt state_kind mode (t,i) st = let fun dest_name (Const (x,_)) = x | dest_name (Free (x,_)) = x | dest_name t = raise TERM ("dest_name",[t]); fun dest_lemAnno ctxt (Const (@{const_name Language.lem},_)$n$c) = let val x = Long_Name.base_name (dest_name n); in (case try (Proof_Context.get_thm ctxt) x of NONE => error ("No lemma: '" ^ x ^ "' found.") | SOME spec => (strip_qnt_vars @{const_name All} (HOLogic.dest_Trueprop (Thm.concl_of spec)),spec)) end | dest_lemAnno _ t = raise TERM ("dest_lemAnno",[t]); val (vars,spec) = dest_lemAnno ctxt (#2 (dest_hoare t)); val rules = [LemAnnoNoAbrupt mode,LemAnno mode]; val rules' = (case vars of [] => rules | xs => adapt_aux_var ctxt true (map fst xs) (map get_aux_tvar rules)); in EVERY [resolve_tac ctxt rules' i, resolve_tac ctxt [spec] (i+1), conseq_simp_tac ctxt state_kind [] i] st end handle TERM _ => no_tac st; fun prems_tac ctxt i = TRY (resolve_tac ctxt (Assumption.all_prems_of ctxt) i); fun mk_proc_assoc thms = let fun name (_,p,_,_,cmode,_) = proc_name cmode p; fun proc_name thm = thm |> Thm.concl_of |> dest_hoare |> #2 |> dest_call |> name; in map (fn thm => (proc_name thm,thm)) thms end; fun mk_hoare_tac cont ctxt mode i (name,tac) = EVERY [trace_tac ctxt ("trying: " ^ name),tac cont ctxt mode i]; (* the main hoare tactic *) fun HoareTac annotate_inv exspecs strip_guards mode state_kind state_space spec_sfx ctxt tac st = let val (P,c,Q,A,_,G,T,F) = dest_hoare (Logic.strip_assums_concl (Logic.get_goal (Thm.prop_of st) 1)); val wp_tacs = #wp_tacs (get_data ctxt); val hoare_tacs = #hoare_tacs (get_data ctxt); val params = (strip_vars (Logic.get_goal (Thm.prop_of st) 1)); val Inv = (if annotate_inv then (* Take the postcondition of the triple as invariant for all *) (* while loops (makes sense for the modifies clause) *) SOME Q else NONE); val exspecthms = map (Proof_Context.get_thm ctxt) exspecs; val asms = try (fn () => mk_proc_assoc (gen_context_thms ctxt mode params G T F @ exspecthms)) () |> the_default []; fun while_annoG_tac (t,i) = whileAnnoG_tac ctxt (annotate_inv orelse strip_guards) mode t i; fun WlpTac tac i = (* WlpTac does not end with subset_refl *) FIRST ([EVERY [resolve_tac ctxt [Seq mode] i,trace_tac ctxt "Seq",HoareRuleTac tac false (i+1)], EVERY [resolve_tac ctxt [Catch mode] i,trace_tac ctxt "Catch",HoareRuleTac tac false (i+1)], EVERY [resolve_tac ctxt [CondCatch mode] i,trace_tac ctxt "CondCatch",HoareRuleTac tac false (i+1)], EVERY [resolve_tac ctxt [BSeq mode] i,trace_tac ctxt "BSeq",HoareRuleTac tac false (i+1)], EVERY [resolve_tac ctxt [FCall mode] i,trace_tac ctxt "FCall"], EVERY [resolve_tac ctxt [GuardsNil mode] i,trace_tac ctxt "GuardsNil"], EVERY [resolve_tac ctxt [GuardsConsGuaranteeStrip mode] i, trace_tac ctxt "GuardsConsGuaranteeStrip"], EVERY [resolve_tac ctxt [GuardsCons mode] i,trace_tac ctxt "GuardsCons"], EVERY [SUBGOAL while_annoG_tac i] ] @ map (mk_hoare_tac (HoareRuleTac tac) ctxt mode i) wp_tacs) and HoareRuleTac tac pre_cond i st = let fun call (t,i) = call_tac (HoareRuleTac tac false) mode state_kind state_space ctxt asms spec_sfx t i fun cond_tac i = if annotate_inv andalso Config.get ctxt use_cond_inv_modifies then EVERY[SUBGOAL (cond_annotate_tac ctxt (the Inv) mode) i, HoareRuleTac tac false (i+4), HoareRuleTac tac false (i+3), BasicSimpTac ctxt state_kind true [] tac (i+2), BasicSimpTac ctxt state_kind true [] tac (i+1) ] else EVERY[resolve_tac ctxt [Cond mode] i,trace_tac ctxt "Cond", HoareRuleTac tac false (i+2), HoareRuleTac tac false (i+1)]; fun switch_tac i = EVERY[resolve_tac ctxt [SwitchNil mode] i, trace_tac ctxt "SwitchNil"] ORELSE EVERY[resolve_tac ctxt [SwitchCons mode] i,trace_tac ctxt "SwitchCons", HoareRuleTac tac false (i+2), HoareRuleTac tac false (i+1)]; fun while_tac' (t,i) = while_tac ctxt state_kind Inv (HoareRuleTac tac true) tac mode t i; in st |> ( (WlpTac tac i THEN HoareRuleTac tac pre_cond i) ORELSE (FIRST([EVERY[resolve_tac ctxt [Skip mode] i,trace_tac ctxt "Skip"], EVERY[resolve_tac ctxt [BasicCond mode] i, trace_tac ctxt "BasicCond", assertion_simp_tac ctxt state_kind [] i], (resolve_tac ctxt [Basic mode] i THEN trace_tac ctxt "Basic") THEN_MAYBE (assertion_simp_tac ctxt state_kind [] i), (* we don't really need simplificaton here. The question is if it is better to simplify the assertion after each Basic step, so that intermediate assertions stay "small", or if we just accumulate the raw assertions and leave the simplification to the final BasicSimpTac *) EVERY[resolve_tac ctxt [Throw mode] i, trace_tac ctxt "Throw"], (resolve_tac ctxt [Raise mode] i THEN trace_tac ctxt "Raise") THEN_MAYBE (assertion_string_eq_simp_tac ctxt state_kind [] i), cond_tac i, switch_tac i, EVERY[resolve_tac ctxt [Block mode] i, trace_tac ctxt "Block", resolve_tac ctxt [allI] (i+2), resolve_tac ctxt [allI] (i+2), HoareRuleTac tac false (i+2), resolve_tac ctxt [allI] (i+1), in_assertion_simp_tac ctxt state_kind [] (i+1), HoareRuleTac tac false (i+1)], SUBGOAL while_tac' i, SUBGOAL (guard_tac ctxt (annotate_inv orelse strip_guards) (HoareRuleTac tac false) mode) i, EVERY[SUBGOAL (specAnno_tac ctxt state_kind (HoareRuleTac tac true) mode) i], EVERY[SUBGOAL (whileAnnoFix_tac ctxt state_kind (HoareRuleTac tac true) mode) i], EVERY[resolve_tac ctxt [SpecIf mode] i, trace_tac ctxt "SpecIf", assertion_simp_tac ctxt state_kind [] i], (resolve_tac ctxt [Spec mode] i THEN trace_tac ctxt "Spec") THEN_MAYBE (assertion_simp_tac ctxt state_kind [@{thm split_conv}] i), EVERY[resolve_tac ctxt [BindR mode] i, trace_tac ctxt "Bind", resolve_tac ctxt [allI] (i+1), HoareRuleTac tac false (i+1)], EVERY[resolve_tac ctxt [DynCom mode] i, trace_tac ctxt "DynCom"], EVERY[trace_tac ctxt "calling call_tac",SUBGOAL call i], EVERY[trace_tac ctxt "LemmaAnno",SUBGOAL (lemAnno_tac ctxt state_kind mode) i]] @ map (mk_hoare_tac (HoareRuleTac tac) ctxt mode i) hoare_tacs) THEN (if pre_cond then EVERY [trace_tac ctxt "pre_cond", TRY (BasicSimpTac ctxt state_kind true [] tac i), (* FIXME: Do we need TRY *) trace_tac ctxt "after BasicSimpTac"] else (resolve_tac ctxt [subset_refl] i)))) end; in ((K (EVERY [REPEAT (resolve_tac ctxt [allI] 1), HoareRuleTac tac true 1])) THEN_ALL_NEW (prems_tac ctxt)) 1 st (*Procedure specifications may have an locale assumption as premise. These are accumulated by the vcg and are be solved afterward by prems_tac *) end; fun prefer_tac i = (Tactic.defer_tac i THEN PRIMITIVE (Thm.permute_prems 0 ~1)); fun HoareStepTac strip_guards mode state_kind state_space spec_sfx ctxt tac st = let val asms = try (fn () => let val (_,_,_,_,_,G,T,F) = dest_hoare (Logic.strip_assums_concl (Logic.get_goal (Thm.prop_of st) 1)); val params = (strip_vars (Logic.get_goal (Thm.prop_of st) 1)); in mk_proc_assoc (gen_context_thms ctxt mode params G T F) end) () |> the_default []; fun result_tac i = TRY (EVERY [resolve_tac ctxt [Basic mode] i, resolve_tac ctxt [subset_refl] i]); fun call (t,i) = call_tac result_tac mode state_kind state_space ctxt asms spec_sfx t i fun final_simp_tac i = EVERY [full_simp_tac (put_simpset HOL_basic_ss ctxt addsimps [mem_Collect_eq]) i, REPEAT (eresolve_tac ctxt [conjE] i), TRY (hyp_subst_tac_thin true ctxt i), BasicSimpTac ctxt state_kind true [] tac i] fun while_annoG_tac (t,i) = whileAnnoG_tac ctxt strip_guards mode t i; in st |> (REPEAT (resolve_tac ctxt [allI] 1) THEN FIRST [resolve_tac ctxt [subset_refl] 1, EVERY[resolve_tac ctxt [Skip mode] 1,TRY (BasicSimpTac ctxt state_kind false [] tac 1)], EVERY[resolve_tac ctxt [BasicCond mode] 1,trace_tac ctxt "BasicCond", TRY (BasicSimpTac ctxt state_kind false [] tac 1)], EVERY[resolve_tac ctxt [Basic mode] 1,TRY (BasicSimpTac ctxt state_kind false [] tac 1)], EVERY[resolve_tac ctxt [Throw mode] 1,TRY (BasicSimpTac ctxt state_kind false [] tac 1)], EVERY[resolve_tac ctxt [Raise mode] 1,TRY (assertion_string_eq_simp_tac ctxt state_kind [] 1)], resolve_tac ctxt [SeqSwap mode] 1 THEN_MAYBE HoareStepTac strip_guards mode state_kind state_space spec_sfx ctxt tac, EVERY[resolve_tac ctxt [BSeq mode] 1, prefer_tac 2 THEN_MAYBE HoareStepTac strip_guards mode state_kind state_space spec_sfx ctxt tac], resolve_tac ctxt [CondSwap mode] 1, resolve_tac ctxt [SwitchNil mode] 1, resolve_tac ctxt [SwitchCons mode] 1, EVERY [SUBGOAL while_annoG_tac 1], EVERY[resolve_tac ctxt [While mode] 1, if mode=Total then wf_tac ctxt 4 else all_tac, BasicSimpTac ctxt state_kind false [] tac 3, if mode=Total then in_rel_simp ctxt 2 THEN (resolve_tac ctxt [allI] 2) else all_tac, BasicSimpTac ctxt state_kind false [] tac 1], resolve_tac ctxt [CatchSwap mode] 1, resolve_tac ctxt [CondCatchSwap mode] 1, EVERY[resolve_tac ctxt [BlockSwap mode] 1, resolve_tac ctxt [allI] 1, resolve_tac ctxt [allI] 1, resolve_tac ctxt [allI] 2, BasicSimpTac ctxt state_kind false [] tac 2], resolve_tac ctxt [GuardsNil mode] 1, resolve_tac ctxt [GuardsConsGuaranteeStrip mode] 1, resolve_tac ctxt [GuardsCons mode] 1, SUBGOAL (guard_tac ctxt strip_guards (fn i => all_tac) mode) 1, EVERY[SUBGOAL (specAnno_tac ctxt state_kind (K all_tac) mode) 1], EVERY[SUBGOAL (whileAnnoFix_tac ctxt state_kind (K all_tac) mode) 1], EVERY[resolve_tac ctxt [SpecIf mode] 1,trace_tac ctxt "SpecIf", TRY (BasicSimpTac ctxt state_kind false [] tac 1)], EVERY[resolve_tac ctxt [Spec mode] 1, TRY (BasicSimpTac ctxt state_kind false [@{thm split_conv}] tac 1)], EVERY[resolve_tac ctxt [BindR mode] 1, resolve_tac ctxt [allI] 2, prefer_tac 2], EVERY[resolve_tac ctxt [FCall mode] 1], EVERY[resolve_tac ctxt [DynCom mode] 1], EVERY[SUBGOAL call 1, BasicSimpTac ctxt state_kind false [] tac 1], EVERY[SUBGOAL (lemAnno_tac ctxt state_kind mode) 1, BasicSimpTac ctxt state_kind false [] tac 1], final_simp_tac 1 ]) end; (*****************************************************************************) (** Generalise verification condition **) (*****************************************************************************) structure RecordSplitState : SPLIT_STATE = struct val globals = @{const_name StateSpace.state.globals}; fun isState (Const _$Abs (s,T,t)) = (case (state_hierarchy T) of ((n,_)::_) => n = "StateSpace.state.state" andalso is_none (try dest_hoare_raw (strip_qnt_body @{const_name All} t)) | _ => false) | isState _ = false; fun isFreeState (Free (_,T)) = (case (state_hierarchy T) of ((n,_)::_) => n = "StateSpace.state.state" | _ => false) | isFreeState _ = false; val abs_state = Option.map snd o first_subterm isFreeState; fun sel_eq (Const (x,_)$_) y = (x=y) | sel_eq t y = raise TERM ("RecordSplitState.sel_eq",[t]); val sel_idx = idx sel_eq; fun bound xs (t as (Const (x,_)$_)) = let val i = sel_idx xs x in if i < 0 then (length xs, xs@[t]) else (i,xs) end | bound xs t = raise TERM ("RecordSplitState.bound",[t]); fun abs_var _ (Const (x,T)$_) = (remdeco' (Long_Name.base_name x),range_type T) | abs_var _ t = raise TERM ("RecordSplitState.abs_var",[t]); fun fld_eq (x, _) y = (x = y) fun fld_idx xs x = idx fld_eq xs x; fun sort_vars ctxt T vars = let val thy = Proof_Context.theory_of ctxt; val (flds,_) = Record.get_recT_fields thy T; val gT = the (AList.lookup (fn (x:string,y) => x=y) flds globals); val (gflds,_) = (Record.get_recT_fields thy gT handle TYPE _ => ([],("",dummyT))); fun compare (Const _$Free _, Const _$(Const _$Free _)) = GREATER | compare (Const (s1,_)$Free _, Const (s2,_)$Free _) = int_ord (fld_idx flds s1, fld_idx flds s2) | compare (Const (s1,_)$(Const _$Free _), Const (s2,_)$(Const _$Free _)) = int_ord (fld_idx gflds s1, fld_idx gflds s2) | compare _ = LESS; in sort (rev_order o compare) vars end; fun fold_state_prop loc glob app abs other inc s res (t as (Const (sel,_)$Free (s',_))) = if s'=s then if is_state_var sel then loc inc res t else raise TERM ("RecordSplitState.fold_state_prop",[t]) else other res t | fold_state_prop loc glob app abs other inc s res (t as ((t1 as (Const (sel,_)))$(t2 as (Const (glb,_)$Free (s',_))))) = if s'=s andalso is_state_var sel andalso (glb=globals) then glob inc res t else let val res1 = fold_state_prop loc glob app abs other inc s res t1 val res2 = fold_state_prop loc glob app abs other inc s res1 t2 in app res1 res2 end | fold_state_prop loc glob app abs other inc s res (t as (Free (s',_))) = if s'=s then raise TERM ("RecordSplitState.fold_state_prop",[t]) else other res t | fold_state_prop loc glob app abs other inc s res (t1$t2) = let val res1 = fold_state_prop loc glob app abs other inc s res t1 val res2 = fold_state_prop loc glob app abs other inc s res1 t2 in app res1 res2 end | fold_state_prop loc glob app abs other inc s res (Abs (x,T,t)) = let val res1 = fold_state_prop loc glob app abs other (inc+1) s res t in abs x T res1 end | fold_state_prop loc glob app abs other inc s res t = other res t fun collect_vars s t = let fun loc _ vars t = snd (bound vars t); fun glob _ vars t = snd (bound vars t); fun app _ vars2 = vars2; fun abs _ _ vars = vars; fun other vars _ = vars; in fold_state_prop loc glob app abs other 0 s [] t end; fun abstract_vars vars s t = let fun loc inc _ t = let val i = fst (bound vars t) in Bound (i+inc) end; fun glob inc _ t = let val i = fst (bound vars t) in Bound (i+inc) end; fun app t1 t2 = t1$t2; fun abs x T t = Abs (x,T,t); fun other _ t = t; val dummy = Bound 0; in fold_state_prop loc glob app abs other 0 s dummy t end; fun split_state ctxt s T t = let val vars = collect_vars s t; val vars' = if Config.get ctxt sort_variables then sort_vars ctxt T vars else vars; in (abstract_vars vars' s t,rev vars') end; fun ex_tac ctxt _ st = Record.split_simp_tac ctxt @{thms simp_thms} (K ~1) 1 st; end; structure FunSplitState : SPLIT_STATE = struct val full_globalsN = @{const_name StateSpace.state.globals}; fun isState (Const _$Abs (s,T,t)) = (case (state_hierarchy T) of ((n,_)::_) => n = "StateSpace.state.state" andalso is_none (try dest_hoare_raw (strip_qnt_body @{const_name All} t)) | _ => false) | isState _ = false; fun isFreeState (Free (_,T)) = (case (state_hierarchy T) of ((n,_)::_) => n = "StateSpace.state.state" | _ => false) | isFreeState _ = false; val abs_state = Option.map snd o first_subterm isFreeState; fun comp_name t = case try (implode o dest_string) t of SOME str => str | NONE => (case t of Free (s,_) => s | Const (s,_) => s | t => raise TERM ("FunSplitState.comp_name",[t])) fun sel_name (Const _$_$name$_) = comp_name name | sel_name t = raise TERM ("FunSplitState.sel_name",[t]); fun sel_raw_name (Const _$_$name$_) = name | sel_raw_name t = raise TERM ("FunSplitState.sel_raw_name",[t]); fun component_type (Const _$_$_$(sel$_)) = range_type (fastype_of sel) | component_type t = raise TERM ("FunSplitState.component_type",[t]); fun component_name (Const _$_$_$((Const (sel,_)$_))) = sel | component_name t = raise TERM ("FunSplitState.component_name",[t]); fun sel_type (Const _$destr$_$_) = range_type (fastype_of destr) | sel_type t = raise TERM ("FunSplitState.sel_type",[t]); fun sel_destr (Const _$destr$_$_) = destr | sel_destr t = raise TERM ("FunSplitState.sel_destr",[t]); fun sel_eq t y = (sel_name t = y) | sel_eq t y = raise TERM ("FunSplitState.sel_eq",[t]); val sel_idx = idx sel_eq; fun bound xs t = let val i = sel_idx xs (sel_name t) in if i < 0 then (length xs, xs@[t]) else (i,xs) end | bound xs t = raise TERM ("FunSplitState.bound",[t]); fun fold_state_prop var app abs other inc s res (t as (Const (@{const_name StateFun.lookup},_)$destr$name$(Const _$Free (s',_)))) = if s'=s then var inc res t else other res t (*raise TERM ("FunSplitState.fold_state_prop",[t])*) | fold_state_prop var app abs other inc s res (t as (Free (s',_))) = if s'=s then raise TERM ("FunSplitState.fold_state_prop",[t]) else other res t | fold_state_prop var app abs other inc s res (t1$t2) = let val res1 = fold_state_prop var app abs other inc s res t1 val res2 = fold_state_prop var app abs other inc s res1 t2 in app res1 res2 end | fold_state_prop var app abs other inc s res (Abs (x,T,t)) = let val res1 = fold_state_prop var app abs other (inc+1) s res t in abs x T res1 end | fold_state_prop var app abs other inc s res t = other res t fun collect_vars s t = let fun var _ vars t = snd (bound vars t); fun app _ vars2 = vars2; fun abs _ _ vars = vars; fun other vars _ = vars; in fold_state_prop var app abs other 0 s [] t end; fun abstract_vars vars s t = let fun var inc _ t = let val i = fst (bound vars t) in Bound (i+inc) end; fun app t1 t2 = t1$t2; fun abs x T t = Abs (x,T,t); fun other _ t = t; val dummy = Bound 0; in fold_state_prop var app abs other 0 s dummy t end; fun sort_vars ctxt vars = let val fld_idx = idx (fn s1:string => fn s2 => s1 = s2); fun compare (_$_$n$(Const (s1,_)$_),_$_$m$(Const (s2,_)$_)) = let val n' = remdeco' (comp_name n); val m' = remdeco' (comp_name m); in if s1 = full_globalsN then if s2 = full_globalsN then string_ord (n',m') else LESS else if s2 = full_globalsN then GREATER else string_ord (n',m') end | compare (t1,t2) = raise TERM ("FunSplitState.sort_vars.compare",[t1,t2]); in sort (rev_order o compare) vars end; fun split_state ctxt s _ t = let val vars = collect_vars s t; val vars' = if Config.get ctxt sort_variables then sort_vars ctxt vars else vars; in (abstract_vars vars' s t,rev vars') end; fun abs_var _ t = (remdeco' (sel_name t), sel_type t); (* Proof for: EX x_1 ... x_n. P x_1 ... x_n * ==> EX s. P (lookup destr_1 "x_1" s) ... (lookup destr_n "x_n" s) * Implementation: * 1. Eliminate existential quantifiers in premise * 2. Instantiate s with: (%x. undefined)("x_1" := constr_1 x_1, ..., "x_n" := constr_n x_n) * 3. Simplify *) local val ss = simpset_of (put_simpset (simpset_of @{theory_context Fun}) @{context} addsimps (@{thm StateFun.lookup_def} :: @{thm StateFun.id_id_cancel} :: @{thms list.inject list.distinct char.inject cong_exp_iff_simps simp_thms}) addsimprocs [Record.simproc, StateFun.lazy_conj_simproc] |> fold Simplifier.add_cong @{thms block_conj_cong}); in fun ex_tac ctxt vs st = let val vs' = rev vs; val (Const (_,exT)$_) = HOLogic.dest_Trueprop (Logic.strip_imp_concl (Logic.get_goal (Thm.prop_of st) 1)); val sT = domain_type (domain_type exT); val s0 = Const (@{const_name HOL.undefined},sT); fun streq (s1:string,s2) = s1=s2 ; fun mk_init [] = [] | mk_init (t::ts) = let val xs = mk_init ts; val n = component_name t; val T = component_type t; in if AList.defined streq xs n then xs else (n,(T,Const (n,sT --> component_type t)$s0))::xs end; fun mk_upd (i,t) xs = let val selN = component_name t; val selT = component_type t; val (_,s) = the (AList.lookup streq xs selN); val strT = domain_type selT; val valT = range_type selT; val constr = destr_to_constr (sel_destr t); val name = (sel_raw_name t); val upd = Const (@{const_name Fun.fun_upd}, (strT --> valT)-->strT-->valT--> (strT --> valT)) $ s $ name $ (constr $ Bound i) in AList.update streq (selN,(selT,upd)) xs end; val upds = fold_index mk_upd vs' (mk_init vs'); val upd = fold (fn (n,(T,upd)) => fn s => Const (n ^ Record.updateN, T --> sT --> sT)$upd$s) upds s0; val inst = fold_rev (Term.abs o (fn t => (sel_name t, sel_type t))) vs upd; fun lift_inst_ex_tac i st = let val rule = Thm.lift_rule (Thm.cprem_of st i) (Drule.incr_indexes st exI); val (_$x) = HOLogic.dest_Trueprop (Logic.strip_assums_concl (hd (Thm.prems_of rule))); val inst_rule = infer_instantiate ctxt [(#1 (dest_Var (head_of x)), Thm.cterm_of ctxt inst)] rule; in (compose_tac ctxt (false,inst_rule, Thm.nprems_of exI) i st) end; in EVERY [REPEAT_DETERM_N (length vs) (eresolve_tac ctxt [exE] 1), lift_inst_ex_tac 1, simp_tac (put_simpset ss ctxt) 1 ] st end end (* Test: What happens when there are no lookups., EX s. True *) end; structure GeneraliseRecord = GeneraliseFun (structure SplitState=RecordSplitState); structure GeneraliseStateFun = GeneraliseFun (structure SplitState=FunSplitState); fun generalise Record = GeneraliseRecord.GENERALISE | generalise Function = GeneraliseStateFun.GENERALISE; (*****************************************************************************) (** record_vanish_tac splits up the records of a verification condition, **) (** trying to generate a predicate without records. **) (** A typical verification condition with a procedure call will have the **) (** form "!!s Z. s=Z ==> ..., where s and Z are records **) (*****************************************************************************) (* FIXME: Check out if removing the useless vars is a performance issue. If so, maybe we can remove all useless vars at once (no iterated simplification) or try to avoid introducing them. Bevore splitting the gaol we can simplifiy the goal with state_simproc this may leed to better performance... *) fun record_vanish_tac ctxt state_kind state_space i = if Config.get ctxt record_vanish then let val rem_useless_vars_simps = [Drule.triv_forall_equality,@{thm Hoare.triv_All_eq},@{thm Hoare.triv_Ex_eq}]; val rem_useless_vars_simpset = empty_simpset ctxt addsimps rem_useless_vars_simps; fun no_spec (t as (Const (@{const_name All},_)$_)) = is_none (try dest_hoare_raw (strip_qnt_body @{const_name All} t)) | no_spec _ = true; fun state_space_no_spec t = if state_space t <> 0 andalso no_spec t then ~1 else 0; in EVERY [trace_tac ctxt "record_vanish_tac -- START --", REPEAT (eresolve_tac ctxt [conjE] i), trace_tac ctxt "record_vanish_tac -- hyp_subst_tac ctxt --", TRY (hyp_subst_tac_thin true ctxt i), full_simp_tac rem_useless_vars_simpset i, (* hyp_subst_tac may have made some state variables unnecessary. We do not want to split them to avoid naming conflicts and increase performance *) trace_tac ctxt "record_vanish_tac -- Splitting records --", if Config.get ctxt use_generalise orelse state_kind = Function then generalise state_kind ctxt i else state_split_simp_tac ctxt rem_useless_vars_simps state_space_no_spec i (*THEN_MAYBE EVERY [trace_tac ctxt "record_vanish_tac -- removing useless vars --", full_simp_tac rem_useless_vars_simpset i, trace_tac ctxt "record_vanish_tac -- STOP --"]*) ] end else all_tac; (* solve_modifies_tac tries to solve modifies-clauses automatically; * The following strategy is followed: * After clar-simplifying the modifies clause we remain with a goal of the form * * EX a b. s(|A := x|) = s(|A:=a,B:=b|) * * (or because of conditional statements conjunctions of these kind of goals) * We split up the state-records and simplify (record_vanish_tac) and get to a goal of the form: * * EX a b. (|A=x,B=B|) = (|A=a,B=b|). * * If the modifies clause was correct we just have to introduce the existential quantifies * and apply reflexivity. * If not we just simplify the goal. *) local val state_fun_update_ss = simpset_of (put_simpset HOL_basic_ss @{context} addsimps ([@{thm StateFun.update_def}, @{thm id_def}, @{thm fun_upd_apply}, @{thm if_True}, @{thm if_False}] @ @{thms list.inject list.distinct char.inject cong_exp_iff_simps simp_thms} @ K_fun_convs) addsimprocs [DistinctTreeProver.distinct_simproc ["distinct_fields", "distinct_fields_globals"]] |> Simplifier.add_cong @{thm imp_cong} (* K_fun_congs FIXME: Stefan fragen*) |> Splitter.add_split @{thm if_split}); in fun solve_modifies_tac ctxt state_kind state_space i st = let val thy = Proof_Context.theory_of ctxt; fun is_split_state (trm as (Const (@{const_name Pure.all},_)$Abs(x,T,t))) = if state_space trm <> 0 then try (fn () => let fun seed (_ $ v $ st) = seed st | seed (_ $ t) = (1,t) (* split only state pair *) | seed t = (~1,t) (* split globals completely *) val all_vars = strip_qnt_vars @{const_name Pure.all} t; val concl = HOLogic.dest_Trueprop (Logic.strip_assums_concl t); val ex_vars = strip_qnt_vars @{const_name Ex} concl; val state = Bound (length all_vars + length ex_vars); val (Const (@{const_name HOL.eq},_)$x_upd$x_upd') = strip_qnt_body @{const_name Ex} concl; val (split,sd) = seed x_upd; in if sd = state then split else 0 end) () |> the_default 0 else 0 | is_split_state t = 0; val simp_ctxt = put_simpset HOL_ss ctxt addsimps (@{thm Ex_True} :: @{thm Ex_False} :: Record.get_extinjects thy); fun try_solve Record i = (*(SOLVE*) (((fn k => (TRY (REPEAT_ALL_NEW (resolve_tac ctxt [conjI, impI, allI]) k))) THEN_ALL_NEW (fn k => EVERY [state_split_simp_tac ctxt [] is_split_state k, simp_tac simp_ctxt k THEN_MAYBE rename_goal ctxt remdeco' k ])) i) (*)*) | try_solve _ i = ((fn k => (TRY (REPEAT_ALL_NEW (resolve_tac ctxt [conjI, impI, allI]) k))) THEN_ALL_NEW (fn k => REPEAT (resolve_tac ctxt [exI] k) THEN resolve_tac ctxt [ext] k THEN simp_tac (put_simpset state_fun_update_ss ctxt) k THEN_MAYBE (REPEAT_ALL_NEW (resolve_tac ctxt [conjI,impI,refl]) k))) i in ((trace_tac ctxt "solve_modifies_tac" THEN clarsimp_tac ((ctxt |> put_claset (claset_of @{theory_context HOL}) |> put_simpset (simpset_of @{theory_context Set})) addsimps ([@{thm Hoare.mex_def},@{thm Hoare.meq_def}]@K_convs) addsimprocs (state_upd_simproc Record::(state_simprocs Record)) |> fold Simplifier.add_cong K_congs) i) THEN_MAYBE try_solve state_kind i) st end; end fun proc_lookup_simp_tac ctxt i st = try (fn () => let val name = (Logic.concl_of_goal (Thm.prop_of st) i) |> dest_hoare |> #2 |> strip_comb |> #2 |> last |> strip_comb |> #2 |> last; (* the$(Gamma$name) or the$(strip$Gamma$name) *) val pname = (unsuffix proc_deco (dest_string' name)); val thms = map_filter I (map (try (Proof_Context.get_thm ctxt)) [suffix bodyP pname, suffix (body_def_sfx^"_def") pname, suffix procL pname^"."^ (suffix (body_def_sfx^"_def") pname)]); in simp_tac (put_simpset HOL_basic_ss ctxt addsimps thms @ strip_simps @ @{thms option.sel}) i st end) () |> the_default (Seq.single st); fun proc_lookup_in_dom_simp_tac ctxt i st = try (fn () => let val _$name$_ = (HOLogic.dest_Trueprop (Logic.concl_of_goal (Thm.prop_of st) i)); (* name : Gamma *) val pname = (unsuffix proc_deco (dest_string' name)); val thms = map_filter I (map (try (Proof_Context.get_thm ctxt)) [suffix bodyP pname]); in simp_tac (put_simpset HOL_basic_ss ctxt addsimps (@{thm Hoare.lookup_Some_in_dom} :: @{thm dom_strip} :: thms)) i st end) () |> the_default (Seq.single st); fun HoareRuleTac ctxt insts fixes st = let val annotate_simp_tac = simp_tac (put_simpset HOL_basic_ss ctxt addsimps (anno_defs@normalize_simps) addsimprocs [@{simproc case_prod_beta}]); fun is_com_eq (Const (@{const_name Trueprop},_)$(Const (@{const_name HOL.eq},T)$_$_)) = (case (binder_types T) of (Type (@{type_name Language.com},_)::_) => true | _ => false) | is_com_eq _ = false; fun annotate_tac i st = if is_com_eq (Logic.concl_of_goal (Thm.prop_of st) i) then annotate_simp_tac i st else all_tac st; in ((fn i => REPEAT (resolve_tac ctxt [allI] i)) THEN' Rule_Insts.res_inst_tac ctxt insts fixes st) THEN_ALL_NEW annotate_tac end; fun HoareCallRuleTac state_kind state_space ctxt thms i st = let fun dest_All (Const (@{const_name All},_)$t) = SOME t | dest_All _ = NONE; fun auxvars t = (case (map_filter ((first_subterm is_hoare) o snd) (max_subterms_dest dest_All t)) of ((vars,_)::_) => vars | _ => []); fun auxtype rule = (case (auxvars (Thm.prop_of rule)) of [] => NONE | vs => (case (last vs) of (_,TVar (z,_)) => SOME (z,rule) | _ => NONE)); val thms' = let val auxvs = map fst (auxvars (Logic.concl_of_goal (Thm.prop_of st) i)); val tvar_thms = map_filter auxtype thms in if length thms = length tvar_thms then adapt_aux_var ctxt true auxvs tvar_thms else thms end; val is_sidecondition = not o can dest_hoare; fun solve_sidecondition_tac (t,i) = if is_sidecondition t then FIRST [CHANGED_PROP (wf_tac ctxt i), (*init_conforms_tac state_kind state_space i,*) post_conforms_tac ctxt state_kind i THEN_MAYBE (if is_modifies_clause t then solve_modifies_tac ctxt state_kind state_space i else all_tac), proc_lookup_in_dom_simp_tac ctxt i ] else in_rel_simp ctxt i THEN simp_tac (put_simpset HOL_basic_ss ctxt addsimps [@{thm Un_empty_left},@{thm Un_empty_right}]) i THEN proc_lookup_simp_tac ctxt i fun basic_tac i = (((resolve_tac ctxt thms') THEN_ALL_NEW (fn k => (SUBGOAL solve_sidecondition_tac k))) i) in (basic_tac ORELSE' (fn k => (REPEAT (resolve_tac ctxt [allI] k)) THEN EVERY [resolve_tac ctxt thms' k])) i st end; (* vcg_polish_tac tries to solve modifies-clauses automatically; for other specifications the * records are only splitted and simplified. *) fun vcg_polish_tac solve_modifies ctxt state_kind state_space i = if solve_modifies then solve_modifies_tac ctxt state_kind state_space i else record_vanish_tac ctxt state_kind state_space i THEN_MAYBE EVERY [rename_goal ctxt remdeco' i(*, simp_tac (HOL_basic_ss addsimps @{thms simp_thms})) i*)]; fun is_funtype (Type ("fun", _)) = true | is_funtype _ = false; fun state_kind_of ctxt T = let val thy = Proof_Context.theory_of ctxt; val (s,sT) = nth (fst (Record.get_recT_fields thy T)) 1; in if Long_Name.base_name s = "locals" andalso is_funtype sT then Function else Record end handle Subscript => Record; fun find_state_space_in_triple ctxt t = try (fn () => (case first_subterm is_hoare t of NONE => NONE | SOME (abs_vars,triple) => let val (_,com,_,_,mode,_,_,_) = dest_hoare_raw triple; val T = fastype_of1 (map snd abs_vars,com) val Type(_,state_spaceT::_) = T; val SOME Tids = stateT_ids state_spaceT; in SOME (Tids,mode, state_kind_of ctxt state_spaceT) end)) () |> Option.join; fun get_state_space_in_subset_eq ctxt t = (* get state type from the following kind of terms: P <= Q, s: P *) try (fn () => let val (subset_eq,_) = (strip_comb o HOLogic.dest_Trueprop o strip_qnt_body @{const_name Pure.all}) t; val Ts = map snd (strip_vars t); val T = fastype_of1 (Ts,subset_eq); val Type (_, [_,Type (_, [Type (_, [state_spaceT]), _])]) = T; (* also works for "in": x : P *) val SOME Tids = stateT_ids state_spaceT; in (Tids,Partial, state_kind_of ctxt state_spaceT) end) (); fun get_state_space ctxt i st = (case try (Logic.concl_of_goal (Thm.prop_of st)) i of SOME t => (case find_state_space_in_triple ctxt t of SOME sp => SOME sp | NONE => get_state_space_in_subset_eq ctxt t) | NONE => NONE); fun mk_hoare_tac hoare_tac finish_tac annotate_inv exnames strip_guards spec_sfx ctxt i st = case get_state_space ctxt i st of SOME (Tids,mode,kind) => SELECT_GOAL (hoare_tac annotate_inv exnames strip_guards mode kind (is_state_space_var Tids) spec_sfx ctxt (finish_tac kind (is_state_space_var Tids))) i st | NONE => no_tac st fun vcg_tac spec_sfx strip_guards exnames ctxt i st = mk_hoare_tac HoareTac (vcg_polish_tac (spec_sfx="_modifies") ctxt) (spec_sfx="_modifies") exnames (strip_guards="true") spec_sfx ctxt i st; fun hoare_tac spec_sfx strip_guards _ ctxt i st = let fun tac state_kind state_space i = if spec_sfx="_modifies" then solve_modifies_tac ctxt state_kind state_space i else all_tac; in mk_hoare_tac HoareTac tac (spec_sfx="_modifies") [] (strip_guards="true") spec_sfx ctxt i st end; fun hoare_raw_tac spec_sfx strip_guards exnames ctxt i st = mk_hoare_tac HoareTac (K (K (K all_tac))) (spec_sfx="_modifies") [] (strip_guards="true") spec_sfx ctxt i st; fun hoare_step_tac spec_sfx strip_guards exnames ctxt i st = mk_hoare_tac (K (K HoareStepTac)) (vcg_polish_tac (spec_sfx="_modifies") ctxt) false [] (strip_guards="true") spec_sfx ctxt i st; fun hoare_rule_tac ctxt thms i st = SUBGOAL (fn _ => (case get_state_space ctxt i st of SOME (Tids,_,kind) => HoareCallRuleTac kind (is_state_space_var Tids) ctxt thms i | NONE => error "could not find proper state space type (structure or record) in goal")) i st; (*** Methods ***) val hoare_rule = Rule_Insts.method HoareRuleTac hoare_rule_tac; val argP = Args.name --| @{keyword "="} -- Args.name val argsP = Scan.repeat argP val default_args = [("spec","spec"),("strip_guards","false")] val vcg_simp_modifiers = [Args.add -- Args.colon >> K (Method.modifier vcg_simp_add \<^here>), Args.del -- Args.colon >> K (Method.modifier vcg_simp_del \<^here>)]; fun assocs2 key = map snd o filter (curry (op =) key o fst); fun gen_simp_method tac = Scan.lift (argsP >> (fn args => args @ default_args)) --| Method.sections vcg_simp_modifiers >> (fn args => fn ctxt => Method.SIMPLE_METHOD' (tac ("_" ^ the (AList.lookup (op =) args "spec")) (the (AList.lookup (op =) args "strip_guards")) (assocs2 "exspec" args) ctxt)); val hoare = gen_simp_method hoare_tac; val hoare_raw = gen_simp_method hoare_raw_tac; val vcg = gen_simp_method vcg_tac; val vcg_step = gen_simp_method hoare_step_tac; val trace_hoare_users = Unsynchronized.ref false fun print_subgoal_tac ctxt s i = SUBGOAL (fn (prem, _) => trace_tac ctxt (s ^ (Syntax.string_of_term ctxt prem))) i fun mk_hoare_thm thm _ ctxt _ i = EVERY [resolve_tac ctxt [thm] i, if !trace_hoare_users then print_subgoal_tac ctxt "Tracing: " i else all_tac] val vcg_hoare_add = Thm.declaration_attribute (fn thm => add_hoare_tacs [(Thm.derivation_name thm, mk_hoare_thm thm)]) exception UNDEF val vcg_hoare_del = Thm.declaration_attribute (fn _ => fn _ => raise UNDEF) (* setup theory *) val _ = Theory.setup (Attrib.setup @{binding vcg_simp} (Attrib.add_del vcg_simp_add vcg_simp_del) "declaration of Simplifier rule for vcg" #> Attrib.setup @{binding vcg_hoare} (Attrib.add_del vcg_hoare_add vcg_hoare_del) "declaration of wp rule for vcg") (*#> add_wp_tacs initial_wp_tacs*) end;