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,2233 @@ 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; + (fn {context = goal_ctxt, prems} => + Local_Defs.unfold0_tac goal_ctxt (map tryext prems) THEN resolve_tac goal_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 (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/Akra_Bazzi/akra_bazzi.ML b/thys/Akra_Bazzi/akra_bazzi.ML --- a/thys/Akra_Bazzi/akra_bazzi.ML +++ b/thys/Akra_Bazzi/akra_bazzi.ML @@ -1,629 +1,630 @@ (* Title: akra_bazzi.ML Author: Manuel Eberl, TU Muenchen *) signature AKRA_BAZZI = sig val master_theorem_tac : string option -> bool -> thm option -> term option -> term option -> term option -> Proof.context -> int -> tactic val master_theorem_function_tac : bool -> Proof.context -> int -> tactic val akra_bazzi_term_tac : Proof.context -> int -> tactic val akra_bazzi_sum_tac : Proof.context -> int -> tactic val akra_bazzi_relation_tac : Proof.context -> int -> tactic val akra_bazzi_measure_tac : Proof.context -> int -> tactic val akra_bazzi_termination_tac : Proof.context -> int -> tactic val setup_master_theorem : Context.generic * Token.T list -> (Proof.context -> Method.method) * (Context.generic * Token.T list) end structure Akra_Bazzi: AKRA_BAZZI = struct fun changed_conv conv ct = let val thm = conv ct in if Thm.is_reflexive thm then raise CTERM ("changed_conv", [ct]) else thm end fun rewrite_tac ctxt thms = CONVERSION (Conv.repeat_conv (changed_conv (Conv.bottom_conv (K (Conv.try_conv (Conv.rewrs_conv (map (fn thm => thm RS @{thm eq_reflection}) thms)))) ctxt))) val term_preps = @{thms akra_bazzi_termination_simps} fun get_term_intros ctxt = Named_Theorems.get ctxt @{named_theorems "akra_bazzi_term_intros"} val master_intros = @{thms master_theorem_function.master1_automation master_theorem_function.master1_bigo_automation master_theorem_function.master1_automation' master_theorem_function.master1_bigo_automation' master_theorem_function.master2_1_automation master_theorem_function.master2_1_automation' master_theorem_function.master2_2_automation master_theorem_function.master2_2_automation' master_theorem_function.master2_3_automation master_theorem_function.master2_3_automation' master_theorem_function.master3_automation master_theorem_function.master3_automation'} fun get_intros NONE b = if not b then master_intros else @{thms master_theorem_function.master1_automation master_theorem_function.master1_bigo_automation master_theorem_function.master1_automation' master_theorem_function.master2_1_automation master_theorem_function.master2_1_automation'} | get_intros (SOME "1") _ = @{thms master_theorem_function.master1_automation master_theorem_function.master1_bigo_automation master_theorem_function.master1_automation' master_theorem_function.master1_bigo_automation'} | get_intros (SOME "2.1") _ = @{thms master_theorem_function.master2_1_automation master_theorem_function.master2_1_automation'} | get_intros (SOME "2.2") _ = @{thms master_theorem_function.master2_2_automation master_theorem_function.master2_2_automation'} | get_intros (SOME "2.3") _ = @{thms master_theorem_function.master2_3_automation master_theorem_function.master2_3_automation'} | get_intros (SOME "3") _ = @{thms master_theorem_function.master3_automation master_theorem_function.master3_automation'} | get_intros _ _ = error "invalid Master theorem case" val allowed_consts = [@{term "(+) :: nat => nat => nat"}, @{term "Suc :: nat => nat"}, @{term "(+) :: real => real => real"}, @{term "(-) :: real => real => real"}, @{term "uminus :: real => real"}, @{term "(*) :: real => real => real"}, @{term "(/) :: real => real => real"}, @{term "inverse :: real => real"}, @{term "(powr) :: real => real => real"}, @{term "real_of_nat"}, @{term "real :: nat => real"}, @{term "Num.numeral :: num => real"}, @{term "Num.Bit0"}, @{term "Num.Bit1"}, @{term "0 :: real"}, @{term "1 :: real"}, @{term "Num.One"}, @{term "Num.BitM"}, @{term "Num.numeral :: num => nat"}, @{term "0 :: nat"}, @{term "1 :: nat"}, @{term "atLeastAtMost :: real => real => real set"}, @{term "greaterThanAtMost :: real => real => real set"}, @{term "atLeastLessThan :: real => real => real set"}, @{term "greaterThanLessThan :: real => real => real set"}, @{term "(=) :: real => real => bool"}, @{term "(\) :: real => real => bool"}, @{term "(<) :: real => real => bool"}, @{term "(\) :: real => real set => bool"}, @{term "(=) :: nat => nat => bool"}, @{term "(\) :: nat => nat => bool"}, @{term "(<) :: nat => nat => bool"}, @{term "Trueprop"}, @{term "HOL.Not"}, @{term "HOL.conj"}, @{term "HOL.disj"}, @{term "HOL.implies"}] fun contains_only_numerals t = let fun go (s $ t) = let val _ = go s in go t end | go (Const c) = if member (op =) allowed_consts (Const c) then () else raise TERM ("contains_only_numerals", [Const c]) | go t = raise TERM ("contains_only_numerals", [t]) val _ = go t in true end handle TERM _ => false fun dest_nat_number t = case HOLogic.dest_number t of (ty, n) => if ty = @{typ "Nat.nat"} then n else raise TERM ("dest_nat_number", [t]) fun num_to_Suc n = let fun go acc 0 = acc | go acc n = go (@{term "Suc :: Nat.nat => Nat.nat"} $ acc) (n - 1) in go @{term "0 :: Nat.nat"} n end (* Termination *) datatype breadcrumb = Fst | Snd (* extract all natural number types from a product type *) fun pick_natT_leaf T = let fun go crumbs T acc = case try HOLogic.dest_prodT T of SOME (T1, T2) => acc |> go (Snd :: crumbs) T2 |> go (Fst :: crumbs) T1 | NONE => if T = HOLogic.natT then crumbs :: acc else acc in go [] T [] end fun breadcrumb Fst = fst | breadcrumb Snd = snd (* construct an extractor function using fst/snd from a path in a product type *) fun mk_extractor T crumbs = let fun aux c f (T', acc) = let val T'' = HOLogic.dest_prodT T' |> f in (T'', Const (c, T' --> T'') $ acc) end fun go [] (_, acc) = Abs ("n", T, acc) | go (Fst :: crumbs) acc = acc |> aux @{const_name fst} fst |> go crumbs | go (Snd :: crumbs) acc = acc |> aux @{const_name snd} snd |> go crumbs in go (rev crumbs) (T, Bound 0) end fun extract_from_prod [] t = t | extract_from_prod (c :: cs) t = extract_from_prod cs t |> HOLogic.dest_prod |> breadcrumb c fun ANY' tacs n = fold (curry op APPEND) (map (fn t => t n) tacs) no_tac (* find recursion variable, set up a measure function for it and apply it to prove termination *) fun akra_bazzi_relation_tac ctxt = case Proof_Context.lookup_fact ctxt "local.termination" of SOME {thms = [thm], ...} => let val prems = Thm.prems_of thm val relT = prems |> hd |> dest_comb |> snd |> dest_comb |> snd |> dest_Var |> snd val T = relT |> HOLogic.dest_setT |> HOLogic.dest_prodT |> fst val prems = prems |> tl |> map Term.strip_all_body val calls = prems |> map (Logic.strip_imp_concl #> HOLogic.dest_Trueprop #> dest_comb #> fst #> dest_comb #> snd #> HOLogic.dest_prod #> fst) val leaves = pick_natT_leaf T val calls = map (fn path => (path, map (extract_from_prod path) calls)) leaves val calls = calls |> filter (snd #> filter is_Bound #> null) |> map fst val measureC = Const (@{const_name Wellfounded.measure}, (T --> HOLogic.natT) --> relT) fun mk_relation f = measureC $ f val relations = map (mk_relation o mk_extractor T) calls in (ANY' (map (Function_Relation.relation_tac ctxt o K) relations) THEN' resolve_tac ctxt @{thms wf_measure}) THEN_ALL_NEW (rewrite_tac ctxt @{thms measure_prod_conv measure_prod_conv'}) end | _ => K no_tac val measure_postproc_simps = @{thms arith_simps of_nat_0 of_nat_1 of_nat_numeral Suc_numeral of_nat_Suc} (* prove that the measure decreases in a recursive call by showing that the recursive call is an Akra-Bazzi term. Apply simple post-processing to make the numbers less ugly *) fun akra_bazzi_measure_tac ctxt = rewrite_tac ctxt term_preps THEN' eresolve_tac ctxt @{thms akra_bazzi_term_measure} THEN' resolve_tac ctxt (get_term_intros ctxt) THEN_ALL_NEW rewrite_tac ctxt measure_postproc_simps (* prove termination of an Akra-Bazzi function *) fun akra_bazzi_termination_tac ctxt = akra_bazzi_relation_tac ctxt THEN_ALL_NEW akra_bazzi_measure_tac ctxt (* General preprocessing *) fun nat_numeral_conv ctxt ct = let val t = Thm.term_of ct val prop = Logic.mk_equals (t, num_to_Suc (dest_nat_number t)) - fun tac {context = ctxt, prems = _ } = - SOLVE (Local_Defs.unfold_tac ctxt @{thms eval_nat_numeral BitM.simps One_nat_def} - THEN resolve_tac ctxt @{thms Pure.reflexive} 1); + fun tac goal_ctxt = + SOLVE (Local_Defs.unfold_tac goal_ctxt @{thms eval_nat_numeral BitM.simps One_nat_def} + THEN resolve_tac goal_ctxt @{thms Pure.reflexive} 1); in - case try (fn () => Goal.prove ctxt [] [] prop tac) () of + case \<^try>\Goal.prove ctxt [] [] prop (tac o #context)\ of NONE => raise CTERM ("nat_numeral_conv", [ct]) | SOME thm => thm end fun akra_bazzi_sum_conv ctxt ct = case Thm.term_of ct of (Const (@{const_name sum}, _) $ _ $ (Const (@{const_name Set_Interval.ord_class.lessThan}, _) $ _)) => Conv.arg_conv (Conv.arg_conv (nat_numeral_conv ctxt)) ct | _ => raise CTERM ("akra_bazzi_sum_conv", [ct]) fun akra_bazzi_sum_tac ctxt i = let fun conv ctxt = Conv.try_conv (akra_bazzi_sum_conv ctxt) val thms = @{thms eval_akra_bazzi_sum eval_akra_bazzi_sum' mult_1_left mult_1_right add_0_left add_0_right} in CHANGED (CONVERSION (Conv.top_conv conv ctxt) i THEN REPEAT (EqSubst.eqsubst_tac ctxt [0] thms i)) end fun akra_bazzi_term_tac ctxt = rewrite_tac ctxt term_preps THEN' (resolve_tac ctxt (get_term_intros ctxt)) fun akra_bazzi_terms_tac ctxt = REPEAT_ALL_NEW (fn i => DETERM (resolve_tac ctxt @{thms akra_bazzi_termsI'} i)) THEN_ALL_NEW akra_bazzi_term_tac ctxt fun analyze_terminal ctxt f x inv t_orig t acc = let val thy = Proof_Context.theory_of ctxt val patvar = ("x",1) val pat = Term.betapply (f, Var (patvar, HOLogic.natT)) fun match t = let val matcher = Pattern.first_order_match thy (pat, t) (Vartab.empty, Vartab.empty) val x' = Vartab.lookup (snd matcher) patvar |> the |> snd in if Term.exists_subterm (fn x' => x = x') x' then x' else raise Pattern.MATCH end fun add_term t (acc, acc_inv, call) = if inv then (acc, t :: acc_inv, call) else (t :: acc, acc_inv, call) fun go t (acc, acc_inv, call) = let val call' = match t in case call of NONE => (acc, acc_inv, SOME call') | SOME _ => raise TERM ("Non-linear occurrence of recursive call", [t_orig]) end handle Pattern.MATCH => add_term t (acc, acc_inv, call) in go t acc end; fun mk_prod [] = @{term "1 :: real"} | mk_prod (x::xs) = fold (fn x => fn acc => @{term "(*) :: real => real => real"} $ x $ acc) xs x fun mk_quot (xs, []) = mk_prod xs | mk_quot ([], ys) = @{term "inverse :: real => real"} $ mk_prod ys | mk_quot (xs, ys) = @{term "(/) :: real => real => real"} $ mk_prod xs $ mk_prod ys fun analyze_prod ctxt f x minus t (acc1, acc2) = let fun go inv (oper $ t1 $ t2) acc = if oper = @{term "(*) :: real => real => real"} then go inv t1 (go inv t2 acc) else if oper = @{term "(/) :: real => real => real"} then go inv t1 (go (not inv) t2 acc) else analyze_terminal ctxt f x inv t (oper $ t1 $ t2) acc | go inv (oper $ t') acc = if oper = @{term "inverse :: real => real"} then go (not inv) t' acc else analyze_terminal ctxt f x inv t (oper $ t') acc | go inv t' acc = analyze_terminal ctxt f x inv t t' acc in case go false t ([], [], NONE) of (_, _, NONE) => (acc1, (if minus then @{term "uminus :: real => real"} $ t else t) :: acc2) | (xs, ys, SOME call) => let val r = mk_quot (xs, ys) val r' = (if minus then @{term "uminus :: real => real"} $ r else r, call) in (r' :: acc1, acc2) end end fun mk_sum [] = @{term "0 :: real"} | mk_sum (x::xs) = let fun go (f $ x) acc = if f = @{term "uminus :: real => real"} then @{term "(-) :: real => real => real"} $ acc $ x else @{term "(+) :: real => real => real"} $ acc $ (f $ x) | go x acc = @{term "(+) :: real => real => real"} $ acc $ x in fold go xs x end fun analyze_sum ctxt f x t = let fun go minus (oper $ t1 $ t2) acc = if oper = @{term "(+) :: real => real => real"} then go minus t1 (go minus t2 acc) else if oper = @{term "(-) :: real => real => real"} then go minus t1 (go (not minus) t2 acc) else analyze_prod ctxt f x minus (oper $ t1 $ t2) acc | go minus (oper $ t) acc = if oper = @{term "uminus :: real => real"} then go (not minus) t acc else analyze_prod ctxt f x minus (oper $ t) acc | go minus t acc = analyze_prod ctxt f x minus t acc in case go false t ([],[]) of (xs,ys) => (xs, mk_sum ys) end fun extract_coeff ctxt t = let val thy = Proof_Context.theory_of ctxt val t' = Raw_Simplifier.rewrite_term thy @{thms akra_bazzi_termination_simps[THEN eq_reflection]} [] t fun aux thm = let val (tmp, pat) = thm |> Thm.concl_of |> HOLogic.dest_Trueprop |> Term.dest_comb val var = tmp |> Term.dest_comb |> snd |> Term.dest_Var |> fst val (_, insts) = Pattern.first_order_match thy (pat, t') (Vartab.empty, Vartab.empty) val coeff = Option.map snd (Vartab.lookup insts var) in coeff end handle Pattern.MATCH => NONE in case get_first aux (get_term_intros ctxt) of NONE => raise TERM ("extract_coeff", [t']) | SOME coeff => coeff end fun get_var_name (Free (x, _)) = x | get_var_name (Var ((x,_),_)) = x | get_var_name t = raise TERM ("get_var_name", [t]) fun analyze_call ctxt f x t = let val (ts, g) = analyze_sum ctxt f x t val k = length ts val a_coeffs = HOLogic.mk_list HOLogic.realT (map fst ts) fun abstract t = Abs (get_var_name x, HOLogic.natT, abstract_over (x, t)) val ts = map (abstract o snd) ts val b_coeffs = HOLogic.mk_list HOLogic.realT (map (extract_coeff ctxt) ts) val ts = HOLogic.mk_list (HOLogic.natT --> HOLogic.natT) ts in {k = HOLogic.mk_nat k, a_coeffs = a_coeffs, b_coeffs = b_coeffs, ts = ts, g = abstract g} end fun instantiate_akra_bazzi ctxt ts = map (Thm.instantiate' [] (map (Option.map (Thm.cterm_of ctxt)) ts)) fun analyze_equation ctxt f eq funvar = let val rhs = eq |> Thm.concl_of |> HOLogic.dest_Trueprop |> HOLogic.dest_eq |> snd val {k, a_coeffs, b_coeffs, ts, g} = analyze_call ctxt f funvar rhs in {k = k, a_coeffs = a_coeffs, b_coeffs = b_coeffs, ts = ts, f = f, g = g} end fun akra_bazzi_instantiate_tac ctxt f eq x0 x1 funvar p' = let val {k, a_coeffs, b_coeffs, ts, g, ...} = analyze_equation ctxt f eq funvar val insts = ([SOME x0, SOME x1, SOME k, SOME a_coeffs, SOME b_coeffs, SOME ts, SOME f, SOME g] @ (filter Option.isSome [p'])) in resolve_tac ctxt o instantiate_akra_bazzi ctxt insts end (* val code_simp_tac = Code_Simp.static_tac {consts = [@{const_name arith_consts}], ctxt = @{context}, simpset = NONE} fun eval_real_numeral_conv ctxt ct = case Thm.term_of ct of @{term "Pure.imp"} $ _ $ _ => Conv.implies_concl_conv (eval_real_numeral_conv ctxt) ct | _ => let val t = Thm.term_of ct val prop = Logic.mk_equals (t, @{term "Trueprop True"}) - val ctxt' = put_simpset (simpset_of @{context}) ctxt - fun tac ({context = ctxt, ...}) = - Local_Defs.unfold_tac ctxt @{thms greaterThanLessThan_iff} - THEN (SOLVE (Simplifier.full_simp_tac ctxt' 1) ORELSE code_simp_tac ctxt 1) + fun tac goal_ctxt = + let val ctxt' = put_simpset (simpset_of @{context}) goal_ctxt in + Local_Defs.unfold_tac ctxt' @{thms greaterThanLessThan_iff} + THEN (SOLVE (Simplifier.full_simp_tac ctxt' 1) ORELSE code_simp_tac ctxt' 1) + end in if contains_only_numerals t then - case try (fn () => Goal.prove ctxt [] [] prop tac) () of + case \<^try>\Goal.prove ctxt [] [] prop (tac o #context)\ of SOME thm => thm | NONE => raise CTERM ("eval_real_numeral_conv", [ct]) else raise CTERM ("eval_real_numeral_conv", [ct]) end *) fun eval_numeral_tac ctxt = SUBGOAL (fn (goal, i) => if contains_only_numerals (Logic.strip_imp_concl goal) then Eval_Numeral.eval_numeral_tac ctxt i else all_tac) fun unfold_simps ctxt = rewrite_tac ctxt @{thms eval_akra_bazzi_le_list_ex eval_length of_nat_numeral of_nat_0 of_nat_1 numeral_One} (* Applying the locale intro rule *) fun master_theorem_function_tac simp ctxt = resolve_tac ctxt @{thms master_theorem_functionI} THEN_ALL_NEW ((fn i => TRY (REPEAT_ALL_NEW (match_tac ctxt @{thms ball_set_intros}) i)) THEN_ALL_NEW ((fn i => TRY (akra_bazzi_terms_tac ctxt i)) THEN_ALL_NEW (fn i => unfold_simps ctxt i THEN TRY (akra_bazzi_sum_tac ctxt i) THEN unfold_simps ctxt i THEN (if simp then TRY (eval_numeral_tac ctxt i) else all_tac) ) ) ) (* Pre-processing of Landau bound in goal and post-processing of Landau bounds in premises *) fun goal_conv conv ct = case Thm.term_of ct of (Const (@{const_name "Pure.imp"}, _) $ _ $ _) => Conv.arg_conv (goal_conv conv) ct | _ => conv ct fun rewrs_conv' thms = goal_conv (Conv.arg_conv (Conv.arg_conv (Conv.rewrs_conv thms))) fun preprocess_tac ctxt = CONVERSION (rewrs_conv' @{thms CLAMP}) THEN' rewrite_tac ctxt @{thms propagate_CLAMP CLAMP_aux} THEN' rewrite_tac ctxt @{thms CLAMP_postproc} THEN' CONVERSION (rewrs_conv' @{thms UNCLAMP'}) fun rewrs_goal_conv thms = goal_conv (Conv.arg_conv (Conv.arg_conv (Conv.rewrs_conv thms))) fun rewrs_conv thms ctxt = Conv.repeat_conv (changed_conv (Conv.bottom_conv (K (Conv.try_conv (Conv.rewrs_conv thms))) ctxt)) fun postprocess_tac ctxt = TRY o ( CONVERSION (rewrs_goal_conv @{thms CLAMP}) THEN' rewrite_tac ctxt @{thms CLAMP_aux} THEN' rewrite_tac ctxt @{thms CLAMP_postproc} THEN' CONVERSION (rewrs_conv @{thms MASTER_BOUND_postproc[THEN eq_reflection]} ctxt) THEN' CONVERSION (rewrs_conv @{thms MASTER_BOUND_UNCLAMP[THEN eq_reflection]} ctxt) THEN' CONVERSION (rewrs_goal_conv @{thms UNCLAMP})) (* The main tactic *) fun intros_tac ctxt thms = DETERM o TRY o REPEAT_ALL_NEW (DETERM o match_tac ctxt thms) fun master_theorem_tac' caseno simp f eq funvar x0 x1 p' ctxt = ( preprocess_tac ctxt THEN' akra_bazzi_instantiate_tac ctxt f eq x0 x1 funvar p' (get_intros caseno (p' <> NONE)) THEN' DETERM o master_theorem_function_tac simp ctxt) THEN_ALL_NEW intros_tac ctxt @{thms ball_set_intros akra_bazzi_p_rel_intros} THEN_ALL_NEW (DETERM o ( unfold_simps ctxt THEN' TRY o akra_bazzi_sum_tac ctxt THEN' unfold_simps ctxt THEN' postprocess_tac ctxt THEN' (if simp then TRY o eval_numeral_tac ctxt else K all_tac) THEN' intros_tac ctxt @{thms allI ballI conjI impI} THEN_ALL_NEW (TRY o eresolve_tac ctxt @{thms atLeastLessThanE}))) fun find_function goal = case goal |> Logic.strip_imp_concl |> HOLogic.dest_Trueprop of @{term "Set.member :: (nat => real) => _ set => bool"} $ f $ _ => f | _ => raise TERM ("find_function", [goal]) val leq_nat = @{term "(<=) :: nat => nat => bool"} val less_nat = @{term "(<) :: nat => nat => bool"} fun find_x1 var eq = let fun get_x1 prem = case prem |> HOLogic.dest_Trueprop of rel $ x1 $ x => if x = var andalso rel = leq_nat then SOME x1 else if x = var andalso rel = less_nat then SOME (@{term "Suc"} $ x1) else NONE | _ => NONE in get_first get_x1 (Thm.prems_of eq) end handle TERM _ => NONE fun analyze_funeq pat var ctxt eq = let val lhs = eq |> Thm.concl_of |> HOLogic.dest_Trueprop |> HOLogic.dest_eq |> fst val matcher = Thm.first_order_match (Thm.cterm_of ctxt lhs, pat) val eq = Thm.instantiate matcher eq in case find_x1 var eq of NONE => NONE | SOME _ => SOME eq end handle Pattern.MATCH => NONE fun find_funeq ctxt f = let val n = ("n", 0) val n' = Var (n, HOLogic.natT) val pat = Thm.cterm_of ctxt (Term.betapply (f, n')) fun get_function_info (Abs (_, _, body)) = get_function_info body | get_function_info (t as f $ _) = (Function.get_info ctxt t handle Empty => get_function_info f) | get_function_info t = Function.get_info ctxt t val {simps = simps, ...} = get_function_info f handle List.Empty => raise TERM ("Could not obtain function info record.", [f]) val simps = case simps of SOME simps => simps | NONE => raise TERM ("Function has no simp rules.", [f]) in case get_first (analyze_funeq pat n' ctxt) simps of SOME s => (s, n') | NONE => raise TERM ("Could not determine recurrence from simp rules.", [f]) end fun master_theorem_tac caseno simp eq x0 x1 p' ctxt = SUBGOAL (fn (goal, i) => Subgoal.FOCUS (fn {context = ctxt, ...} => let val f = find_function goal val (eq, funvar) = case eq of SOME eq => ( let val n = Var (("n", 0), HOLogic.natT) in case (analyze_funeq (Thm.cterm_of ctxt (f $ n)) n ctxt eq) of SOME eq => (eq, n) | NONE => raise TERM ("", []) end handle TERM _ => raise TERM ( "Could not determine recursion variable from recurrence theorem.", [Thm.prop_of eq])) | NONE => find_funeq ctxt f val x0 = Option.getOpt (x0, @{term "0::nat"}) val x1 = if Option.isSome x1 then x1 else find_x1 funvar eq val x1 = case x1 of SOME x1 => x1 | NONE => raise TERM ("Could not determine x1 from recurrence theorem", [goal, Thm.prop_of eq, funvar]) val p' = if caseno = NONE orelse caseno = SOME "1" orelse caseno = SOME "2.1" then p' else NONE in HEADGOAL (master_theorem_tac' caseno simp f eq funvar x0 x1 p' ctxt) end ) ctxt i) handle TERM _ => K no_tac fun read_term_with_type ctxt T = Syntax.check_term ctxt o Type.constraint T o Syntax.parse_term ctxt fun parse_assoc key scan = Scan.lift (Args.$$$ key -- Args.colon) -- scan >> snd val parse_param = let fun term_with_type T = Scan.peek (fn context => Args.embedded_inner_syntax >> read_term_with_type (Context.proof_of context) T); in parse_assoc "recursion" Attrib.thm >> (fn x => (SOME x, NONE, NONE, NONE)) || parse_assoc "x0" (term_with_type HOLogic.natT) >> (fn x => (NONE, SOME x, NONE, NONE)) || parse_assoc "x1" (term_with_type HOLogic.natT) >> (fn x => (NONE, NONE, SOME x, NONE)) || parse_assoc "p'" (term_with_type HOLogic.realT) >> (fn x => (NONE, NONE, NONE, SOME x)) end val parse_params = let fun add_option NONE y = y | add_option x _ = x fun go (a,b,c,d) (a',b',c',d') = (add_option a a', add_option b b', add_option c c', add_option d d') in Scan.repeat parse_param >> (fn xs => fold go xs (NONE, NONE, NONE, NONE)) end val setup_master_theorem = Scan.option (Scan.lift (Parse.number || Parse.float_number)) -- Scan.lift (Args.parens (Args.$$$ "nosimp") >> K false || Scan.succeed true) -- parse_params >> (fn ((caseno, simp), (thm, x0, x1, p')) => fn ctxt => let val _ = case caseno of SOME caseno => if member (op =) ["1","2.1","2.2","2.3","3"] caseno then () else cat_error "illegal Master theorem case: " caseno | NONE => () in SIMPLE_METHOD' (master_theorem_tac caseno simp thm x0 x1 p' ctxt) end) 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,1377 @@ (* 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; + fun tac goal_ctxt = + HEADGOAL (CONVERSION unfold_comb_conv THEN' + Raw_Simplifier.rewrite_goal_tac goal_ctxt red_thms THEN' + resolve_tac goal_ctxt [@{thm refl}]); + in + singleton (Variable.export ctxt' ctxt) (Goal.prove ctxt' [] [] lifted_prop (tac o #context)) + 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 (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; + fun tac goal_ctxt = + HEADGOAL (Raw_Simplifier.rewrite_goal_tac goal_ctxt [B_intro, merge_rule] THEN' + resolve_tac goal_ctxt [@{thm refl}]); + in + singleton (Variable.export ctxt' ctxt) (Goal.prove ctxt' [] [] prop (tac o #context)) + 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); + fun tac goal_ctxt = + HEADGOAL (Raw_Simplifier.rewrite_goal_tac goal_ctxt [Thm.symmetric merge_thm] THEN' + resolve_tac goal_ctxt [T_red]); + val weak_red = + singleton (Variable.export ctxt' ctxt) (Goal.prove ctxt' [] [] weak_prop (tac o #context)); 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 tac goal_ctxt = + HEADGOAL (normalize_wrapper_tac goal_ctxt (SOME af) THEN' + Raw_Simplifier.rewrite_goal_tac goal_ctxt comb_unfolds THEN' + resolve_tac goal_ctxt [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); + val hol_thm = + singleton (Variable.export ctxt' ctxt) (Goal.prove ctxt' [] [] prop (tac o #context)); 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/Automated_Stateful_Protocol_Verification/trac/trac.thy b/thys/Automated_Stateful_Protocol_Verification/trac/trac.thy --- a/thys/Automated_Stateful_Protocol_Verification/trac/trac.thy +++ b/thys/Automated_Stateful_Protocol_Verification/trac/trac.thy @@ -1,1947 +1,1947 @@ (* (C) Copyright Andreas Viktor Hess, DTU, 2020 (C) Copyright Sebastian A. Mödersheim, DTU, 2020 (C) Copyright Achim D. Brucker, University of Exeter, 2020 (C) Copyright Anders Schlichtkrull, DTU, 2020 All Rights Reserved. Redistribution and use in source and binary forms, with or without modification, are permitted provided that the following conditions are met: - Redistributions of source code must retain the above copyright notice, this list of conditions and the following disclaimer. - Redistributions in binary form must reproduce the above copyright notice, this list of conditions and the following disclaimer in the documentation and/or other materials provided with the distribution. - Neither the name of the copyright holder nor the names of its contributors may be used to endorse or promote products derived from this software without specific prior written permission. THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS "AS IS" AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT LIMITED TO, THE IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR A PARTICULAR PURPOSE ARE DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT OWNER OR CONTRIBUTORS BE LIABLE FOR ANY DIRECT, INDIRECT, INCIDENTAL, SPECIAL, EXEMPLARY, OR CONSEQUENTIAL DAMAGES (INCLUDING, BUT NOT LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR SERVICES; LOSS OF USE, DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY, OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE. *) (* Title: trac.thy Author: Andreas Viktor Hess, DTU Author: Sebastian A. Mödersheim, DTU Author: Achim D. Brucker, University of Exeter Author: Anders Schlichtkrull, DTU *) section\Support for the Trac Format\ theory "trac" imports trac_fp_parser trac_protocol_parser keywords "trac" :: thy_decl and "trac_import" :: thy_decl and "trac_trac" :: thy_decl and "trac_import_trac" :: thy_decl and "protocol_model_setup" :: thy_decl and "protocol_security_proof" :: thy_decl and "manual_protocol_model_setup" :: thy_decl and "manual_protocol_security_proof" :: thy_decl and "compute_fixpoint" :: thy_decl and "compute_SMP" :: thy_decl and "setup_protocol_model'" :: thy_decl and "protocol_security_proof'" :: thy_decl and "setup_protocol_checks" :: thy_decl begin ML \ (* Some of this is based on code from the following files distributed with Isabelle 2018: * HOL/Tools/value_command.ML * HOL/Code_Evaluation.thy * Pure.thy *) fun protocol_model_interpretation_defs name = let fun f s = (Binding.empty_atts:Attrib.binding, ((Binding.name s, NoSyn), name ^ "." ^ s)) in (map f [ "public", "arity", "Ana", "\", "\\<^sub>v", "timpls_transformable_to", "intruder_synth_mod_timpls", "analyzed_closed_mod_timpls", "timpls_transformable_to'", "intruder_synth_mod_timpls'", "analyzed_closed_mod_timpls'", "admissible_transaction_terms", "admissible_transaction", "abs_substs_set", "abs_substs_fun", "in_trancl", "transaction_poschecks_comp", "transaction_negchecks_comp", "transaction_check_comp", "transaction_check", "transaction_check_pre", "transaction_check_post", "compute_fixpoint_fun'", "compute_fixpoint_fun", "attack_notin_fixpoint", "protocol_covered_by_fixpoint", "analyzed_fixpoint", "wellformed_protocol'", "wellformed_protocol", "wellformed_fixpoint", "wellformed_composable_protocols", "composable_protocols" ]):string Interpretation.defines end fun protocol_model_interpretation_params name = let fun f s = name ^ "_" ^ s in map SOME [f "arity", "\_. 0", f "public", f "Ana", f "\", "0::nat", "1::nat"] end fun declare_thm_attr attribute name print lthy = let val arg = [(Facts.named name, [[Token.make_string (attribute, Position.none)]])] val (_, lthy') = Specification.theorems_cmd "" [(Binding.empty_atts, arg)] [] print lthy in lthy' end fun declare_def_attr attribute name = declare_thm_attr attribute (name ^ "_def") val declare_code_eqn = declare_def_attr "code" val declare_protocol_check = declare_def_attr "protocol_checks" fun declare_protocol_checks print = declare_protocol_check "attack_notin_fixpoint" print #> declare_protocol_check "protocol_covered_by_fixpoint" print #> declare_protocol_check "analyzed_fixpoint" print #> declare_protocol_check "wellformed_protocol'" print #> declare_protocol_check "wellformed_protocol" print #> declare_protocol_check "wellformed_fixpoint" print #> declare_protocol_check "compute_fixpoint_fun" print fun eval_define (name, raw_t) lthy = let val t = Code_Evaluation.dynamic_value_strict lthy (Syntax.read_term lthy raw_t) val arg = ((Binding.name name, NoSyn), ((Binding.name (name ^ "_def"),[]), t)) val (_, lthy') = Local_Theory.define arg lthy in (t, lthy') end fun eval_define_declare (name, raw_t) print = eval_define (name, raw_t) ##> declare_code_eqn name print val _ = Outer_Syntax.local_theory' @{command_keyword "compute_fixpoint"} "evaluate and define protocol fixpoint" (Parse.name -- Parse.name >> (fn (protocol, fixpoint) => fn print => snd o eval_define_declare (fixpoint, "compute_fixpoint_fun " ^ protocol) print)); val _ = Outer_Syntax.local_theory' @{command_keyword "compute_SMP"} "evaluate and define a finite representation of the sub-message patterns of a protocol" ((Scan.optional (\<^keyword>\[\ |-- Parse.name --| \<^keyword>\]\) "no_optimizations") -- Parse.name -- Parse.name >> (fn ((opt,protocol), smp) => fn print => let val rmd = "List.remdups" val f = "Stateful_Strands.trms_list\<^sub>s\<^sub>s\<^sub>t" val g = "(\T. " ^ f ^ " T@map (pair' prot_fun.Pair) (Stateful_Strands.setops_list\<^sub>s\<^sub>s\<^sub>t T))" fun s trms = "(" ^ rmd ^ " (List.concat (List.map (" ^ trms ^ " \ Labeled_Strands.unlabel \ transaction_strand) " ^ protocol ^ ")))" val opt1 = "remove_superfluous_terms \" val opt2 = "generalize_terms \ is_Var" val gsmp_opt = "generalize_terms \ (\t. is_Var t \ t \ TAtom AttackType \ " ^ "t \ TAtom SetType \ t \ TAtom OccursSecType \ \is_Atom (the_Var t))" val smp_fun = "SMP0 Ana \" fun smp_fun' opts = "(\T. let T' = (" ^ rmd ^ " \ " ^ opts ^ " \ " ^ smp_fun ^ ") T in List.map (\t. t \ Typed_Model.var_rename (Typed_Model.max_var_set " ^ "(Messages.fv\<^sub>s\<^sub>e\<^sub>t (set (T@T'))))) T')" val cmd = if opt = "no_optimizations" then smp_fun ^ " " ^ s f else if opt = "optimized" then smp_fun' (opt1 ^ " \ " ^ opt2) ^ " " ^ s f else if opt = "GSMP" then smp_fun' (opt1 ^ " \ " ^ gsmp_opt) ^ " " ^ s g else error ("Invalid option: " ^ opt) in snd o eval_define_declare (smp, cmd) print end)); val _ = Outer_Syntax.local_theory' @{command_keyword "setup_protocol_checks"} "setup protocol checks" (Parse.name -- Parse.name >> (fn (protocol_model, protocol_name) => fn print => let val a1 = "coverage_check_intro_lemmata" val a2 = "coverage_check_unfold_lemmata" val a3 = "coverage_check_unfold_protocol_lemma" in declare_protocol_checks print #> declare_thm_attr a1 (protocol_model ^ ".protocol_covered_by_fixpoint_intros") print #> declare_def_attr a2 (protocol_model ^ ".protocol_covered_by_fixpoint") print #> declare_def_attr a3 protocol_name print end )); val _ = Outer_Syntax.local_theory_to_proof \<^command_keyword>\setup_protocol_model'\ "prove interpretation of protocol model locale into global theory" (Parse.!!! (Parse.name -- Parse_Spec.locale_expression) >> (fn (prefix,expr) => fn lthy => let fun f x y z = ([(x,(y,(Expression.Positional z,[])))],[]) val (a,(b,c)) = nth (fst expr) 0 val name = fst b val _ = case c of (Expression.Named [],[]) => () | _ => error "Invalid arguments" val pexpr = f a b (protocol_model_interpretation_params prefix) val pdefs = protocol_model_interpretation_defs name in if name = "" then error "No name given" else Interpretation.global_interpretation_cmd pexpr pdefs lthy end)); val _ = Outer_Syntax.local_theory_to_proof' \<^command_keyword>\protocol_security_proof'\ "prove interpretation of secure protocol locale into global theory" (Parse.!!! (Parse.name -- Parse_Spec.locale_expression) >> (fn (prefix,expr) => fn print => let fun f x y z = ([(x,(y,(Expression.Positional z,[])))],[]) val (a,(b,c)) = nth (fst expr) 0 val d = case c of (Expression.Positional ps,[]) => ps | _ => error "Invalid arguments" val pexpr = f a b (protocol_model_interpretation_params prefix@d) in declare_protocol_checks print #> Interpretation.global_interpretation_cmd pexpr [] end )); \ ML\ structure ml_isar_wrapper = struct fun define_constant_definition (constname, trm) lthy = let val arg = ((Binding.name constname, NoSyn), ((Binding.name (constname^"_def"),[]), trm)) val ((_, (_ , thm)), lthy') = Local_Theory.define arg lthy in (thm, lthy') end fun define_constant_definition' (constname, trm) print lthy = let val arg = ((Binding.name constname, NoSyn), ((Binding.name (constname^"_def"),[]), trm)) val ((_, (_ , thm)), lthy') = Local_Theory.define arg lthy val lthy'' = declare_code_eqn constname print lthy' in (thm, lthy'') end fun define_simple_abbrev (constname, trm) lthy = let val arg = ((Binding.name constname, NoSyn), trm) val ((_, _), lthy') = Local_Theory.abbrev Syntax.mode_default arg lthy in lthy' end fun define_simple_type_synonym (name, typedecl) lthy = let val (_, lthy') = Typedecl.abbrev_global (Binding.name name, [], NoSyn) typedecl lthy in lthy' end fun define_simple_datatype (dt_tyargs, dt_name) constructors = let val options = Plugin_Name.default_filter fun lift_c (tyargs, name) = (((Binding.empty, Binding.name name), map (fn t => (Binding.empty, t)) tyargs), NoSyn) val c_spec = map lift_c constructors val datatyp = ((map (fn ty => (NONE, ty)) dt_tyargs, Binding.name dt_name), NoSyn) val dtspec = ((options,false), [(((datatyp, c_spec), (Binding.empty, Binding.empty, Binding.empty)), [])]) in BNF_FP_Def_Sugar.co_datatypes BNF_Util.Least_FP BNF_LFP.construct_lfp dtspec end fun define_simple_primrec pname precs lthy = let val rec_eqs = map (fn (lhs,rhs) => (((Binding.empty,[]), HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs,rhs))),[],[])) precs in snd (BNF_LFP_Rec_Sugar.primrec false [] [(Binding.name pname, NONE, NoSyn)] rec_eqs lthy) end fun define_simple_fun pname precs lthy = let val rec_eqs = map (fn (lhs,rhs) => (((Binding.empty,[]), HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs,rhs))),[],[])) precs in Function_Fun.add_fun [(Binding.name pname, NONE, NoSyn)] rec_eqs Function_Common.default_config lthy end fun prove_simple name stmt tactic lthy = let - val thm = Goal.prove lthy [] [] stmt (fn {context, ...} => tactic context) + val thm = Goal.prove lthy [] [] stmt (tactic o #context) |> Goal.norm_result lthy |> Goal.check_finished lthy in lthy |> snd o Local_Theory.note ((Binding.name name, []), [thm]) end fun prove_state_simple method proof_state = Seq.the_result "error in proof state" ( (Proof.refine method proof_state)) |> Proof.global_done_proof end \ ML\ structure trac_definitorial_package = struct (* constant names *) open Trac_Utils val enum_constsN="enum_consts" val setsN="sets" val funN="fun" val atomN="atom" val arityN="arity" val publicN = "public" val gammaN = "\" val anaN = "Ana" val valN = "val" val timpliesN = "timplies" val occursN = "occurs" val enumN = "enum" val priv_fun_secN = "PrivFunSec" val secret_typeN = "SecretType" val enum_typeN = "EnumType" val other_pubconsts_typeN = "PubConstType" val types = [enum_typeN, secret_typeN] val special_funs = ["occurs", "zero", valN, priv_fun_secN] fun mk_listT T = Type ("List.list", [T]) val mk_setT = HOLogic.mk_setT val boolT = HOLogic.boolT val natT = HOLogic.natT val mk_tupleT = HOLogic.mk_tupleT val mk_prodT = HOLogic.mk_prodT val mk_set = HOLogic.mk_set val mk_list = HOLogic.mk_list val mk_nat = HOLogic.mk_nat val mk_eq = HOLogic.mk_eq val mk_Trueprop = HOLogic.mk_Trueprop val mk_tuple = HOLogic.mk_tuple val mk_prod = HOLogic.mk_prod fun mkN (a,b) = a^"_"^b val info = Output.information fun rm_special_funs sel l = list_minus (list_rm_pair sel) l special_funs fun is_priv_fun (trac:TracProtocol.protocol) f = let val funs = #private (Option.valOf (#function_spec trac)) in (* not (List.find (fn g => fst g = f) funs = NONE) *) List.exists (fn (g,n) => f = g andalso n <> "0") funs end fun full_name name lthy = Local_Theory.full_name lthy (Binding.name name) fun full_name' n (trac:TracProtocol.protocol) lthy = full_name (mkN (#name trac, n)) lthy fun mk_prot_type name targs (trac:TracProtocol.protocol) lthy = Term.Type (full_name' name trac lthy, targs) val enum_constsT = mk_prot_type enum_constsN [] fun mk_enum_const a trac lthy = Term.Const (full_name' enum_constsN trac lthy ^ "." ^ a, enum_constsT trac lthy) val databaseT = mk_prot_type setsN [] val funT = mk_prot_type funN [] val atomT = mk_prot_type atomN [] fun messageT (trac:TracProtocol.protocol) lthy = Term.Type ("Transactions.prot_term", [funT trac lthy, atomT trac lthy, databaseT trac lthy]) fun message_funT (trac:TracProtocol.protocol) lthy = Term.Type ("Transactions.prot_fun", [funT trac lthy, atomT trac lthy, databaseT trac lthy]) fun message_varT (trac:TracProtocol.protocol) lthy = Term.Type ("Transactions.prot_var", [funT trac lthy, atomT trac lthy, databaseT trac lthy]) fun message_term_typeT (trc:TracProtocol.protocol) lthy = Term.Type ("Transactions.prot_term_type", [funT trc lthy, atomT trc lthy, databaseT trc lthy]) fun message_atomT (trac:TracProtocol.protocol) lthy = Term.Type ("Transactions.prot_atom", [atomT trac lthy]) fun messageT' varT (trac:TracProtocol.protocol) lthy = Term.Type ("Term.term", [message_funT trac lthy, varT]) fun message_listT (trac:TracProtocol.protocol) lthy = mk_listT (messageT trac lthy) fun message_listT' varT (trac:TracProtocol.protocol) lthy = mk_listT (messageT' varT trac lthy) fun absT (trac:TracProtocol.protocol) lthy = mk_setT (databaseT trac lthy) fun abssT (trac:TracProtocol.protocol) lthy = mk_setT (absT trac lthy) val poscheckvariantT = Term.Type ("Strands_and_Constraints.poscheckvariant", []) val strand_labelT = Term.Type ("Labeled_Strands.strand_label", [natT]) fun strand_stepT (trac:TracProtocol.protocol) lthy = Term.Type ("Stateful_Strands.stateful_strand_step", [message_funT trac lthy, message_varT trac lthy]) fun labeled_strand_stepT (trac:TracProtocol.protocol) lthy = mk_prodT (strand_labelT, strand_stepT trac lthy) fun prot_strandT (trac:TracProtocol.protocol) lthy = mk_listT (labeled_strand_stepT trac lthy) fun prot_transactionT (trac:TracProtocol.protocol) lthy = Term.Type ("Transactions.prot_transaction", [funT trac lthy, atomT trac lthy, databaseT trac lthy, natT]) val mk_star_label = Term.Const ("Labeled_Strands.strand_label.LabelS", strand_labelT) fun mk_prot_label (lbl:int) = Term.Const ("Labeled_Strands.strand_label.LabelN", natT --> strand_labelT) $ mk_nat lbl fun mk_labeled_step (label:term) (step:term) = mk_prod (label, step) fun mk_Send_step (trac:TracProtocol.protocol) lthy (label:term) (msg:term) = mk_labeled_step label (Term.Const ("Stateful_Strands.stateful_strand_step.Send", messageT trac lthy --> strand_stepT trac lthy) $ msg) fun mk_Receive_step (trac:TracProtocol.protocol) lthy (label:term) (msg:term) = mk_labeled_step label (Term.Const ("Stateful_Strands.stateful_strand_step.Receive", messageT trac lthy --> strand_stepT trac lthy) $ msg) fun mk_InSet_step (trac:TracProtocol.protocol) lthy (label:term) (elem:term) (set:term) = let val psT = [poscheckvariantT, messageT trac lthy, messageT trac lthy] in mk_labeled_step label (Term.Const ("Stateful_Strands.stateful_strand_step.InSet", psT ---> strand_stepT trac lthy) $ Term.Const ("Strands_and_Constraints.poscheckvariant.Check", poscheckvariantT) $ elem $ set) end fun mk_NotInSet_step (trac:TracProtocol.protocol) lthy (label:term) (elem:term) (set:term) = let val varT = message_varT trac lthy val trm_prodT = mk_prodT (messageT trac lthy, messageT trac lthy) val psT = [mk_listT varT, mk_listT trm_prodT, mk_listT trm_prodT] in mk_labeled_step label (Term.Const ("Stateful_Strands.stateful_strand_step.NegChecks", psT ---> strand_stepT trac lthy) $ mk_list varT [] $ mk_list trm_prodT [] $ mk_list trm_prodT [mk_prod (elem,set)]) end fun mk_Inequality_step (trac:TracProtocol.protocol) lthy (label:term) (t1:term) (t2:term) = let val varT = message_varT trac lthy val trm_prodT = mk_prodT (messageT trac lthy, messageT trac lthy) val psT = [mk_listT varT, mk_listT trm_prodT, mk_listT trm_prodT] in mk_labeled_step label (Term.Const ("Stateful_Strands.stateful_strand_step.NegChecks", psT ---> strand_stepT trac lthy) $ mk_list varT [] $ mk_list trm_prodT [mk_prod (t1,t2)] $ mk_list trm_prodT []) end fun mk_Insert_step (trac:TracProtocol.protocol) lthy (label:term) (elem:term) (set:term) = mk_labeled_step label (Term.Const ("Stateful_Strands.stateful_strand_step.Insert", [messageT trac lthy, messageT trac lthy] ---> strand_stepT trac lthy) $ elem $ set) fun mk_Delete_step (trac:TracProtocol.protocol) lthy (label:term) (elem:term) (set:term) = mk_labeled_step label (Term.Const ("Stateful_Strands.stateful_strand_step.Delete", [messageT trac lthy, messageT trac lthy] ---> strand_stepT trac lthy) $ elem $ set) fun mk_Transaction (trac:TracProtocol.protocol) lthy S1 S2 S3 S4 S5 S6 = let val varT = message_varT trac lthy val msgT = messageT trac lthy val var_listT = mk_listT varT val msg_listT = mk_listT msgT val trT = prot_transactionT trac lthy (* val decl_elemT = mk_prodT (varT, mk_listT msgT) val declT = mk_listT decl_elemT *) val stepT = labeled_strand_stepT trac lthy val strandT = prot_strandT trac lthy val strandsT = mk_listT strandT val paramsT = [(* declT, *)var_listT, strandT, strandT, strandT, strandT, strandT] in Term.Const ("Transactions.prot_transaction.Transaction", paramsT ---> trT) $ (* mk_list decl_elemT [] $ *) (if null S4 then mk_list varT [] else (Term.Const (@{const_name "map"}, [msgT --> varT, msg_listT] ---> var_listT) $ Term.Const (@{const_name "the_Var"}, msgT --> varT) $ mk_list msgT S4)) $ mk_list stepT S1 $ mk_list stepT [] $ (if null S3 then mk_list stepT S2 else (Term.Const (@{const_name "append"}, [strandT,strandT] ---> strandT) $ mk_list stepT S2 $ (Term.Const (@{const_name "concat"}, strandsT --> strandT) $ mk_list strandT S3))) $ mk_list stepT S5 $ mk_list stepT S6 end fun get_funs (trac:TracProtocol.protocol) = let fun append_sec fs = fs@[(priv_fun_secN, "0")] val filter_funs = filter (fn (_,n) => n <> "0") val filter_consts = filter (fn (_,n) => n = "0") fun inc_ar (s,n) = (s, Int.toString (1+Option.valOf (Int.fromString n))) in case (#function_spec trac) of NONE => ([],[],[]) | SOME ({public=pub, private=priv}) => let val pub_symbols = rm_special_funs fst (pub@map inc_ar (filter_funs priv)) val pub_funs = filter_funs pub_symbols val pub_consts = filter_consts pub_symbols val priv_consts = append_sec (rm_special_funs fst (filter_consts priv)) in (pub_funs, pub_consts, priv_consts) end end fun get_set_spec (trac:TracProtocol.protocol) = mk_unique (map (fn (s,n) => (s,Option.valOf (Int.fromString n))) (#set_spec trac)) fun set_arity (trac:TracProtocol.protocol) s = case List.find (fn x => fst x = s) (get_set_spec trac) of SOME (_,n) => SOME n | NONE => NONE fun get_enums (trac:TracProtocol.protocol) = mk_unique (TracProtocol.extract_Consts (#type_spec trac)) fun flatten_type_spec (trac:TracProtocol.protocol) = let fun find_type taus tau = case List.find (fn x => fst x = tau) taus of SOME x => snd x | NONE => error ("Type " ^ tau ^ " has not been declared") fun step taus (s,e) = case e of TracProtocol.Union ts => let val es = map (find_type taus) ts fun f es' = mk_unique (List.concat (map TracProtocol.the_Consts es')) in if List.all TracProtocol.is_Consts es then (s,TracProtocol.Consts (f es)) else (s,TracProtocol.Union ts) end | c => (s,c) fun loop taus = let val taus' = map (step taus) taus in if taus = taus' then taus else loop taus' end val flat_type_spec = let val x = loop (#type_spec trac) val errpre = "Couldn't flatten the enumeration types: " in if List.all (fn (_,e) => TracProtocol.is_Consts e) x then let val y = map (fn (s,e) => (s,TracProtocol.the_Consts e)) x in if List.all (not o List.null o snd) y then y else error (errpre ^ "does every type have at least one value?") end else error (errpre ^ "have all types been declared?") end in flat_type_spec end fun is_attack_transaction (tr:TracProtocol.cTransaction) = not (null (#attack_actions tr)) fun get_transaction_name (tr:TracProtocol.cTransaction) = #1 (#transaction tr) fun get_fresh_value_variables (tr:TracProtocol.cTransaction) = map_filter (TracProtocol.maybe_the_Fresh o snd) (#fresh_actions tr) fun get_nonfresh_value_variables (tr:TracProtocol.cTransaction) = map fst (filter (fn x => snd x = "value") (#2 (#transaction tr))) fun get_value_variables (tr:TracProtocol.cTransaction) = get_nonfresh_value_variables tr@get_fresh_value_variables tr fun get_enum_variables (tr:TracProtocol.cTransaction) = mk_unique (filter (fn x => snd x <> "value") (#2 (#transaction tr))) fun get_variable_restrictions (tr:TracProtocol.cTransaction) = let val enum_vars = get_enum_variables tr val value_vars = get_value_variables tr fun enum_member x = List.exists (fn y => x = fst y) fun value_member x = List.exists (fn y => x = y) fun aux [] = ([],[]) | aux ((a,b)::rs) = if enum_member a enum_vars andalso enum_member b enum_vars then let val (es,vs) = aux rs in ((a,b)::es,vs) end else if value_member a value_vars andalso value_member b value_vars then let val (es,vs) = aux rs in (es,(a,b)::vs) end else error ("Ill-formed or ill-typed variable restriction: " ^ a ^ " != " ^ b) in aux (#3 (#transaction tr)) end fun conv_enum_consts trac (t:Trac_Term.cMsg) = let open Trac_Term val enums = get_enums trac fun aux (cFun (f,ts)) = if List.exists (fn x => x = f) enums then if null ts then cEnum f else error ("Enum constant " ^ f ^ " should not have a parameter list") else cFun (f,map aux ts) | aux (cConst c) = if List.exists (fn x => x = c) enums then cEnum c else cConst c | aux (cSet (s,ts)) = cSet (s,map aux ts) | aux (cOccursFact bs) = cOccursFact (aux bs) | aux t = t in aux t end fun val_to_abs_list vs = let open Trac_Term fun aux t = case t of cEnum b => b | _ => error "Invalid val parameter list" in case vs of [] => [] | (cConst "0"::ts) => val_to_abs_list ts | (cFun (s,ps)::ts) => (s, map aux ps)::val_to_abs_list ts | (cSet (s,ps)::ts) => (s, map aux ps)::val_to_abs_list ts | _ => error "Invalid val parameter list" end fun val_to_abs (t:Trac_Term.cMsg) = let open Trac_Term fun aux t = case t of cEnum b => b | _ => error "Invalid val parameter list" fun val_to_abs_list [] = [] | val_to_abs_list (cConst "0"::ts) = val_to_abs_list ts | val_to_abs_list (cFun (s,ps)::ts) = (s, map aux ps)::val_to_abs_list ts | val_to_abs_list (cSet (s,ps)::ts) = (s, map aux ps)::val_to_abs_list ts | val_to_abs_list _ = error "Invalid val parameter list" in case t of cFun (f,ts) => if f = valN then cAbs (val_to_abs_list ts) else cFun (f,map val_to_abs ts) | cSet (s,ts) => cSet (s,map val_to_abs ts) | cOccursFact bs => cOccursFact (val_to_abs bs) | t => t end fun occurs_enc t = let open Trac_Term fun aux [cVar x] = cVar x | aux [cAbs bs] = cAbs bs | aux _ = error "Invalid occurs parameter list" fun enc (cFun (f,ts)) = ( if f = occursN then cOccursFact (aux ts) else cFun (f,map enc ts)) | enc (cSet (s,ts)) = cSet (s,map enc ts) | enc (cOccursFact bs) = cOccursFact (enc bs) | enc t = t in enc t end fun priv_fun_enc trac (Trac_Term.cFun (f,ts)) = ( if is_priv_fun trac f andalso (case ts of Trac_Term.cPrivFunSec::_ => false | _ => true) then Trac_Term.cFun (f,Trac_Term.cPrivFunSec::map (priv_fun_enc trac) ts) else Trac_Term.cFun (f,map (priv_fun_enc trac) ts)) | priv_fun_enc _ t = t fun transform_cMsg trac = priv_fun_enc trac o occurs_enc o val_to_abs o conv_enum_consts trac fun check_no_vars_and_consts (fp:Trac_Term.cMsg list) = let open Trac_Term fun aux (cVar _) = false | aux (cConst _) = false | aux (cFun (_,ts)) = List.all aux ts | aux (cSet (_,ts)) = List.all aux ts | aux (cOccursFact bs) = aux bs | aux _ = true in if List.all aux fp then fp else error "There shouldn't be any cVars and cConsts at this point in the fixpoint translation" end fun split_fp (fp:Trac_Term.cMsg list) = let open Trac_Term fun fa t = case t of cFun (s,_) => s <> timpliesN | _ => true fun fb (t,ts) = case t of cOccursFact (cAbs bs) => bs::ts | _ => ts fun fc (cFun (s, [cAbs bs, cAbs cs]),ts) = if s = timpliesN then (bs,cs)::ts else ts | fc (_,ts) = ts val eq = eq_set (fn ((s,xs),(t,ys)) => s = t andalso eq_set (op =) (xs,ys)) fun eq_pairs ((a,b),(c,d)) = eq (a,c) andalso eq (b,d) val timplies_trancl = let fun trans_step ts = let fun aux (s,t) = map (fn (_,u) => (s,u)) (filter (fn (v,_) => eq (t,v)) ts) in distinct eq_pairs (filter (not o eq) (ts@List.concat (map aux ts))) end fun loop ts = let val ts' = trans_step ts in if eq_set eq_pairs (ts,ts') then ts else loop ts' end in loop end val ti = List.foldl fc [] fp in (filter fa fp, distinct eq (List.foldl fb [] fp@map snd ti), timplies_trancl ti) end fun mk_enum_substs trac (vars:(string * Trac_Term.VarType) list) = let open Trac_Term val flat_type_spec = flatten_type_spec trac val deltas = let fun f (s,EnumType tau) = ( case List.find (fn x => fst x = tau) flat_type_spec of SOME x => map (fn c => (s,c)) (snd x) | NONE => error ("Type " ^ tau ^ " was not found in the type specification")) | f (s,_) = error ("Variable " ^ s ^ " is not of enum type") in list_product (map f vars) end in map (fn d => map (fn (x,t) => (x,cEnum t)) d) deltas end fun ground_enum_variables trac (fp:Trac_Term.cMsg list) = let open Trac_Term fun do_grounding t = map (fn d => subst_apply d t) (mk_enum_substs trac (fv_cMsg t)) in List.concat (map do_grounding fp) end fun transform_fp trac (fp:Trac_Term.cMsg list) = fp |> ground_enum_variables trac |> map (transform_cMsg trac) |> check_no_vars_and_consts |> split_fp fun database_to_hol (db:string * Trac_Term.cMsg list) (trac:TracProtocol.protocol) lthy = let open Trac_Term val errmsg = "Invalid database parameter" fun mkN' n = mkN (#name trac, n) val s_prefix = full_name (mkN' setsN) lthy ^ "." val e_prefix = full_name (mkN' enum_constsN) lthy ^ "." val (s,es) = db val tau = enum_constsT trac lthy val databaseT = databaseT trac lthy val a = Term.Const (s_prefix ^ s, map (fn _ => tau) es ---> databaseT) fun param_to_hol (cVar (x,EnumType _)) = Term.Free (x, tau) | param_to_hol (cVar (x,Untyped)) = Term.Free (x, tau) | param_to_hol (cEnum e) = Term.Const (e_prefix ^ e, tau) | param_to_hol (cConst c) = error (errmsg ^ ": cConst " ^ c) | param_to_hol (cVar (x,ValueType)) = error (errmsg ^ ": cVar (" ^ x ^ ",ValueType)") | param_to_hol _ = error errmsg in fold (fn e => fn b => b $ param_to_hol e) es a end fun abs_to_hol (bs:(string * string list) list) (trac:TracProtocol.protocol) lthy = let val databaseT = databaseT trac lthy fun db_params_to_cEnum (a,cs) = (a, map Trac_Term.cEnum cs) in mk_set databaseT (map (fn db => database_to_hol (db_params_to_cEnum db) trac lthy) bs) end fun cMsg_to_hol (t:Trac_Term.cMsg) lbl varT var_map free_enum_var trac lthy = let open Trac_Term val tT = messageT' varT trac lthy val fT = message_funT trac lthy val enum_constsT = enum_constsT trac lthy val tsT = message_listT' varT trac lthy val VarT = varT --> tT val FunT = [fT, tsT] ---> tT val absT = absT trac lthy val databaseT = databaseT trac lthy val AbsT = absT --> fT val funT = funT trac lthy val FuT = funT --> fT val SetT = databaseT --> fT val enumT = enum_constsT --> funT val VarC = Term.Const (@{const_name "Var"}, VarT) val FunC = Term.Const (@{const_name "Fun"}, FunT) val NilC = Term.Const (@{const_name "Nil"}, tsT) val prot_label = mk_nat lbl fun full_name'' n = full_name' n trac lthy fun mk_enum_const' a = mk_enum_const a trac lthy fun mk_prot_fun_trm f tau = Term.Const ("Transactions.prot_fun." ^ f, tau) fun mk_enum_trm etrm = mk_prot_fun_trm "Fu" FuT $ (Term.Const (full_name'' funN ^ "." ^ enumN, enumT) $ etrm) fun mk_Fu_trm f = mk_prot_fun_trm "Fu" FuT $ Term.Const (full_name'' funN ^ "." ^ f, funT) fun c_to_h s = cMsg_to_hol s lbl varT var_map free_enum_var trac lthy fun c_list_to_h ts = mk_list tT (map c_to_h ts) in case t of cVar x => if free_enum_var x then FunC $ mk_enum_trm (Term.Free (fst x, enum_constsT)) $ NilC else VarC $ var_map x | cConst f => FunC $ mk_Fu_trm f $ NilC | cFun (f,ts) => FunC $ mk_Fu_trm f $ c_list_to_h ts | cSet (s,ts) => FunC $ (mk_prot_fun_trm "Set" SetT $ database_to_hol (s,ts) trac lthy) $ NilC | cAttack => FunC $ (mk_prot_fun_trm "Attack" (natT --> fT) $ prot_label) $ NilC | cAbs bs => FunC $ (mk_prot_fun_trm "Abs" AbsT $ abs_to_hol bs trac lthy) $ NilC | cOccursFact bs => FunC $ mk_prot_fun_trm "OccursFact" fT $ mk_list tT [ FunC $ mk_prot_fun_trm "OccursSec" fT $ NilC, c_to_h bs] | cPrivFunSec => FunC $ mk_Fu_trm priv_fun_secN $ NilC | cEnum a => FunC $ mk_enum_trm (mk_enum_const' a) $ NilC end fun ground_cMsg_to_hol t lbl trac lthy = cMsg_to_hol t lbl (message_varT trac lthy) (fn _ => error "Term not ground") (fn _ => false) trac lthy fun ana_cMsg_to_hol inc_vars t (ana_var_map:string list) = let open Trac_Term fun var_map (x,Untyped) = ( case list_find (fn y => x = y) ana_var_map of SOME (_,n) => if inc_vars then mk_nat (1+n) else mk_nat n | NONE => error ("Analysis variable " ^ x ^ " not found")) | var_map _ = error "Analysis variables must be untyped" val lbl = 0 (* There's no constants in analysis messages requiring labels anyway *) in cMsg_to_hol t lbl natT var_map (fn _ => false) end fun transaction_cMsg_to_hol t lbl (transaction_var_map:string list) trac lthy = let open Trac_Term val varT = message_varT trac lthy val atomT = message_atomT trac lthy val term_typeT = message_term_typeT trac lthy fun TAtom_Value_var n = let val a = Term.Const (@{const_name "Var"}, atomT --> term_typeT) $ Term.Const ("Transactions.prot_atom.Value", atomT) in HOLogic.mk_prod (a, mk_nat n) end fun var_map_err_prefix x = "Transaction variable " ^ x ^ " should be value typed but is actually " fun var_map (x,ValueType) = ( case list_find (fn y => x = y) transaction_var_map of SOME (_,n) => TAtom_Value_var n | NONE => error ("Transaction variable " ^ x ^ " not found")) | var_map (x,EnumType e) = error (var_map_err_prefix x ^ "of enum type " ^ e) | var_map (x,Untyped) = error (var_map_err_prefix x ^ "untyped") in cMsg_to_hol t lbl varT var_map (fn (_,t) => case t of EnumType _ => true | _ => false) trac lthy end fun fp_triple_to_hol (fp,occ,ti) trac lthy = let val prot_label = 0 val tau_abs = absT trac lthy val tau_fp_elem = messageT trac lthy val tau_occ_elem = tau_abs val tau_ti_elem = mk_prodT (tau_abs, tau_abs) fun a_to_h bs = abs_to_hol bs trac lthy fun c_to_h t = ground_cMsg_to_hol t prot_label trac lthy val fp' = mk_list tau_fp_elem (map c_to_h fp) val occ' = mk_list tau_occ_elem (map a_to_h occ) val ti' = mk_list tau_ti_elem (map (mk_prod o map_prod a_to_h) ti) in mk_tuple [fp', occ', ti'] end fun abstract_over_enum_vars enum_vars enum_ineqs trm flat_type_spec trac lthy = let val enum_constsT = enum_constsT trac lthy fun enumlistelemT n = mk_tupleT (replicate n enum_constsT) fun enumlistT n = mk_listT (enumlistelemT n) fun mk_enum_const' a = mk_enum_const a trac lthy fun absfreeprod xs trm = let val tau = enum_constsT val tau_out = Term.fastype_of trm fun absfree' x = absfree (x,enum_constsT) fun aux _ [] = trm | aux _ [x] = absfree' x trm | aux len (x::y::xs) = Term.Const (@{const_name "case_prod"}, [[tau,mk_tupleT (replicate (len-1) tau)] ---> tau_out, mk_tupleT (replicate len tau)] ---> tau_out) $ absfree' x (aux (len-1) (y::xs)) in aux (length xs) xs end fun mk_enum_neq (a,b) = (HOLogic.mk_not o HOLogic.mk_eq) (Term.Free (a, enum_constsT), Term.Free (b, enum_constsT)) fun mk_enum_neqs_list [] = Term.Const (@{const_name "True"}, HOLogic.boolT) | mk_enum_neqs_list [x] = mk_enum_neq x | mk_enum_neqs_list (x::y::xs) = HOLogic.mk_conj (mk_enum_neq x, mk_enum_neqs_list (y::xs)) val enum_types = let fun aux t = if t = "" then get_enums trac else case List.find (fn (s,_) => t = s) flat_type_spec of SOME (_,cs) => cs | NONE => error ("Not an enum type: " ^ t ^ "?") in map (aux o snd) enum_vars end val enumlist_product = let fun mk_enumlist ns = mk_list enum_constsT (map mk_enum_const' ns) fun aux _ [] = mk_enumlist [] | aux _ [ns] = mk_enumlist ns | aux len (ns::ms::elists) = Term.Const ("List.product", [enumlistT 1, enumlistT (len-1)] ---> enumlistT len) $ mk_enumlist ns $ aux (len-1) (ms::elists) in aux (length enum_types) enum_types end val absfp = absfreeprod (map fst enum_vars) trm val eptrm = enumlist_product val typof = Term.fastype_of val evseT = enumlistelemT (length enum_vars) val evslT = enumlistT (length enum_vars) val eneqs = absfreeprod (map fst enum_vars) (mk_enum_neqs_list enum_ineqs) in if null enum_vars then mk_list (typof trm) [trm] else if null enum_ineqs then Term.Const(@{const_name "map"}, [typof absfp, typof eptrm] ---> mk_listT (typof trm)) $ absfp $ eptrm else Term.Const(@{const_name "map"}, [typof absfp, typof eptrm] ---> mk_listT (typof trm)) $ absfp $ (Term.Const(@{const_name "filter"}, [evseT --> HOLogic.boolT, evslT] ---> evslT) $ eneqs $ eptrm) end fun mk_type_of_name lthy pname name ty_args = Type(Local_Theory.full_name lthy (Binding.name (mkN(pname, name))), ty_args) fun mk_mt_list t = Term.Const (@{const_name "Nil"}, mk_listT t) fun name_of_typ (Type (s, _)) = s | name_of_typ (TFree _) = error "name_of_type: unexpected TFree" | name_of_typ (TVar _ ) = error "name_of_type: unexpected TVAR" fun prove_UNIV name typ elems thmsN lthy = let val rhs = mk_set typ elems val lhs = Const("Set.UNIV",mk_setT typ) val stmt = mk_Trueprop (mk_eq (lhs,rhs)) val fq_tname = name_of_typ typ fun inst_and_prove_enum thy = let val _ = writeln("Inst enum: "^name) val lthy = Class.instantiation ([fq_tname], [], @{sort enum}) thy val enum_eq = Const("Pure.eq",mk_listT typ --> mk_listT typ --> propT) $Const(@{const_name "enum_class.enum"},mk_listT typ) $(mk_list typ elems) val ((_, (_, enum_def')), lthy) = Specification.definition NONE [] [] ((Binding.name ("enum_"^name),[]), enum_eq) lthy val ctxt_thy = Proof_Context.init_global (Proof_Context.theory_of lthy) val enum_def = singleton (Proof_Context.export lthy ctxt_thy) enum_def' val enum_all_eq = Const("Pure.eq", boolT --> boolT --> propT) $(Const(@{const_name "enum_class.enum_all"},(typ --> boolT) --> boolT) $Free("P",typ --> boolT)) $(Const(@{const_name "list_all"},(typ --> boolT) --> (mk_listT typ) --> boolT) $Free("P",typ --> boolT)$(mk_list typ elems)) val ((_, (_, enum_all_def')), lthy) = Specification.definition NONE [] [] ((Binding.name ("enum_all_"^name),[]), enum_all_eq) lthy val ctxt_thy = Proof_Context.init_global (Proof_Context.theory_of lthy) val enum_all_def = singleton (Proof_Context.export lthy ctxt_thy) enum_all_def' val enum_ex_eq = Const("Pure.eq", boolT --> boolT --> propT) $(Const(@{const_name "enum_class.enum_ex"},(typ --> boolT) --> boolT) $Free("P",typ --> boolT)) $(Const(@{const_name "list_ex"},(typ --> boolT) --> (mk_listT typ) --> boolT) $Free("P",typ --> boolT)$(mk_list typ elems)) val ((_, (_, enum_ex_def')), lthy) = Specification.definition NONE [] [] ((Binding.name ("enum_ex_"^name),[]), enum_ex_eq) lthy val ctxt_thy = Proof_Context.init_global (Proof_Context.theory_of lthy) val enum_ex_def = singleton (Proof_Context.export lthy ctxt_thy) enum_ex_def' in Class.prove_instantiation_exit (fn ctxt => (Class.intro_classes_tac ctxt []) THEN ALLGOALS (simp_tac (ctxt addsimps [Proof_Context.get_thm ctxt (name^"_UNIV"), enum_def, enum_all_def, enum_ex_def]) ) )lthy end fun inst_and_prove_finite thy = let val lthy = Class.instantiation ([fq_tname], [], @{sort finite}) thy in Class.prove_instantiation_exit (fn ctxt => (Class.intro_classes_tac ctxt []) THEN (simp_tac (ctxt addsimps[Proof_Context.get_thm ctxt (name^"_UNIV")])) 1) lthy end in lthy |> ml_isar_wrapper.prove_simple (name^"_UNIV") stmt (fn c => (safe_tac c) THEN (ALLGOALS(simp_tac c)) THEN (ALLGOALS(Metis_Tactic.metis_tac ["full_types"] "combs" c (map (Proof_Context.get_thm c) thmsN))) ) |> Local_Theory.raw_theory inst_and_prove_finite |> Local_Theory.raw_theory inst_and_prove_enum end fun def_types (trac:TracProtocol.protocol) lthy = let val pname = #name trac val defname = mkN(pname, enum_constsN) val _ = info(" Defining "^defname) val tnames = get_enums trac val types = map (fn x => ([],x)) tnames in ([defname], ml_isar_wrapper.define_simple_datatype ([], defname) types lthy) end fun def_sets (trac:TracProtocol.protocol) lthy = let val pname = #name trac val defname = mkN(pname, setsN) val _ = info (" Defining "^defname) val sspec = get_set_spec trac val tfqn = Local_Theory.full_name lthy (Binding.name (mkN(pname, enum_constsN))) val ttyp = Type(tfqn, []) val types = map (fn (x,n) => (replicate n ttyp,x)) sspec in lthy |> ml_isar_wrapper.define_simple_datatype ([], defname) types end fun def_funs (trac:TracProtocol.protocol) lthy = let val pname = #name trac val (pub_f, pub_c, priv) = get_funs trac val pub = pub_f@pub_c fun def_atom lthy = let val def_atomname = mkN(pname, atomN) val types = if null pub_c then types else types@[other_pubconsts_typeN] fun define_atom_dt lthy = let val _ = info(" Defining "^def_atomname) in lthy |> ml_isar_wrapper.define_simple_datatype ([], def_atomname) (map (fn x => ([],x)) types) end fun prove_UNIV_atom lthy = let val _ = info (" Proving "^def_atomname^"_UNIV") val thmsN = [def_atomname^".exhaust"] val fqn = Local_Theory.full_name lthy (Binding.name (mkN(pname, atomN))) val typ = Type(fqn, []) in lthy |> prove_UNIV (def_atomname) typ (map (fn c => Const(fqn^"."^c,typ)) types) thmsN end in lthy |> define_atom_dt |> prove_UNIV_atom end fun def_fun_dt lthy = let val def_funname = mkN(pname, funN) val _ = info(" Defining "^def_funname) val types = map (fn x => ([],x)) (map fst (pub@priv)) val ctyp = Type(Local_Theory.full_name lthy (Binding.name (mkN(pname, enum_constsN))), []) in ml_isar_wrapper.define_simple_datatype ([], def_funname) (types@[([ctyp],enumN)]) lthy end fun def_fun_arity lthy = let val fqn_name = Local_Theory.full_name lthy (Binding.name (mkN(pname, funN))) val ctyp = Type(fqn_name, []) fun mk_rec_eq name (fname,arity) = (Free(name,ctyp --> natT) $Const(fqn_name^"."^fname,ctyp), mk_nat((Option.valOf o Int.fromString) arity)) val name = mkN(pname, arityN) val _ = info(" Defining "^name) val ctyp' = Type(Local_Theory.full_name lthy (Binding.name (mkN(pname, enum_constsN))), []) in ml_isar_wrapper.define_simple_fun name ((map (mk_rec_eq name) (pub@priv))@[ (Free(name, ctyp --> natT) $(Const(fqn_name^"."^enumN, ctyp' --> ctyp)$(Term.dummy_pattern ctyp')), mk_nat(0))]) lthy end fun def_public lthy = let val fqn_name = Local_Theory.full_name lthy (Binding.name (mkN(pname, funN))) val ctyp = Type(fqn_name, []) fun mk_rec_eq name t fname = (Free(name, ctyp --> boolT) $Const(fqn_name^"."^fname,ctyp), t) val name = mkN(pname, publicN) val _ = info(" Defining "^name) val ctyp' = Type(Local_Theory.full_name lthy (Binding.name (mkN(pname, enum_constsN))), []) in ml_isar_wrapper.define_simple_fun name ((map (mk_rec_eq name (@{term "False"})) (map fst priv)) @(map (mk_rec_eq name (@{term "True"})) (map fst pub)) @[(Free(name, ctyp --> boolT) $(Const(fqn_name^"."^enumN, ctyp' --> ctyp)$(Term.dummy_pattern ctyp')), @{term "True"})]) lthy end fun def_gamma lthy = let fun optionT t = Type (@{type_name "option"}, [t]) fun mk_Some t = Const (@{const_name "Some"}, t --> optionT t) fun mk_None t = Const (@{const_name "None"}, optionT t) val fqn_name = Local_Theory.full_name lthy (Binding.name (mkN(pname, funN))) val ctyp = Type(fqn_name, []) val atomFQN = Local_Theory.full_name lthy (Binding.name (mkN(pname, atomN))) val atomT = Type(atomFQN, []) fun mk_rec_eq name t fname = (Free(name, ctyp --> optionT atomT) $Const(fqn_name^"."^fname,ctyp), t) val name = mkN(pname, gammaN) val _ = info(" Defining "^name) val ctyp' = Type(Local_Theory.full_name lthy (Binding.name (mkN(pname, enum_constsN))), []) in ml_isar_wrapper.define_simple_fun name ((map (mk_rec_eq name ((mk_Some atomT)$(Const(atomFQN^"."^secret_typeN, atomT)))) (map fst priv)) @(map (mk_rec_eq name ((mk_Some atomT)$(Const(atomFQN^"."^other_pubconsts_typeN, atomT)))) (map fst pub_c)) @[(Free(name, ctyp --> optionT atomT) $(Const(fqn_name^"."^enumN, ctyp' --> ctyp)$(Term.dummy_pattern ctyp')), (mk_Some atomT)$(Const(atomFQN^"."^enum_typeN,atomT)))] @(map (mk_rec_eq name (mk_None atomT)) (map fst pub_f)) ) lthy end fun def_ana lthy = let val pname = #name trac val (pub_f, pub_c, priv) = get_funs trac val pub = pub_f@pub_c val keyT = messageT' natT trac lthy val fqn_name = Local_Theory.full_name lthy (Binding.name (mkN(pname, funN))) val ctyp = Type(fqn_name, []) val ana_outputT = mk_prodT (mk_listT keyT, mk_listT natT) val default_output = mk_prod (mk_list keyT [], mk_list natT []) fun mk_ana_output ks rs = mk_prod (mk_list keyT ks, mk_list natT rs) fun mk_rec_eq name t fname = (Free(name, ctyp --> ana_outputT) $Term.Const(fqn_name^"."^fname,ctyp), t) val name = mkN(pname, anaN) val _ = info(" Defining "^name) val ctyp' = Type(Local_Theory.full_name lthy (Binding.name (mkN(pname, enum_constsN))), []) val ana_spec = let val toInt = Option.valOf o Int.fromString fun ana_arity (f,n) = (if is_priv_fun trac f then (toInt n)-1 else toInt n) fun check_valid_arity ((f,ps),ks,rs) = case List.find (fn g => f = fst g) pub_f of SOME (f',n) => if length ps <> ana_arity (f',n) then error ("Invalid number of parameters in the analysis rule for " ^ f ^ " (expected " ^ Int.toString (ana_arity (f',n)) ^ " but got " ^ Int.toString (length ps) ^ ")") else ((f,ps),ks,rs) | NONE => error (f ^ " is not a declared function symbol of arity greater than zero") val transform_cMsg = transform_cMsg trac val rm_special_funs = rm_special_funs (fn ((f,_),_,_) => f) fun var_to_nat f xs x = let val n = snd (Option.valOf ((list_find (fn y => y = x) xs))) in if is_priv_fun trac f then mk_nat (1+n) else mk_nat n end fun c_to_h f xs t = ana_cMsg_to_hol (is_priv_fun trac f) t xs trac lthy fun keys f ps ks = map (c_to_h f ps o transform_cMsg o Trac_Term.certifyMsg [] []) ks fun results f ps rs = map (var_to_nat f ps) rs fun aux ((f,ps),ks,rs) = (f, mk_ana_output (keys f ps ks) (results f ps rs)) in map (aux o check_valid_arity) (rm_special_funs (#analysis_spec trac)) end val other_funs = filter (fn f => not (List.exists (fn g => f = g) (map fst ana_spec))) (map fst (pub@priv)) in ml_isar_wrapper.define_simple_fun name ((map (fn (f,out) => mk_rec_eq name out f) ana_spec) @(map (mk_rec_eq name default_output) other_funs) @[(Free(name, ctyp --> ana_outputT) $(Term.Const(fqn_name^"."^enumN, ctyp' --> ctyp)$(Term.dummy_pattern ctyp')), default_output)]) lthy end in lthy |> def_atom |> def_fun_dt |> def_fun_arity |> def_public |> def_gamma |> def_ana end fun define_term_model (trac:TracProtocol.protocol) lthy = let val _ = info("Defining term model") in lthy |> snd o def_types trac |> def_sets trac |> def_funs trac end fun define_fixpoint fp trac print lthy = let val fp_name = mkN (#name trac, "fixpoint") val _ = info("Defining fixpoint") val _ = info(" Defining "^fp_name) val fp_triple = transform_fp trac fp val fp_triple_trm = fp_triple_to_hol fp_triple trac lthy val trac = TracProtocol.update_fixed_point trac (SOME fp_triple) in (trac, #2 (ml_isar_wrapper.define_constant_definition' (fp_name, fp_triple_trm) print lthy)) end fun define_protocol print ((trac:TracProtocol.protocol), lthy) = let val _ = if length (#transaction_spec trac) > 1 then info("Defining protocols") else info("Defining protocol") val pname = #name trac val flat_type_spec = flatten_type_spec trac val mk_Transaction = mk_Transaction trac lthy val mk_Send = mk_Send_step trac lthy val mk_Receive = mk_Receive_step trac lthy val mk_InSet = mk_InSet_step trac lthy val mk_NotInSet = mk_NotInSet_step trac lthy val mk_Inequality = mk_Inequality_step trac lthy val mk_Insert = mk_Insert_step trac lthy val mk_Delete = mk_Delete_step trac lthy val star_label = mk_star_label val prot_label = mk_prot_label val certify_transation = TracProtocol.certifyTransaction fun mk_tname i (tr:TracProtocol.transaction_name) = let val x = #1 tr val y = case i of NONE => x | SOME n => mkN(n, x) val z = mkN("transaction", y) in mkN(pname, z) end fun def_transaction name_prefix prot_num (transaction:TracProtocol.cTransaction) lthy = let val defname = mk_tname name_prefix (#transaction transaction) val _ = info(" Defining "^defname) val receives = #receive_actions transaction val checkssingle = #checksingle_actions transaction val checksall = #checkall_actions transaction val updates = #update_actions transaction val sends = #send_actions transaction val fresh = get_fresh_value_variables transaction val attack_signals = #attack_actions transaction val nonfresh_value_vars = get_nonfresh_value_variables transaction val value_vars = get_value_variables transaction val enum_vars = get_enum_variables transaction val (enum_ineqs, value_ineqs) = get_variable_restrictions transaction val transform_cMsg = transform_cMsg trac fun c_to_h trm = transaction_cMsg_to_hol (transform_cMsg trm) prot_num value_vars trac lthy val abstract_over_enum_vars = fn x => fn y => fn z => abstract_over_enum_vars x y z flat_type_spec trac lthy fun mk_transaction_term (rcvs, chcksingle, chckall, upds, snds, frsh, atcks) = let open Trac_Term fun action_filter f (lbl,a) = case f a of SOME x => SOME (lbl,x) | NONE => NONE fun lbl_to_h (TracProtocol.LabelS) = star_label | lbl_to_h (TracProtocol.LabelN) = prot_label prot_num fun lbl_trm_to_h f (lbl,t) = f (lbl_to_h lbl) (c_to_h t) val S1 = map (lbl_trm_to_h mk_Receive) (map_filter (action_filter TracProtocol.maybe_the_Receive) rcvs) val S2 = let fun aux (lbl,TracProtocol.cInequality (x,y)) = SOME (mk_Inequality (lbl_to_h lbl) (c_to_h x) (c_to_h y)) | aux (lbl,TracProtocol.cInSet (e,s)) = SOME (mk_InSet (lbl_to_h lbl) (c_to_h e) (c_to_h s)) | aux (lbl,TracProtocol.cNotInSet (e,s)) = SOME (mk_NotInSet (lbl_to_h lbl) (c_to_h e) (c_to_h s)) | aux _ = NONE in map_filter aux chcksingle end val S3 = let fun arity s = case set_arity trac s of SOME n => n | NONE => error ("Not a set family: " ^ s) fun mk_evs s = map (fn n => ("X" ^ Int.toString n, "")) (0 upto ((arity s) -1)) fun mk_trm (lbl,e,s) = let val ps = map (fn x => cVar (x,Untyped)) (map fst (mk_evs s)) in mk_NotInSet (lbl_to_h lbl) (c_to_h e) (c_to_h (cSet (s,ps))) end fun mk_trms (lbl,(e,s)) = abstract_over_enum_vars (mk_evs s) [] (mk_trm (lbl,e,s)) in map mk_trms (map_filter (action_filter TracProtocol.maybe_the_NotInAny) chckall) end val S4 = map (c_to_h o mk_Value_cVar) frsh val S5 = let fun aux (lbl,TracProtocol.cInsert (e,s)) = SOME (mk_Insert (lbl_to_h lbl) (c_to_h e) (c_to_h s)) | aux (lbl,TracProtocol.cDelete (e,s)) = SOME (mk_Delete (lbl_to_h lbl) (c_to_h e) (c_to_h s)) | aux _ = NONE in map_filter aux upds end val S6 = let val snds' = map_filter (action_filter TracProtocol.maybe_the_Send) snds in map (lbl_trm_to_h mk_Send) (snds'@map (fn (lbl,_) => (lbl,cAttack)) atcks) end in abstract_over_enum_vars enum_vars enum_ineqs (mk_Transaction S1 S2 S3 S4 S5 S6) end fun def_trm trm print lthy = #2 (ml_isar_wrapper.define_constant_definition' (defname, trm) print lthy) val additional_value_ineqs = let open Trac_Term open TracProtocol val poschecks = map_filter (maybe_the_InSet o snd) checkssingle val negchecks_single = map_filter (maybe_the_NotInSet o snd) checkssingle val negchecks_all = map_filter (maybe_the_NotInAny o snd) checksall fun aux' (cVar (x,ValueType),s) (cVar (y,ValueType),t) = if s = t then SOME (x,y) else NONE | aux' _ _ = NONE fun aux (x,cSet (s,ps)) = SOME ( map_filter (aux' (x,cSet (s,ps))) negchecks_single@ map_filter (aux' (x,s)) negchecks_all ) | aux _ = NONE in List.concat (map_filter aux poschecks) end val all_value_ineqs = mk_unique (value_ineqs@additional_value_ineqs) val valvarsprod = filter (fn p => not (List.exists (fn q => p = q orelse swap p = q) all_value_ineqs)) (list_triangle_product (fn x => fn y => (x,y)) nonfresh_value_vars) val transaction_trm0 = mk_transaction_term (receives, checkssingle, checksall, updates, sends, fresh, attack_signals) in if null valvarsprod then def_trm transaction_trm0 print lthy else let val partitions = list_partitions nonfresh_value_vars all_value_ineqs val ps = filter (not o null) (map (filter (fn x => length x > 1)) partitions) fun mk_subst ps = let open Trac_Term fun aux [] = NONE | aux (x::xs) = SOME (map (fn y => (y,cVar (x,ValueType))) xs) in List.concat (map_filter aux ps) end fun apply d = let val ap = TracProtocol.subst_apply_actions d fun f (TracProtocol.cInequality (x,y)) = x <> y | f _ = true val checksingle' = filter (f o snd) (ap checkssingle) in (ap receives, checksingle', ap checksall, ap updates, ap sends, fresh, attack_signals) end val transaction_trms = transaction_trm0::map (mk_transaction_term o apply o mk_subst) ps val transaction_typ = Term.fastype_of transaction_trm0 fun mk_concat_trm tau trms = Term.Const (@{const_name "concat"}, mk_listT tau --> tau) $ mk_list tau trms in def_trm (mk_concat_trm transaction_typ transaction_trms) print lthy end end val def_transactions = let val prots = map (fn (n,pr) => map (fn tr => (n,tr)) pr) (#transaction_spec trac) val lbls = list_upto (length prots) val lbl_prots = List.concat (map (fn i => map (fn tr => (i,tr)) (nth prots i)) lbls) val f = fold (fn (i,(n,tr)) => def_transaction n i (certify_transation tr)) in f lbl_prots end fun def_protocols lthy = let fun mk_prot_def (name,trm) lthy = let val _ = info(" Defining "^name) in #2 (ml_isar_wrapper.define_constant_definition' (name,trm) print lthy) end val prots = #transaction_spec trac val num_prots = length prots val pdefname = mkN(pname, "protocol") fun mk_tnames i = let val trs = case nth prots i of (j,prot) => map (fn tr => (j,tr)) prot in map (fn (j,s) => full_name (mk_tname j (#transaction s)) lthy) trs end val tnames = List.concat (map mk_tnames (list_upto num_prots)) val pnames = let val f = fn i => (Int.toString i,nth prots i) val g = fn (i,(n,_)) => case n of NONE => i | SOME m => m val h = fn s => mkN (pdefname,s) in map (h o g o f) (list_upto num_prots) end val trtyp = prot_transactionT trac lthy val trstyp = mk_listT trtyp fun mk_prot_trm names = Term.Const (@{const_name "concat"}, mk_listT trstyp --> trstyp) $ mk_list trstyp (map (fn x => Term.Const (x, trstyp)) names) val lthy = if num_prots > 1 then fold (fn (i,pname) => mk_prot_def (pname, mk_prot_trm (mk_tnames i))) (map (fn i => (i, nth pnames i)) (list_upto num_prots)) lthy else lthy val pnames' = map (fn n => full_name n lthy) pnames fun mk_prot_trm_with_star i = let fun f j = if j = i then Term.Const (nth pnames' j, trstyp) else (Term.Const (@{const_name "map"}, [trtyp --> trtyp, trstyp] ---> trstyp) $ Term.Const ("Transactions.transaction_star_proj", trtyp --> trtyp) $ Term.Const (nth pnames' j, trstyp)) in Term.Const (@{const_name "concat"}, mk_listT trstyp --> trstyp) $ mk_list trstyp (map f (list_upto num_prots)) end val lthy = if num_prots > 1 then fold (fn (i,pname) => mk_prot_def (pname, mk_prot_trm_with_star i)) (map (fn i => (i, nth pnames i ^ "_with_star")) (list_upto num_prots)) lthy else lthy in mk_prot_def (pdefname, mk_prot_trm (if num_prots > 1 then pnames' else tnames)) lthy end in (trac, lthy |> def_transactions |> def_protocols) end end \ ML\ structure trac = struct open Trac_Term val info = Output.information (* Define global configuration option "trac" *) (* val trac_fp_compute_binary_cfg = let val (trac_fp_compute_path_config, trac_fp_compute_path_setup) = Attrib.config_string (Binding.name "trac_fp_compute") (K "trac_fp_compute") in Context.>>(Context.map_theory trac_fp_compute_path_setup); trac_fp_compute_path_config end val trac_eval_cfg = let val (trac_fp_compute_eval_config, trac_fp_compute_eval) = Attrib.config_bool (Binding.name "trac_fp_compute_eval") (K false) in Context.>>(Context.map_theory trac_fp_compute_eval); trac_fp_compute_eval_config end *) type hide_tvar_tab = (TracProtocol.protocol) Symtab.table fun trac_eq (a, a') = (#name a) = (#name a') fun merge_trac_tab (tab,tab') = Symtab.merge trac_eq (tab,tab') structure Data = Generic_Data ( type T = hide_tvar_tab val empty = Symtab.empty:hide_tvar_tab val extend = I fun merge(t1,t2) = merge_trac_tab (t1, t2) ); fun update p thy = Context.theory_of ((Data.map (fn tab => Symtab.update (#name p, p) tab) (Context.Theory thy))) fun lookup name thy = (Symtab.lookup ((Data.get o Context.Theory) thy) name,thy) fun mk_abs_filename thy filename = let val filename = Path.explode filename val master_dir = Resources.master_directory thy in Path.implode (if (Path.is_absolute filename) then filename else master_dir + filename) end (* fun exec {trac_path, error_detail} filename = let open OS.FileSys OS.Process val tmpname = tmpName() val err_tmpname = tmpName() fun plural 1 = "" | plural _ = "s" val trac = case trac_path of SOME s => s | NONE => raise error ("trac_fp_compute_path not specified") val cmdline = trac ^ " \"" ^ filename ^ "\" > " ^ tmpname ^ " 2> " ^ err_tmpname in if isSuccess (system cmdline) then (OS.FileSys.remove err_tmpname; tmpname) else let val _ = OS.FileSys.remove tmpname val (msg, rest) = File.read_lines (Path.explode err_tmpname) |> chop error_detail val _ = OS.FileSys.remove err_tmpname val _ = warning ("trac failed on " ^ filename ^ "\nCommand: " ^ cmdline ^ "\n\nOutput:\n" ^ cat_lines (msg @ (if null rest then [] else ["(... " ^ string_of_int (length rest) ^ " more line" ^ plural (length rest) ^ ")"]))) in raise error ("trac failed on " ^ filename) end end *) fun lookup_trac (pname:string) lthy = Option.valOf (fst (lookup pname (Proof_Context.theory_of lthy))) fun def_fp fp_str print (trac, lthy) = let val fp = TracFpParser.parse_str fp_str val (trac,lthy) = trac_definitorial_package.define_fixpoint fp trac print lthy val lthy = Local_Theory.raw_theory (update trac) lthy in (trac, lthy) end fun def_fp_file filename print (trac, lthy) = let val thy = Proof_Context.theory_of lthy val abs_filename = mk_abs_filename thy filename val fp = TracFpParser.parse_file abs_filename val (trac,lthy) = trac_definitorial_package.define_fixpoint fp trac print lthy val lthy = Local_Theory.raw_theory (update trac) lthy in (trac, lthy) end fun def_fp_trac fp_filename print (trac, lthy) = let open OS.FileSys OS.Process val _ = info("Checking protocol specification with trac.") val thy = Proof_Context.theory_of lthy (* val trac = Config.get_global thy trac_binary_cfg *) val abs_filename = mk_abs_filename thy fp_filename (* val fp_file = exec {error_detail=10, trac_path = SOME trac} abs_filename *) (* val fp_raw = File.read (Path.explode fp_file) *) val fp_raw = File.read (Path.explode abs_filename) val fp = TracFpParser.parse_str fp_raw (* val _ = OS.FileSys.remove fp_file *) val _ = if TracFpParser.attack fp then error (" ATTACK found, skipping generating of Isabelle/HOL definitions.\n\n") else info(" No attack found, continue with generating Isabelle/HOL definitions.") val (trac,lthy) = trac_definitorial_package.define_fixpoint fp trac print lthy val lthy = Local_Theory.raw_theory (update trac) lthy in (trac, lthy) end fun def_trac_term_model str lthy = let val trac = TracProtocolParser.parse_str str val lthy = Local_Theory.raw_theory (update trac) lthy val lthy = trac_definitorial_package.define_term_model trac lthy in (trac, lthy) end val def_trac_protocol = trac_definitorial_package.define_protocol fun def_trac str print = def_trac_protocol print o def_trac_term_model str fun def_trac_file filename print lthy = let val trac_raw = File.read (Path.explode filename) val (trac,lthy) = def_trac trac_raw print lthy val lthy = Local_Theory.raw_theory (update trac) lthy in (trac, lthy) end fun def_trac_fp_trac trac_str print lthy = let open OS.FileSys OS.Process val (trac,lthy) = def_trac trac_str print lthy val tmpname = tmpName() val _ = File.write (Path.explode tmpname) trac_str val (trac,lthy) = def_fp_trac tmpname print (trac, lthy) val _ = OS.FileSys.remove tmpname val lthy = Local_Theory.raw_theory (update trac) lthy in lthy end end \ ML\ val fileNameP = Parse.name -- Parse.name val _ = Outer_Syntax.local_theory' @{command_keyword "trac_import"} "Import protocol and fixpoint from trac files." (fileNameP >> (fn (trac_filename, fp_filename) => fn print => trac.def_trac_file trac_filename print #> trac.def_fp_file fp_filename print #> snd)); val _ = Outer_Syntax.local_theory' @{command_keyword "trac_import_trac"} "Import protocol from trac file and compute fixpoint with trac." (fileNameP >> (fn (trac_filename, fp_filename) => fn print => trac.def_trac trac_filename print #> trac.def_fp_trac fp_filename print #> snd)); val _ = Outer_Syntax.local_theory' @{command_keyword "trac_trac"} "Define protocol using trac format and compute fixpoint with trac." (Parse.cartouche >> (fn trac => fn print => trac.def_trac_fp_trac trac print)); val _ = Outer_Syntax.local_theory' @{command_keyword "trac"} "Define protocol and (optionally) fixpoint using trac format." (Parse.cartouche -- Scan.optional Parse.cartouche "" >> (fn (trac,fp) => fn print => if fp = "" then trac.def_trac trac print #> snd else trac.def_trac trac print #> trac.def_fp fp print #> snd)); \ ML\ val name_prefix_parser = Parse.!!! (Parse.name --| Parse.$$$ ":" -- Parse.name) (* Original definition (opt_evaluator) copied from value_command.ml *) val opt_proof_method_choice = Scan.optional (\<^keyword>\[\ |-- Parse.name --| \<^keyword>\]\) "safe"; (* Original definition (locale_expression) copied from parse_spec.ML *) val opt_defs_list = Scan.optional (\<^keyword>\for\ |-- Scan.repeat1 Parse.name >> (fn xs => if length xs > 3 then error "Too many optional arguments" else xs)) []; val security_proof_locale_parser = name_prefix_parser -- opt_defs_list val security_proof_locale_parser_with_method_choice = opt_proof_method_choice -- name_prefix_parser -- opt_defs_list fun protocol_model_setup_proof_state name prefix lthy = let fun f x y z = ([((x,Position.none),((y,true),(Expression.Positional z,[])))],[]) val _ = if name = "" then error "No name given" else () val pexpr = f "stateful_protocol_model" name (protocol_model_interpretation_params prefix) val pdefs = protocol_model_interpretation_defs name val proof_state = Interpretation.global_interpretation_cmd pexpr pdefs lthy in proof_state end fun protocol_security_proof_proof_state manual_proof name prefix opt_defs print lthy = let fun f x y z = ([((x,Position.none),((y,true),(Expression.Positional z,[])))],[]) val _ = if name = "" then error "No name given" else () val num_defs = length opt_defs val pparams = protocol_model_interpretation_params prefix val default_defs = [prefix ^ "_" ^ "protocol", prefix ^ "_" ^ "fixpoint"] fun g locale_name extra_params = f locale_name name (pparams@map SOME extra_params) val (prot_fp_smp_names, pexpr) = if manual_proof then (case num_defs of 0 => (default_defs, g "secure_stateful_protocol'" default_defs) | 1 => (opt_defs, g "secure_stateful_protocol''" opt_defs) | 2 => (opt_defs, g "secure_stateful_protocol'" opt_defs) | _ => (opt_defs, g "secure_stateful_protocol" opt_defs)) else (case num_defs of 0 => (default_defs, g "secure_stateful_protocol''''" default_defs) | 1 => (opt_defs, g "secure_stateful_protocol''" opt_defs) | 2 => (opt_defs, g "secure_stateful_protocol''''" opt_defs) | _ => (opt_defs, g "secure_stateful_protocol'''" opt_defs)) val proof_state = lthy |> declare_protocol_checks print |> Interpretation.global_interpretation_cmd pexpr [] in (prot_fp_smp_names, proof_state) end val _ = Outer_Syntax.local_theory \<^command_keyword>\protocol_model_setup\ "prove interpretation of protocol model locale into global theory" (name_prefix_parser >> (fn (name,prefix) => fn lthy => let val proof_state = protocol_model_setup_proof_state name prefix lthy val meth = let val m = "protocol_model_interpretation" val _ = Output.information ( "Proving protocol model locale instance with proof method " ^ m) in Method.Source (Token.make_src (m, Position.none) []) end in ml_isar_wrapper.prove_state_simple meth proof_state end)); val _ = Outer_Syntax.local_theory_to_proof \<^command_keyword>\manual_protocol_model_setup\ "prove interpretation of protocol model locale into global theory" (name_prefix_parser >> (fn (name,prefix) => fn lthy => let val proof_state = protocol_model_setup_proof_state name prefix lthy val subgoal_proof = " subgoal by protocol_model_subgoal\n" val _ = Output.information ("Example proof:\n" ^ Active.sendback_markup_command (" apply unfold_locales\n"^ subgoal_proof^ subgoal_proof^ subgoal_proof^ subgoal_proof^ subgoal_proof^ " done\n")) in proof_state end)); val _ = Outer_Syntax.local_theory' \<^command_keyword>\protocol_security_proof\ "prove interpretation of secure protocol locale into global theory" (security_proof_locale_parser_with_method_choice >> (fn params => fn print => fn lthy => let val ((opt_meth_level,(name,prefix)),opt_defs) = params val (defs, proof_state) = protocol_security_proof_proof_state false name prefix opt_defs print lthy val num_defs = length defs val meth = let val m = case opt_meth_level of "safe" => "check_protocol" ^ "'" (* (if num_defs = 1 then "'" else "") *) | "unsafe" => "check_protocol_unsafe" ^ "'" (* (if num_defs = 1 then "'" else "") *) | _ => error ("Invalid option: " ^ opt_meth_level) val _ = Output.information ( "Proving security of protocol " ^ nth defs 0 ^ " with proof method " ^ m) val _ = if num_defs > 1 then Output.information ("Using fixpoint " ^ nth defs 1) else () val _ = if num_defs > 2 then Output.information ("Using SMP set " ^ nth defs 2) else () in Method.Source (Token.make_src (m, Position.none) []) end in ml_isar_wrapper.prove_state_simple meth proof_state end )); val _ = Outer_Syntax.local_theory_to_proof' \<^command_keyword>\manual_protocol_security_proof\ "prove interpretation of secure protocol locale into global theory" (security_proof_locale_parser >> (fn params => fn print => fn lthy => let val ((name,prefix),opt_defs) = params val (defs, proof_state) = protocol_security_proof_proof_state true name prefix opt_defs print lthy val subgoal_proof = let val m = "code_simp" (* case opt_meth_level of "safe" => "code_simp" | "unsafe" => "eval" | _ => error ("Invalid option: " ^ opt_meth_level) *) in " subgoal by " ^ m ^ "\n" end val _ = Output.information ("Example proof:\n" ^ Active.sendback_markup_command (" apply check_protocol_intro\n"^ subgoal_proof^ (if length defs = 1 then "" else subgoal_proof^ subgoal_proof^ subgoal_proof^ subgoal_proof)^ " done\n")) in proof_state end )); \ end diff --git a/thys/Automatic_Refinement/Lib/Cond_Rewr_Conv.ML b/thys/Automatic_Refinement/Lib/Cond_Rewr_Conv.ML --- a/thys/Automatic_Refinement/Lib/Cond_Rewr_Conv.ML +++ b/thys/Automatic_Refinement/Lib/Cond_Rewr_Conv.ML @@ -1,30 +1,30 @@ signature COND_REWR_CONV = sig - val cond_rewr_conv: tactic -> thm -> Proof.context -> conv - val cond_rewrs_conv: tactic -> thm list -> Proof.context -> conv + val cond_rewr_conv: (Proof.context -> tactic) -> thm -> Proof.context -> conv + val cond_rewrs_conv: (Proof.context -> tactic) -> thm list -> Proof.context -> conv end structure Cond_Rewr_Conv :COND_REWR_CONV = struct open Refine_Util (* Conditional rewrite rule. tac used to discharge conditions *) fun cond_rewr_conv_aux tac thm ctxt ct = let val lhs = thm |> Thm.concl_of |> Logic.dest_equals |> #1 |> Thm.cterm_of ctxt val inst = Thm.match (lhs,ct) handle Pattern.MATCH => raise CTERM ("dis_rewr_conv: MATCH",[lhs,ct]) val thm' = Thm.instantiate inst thm val dprems = Thm.prems_of thm' val dthms = map (fn t => - (Goal.prove ctxt [] [] t (K tac)) handle ERROR s + (Goal.prove ctxt [] [] t (tac o #context)) handle ERROR s => raise TERM ("dis_rew_conv: "^ s,[t])) dprems val res = thm' OF dthms in res end; fun cond_rewr_conv tac thm ctxt = fix_conv ctxt (cond_rewr_conv_aux tac thm ctxt) (*fun first_conv [] ct = Conv.no_conv ct | first_conv (cv::cvs) ct = (cv else_conv first_conv cvs) ct*) fun cond_rewrs_conv tac thms ctxt = Conv.first_conv (map (fn thm => cond_rewr_conv tac thm ctxt) thms) 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 + val f_tac_conv: Proof.context -> (term -> term) -> (Proof.context -> 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' = Vars.map (fn ((name, _), _) => tag_ct ctxt name) tm; val ct' = Thm.instantiate_cterm (tym, tm') pat; + val goal = Logic.mk_equals (apply2 Thm.term_of (ct, ct')) + val goal_ctxt = Variable.declare_term goal ctxt 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.prove_internal goal_ctxt [] (Thm.cterm_of ctxt goal) + (K (simp_tac (put_simpset HOL_basic_ss goal_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) + val goal_ctxt = Variable.declare_term goal ctxt + val thm = Goal.prove_internal ctxt [] (Thm.cterm_of ctxt goal) (K (tac goal_ctxt)) 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 (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/Automatic_Refinement/Parametricity/Param_Tool.thy b/thys/Automatic_Refinement/Parametricity/Param_Tool.thy --- a/thys/Automatic_Refinement/Parametricity/Param_Tool.thy +++ b/thys/Automatic_Refinement/Parametricity/Param_Tool.thy @@ -1,348 +1,348 @@ section \Basic Parametricity Reasoning\ theory Param_Tool imports Relators begin subsection \Auxiliary Lemmas\ lemma tag_both: "\ (Let x f,Let x' f')\R \ \ (f x,f' x')\R" by simp lemma tag_rhs: "\ (c,Let x f)\R \ \ (c,f x)\R" by simp lemma tag_lhs: "\ (Let x f,a)\R \ \ (f x,a)\R" by simp lemma tagged_fun_relD_both: "\ (f,f')\A\B; (x,x')\A \ \ (Let x f,Let x' f')\B" and tagged_fun_relD_rhs: "\ (f,f')\A\B; (x,x')\A \ \ (f x,Let x' f')\B" and tagged_fun_relD_lhs: "\ (f,f')\A\B; (x,x')\A \ \ (Let x f,f' x')\B" and tagged_fun_relD_none: "\ (f,f')\A\B; (x,x')\A \ \ (f x,f' x')\B" by (simp_all add: fun_relD) subsection \ML-Setup\ ML \ signature PARAMETRICITY = sig type param_ruleT = { lhs: term, rhs: term, R: term, rhs_head: term, arity: int } val dest_param_term: term -> param_ruleT val dest_param_rule: thm -> param_ruleT val dest_param_goal: int -> thm -> param_ruleT val safe_fun_relD_tac: Proof.context -> tactic' val adjust_arity: int -> thm -> thm val adjust_arity_tac: int -> Proof.context -> tactic' val unlambda_tac: Proof.context -> tactic' val prepare_tac: Proof.context -> tactic' val fo_rule: thm -> thm (*** Basic tactics ***) val param_rule_tac: Proof.context -> thm -> tactic' val param_rules_tac: Proof.context -> thm list -> tactic' val asm_param_tac: Proof.context -> tactic' (*** Nets of parametricity rules ***) type param_net val net_empty: param_net val net_add: thm -> param_net -> param_net val net_del: thm -> param_net -> param_net val net_add_int: Context.generic -> thm -> param_net -> param_net val net_del_int: Context.generic -> thm -> param_net -> param_net val net_tac: param_net -> Proof.context -> tactic' (*** Default parametricity rules ***) val add_dflt: thm -> Context.generic -> Context.generic val add_dflt_attr: attribute val del_dflt: thm -> Context.generic -> Context.generic val del_dflt_attr: attribute val get_dflt: Proof.context -> param_net (** Configuration **) val cfg_use_asm: bool Config.T val cfg_single_step: bool Config.T (** Setup **) val setup: theory -> theory end structure Parametricity : PARAMETRICITY = struct type param_ruleT = { lhs: term, rhs: term, R: term, rhs_head: term, arity: int } fun dest_param_term t = case strip_all_body t |> Logic.strip_imp_concl |> HOLogic.dest_Trueprop of @{mpat "(?lhs,?rhs):?R"} => let val (rhs_head,arity) = case strip_comb rhs of (c as Const _,l) => (c,length l) | (c as Free _,l) => (c,length l) | (c as Abs _,l) => (c,length l) | _ => raise TERM ("dest_param_term: Head",[t]) in { lhs = lhs, rhs = rhs, R=R, rhs_head = rhs_head, arity = arity } end | t => raise TERM ("dest_param_term: Expected (_,_):_",[t]) val dest_param_rule = dest_param_term o Thm.prop_of fun dest_param_goal i st = if i > Thm.nprems_of st then raise THM ("dest_param_goal",i,[st]) else dest_param_term (Logic.concl_of_goal (Thm.prop_of st) i) fun safe_fun_relD_tac ctxt = let fun t a b = fo_resolve_tac [a] ctxt THEN' resolve_tac ctxt [b] in DETERM o ( t @{thm tag_both} @{thm tagged_fun_relD_both} ORELSE' t @{thm tag_rhs} @{thm tagged_fun_relD_rhs} ORELSE' t @{thm tag_lhs} @{thm tagged_fun_relD_lhs} ORELSE' resolve_tac ctxt @{thms tagged_fun_relD_none} ) end fun adjust_arity i thm = if i = 0 then thm else if i<0 then funpow (~i) (fn thm => thm RS @{thm fun_relI}) thm else funpow i (fn thm => thm RS @{thm fun_relD}) thm fun NTIMES k tac = if k <= 0 then K all_tac else tac THEN' NTIMES (k-1) tac fun adjust_arity_tac n ctxt i st = (if n = 0 then K all_tac else if n>0 then NTIMES n (DETERM o resolve_tac ctxt @{thms fun_relI}) else NTIMES (~n) (safe_fun_relD_tac ctxt)) i st fun unlambda_tac ctxt i st = case try (dest_param_goal i) st of NONE => Seq.empty | SOME g => let val n = Term.strip_abs (#rhs_head g) |> #1 |> length in NTIMES n (resolve_tac ctxt @{thms fun_relI}) i st end fun prepare_tac ctxt = Subgoal.FOCUS (K (PRIMITIVE (Drule.eta_contraction_rule))) ctxt THEN' unlambda_tac ctxt fun could_param_rl rl i st = if i > Thm.nprems_of st then NONE else ( case (try (dest_param_goal i) st, try dest_param_term rl) of (SOME g, SOME r) => if Term.could_unify (#rhs_head g, #rhs_head r) then SOME (#arity r - #arity g) else NONE | _ => NONE ) fun param_rule_tac_aux ctxt rl i st = case could_param_rl (Thm.prop_of rl) i st of SOME adj => (adjust_arity_tac adj ctxt THEN' resolve_tac ctxt [rl]) i st | _ => Seq.empty fun param_rule_tac ctxt rl = prepare_tac ctxt THEN' param_rule_tac_aux ctxt rl fun param_rules_tac ctxt rls = prepare_tac ctxt THEN' FIRST' (map (param_rule_tac_aux ctxt) rls) fun asm_param_tac_aux ctxt i st = if i > Thm.nprems_of st then Seq.empty else let val prems = Logic.prems_of_goal (Thm.prop_of st) i |> tag_list 1 fun tac (n,t) i st = case could_param_rl t i st of SOME adj => (adjust_arity_tac adj ctxt THEN' rprem_tac n ctxt) i st | NONE => Seq.empty in FIRST' (map tac prems) i st end fun asm_param_tac ctxt = prepare_tac ctxt THEN' asm_param_tac_aux ctxt type param_net = (param_ruleT * thm) Item_Net.T local val param_get_key = single o #rhs_head o #1 in val net_empty = Item_Net.init (Thm.eq_thm o apply2 #2) param_get_key end fun wrap_pr_op f context thm = case try (`dest_param_rule) thm of NONE => let val msg = "Ignoring invalid parametricity theorem: " ^ Thm.string_of_thm (Context.proof_of context) thm val _ = warning msg in I end | SOME p => f p val net_add_int = wrap_pr_op Item_Net.update val net_del_int = wrap_pr_op Item_Net.remove val net_add = Item_Net.update o `dest_param_rule val net_del = Item_Net.remove o `dest_param_rule fun net_tac_aux net ctxt i st = if i > Thm.nprems_of st then Seq.empty else let val g = dest_param_goal i st val rls = Item_Net.retrieve net (#rhs_head g) fun tac (r,thm) = adjust_arity_tac (#arity r - #arity g) ctxt THEN' DETERM o resolve_tac ctxt [thm] in FIRST' (map tac rls) i st end fun net_tac net ctxt = prepare_tac ctxt THEN' net_tac_aux net ctxt structure dflt_rules = Generic_Data ( type T = param_net val empty = net_empty val extend = I val merge = Item_Net.merge ) fun add_dflt thm context = dflt_rules.map (net_add_int context thm) context fun del_dflt thm context = dflt_rules.map (net_del_int context thm) context val add_dflt_attr = Thm.declaration_attribute add_dflt val del_dflt_attr = Thm.declaration_attribute del_dflt val get_dflt = dflt_rules.get o Context.Proof val cfg_use_asm = Attrib.setup_config_bool @{binding param_use_asm} (K true) val cfg_single_step = Attrib.setup_config_bool @{binding param_single_step} (K false) local open Refine_Util val param_modifiers = [Args.add -- Args.colon >> K (Method.modifier add_dflt_attr \<^here>), Args.del -- Args.colon >> K (Method.modifier del_dflt_attr \<^here>), Args.$$$ "only" -- Args.colon >> K {init = Context.proof_map (dflt_rules.map (K net_empty)), attribute = add_dflt_attr, pos = \<^here>}] val param_flags = parse_bool_config "use_asm" cfg_use_asm || parse_bool_config "single_step" cfg_single_step in val parametricity_method = parse_paren_lists param_flags |-- Method.sections param_modifiers >> (fn _ => fn ctxt => let val net2 = get_dflt ctxt val asm_tac = if Config.get ctxt cfg_use_asm then asm_param_tac ctxt else K no_tac val RPT = if Config.get ctxt cfg_single_step then I else REPEAT_ALL_NEW_FWD in SIMPLE_METHOD' ( RPT ( (assume_tac ctxt ORELSE' net_tac net2 ctxt ORELSE' asm_tac) ) ) end ) end fun fo_rule thm = case Thm.concl_of thm of @{mpat "Trueprop ((_,_)\_\_)"} => fo_rule (thm RS @{thm fun_relD}) | _ => thm val param_fo_attr = Scan.succeed (Thm.rule_attribute [] (K fo_rule)) val setup = I #> Attrib.setup @{binding param} (Attrib.add_del add_dflt_attr del_dflt_attr) "declaration of parametricity theorem" #> Global_Theory.add_thms_dynamic (@{binding param}, map #2 o Item_Net.content o dflt_rules.get) #> Method.setup @{binding parametricity} parametricity_method "Parametricity solver" #> Attrib.setup @{binding param_fo} param_fo_attr "Parametricity: Rule in first-order form" end \ setup Parametricity.setup subsection \Convenience Tools\ ML \ (* Prefix p_ or wrong type supresses generation of relAPP *) fun cnv_relAPP t = let fun consider (Var ((name,_),T)) = if String.isPrefix "p_" name then false else ( case T of Type(@{type_name set},[Type(@{type_name prod},_)]) => true | _ => false) | consider _ = true fun strip_rcomb u : term * term list = let fun stripc (x as (f$t, ts)) = if consider t then stripc (f, t::ts) else x | stripc x = x in stripc(u,[]) end; val (f,a) = strip_rcomb t in Relators.list_relAPP a f end fun to_relAPP_conv ctxt = Refine_Util.f_tac_conv ctxt cnv_relAPP - (ALLGOALS (simp_tac - (put_simpset HOL_basic_ss ctxt addsimps @{thms relAPP_def}))) + (fn goal_ctxt => ALLGOALS (simp_tac + (put_simpset HOL_basic_ss goal_ctxt addsimps @{thms relAPP_def}))) val to_relAPP_attr = Thm.rule_attribute [] (fn context => let val ctxt = Context.proof_of context in Conv.fconv_rule (Conv.arg1_conv (to_relAPP_conv ctxt)) end) \ attribute_setup to_relAPP = \Scan.succeed (to_relAPP_attr)\ "Convert relator definition to prefix-form" end diff --git a/thys/Automatic_Refinement/Tool/Autoref_Fix_Rel.thy b/thys/Automatic_Refinement/Tool/Autoref_Fix_Rel.thy --- a/thys/Automatic_Refinement/Tool/Autoref_Fix_Rel.thy +++ b/thys/Automatic_Refinement/Tool/Autoref_Fix_Rel.thy @@ -1,984 +1,984 @@ section \Relator Fixing\ theory Autoref_Fix_Rel imports Autoref_Id_Ops begin ML_val "2 upto 5" subsubsection \Priority tags\ text \ Priority tags are used to influence the ordering of refinement theorems. A priority tag defines two numeric priorities, a major and a minor priority. The major priority is considered first, the minor priority last, i.e., after the homogenity and relator-priority criteria. The default value for both priorities is 0. \ definition PRIO_TAG :: "int \ int \ bool" where [simp]: "PRIO_TAG ma mi \ True" lemma PRIO_TAGI: "PRIO_TAG ma mi" by simp abbreviation "MAJOR_PRIO_TAG i \ PRIO_TAG i 0" abbreviation "MINOR_PRIO_TAG i \ PRIO_TAG 0 i" abbreviation "DFLT_PRIO_TAG \ PRIO_TAG 0 0" text \Some standard tags\ abbreviation "PRIO_TAG_OPTIMIZATION \ MINOR_PRIO_TAG 10" \ \Optimized version of an algorithm, with additional side-conditions\ abbreviation "PRIO_TAG_GEN_ALGO \ MINOR_PRIO_TAG (- 10)" \ \Generic algorithm, considered to be less efficient than default algorithm\ subsection \Solving Relator Constraints\ text \ In this phase, we try to instantiate the annotated relators, using the available refinement rules. \ definition CONSTRAINT :: "'a \ ('c\'a) set \ bool" where [simp]: "CONSTRAINT f R \ True" lemma CONSTRAINTI: "CONSTRAINT f R" by auto ML \ structure Autoref_Rules = Named_Thms ( val name = @{binding autoref_rules_raw} val description = "Refinement Framework: " ^ "Automatic refinement rules" ); \ setup Autoref_Rules.setup text \Generic algorithm tags have to be defined here, as we need them for relator fixing !\ definition PREFER_tag :: "bool \ bool" where [simp, autoref_tag_defs]: "PREFER_tag x \ x" definition DEFER_tag :: "bool \ bool" where [simp, autoref_tag_defs]: "DEFER_tag x \ x" lemma PREFER_tagI: "P \ PREFER_tag P" by simp lemma DEFER_tagI: "P \ DEFER_tag P" by simp lemmas SIDEI = PREFER_tagI DEFER_tagI definition [simp, autoref_tag_defs]: "GEN_OP_tag P == P" lemma GEN_OP_tagI: "P ==> GEN_OP_tag P" by simp abbreviation "SIDE_GEN_OP P == PREFER_tag (GEN_OP_tag P)" text \Shortcut for assuming an operation in a generic algorithm lemma\ abbreviation "GEN_OP c a R \ SIDE_GEN_OP ((c,OP a ::: R) \ R)" definition TYREL :: "('a\'b) set \ bool" where [simp]: "TYREL R \ True" definition TYREL_DOMAIN :: "'a itself \ bool" where [simp]: "TYREL_DOMAIN i \ True" lemma TYREL_RES: "\ TYREL_DOMAIN TYPE('a); TYREL (R::(_\'a) set) \ \ TYREL R" . lemma DOMAIN_OF_TYREL: "TYREL (R::(_\'a) set) \ TYREL_DOMAIN TYPE('a)" by simp lemma TYRELI: "TYREL (R::(_\'a) set)" by simp lemma ty_REL: "TYREL (R::(_\'a) set)" by simp ML \ signature AUTOREF_FIX_REL = sig type constraint = (term * term) list * (term * term) type thm_pairs = (constraint option * thm) list type hom_net = (int * thm) Net.net val thm_pairsD_init: Proof.context -> Proof.context val thm_pairsD_get: Proof.context -> thm_pairs val constraints_of_term: term -> (term * term) list val constraints_of_goal: int -> thm -> (term * term) list val mk_CONSTRAINT: term * term -> term val mk_CONSTRAINT_rl: Proof.context -> constraint -> thm val insert_CONSTRAINTS_tac: Proof.context -> tactic' val constraint_of_thm: Proof.context -> thm -> constraint datatype prio_relpos = PR_FIRST | PR_LAST | PR_BEFORE of string | PR_AFTER of string val declare_prio: string -> term -> prio_relpos -> local_theory -> local_theory val delete_prio: string -> local_theory -> local_theory val print_prios: Proof.context -> unit val compute_hom_net: thm_pairs -> Proof.context -> hom_net val add_hom_rule: thm -> Context.generic -> Context.generic val del_hom_rule: thm -> Context.generic -> Context.generic val get_hom_rules: Proof.context -> thm list val add_tyrel_rule: thm -> Context.generic -> Context.generic val del_tyrel_rule: thm -> Context.generic -> Context.generic val get_tyrel_rules: Proof.context -> thm list val insert_tyrel_tac : Proof.context -> int -> int -> tactic' val solve_tyrel_tac : Proof.context -> tactic' val tyrel_tac : Proof.context -> itactic val internal_hom_tac: Proof.context -> itactic val internal_spec_tac: Proof.context -> itactic val internal_solve_tac: Proof.context -> itactic val guess_relators_tac: Proof.context -> itactic val pretty_constraint: Proof.context -> constraint -> Pretty.T val pretty_constraints: Proof.context -> constraint list -> Pretty.T val pretty_thm_pair: Proof.context -> (constraint option * thm) -> Pretty.T val pretty_thm_pairs: Proof.context -> thm_pairs -> Pretty.T val analyze: Proof.context -> int -> int -> thm -> bool val pretty_failure: Proof.context -> int -> int -> thm -> Pretty.T val try_solve_tac: Proof.context -> tactic' val solve_step_tac: Proof.context -> tactic' val phase: Autoref_Phases.phase val setup: theory -> theory end structure Autoref_Fix_Rel :AUTOREF_FIX_REL = struct type constraint = (term * term) list * (term * term) type thm_pairs = (constraint option * thm) list type hom_net = (int * thm) Net.net (*********************) (* Constraints *) (*********************) local fun fix_loose_bvars env t = if Term.is_open t then let val frees = tag_list 0 env |> map (fn (i,(n,T)) => Free (":"^string_of_int i ^ "_" ^ n,T)) in subst_bounds (frees, t) end else t fun constraints env @{mpat "OP ?f ::: ?R"} = ( Term.is_open R andalso raise TERM ("Loose bvar in relator",[R]); [(fix_loose_bvars env f,R)] ) | constraints _ (Free _) = [] | constraints _ (Bound _) = [] | constraints env @{mpat "?f ::: _"} = constraints env f | constraints env @{mpat "?f$?x"} = constraints env x @ constraints env f | constraints env @{mpat "PROTECT (\x. PROTECT ?t)"} = constraints ((x,x_T)::env) t | constraints _ @{mpat "PROTECT PROTECT"} = [] | constraints _ t = raise TERM ("constraints_of_term",[t]) in val constraints_of_term = constraints [] end fun constraints_of_goal i st = case Logic.concl_of_goal (Thm.prop_of st) i of @{mpat "Trueprop ((_,?a)\_)"} => constraints_of_term a | _ => raise THM ("constraints_of_goal",i,[st]) fun mk_CONSTRAINT (f,R) = let val fT = fastype_of f val RT = fastype_of R val res = Const (@{const_name CONSTRAINT},fT --> RT --> HOLogic.boolT) $f$R in res end; (* Types of f and R must match! *) fun mk_CONSTRAINT_rl ctxt (ps,c) = let val ps = map (mk_CONSTRAINT #> HOLogic.mk_Trueprop) ps val c = mk_CONSTRAINT c |> HOLogic.mk_Trueprop val g = Logic.list_implies (ps,c) in Goal.prove ctxt [] [] g - (K (resolve_tac ctxt @{thms CONSTRAINTI} 1)) + (fn {context = goal_ctxt, ...} => resolve_tac goal_ctxt @{thms CONSTRAINTI} 1) end; (* Internal use for hom-patterns, f and R are unified *) fun mk_CONSTRAINT_rl_atom ctxt (f,R) = let val ts = map (SOME o Thm.cterm_of ctxt) [f,R] val idx = Term.maxidx_term f (Term.maxidx_of_term R) + 1 in infer_instantiate' ctxt ts (Thm.incr_indexes idx @{thm CONSTRAINTI}) end; fun insert_CONSTRAINTS_tac ctxt i st = let val cs = constraints_of_goal i st |> map (mk_CONSTRAINT #> HOLogic.mk_Trueprop #> Thm.cterm_of ctxt) in Refine_Util.insert_subgoals_tac cs i st end fun constraint_of_thm ctxt thm = let exception NO_REL of term open Autoref_Tagging fun extract_entry t = case Logic.strip_imp_concl (strip_all_body t) of @{mpat "Trueprop ((_,?f)\_)"} => SOME (fst (strip_app f),t) | _ => NONE fun relator_of t = let (*val _ = tracing (Syntax.string_of_term @{context} t)*) val t = strip_all_body t val prems = Logic.strip_imp_prems t val concl = Logic.strip_imp_concl t in case concl of @{mpat "Trueprop ((_,?t)\?R)"} => let val (f,args) = strip_app t in case f of @{mpat "OP ?f:::?rel"} => (f,rel) | _ => let val rels = map_filter extract_entry prems fun find_rel t = case filter (fn (t',_) => t=t') rels of [(_,t)] => snd (relator_of t) | _ => raise NO_REL t val argrels = map find_rel args val rel = fold Relators.mk_fun_rel (rev argrels) R in (f,rel) end end | _ => raise THM ("constraint_of_thm: Invalid concl",~1,[thm]) end val (f,rel) = relator_of (Thm.prop_of thm) handle exc as (NO_REL t) => ( warning ( "Could not infer unique higher-order relator for " ^ "refinement rule: \n" ^ Thm.string_of_thm ctxt thm ^ "\n for argument: " ^ Syntax.string_of_term ctxt t ); Exn.reraise exc) (* Extract GEN_OP-tags *) fun genop_cs @{mpat "Trueprop (SIDE_GEN_OP ((_,OP ?f ::: _) \ ?R))"} = if has_Var f then NONE else SOME (f,R) | genop_cs _ = NONE val gen_ops = Thm.prems_of thm |> map_filter genop_cs in (gen_ops,(f,rel)) end (*********************) (* Priorities *) (*********************) structure Rel_Prio_List = Prio_List ( type item = string * term val eq = (op =) o apply2 fst ) structure Rel_Prio = Generic_Data ( type T = Rel_Prio_List.T val empty = Rel_Prio_List.empty val merge = Rel_Prio_List.merge val extend = I ) fun pretty_rel_prio ctxt (s,t) = Pretty.block [ Pretty.str s, Pretty.str ":", Pretty.brk 1, Syntax.pretty_term ctxt t ] fun print_prios ctxt = let val rpl = Rel_Prio.get (Context.Proof ctxt) in (map (pretty_rel_prio ctxt) rpl) |> Pretty.big_list "Relator Priorities" |> Pretty.string_of |> warning end datatype prio_relpos = PR_FIRST | PR_LAST | PR_BEFORE of string | PR_AFTER of string fun declare_prio name pat0 relpos lthy = let val pat1 = Proof_Context.cert_term lthy pat0 val pat2 = singleton (Variable.export_terms (Proof_Context.augment pat1 lthy) lthy) pat1 in lthy |> Local_Theory.declaration {syntax = false, pervasive = false} (fn phi => let val item = (name, Morphism.term phi pat2) in Rel_Prio.map (fn rpl => case relpos of PR_FIRST => Rel_Prio_List.add_first rpl item | PR_LAST => Rel_Prio_List.add_last rpl item | PR_BEFORE n => Rel_Prio_List.add_before rpl item (n,Term.dummy) | PR_AFTER n => Rel_Prio_List.add_after rpl item (n,Term.dummy) ) end) end fun delete_prio name = Local_Theory.declaration {syntax = false, pervasive = false} (fn phi => Rel_Prio.map (Rel_Prio_List.delete (name, Term.dummy))) local fun relators_of R = let fun f @{mpat "?R1.0\?R2.0"} = f R1 @ f R2 | f R = [R] in f R |> map Refine_Util.anorm_term |> distinct (op =) end fun dest_prio_tag @{mpat "Trueprop (PRIO_TAG ?ma ?mi)"} = apply2 (#2 o HOLogic.dest_number) (ma,mi) | dest_prio_tag t = raise TERM ("dest_prio_tag",[t]) fun get_tagged_prios thm = let val prems = Thm.prems_of thm fun r [] = (0,0) | r (prem::prems) = ( case try dest_prio_tag prem of NONE => r prems | SOME p => p ) in r prems end fun prio_order_of ctxt (SOME (_,(_,R)),thm) = let val rels = relators_of R val hom = length rels val (major_prio,minor_prio) = get_tagged_prios thm val rpl = Rel_Prio.get (Context.Proof ctxt) val matches = Pattern.matches (Proof_Context.theory_of ctxt) fun prefer ((_,p1),(_,p2)) = matches (p2,p1) fun prio_of R = Rel_Prio_List.prio_of (fn (_,pat) => matches (pat,R)) prefer rpl + 1 val prio = fold (fn R => fn s => prio_of R + s) rels 0 in (major_prio, (hom,(prio,minor_prio))) end | prio_order_of _ _ = raise Match val prio_order = prod_ord (rev_order o int_ord) (prod_ord int_ord (prod_ord (rev_order o int_ord) (rev_order o int_ord))) fun annotate_thm_pair ctxt (SOME (ps,(f,R)),thm) = let open Autoref_Tagging Conv fun warn () = warning ("Error annotating refinement theorem: " ^ Thm.string_of_thm ctxt thm ) val R_cert = Thm.cterm_of ctxt R fun cnv ctxt ct = (case Thm.term_of ct of @{mpat "OP _ ::: _"} => all_conv | @{mpat "OP _"} => mk_rel_ANNOT_conv ctxt R_cert | @{mpat "_ $ _"} => arg1_conv (cnv ctxt) | _ => mk_OP_conv then_conv mk_rel_ANNOT_conv ctxt R_cert ) ct (*val _ = tracing ("ANNOT: " ^ @{make_string} thm)*) val thm = (fconv_rule (rhs_conv cnv ctxt)) thm val thm = case try (fconv_rule (rhs_conv cnv ctxt)) thm of NONE => (warn (); thm) | SOME thm => thm (*val _ = tracing ("RES: " ^ @{make_string} thm)*) in (SOME (ps,(f,R)),thm) end | annotate_thm_pair _ p = p in fun compute_thm_pairs ctxt = let val rules = Autoref_Rules.get ctxt fun add_o p = (prio_order_of ctxt p,p) val pairs = rules |> map (fn thm => (try (constraint_of_thm ctxt) thm,thm)) val spairs = filter (is_some o #1) pairs |> map add_o |> sort (prio_order o apply2 #1) |> map #2 val npairs = filter (is_none o #1) pairs in spairs@npairs |> map (annotate_thm_pair ctxt) end end structure thm_pairsD = Autoref_Data ( type T = thm_pairs val compute = compute_thm_pairs val prereq = [] ) val thm_pairsD_init = thm_pairsD.init val thm_pairsD_get = thm_pairsD.get structure hom_rules = Named_Sorted_Thms ( val name = @{binding autoref_hom} val description = "Autoref: Homogenity rules" val sort = K I val transform = K ( fn thm => case Thm.concl_of thm of @{mpat "Trueprop (CONSTRAINT _ _)"} => [thm] | _ => raise THM ("Invalid homogenity rule",~1,[thm]) ) ) val add_hom_rule = hom_rules.add_thm val del_hom_rule = hom_rules.del_thm val get_hom_rules = hom_rules.get local open Relators fun repl @{mpat "?R\?S"} ctab = let val (R,ctab) = repl R ctab val (S,ctab) = repl S ctab in (mk_fun_rel R S,ctab) end | repl R ctab = let val (args,R) = strip_relAPP R val (args,ctab) = fold_map repl args ctab val (ctxt,tab) = ctab val (R,(ctxt,tab)) = case Termtab.lookup tab R of SOME R => (R,(ctxt,tab)) | NONE => let val aT = fastype_of R |> strip_type |> #2 |> HOLogic.dest_setT |> HOLogic.dest_prodT |> #2 val (cT,ctxt) = yield_singleton Variable.invent_types @{sort type} ctxt val cT = TFree cT val T = map fastype_of args ---> HOLogic.mk_setT (HOLogic.mk_prodT (cT,aT)) val (R',ctxt) = yield_singleton Variable.variant_fixes "R" ctxt val R' = list_relAPP args (Free (R',T)) val tab = Termtab.update (R,R') tab in (R',(ctxt,tab)) end in (R,(ctxt,tab)) end fun hom_pat_of_rel ctxt R = let val (R,(ctxt',_)) = repl R (ctxt,Termtab.empty) val R = singleton (Variable.export_terms ctxt' ctxt) R in Refine_Util.anorm_term R end in fun compute_hom_net pairs ctxt = let val cs = map_filter #1 pairs val cs' = map (fn (_,(f,R)) => (f,hom_pat_of_rel ctxt R)) cs val thms = get_hom_rules ctxt @ map (mk_CONSTRAINT_rl_atom ctxt) cs' val thms = map (Thm.cprop_of #> Thm.trivial) thms val net = Tactic.build_net thms in net end end structure hom_netD = Autoref_Data ( type T = hom_net fun compute ctxt = compute_hom_net (thm_pairsD.get ctxt) ctxt val prereq = [ thm_pairsD.init ] ) structure tyrel_rules = Named_Sorted_Thms ( val name = @{binding autoref_tyrel} val description = "Autoref: Type-based relator fixing rules" val sort = K I val transform = K ( fn thm => case Thm.prop_of thm of @{mpat "Trueprop (TYREL _)"} => [thm] | _ => raise THM ("Invalid tyrel-rule",~1,[thm]) ) ) val add_tyrel_rule = tyrel_rules.add_thm val del_tyrel_rule = tyrel_rules.del_thm val get_tyrel_rules = tyrel_rules.get local (*fun rel_annots @{mpat "_ ::: ?R"} = [R] | rel_annots @{mpat "?f$?x"} = rel_annots f @ rel_annots x | rel_annots @{mpat "PROTECT (\_. PROTECT ?t)"} = rel_annots t | rel_annots @{mpat "PROTECT PROTECT"} = [] | rel_annots (Free _) = [] | rel_annots (Bound _) = [] | rel_annots t = raise TERM ("rel_annots",[t]) *) fun add_relators t acc = let open Relators val (args,_) = strip_relAPP t val res = fold add_relators args acc val res = insert (op =) t res in res end fun add_relators_of_subgoal st i acc = case Logic.concl_of_goal (Thm.prop_of st) i of @{mpat "Trueprop (CONSTRAINT _ ?R)"} => add_relators R acc | _ => acc in fun insert_tyrel_tac ctxt i j k st = let fun get_constraint t = let val T = fastype_of t val res = Const (@{const_name TYREL}, T --> HOLogic.boolT) $ t in res |> HOLogic.mk_Trueprop |> Thm.cterm_of ctxt end val relators = fold (add_relators_of_subgoal st) (i upto j) [] val tyrels = map get_constraint relators in Refine_Util.insert_subgoals_tac tyrels k st end end fun solve_tyrel_tac ctxt = let fun mk_tac rl = resolve_tac ctxt @{thms TYREL_RES} THEN' match_tac ctxt [rl RS @{thm DOMAIN_OF_TYREL}] THEN' resolve_tac ctxt [rl] val tac = FIRST' (map mk_tac (tyrel_rules.get ctxt)) in DETERM o tac ORELSE' (TRY o resolve_tac ctxt @{thms TYRELI}) end fun tyrel_tac ctxt i j = (insert_tyrel_tac ctxt i j THEN_ALL_NEW_FWD solve_tyrel_tac ctxt) i fun internal_hom_tac ctxt = let val hom_net = hom_netD.get ctxt in Seq.INTERVAL (TRY o DETERM o resolve_from_net_tac ctxt hom_net) end fun internal_spec_tac ctxt = let val pairs = thm_pairsD.get ctxt val net = pairs |> map_filter (fst #> map_option (snd #> mk_CONSTRAINT_rl_atom ctxt)) |> Tactic.build_net in fn i => fn j => REPEAT (CHANGED (Seq.INTERVAL (DETERM o Anti_Unification.specialize_net_tac ctxt net) i j) ) end fun apply_to_constraints tac = let fun no_CONSTRAINT_tac i st = case Logic.concl_of_goal (Thm.prop_of st) i of @{mpat "Trueprop (CONSTRAINT _ _)"} => Seq.empty | _ => Seq.single st in Seq.INTERVAL (no_CONSTRAINT_tac ORELSE' tac) end fun internal_solve_tac ctxt = let val pairs = thm_pairsD.get ctxt val net = pairs |> map_filter (fst #> map_option (mk_CONSTRAINT_rl ctxt)) |> Tactic.build_net val s_tac = SOLVED' (REPEAT_ALL_NEW (resolve_from_net_tac ctxt net)) in apply_to_constraints s_tac ORELSE_INTERVAL apply_to_constraints (TRY o DETERM o s_tac) end fun guess_relators_tac ctxt = let val pairs = thm_pairsD.get ctxt val net = pairs |> map_filter (fst #> map_option (mk_CONSTRAINT_rl ctxt)) |> Tactic.build_net val hom_net = hom_netD.get ctxt fun hom_tac i j = Seq.INTERVAL (TRY o DETERM o resolve_from_net_tac ctxt hom_net) i j fun spec_tac i j = REPEAT (CHANGED (Seq.INTERVAL (DETERM o Anti_Unification.specialize_net_tac ctxt net) i j) ) val solve_tac = let val s_tac = SOLVED' (REPEAT_ALL_NEW (resolve_from_net_tac ctxt net)) in apply_to_constraints s_tac ORELSE_INTERVAL apply_to_constraints (TRY o DETERM o s_tac) end in Seq.INTERVAL (insert_CONSTRAINTS_tac ctxt) THEN_INTERVAL hom_tac THEN_INTERVAL spec_tac THEN_INTERVAL (tyrel_tac ctxt) THEN_INTERVAL solve_tac end (*********************) (* Pretty Printing *) (*********************) fun pretty_constraint_atom ctxt (f,R) = Pretty.block [ Syntax.pretty_term ctxt f, Pretty.str " :: ", Syntax.pretty_typ ctxt (fastype_of f), Pretty.str " ::: ", Syntax.pretty_term ctxt R] fun pretty_constraint ctxt (ps,(f,R)) = case ps of [] => pretty_constraint_atom ctxt (f,R) | _ => Pretty.block [ map (pretty_constraint_atom ctxt) ps |> Pretty.separate "; " |> Pretty.enclose "\" "\", Pretty.brk 1, Pretty.str "\", Pretty.brk 1, pretty_constraint_atom ctxt (f,R) ] fun pretty_constraints ctxt l = Pretty.big_list "Constraints" (map (pretty_constraint ctxt) l) fun pretty_thm_pair ctxt (c,thm) = Pretty.block [ case c of NONE => Pretty.str "NONE" | SOME c => pretty_constraint ctxt c, Pretty.brk 2, Pretty.str "---", Pretty.brk 2, Thm.pretty_thm ctxt thm ] fun pretty_thm_pairs ctxt pairs = Pretty.big_list "Thm-Pairs" (map (pretty_thm_pair ctxt) pairs) local fun unifies ctxt (t1, t2) = Term.could_unify (t1, t2) andalso let val idx1 = Term.maxidx_of_term t1 val t2 = Logic.incr_indexes ([], [], idx1 + 1) t2 val idx2 = Term.maxidx_of_term t2 in can (Pattern.unify (Context.Proof ctxt) (t1,t2)) (Envir.empty idx2) end fun analyze_possible_problems ctxt (f,R) = let fun strange_aux sf R = ( if sf then let val T = fastype_of R in case try (HOLogic.dest_prodT o HOLogic.dest_setT) T of SOME _ => [] | NONE => [Pretty.block [ Pretty.str "Strange relator type, expected plain relation: ", Syntax.pretty_term (Config.put show_types true ctxt) R ]] end else [] ) @ ( case R of @{mpat "\?R\?S"} => strange_aux true R @ strange_aux false S | Var (_,T) => ( case try (HOLogic.dest_prodT o HOLogic.dest_setT) (#2 (strip_type T)) of SOME (TFree _,_) => [Pretty.block [ Pretty.str "Fixed concrete type on schematic relator: ", Syntax.pretty_term (Config.put show_types true ctxt) R ]] | _ => [] ) | _ => [] ) val strange = case strange_aux true R of [] => NONE | l => SOME (Pretty.block l) val folded_relator = let fun match (Type (name,args)) R = let val (Rargs,Rhd) = Relators.strip_relAPP R in if is_Var Rhd then [] else if length args <> length Rargs then [Pretty.block [ Pretty.str "Type/relator arity mismatch:", Pretty.brk 1, Pretty.block [ Pretty.str name, Pretty.str "/", Pretty.str (string_of_int (length args)) ], Pretty.brk 1,Pretty.str "vs.",Pretty.brk 1, Pretty.block [ Syntax.pretty_term ctxt Rhd, Pretty.str "/", Pretty.str (string_of_int (length Rargs)) ] ]] else args ~~ Rargs |> map (uncurry match) |> flat end | match _ _ = [] in case match (fastype_of f) R of [] => NONE | l => SOME (Pretty.block (Pretty.fbreaks l @ [Pretty.fbrk, Pretty.str ("Explanation: This may be due to using polymorphic " ^ "relators like Id on non-terminal types." ^ "A problem usually occurs when " ^ "this relator has to be matched against a fully unfolded one." ^ "This warning is also issued on partially parametric relators " ^ "like br. However, the refinement rules are usually set up to " ^ "compensate for this, so this is probably not the cause for an " ^ "unsolved constraint") ])) end val issues = [strange, folded_relator] |> map_filter I in case issues of [] => NONE | l => SOME (Pretty.big_list "Possible problems" l) end fun pretty_try_candidates ctxt i st = if i > Thm.nprems_of st then Pretty.str "Goal number out of range" else case Logic.concl_of_goal (Thm.prop_of st) i of @{mpat "Trueprop (CONSTRAINT ?f ?R)"} => let val pairs = thm_pairsD.get ctxt val st = Drule.zero_var_indexes st val pt_hd = Pretty.block [ Pretty.str "Head: ", Pretty.fbrk, pretty_constraint_atom ctxt (f,R) ] fun isc (SOME (ps,(fp,R)),_) = if unifies ctxt (f,fp) then SOME (ps,(fp,R)) else NONE | isc _ = NONE val candidates = pairs |> map_filter isc fun try_c c = let val pt1 = Pretty.block [ Pretty.str "Trying ", pretty_constraint ctxt c ] val rl = mk_CONSTRAINT_rl ctxt c |> Drule.zero_var_indexes val res = (SOLVED' (resolve_tac ctxt [rl])) i st |> Seq.pull |> is_some val pt2 = (if res then Pretty.str "OK" else Pretty.str "ERR") in Pretty.block [pt1,Pretty.fbrk,pt2] end val res = Pretty.block ( Pretty.fbreaks [pt_hd, Pretty.big_list "Solving Attempts" (map try_c candidates)] ) in res end | _ => Pretty.str "Unexpected goal format" exception ERR of Pretty.T fun analyze' ctxt i j st = let val As = Logic.strip_horn (Thm.prop_of st) |> #1 |> drop (i-1) |> take (j-i+1) |> map (strip_all_body #> Logic.strip_imp_concl) val Cs = map_filter ( fn @{mpat "Trueprop (CONSTRAINT ?f ?R)"} => SOME (f,R) | @{mpat "Trueprop ((_,_)\_)"} => NONE | t => raise ERR (Pretty.block [ Pretty.str "Internal: Unexpected goal format: ", Syntax.pretty_term ctxt t ]) ) As val Cs_problems = map (fn c => case analyze_possible_problems ctxt c of NONE => pretty_constraint_atom ctxt c | SOME p => Pretty.block [pretty_constraint_atom ctxt c,Pretty.fbrk,p] ) Cs val Cs_pretty = Pretty.big_list "Constraints" Cs_problems in case Cs of [] => () | _ => raise ERR (Pretty.block [ Pretty.str "Could not infer all relators, some constraints remaining", Pretty.fbrk, Cs_pretty, Pretty.fbrk, Pretty.block [ Pretty.str "Trying to solve first constraint", Pretty.fbrk, pretty_try_candidates ctxt i st ] ]) end in fun analyze ctxt i j st = can (analyze' ctxt i j) st fun pretty_failure ctxt i j st = (analyze' ctxt i j st; Pretty.str "No failure") handle ERR p => p fun try_solve_tac ctxt i st = if i > Thm.nprems_of st then (tracing "Goal number out of range"; Seq.empty) else case Logic.concl_of_goal (Thm.prop_of st) i of @{mpat "Trueprop (CONSTRAINT ?f ?R)"} => let val pairs = thm_pairsD.get ctxt val st = Drule.zero_var_indexes st val pt = Pretty.block [ Pretty.str "Head: ", Pretty.fbrk, pretty_constraint_atom ctxt (f,R) ] val _ = tracing (Pretty.string_of pt) val _ = case analyze_possible_problems ctxt (f,R) of NONE => () | SOME p => tracing (Pretty.string_of p) fun isc (SOME (ps,(fp,R)),_) = if unifies ctxt (f,fp) then SOME (ps,(fp,R)) else NONE | isc _ = NONE val net = pairs |> map_filter (fst #> map_option (mk_CONSTRAINT_rl ctxt)) |> Tactic.build_net val candidates = pairs |> map_filter isc fun try_c c = let val _ = Pretty.block [ Pretty.str "Trying ", pretty_constraint ctxt c ] |> Pretty.string_of |> tracing val rl = mk_CONSTRAINT_rl ctxt c |> Drule.zero_var_indexes val res = (SOLVED' (resolve_tac ctxt [rl] THEN_ALL_NEW (REPEAT_ALL_NEW (resolve_from_net_tac ctxt net))) ) i st |> Seq.pull |> is_some val _ = (if res then Pretty.str "OK" else Pretty.str "ERR") |> Pretty.string_of |> tracing in () end val _ = map try_c candidates in Seq.single st end | _ => Seq.empty end fun solve_step_tac ctxt = let val pairs = thm_pairsD.get ctxt val net = pairs |> map_filter (fst #> map_option (mk_CONSTRAINT_rl ctxt)) |> Tactic.build_net in resolve_from_net_tac ctxt net end val phase = { init = thm_pairsD.init #> hom_netD.init, tac = guess_relators_tac, analyze = analyze, pretty_failure = pretty_failure } val setup = hom_rules.setup #> tyrel_rules.setup end \ setup Autoref_Fix_Rel.setup end diff --git a/thys/Automatic_Refinement/Tool/Autoref_Id_Ops.thy b/thys/Automatic_Refinement/Tool/Autoref_Id_Ops.thy --- a/thys/Automatic_Refinement/Tool/Autoref_Id_Ops.thy +++ b/thys/Automatic_Refinement/Tool/Autoref_Id_Ops.thy @@ -1,692 +1,693 @@ section \Operation Identification\ theory Autoref_Id_Ops imports "../Lib/Refine_Lib" Autoref_Phases Autoref_Data Autoref_Tagging "../Parametricity/Parametricity" begin subsection \Main Tool\ typedecl interface definition intfAPP :: "(interface \ _) \ interface \ _" where "intfAPP f x \ f x" syntax "_intf_APP" :: "args \ 'a \ 'b" ("\_\\<^sub>i_" [0,900] 900) translations "\x,xs\\<^sub>iR" == "\xs\\<^sub>i(CONST intfAPP R x)" "\x\\<^sub>iR" == "CONST intfAPP R x" consts i_fun :: "interface \ interface \ interface" abbreviation i_fun_app (infixr "\\<^sub>i" 60) where "i1\\<^sub>ii2 \ \i1,i2\\<^sub>ii_fun" consts i_annot :: "interface \ annot" abbreviation i_ANNOT :: "'a \ interface \ 'a" (infixr ":::\<^sub>i" 10) where "t:::\<^sub>iI \ ANNOT t (i_annot I)" text \Declaration of interface-type for constant\ definition CONST_INTF :: "'a \ interface \ bool" (infixr "::\<^sub>i" 10) where [simp]: "c::\<^sub>i I \ True" text \ Predicate for operation identification. \ID_OP t t' I\ means that term \t\ has been annotated as \t'\, and its interface is \I\. \ definition ID_OP :: "'a \ 'a \ interface \ bool" where [simp]: "ID_OP t t' I \ t=t'" text \ Interface inference rules. Caution: Some of these must be applied with custom unification! \ lemma ID_abs: \ \Tag abs first\ "\ \x. ID_OP x x I1 \ ID_OP (f x) (f' x) I2 \ \ ID_OP (\'x. f x) (\'x. f' x) (I1\\<^sub>iI2)" by simp lemma ID_app: \ \Tag app first\ "\ INDEP I1; ID_OP x x' I1; ID_OP f f' (I1\\<^sub>iI2) \ \ ID_OP (f$x) (f'$x') I2" by simp lemma ID_const: \ \Only if c is constant or free variable\ "\ c ::\<^sub>i I \ \ ID_OP c (OP c :::\<^sub>i I) I" by simp definition [simp]: "ID_TAG x \ x" lemma ID_const_any: \ \Only if no typing for constant exists\ "ID_OP c (OP (ID_TAG c) :::\<^sub>i I) I" by simp lemma ID_const_check_known: "\ c ::\<^sub>i I' \ \ ID_OP c c I" by simp lemma ID_tagged_OP: \ \Try first\ "ID_OP (OP f :::\<^sub>i I) (OP f :::\<^sub>i I) I" by simp lemma ID_is_tagged_OP: "ID_OP (OP c) t' I \ ID_OP (OP c) t' I" . lemma ID_tagged_OP_no_annot: "c ::\<^sub>i I \ ID_OP (OP c) (OP c :::\<^sub>i I) I" by simp lemmas ID_tagged = ID_tagged_OP ID_abs ID_app lemma ID_annotated: \ \Try second\ "ID_OP t t' I \ ID_OP (t :::\<^sub>i I) t' I" "ID_OP t t' I \ ID_OP (ANNOT t A) (ANNOT t' A) I" by simp_all lemma ID_init: assumes "ID_OP a a' I" assumes "(c,a')\R" shows "(c,a)\R" using assms by auto lemma itypeI: "(c::'t) ::\<^sub>i I" by simp consts depth_limit_dummy :: 'a notation (output) depth_limit_dummy ("\") ML \ fun limit_depth _ (t as Const _) = t | limit_depth _ (t as Var _) = t | limit_depth _ (t as Free _) = t | limit_depth _ (t as Bound _) = t | limit_depth 0 t = Const (@{const_name depth_limit_dummy},fastype_of t) | limit_depth i (t as _$_) = let val (f,args) = strip_comb t val f = limit_depth (i - 1) f val args = map (limit_depth (i - 1)) args in list_comb (f,args) end | limit_depth i (Abs (x,T,t)) = Abs (x,T,limit_depth (i - 1) t) fun depth_of (t as _$_) = let val (f,args) = strip_comb t in Integer.max (depth_of f) (fold (Integer.max o depth_of) args 0) + 1 end | depth_of (Abs (_,_,t)) = depth_of t + 1 | depth_of _ = 0 val depth_of_lhs = depth_of o Thm.term_of o Thm.lhs_of val depth_of_rhs = depth_of o Thm.term_of o Thm.rhs_of fun pretty_rewrite ctxt thm rthm = let val lhsd = depth_of_lhs thm val t = Thm.lhs_of rthm |> Thm.term_of |> limit_depth lhsd val rhsd = depth_of_rhs thm val t' = Thm.rhs_of rthm |> Thm.term_of |> limit_depth rhsd in Pretty.block [ Syntax.pretty_term ctxt t, Pretty.brk 1, Pretty.str "->", Pretty.brk 1, Syntax.pretty_term ctxt t' ] end \ ML_val \ depth_of @{term "f [1] [2] []"}; limit_depth 2 @{term "[1,2,3,4,5,6,7]"} |> Thm.cterm_of @{context} \ ML \ fun index_rewr_thms thms = let fun lhs thm = case Thm.concl_of thm of @{mpat "?lhs == _"} => [lhs] | _ => [] val net = Item_Net.init Thm.eq_thm_prop lhs val net = fold_rev Item_Net.update thms net in net end fun net_rewr_tac net get_pat frame_conv ctxt = IF_EXGOAL ( fn i => fn st => let val g = Logic.concl_of_goal (Thm.prop_of st) i |> get_pat val thms = Item_Net.retrieve net g val cnv = map (fn thm => CONVERSION (frame_conv (Conv.rewr_conv thm) ctxt)) thms |> APPEND_LIST' in cnv i st end ) \ ML \ signature AUTOREF_ID_OPS = sig val id_tac: Proof.context -> tactic' val id_phase: Autoref_Phases.phase val mk_const_intf: term -> term -> term val mk_const_intf_thm: Proof.context -> term -> term -> thm val dest_const_intf: term -> term * term val dest_const_intf_thm: thm -> term * term val cfg_trace_intf_unif: bool Config.T val cfg_trace_failed_id: bool Config.T val cfg_ss_id_op: bool Config.T val cfg_trace_patterns: bool Config.T val cfg_use_id_tags: bool Config.T val cfg_trace_id_tags: bool Config.T val typ_thms_of_seq: Proof.context -> term -> thm Seq.seq val has_typ_thms: Proof.context -> term -> bool val decl_derived_typing: bool -> term -> term -> Context.generic -> Context.generic val setup: theory -> theory end structure Autoref_Id_Ops :AUTOREF_ID_OPS = struct open Refine_Util Autoref_Tagging fun mk_const_intf c I = let val Tc = fastype_of c val T = Tc --> @{typ interface} --> @{typ bool} in Const (@{const_name CONST_INTF},T)$c$I end fun mk_const_intf_thm ctxt f I = let val fT = fastype_of f |> Thm.ctyp_of ctxt val f = Thm.cterm_of ctxt f val I = Thm.cterm_of ctxt I val thm = Thm.instantiate' [SOME fT] [SOME f, SOME I] @{thm itypeI} in thm end fun dest_const_intf @{mpat "?c::\<^sub>i?I"} = (c,I) | dest_const_intf t = raise TERM ("dest_const_intf",[t]) val dest_const_intf_thm = Thm.concl_of #> HOLogic.dest_Trueprop #> dest_const_intf fun LHS_COND' P = CONCL_COND' (fn @{mpat "Trueprop (ID_OP ?lhs _ _)"} => P lhs | _ => false) local open Conv in fun id_op_lhs_conv cnv ct = case Thm.term_of ct of @{mpat "ID_OP _ _ _"} => (fun_conv (fun_conv (arg_conv cnv))) ct | _ => raise CTERM ("id_op_lhs_conv",[ct]) end structure intf_types = Named_Thms ( val name = @{binding autoref_itype} val description = "Interface type declaration" ) structure op_patterns = Named_Thms ( val name = @{binding autoref_op_pat} val description = "Operation patterns" ) structure op_patterns_def = Named_Thms ( val name = @{binding autoref_op_pat_def} val description = "Definitive operation patterns" ) val cfg_trace_intf_unif = Attrib.setup_config_bool @{binding autoref_trace_intf_unif} (K false) val cfg_trace_failed_id = Attrib.setup_config_bool @{binding autoref_trace_failed_id} (K false) val cfg_ss_id_op = Attrib.setup_config_bool @{binding autoref_ss_id_op} (K false) val cfg_trace_patterns = Attrib.setup_config_bool @{binding autoref_trace_pat} (K false) val cfg_use_id_tags = Attrib.setup_config_bool @{binding autoref_use_id_tags} (K false) val cfg_trace_id_tags = Attrib.setup_config_bool @{binding autoref_trace_id_tags} (K false) fun get_typ_net ctxt = let val thy = Proof_Context.theory_of ctxt val typ_net = intf_types.get ctxt |> Refine_Util.subsume_sort Thm.concl_of thy |> Tactic.build_net in typ_net end fun typ_thms_of_seq' ctxt typ_net c = let val idx = Term.maxidx_of_term c + 1 val typ_thms = mk_const_intf c (Var (("I",idx),@{typ interface})) |> HOLogic.mk_Trueprop |> Thm.cterm_of ctxt |> Goal.init |> resolve_from_net_tac ctxt typ_net 1 |> Seq.map Goal.conclude in typ_thms end fun typ_thms_of_seq ctxt = typ_thms_of_seq' ctxt (get_typ_net ctxt) fun has_typ_thms' thy typ_net = typ_thms_of_seq' thy typ_net #> Seq.pull #> is_some fun has_typ_thms ctxt = has_typ_thms' ctxt (get_typ_net ctxt) fun id_tac ctxt = let val thy = Proof_Context.theory_of ctxt val typ_net = get_typ_net ctxt val typ_thms_seq_of = typ_thms_of_seq' ctxt typ_net val typ_thms_of = typ_thms_seq_of #> Seq.list_of fun pretty_typ_thms l = l |> map (Thm.pretty_thm ctxt) |> Pretty.fbreaks |> Pretty.block val tr_iu = if Config.get ctxt cfg_trace_intf_unif then fn i => fn st => (( case Logic.concl_of_goal (Thm.prop_of st) i of (t as @{mpat "Trueprop (?c::\<^sub>i_)"}) => ( case typ_thms_of c of [] => () | tts => Pretty.block [ Pretty.block [ Pretty.str "Interface unification failed:", Pretty.brk 1, Syntax.pretty_term ctxt t ], Pretty.fbrk, Pretty.str " ", Pretty.block [ Pretty.str "Candidates: ", Pretty.fbrk, pretty_typ_thms tts ] ] |> Pretty.string_of |> tracing ) | _ => () ); Seq.empty) else K no_tac val id_typ = resolve_from_net_tac ctxt typ_net ORELSE' tr_iu val pat_net = op_patterns.get ctxt |> Refine_Util.subsume_sort (Thm.concl_of #> Logic.dest_equals #> #1) thy |> index_rewr_thms val def_pat_net = op_patterns_def.get ctxt |> Refine_Util.subsume_sort (Thm.concl_of #> Logic.dest_equals #> #1) thy |> index_rewr_thms val id_abs = CONVERSION (HOL_concl_conv (fn ctxt => id_op_lhs_conv (mk_ABS_conv ctxt)) ctxt) THEN' resolve_tac ctxt @{thms ID_abs} val id_app = CONVERSION (HOL_concl_conv (fn _ => id_op_lhs_conv (mk_APP_conv)) ctxt) THEN' resolve_tac ctxt @{thms ID_app} val id_tag_tac = let val trace = Config.get ctxt cfg_trace_id_tags in if trace then IF_EXGOAL (fn i => fn st => let fun tr_tac _ st' = let val goal = Logic.get_goal (Thm.prop_of st) i val concl = Logic.concl_of_goal (Thm.prop_of st) i val _ = case concl of @{mpat "Trueprop (ID_OP ?lhs _ _)"} => tracing ("ID_TAG: " ^ @{make_string} lhs) | _ => tracing "ID_TAG: LHS???" val _ = Pretty.block [ Pretty.str "ID_TAG: ", Pretty.brk 1, Syntax.pretty_term ctxt goal ] |> Pretty.string_of |> tracing in Seq.single st' end in (resolve_tac ctxt @{thms ID_const_any} THEN' tr_tac) i st end) else resolve_tac ctxt @{thms ID_const_any} end val id_const = LHS_COND' (fn t => is_Const t orelse is_Free t) THEN' ( (resolve_tac ctxt @{thms ID_const} THEN' id_typ) (* Try to find type *) ORELSE' ( if Config.get ctxt cfg_use_id_tags then CAN' (resolve_tac ctxt @{thms ID_const_check_known} THEN' id_typ) THEN_ELSE' ( K no_tac, id_tag_tac ) else K no_tac ) ) val id_tagged = resolve_tac ctxt @{thms ID_tagged} val id_annotated = resolve_tac ctxt @{thms ID_annotated} (* val traced_rewr_conv = if Config.get ctxt cfg_trace_patterns then fn ctxt => fn thm => fn ct => let val rthm = Conv.rewr_conv thm ct val _ = Pretty.block [ Pretty.str "Trying (-pat:)", Pretty.brk 1, pretty_rewrite ctxt thm rthm ] |> Pretty.string_of |> tracing in rthm end else K Conv.rewr_conv *) fun get_rewr_pat @{mpat "Trueprop (ID_OP ?lhs _ _)"} = lhs | get_rewr_pat t = t fun rewr_frame_conv conv = HOL_concl_conv (fn _ => id_op_lhs_conv conv) val def_id_pat = DETERM o net_rewr_tac def_pat_net get_rewr_pat rewr_frame_conv ctxt val id_pat = net_rewr_tac pat_net get_rewr_pat rewr_frame_conv ctxt val id_dflt = FIRST' [id_app,id_const,id_abs] val id_fail = if Config.get ctxt cfg_trace_failed_id then IF_EXGOAL (fn i => fn st => let val pat = Logic.concl_of_goal (Thm.prop_of st) i |> get_rewr_pat val _ = Pretty.block [ Pretty.str "Failed to identify: ", Syntax.pretty_term ctxt pat ] |> Pretty.string_of |> tracing in Seq.empty end ) else K no_tac val ss = Config.get ctxt cfg_ss_id_op val step_tac = FIRST' [ assume_tac ctxt, id_tagged, resolve_tac ctxt @{thms ID_is_tagged_OP} THEN_ELSE' ( resolve_tac ctxt @{thms ID_tagged_OP_no_annot} THEN' id_typ, FIRST' [ Indep_Vars.indep_tac ctxt, id_annotated, def_id_pat, id_pat APPEND' id_dflt, id_fail ] ) ] fun rec_tac i st = ( step_tac THEN_ALL_NEW_FWD (if ss then K all_tac else rec_tac) ) i st in rec_tac end fun id_analyze _ i j _ = (i = j) fun id_pretty_failure _ i j _ = if i = j then Pretty.str "No failure" else Pretty.str "Interface typing error. Enable tracing for more information" val id_phase = { init = I, tac = (fn ctxt => Seq.INTERVAL (resolve_tac ctxt @{thms ID_init} THEN' id_tac ctxt)), analyze = id_analyze, pretty_failure = id_pretty_failure } fun decl_derived_typing overl c I context = let val ctxt = Context.proof_of context val typ_thms = intf_types.get ctxt (* TODO: Use net cached in ctxt here! *) val thm = mk_const_intf_thm ctxt c I val st = Thm.cprop_of thm |> Goal.init val has_t = SOLVED' (match_tac ctxt typ_thms) 1 st |> Seq.pull |> is_some in if has_t then context else ( not overl andalso has_typ_thms ctxt c andalso (warning ( Pretty.block [ Pretty.str "Adding overloaded interface type to constant:", Pretty.brk 1, Thm.pretty_thm ctxt thm ] |> Pretty.string_of ); true); intf_types.add_thm thm context ) end val setup = I #> intf_types.setup #> op_patterns.setup #> op_patterns_def.setup end \ setup Autoref_Id_Ops.setup definition IND_FACT :: "rel_name \ ('c \ 'a) set \ bool" ("#_=_" 10) where [simp]: "#name=R \ True" lemma REL_INDIRECT: "#name=R" by simp definition CNV_ANNOT :: "'a \ 'a \ (_\'a) set \ bool" where [simp]: "CNV_ANNOT t t' R \ t=t'" definition REL_OF_INTF :: "interface \ ('c\'a) set \ bool" where [simp]: "REL_OF_INTF I R \ True" definition [simp]: "REL_OF_INTF_P I R \ True" \ \Version to resolve relator arguments\ lemma CNV_ANNOT: "\f f' a a'. \ CNV_ANNOT a a' Ra; CNV_ANNOT f f' (Ra\Rr) \ \ CNV_ANNOT (f$a) (f'$a') (Rr)" "\f f'. \ \x. CNV_ANNOT x x Ra \ CNV_ANNOT (f x) (f' x) Rr \ \ CNV_ANNOT (\'x. f x) (\'x. f' x) (Ra\Rr)" "\f f I R. \undefined (''Id tag not yet supported'',f)\ \ CNV_ANNOT (OP (ID_TAG f) :::\<^sub>i I) f R" "\f I R. \ INDEP R; REL_OF_INTF I R \ \ CNV_ANNOT (OP f :::\<^sub>i I) (OP f ::: R) R" "\t t' R. CNV_ANNOT t t' R \ CNV_ANNOT (t ::: R) t' R" "\t t' name R. \ #name=R; CNV_ANNOT t t' R \ \ CNV_ANNOT (t ::#name) t' R" by simp_all consts i_of_rel :: "'a \ 'b" lemma ROI_P_app: \ \Only if interface is really application\ "REL_OF_INTF_P I R \ REL_OF_INTF I R" by auto lemma ROI_app: \ \Only if interface is really application\ "\ REL_OF_INTF I R; REL_OF_INTF_P J S \ \ REL_OF_INTF_P (\I\\<^sub>iJ) (\R\S)" by auto lemma ROI_i_of_rel: "REL_OF_INTF_P (i_of_rel S) S" "REL_OF_INTF (i_of_rel R) R" by auto lemma ROI_const: "REL_OF_INTF_P J S" "REL_OF_INTF I R" by auto lemma ROI_init: assumes "CNV_ANNOT a a' R" assumes "(c,a')\R" shows "(c,a)\R" using assms by simp lemma REL_OF_INTF_I: "REL_OF_INTF I R" by simp ML \ signature AUTOREF_REL_INF = sig val roi_tac: Proof.context -> tactic' val roi_step_tac: Proof.context -> tactic' val roi_phase: Autoref_Phases.phase val cfg_sbias: int Config.T val setup: theory -> theory end structure Autoref_Rel_Inf :AUTOREF_REL_INF = struct val cfg_sbias = Attrib.setup_config_int @{binding autoref_sbias} (K 100) structure rel_indirect = Named_Thms ( val name = @{binding autoref_rel_indirect} val description = "Indirect relator bindings" ) fun rel_of_intf_thm ctxt I = let fun roi (Free (n,_)) ctxt = let val (Rn,ctxt) = yield_singleton Variable.variant_fixes ("R_" ^ n) ctxt in (Free (Rn,dummyT),ctxt) end | roi (Const (n,_)) ctxt = let val n = Long_Name.base_name n val (Rn,ctxt) = yield_singleton Variable.variant_fixes ("R_" ^ n) ctxt in (Free (Rn,dummyT),ctxt) end | roi @{mpat "\?Ia\\<^sub>i?Ib"} ctxt = let val (Ra,ctxt) = roi Ia ctxt val (Rb,ctxt) = roi Ib ctxt in (Const (@{const_name relAPP},dummyT)$Rb$Ra,ctxt) end | roi @{mpat "i_of_rel ?R"} ctxt = (R,ctxt) | roi t _ = raise TERM ("rel_of_intf: Unexpected interface", [t]) val orig_ctxt = ctxt val (I,ctxt) = yield_singleton (Variable.import_terms true) I ctxt val (R,ctxt) = roi I ctxt val res = Const (@{const_name REL_OF_INTF},dummyT)$I$R val res = Syntax.check_term ctxt res val res = singleton (Variable.export_terms ctxt orig_ctxt) res |> HOLogic.mk_Trueprop - |> Thm.cterm_of ctxt - - val thm = Goal.prove_internal ctxt [] res (fn _ => resolve_tac ctxt @{thms REL_OF_INTF_I} 1) + val goal_ctxt = Variable.declare_term res ctxt + val thm = + Goal.prove_internal ctxt [] (Thm.cterm_of ctxt res) + (K (resolve_tac goal_ctxt @{thms REL_OF_INTF_I} 1)) in thm end fun roi_step_tac ctxt = let val ind_net = rel_indirect.get ctxt |> Tactic.build_net in IF_EXGOAL ( assume_tac ctxt ORELSE' Indep_Vars.indep_tac ctxt ORELSE' resolve_from_net_tac ctxt ind_net ORELSE' (fn i => fn st => case Logic.concl_of_goal (Thm.prop_of st) i of @{mpat "Trueprop (CNV_ANNOT _ _ _)"} => resolve_tac ctxt @{thms CNV_ANNOT} i st | @{mpat "Trueprop (REL_OF_INTF ?I _)"} => resolve_tac ctxt [rel_of_intf_thm ctxt I] i st | _ => Seq.empty (* | @{mpat "Trueprop (REL_OF_INTF (\_\\<^sub>i_) _)"} => rtac @{thm ROI_P_app} i st | @{mpat "Trueprop (REL_OF_INTF_P (\_\\<^sub>i_) _)"} => rtac @{thm ROI_app} i st | @{mpat "Trueprop (REL_OF_INTF_P (i_of_rel _) _)"} => resolve_tac ctxt @{thms ROI_i_of_rel} i st | @{mpat "Trueprop (REL_OF_INTF (i_of_rel _) _)"} => resolve_tac ctxt @{thms ROI_i_of_rel} i st | _ => resolve_tac ctxt @{thms ROI_const} i st *) ) ) end fun roi_tac ctxt = REPEAT_ALL_NEW_FWD (DETERM o roi_step_tac ctxt) fun roi_analyze _ i j _ = (i = j) fun roi_pretty_failure _ i j _ = if i = j then Pretty.str "No failure" else Pretty.str "Relator inference error" val roi_phase = { init = I, tac = (fn ctxt => Seq.INTERVAL (resolve_tac ctxt @{thms ROI_init} THEN' roi_tac ctxt)), analyze = roi_analyze, pretty_failure = roi_pretty_failure } val setup = rel_indirect.setup end \ setup Autoref_Rel_Inf.setup end diff --git a/thys/Automatic_Refinement/Tool/Autoref_Tool.thy b/thys/Automatic_Refinement/Tool/Autoref_Tool.thy --- a/thys/Automatic_Refinement/Tool/Autoref_Tool.thy +++ b/thys/Automatic_Refinement/Tool/Autoref_Tool.thy @@ -1,225 +1,225 @@ section \Automatic Refinement Tool\ theory Autoref_Tool imports Autoref_Translate Autoref_Gen_Algo Autoref_Relator_Interface begin subsection \Standard setup\ text \Declaration of standard phases\ declaration \fn phi => let open Autoref_Phases in I #> register_phase "id_op" 10 Autoref_Id_Ops.id_phase phi #> register_phase "rel_inf" 20 Autoref_Rel_Inf.roi_phase phi #> register_phase "fix_rel" 22 Autoref_Fix_Rel.phase phi #> register_phase "trans" 30 Autoref_Translate.trans_phase phi end \ (* declaration {* fn phi => let open Autoref_Phases in I #> register_phase "id_op" 10 Autoref_Id_Ops.id_ops_phase phi #> register_phase "rel_inf" 20 Autoref_Rel_Inf.hm_infer_rel_phase phi #> register_phase "fix_rel" 21 Autoref_Fix_Rel.phase phi #> register_phase "trans" 30 Autoref_Translate.trans_phase phi end *} *) text \Main method\ method_setup autoref = \let open Refine_Util val autoref_flags = parse_bool_config "trace" Autoref_Phases.cfg_trace || parse_bool_config "debug" Autoref_Phases.cfg_debug || parse_bool_config "keep_goal" Autoref_Phases.cfg_keep_goal val autoref_phases = Args.$$$ "phases" |-- Args.colon |-- Scan.repeat1 Args.name in parse_paren_lists autoref_flags |-- Scan.option (Scan.lift (autoref_phases)) >> ( fn phases => fn ctxt => SIMPLE_METHOD' ( ( case phases of NONE => Autoref_Phases.all_phases_tac | SOME names => Autoref_Phases.phases_tacN names ) (Autoref_Phases.init_data ctxt) (* TODO: If we want more fine-grained initialization here, solvers have to depend on phases, or on data that they initialize if necessary *) )) end \ "Automatic Refinement" subsection \Tools\ setup \ let fun higher_order_rl_of ctxt thm = case Thm.concl_of thm of @{mpat "Trueprop ((_,?t)\_)"} => let val (f,args) = strip_comb t in if length args = 0 then thm else let val cT = TVar(("'c",0), @{sort type}) val c = Var (("c",0),cT) val R = Var (("R",0), HOLogic.mk_setT (HOLogic.mk_prodT (cT, fastype_of f))) val goal = HOLogic.mk_mem (HOLogic.mk_prod (c,f), R) |> HOLogic.mk_Trueprop - |> Thm.cterm_of ctxt - - val res_thm = Goal.prove_internal ctxt [] goal (fn _ => - REPEAT (resolve_tac ctxt @{thms fun_relI} 1) - THEN (resolve_tac ctxt [thm] 1) - THEN (ALLGOALS (assume_tac ctxt)) - ) + val goal_ctxt = Variable.declare_term goal ctxt + val res_thm = + Goal.prove_internal ctxt [] (Thm.cterm_of ctxt goal) + (fn _ => + REPEAT (resolve_tac goal_ctxt @{thms fun_relI} 1) + THEN (resolve_tac goal_ctxt [thm] 1) + THEN (ALLGOALS (assume_tac goal_ctxt))) in res_thm end end | _ => raise THM("Expected autoref rule",~1,[thm]) val higher_order_rl_attr = Thm.rule_attribute [] (higher_order_rl_of o Context.proof_of) in Attrib.setup @{binding autoref_higher_order_rule} (Scan.succeed higher_order_rl_attr) "Autoref: Convert rule to higher-order form" end \ subsection \Advanced Debugging\ method_setup autoref_trans_step_keep = \ Scan.succeed (fn ctxt => SIMPLE_METHOD' ( Autoref_Translate.trans_dbg_step_tac (Autoref_Phases.init_data ctxt) )) \ "Single translation step, leaving unsolved side-coditions" method_setup autoref_trans_step = \ Scan.succeed (fn ctxt => SIMPLE_METHOD' ( Autoref_Translate.trans_step_tac (Autoref_Phases.init_data ctxt) )) \ "Single translation step" method_setup autoref_trans_step_only = \ Scan.succeed (fn ctxt => SIMPLE_METHOD' ( Autoref_Translate.trans_step_only_tac (Autoref_Phases.init_data ctxt) )) \ "Single translation step, not attempting to solve side-coditions" method_setup autoref_side = \ Scan.succeed (fn ctxt => SIMPLE_METHOD' ( Autoref_Translate.side_dbg_tac (Autoref_Phases.init_data ctxt) )) \ "Solve side condition, leave unsolved subgoals" method_setup autoref_try_solve = \ Scan.succeed (fn ctxt => SIMPLE_METHOD' ( Autoref_Fix_Rel.try_solve_tac (Autoref_Phases.init_data ctxt) )) \ "Try to solve constraint and trace debug information" method_setup autoref_solve_step = \ Scan.succeed (fn ctxt => SIMPLE_METHOD' ( Autoref_Fix_Rel.solve_step_tac (Autoref_Phases.init_data ctxt) )) \ "Single-step of constraint solver" method_setup autoref_id_op = \ Scan.succeed (fn ctxt => SIMPLE_METHOD' ( Autoref_Id_Ops.id_tac ctxt )) \ method_setup autoref_solve_id_op = \ Scan.succeed (fn ctxt => SIMPLE_METHOD' ( Autoref_Id_Ops.id_tac (Config.put Autoref_Id_Ops.cfg_ss_id_op false ctxt) )) \ ML \ structure Autoref_Debug = struct fun print_thm_pairs ctxt = let val ctxt = Autoref_Phases.init_data ctxt val p = Autoref_Fix_Rel.thm_pairsD_get ctxt |> Autoref_Fix_Rel.pretty_thm_pairs ctxt |> Pretty.string_of in warning p end fun print_thm_pairs_matching ctxt cpat = let val pat = Thm.term_of cpat val ctxt = Autoref_Phases.init_data ctxt val thy = Proof_Context.theory_of ctxt fun matches NONE = false | matches (SOME (_,(f,_))) = Pattern.matches thy (pat,f) val p = Autoref_Fix_Rel.thm_pairsD_get ctxt |> filter (matches o #1) |> Autoref_Fix_Rel.pretty_thm_pairs ctxt |> Pretty.string_of in warning p end end \ text \General casting-tag, that allows type-casting on concrete level, while being identity on abstract level.\ definition [simp]: "CAST \ id" lemma [autoref_itype]: "CAST ::\<^sub>i I \\<^sub>i I" by simp (* TODO: This idea does currently not work, b/c a homogeneity rule will be created from the (\x. x, CAST)\R \ R rule, which will always be applied first! As a workaround, we make the cast mandatory! *) (*lemma [autoref_rules]: assumes "PRIO_TAG_GEN_ALGO" shows "(\x. x,CAST) \ R \ R" by auto *) text \Hide internal stuff\ notation (input) rel_ANNOT (infix ":::\<^sub>r" 10) notation (input) ind_ANNOT (infix "::#\<^sub>r" 10) locale autoref_syn begin notation (input) APP (infixl "$" 900) notation (input) rel_ANNOT (infix ":::" 10) notation (input) ind_ANNOT (infix "::#" 10) notation OP ("OP") notation (input) ABS (binder "\''" 10) end no_notation (input) APP (infixl "$" 900) no_notation (input) ABS (binder "\''" 10) no_notation (input) rel_ANNOT (infix ":::" 10) no_notation (input) ind_ANNOT (infix "::#" 10) hide_const (open) PROTECT ANNOT OP APP ABS ID_FAIL rel_annot ind_annot end diff --git a/thys/Automatic_Refinement/Tool/Autoref_Translate.thy b/thys/Automatic_Refinement/Tool/Autoref_Translate.thy --- a/thys/Automatic_Refinement/Tool/Autoref_Translate.thy +++ b/thys/Automatic_Refinement/Tool/Autoref_Translate.thy @@ -1,327 +1,327 @@ theory Autoref_Translate imports Autoref_Fix_Rel begin subsection \Definitions\ subsubsection \APP and ABS Rules\ lemma autoref_ABS: "\ \x x'. (x,x')\Ra \ (c x, a x')\Rr \ \ (c, \'x. a x)\Ra\Rr" by auto lemma autoref_APP: "\ (c,a)\Ra\Rr; (x,x')\Ra \ \ (c$x, a $ x')\Rr" by (auto dest: fun_relD) (*lemma autoref_ANNOT: "\ (x,x')\R \ \ (x,x':::R)\R" by (auto simp: ANNOT_def) *) lemma autoref_beta: assumes "(c,a x)\R" shows "(c,(\'x. a x)$x)\R" using assms by auto lemmas dflt_trans_rules = autoref_beta autoref_ABS autoref_APP subsubsection \Side Conditions\ text \ Rules can have prefer and defer side-conditions. Prefer conditions must be solvable in order for the rule to apply, and defer conditions must hold after the rule has been applied and the recursive translations have been performed. Thus, prefer-conditions typically restrict on the abstract expression, while defer conditions restrict the translated expression. In order to solve the actual side conditions, we use the \Tagged_Solver\-infrastructure. The solvers are applied after the \PREFER\/\DEFER\ tag has been removed. \ text \ Tag to remove internal stuff from term. Before a prefer/defer side condition is evaluated, all terms inside these tags are purged from autoref-specific annotations, i.e., operator-annotations, relator annotations, and tagged applications. \ definition [simp, autoref_tag_defs]: "REMOVE_INTERNAL x \ x" text \Useful abbreviation to require some property that is not related to teh refinement framework\ abbreviation "PREFER nt \ \ PREFER_tag (nt (REMOVE_INTERNAL \))" abbreviation "DEFER nt \ \ DEFER_tag (nt (REMOVE_INTERNAL \))" definition [simp]: "REMOVE_INTERNAL_EQ a b \ a=b" lemma REMOVE_INTERNAL_EQI: "REMOVE_INTERNAL_EQ a a" by simp lemma autoref_REMOVE_INTERNAL_EQ: assumes "(c,a)\R" assumes "REMOVE_INTERNAL_EQ c c'" shows "(c',a)\R" using assms by simp ML \ signature AUTOREF_TACTICALS = sig val is_prefer_cond: int -> thm -> bool val is_defer_cond: int -> thm -> bool val REPEAT_INTERVAL: tactic' -> itactic val COND'': (int -> thm -> bool) -> tactic' -> tactic' -> tactic' val REPEAT_ON_SUBGOAL: tactic' -> tactic' val IF_SOLVED: tactic' -> tactic' -> tactic' -> tactic' end structure Autoref_Tacticals :AUTOREF_TACTICALS = struct fun is_prefer_cond i st = case Logic.concl_of_goal (Thm.prop_of st) i of @{mpat "Trueprop (PREFER_tag _)"} => true | _ => false fun is_defer_cond i st = case Logic.concl_of_goal (Thm.prop_of st) i of @{mpat "Trueprop (DEFER_tag _)"} => true | _ => false fun REPEAT_INTERVAL tac i j st = let val n = Thm.nprems_of st - (j - i) fun r st = ( COND (has_fewer_prems n) (all_tac) (tac i THEN_ELSE (r,all_tac)) ) st in r st end fun COND'' cond tac1 tac2 = IF_EXGOAL (fn i => fn st => if cond i st then tac1 i st else tac2 i st ) fun REPEAT_ON_SUBGOAL tac i = REPEAT_INTERVAL tac i i fun IF_SOLVED tac tac1 tac2 i st = let val n = Thm.nprems_of st in (tac i THEN COND (has_fewer_prems n) (tac1 i) (tac2 i)) st end end signature AUTOREF_TRANSLATE = sig type trans_net = (int * thm) Net.net val add_post_rule: thm -> Context.generic -> Context.generic val delete_post_rule: thm -> Context.generic -> Context.generic val get_post_rules: Proof.context -> thm list val compute_trans_net: Autoref_Fix_Rel.thm_pairs -> Proof.context -> trans_net val side_tac: Proof.context -> tactic' val side_dbg_tac: Proof.context -> tactic' val trans_step_only_tac: Proof.context -> tactic' val trans_dbg_step_tac: Proof.context -> tactic' val trans_step_tac: Proof.context -> tactic' val trans_tac: Proof.context -> itactic val trans_analyze: Proof.context -> int -> int -> thm -> bool val trans_pretty_failure: Proof.context -> int -> int -> thm -> Pretty.T val trans_phase: Autoref_Phases.phase val setup: theory -> theory end structure Autoref_Translate :AUTOREF_TRANSLATE = struct type trans_net = (int * thm) Net.net (**************************) (* Translation *) (**************************) structure autoref_post_simps = Named_Thms ( val name = @{binding autoref_post_simps} val description = "Refinement Framework: " ^ "Automatic refinement post simplification rules" ); val add_post_rule = autoref_post_simps.add_thm val delete_post_rule = autoref_post_simps.del_thm val get_post_rules = autoref_post_simps.get fun compute_trans_net thm_pairs _ = thm_pairs |> map #2 |> (fn thms => thms @ @{thms dflt_trans_rules}) |> Tactic.build_net structure trans_netD = Autoref_Data ( type T = trans_net fun compute ctxt = compute_trans_net (Autoref_Fix_Rel.thm_pairsD_get ctxt) ctxt val prereq = [Autoref_Fix_Rel.thm_pairsD_init] ) - fun REMOVE_INTERNAL_conv ctxt = + val REMOVE_INTERNAL_conv = Conv.top_sweep_conv ( - fn _ => fn ct => (case Thm.term_of ct of + fn ctxt => fn ct => (case Thm.term_of ct of @{mpat "REMOVE_INTERNAL _"} => Conv.rewr_conv @{thm REMOVE_INTERNAL_def} then_conv Autoref_Tagging.untag_conv ctxt | _ => Conv.no_conv) ct - ) ctxt + ) fun REMOVE_INTERNAL_tac ctxt = CONVERSION (REMOVE_INTERNAL_conv ctxt) fun side_tac ctxt = resolve_tac ctxt @{thms SIDEI} THEN' REMOVE_INTERNAL_tac ctxt THEN' Tagged_Solver.solve_tac ctxt fun side_dbg_tac ctxt = let val ctxt = Config.put Tagged_Solver.cfg_keep true ctxt in resolve_tac ctxt @{thms SIDEI} THEN' REMOVE_INTERNAL_tac ctxt THEN' TRY o Tagged_Solver.solve_tac ctxt end local open Autoref_Tacticals fun trans_rule_tac ctxt net = resolve_from_net_tac ctxt net THEN_ALL_NEW (TRY o match_tac ctxt [@{thm PRIO_TAGI}]) in (* Do not even attempt to solve side conditions *) fun trans_step_only_tac ctxt = let val net = trans_netD.get ctxt in ( COND'' is_defer_cond (K no_tac) (assume_tac ctxt ORELSE' trans_rule_tac ctxt net) ) end (* Leave unsolved side conditions *) fun trans_dbg_step_tac ctxt = let val net = trans_netD.get ctxt val s_tac = side_tac ctxt in DETERM o ( COND'' is_defer_cond (SOLVED' s_tac) ( assume_tac ctxt ORELSE' (trans_rule_tac ctxt net THEN_ALL_NEW_FWD COND'' is_prefer_cond (TRY o DETERM o SOLVED' s_tac) (K all_tac) ) ) ) end fun trans_step_tac ctxt = let val net = trans_netD.get ctxt val s_tac = side_tac ctxt in DETERM o ( COND'' is_defer_cond (SOLVED' s_tac) ( assume_tac ctxt ORELSE' (trans_rule_tac ctxt net THEN_ALL_NEW_FWD COND'' is_prefer_cond (DETERM o SOLVED' s_tac) (K all_tac) ) ) ) end fun trans_tac ctxt = let val ss = put_simpset HOL_basic_ss ctxt addsimps @{thms APP_def PROTECT_def ANNOT_def} addsimps get_post_rules ctxt val trans_opt_tac = resolve_tac ctxt @{thms autoref_REMOVE_INTERNAL_EQ} THEN' IF_SOLVED (REPEAT_ON_SUBGOAL (trans_step_tac ctxt)) (simp_tac ss THEN' resolve_tac ctxt @{thms REMOVE_INTERNAL_EQI}) (K all_tac) in Seq.INTERVAL trans_opt_tac end end fun trans_analyze _ i j _ = j < i fun trans_pretty_failure ctxt i j st = if j < i then Pretty.str "No failure" else let val goal = Logic.get_goal (Thm.prop_of st) i val concl = strip_all_body goal |> Logic.strip_imp_concl val msg = case concl of @{mpat "Trueprop (DEFER_tag _)"} => Pretty.str "Could not solve defered side condition" | @{mpat "Trueprop ((_,_)\_)"} => Pretty.str "Could not refine subterm" | _ => Pretty.str "Internal: Unexpected goal" in Pretty.block [ msg, Pretty.fbrk, Syntax.pretty_term ctxt goal ] end val trans_phase = { init = trans_netD.init, tac = trans_tac, analyze = trans_analyze, pretty_failure = trans_pretty_failure } val setup = I #> autoref_post_simps.setup end \ setup Autoref_Translate.setup subsubsection \Standard side-tactics\ text \Preconditions\ definition [simp, autoref_tag_defs]: "PRECOND_tag P == P" lemma PRECOND_tagI: "P ==> PRECOND_tag P" by simp abbreviation "SIDE_PRECOND P == PREFER PRECOND_tag P" declaration \ Tagged_Solver.declare_solver @{thms PRECOND_tagI} @{binding PRECOND} "Refinement: Solve preconditions" ( fn ctxt => SOLVED' ( SELECT_GOAL (auto_tac ctxt) ) ) \ text \Optional preconditions\ definition [simp, autoref_tag_defs]: "PRECOND_OPT_tag P == P" lemma PRECOND_OPT_tagI: "P ==> PRECOND_OPT_tag P" by simp abbreviation "SIDE_PRECOND_OPT P == PREFER PRECOND_OPT_tag P" declaration \ Tagged_Solver.declare_solver @{thms PRECOND_OPT_tagI} @{binding PRECOND_OPT} "Refinement: Solve optional preconditions" ( fn ctxt => SOLVED' (asm_full_simp_tac ctxt)) \ end diff --git a/thys/Bertrands_Postulate/Bertrand.thy b/thys/Bertrands_Postulate/Bertrand.thy --- a/thys/Bertrands_Postulate/Bertrand.thy +++ b/thys/Bertrands_Postulate/Bertrand.thy @@ -1,1854 +1,1859 @@ (* File: Bertrand.thy Authors: Julian Biendarra, Manuel Eberl , Larry Paulson A proof of Bertrand's postulate (based on John Harrison's HOL Light proof). Uses reflection and the approximation tactic. *) theory Bertrand imports Complex_Main "HOL-Number_Theory.Number_Theory" "HOL-Library.Discrete" "HOL-Decision_Procs.Approximation_Bounds" "HOL-Library.Code_Target_Numeral" Pratt_Certificate.Pratt_Certificate begin subsection \Auxiliary facts\ lemma ln_2_le: "ln 2 \ 355 / (512 :: real)" proof - have "ln 2 \ real_of_float (ub_ln2 12)" by (rule ub_ln2) also have "ub_ln2 12 = Float 5680 (- 13)" by code_simp finally show ?thesis by simp qed lemma ln_2_ge: "ln 2 \ (5677 / 8192 :: real)" proof - have "ln 2 \ real_of_float (lb_ln2 12)" by (rule lb_ln2) also have "lb_ln2 12 = Float 5677 (-13)" by code_simp finally show ?thesis by simp qed lemma ln_2_ge': "ln (2 :: real) \ 2/3" and ln_2_le': "ln (2 :: real) \ 16/23" using ln_2_le ln_2_ge by simp_all lemma of_nat_ge_1_iff: "(of_nat x :: 'a :: linordered_semidom) \ 1 \ x \ 1" using of_nat_le_iff[of 1 x] by (subst (asm) of_nat_1) lemma floor_conv_div_nat: "of_int (floor (real m / real n)) = real (m div n)" by (subst floor_divide_of_nat_eq) simp lemma frac_conv_mod_nat: "frac (real m / real n) = real (m mod n) / real n" by (cases "n = 0") (simp_all add: frac_def floor_conv_div_nat field_simps of_nat_mult [symmetric] of_nat_add [symmetric] del: of_nat_mult of_nat_add) lemma of_nat_prod_mset: "prod_mset (image_mset of_nat A) = of_nat (prod_mset A)" by (induction A) simp_all lemma prod_mset_pos: "(\x :: 'a :: linordered_semidom. x \# A \ x > 0) \ prod_mset A > 0" by (induction A) simp_all lemma ln_msetprod: assumes "\x. x \#I \ x > 0" shows "(\p::nat\#I. ln p) = ln (\p\#I. p)" using assms by (induction I) (simp_all add: of_nat_prod_mset ln_mult prod_mset_pos) lemma ln_fact: "ln (fact n) = (\d=1..n. ln d)" by (induction n) (simp_all add: ln_mult) lemma overpower_lemma: fixes f g :: "real \ real" assumes "f a \ g a" assumes "\x. a \ x \ ((\x. g x - f x) has_real_derivative (d x)) (at x)" assumes "\x. a \ x \ d x \ 0" assumes "a \ x" shows "f x \ g x" proof (cases "a < x") case True with assms have "\z. z > a \ z < x \ g x - f x - (g a - f a) = (x - a) * d z" by (intro MVT2) auto then obtain z where z: "z > a" "z < x" "g x - f x - (g a - f a) = (x - a) * d z" by blast hence "f x = g x + (f a - g a) + (a - x) * d z" by (simp add: algebra_simps) also from assms have "f a - g a \ 0" by (simp add: algebra_simps) also from assms z have "(a - x) * d z \ 0 * d z" by (intro mult_right_mono) simp_all finally show ?thesis by simp qed (insert assms, auto) subsection \Preliminary definitions\ definition primepow_even :: "nat \ bool" where "primepow_even q \ (\ p k. 1 \ k \ prime p \ q = p^(2*k))" definition primepow_odd :: "nat \ bool" where "primepow_odd q \ (\ p k. 1 \ k \ prime p \ q = p^(2*k+1))" abbreviation (input) isprimedivisor :: "nat \ nat \ bool" where "isprimedivisor q p \ prime p \ p dvd q" definition pre_mangoldt :: "nat \ nat" where "pre_mangoldt d = (if primepow d then aprimedivisor d else 1)" definition mangoldt_even :: "nat \ real" where "mangoldt_even d = (if primepow_even d then ln (real (aprimedivisor d)) else 0)" definition mangoldt_odd :: "nat \ real" where "mangoldt_odd d = (if primepow_odd d then ln (real (aprimedivisor d)) else 0)" definition mangoldt_1 :: "nat \ real" where "mangoldt_1 d = (if prime d then ln d else 0)" definition psi :: "nat \ real" where "psi n = (\d=1..n. mangoldt d)" definition psi_even :: "nat \ real" where "psi_even n = (\d=1..n. mangoldt_even d)" definition psi_odd :: "nat \ real" where "psi_odd n = (\d=1..n. mangoldt_odd d)" abbreviation (input) psi_even_2 :: "nat \ real" where "psi_even_2 n \ (\d=2..n. mangoldt_even d)" abbreviation (input) psi_odd_2 :: "nat \ real" where "psi_odd_2 n \ (\d=2..n. mangoldt_odd d)" definition theta :: "nat \ real" where "theta n = (\p=1..n. if prime p then ln (real p) else 0)" subsection \Properties of prime powers\ lemma primepow_even_imp_primepow: assumes "primepow_even n" shows "primepow n" proof - from assms obtain p k where "1 \ k" "prime p" "n = p ^ (2 * k)" unfolding primepow_even_def by blast moreover from \1 \ k\ have "2 * k > 0" by simp ultimately show ?thesis unfolding primepow_def by blast qed lemma primepow_odd_imp_primepow: assumes "primepow_odd n" shows "primepow n" proof - from assms obtain p k where "1 \ k" "prime p" "n = p ^ (2 * k + 1)" unfolding primepow_odd_def by blast moreover from \1 \ k\ have "Suc (2 * k) > 0" by simp ultimately show ?thesis unfolding primepow_def by (auto simp del: power_Suc) qed lemma primepow_odd_altdef: "primepow_odd n \ primepow n \ odd (multiplicity (aprimedivisor n) n) \ multiplicity (aprimedivisor n) n > 1" proof (intro iffI conjI; (elim conjE)?) assume "primepow_odd n" then obtain p k where n: "k \ 1" "prime p" "n = p ^ (2 * k + 1)" by (auto simp: primepow_odd_def) thus "odd (multiplicity (aprimedivisor n) n)" "multiplicity (aprimedivisor n) n > 1" by (simp_all add: aprimedivisor_primepow prime_elem_multiplicity_mult_distrib) next assume A: "primepow n" and B: "odd (multiplicity (aprimedivisor n) n)" and C: "multiplicity (aprimedivisor n) n > 1" from A obtain p k where n: "k \ 1" "prime p" "n = p ^ k" by (auto simp: primepow_def Suc_le_eq) with B C have "odd k" "k > 1" by (simp_all add: aprimedivisor_primepow prime_elem_multiplicity_mult_distrib) then obtain j where j: "k = 2 * j + 1" "j > 0" by (auto elim!: oddE) with n show "primepow_odd n" by (auto simp: primepow_odd_def intro!: exI[of _ p, OF exI[of _ j]]) qed (auto dest: primepow_odd_imp_primepow) lemma primepow_even_altdef: "primepow_even n \ primepow n \ even (multiplicity (aprimedivisor n) n)" proof (intro iffI conjI; (elim conjE)?) assume "primepow_even n" then obtain p k where n: "k \ 1" "prime p" "n = p ^ (2 * k)" by (auto simp: primepow_even_def) thus "even (multiplicity (aprimedivisor n) n)" by (simp_all add: aprimedivisor_primepow prime_elem_multiplicity_mult_distrib) next assume A: "primepow n" and B: "even (multiplicity (aprimedivisor n) n)" from A obtain p k where n: "k \ 1" "prime p" "n = p ^ k" by (auto simp: primepow_def Suc_le_eq) with B have "even k" by (simp_all add: aprimedivisor_primepow prime_elem_multiplicity_mult_distrib) then obtain j where j: "k = 2 * j" by (auto elim!: evenE) from j n have "j \ 0" by (intro notI) simp_all with j n show "primepow_even n" by (auto simp: primepow_even_def intro!: exI[of _ p, OF exI[of _ j]]) qed (auto dest: primepow_even_imp_primepow) lemma primepow_odd_mult: assumes "d > Suc 0" shows "primepow_odd (aprimedivisor d * d) \ primepow_even d" using assms by (auto simp: primepow_odd_altdef primepow_even_altdef primepow_mult_aprimedivisorI aprimedivisor_primepow prime_aprimedivisor' aprimedivisor_dvd' prime_elem_multiplicity_mult_distrib prime_elem_aprimedivisor_nat dest!: primepow_multD) lemma pre_mangoldt_primepow: assumes "primepow n" "aprimedivisor n = p" shows "pre_mangoldt n = p" using assms by (simp add: pre_mangoldt_def) lemma pre_mangoldt_notprimepow: assumes "\primepow n" shows "pre_mangoldt n = 1" using assms by (simp add: pre_mangoldt_def) lemma primepow_cases: "primepow d \ ( primepow_even d \ \ primepow_odd d \ \ prime d) \ (\ primepow_even d \ primepow_odd d \ \ prime d) \ (\ primepow_even d \ \ primepow_odd d \ prime d)" by (auto simp: primepow_even_altdef primepow_odd_altdef multiplicity_aprimedivisor_Suc_0_iff elim!: oddE intro!: Nat.gr0I) subsection \Deriving a recurrence for the psi function\ lemma ln_fact_bounds: assumes "n > 0" shows "abs(ln (fact n) - n * ln n + n) \ 1 + ln n" proof - have "\n\{0<..}. \z>real n. z < real (n + 1) \ real (n + 1) * ln (real (n + 1)) - real n * ln (real n) = (real (n + 1) - real n) * (ln z + 1)" by (intro ballI MVT2) (auto intro!: derivative_eq_intros) hence "\n\{0<..}. \z>real n. z < real (n + 1) \ real (n + 1) * ln (real (n + 1)) - real n * ln (real n) = (ln z + 1)" by (simp add: algebra_simps) from bchoice[OF this] obtain k :: "nat \ real" where lb: "real n < k n" and ub: "k n < real (n + 1)" and mvt: "real (n+1) * ln (real (n+1)) - real n * ln (real n) = ln (k n) + 1" if "n > 0" for n::nat by blast have *: "(n + 1) * ln (n + 1) = (\i=1..n. ln(k i) + 1)" for n::nat proof (induction n) case (Suc n) have "(\i = 1..n+1. ln (k i) + 1) = (\i = 1..n. ln (k i) + 1) + ln (k (n+1)) + 1" by simp also from Suc.IH have "(\i = 1..n. ln (k i) + 1) = real (n+1) * ln (real (n+1))" .. also from mvt[of "n+1"] have "\ = real (n+2) * ln (real (n+2)) - ln (k (n+1)) - 1" by simp finally show ?case by simp qed simp have **: "abs((\i=1..n+1. ln i) - ((n+1) * ln (n+1) - (n+1))) \ 1 + ln(n+1)" for n::nat proof - have "(\i=1..n+1. ln i) \ (\i=1..n. ln i) + ln (n+1)" by simp also have "(\i=1..n. ln i) \ (\i=1..n. ln (k i))" by (intro sum_mono, subst ln_le_cancel_iff) (auto simp: Suc_le_eq dest: lb ub) also have "\ = (\i=1..n. ln (k i) + 1) - n" by (simp add: sum.distrib) also from * have "\ = (n+1) * ln (n+1) - n" by simp finally have a_minus_b: "(\i=1..n+1. ln i) - ((n+1) * ln (n+1) - (n+1)) \ 1 + ln (n+1)" by simp from * have "(n+1) * ln (n+1) - n = (\i=1..n. ln (k i) + 1) - n" by simp also have "\ = (\i=1..n. ln (k i))" by (simp add: sum.distrib) also have "\ \ (\i=1..n. ln (i+1))" by (intro sum_mono, subst ln_le_cancel_iff) (auto simp: Suc_le_eq dest: lb ub) also from sum.shift_bounds_cl_nat_ivl[of "ln" 1 1 n] have "\ = (\i=1+1..n+1. ln i)" .. also have "\ = (\i=1..n+1. ln i)" by (rule sum.mono_neutral_left) auto finally have b_minus_a: "((n+1) * ln (n+1) - (n+1)) - (\i=1..n+1. ln i) \ 1" by simp have "0 \ ln (n+1)" by simp with b_minus_a have "((n+1) * ln (n+1) - (n+1)) - (\i=1..n+1. ln i) \ 1 + ln (n+1)" by linarith with a_minus_b show ?thesis by linarith qed from \n > 0\ have "n \ 1" by simp thus ?thesis proof (induction n rule: dec_induct) case base then show ?case by simp next case (step n) from ln_fact[of "n+1"] **[of n] show ?case by simp qed qed lemma ln_fact_diff_bounds: "abs(ln (fact n) - 2 * ln (fact (n div 2)) - n * ln 2) \ 4 * ln (if n = 0 then 1 else n) + 3" proof (cases "n div 2 = 0") case True hence "n \ 1" by simp with ln_le_minus_one[of "2::real"] show ?thesis by (cases n) simp_all next case False then have "n > 1" by simp let ?a = "real n * ln 2" let ?b = "4 * ln (real n) + 3" let ?l1 = "ln (fact (n div 2))" let ?a1 = "real (n div 2) * ln (real (n div 2)) - real (n div 2)" let ?b1 = "1 + ln (real (n div 2))" let ?l2 = "ln (fact n)" let ?a2 = "real n * ln (real n) - real n" let ?b2 = "1 + ln (real n)" have abs_a: "abs(?a - (?a2 - 2 * ?a1)) \ ?b - 2 * ?b1 - ?b2" proof (cases "even n") case True then have "real (2 * (n div 2)) = real n" by simp then have n_div_2: "real (n div 2) = real n / 2" by simp from \n > 1\ have *: "abs(?a - (?a2 - 2 * ?a1)) = 0" by (simp add: n_div_2 ln_div algebra_simps) from \even n\ and \n > 1\ have "0 \ ln (real n) - ln (real (n div 2))" by (auto elim: evenE) also have "2 * \ \ 3 * ln (real n) - 2 * ln (real (n div 2))" using \n > 1\ by (auto intro!: ln_ge_zero) also have "\ = ?b - 2 * ?b1 - ?b2" by simp finally show ?thesis using * by simp next case False then have "real (2 * (n div 2)) = real (n - 1)" by simp with \n > 1\ have n_div_2: "real (n div 2) = (real n - 1) / 2" by simp from \odd n\ \n div 2 \ 0\ have "n \ 3" by presburger have "?a - (?a2 - 2 * ?a1) = real n * ln 2 - real n * ln (real n) + real n + 2 * real (n div 2) * ln (real (n div 2)) - 2* real (n div 2)" by (simp add: algebra_simps) also from n_div_2 have "2 * real (n div 2) = real n - 1" by simp also have "real n * ln 2 - real n * ln (real n) + real n + (real n - 1) * ln (real (n div 2)) - (real n - 1) = real n * (ln (real n - 1) - ln (real n)) - ln (real (n div 2)) + 1" using \n > 1\ by (simp add: algebra_simps n_div_2 ln_div) finally have lhs: "abs(?a - (?a2 - 2 * ?a1)) = abs(real n * (ln (real n - 1) - ln (real n)) - ln (real (n div 2)) + 1)" by simp from \n > 1\ have "real n * (ln (real n - 1) - ln (real n)) \ 0" by (simp add: algebra_simps mult_left_mono) moreover from \n > 1\ have "ln (real (n div 2)) \ ln (real n)" by simp moreover { have "exp 1 \ (3::real)" by (rule exp_le) also from \n \ 3\ have "\ \ exp (ln (real n))" by simp finally have "ln (real n) \ 1" by simp } ultimately have ub: "real n * (ln (real n - 1) - ln (real n)) - ln(real (n div 2)) + 1 \ 3 * ln (real n) - 2 * ln(real (n div 2))" by simp have mon: "real n' * (ln (real n') - ln (real n' - 1)) \ real n * (ln (real n) - ln (real n - 1))" if "n \ 3" "n' \ n" for n n'::nat proof (rule DERIV_nonpos_imp_nonincreasing[where f = "\x. x * (ln x - ln (x - 1))"]) fix t assume t: "real n \ t" "t \ real n'" with that have "1 / (t - 1) \ ln (1 + 1/(t - 1))" by (intro ln_add_one_self_le_self) simp_all also from t that have "ln (1 + 1/(t - 1)) = ln t- ln (t - 1)" by (simp add: ln_div [symmetric] field_simps) finally have "ln t - ln (t - 1) \ 1 / (t - 1)" . with that t show "\y. ((\x. x * (ln x - ln (x - 1))) has_field_derivative y) (at t) \ y \ 0" by (intro exI[of _ "1 / (1 - t) + ln t - ln (t - 1)"]) (force intro!: derivative_eq_intros simp: field_simps)+ qed (use that in simp_all) from \n > 1\ have "ln 2 = ln (real n) - ln (real n / 2)" by (simp add: ln_div) also from \n > 1\ have "\ \ ln (real n) - ln (real (n div 2))" by simp finally have *: "3*ln 2 + ln(real (n div 2)) \ 3* ln(real n) - 2* ln(real (n div 2))" by simp have "- real n * (ln (real n - 1) - ln (real n)) + ln(real (n div 2)) - 1 = real n * (ln (real n) - ln (real n - 1)) - 1 + ln(real (n div 2))" by (simp add: algebra_simps) also have "real n * (ln (real n) - ln (real n - 1)) \ 3 * (ln 3 - ln (3 - 1))" using mon[OF _ \n \ 3\] by simp also { have "Some (Float 3 (-1)) = ub_ln 1 3" by code_simp from ub_ln(1)[OF this] have "ln 3 \ (1.6 :: real)" by simp also have "1.6 - 1 / 3 \ 2 * (2/3 :: real)" by simp also have "2/3 \ ln (2 :: real)" by (rule ln_2_ge') finally have "ln 3 - 1 / 3 \ 2 * ln (2 :: real)" by simp } hence "3 * (ln 3 - ln (3 - 1)) - 1 \ 3 * ln (2 :: real)" by simp also note * finally have "- real n * (ln (real n - 1) - ln (real n)) + ln(real (n div 2)) - 1 \ 3 * ln (real n) - 2 * ln (real (n div 2))" by simp hence lhs': "abs(real n * (ln (real n - 1) - ln (real n)) - ln(real (n div 2)) + 1) \ 3 * ln (real n) - 2 * ln (real (n div 2))" using ub by simp have rhs: "?b - 2 * ?b1 - ?b2 = 3* ln (real n) - 2 * ln (real (n div 2))" by simp from \n > 1\ have "ln (real (n div 2)) \ 3* ln (real n) - 2* ln (real (n div 2))" by simp with rhs lhs lhs' show ?thesis by simp qed then have minus_a: "-?a \ ?b - 2 * ?b1 - ?b2 - (?a2 - 2 * ?a1)" by simp from abs_a have a: "?a \ ?b - 2 * ?b1 - ?b2 + ?a2 - 2 * ?a1" by (simp) from ln_fact_bounds[of "n div 2"] False have abs_l1: "abs(?l1 - ?a1) \ ?b1" by (simp add: algebra_simps) then have minus_l1: "?a1 - ?l1 \ ?b1" by linarith from abs_l1 have l1: "?l1 - ?a1 \ ?b1" by linarith from ln_fact_bounds[of n] False have abs_l2: "abs(?l2 - ?a2) \ ?b2" by (simp add: algebra_simps) then have l2: "?l2 - ?a2 \ ?b2" by simp from abs_l2 have minus_l2: "?a2 - ?l2 \ ?b2" by simp from minus_a minus_l1 l2 have "?l2 - 2 * ?l1 - ?a \ ?b" by simp moreover from a l1 minus_l2 have "- ?l2 + 2 * ?l1 + ?a \ ?b" by simp ultimately have "abs((?l2 - 2*?l1) - ?a) \ ?b" by simp then show ?thesis by simp qed lemma ln_primefact: assumes "n \ (0::nat)" shows "ln n = (\d=1..n. if primepow d \ d dvd n then ln (aprimedivisor d) else 0)" (is "?lhs = ?rhs") proof - have "?rhs = (\d\{x \ {1..n}. primepow x \ x dvd n}. ln (real (aprimedivisor d)))" unfolding primepow_factors_def by (subst sum.inter_filter [symmetric]) simp_all also have "{x \ {1..n}. primepow x \ x dvd n} = primepow_factors n" using assms by (auto simp: primepow_factors_def dest: dvd_imp_le primepow_gt_Suc_0) finally have *: "(\d\primepow_factors n. ln (real (aprimedivisor d))) = ?rhs" .. from in_prime_factors_imp_prime prime_gt_0_nat have pf_pos: "\p. p\#prime_factorization n \ p > 0" by blast from ln_msetprod[of "prime_factorization n", OF pf_pos] assms have "ln n = (\p\#prime_factorization n. ln p)" by (simp add: of_nat_prod_mset) also from * sum_prime_factorization_conv_sum_primepow_factors[of n ln, OF assms(1)] have "\ = ?rhs" by simp finally show ?thesis . qed context begin private lemma divisors: fixes x d::nat assumes "x \ {1..n}" assumes "d dvd x" shows "\k\{1..n div d}. x = d * k" proof - from assms have "x \ n" by simp then have ub: "x div d \ n div d" by (simp add: div_le_mono \x \ n\) from assms have "1 \ x div d" by (auto elim!: dvdE) with ub have "x div d \ {1..n div d}" by simp with \d dvd x\ show ?thesis by (auto intro!: bexI[of _ "x div d"]) qed lemma ln_fact_conv_mangoldt: "ln (fact n) = (\d=1..n. mangoldt d * floor (n / d))" proof - have *: "(\da=1..n. if primepow da \ da dvd d then ln (aprimedivisor da) else 0) = (\(da::nat)=1..d. if primepow da \ da dvd d then ln (aprimedivisor da) else 0)" if d: "d \ {1..n}" for d by (rule sum.mono_neutral_right, insert d) (auto dest: dvd_imp_le) have "(\d=1..n. \da=1..d. if primepow da \ da dvd d then ln (aprimedivisor da) else 0) = (\d=1..n. \da=1..n. if primepow da \ da dvd d then ln (aprimedivisor da) else 0)" by (rule sum.cong) (insert *, simp_all) also have "\ = (\da=1..n. \d=1..n. if primepow da \ da dvd d then ln (aprimedivisor da) else 0)" by (rule sum.swap) also have "\ = sum (\d. mangoldt d * floor (n/d)) {1..n}" proof (rule sum.cong) fix d assume d: "d \ {1..n}" have "(\da = 1..n. if primepow d \ d dvd da then ln (real (aprimedivisor d)) else 0) = (\da = 1..n. if d dvd da then mangoldt d else 0)" by (intro sum.cong) (simp_all add: mangoldt_def) also have "\ = mangoldt d * real (card {x. x \ {1..n} \ d dvd x})" by (subst sum.inter_filter [symmetric]) (simp_all add: algebra_simps) also { have "{x. x \ {1..n} \ d dvd x} = {x. \k \{1..n div d}. x=k*d}" proof safe fix x assume "x \ {1..n}" "d dvd x" thus "\k\{1..n div d}. x = k * d" using divisors[of x n d] by auto next fix x k assume k: "k \ {1..n div d}" from k have "k * d \ n div d * d" by (intro mult_right_mono) simp_all also have "n div d * d \ n div d * d + n mod d" by (rule le_add1) also have "\ = n" by simp finally have "k * d \ n" . thus "k * d \ {1..n}" using d k by auto qed auto also have "\ = (\k. k*d) ` {1..n div d}" by fast also have "card \ = card {1..n div d}" by (rule card_image) (simp add: inj_on_def) also have "\ = n div d" by simp also have "... = \n / d\" by (simp add: floor_divide_of_nat_eq) finally have "real (card {x. x \ {1..n} \ d dvd x}) = real_of_int \n / d\" by force } finally show "(\da = 1..n. if primepow d \ d dvd da then ln (real (aprimedivisor d)) else 0) = mangoldt d * real_of_int \real n / real d\" . qed simp_all finally have "(\d=1..n. \da=1..d. if primepow da \ da dvd d then ln (aprimedivisor da) else 0) = sum (\d. mangoldt d * floor (n/d)) {1..n}" . with ln_primefact have "(\d=1..n. ln d) = (\d=1..n. mangoldt d * floor (n/d))" by simp with ln_fact show ?thesis by simp qed end context begin private lemma div_2_mult_2_bds: fixes n d :: nat assumes "d > 0" shows "0 \ \n / d\ - 2 * \(n div 2) / d\" "\n / d\ - 2 * \(n div 2) / d\ \ 1" proof - have "\2::real\ * \(n div 2) / d\ \ \2 * ((n div 2) / d)\" by (rule le_mult_floor) simp_all also from assms have "\ \ \n / d\" by (intro floor_mono) (simp_all add: field_simps) finally show "0 \ \n / d\ - 2 * \(n div 2) / d\" by (simp add: algebra_simps) next have "real (n div d) \ real (2 * ((n div 2) div d) + 1)" by (subst div_mult2_eq [symmetric], simp only: mult.commute, subst div_mult2_eq) simp thus "\n / d\ - 2 * \(n div 2) / d\ \ 1" unfolding of_nat_add of_nat_mult floor_conv_div_nat [symmetric] by simp_all qed private lemma n_div_d_eq_1: "d \ {n div 2 + 1..n} \ \real n / real d\ = 1" by (cases "n = d") (auto simp: field_simps intro: floor_eq) lemma psi_bounds_ln_fact: shows "ln (fact n) - 2 * ln (fact (n div 2)) \ psi n" "psi n - psi (n div 2) \ ln (fact n) - 2 * ln (fact (n div 2))" proof - fix n::nat let ?k = "n div 2" and ?d = "n mod 2" have *: "\?k / d\ = 0" if "d > ?k" for d proof - from that div_less have "0 = ?k div d" by simp also have "\ = \?k / d\" by (rule floor_divide_of_nat_eq [symmetric]) finally show "\?k / d\ = 0" by simp qed have sum_eq: "(\d=1..2*?k+?d. mangoldt d * \?k / d\) = (\d=1..?k. mangoldt d * \?k / d\)" by (intro sum.mono_neutral_right) (auto simp: *) from ln_fact_conv_mangoldt have "ln (fact n) = (\d=1..n. mangoldt d * \n / d\)" . also have "\ = (\d=1..n. mangoldt d * \(2 * (n div 2) + n mod 2) / d\)" by simp also have "\ \ (\d=1..n. mangoldt d * (2 * \?k / d\ + 1))" using div_2_mult_2_bds(2)[of _ n] by (intro sum_mono mult_left_mono, subst of_int_le_iff) (auto simp: algebra_simps mangoldt_nonneg) also have "\ = 2 * (\d=1..n. mangoldt d * \(n div 2) / d\) + (\d=1..n. mangoldt d)" by (simp add: algebra_simps sum.distrib sum_distrib_left) also have "\ = 2 * (\d=1..2*?k+?d. mangoldt d * \(n div 2) / d\) + (\d=1..n. mangoldt d)" by presburger also from sum_eq have "\ = 2 * (\d=1..?k. mangoldt d * \(n div 2) / d\) + (\d=1..n. mangoldt d)" by presburger also from ln_fact_conv_mangoldt psi_def have "\ = 2 * ln (fact ?k) + psi n" by presburger finally show "ln (fact n) - 2 * ln (fact (n div 2)) \ psi n" by simp next fix n::nat let ?k = "n div 2" and ?d = "n mod 2" from psi_def have "psi n - psi ?k = (\d=1..2*?k+?d. mangoldt d) - (\d=1..?k. mangoldt d)" by presburger also have "\ = sum mangoldt ({1..2 * (n div 2) + n mod 2} - {1..n div 2})" by (subst sum_diff) simp_all also have "\ = (\d\({1..2 * (n div 2) + n mod 2} - {1..n div 2}). (if d \ ?k then 0 else mangoldt d))" by (intro sum.cong) simp_all also have "\ = (\d=1..2*?k+?d. (if d \ ?k then 0 else mangoldt d))" by (intro sum.mono_neutral_left) auto also have "\ = (\d=1..n. (if d \ ?k then 0 else mangoldt d))" by presburger also have "\ = (\d=1..n. (if d \ ?k then mangoldt d * 0 else mangoldt d))" by (intro sum.cong) simp_all also from div_2_mult_2_bds(1) have "\ \ (\d=1..n. (if d \ ?k then mangoldt d * (\n/d\ - 2 * \?k/d\) else mangoldt d))" by (intro sum_mono) (auto simp: algebra_simps mangoldt_nonneg intro!: mult_left_mono simp del: of_int_mult) also from n_div_d_eq_1 have "\ = (\d=1..n. (if d \ ?k then mangoldt d * (\n/d\ - 2 * \?k/d\) else mangoldt d * \n/d\))" by (intro sum.cong refl) auto also have "\ = (\d=1..n. mangoldt d * real_of_int (\real n / real d\) - (if d \ ?k then 2 * mangoldt d * real_of_int \real ?k / real d\ else 0))" by (intro sum.cong refl) (auto simp: algebra_simps) also have "\ = (\d=1..n. mangoldt d * real_of_int (\real n / real d\)) - (\d=1..n. (if d \ ?k then 2 * mangoldt d * real_of_int \real ?k / real d\ else 0))" by (rule sum_subtractf) also have "(\d=1..n. (if d \ ?k then 2 * mangoldt d * real_of_int \real ?k / real d\ else 0)) = (\d=1..?k. (if d \ ?k then 2 * mangoldt d * real_of_int \real ?k / real d\ else 0))" by (intro sum.mono_neutral_right) auto also have "\ = (\d=1..?k. 2 * mangoldt d * real_of_int \real ?k / real d\)" by (intro sum.cong) simp_all also have "\ = 2 * (\d=1..?k. mangoldt d * real_of_int \real ?k / real d\)" by (simp add: sum_distrib_left mult_ac) also have "(\d = 1..n. mangoldt d * real_of_int \real n / real d\) - \ = ln (fact n) - 2 * ln (fact (n div 2))" by (simp add: ln_fact_conv_mangoldt) finally show "psi n - psi (n div 2) \ ln (fact n) - 2 * ln (fact (n div 2))" . qed end lemma psi_bounds_induct: "real n * ln 2 - (4 * ln (real (if n = 0 then 1 else n)) + 3) \ psi n" "psi n - psi (n div 2) \ real n * ln 2 + (4 * ln (real (if n = 0 then 1 else n)) + 3)" proof - from le_imp_neg_le[OF ln_fact_diff_bounds] have "n * ln 2 - (4 * ln (if n = 0 then 1 else n) + 3) \ n * ln 2 - abs(ln (fact n) - 2 * ln (fact (n div 2)) - n * ln 2)" by simp also have "\ \ ln (fact n) - 2 * ln (fact (n div 2))" by simp also from psi_bounds_ln_fact (1) have "\ \ psi n" by simp finally show "real n * ln 2 - (4 * ln (real (if n = 0 then 1 else n)) + 3) \ psi n" . next from psi_bounds_ln_fact (2) have "psi n - psi (n div 2) \ ln (fact n) - 2 * ln (fact (n div 2))" . also have "\ \ n * ln 2 + abs(ln (fact n) - 2 * ln (fact (n div 2)) - n * ln 2)" by simp also from ln_fact_diff_bounds [of n] have "abs(ln (fact n) - 2 * ln (fact (n div 2)) - n * ln 2) \ (4 * ln (real (if n = 0 then 1 else n)) + 3)" by simp finally show "psi n - psi (n div 2) \ real n * ln 2 + (4 * ln (real (if n = 0 then 1 else n)) + 3)" by simp qed subsection \Bounding the psi function\ text \ In this section, we will first prove the relatively tight estimate @{prop "psi n \ 3 / 2 + ln 2 * n"} for @{term "n \ 128"} and then use the recurrence we have just derived to extend it to @{prop "psi n \ 551 / 256"} for @{term "n \ 1024"}, at which point applying the recurrence can be used to prove the same bound for arbitrarily big numbers. First of all, we will prove the bound for @{term "n <= 128"} using reflection and approximation. \ context begin private lemma Ball_insertD: assumes "\x\insert y A. P x" shows "P y" "\x\A. P x" using assms by auto private lemma meta_eq_TrueE: "PROP A \ Trueprop True \ PROP A" by simp private lemma pre_mangoldt_pos: "pre_mangoldt n > 0" unfolding pre_mangoldt_def by (auto simp: primepow_gt_Suc_0) private lemma psi_conv_pre_mangoldt: "psi n = ln (real (prod pre_mangoldt {1..n}))" by (auto simp: psi_def mangoldt_def pre_mangoldt_def ln_prod primepow_gt_Suc_0 intro!: sum.cong) private lemma eval_psi_aux1: "psi 0 = ln (real (numeral Num.One))" by (simp add: psi_def) private lemma eval_psi_aux2: assumes "psi m = ln (real (numeral x))" "pre_mangoldt n = y" "m + 1 = n" "numeral x * y = z" shows "psi n = ln (real z)" proof - from assms(2) [symmetric] have [simp]: "y > 0" by (simp add: pre_mangoldt_pos) have "psi n = psi (Suc m)" by (simp add: assms(3) [symmetric]) also have "\ = ln (real y * (\x = Suc 0..m. real (pre_mangoldt x)))" using assms(2,3) [symmetric] by (simp add: psi_conv_pre_mangoldt prod.nat_ivl_Suc' mult_ac) also have "\ = ln (real y) + psi m" by (subst ln_mult) (simp_all add: pre_mangoldt_pos prod_pos psi_conv_pre_mangoldt) also have "psi m = ln (real (numeral x))" by fact also have "ln (real y) + \ = ln (real (numeral x * y))" by (simp add: ln_mult) finally show ?thesis by (simp add: assms(4) [symmetric]) qed private lemma Ball_atLeast0AtMost_doubleton: assumes "psi 0 \ 3 / 2 * ln 2 * real 0" assumes "psi 1 \ 3 / 2 * ln 2 * real 1" shows "(\x\{0..1}. psi x \ 3 / 2 * ln 2 * real x)" using assms unfolding One_nat_def atLeast0_atMost_Suc ball_simps by auto private lemma Ball_atLeast0AtMost_insert: assumes "(\x\{0..m}. psi x \ 3 / 2 * ln 2 * real x)" assumes "psi (numeral n) \ 3 / 2 * ln 2 * real (numeral n)" "m = pred_numeral n" shows "(\x\{0..numeral n}. psi x \ 3 / 2 * ln 2 * real x)" using assms by (subst numeral_eq_Suc[of n], subst atLeast0_atMost_Suc, subst ball_simps, simp only: numeral_eq_Suc [symmetric]) private lemma eval_psi_ineq_aux: assumes "psi n = x" "x \ 3 / 2 * ln 2 * n" shows "psi n \ 3 / 2 * ln 2 * n" using assms by simp_all private lemma eval_psi_ineq_aux2: assumes "numeral m ^ 2 \ (2::nat) ^ (3 * n)" shows "ln (real (numeral m)) \ 3 / 2 * ln 2 * real n" proof - have "ln (real (numeral m)) \ 3 / 2 * ln 2 * real n \ 2 * log 2 (real (numeral m)) \ 3 * real n" by (simp add: field_simps log_def) also have "2 * log 2 (real (numeral m)) = log 2 (real (numeral m ^ 2))" by (subst of_nat_power, subst log_nat_power) simp_all also have "\ \ 3 * real n \ real ((numeral m) ^ 2) \ 2 powr real (3 * n)" by (subst Transcendental.log_le_iff) simp_all also have "2 powr (3 * n) = real (2 ^ (3 * n))" by (simp add: powr_realpow [symmetric]) also have "real ((numeral m) ^ 2) \ \ \ numeral m ^ 2 \ (2::nat) ^ (3 * n)" by (rule of_nat_le_iff) finally show ?thesis using assms by blast qed private lemma eval_psi_ineq_aux_mono: assumes "psi n = x" "psi m = x" "psi n \ 3 / 2 * ln 2 * n" "n \ m" shows "psi m \ 3 / 2 * ln 2 * m" proof - from assms have "psi m = psi n" by simp also have "\ \ 3 / 2 * ln 2 * n" by fact also from \n \ m\ have "\ \ 3 / 2 * ln 2 * m" by simp finally show ?thesis . qed lemma not_primepow_1_nat: "\primepow (1 :: nat)" by auto ML_file \bertrand.ML\ (* This should not take more than 1 minute *) -local_setup \fn ctxt => +local_setup \fn lthy => let - fun tac {context = ctxt, ...} = + fun tac ctxt = let val psi_cache = Bertrand.prove_psi ctxt 129 fun prove_psi_ineqs ctxt = let - fun tac {context = ctxt, ...} = - HEADGOAL (resolve_tac ctxt @{thms eval_psi_ineq_aux2} THEN' Simplifier.simp_tac ctxt) + fun tac goal_ctxt = + HEADGOAL (resolve_tac goal_ctxt @{thms eval_psi_ineq_aux2} THEN' + Simplifier.simp_tac goal_ctxt) fun prove_by_approx n thm = let val thm = thm RS @{thm eval_psi_ineq_aux} val [prem] = Thm.prems_of thm - val prem = Goal.prove ctxt [] [] prem tac + val prem = Goal.prove ctxt [] [] prem (tac o #context) in prem RS thm end fun prove_by_mono last_thm last_thm' thm = let val thm = @{thm eval_psi_ineq_aux_mono} OF [last_thm, thm, last_thm'] val [prem] = Thm.prems_of thm - val prem = Goal.prove ctxt [] [] prem (K (HEADGOAL (Simplifier.simp_tac ctxt))) + val prem = + Goal.prove ctxt [] [] prem (fn {context = goal_ctxt, ...} => + HEADGOAL (Simplifier.simp_tac goal_ctxt)) in prem RS thm end fun go _ acc [] = acc | go last acc ((n, x, thm) :: xs) = let val thm' = case last of NONE => prove_by_approx n thm | SOME (last_x, last_thm, last_thm') => if last_x = x then prove_by_mono last_thm last_thm' thm else prove_by_approx n thm in go (SOME (x, thm, thm')) (thm' :: acc) xs end in rev o go NONE [] end val psi_ineqs = prove_psi_ineqs ctxt psi_cache fun prove_ball ctxt (thm1 :: thm2 :: thms) = let val thm = @{thm Ball_atLeast0AtMost_doubleton} OF [thm1, thm2] fun solve_prem thm = let - fun tac {context = ctxt, ...} = HEADGOAL (Simplifier.simp_tac ctxt) - val thm' = Goal.prove ctxt [] [] (Thm.cprem_of thm 1 |> Thm.term_of) tac + val thm' = + Goal.prove ctxt [] [] (Thm.cprem_of thm 1 |> Thm.term_of) + (fn {context = goal_ctxt, ...} => + HEADGOAL (Simplifier.simp_tac goal_ctxt)) in thm' RS thm end fun go thm thm' = (@{thm Ball_atLeast0AtMost_insert} OF [thm', thm]) |> solve_prem in fold go thms thm end | prove_ball _ _ = raise Match in HEADGOAL (resolve_tac ctxt [prove_ball ctxt psi_ineqs]) end - val thm = Goal.prove @{context} [] [] @{prop "\n\{0..128}. psi n \ 3 / 2 * ln 2 * n"} tac + val thm = Goal.prove lthy [] [] @{prop "\n\{0..128}. psi n \ 3 / 2 * ln 2 * n"} (tac o #context) in - Local_Theory.note ((@{binding psi_ubound_log_128}, []), [thm]) ctxt |> snd + Local_Theory.note ((@{binding psi_ubound_log_128}, []), [thm]) lthy |> snd end \ end context begin private lemma psi_ubound_aux: defines "f \ \x::real. (4 * ln x + 3) / (ln 2 * x)" assumes "x \ 2" "x \ y" shows "f x \ f y" using assms(3) proof (rule DERIV_nonpos_imp_nonincreasing, goal_cases) case (1 t) define f' where "f' = (\x. (1 - 4 * ln x) / x^2 / ln 2 :: real)" from 1 assms(2) have "(f has_real_derivative f' t) (at t)" unfolding f_def f'_def by (auto intro!: derivative_eq_intros simp: field_simps power2_eq_square) moreover { from ln_2_ge have "1/4 \ ln (2::real)" by simp also from assms(2) 1 have "\ \ ln t" by simp finally have "ln t \ 1/4" . } with 1 assms(2) have "f' t \ 0" by (simp add: f'_def field_simps) ultimately show ?case by (intro exI[of _ "f' t"]) simp_all qed text \ These next rules are used in combination with @{thm psi_bounds_induct} and @{thm psi_ubound_log_128} to extend the upper bound for @{term "psi"} from values no greater than 128 to values no greater than 1024. The constant factor of the upper bound changes every time, but once we have reached 1024, the recurrence is self-sustaining in the sense that we do not have to adjust the constant factor anymore in order to double the range. \ lemma psi_ubound_log_double_cases': assumes "\n. n \ m \ psi n \ c * ln 2 * real n" "n \ m'" "m' = 2*m" "c \ c'" "c \ 0" "m \ 1" "c' \ 1 + c/2 + (4 * ln (m+1) + 3) / (ln 2 * (m+1))" shows "psi n \ c' * ln 2 * real n" proof (cases "n > m") case False hence "psi n \ c * ln 2 * real n" by (intro assms) simp_all also have "c \ c'" by fact finally show ?thesis by - (simp_all add: mult_right_mono) next case True hence n: "n \ m+1" by simp from psi_bounds_induct(2)[of n] True have "psi n \ real n * ln 2 + 4 * ln (real n) + 3 + psi (n div 2)" by simp also from assms have "psi (n div 2) \ c * ln 2 * real (n div 2)" by (intro assms) simp_all also have "real (n div 2) \ real n / 2" by simp also have "c * ln 2 * \ = c / 2 * ln 2 * real n" by simp also have "real n * ln 2 + 4 * ln (real n) + 3 + \ = (1 + c/2) * ln 2 * real n + (4 * ln (real n) + 3)" by (simp add: field_simps) also { have "(4 * ln (real n) + 3) / (ln 2 * (real n)) \ (4 * ln (m+1) + 3) / (ln 2 * (m+1))" using n assms by (intro psi_ubound_aux) simp_all also from assms have "(4 * ln (m+1) + 3) / (ln 2 * (m+1)) \ c' - 1 - c/2" by (simp add: algebra_simps) finally have "4 * ln (real n) + 3 \ (c' - 1 - c/2) * ln 2 * real n" using n by (simp add: field_simps) } also have "(1 + c / 2) * ln 2 * real n + (c' - 1 - c / 2) * ln 2 * real n = c' * ln 2 * real n" by (simp add: field_simps) finally show ?thesis using \c \ 0\ by (simp_all add: mult_left_mono) qed end lemma psi_ubound_log_double_cases: assumes "\n\m. psi n \ c * ln 2 * real n" "c' \ 1 + c/2 + (4 * ln (m+1) + 3) / (ln 2 * (m+1))" "m' = 2*m" "c \ c'" "c \ 0" "m \ 1" shows "\n\m'. psi n \ c' * ln 2 * real n" using assms(1) by (intro allI impI assms psi_ubound_log_double_cases'[of m c _ m' c']) auto lemma psi_ubound_log_1024: "\n\1024. psi n \ 551 / 256 * ln 2 * real n" proof - from psi_ubound_log_128 have "\n\128. psi n \ 3 / 2 * ln 2 * real n" by simp hence "\n\256. psi n \ 1025 / 512 * ln 2 * real n" proof (rule psi_ubound_log_double_cases, goal_cases) case 1 have "Some (Float 624 (- 7)) = ub_ln 9 129" by code_simp from ub_ln(1)[OF this] and ln_2_ge show ?case by (simp add: field_simps) qed simp_all hence "\n\512. psi n \ 549 / 256 * ln 2 * real n" proof (rule psi_ubound_log_double_cases, goal_cases) case 1 have "Some (Float 180 (- 5)) = ub_ln 7 257" by code_simp from ub_ln(1)[OF this] and ln_2_ge show ?case by (simp add: field_simps) qed simp_all thus "\n\1024. psi n \ 551 / 256 * ln 2 * real n" proof (rule psi_ubound_log_double_cases, goal_cases) case 1 have "Some (Float 203 (- 5)) = ub_ln 7 513" by code_simp from ub_ln(1)[OF this] and ln_2_ge show ?case by (simp add: field_simps) qed simp_all qed lemma psi_bounds_sustained_induct: assumes "4 * ln (1 + 2 ^ j) + 3 \ d * ln 2 * (1 + 2^j)" assumes "4 / (1 + 2^j) \ d * ln 2" assumes "0 \ c" assumes "c / 2 + d + 1 \ c" assumes "j \ k" assumes "\n. n \ 2^k \ psi n \ c * ln 2 * n" assumes "n \ 2^(Suc k)" shows "psi n \ c * ln 2 * n" proof (cases "n \ 2^k") case True with assms(6) show ?thesis . next case False from psi_bounds_induct(2) have "psi n - psi (n div 2) \ real n * ln 2 + (4 * ln (real (if n = 0 then 1 else n)) + 3)" . also from False have "(if n = 0 then 1 else n) = n" by simp finally have "psi n \ real n * ln 2 + (4 * ln (real n) + 3) + psi (n div 2)" by simp also from assms(6,7) have "psi (n div 2) \ c * ln 2 * (n div 2)" by simp also have "real (n div 2) \ real n / 2" by simp also have "real n * ln 2 + (4 * ln (real n) + 3) + c * ln 2 * (n / 2) \ c * ln 2 * real n" proof (rule overpower_lemma[of "\x. x * ln 2 + (4 * ln x + 3) + c * ln 2 * (x / 2)" "1+2^j" "\x. c * ln 2 * x" "\x. c * ln 2 - ln 2 - 4 / x - c / 2 * ln 2" "real n"]) from assms(1) have "4 * ln (1 + 2^j) + 3 \ d * ln 2 * (1 + 2^j)" . also from assms(4) have "d \ c - c/2 - 1" by simp also have "(\) * ln 2 * (1 + 2 ^ j) = c * ln 2 * (1 + 2 ^ j) - c / 2 * ln 2 * (1 + 2 ^ j) - (1 + 2 ^ j) * ln 2" by (simp add: left_diff_distrib) finally have "4 * ln (1 + 2^j) + 3 \ c * ln 2 * (1 + 2 ^ j) - c / 2 * ln 2 * (1 + 2 ^ j) - (1 + 2 ^ j) * ln 2" by (simp add: add_pos_pos) then show "(1 + 2 ^ j) * ln 2 + (4 * ln (1 + 2 ^ j) + 3) + c * ln 2 * ((1 + 2 ^ j) / 2) \ c * ln 2 * (1 + 2 ^ j)" by simp next fix x::real assume x: "1 + 2^j \ x" moreover have "1 + 2 ^ j > (0::real)" by (simp add: add_pos_pos) ultimately have x_pos: "x > 0" by linarith show "((\x. c * ln 2 * x - (x * ln 2 + (4 * ln x + 3) + c * ln 2 * (x / 2))) has_real_derivative c * ln 2 - ln 2 - 4 / x - c / 2 * ln 2) (at x)" by (rule derivative_eq_intros refl | simp add: \0 < x\)+ from \0 < x\ \0 < 1 + 2^j\ have "0 < x * (1 + 2^j)" by (rule mult_pos_pos) have "4 / x \ 4 / (1 + 2^j)" by (intro divide_left_mono mult_pos_pos add_pos_pos x x_pos) simp_all also from assms(2) have "4 / (1 + 2^j) \ d * ln 2" . also from assms(4) have "d \ c - c/2 - 1" by simp also have "\ * ln 2 = c * ln 2 - c/2 * ln 2 - ln 2" by (simp add: algebra_simps) finally show "0 \ c * ln 2 - ln 2 - 4 / x - c / 2 * ln 2" by simp next have "1 + 2^j = real (1 + 2^j)" by simp also from assms(5) have "\ \ real (1 + 2^k)" by simp also from False have "2^k \ n - 1" by simp finally show "1 + 2^j \ real n" using False by simp qed finally show ?thesis using assms by - (simp_all add: mult_left_mono) qed lemma psi_bounds_sustained: assumes "\n. n \ 2^k \ psi n \ c * ln 2 * n" assumes "4 * ln (1 + 2^k) + 3 \ (c/2 - 1) * ln 2 * (1 + 2^k)" assumes "4 / (1 + 2^k) \ (c/2 - 1) * ln 2" assumes "c \ 0" shows "psi n \ c * ln 2 * n" proof - have "psi n \ c * ln 2 * n" if "n \ 2^j" for j n using that proof (induction j arbitrary: n) case 0 with assms(4) 0 show ?case unfolding psi_def mangoldt_def by (cases n) auto next case (Suc j) show ?case proof (cases "k \ j") case True from assms(4) have c_div_2: "c/2 + (c/2 - 1) + 1 \ c" by simp from psi_bounds_sustained_induct[of k "c/2 -1" c j, OF assms(2) assms(3) assms(4) c_div_2 True Suc.IH Suc.prems] show ?thesis by simp next case False then have j_lt_k: "Suc j \ k" by simp from Suc.prems have "n \ 2 ^ Suc j" . also have "(2::nat) ^ Suc j \ 2 ^ k" using power_increasing[of "Suc j" k "2::nat", OF j_lt_k] by simp finally show ?thesis using assms(1) by simp qed qed from less_exp this [of n n] show ?thesis by simp qed lemma psi_ubound_log: "psi n \ 551 / 256 * ln 2 * n" proof (rule psi_bounds_sustained) show "0 \ 551 / (256 :: real)" by simp next fix n :: nat assume "n \ 2 ^ 10" with psi_ubound_log_1024 show "psi n \ 551 / 256 * ln 2 * real n" by auto next have "4 / (1 + 2 ^ 10) \ (551 / 256 / 2 - 1) * (2/3 :: real)" by simp also have "\ \ (551 / 256 / 2 - 1) * ln 2" by (intro mult_left_mono ln_2_ge') simp_all finally show "4 / (1 + 2 ^ 10) \ (551 / 256 / 2 - 1) * ln (2 :: real)" . next have "Some (Float 16 (-1)) = ub_ln 3 1025" by code_simp from ub_ln(1)[OF this] and ln_2_ge have "2048 * ln 1025 + 1536 \ 39975 * (ln 2::real)" by simp thus "4 * ln (1 + 2 ^ 10) + 3 \ (551 / 256 / 2 - 1) * ln 2 * (1 + 2 ^ 10 :: real)" by simp qed lemma psi_ubound_3_2: "psi n \ 3/2 * n" proof - have "(551 / 256) * ln 2 \ (551 / 256) * (16/23 :: real)" by (intro mult_left_mono ln_2_le') auto also have "\ \ 3 / 2" by simp finally have "551 / 256 * ln 2 \ 3/(2::real)" . with of_nat_0_le_iff mult_right_mono have "551 / 256 * ln 2 * n \ 3/2 * n" by blast with psi_ubound_log[of "n"] show ?thesis by linarith qed subsection \Doubling psi and theta\ lemma psi_residues_compare_2: "psi_odd_2 n \ psi_even_2 n" proof - have "psi_odd_2 n = (\d\{d. d \ {2..n} \ primepow_odd d}. mangoldt_odd d)" unfolding mangoldt_odd_def by (rule sum.mono_neutral_right) auto also have "\ = (\d\{d. d \ {2..n} \ primepow_odd d}. ln (real (aprimedivisor d)))" by (intro sum.cong refl) (simp add: mangoldt_odd_def) also have "\ \ (\d\{d. d \ {2..n} \ primepow_even d}. ln (real (aprimedivisor d)))" proof (rule sum_le_included [where i = "\y. y * aprimedivisor y"]; clarify?) fix d :: nat assume "d \ {2..n}" "primepow_odd d" note d = this then obtain p k where d': "k \ 1" "prime p" "d = p ^ (2*k+1)" by (auto simp: primepow_odd_def) from d' have "p ^ (2 * k) \ p ^ (2 * k + 1)" by (subst power_increasing_iff) (auto simp: prime_gt_Suc_0_nat) also from d d' have "\ \ n" by simp finally have "p ^ (2 * k) \ n" . moreover from d' have "p ^ (2 * k) > 1" by (intro one_less_power) (simp_all add: prime_gt_Suc_0_nat) ultimately have "p ^ (2 * k) \ {2..n}" by simp moreover from d' have "primepow_even (p ^ (2 * k))" by (auto simp: primepow_even_def) ultimately show "\y\{d \ {2..n}. primepow_even d}. y * aprimedivisor y = d \ ln (real (aprimedivisor d)) \ ln (real (aprimedivisor y))" using d' by (intro bexI[of _ "p ^ (2 * k)"]) (auto simp: aprimedivisor_prime_power aprimedivisor_primepow) qed (simp_all add: of_nat_ge_1_iff Suc_le_eq) also have "\ = (\d\{d. d \ {2..n} \ primepow_even d}. mangoldt_even d)" by (intro sum.cong refl) (simp add: mangoldt_even_def) also have "\ = psi_even_2 n" unfolding mangoldt_even_def by (rule sum.mono_neutral_left) auto finally show ?thesis . qed lemma psi_residues_compare: "psi_odd n \ psi_even n" proof - have "\ primepow_odd 1" by (simp add: primepow_odd_def) hence *: "mangoldt_odd 1 = 0" by (simp add: mangoldt_odd_def) have "\ primepow_even 1" using primepow_gt_Suc_0[OF primepow_even_imp_primepow, of 1] by auto with mangoldt_even_def have **: "mangoldt_even 1 = 0" by simp from psi_odd_def have "psi_odd n = (\d=1..n. mangoldt_odd d)" by simp also from * have "\ = psi_odd_2 n" by (cases "n \ 1") (simp_all add: eval_nat_numeral sum.atLeast_Suc_atMost) also from psi_residues_compare_2 have "\ \ psi_even_2 n" . also from ** have "\ = psi_even n" by (cases "n \ 1") (simp_all add: eval_nat_numeral sum.atLeast_Suc_atMost psi_even_def) finally show ?thesis . qed lemma primepow_iff_even_sqr: "primepow n \ primepow_even (n^2)" by (cases "n = 0") (auto simp: primepow_even_altdef aprimedivisor_primepow_power primepow_power_iff_nat prime_elem_multiplicity_power_distrib prime_aprimedivisor' prime_imp_prime_elem unit_factor_nat_def primepow_gt_0_nat dest: primepow_gt_Suc_0) lemma psi_sqrt: "psi (Discrete.sqrt n) = psi_even n" proof (induction n) case 0 with psi_def psi_even_def show ?case by simp next case (Suc n) then show ?case proof cases assume asm: "\ m. Suc n = m^2" with sqrt_Suc have sqrt_seq: "Discrete.sqrt(Suc n) = Suc(Discrete.sqrt n)" by simp from asm obtain "m" where " Suc n = m^2" by blast with sqrt_seq have "Suc(Discrete.sqrt n) = m" by simp with \Suc n = m^2\ have suc_sqrt_n_sqrt: "(Suc(Discrete.sqrt n))^2 = Suc n" by simp from sqrt_seq have "psi (Discrete.sqrt (Suc n)) = psi (Suc (Discrete.sqrt n))" by simp also from psi_def have "\ = psi (Discrete.sqrt n) + mangoldt (Suc (Discrete.sqrt n))" by simp also from Suc.IH have "psi (Discrete.sqrt n) = psi_even n" . also have "mangoldt (Suc (Discrete.sqrt n)) = mangoldt_even (Suc n)" proof (cases "primepow (Suc(Discrete.sqrt n))") case True with primepow_iff_even_sqr have True2: "primepow_even ((Suc(Discrete.sqrt n))^2)" by simp from suc_sqrt_n_sqrt have "mangoldt_even (Suc n) = mangoldt_even ((Suc(Discrete.sqrt n))^2)" by simp also from mangoldt_even_def True2 have "\ = ln (aprimedivisor ((Suc (Discrete.sqrt n))^2))" by simp also from True have "aprimedivisor ((Suc (Discrete.sqrt n))^2) = aprimedivisor (Suc (Discrete.sqrt n))" by (simp add: aprimedivisor_primepow_power) also from True have "ln (\) = mangoldt (Suc (Discrete.sqrt n))" by (simp add: mangoldt_def) finally show ?thesis .. next case False with primepow_iff_even_sqr have False2: "\ primepow_even ((Suc(Discrete.sqrt n))^2)" by simp from suc_sqrt_n_sqrt have "mangoldt_even (Suc n) = mangoldt_even ((Suc(Discrete.sqrt n))^2)" by simp also from mangoldt_even_def False2 have "\ = 0" by simp also from False have "\ = mangoldt (Suc (Discrete.sqrt n))" by (simp add: mangoldt_def) finally show ?thesis .. qed also from psi_even_def have "psi_even n + mangoldt_even (Suc n) = psi_even (Suc n)" by simp finally show ?case . next assume asm: "\(\m. Suc n = m^2)" with sqrt_Suc have sqrt_eq: "Discrete.sqrt (Suc n) = Discrete.sqrt n" by simp then have lhs: "psi (Discrete.sqrt (Suc n)) = psi (Discrete.sqrt n)" by simp have "\ primepow_even (Suc n)" proof assume "primepow_even (Suc n)" with primepow_even_def obtain "p" "k" where "1 \ k \ prime p \ Suc n = p ^ (2 * k)" by blast with power_even_eq have "Suc n = (p ^ k)^2" by simp with asm show False by blast qed with psi_even_def mangoldt_even_def have rhs: "psi_even (Suc n) = psi_even n" by simp from Suc.IH lhs rhs show ?case by simp qed qed lemma mangoldt_split: "mangoldt d = mangoldt_1 d + mangoldt_even d + mangoldt_odd d" proof (cases "primepow d") case False thus ?thesis by (auto simp: mangoldt_def mangoldt_1_def mangoldt_even_def mangoldt_odd_def dest: primepow_even_imp_primepow primepow_odd_imp_primepow) next case True thus ?thesis by (auto simp: mangoldt_def mangoldt_1_def mangoldt_even_def mangoldt_odd_def primepow_cases) qed lemma psi_split: "psi n = theta n + psi_even n + psi_odd n" by (induction n) (simp_all add: psi_def theta_def psi_even_def psi_odd_def mangoldt_1_def mangoldt_split) lemma psi_mono: "m \ n \ psi m \ psi n" unfolding psi_def by (intro sum_mono2 mangoldt_nonneg) auto lemma psi_pos: "0 \ psi n" by (auto simp: psi_def intro!: sum_nonneg mangoldt_nonneg) lemma mangoldt_odd_pos: "0 \ mangoldt_odd d" using aprimedivisor_gt_Suc_0[of d] by (auto simp: mangoldt_odd_def of_nat_le_iff[of 1, unfolded of_nat_1] Suc_le_eq intro!: ln_ge_zero dest!: primepow_odd_imp_primepow primepow_gt_Suc_0) lemma psi_odd_mono: "m \ n \ psi_odd m \ psi_odd n" using mangoldt_odd_pos sum_mono2[of "{1..n}" "{1..m}" "mangoldt_odd"] by (simp add: psi_odd_def) lemma psi_odd_pos: "0 \ psi_odd n" by (auto simp: psi_odd_def intro!: sum_nonneg mangoldt_odd_pos) lemma psi_theta: "theta n + psi (Discrete.sqrt n) \ psi n" "psi n \ theta n + 2 * psi (Discrete.sqrt n)" using psi_odd_pos[of n] psi_residues_compare[of n] psi_sqrt[of n] psi_split[of n] by simp_all context begin private lemma sum_minus_one: "(\x \ {1..y}. (- 1 :: real) ^ (x + 1)) = (if odd y then 1 else 0)" by (induction y) simp_all private lemma div_invert: fixes x y n :: nat assumes "x > 0" "y > 0" "y \ n div x" shows "x \ n div y" proof - from assms(1,3) have "y * x \ (n div x) * x" by simp also have "\ \ n" by (simp add: minus_mod_eq_div_mult[symmetric]) finally have "y * x \ n" . with assms(2) show ?thesis using div_le_mono[of "y*x" n y] by simp qed lemma sum_expand_lemma: "(\d=1..n. (-1) ^ (d + 1) * psi (n div d)) = (\d = 1..n. (if odd (n div d) then 1 else 0) * mangoldt d)" proof - have **: "x \ n" if "x \ n div y" for x y using div_le_dividend order_trans that by blast have "(\d=1..n. (-1)^(d+1) * psi (n div d)) = (\d=1..n. (-1)^(d+1) * (\e=1..n div d. mangoldt e))" by (simp add: psi_def) also have "\ = (\d = 1..n. \e = 1..n div d. (-1)^(d+1) * mangoldt e)" by (simp add: sum_distrib_left) also from ** have "\ = (\d = 1..n. \e\{y\{1..n}. y \ n div d}. (-1)^(d+1) * mangoldt e)" by (intro sum.cong) auto also have "\ = (\y = 1..n. \x | x \ {1..n} \ y \ n div x. (- 1) ^ (x + 1) * mangoldt y)" by (rule sum.swap_restrict) simp_all also have "\ = (\y = 1..n. \x | x \ {1..n} \ x \ n div y. (- 1) ^ (x + 1) * mangoldt y)" by (intro sum.cong) (auto intro: div_invert) also from ** have "\ = (\y = 1..n. \x \ {1..n div y}. (- 1) ^ (x + 1) * mangoldt y)" by (intro sum.cong) auto also have "\ = (\y = 1..n. (\x \ {1..n div y}. (- 1) ^ (x + 1)) * mangoldt y)" by (intro sum.cong) (simp_all add: sum_distrib_right) also have "\ = (\y = 1..n. (if odd (n div y) then 1 else 0) * mangoldt y)" by (intro sum.cong refl) (simp_all only: sum_minus_one) finally show ?thesis . qed private lemma floor_half_interval: fixes n d :: nat assumes "d \ 0" shows "real (n div d) - real (2 * ((n div 2) div d)) = (if odd (n div d) then 1 else 0)" proof - have "((n div 2) div d) = (n div (2 * d))" by (rule div_mult2_eq[symmetric]) also have "\ = ((n div d) div 2)" by (simp add: mult_ac div_mult2_eq) also have "real (n div d) - real (2 * \) = (if odd (n div d) then 1 else 0)" by (cases "odd (n div d)", cases "n div d = 0 ", simp_all) finally show ?thesis by simp qed lemma fact_expand_psi: "ln (fact n) - 2 * ln (fact (n div 2)) = (\d=1..n. (-1)^(d+1) * psi (n div d))" proof - have "ln (fact n) - 2 * ln (fact (n div 2)) = (\d=1..n. mangoldt d * \n / d\) - 2 * (\d=1..n div 2. mangoldt d * \(n div 2) / d\)" by (simp add: ln_fact_conv_mangoldt) also have "(\d=1..n div 2. mangoldt d * \real (n div 2) / d\) = (\d=1..n. mangoldt d * \real (n div 2) / d\)" by (rule sum.mono_neutral_left) (auto simp: floor_unique[of 0]) also have "2 * \ = (\d=1..n. mangoldt d * 2 * \real (n div 2) / d\)" by (simp add: sum_distrib_left mult_ac) also have "(\d=1..n. mangoldt d * \n / d\) - \ = (\d=1..n. (mangoldt d * \n / d\ - mangoldt d * 2 * \real (n div 2) / d\))" by (simp add: sum_subtractf) also have "\ = (\d=1..n. mangoldt d * (\n / d\ - 2 * \real (n div 2) / d\))" by (simp add: algebra_simps) also have "\ = (\d=1..n. mangoldt d * (if odd(n div d) then 1 else 0))" by (intro sum.cong refl) (simp_all add: floor_conv_div_nat [symmetric] floor_half_interval [symmetric]) also have "\ = (\d=1..n. (if odd(n div d) then 1 else 0) * mangoldt d)" by (simp add: mult_ac) also from sum_expand_lemma[symmetric] have "\ = (\d=1..n. (-1)^(d+1) * psi (n div d))" . finally show ?thesis . qed end lemma psi_expansion_cutoff: assumes "m \ p" shows "(\d=1..2*m. (-1)^(d+1) * psi (n div d)) \ (\d=1..2*p. (-1)^(d+1) * psi (n div d))" "(\d=1..2*p+1. (-1)^(d+1) * psi (n div d)) \ (\d=1..2*m+1. (-1)^(d+1) * psi (n div d))" using assms proof (induction m rule: inc_induct) case (step k) have "(\d = 1..2 * k. (-1)^(d + 1) * psi (n div d)) \ (\d = 1..2 * Suc k. (-1)^(d + 1) * psi (n div d))" by (simp add: psi_mono div_le_mono2) with step.IH(1) show "(\d = 1..2 * k. (-1)^(d + 1) * psi (n div d)) \ (\d = 1..2 * p. (-1)^(d + 1) * psi (n div d))" by simp from step.IH(2) have "(\d = 1..2 * p + 1. (-1)^(d + 1) * psi (n div d)) \ (\d = 1..2 * Suc k + 1. (-1)^(d + 1) * psi (n div d))" . also have "\ \ (\d = 1..2 * k + 1. (-1)^(d + 1) * psi (n div d))" by (simp add: psi_mono div_le_mono2) finally show "(\d = 1..2 * p + 1. (-1)^(d + 1) * psi (n div d)) \ (\d = 1..2 * k + 1. (-1)^(d + 1) * psi (n div d))" . qed simp_all lemma fact_psi_bound_even: assumes "even k" shows "(\d=1..k. (-1)^(d+1) * psi (n div d)) \ ln (fact n) - 2 * ln (fact (n div 2))" proof - have "(\d=1..k. (-1)^(d+1) * psi (n div d)) \ (\d = 1..n. (- 1) ^ (d + 1) * psi (n div d))" proof (cases "k \ n") case True with psi_expansion_cutoff(1)[of "k div 2" "n div 2" n] have "(\d=1..2*(k div 2). (-1)^(d+1) * psi (n div d)) \ (\d = 1..2*(n div 2). (- 1) ^ (d + 1) * psi (n div d))" by simp also from assms have "2*(k div 2) = k" by simp also have "(\d = 1..2*(n div 2). (- 1) ^ (d + 1) * psi (n div d)) \ (\d = 1..n. (- 1) ^ (d + 1) * psi (n div d))" proof (cases "even n") case True then show ?thesis by simp next case False from psi_pos have "(\d = 1..2*(n div 2). (- 1) ^ (d + 1) * psi (n div d)) \ (\d = 1..2*(n div 2) + 1. (- 1) ^ (d + 1) * psi (n div d))" by simp with False show ?thesis by simp qed finally show ?thesis . next case False hence *: "n div 2 \ (k-1) div 2" by simp have "(\d=1..k. (-1)^(d+1) * psi (n div d)) \ (\d=1..2*((k-1) div 2) + 1. (-1)^(d+1) * psi (n div d))" proof (cases "k = 0") case True with psi_pos show ?thesis by simp next case False with sum.cl_ivl_Suc[of "\d. (-1)^(d+1) * psi (n div d)" 1 "k-1"] have "(\d=1..k. (-1)^(d+1) * psi (n div d)) = (\d=1..k-1. (-1)^(d+1) * psi (n div d)) + (-1)^(k+1) * psi (n div k)" by simp also from assms psi_pos have "(-1)^(k+1) * psi (n div k) \ 0" by simp also from assms False have "k-1 = 2*((k-1) div 2) + 1" by presburger finally show ?thesis by simp qed also from * psi_expansion_cutoff(2)[of "n div 2" "(k-1) div 2" n] have "\ \ (\d=1..2*(n div 2) + 1. (-1)^(d+1) * psi (n div d))" by blast also have "\ \ (\d = 1..n. (- 1) ^ (d + 1) * psi (n div d))" by (cases "even n") (simp_all add: psi_def) finally show ?thesis . qed also from fact_expand_psi have "\ = ln (fact n) - 2 * ln (fact (n div 2))" .. finally show ?thesis . qed lemma fact_psi_bound_odd: assumes "odd k" shows "ln (fact n) - 2 * ln (fact (n div 2)) \ (\d=1..k. (-1)^(d+1) * psi (n div d))" proof - from fact_expand_psi have "ln (fact n) - 2 * ln (fact (n div 2)) = (\d = 1..n. (- 1) ^ (d + 1) * psi (n div d))" . also have "\ \ (\d=1..k. (-1)^(d+1) * psi (n div d))" proof (cases "k \ n") case True have "(\d=1..n. (-1)^(d+1) * psi (n div d)) \ ( \d=1..2*(n div 2)+1. (-1)^(d+1) * psi (n div d))" by (cases "even n") (simp_all add: psi_pos) also from True assms psi_expansion_cutoff(2)[of "k div 2" "n div 2" n] have "\ \ (\d=1..k. (-1)^(d+1) * psi (n div d))" by simp finally show ?thesis . next case False have "(\d=1..n. (-1)^(d+1) * psi (n div d)) \ (\d=1..2*((n+1) div 2). (-1)^(d+1) * psi (n div d))" by (cases "even n") (simp_all add: psi_def) also from False assms psi_expansion_cutoff(1)[of "(n+1) div 2" "k div 2" n] have "(\d=1..2*((n+1) div 2). (-1)^(d+1) * psi (n div d)) \ (\d=1..2*(k div 2). (-1)^(d+1) * psi (n div d))" by simp also from assms have "\ \ (\d=1..k. (-1)^(d+1) * psi (n div d))" by (auto elim: oddE simp: psi_pos) finally show ?thesis . qed finally show ?thesis . qed lemma fact_psi_bound_2_3: "psi n - psi (n div 2) \ ln (fact n) - 2 * ln (fact (n div 2))" "ln (fact n) - 2 * ln (fact (n div 2)) \ psi n - psi (n div 2) + psi (n div 3)" proof - show "psi n - psi (n div 2) \ ln (fact n) - 2 * ln (fact (n div 2))" by (rule psi_bounds_ln_fact (2)) next from fact_psi_bound_odd[of 3 n] have "ln (fact n) - 2 * ln (fact (n div 2)) \ (\d = 1..3. (- 1) ^ (d + 1) * psi (n div d))" by simp also have "\ = psi n - psi (n div 2) + psi (n div 3)" by (simp add: sum.atLeast_Suc_atMost numeral_2_eq_2) finally show "ln (fact n) - 2 * ln (fact (n div 2)) \ psi n - psi (n div 2) + psi (n div 3)" . qed lemma ub_ln_1200: "ln 1200 \ 57 / (8 :: real)" proof - have "Some (Float 57 (-3)) = ub_ln 8 1200" by code_simp from ub_ln(1)[OF this] show ?thesis by simp qed lemma psi_double_lemma: assumes "n \ 1200" shows "real n / 6 \ psi n - psi (n div 2)" proof - from ln_fact_diff_bounds have "\ln (fact n) - 2 * ln (fact (n div 2)) - real n * ln 2\ \ 4 * ln (real (if n = 0 then 1 else n)) + 3" . with assms have "ln (fact n) - 2 * ln (fact (n div 2)) \ real n * ln 2 - 4 * ln (real n) - 3" by simp moreover have "real n * ln 2 - 4 * ln (real n) - 3 \ 2 / 3 * n" proof (rule overpower_lemma[of "\n. 2/3 * n" 1200]) show "2 / 3 * 1200 \ 1200 * ln 2 - 4 * ln 1200 - (3::real)" using ub_ln_1200 ln_2_ge by linarith next fix x::real assume "1200 \ x" then have "0 < x" by simp show "((\x. x * ln 2 - 4 * ln x - 3 - 2 / 3 * x) has_real_derivative ln 2 - 4 / x - 2 / 3) (at x)" by (rule derivative_eq_intros refl | simp add: \0 < x\)+ next fix x::real assume "1200 \ x" then have "12 / x \ 12 / 1200" by simp then have "0 \ 0.67 - 4 / x - 2 / 3" by simp also have "0.67 \ ln (2::real)" using ln_2_ge by simp finally show "0 \ ln 2 - 4 / x - 2 / 3" by simp next from assms show "1200 \ real n" by simp qed ultimately have "2 / 3 * real n \ ln (fact n) - 2 * ln (fact (n div 2))" by simp with psi_ubound_3_2[of "n div 3"] have "n/6 + psi (n div 3) \ ln (fact n) - 2 * ln (fact (n div 2))" by simp with fact_psi_bound_2_3[of "n"] show ?thesis by simp qed lemma theta_double_lemma: assumes "n \ 1200" shows "theta (n div 2) < theta n" proof - from psi_theta[of "n div 2"] psi_pos[of "Discrete.sqrt (n div 2)"] have theta_le_psi_n_2: "theta (n div 2) \ psi (n div 2)" by simp have "(Discrete.sqrt n * 18)^2 \ 324 * n" by simp from mult_less_cancel2[of "324" "n" "n"] assms have "324 * n < n^2" by (simp add: power2_eq_square) with \(Discrete.sqrt n * 18)^2 \ 324 * n\ have "(Discrete.sqrt n*18)^2 < n^2" by presburger with power2_less_imp_less assms have "Discrete.sqrt n * 18 < n" by blast with psi_ubound_3_2[of "Discrete.sqrt n"] have "2 * psi (Discrete.sqrt n) < n / 6" by simp with psi_theta[of "n"] have psi_lt_theta_n: "psi n - n / 6 < theta n" by simp from psi_double_lemma[OF assms(1)] have "psi (n div 2) \ psi n - n / 6" by simp with theta_le_psi_n_2 psi_lt_theta_n show ?thesis by simp qed subsection \Proof of the main result\ lemma theta_mono: "mono theta" by (auto simp: theta_def [abs_def] intro!: monoI sum_mono2) lemma theta_lessE: assumes "theta m < theta n" "m \ 1" obtains p where "p \ {m<..n}" "prime p" proof - from mono_invE[OF theta_mono assms(1)] have "m \ n" by blast hence "theta n = theta m + (\p\{m<..n}. if prime p then ln (real p) else 0)" unfolding theta_def using assms(2) by (subst sum.union_disjoint [symmetric]) (auto simp: ivl_disj_un) also note assms(1) finally have "(\p\{m<..n}. if prime p then ln (real p) else 0) \ 0" by simp then obtain p where "p \ {m<..n}" "(if prime p then ln (real p) else 0) \ 0" by (rule sum.not_neutral_contains_not_neutral) thus ?thesis using that[of p] by (auto intro!: exI[of _ p] split: if_splits) qed theorem bertrand: fixes n :: nat assumes "n > 1" shows "\p\{n<..<2*n}. prime p" proof cases assume n_less: "n < 600" define prime_constants where "prime_constants = {2, 3, 5, 7, 13, 23, 43, 83, 163, 317, 631::nat}" from \n > 1\ n_less have "\p \ prime_constants. n < p \ p < 2 * n" unfolding bex_simps greaterThanLessThan_iff prime_constants_def by presburger moreover have "\p\prime_constants. prime p" unfolding prime_constants_def ball_simps HOL.simp_thms by (intro conjI; pratt (silent)) ultimately show ?thesis unfolding greaterThanLessThan_def greaterThan_def lessThan_def by blast next assume n: "\(n < 600)" from n have "theta n < theta (2 * n)" using theta_double_lemma[of "2 * n"] by simp with assms obtain p where "p \ {n<..2*n}" "prime p" by (auto elim!: theta_lessE) moreover from assms have "\prime (2*n)" by (auto dest!: prime_product) with \prime p\ have "p \ 2 * n" by auto ultimately show ?thesis by auto qed subsection \Proof of Mertens' first theorem\ text \ The following proof of Mertens' first theorem was ported from John Harrison's HOL Light proof by Larry Paulson: \ lemma sum_integral_ubound_decreasing': fixes f :: "real \ real" assumes "m \ n" and der: "\x. x \ {of_nat m - 1..of_nat n} \ (g has_field_derivative f x) (at x)" and le: "\x y. \real m - 1 \ x; x \ y; y \ real n\ \ f y \ f x" shows "(\k = m..n. f (of_nat k)) \ g (of_nat n) - g (of_nat m - 1)" proof - have "(\k = m..n. f (of_nat k)) \ (\k = m..n. g (of_nat(Suc k) - 1) - g (of_nat k - 1))" proof (rule sum_mono, clarsimp) fix r assume r: "m \ r" "r \ n" hence "\z>real r - 1. z < real r \ g (real r) - g (real r - 1) = (real r - (real r - 1)) * f z" using assms by (intro MVT2) auto hence "\z\{of_nat r - 1..of_nat r}. g (real r) - g (real r - 1) = f z" by auto then obtain u::real where u: "u \ {of_nat r - 1..of_nat r}" and eq: "g r - g (of_nat r - 1) = f u" by blast have "real m \ u + 1" using r u by auto then have "f (of_nat r) \ f u" using r(2) and u by (intro le) auto then show "f (of_nat r) \ g r - g (of_nat r - 1)" by (simp add: eq) qed also have "\ \ g (of_nat n) - g (of_nat m - 1)" using \m \ n\ by (subst sum_Suc_diff) auto finally show ?thesis . qed lemma Mertens_lemma: assumes "n \ 0" shows "\(\d = 1..n. mangoldt d / real d) - ln n\ \ 4" proof - have *: "\abs(s' - nl + n) \ a; abs(s' - s) \ (k - 1) * n - a\ \ abs(s - nl) \ n * k" for s' s k nl a::real by (auto simp: algebra_simps abs_if split: if_split_asm) have le: "\(\d=1..n. mangoldt d * floor (n / d)) - n * ln n + n\ \ 1 + ln n" using ln_fact_bounds ln_fact_conv_mangoldt assms by simp have "\real n * ((\d = 1..n. mangoldt d / real d) - ln n)\ = \((\d = 1..n. real n * mangoldt d / real d) - n * ln n)\" by (simp add: algebra_simps sum_distrib_left) also have "\ \ real n * 4" proof (rule * [OF le]) have "\(\d = 1..n. mangoldt d * \n / d\) - (\d = 1..n. n * mangoldt d / d)\ = \\d = 1..n. mangoldt d * (\n / d\ - n / d)\" by (simp add: sum_subtractf algebra_simps) also have "\ \ psi n" (is "\?sm\ \ ?rhs") proof - have "-?sm = (\d = 1..n. mangoldt d * (n/d - \n/d\))" by (simp add: sum_subtractf algebra_simps) also have "\ \ (\d = 1..n. mangoldt d * 1)" by (intro sum_mono mult_left_mono mangoldt_nonneg) linarith+ finally have "-?sm \ ?rhs" by (simp add: psi_def) moreover have "?sm \ 0" using mangoldt_nonneg by (simp add: mult_le_0_iff sum_nonpos) ultimately show ?thesis by (simp add: abs_if) qed also have "\ \ 3/2 * real n" by (rule psi_ubound_3_2) also have "\\ (4 - 1) * real n - (1 + ln n)" using ln_le_minus_one [of n] assms by (simp add: divide_simps) finally show "\(\d = 1..n. mangoldt d * real_of_int \real n / real d\) - (\d = 1..n. real n * mangoldt d / real d)\ \ (4 - 1) * real n - (1 + ln n)" . qed finally have "\real n * ((\d = 1..n. mangoldt d / real d) - ln n)\ \ real n * 4" . then show ?thesis using assms mult_le_cancel_left_pos by (simp add: abs_mult) qed lemma Mertens_mangoldt_versus_ln: assumes "I \ {1..n}" shows "\(\i\I. mangoldt i / i) - (\p | prime p \ p \ I. ln p / p)\ \ 3" (is "\?lhs\ \ 3") proof (cases "n = 0") case True with assms show ?thesis by simp next case False have "finite I" using assms finite_subset by blast have "0 \ (\i\I. mangoldt i / i - (if prime i then ln i / i else 0))" using mangoldt_nonneg by (intro sum_nonneg) simp_all moreover have "\ \ (\i = 1..n. mangoldt i / i - (if prime i then ln i / i else 0))" using assms by (intro sum_mono2) (auto simp: mangoldt_nonneg) ultimately have *: "\\i\I. mangoldt i / i - (if prime i then ln i / i else 0)\ \ \\i = 1..n. mangoldt i / i - (if prime i then ln i / i else 0)\" by linarith moreover have "?lhs = (\i\I. mangoldt i / i - (if prime i then ln i / i else 0))" "(\i = 1..n. mangoldt i / i - (if prime i then ln i / i else 0)) = (\d = 1..n. mangoldt d / d) - (\p | prime p \ p \ {1..n}. ln p / p)" using sum.inter_restrict [of _ "\i. ln (real i) / i" "Collect prime", symmetric] by (force simp: sum_subtractf \finite I\ intro: sum.cong)+ ultimately have "\?lhs\ \ \(\d = 1..n. mangoldt d / d) - (\p | prime p \ p \ {1..n}. ln p / p)\" by linarith also have "\ \ 3" proof - have eq_sm: "(\i = 1..n. mangoldt i / i) = (\i \ {p^k |p k. prime p \ p^k \ n \ k \ 1}. mangoldt i / i)" proof (intro sum.mono_neutral_right ballI, goal_cases) case (3 i) hence "\primepow i" by (auto simp: primepow_def Suc_le_eq) thus ?case by (simp add: mangoldt_def) qed (auto simp: Suc_le_eq prime_gt_0_nat) have "(\i = 1..n. mangoldt i / i) - (\p | prime p \ p \ {1..n}. ln p / p) = (\i \ {p^k |p k. prime p \ p^k \ n \ k \ 2}. mangoldt i / i)" proof - have eq: "{p ^ k |p k. prime p \ p ^ k \ n \ 1 \ k} = {p ^ k |p k. prime p \ p ^ k \ n \ 2 \ k} \ {p. prime p \ p \ {1..n}}" (is "?A = ?B \ ?C") proof (intro equalityI subsetI; (elim UnE)?) fix x assume "x \ ?A" then obtain p k where "x = p ^ k" "prime p" "p ^ k \ n" "k \ 1" by auto thus "x \ ?B \ ?C" by (cases "k \ 2") (auto simp: prime_power_iff Suc_le_eq) next fix x assume "x \ ?B" then obtain p k where "x = p ^ k" "prime p" "p ^ k \ n" "k \ 1" by auto thus "x \ ?A" by (auto simp: prime_power_iff Suc_le_eq) next fix x assume "x \ ?C" then obtain p where "x = p ^ 1" "1 \ (1::nat)" "prime p" "p ^ 1 \ n" by auto thus "x \ ?A" by blast qed have eqln: "(\p | prime p \ p \ {1..n}. ln p / p) = (\p | prime p \ p \ {1..n}. mangoldt p / p)" by (rule sum.cong) auto have "(\i \ {p^k |p k. prime p \ p^k \ n \ k \ 1}. mangoldt i / i) = (\i \ {p ^ k |p k. prime p \ p ^ k \ n \ 2 \ k} \ {p. prime p \ p \ {1..n}}. mangoldt i / i)" by (subst eq) simp_all also have "\ = (\i \ {p^k |p k. prime p \ p^k \ n \ k \ 2}. mangoldt i / i) + (\p | prime p \ p \ {1..n}. mangoldt p / p)" by (intro sum.union_disjoint) (auto simp: prime_power_iff finite_nat_set_iff_bounded_le) also have "\ = (\i \ {p^k |p k. prime p \ p^k \ n \ k \ 2}. mangoldt i / i) + (\p | prime p \ p \ {1..n}. ln p / p)" by (simp only: eqln) finally show ?thesis using eq_sm by auto qed have "(\p | prime p \ p \ {1..n}. ln p / p) \ (\p | prime p \ p \ {1..n}. mangoldt p / p)" using mangoldt_nonneg by (auto intro: sum_mono) also have "\ \ (\i = Suc 0..n. mangoldt i / i)" by (intro sum_mono2) (auto simp: mangoldt_nonneg) finally have "0 \ (\i = 1..n. mangoldt i / i) - (\p | prime p \ p \ {1..n}. ln p / p)" by simp moreover have "(\i = 1..n. mangoldt i / i) - (\p | prime p \ p \ {1..n}. ln p / p) \ 3" (is "?M - ?L \ 3") proof - have *: "\q. \j\{1..n}. prime q \ 1 \ q \ q \ n \ (q ^ j = p ^ k \ mangoldt (p ^ k) / real p ^ k \ ln (real q) / real q ^ j)" if "prime p" "p ^ k \ n" "1 \ k" for p k proof - have "mangoldt (p ^ k) / real p ^ k \ ln p / p ^ k" using that by (simp add: divide_simps) moreover have "p \ n" using that self_le_power[of p k] by (simp add: prime_ge_Suc_0_nat) moreover have "k \ n" proof - have "k < 2^k" using of_nat_less_two_power of_nat_less_numeral_power_cancel_iff by blast also have "\ \ p^k" by (simp add: power_mono prime_ge_2_nat that) also have "\ \ n" by (simp add: that) finally show ?thesis by (simp add: that) qed ultimately show ?thesis using prime_ge_1_nat that by auto (use atLeastAtMost_iff in blast) qed have finite: "finite {p ^ k |p k. prime p \ p ^ k \ n \ 1 \ k}" by (rule finite_subset[of _ "{..n}"]) auto have "?M \ (\(x, k)\{p. prime p \ p \ {1..n}} \ {1..n}. ln (real x) / real x ^ k)" by (subst eq_sm, intro sum_le_included [where i = "\(p,k). p^k"]) (insert * finite, auto) also have "\ = (\p | prime p \ p \ {1..n}. (\k = 1..n. ln p / p^k))" by (subst sum.Sigma) auto also have "\ = ?L + (\p | prime p \ p \ {1..n}. (\k = 2..n. ln p / p^k))" by (simp add: comm_monoid_add_class.sum.distrib sum.atLeast_Suc_atMost numeral_2_eq_2) finally have "?M - ?L \ (\p | prime p \ p \ {1..n}. (\k = 2..n. ln p / p^k))" by (simp add: algebra_simps) also have "\ = (\p | prime p \ p \ {1..n}. ln p * (\k = 2..n. inverse p ^ k))" by (simp add: field_simps sum_distrib_left) also have "\ = (\p | prime p \ p \ {1..n}. ln p * (((inverse p)\<^sup>2 - inverse p ^ Suc n) / (1 - inverse p)))" by (intro sum.cong refl) (simp add: sum_gp) also have "\ \ (\p | prime p \ p \ {1..n}. ln p * inverse (real (p * (p - 1))))" by (intro sum_mono mult_left_mono) (auto simp: divide_simps power2_eq_square of_nat_diff mult_less_0_iff) also have "\ \ (\p = 2..n. ln p * inverse (real (p * (p - 1))))" by (rule sum_mono2) (use prime_ge_2_nat in auto) also have "\ \ (\i = 2..n. ln i / (i - 1)\<^sup>2)" unfolding divide_inverse power2_eq_square mult.assoc by (auto intro: sum_mono mult_left_mono mult_right_mono) also have "\ \ 3" proof (cases "n \ 3") case False then show ?thesis proof (cases "n \ 2") case False then show ?thesis by simp next case True then have "n = 2" using False by linarith with ln_le_minus_one [of 2] show ?thesis by simp qed next case True have "(\i = 3..n. ln (real i) / (real (i - Suc 0))\<^sup>2) \ (ln (of_nat n - 1)) - (ln (of_nat n)) - (ln (of_nat n) / (of_nat n - 1)) + 2 * ln 2" proof - have 1: "((\z. ln (z - 1) - ln z - ln z / (z - 1)) has_field_derivative ln x / (x - 1)\<^sup>2) (at x)" if x: "x \ {2..real n}" for x by (rule derivative_eq_intros | rule refl | (use x in \force simp: power2_eq_square divide_simps\))+ have 2: "ln y / (y - 1)\<^sup>2 \ ln x / (x - 1)\<^sup>2" if xy: "2 \ x" "x \ y" "y \ real n" for x y proof (cases "x = y") case False define f' :: "real \ real" where "f' = (\u. ((u - 1)\<^sup>2 / u - ln u * (2 * u - 2)) / (u - 1) ^ 4)" have f'_altdef: "f' u = inverse u * inverse ((u - 1)\<^sup>2) - 2 * ln u / (u - 1) ^ 3" if u: "u \ {x..y}" for u::real unfolding f'_def using u (* TODO ugly *) by (simp add: eval_nat_numeral divide_simps) (simp add: algebra_simps)? have deriv: "((\z. ln z / (z - 1)\<^sup>2) has_field_derivative f' u) (at u)" if u: "u \ {x..y}" for u::real unfolding f'_def by (rule derivative_eq_intros refl | (use u xy in \force simp: divide_simps\))+ hence "\z>x. z < y \ ln y / (y - 1)\<^sup>2 - ln x / (x - 1)\<^sup>2 = (y - x) * f' z" using xy and \x \ y\ by (intro MVT2) auto then obtain \::real where "x < \" "\ < y" and \: "ln y / (y - 1)\<^sup>2 - ln x / (x - 1)\<^sup>2 = (y - x) * f' \" by blast have "f' \ \ 0" proof - have "2/3 \ ln (2::real)" by (fact ln_2_ge') also have "\ \ ln \" using \x < \\ xy by auto finally have "1 \ 2 * ln \" by simp then have *: "\ \ \ * (2 * ln \)" using \x < \\ xy by auto hence "\ - 1 \ ln \ * 2 * \" by (simp add: algebra_simps) hence "1 / (\ * (\ - 1)\<^sup>2) \ ln \ * 2 / (\ - 1) ^ 3" using xy \x < \\ by (simp add: divide_simps power_eq_if) thus ?thesis using xy \x < \\ \\ < y\ by (subst f'_altdef) (auto simp: divide_simps) qed then have "(ln y / (y - 1)\<^sup>2 - ln x / (x - 1)\<^sup>2) \ 0" using \x \ y\ by (simp add: mult_le_0_iff \) then show ?thesis by simp qed simp_all show ?thesis using sum_integral_ubound_decreasing' [OF \3 \ n\, of "\z. ln(z-1) - ln z - ln z / (z - 1)" "\z. ln z / (z-1)\<^sup>2"] 1 2 \3 \ n\ by (auto simp: in_Reals_norm of_nat_diff) qed also have "\ \ 2" proof - have "ln (real n - 1) - ln n \ 0" "0 \ ln n / (real n - 1)" using \3 \ n\ by auto then have "ln (real n - 1) - ln n - ln n / (real n - 1) \ 0" by linarith with ln_2_less_1 show ?thesis by linarith qed also have "\ \ 3 - ln 2" using ln_2_less_1 by (simp add: algebra_simps) finally show ?thesis using True by (simp add: algebra_simps sum.atLeast_Suc_atMost [of 2 n]) qed finally show ?thesis . qed ultimately show ?thesis by linarith qed finally show ?thesis . qed proposition Mertens: assumes "n \ 0" shows "\(\p | prime p \ p \ n. ln p / of_nat p) - ln n\ \ 7" proof - have "\(\d = 1..n. mangoldt d / real d) - (\p | prime p \ p \ {1..n}. ln (real p) / real p)\ \ 7 - 4" using Mertens_mangoldt_versus_ln [of "{1..n}" n] by simp_all also have "{p. prime p \ p \ {1..n}} = {p. prime p \ p \ n}" using atLeastAtMost_iff prime_ge_1_nat by blast finally have "\(\d = 1..n. mangoldt d / real d) - (\p\\. ln (real p) / real p)\ \ 7 - 4" . moreover from assms have "\(\d = 1..n. mangoldt d / real d) - ln n\ \ 4" by (rule Mertens_lemma) ultimately show ?thesis by linarith qed end diff --git a/thys/Bertrands_Postulate/bertrand.ML b/thys/Bertrands_Postulate/bertrand.ML --- a/thys/Bertrands_Postulate/bertrand.ML +++ b/thys/Bertrands_Postulate/bertrand.ML @@ -1,159 +1,157 @@ signature BERTRAND = sig type cache = (int * thm) list datatype primepow_thm = PrimepowThm of thm * thm | NotPrimepowThm of thm val prime_conv : Proof.context -> cache -> conv val primepow_conv : Proof.context -> cache -> conv val pre_mangoldt_conv : Proof.context -> cache -> conv val prove_primepow : Proof.context -> cache -> term -> primepow_thm val prove_pre_mangoldt : Proof.context -> cache -> term -> thm val prove_psi : Proof.context -> int -> (int * int * thm) list val mk_prime_cache : Proof.context -> int -> cache end structure Bertrand : BERTRAND = struct type cache = (int * thm) list datatype primepow_cert = Primepow of int * int | NotPrimepow of int * int datatype primepow_thm = PrimepowThm of thm * thm | NotPrimepowThm of thm val mk_nat = HOLogic.mk_number \<^Type>\nat\ fun mk_primepow_cert n = let fun find_prime_divisor n = let fun go k = if n mod k = 0 then k else go (k + 1) in go 2 end fun divide_out p m acc = if m mod p = 0 then divide_out p (m div p) (acc + 1) else (acc, m) val p = find_prime_divisor n val (k, m) = divide_out p n 0 in if m = 1 then Primepow (p, k) else NotPrimepow (p, find_prime_divisor m) end fun bool_thm_to_eq_thm thm = case HOLogic.dest_Trueprop (Thm.concl_of thm) of \<^Const_>\Not for _\ => thm RS @{thm Eq_FalseI} | _ => thm RS @{thm Eq_TrueI} fun prime_conv ctxt cache ct = case Thm.term_of ct of \<^Const_>\Trueprop for \<^Const>\prime \<^Type>\nat\ for p\\ => Pratt.prove_prime cache (snd (HOLogic.dest_number p)) ctxt |> fst |> the |> bool_thm_to_eq_thm | _ => raise CTERM ("prime_conv", [ct]) fun prime_sieve n = let fun go ps k = if k > n then rev ps else if exists (fn p => k mod p = 0) ps then go ps (k + 1) else go (k :: ps) (k + 1) in go [] 2 end fun fold_accum f xs acc = fold (fn x => fn (ys, acc) => case f x acc of (y, acc') => (y :: ys, acc')) xs ([], acc) |> apfst rev fun mk_prime_cache ctxt n = let val ps = prime_sieve n in fold_accum (fn p => fn cache => Pratt.prove_prime cache p ctxt) ps [] |> snd end fun prove_primepow ctxt cache t = let val (_, n) = HOLogic.dest_number t fun inst xs = Drule.infer_instantiate' ctxt (map (SOME o Thm.cterm_of ctxt o mk_nat) xs) - val prove = Goal.prove ctxt [] [] fun prove_prime' p = the (fst (Pratt.prove_prime cache p ctxt)) in if n <= 0 then raise TERM ("prove_primepow", [t]) else if n = 1 then NotPrimepowThm @{thm not_primepow_1_nat} else case mk_primepow_cert n of Primepow (p, k) => let val thm = prove_prime' p RS inst [p, k, n] @{thm primepowI} - val thm' = prove (Thm.concl_of thm) (fn {context = ctxt, ...} => + val thm' = Goal.prove ctxt [] [] (Thm.concl_of thm) (fn {context = ctxt, ...} => HEADGOAL (resolve_tac ctxt [thm]) THEN ALLGOALS (Simplifier.simp_tac ctxt)) in PrimepowThm (thm' RS @{thm conjunct1}, thm' RS @{thm conjunct2}) end | NotPrimepow (p, q) => let val thm = inst [p, q, n] @{thm not_primepowI} OF (map prove_prime' [p, q]) - val thm' = prove (Thm.concl_of thm) (fn {context = ctxt, ...} => + val thm' = Goal.prove ctxt [] [] (Thm.concl_of thm) (fn {context = ctxt, ...} => HEADGOAL (resolve_tac ctxt [thm]) THEN ALLGOALS (Simplifier.simp_tac ctxt)) in NotPrimepowThm thm' end end fun primepow_conv ctxt cache ct = case prove_primepow ctxt cache (Thm.term_of ct) of PrimepowThm (thm, _) => bool_thm_to_eq_thm thm | NotPrimepowThm thm => bool_thm_to_eq_thm thm fun prove_pre_mangoldt ctxt cache t = case prove_primepow ctxt cache t of PrimepowThm (thm1, thm2) => @{thm pre_mangoldt_primepow} OF [thm1, thm2] | NotPrimepowThm thm => thm RS @{thm pre_mangoldt_notprimepow} fun pre_mangoldt_conv ctxt cache = (fn thm => thm RS @{thm eq_reflection}) o prove_pre_mangoldt ctxt cache o Thm.term_of val nat_eq_ss = simpset_of (put_simpset HOL_basic_ss @{context} addsimps @{thms Numeral_Simprocs.semiring_norm}) fun prove_nat_eq ctxt (t1, t2) = let val goal = HOLogic.mk_eq (t1, t2) |> HOLogic.mk_Trueprop - fun tac {context = ctxt, ...} = - HEADGOAL (resolve_tac ctxt @{thms mult_1_left mult_1_right}) - ORELSE HEADGOAL (Simplifier.simp_tac (put_simpset nat_eq_ss ctxt)) in - Goal.prove ctxt [] [] goal tac + Goal.prove ctxt [] [] goal (fn {context = goal_ctxt, ...} => + HEADGOAL (resolve_tac goal_ctxt @{thms mult_1_left mult_1_right}) + ORELSE HEADGOAL (Simplifier.simp_tac (put_simpset nat_eq_ss goal_ctxt))) end fun prove_psi ctxt n = let val cache = mk_prime_cache ctxt n fun go thm x x' k acc = if k > n then rev acc else let val pre_mangoldt_thm = prove_pre_mangoldt ctxt cache (mk_nat k) val y = pre_mangoldt_thm |> Thm.concl_of |> HOLogic.dest_Trueprop |> HOLogic.dest_eq |> snd val y' = HOLogic.dest_number y |> snd val eq_thm1 = prove_nat_eq ctxt (\<^Const>\plus \<^Type>\nat\\ $ mk_nat (k - 1) $ @{term "1 :: nat"}, mk_nat k) val z = if y' = 1 then x else mk_nat (x' * y') val eq_thm2 = prove_nat_eq ctxt (\<^Const>\times \<^Type>\nat\ for x y\, z) val thm' = @{thm eval_psi_aux2} OF [thm, pre_mangoldt_thm, eq_thm1, eq_thm2] in go thm' z (x' * y') (k + 1) ((k - 1, x', thm) :: acc) end in go @{thm eval_psi_aux1} @{term "numeral Num.One :: nat"} 1 1 [] end end diff --git a/thys/Case_Labeling/Examples/Hoare/labeled_hoare_tac.ML b/thys/Case_Labeling/Examples/Hoare/labeled_hoare_tac.ML --- a/thys/Case_Labeling/Examples/Hoare/labeled_hoare_tac.ML +++ b/thys/Case_Labeling/Examples/Hoare/labeled_hoare_tac.ML @@ -1,222 +1,224 @@ (* Title: labeled_hoare_tac.ML Author: Leonor Prensa Nieto & Tobias Nipkow & Lars Noschinski Derivation of the proof rules and, most importantly, the VCG tactic. *) signature LABELED_HOARE = sig val wrap_label_tac: Proof.context -> (int -> tactic) -> (int -> tactic) val hoare_rule_tac: Proof.context -> term list * thm -> (int -> tactic) -> bool -> int -> tactic val hoare_tac: Proof.context -> (int -> tactic) -> int -> tactic end; structure Labeled_Hoare: LABELED_HOARE = struct (*** The tactics ***) val to_prems_ths = @{thms LABELs_to_prems} val to_concl_ths = @{thms LABELs_to_concl} infix 0 THEN_ELSE'; fun (tac THEN_ELSE' (tac1, tac2)) x = tac x THEN_ELSE (tac1 x, tac2 x) fun wrap_label_tac ctxt tac = let fun wrap_tac i st = (resolve_tac ctxt to_prems_ths THEN_ELSE' (wrap_tac THEN_ALL_NEW eresolve_tac ctxt to_concl_ths, tac) ) i st in wrap_tac end (*****************************************************************************) (** The function Mset makes the theorem **) (** "?Mset <= {(x1,...,xn). ?P (x1,...,xn)} ==> ?Mset <= {s. ?P s}", **) (** where (x1,...,xn) are the variables of the particular program we are **) (** working on at the moment of the call **) (*****************************************************************************) local (** maps (%x1 ... xn. t) to [x1,...,xn] **) fun abs2list (Const (@{const_name case_prod}, _) $ Abs (x, T, t)) = Free (x, T) :: abs2list t | abs2list (Abs (x, T, _)) = [Free (x, T)] | abs2list _ = []; (** maps {(x1,...,xn). t} to [x1,...,xn] **) fun mk_vars (Const (@{const_name Collect},_) $ T) = abs2list T | mk_vars _ = []; (** abstraction of body over a tuple formed from a list of free variables. Types are also built **) fun mk_abstupleC [] body = absfree ("x", HOLogic.unitT) body | mk_abstupleC [v] body = absfree (dest_Free v) body | mk_abstupleC (v :: w) body = let val (x, T) = dest_Free v; val z = mk_abstupleC w body; val T2 = (case z of Abs (_, T, _) => T | Const (_, Type (_, [_, Type (_, [T, _])])) $ _ => T); in Const (@{const_name case_prod}, (T --> T2 --> HOLogic.boolT) --> HOLogic.mk_prodT (T, T2) --> HOLogic.boolT) $ absfree (x, T) z end; (** maps [x1,...,xn] to (x1,...,xn) and types**) fun mk_bodyC [] = HOLogic.unit | mk_bodyC [x] = x | mk_bodyC (x :: xs) = let val (_, T) = dest_Free x; val z = mk_bodyC xs; val T2 = (case z of Free (_, T) => T | Const (@{const_name Pair}, Type ("fun", [_, Type ("fun", [_, T])])) $ _ $ _ => T); in Const (@{const_name Pair}, [T, T2] ---> HOLogic.mk_prodT (T, T2)) $ x $ z end; (** maps a subgoal of the form: VARS x1 ... xn {._.} _ {._.} or to [x1,...,xn] **) fun get_vars c = let val d = Logic.strip_assums_concl c; val Const _ $ pre $ _ $ _ $ _ = HOLogic.dest_Trueprop d; in mk_vars pre end; fun mk_CollectC tm = let val T as Type ("fun",[t,_]) = fastype_of tm; in HOLogic.Collect_const t $ tm end; fun inclt ty = Const (@{const_name Orderings.less_eq}, [ty,ty] ---> HOLogic.boolT); in fun Mset ctxt prop = let val [(Mset, _), (P, _)] = Variable.variant_frees ctxt [] [("Mset", ()), ("P", ())]; val vars = get_vars prop; val varsT = fastype_of (mk_bodyC vars); val big_Collect = mk_CollectC (mk_abstupleC vars (Free (P, varsT --> HOLogic.boolT) $ mk_bodyC vars)); val small_Collect = mk_CollectC (Abs ("x", varsT, Free (P, varsT --> HOLogic.boolT) $ Bound 0)); val MsetT = fastype_of big_Collect; fun Mset_incl t = HOLogic.mk_Trueprop (inclt MsetT $ Free (Mset, MsetT) $ t); val impl = Logic.mk_implies (Mset_incl big_Collect, Mset_incl small_Collect); - val th = Goal.prove ctxt [Mset, P] [] impl (fn _ => blast_tac ctxt 1); + val th = + Goal.prove ctxt [Mset, P] [] impl + (fn {context = goal_ctxt, ...} => blast_tac goal_ctxt 1); in (vars, th) end; end; (*****************************************************************************) (** Simplifying: **) (** Some useful lemmata, lists and simplification tactics to control which **) (** theorems are used to simplify at each moment, so that the original **) (** input does not suffer any unexpected transformation **) (*****************************************************************************) (**Simp_tacs**) fun before_set2pred_simp_tac ctxt = simp_tac (put_simpset HOL_basic_ss ctxt addsimps [Collect_conj_eq RS sym, @{thm Compl_Collect}]); fun split_simp_tac ctxt = simp_tac (put_simpset HOL_basic_ss ctxt addsimps [@{thm split_conv}]); (*****************************************************************************) (** set_to_pred_tac 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 (split_all_tac) **) (*****************************************************************************) fun set_to_pred_tac ctxt var_names = SUBGOAL (fn (_, i) => before_set2pred_simp_tac ctxt i THEN_MAYBE EVERY [ resolve_tac ctxt [subsetI] i, resolve_tac ctxt [CollectI] i, dresolve_tac ctxt [CollectD] i, TRY (split_all_tac ctxt i) THEN_MAYBE (rename_tac var_names i THEN full_simp_tac (put_simpset HOL_basic_ss ctxt addsimps [@{thm split_conv}]) i)]); (*******************************************************************************) (** basic_simp_tac is called to simplify all verification conditions. It does **) (** a light simplification by applying "mem_Collect_eq", then it calls **) (** max_simp_tac, which solves subgoals of the form "A <= A", **) (** and transforms any other into predicates, applying then **) (** the tactic chosen by the user, which may solve the subgoal completely. **) (*******************************************************************************) fun max_simp_tac ctxt var_names tac = FIRST' [resolve_tac ctxt [subset_refl], set_to_pred_tac ctxt var_names THEN_MAYBE' tac]; fun basic_simp_tac ctxt var_names tac = wrap_label_tac ctxt ( simp_tac (put_simpset HOL_basic_ss ctxt addsimps [mem_Collect_eq, @{thm split_conv}] addsimprocs [Record.simproc]) THEN_MAYBE' max_simp_tac ctxt var_names tac ); (** hoare_rule_tac **) fun hoare_rule_tac ctxt (vars, Mlem) tac = let val var_names = map (fst o dest_Free) vars; fun wlp_tac i = resolve_tac ctxt @{thms LSeqRule} i THEN rule_tac false (i + 1) and rule_tac pre_cond i st = st |> (*abstraction over st prevents looping*) ((wlp_tac i THEN rule_tac pre_cond i) ORELSE TRY (FIRST [ resolve_tac ctxt @{thms LSkipRule} i, resolve_tac ctxt @{thms LAbortRule} i, EVERY [ resolve_tac ctxt @{thms LBasicRule} i, wrap_label_tac ctxt (resolve_tac ctxt [Mlem]) i, split_simp_tac ctxt i], EVERY [ resolve_tac ctxt @{thms LCondRule} i, rule_tac false (i + 2), rule_tac false (i + 1)], EVERY [ resolve_tac ctxt @{thms LWhileRule} i, basic_simp_tac ctxt var_names tac (i + 2), rule_tac true (i + 1)]] THEN (if pre_cond then basic_simp_tac ctxt var_names tac i else wrap_label_tac ctxt (resolve_tac ctxt [subset_refl]) i) ) ); in rule_tac end; (** tac is the tactic the user chooses to solve or simplify **) (** the final verification conditions **) fun hoare_tac ctxt tac = SUBGOAL (fn (goal, i) => let val mset = Mset ctxt goal in SELECT_GOAL (resolve_tac ctxt @{thms Initial_Label} 1 THEN hoare_rule_tac ctxt mset tac true 1) i end); end; diff --git a/thys/Containers/containers_generator.ML b/thys/Containers/containers_generator.ML --- a/thys/Containers/containers_generator.ML +++ b/thys/Containers/containers_generator.ML @@ -1,130 +1,131 @@ (* Author: René Thiemann, UIBK *) (* The generators have been written as part of the IsaFoR/CeTA formalization. *) signature CONTAINERS_GENERATOR = sig val mk_is_c_dots : typ -> string -> term (* const sort choices *) val derive_set_map_impl : string -> sort -> (string * term) list -> string -> string -> theory -> theory (* const_name, sort, mk_none, typ_name *) val derive_none : string -> sort -> (typ -> term) -> string -> theory -> theory (* const_name base_name *) val register_is_c_dots_lemma : string -> string -> thm -> theory -> theory (* typ const_name defs base_name *) val derive_is_c_dots_lemma : typ -> string -> thm list -> string -> theory -> theory val mk_Some : term -> term (* p1 /\ ... /\ pn *) val HOLogic_list_conj : term list -> term val all_tys : term -> typ list -> term -> term val is_class_instance : theory -> string -> sort -> bool end structure Containers_Generator : CONTAINERS_GENERATOR = struct open Generator_Aux fun is_class_instance thy tname class = Proof_Context.read_type_name {proper = true, strict = true} (Proof_Context.init_global thy) tname |> (fn T => Sign.of_sort thy (T, class)) fun all_tys comp free_types = let val Ts = fastype_of comp |> strip_type |> fst |> drop_last |> List.last |> dest_Type |> snd in rename_types (Ts ~~ free_types) end fun HOLogic_list_conj [] = @{term True} | HOLogic_list_conj [x] = x | HOLogic_list_conj (x :: xs) = HOLogic.mk_conj (x, HOLogic_list_conj xs) fun mk_Some t = let val ty = fastype_of t in Const (@{const_name Some}, ty --> Type (@{type_name option}, [ty])) $ t end fun mk_set_map_impl choices ty choice thy = let val smi = case AList.lookup (op =) choices choice of SOME smi => smi | NONE => if choice = "" then error "you must provide some constant as parameter" else Syntax.read_term (Proof_Context.init_global thy) choice val smi_ty = fastype_of smi val pty = Type (@{type_name phantom},[ty,smi_ty]) val ph = Const (@{const_name phantom}, smi_ty --> pty) val res = ph $ smi in res end fun derive_set_map_impl smi_const smi_sort choices typ_name choice thy = let val base_name = Long_Name.base_name typ_name val smi_name = Long_Name.base_name smi_const val _ = writeln ("use " ^ choice ^ " as " ^ smi_name ^ " for type " ^ base_name) val (ty,vs) = Generator_Aux.typ_and_vs_of_typname thy typ_name @{sort type} val set_map_impl_rhs = mk_set_map_impl choices ty choice thy val set_map_impl_ty = Term.fastype_of set_map_impl_rhs val set_map_impl_def = Generator_Aux.mk_def set_map_impl_ty smi_const set_map_impl_rhs val (_,lthy) = Class.instantiation ([typ_name],vs,smi_sort) thy |> Generator_Aux.define_overloaded (smi_name ^ "_" ^ base_name ^ "_def", set_map_impl_def) val thy' = Class.prove_instantiation_exit (fn ctxt => Class.intro_classes_tac ctxt []) lthy val _ = writeln ("registered " ^ base_name ^ " in class " ^ smi_name) in thy' end fun derive_none const sort mk_none typ_name thy = let val base_name = Long_Name.base_name typ_name val const_name = Long_Name.base_name const val sort_name = hd sort |> Long_Name.base_name val _ = writeln ("use None as trivial implementation of " ^ sort_name ^ " for type " ^ base_name) val (ty,vs) = Generator_Aux.typ_and_vs_of_typname thy typ_name @{sort type} val none_rhs = mk_none ty val none_ty = Term.fastype_of none_rhs val const_def = Generator_Aux.mk_def none_ty const none_rhs val (none_thm,lthy) = Class.instantiation ([typ_name],vs,sort) thy |> Generator_Aux.define_overloaded (const_name ^ "_" ^ base_name ^ "_def", const_def) val thy' = Class.prove_instantiation_exit (fn ctxt => Class.intro_classes_tac ctxt [] THEN unfold_tac ctxt [none_thm] THEN (REPEAT (force_tac ctxt 1))) lthy val _ = writeln ("registered " ^ base_name ^ " in class " ^ sort_name) in thy' end fun mk_is_c_dots ty dots = let val ity = Type (@{type_name itself}, [ty]) val it = Const (@{const_name Pure.type},ity) val res = Const (dots, ity --> @{typ bool}) $ it in res end fun register_is_c_dots_lemma const_name base_name thm thy = let val name = Long_Name.base_name const_name ^ "_" ^ base_name val lthy_map = Local_Theory.note ((Binding.name name, @{attributes [simp,code_post]}), [thm]) #> #2 val thy' = Named_Target.theory_map lthy_map thy val _ = writeln ("derived " ^ name ^ "-lemma") in thy' end fun derive_is_c_dots_lemma ty const_name defs base_name thy = let val is_c_dots = HOLogic.mk_Trueprop (mk_is_c_dots ty const_name) - val thm = Goal.prove_future (Proof_Context.init_global thy) [] [] is_c_dots (fn {prems = _, context = ctxt} => - unfold_tac ctxt (defs @ @{thms ID_Some option.simps}) - THEN blast_tac ctxt 1 - ) + val thm = Goal.prove_future (Proof_Context.init_global thy) [] [] is_c_dots + (fn {prems = _, context = ctxt} => + unfold_tac ctxt (defs @ @{thms ID_Some option.simps}) + THEN blast_tac ctxt 1 + ) in register_is_c_dots_lemma const_name base_name thm thy end end diff --git a/thys/Deriving/Comparator_Generator/comparator_generator.ML b/thys/Deriving/Comparator_Generator/comparator_generator.ML --- a/thys/Deriving/Comparator_Generator/comparator_generator.ML +++ b/thys/Deriving/Comparator_Generator/comparator_generator.ML @@ -1,641 +1,643 @@ (* Title: Deriving class instances for datatypes Author: Christian Sternagel and René Thiemann Maintainer: Christian Sternagel and René Thiemann License: LGPL *) signature COMPARATOR_GENERATOR = sig type info = {map : term, (* take % x. x, if there is no map *) pcomp : term, (* partial comparator *) comp : term, (* full comparator *) comp_def : thm option, (* definition of comparator, important for nesting *) map_comp : thm option, (* compositionality of map, important for nesting *) partial_comp_thms : thm list, (* first eq, then sym, finally trans *) comp_thm : thm, (* comparator acomp \ \ \ comparator (full_comp acomp \) *) used_positions : bool list} (* registers @{term comparator_of :: "some_type :: linorder comparator"} where some_type must just be a type without type-arguments *) val register_comparator_of : string -> local_theory -> local_theory val register_foreign_comparator : typ -> (* type-constant without type-variables *) term -> (* comparator for type *) thm -> (* comparator thm for provided comparator *) local_theory -> local_theory val register_foreign_partial_and_full_comparator : string -> (* long type name *) term -> (* map function, should be \x. x, if there is no map *) term -> (* partial comparator of type ('a => order, 'b)ty => ('a,'b)ty => order, where 'a is used, 'b is unused *) term -> (* (full) comparator of type ('a \ 'a \ order) \ ('a,'b)ty \ ('a,'b)ty \ order, where 'a is used, 'b is unused *) thm option -> (* comp_def, should be full_comp = pcomp o map acomp ..., important for nesting *) thm option -> (* map compositionality, important for nesting *) thm -> (* partial eq thm for full comparator *) thm -> (* partial sym thm for full comparator *) thm -> (* partial trans thm for full comparator *) thm -> (* full thm: comparator a-comp => comparator (full_comp a-comp) *) bool list -> (*used positions*) local_theory -> local_theory datatype comparator_type = Linorder | BNF val generate_comparators_from_bnf_fp : string -> (* name of type *) local_theory -> ((term * thm list) list * (* partial comparators + simp-rules *) (term * thm) list) * (* non-partial comparator + def_rule *) local_theory val generate_comparator : comparator_type -> string -> (* name of type *) local_theory -> local_theory val get_info : Proof.context -> string -> info option (* ensures that the info will be available on later requests *) val ensure_info : comparator_type -> string -> local_theory -> local_theory end structure Comparator_Generator : COMPARATOR_GENERATOR = struct open Generator_Aux datatype comparator_type = BNF | Linorder val debug = false fun debug_out s = if debug then writeln s else () val orderT = @{typ order} fun compT T = T --> T --> orderT val orderify = map_atyps (fn T => T --> orderT) fun pcompT T = orderify T --> T --> orderT type info = {map : term, pcomp : term, comp : term, comp_def : thm option, map_comp : thm option, partial_comp_thms : thm list, comp_thm : thm, used_positions : bool list}; structure Data = Generic_Data ( type T = info Symtab.table; val empty = Symtab.empty; val extend = I; val merge = Symtab.merge (fn (info1 : info, info2 : info) => #comp info1 = #comp info2); ); fun add_info T info = Data.map (Symtab.update_new (T, info)) val get_info = Context.Proof #> Data.get #> Symtab.lookup fun the_info ctxt tyco = (case get_info ctxt tyco of SOME info => info | NONE => error ("no comparator information available for type " ^ quote tyco)) fun declare_info tyco m p c c_def m_comp p_thms c_thm used_pos = Local_Theory.declaration {syntax = false, pervasive = false} (fn phi => add_info tyco {map = Morphism.term phi m, pcomp = Morphism.term phi p, comp = Morphism.term phi c, comp_def = Option.map (Morphism.thm phi) c_def, map_comp = Option.map (Morphism.thm phi) m_comp, partial_comp_thms = Morphism.fact phi p_thms, comp_thm = Morphism.thm phi c_thm, used_positions = used_pos}) val EQ = 0 val SYM = 1 val TRANS = 2 fun register_foreign_partial_and_full_comparator tyco m p c c_def m_comp eq_thm sym_thm trans_thm c_thm = declare_info tyco m p c c_def m_comp [eq_thm, sym_thm, trans_thm] c_thm fun mk_infer_const name ctxt c = infer_type ctxt (Const (name, dummyT) $ c) val mk_eq_comp = mk_infer_const @{const_name eq_comp} val mk_peq_comp = mk_infer_const @{const_name peq_comp} val mk_sym_comp = mk_infer_const @{const_name sym_comp} val mk_psym_comp = mk_infer_const @{const_name psym_comp} val mk_trans_comp = mk_infer_const @{const_name trans_comp} val mk_ptrans_comp = mk_infer_const @{const_name ptrans_comp} val mk_comp = mk_infer_const @{const_name comparator} fun default_comp T = absdummy T (absdummy T @{term Eq}) (*%_ _. Eq*) fun register_foreign_comparator T comp comp_thm lthy = let val tyco = (case T of Type (tyco, []) => tyco | _ => error "expected type constant") val eq = @{thm comp_to_peq_comp} OF [comp_thm] val sym = @{thm comp_to_psym_comp} OF [comp_thm] val trans = @{thm comp_to_ptrans_comp} OF [comp_thm] in register_foreign_partial_and_full_comparator tyco (HOLogic.id_const T) comp comp NONE NONE eq sym trans comp_thm [] lthy end fun register_comparator_of tyco lthy = let val T = Type (tyco, []) val comp = Const (@{const_name comparator_of}, compT T) val comp_thm = Thm.instantiate' [SOME (Thm.ctyp_of lthy T)] [] @{thm comparator_of} in register_foreign_comparator T comp comp_thm lthy end fun generate_comparators_from_bnf_fp tyco lthy = let val (tycos, Ts) = mutual_recursive_types tyco lthy val _ = map (fn tyco => "generating comparator for type " ^ quote tyco) tycos |> cat_lines |> writeln val (tfrees, used_tfrees) = type_parameters (hd Ts) lthy val used_positions = map (member (op =) used_tfrees o TFree) tfrees val cs = map (subT "comp") used_tfrees val comp_Ts = map compT used_tfrees val arg_comps = map Free (cs ~~ comp_Ts) val dep_tycos = fold (add_used_tycos lthy) tycos [] val XTys = Bnf_Access.bnf_types lthy tycos val inst_types = typ_subst_atomic (XTys ~~ Ts) val cTys = map (map (map inst_types)) (Bnf_Access.constr_argument_types lthy tycos) val map_simps = Bnf_Access.map_simps lthy tycos val case_simps = Bnf_Access.case_simps lthy tycos val maps = Bnf_Access.map_terms lthy tycos val map_comp_thms = Bnf_Access.map_comps lthy tycos val t_ixs = 0 upto (length Ts - 1) val compNs = (*TODO: clashes in presence of same type names in different theories*) map (Long_Name.base_name) tycos |> map (fn s => "comparator_" ^ s) fun gen_vars prefix = map (fn (i, pty) => Free (prefix ^ ints_to_subscript [i], pty)) (t_ixs ~~ Ts) (* primrec definitions of partial comparators *) fun mk_pcomp (tyco, T) = ("partial_comparator_" ^ Long_Name.base_name tyco, pcompT T) fun constr_terms lthy = Bnf_Access.constr_terms lthy #> map (apsnd (map freeify_tvars o fst o strip_type) o dest_Const) fun generate_pcomp_eqs lthy (tyco, T) = let val constrs = constr_terms lthy tyco fun comp_arg T x y = let val m = Generator_Aux.create_map default_comp (K o Free o mk_pcomp) () (K false) (#used_positions oo the_info) (#map oo the_info) (K o #pcomp oo the_info) tycos ((K o K) ()) T lthy val p = Generator_Aux.create_partial () (K false) (#used_positions oo the_info) (#map oo the_info) (K o #pcomp oo the_info) tycos ((K o K) ()) T lthy in p $ (m $ x) $ y |> infer_type lthy end fun generate_eq lthy (c_T as (cN, Ts)) = let val arg_Ts' = map orderify Ts val c = Const (cN, arg_Ts' ---> orderify T) val (y, (xs, ys)) = Name.variant "y" (Variable.names_of lthy) |>> Free o rpair T ||> (fn ctxt => Name.invent_names ctxt "x" (arg_Ts' @ Ts) |> map Free) ||> chop (length Ts) val k = find_index (curry (op =) c_T) constrs val cases = constrs |> map_index (fn (i, (_, Ts')) => if i < k then fold_rev absdummy Ts' @{term Gt} else if k < i then fold_rev absdummy Ts' @{term Lt} else @{term comp_lex} $ HOLogic.mk_list orderT (@{map 3} comp_arg Ts xs ys) |> lambdas ys) val lhs = Free (mk_pcomp (tyco, T)) $ list_comb (c, xs) $ y val rhs = list_comb (singleton (Bnf_Access.case_consts lthy) tyco, cases) $ y in HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs)) |> infer_type lthy end in map (generate_eq lthy) constrs end val eqs = map (generate_pcomp_eqs lthy) (tycos ~~ Ts) |> flat val bindings = tycos ~~ Ts |> map mk_pcomp |> map (fn (name, T) => (Binding.name name, SOME T, NoSyn)) val ((pcomps, pcomp_simps), lthy) = lthy |> Local_Theory.begin_nested |> snd |> (BNF_LFP_Rec_Sugar.primrec false [] bindings (map (fn t => ((Binding.empty_atts, t), [], [])) eqs)) |> Local_Theory.end_nested_result (fn phi => fn (pcomps, _, pcomp_simps) => (map (Morphism.term phi) pcomps, map (Morphism.fact phi) pcomp_simps)) (* definitions of comparators via partial comparators and maps *) fun generate_comp_def tyco lthy = let val cs = map (subT "comp") used_tfrees val arg_Ts = map compT used_tfrees val args = map Free (cs ~~ arg_Ts) val (pcomp, m) = AList.lookup (op =) (tycos ~~ (pcomps ~~ maps)) tyco |> the val ts = tfrees |> map TFree |> map (fn T => AList.lookup (op =) (used_tfrees ~~ args) T |> the_default (default_comp T)) val rhs = HOLogic.mk_comp (pcomp, list_comb (m, ts)) |> infer_type lthy val abs_def = lambdas args rhs val name = "comparator_" ^ Long_Name.base_name tyco val ((comp, (_, prethm)), lthy) = Local_Theory.define ((Binding.name name, NoSyn), (Binding.empty_atts, abs_def)) lthy val eq = Logic.mk_equals (list_comb (comp, args), rhs) - val thm = Goal.prove lthy (map (fst o dest_Free) args) [] eq (K (unfold_tac lthy [prethm])) + val thm = + Goal.prove lthy (map (fst o dest_Free) args) [] eq + (fn {context = ctxt, ...} => unfold_tac ctxt [prethm]) in Local_Theory.note ((Binding.name (name ^ "_def"), []), [thm]) lthy |>> the_single o snd |>> `(K comp) end val ((comps, comp_defs), lthy) = lthy |> Local_Theory.begin_nested |> snd |> fold_map generate_comp_def tycos |>> split_list |> Local_Theory.end_nested_result (fn phi => fn (comps, comp_defs) => (map (Morphism.term phi) comps, map (Morphism.thm phi) comp_defs)) (* alternative simp-rules for comparators *) val full_comps = map (list_comb o rpair arg_comps) comps fun generate_comp_simps (tyco, T) lthy = let val constrs = constr_terms lthy tyco fun comp_arg T x y = let fun create_comp (T as TFree _) = AList.lookup (op =) (used_tfrees ~~ arg_comps) T |> the_default (HOLogic.id_const dummyT) | create_comp (Type (tyco, Ts)) = (case AList.lookup (op =) (tycos ~~ comps) tyco of SOME c => list_comb (c, arg_comps) | NONE => let val {comp = c, used_positions = up, ...} = the_info lthy tyco val ts = (up ~~ Ts) |> map_filter (fn (b, T) => if b then SOME (create_comp T) else NONE) in list_comb (c, ts) end) | create_comp T = error ("unexpected schematic variable " ^ quote (Syntax.string_of_typ lthy T)) val comp = create_comp T in comp $ x $ y |> infer_type lthy end fun generate_eq_thm lthy (c_T as (_, Ts)) = let val (xs, ctxt) = Variable.names_of lthy |> fold_map (fn T => Name.variant "x" #>> Free o rpair T) Ts fun mk_const (c, Ts) = Const (c, Ts ---> T) val comp_const = AList.lookup (op =) (tycos ~~ comps) tyco |> the val lhs = list_comb (comp_const, arg_comps) $ list_comb (mk_const c_T, xs) val k = find_index (curry (op =) c_T) constrs fun mk_eq c ys rhs = let val y = list_comb (mk_const c, ys) val eq = HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs $ y, rhs)) in (ys, eq |> infer_type lthy) end val ((ys, eqs), _) = fold_map (fn (i, c as (_, Ts')) => fn ctxt => let val (ys, ctxt) = fold_map (fn T => Name.variant "y" #>> Free o rpair T) Ts' ctxt in (if i < k then mk_eq c ys @{term Gt} else if k < i then mk_eq c ys @{term Lt} else @{term comp_lex} $ HOLogic.mk_list orderT (@{map 3} comp_arg Ts xs ys) |> mk_eq c ys, ctxt) end) (tag_list 0 constrs) ctxt |> apfst (apfst flat o split_list) val dep_comp_defs = map_filter (#comp_def o the_info lthy) dep_tycos val dep_map_comps = map_filter (#map_comp o the_info lthy) dep_tycos val thms = prove_multi_future lthy (map (fst o dest_Free) (xs @ ys) @ cs) [] eqs (fn {context = ctxt, ...} => Goal.conjunction_tac 1 THEN unfold_tac ctxt (@{thms id_apply o_def} @ flat case_simps @ flat pcomp_simps @ dep_map_comps @ comp_defs @ dep_comp_defs @ flat map_simps)) in thms end val thms = map (generate_eq_thm lthy) constrs |> flat val simp_thms = map (Local_Defs.unfold lthy @{thms comp_lex_unfolds}) thms val name = "comparator_" ^ Long_Name.base_name tyco in lthy |> Local_Theory.note ((Binding.name (name ^ "_simps"), @{attributes [simp, code]}), simp_thms) |> snd |> (fn lthy => (thms, lthy)) end val (comp_simps, lthy) = lthy |> Local_Theory.begin_nested |> snd |> fold_map generate_comp_simps (tycos ~~ Ts) |> Local_Theory.end_nested_result (fn phi => map (Morphism.fact phi)) (* partial theorems *) val set_funs = Bnf_Access.set_terms lthy tycos val x_vars = gen_vars "x" val free_names = map (fst o dest_Free) (x_vars @ arg_comps) val xi_vars = map_index (fn (i, _) => map_index (fn (j, pty) => Free ("x" ^ ints_to_subscript [i, j], pty)) used_tfrees) Ts fun mk_eq_sym_trans_thm' mk_eq_sym_trans' = map_index (fn (i, ((set_funs, x), xis)) => let fun create_cond ((set_t, xi), c) = let val rhs = mk_eq_sym_trans' lthy c $ xi |> HOLogic.mk_Trueprop val lhs = HOLogic.mk_mem (xi, set_t $ x) |> HOLogic.mk_Trueprop in Logic.all xi (Logic.mk_implies (lhs, rhs)) end val used_sets = map (the o AList.lookup (op =) (map TFree tfrees ~~ set_funs)) used_tfrees val conds = map create_cond (used_sets ~~ xis ~~ arg_comps) val concl = mk_eq_sym_trans' lthy (nth full_comps i) $ x |> HOLogic.mk_Trueprop in Logic.list_implies (conds, concl) |> infer_type lthy end) (set_funs ~~ x_vars ~~ xi_vars) val induct_thms = Bnf_Access.induct_thms lthy tycos val set_simps = Bnf_Access.set_simps lthy tycos val case_thms = Bnf_Access.case_thms lthy tycos val distinct_thms = Bnf_Access.distinct_thms lthy tycos val inject_thms = Bnf_Access.inject_thms lthy tycos val rec_info = (the_info lthy, #used_positions, tycos) val split_IHs = split_IHs rec_info val unknown_value = false (* effect of choosing false / true not yet visible *) fun induct_tac ctxt f = ((DETERM o Induction.induction_tac ctxt false (map (fn x => [SOME (NONE, (x, unknown_value))]) x_vars) [] [] (SOME induct_thms) []) THEN_ALL_NEW (fn i => Subgoal.SUBPROOF (fn {context = ctxt, prems = prems, params = iparams, ...} => f (i - 1) ctxt prems iparams) ctxt i)) 1 fun recursor_tac kind = std_recursor_tac rec_info used_tfrees (fn info => nth (#partial_comp_thms info) kind) fun instantiate_IHs IHs pre_conds = map (fn IH => OF_option IH (replicate (Thm.nprems_of IH - length pre_conds) NONE @ map SOME pre_conds)) IHs fun get_v_i vs k = nth vs k |> snd |> SOME (* partial eq-theorem *) val _ = debug_out "Partial equality" val eq_thms' = mk_eq_sym_trans_thm' mk_peq_comp fun eq_solve_tac i ctxt IH_prems xs = let val (i, j) = ind_case_to_idxs cTys i val k = length IH_prems - length arg_comps val pre_conds = drop k IH_prems val IH = take k IH_prems val comp_simps = nth comp_simps i val case_thm = nth case_thms i val distinct_thms = nth distinct_thms i val inject_thms = nth inject_thms i val set_thms = nth set_simps i in (* after induction *) resolve_tac ctxt @{thms peq_compI} 1 THEN Subgoal.FOCUS (fn focus => let val y = #params focus |> hd val yt = y |> snd |> Thm.term_of val ctxt = #context focus val pre_cond = map (fn pre_cond => Local_Defs.unfold ctxt set_thms pre_cond) pre_conds val IH = instantiate_IHs IH pre_cond val xs_tys = map (fastype_of o Thm.term_of o snd) xs val IHs = split_IHs xs_tys IH fun sub_case_tac j' (ctxt, y_simp, _) = if j = j' then unfold_tac ctxt (y_simp @ comp_simps) THEN unfold_tac ctxt @{thms comp_lex_eq} THEN unfold_tac ctxt (@{thms in_set_simps} @ inject_thms @ @{thms refl_True}) THEN conjI_tac @{thms conj_weak_cong} ctxt xs (fn ctxt' => fn k => resolve_tac ctxt @{thms peq_compD} 1 THEN recursor_tac EQ pre_cond (nth xs_tys k) (nth IHs k) ctxt') else (* different constructors *) unfold_tac ctxt (y_simp @ distinct_thms @ comp_simps @ @{thms order.simps}) in mk_case_tac ctxt [[SOME yt]] case_thm sub_case_tac end ) ctxt 1 end val eq_thms' = prove_multi_future lthy free_names [] eq_thms' (fn {context = ctxt, ...} => induct_tac ctxt eq_solve_tac) val _ = debug_out (@{make_string} eq_thms') (* partial symmetry-theorem *) val _ = debug_out "Partial symmetry" val sym_thms' = mk_eq_sym_trans_thm' mk_psym_comp fun sym_solve_tac i ctxt IH_prems xs = let val (i, j) = ind_case_to_idxs cTys i val k = length IH_prems - length arg_comps val pre_conds = drop k IH_prems val IH = take k IH_prems val comp_simps = nth comp_simps i val case_thm = nth case_thms i val set_thms = nth set_simps i in (* after induction *) resolve_tac ctxt @{thms psym_compI} 1 THEN Subgoal.FOCUS (fn focus => let val y = #params focus |> hd val yt = y |> snd |> Thm.term_of val ctxt = #context focus val pre_cond = map (fn pre_cond => Local_Defs.unfold ctxt set_thms pre_cond) pre_conds val IH = instantiate_IHs IH pre_cond val xs_tys = map (fastype_of o Thm.term_of o snd) xs val IHs = split_IHs xs_tys IH fun sub_case_tac j' (ctxt, y_simp, ys) = if j = j' then unfold_tac ctxt (y_simp @ comp_simps) THEN resolve_tac ctxt @{thms comp_lex_sym} 1 THEN unfold_tac ctxt (@{thms length_nth_simps forall_finite}) THEN conjI_tac @{thms conjI} ctxt xs (fn ctxt' => fn k => resolve_tac ctxt' [infer_instantiate' ctxt' [NONE, get_v_i xs k, get_v_i ys k] @{thm psym_compD}] 1 THEN recursor_tac SYM pre_cond (nth xs_tys k) (nth IHs k) ctxt') else (* different constructors *) unfold_tac ctxt (y_simp @ comp_simps @ @{thms invert_order.simps}) in mk_case_tac ctxt [[SOME yt]] case_thm sub_case_tac end ) ctxt 1 end val sym_thms' = prove_multi_future lthy free_names [] sym_thms' (fn {context = ctxt, ...} => induct_tac ctxt sym_solve_tac) val _ = debug_out (@{make_string} sym_thms') (* partial transitivity-theorem *) val _ = debug_out "Partial transitivity" val trans_thms' = mk_eq_sym_trans_thm' mk_ptrans_comp fun trans_solve_tac i ctxt IH_prems xs = let val (i, j) = ind_case_to_idxs cTys i val k = length IH_prems - length arg_comps val pre_conds = drop k IH_prems val IH = take k IH_prems val comp_simps = nth comp_simps i val case_thm = nth case_thms i val set_thms = nth set_simps i in (* after induction *) resolve_tac ctxt @{thms ptrans_compI} 1 THEN Subgoal.FOCUS (fn focus => let val y = nth (#params focus) 0 val z = nth (#params focus) 1 val yt = y |> snd |> Thm.term_of val zt = z |> snd |> Thm.term_of val ctxt = #context focus val pre_cond = map (fn pre_cond => Local_Defs.unfold ctxt set_thms pre_cond) pre_conds val IH = instantiate_IHs IH pre_cond val xs_tys = map (fastype_of o Thm.term_of o snd) xs val IHs = split_IHs xs_tys IH fun sub_case_tac j' (ctxt, y_simp, ys) = let fun sub_case_tac' j'' (ctxt, z_simp, zs) = if j = j' andalso j = j'' then unfold_tac ctxt (y_simp @ z_simp @ comp_simps) THEN resolve_tac ctxt @{thms comp_lex_trans} 1 THEN unfold_tac ctxt (@{thms length_nth_simps forall_finite}) THEN conjI_tac @{thms conjI} ctxt xs (fn ctxt' => fn k => resolve_tac ctxt' [infer_instantiate' ctxt' [NONE, get_v_i xs k, get_v_i ys k, get_v_i zs k] @{thm ptrans_compD}] 1 THEN recursor_tac TRANS pre_cond (nth xs_tys k) (nth IHs k) ctxt') else (* different constructors *) unfold_tac ctxt (y_simp @ z_simp @ comp_simps @ @{thms trans_order_different}) in mk_case_tac ctxt [[SOME zt]] case_thm sub_case_tac' end in mk_case_tac ctxt [[SOME yt]] case_thm sub_case_tac end ) ctxt 1 end val trans_thms' = prove_multi_future lthy free_names [] trans_thms' (fn {context = ctxt, ...} => induct_tac ctxt trans_solve_tac) val _ = debug_out (@{make_string} trans_thms') (* total theorems *) fun mk_eq_sym_trans_thm mk_eq_sym_trans compI2 compE2 thms' = let val conds = map (fn c => mk_eq_sym_trans lthy c |> HOLogic.mk_Trueprop) arg_comps val thms = map (fn i => mk_eq_sym_trans lthy (nth full_comps i) |> HOLogic.mk_Trueprop |> (fn concl => Logic.list_implies (conds,concl))) t_ixs val thms = prove_multi_future lthy free_names [] thms (fn {context = ctxt, ...} => ALLGOALS Goal.conjunction_tac THEN Method.intros_tac ctxt (@{thm conjI} :: compI2 :: thms') [] THEN ALLGOALS (eresolve_tac ctxt [compE2])) in thms end val eq_thms = mk_eq_sym_trans_thm mk_eq_comp @{thm eq_compI2} @{thm eq_compD2} eq_thms' val sym_thms = mk_eq_sym_trans_thm mk_sym_comp @{thm sym_compI2} @{thm sym_compD2} sym_thms' val trans_thms = mk_eq_sym_trans_thm mk_trans_comp @{thm trans_compI2} @{thm trans_compD2} trans_thms' val _ = debug_out "full comparator thms" fun mk_comp_thm (i, ((e, s), t)) = let val conds = map (fn c => mk_comp lthy c |> HOLogic.mk_Trueprop) arg_comps fun from_comp thm i = thm OF replicate (Thm.prems_of thm |> length) (nth @{thms comparator_imp_eq_sym_trans} i) val nearly_thm = @{thm eq_sym_trans_imp_comparator} OF [from_comp e EQ, from_comp s SYM, from_comp t TRANS] val thm = mk_comp lthy (nth full_comps i) |> HOLogic.mk_Trueprop |> (fn concl => Logic.list_implies (conds, concl)) in Goal.prove_future lthy free_names [] thm (K (resolve_tac lthy [nearly_thm] 1 THEN ALLGOALS (assume_tac lthy))) end val comp_thms = map_index mk_comp_thm (eq_thms ~~ sym_thms ~~ trans_thms) val (_, lthy) = fold_map (fn (thm, cname) => Local_Theory.note ((Binding.name cname, []), [thm])) (comp_thms ~~ compNs) lthy val _ = debug_out (@{make_string} comp_thms) val pcomp_thms = map (fn ((e, s), t) => [e, s, t]) (eq_thms' ~~ sym_thms' ~~ trans_thms') val (_, lthy) = fold_map (fn (thms, cname) => Local_Theory.note ((Binding.name (cname ^ "_pointwise"), []), thms)) (pcomp_thms ~~ compNs) lthy in ((pcomps ~~ pcomp_simps, comps ~~ comp_defs), lthy) ||> fold (fn (((((((tyco, map), pcomp), comp), comp_def), map_comp), pcomp_thms), comp_thm) => declare_info tyco map pcomp comp (SOME comp_def) (SOME map_comp) pcomp_thms comp_thm used_positions) (tycos ~~ maps ~~ pcomps ~~ comps ~~ comp_defs ~~ map_comp_thms ~~ pcomp_thms ~~ comp_thms) end fun generate_comparator gen_type tyco lthy = let val _ = is_some (get_info lthy tyco) andalso error ("type " ^ quote tyco ^ " does already have a comparator") in case gen_type of BNF => generate_comparators_from_bnf_fp tyco lthy |> snd | Linorder => register_comparator_of tyco lthy end fun ensure_info gen_type tyco lthy = (case get_info lthy tyco of SOME _ => lthy | NONE => generate_comparator gen_type tyco lthy) fun generate_comparator_cmd tyco param = Named_Target.theory_map ( if param = "linorder" then generate_comparator Linorder tyco else if param = "" then generate_comparator BNF tyco else error ("unknown parameter, expecting no parameter for BNF-datatypes, " ^ "or \"linorder\" for types which are already in linorder")) val _ = Theory.setup (Derive_Manager.register_derive "comparator" "generate comparators for given types, options: (linorder) or ()" generate_comparator_cmd) end diff --git a/thys/Deriving/Equality_Generator/equality_generator.ML b/thys/Deriving/Equality_Generator/equality_generator.ML --- a/thys/Deriving/Equality_Generator/equality_generator.ML +++ b/thys/Deriving/Equality_Generator/equality_generator.ML @@ -1,517 +1,519 @@ (* Title: Deriving class instances for datatypes Author: Christian Sternagel and René Thiemann Maintainer: Christian Sternagel and René Thiemann License: LGPL *) signature EQUALITY_GENERATOR = sig type info = {map : term, (* take % x. x, if there is no map *) pequality : term, (* partial equality *) equality : term, (* full equality *) equality_def : thm option, (* definition of equality, important for nesting *) map_comp : thm option, (* compositionality of map, important for nesting *) partial_equality_thm : thm, (* partial version of equality thm *) equality_thm : thm, (* equality acomp \ \ \ equality (full_comp acomp \) *) used_positions : bool list} (* registers @{term equality_of :: "some_type :: linorder equality"} where some_type must just be a type without type-arguments *) val register_equality_of : string -> local_theory -> local_theory val register_foreign_equality : typ -> (* type-constant without type-variables *) term -> (* equality for type *) thm -> (* equality thm for provided equality *) local_theory -> local_theory val register_foreign_partial_and_full_equality : string -> (* long type name *) term -> (* map function, should be \x. x, if there is no map *) term -> (* partial equality of type ('a => order, 'b)ty => ('a,'b)ty => order, where 'a is used, 'b is unused *) term -> (* (full) equality of type ('a \ 'a \ order) \ ('a,'b)ty \ ('a,'b)ty \ order, where 'a is used, 'b is unused *) thm option -> (* comp_def, should be full_comp = pcomp o map acomp ..., important for nesting *) thm option -> (* map compositionality, important for nesting *) thm -> (* partial eq thm for full equality *) thm -> (* full thm: equality a-comp => equality (full_comp a-comp) *) bool list -> (*used positions*) local_theory -> local_theory datatype equality_type = EQ | BNF val generate_equalitys_from_bnf_fp : string -> (* name of type *) local_theory -> ((term * thm list) list * (* partial equalitys + simp-rules *) (term * thm) list) * (* non-partial equality + def_rule *) local_theory val generate_equality : equality_type -> string -> (* name of type *) local_theory -> local_theory val get_info : Proof.context -> string -> info option (* ensures that the info will be available on later requests *) val ensure_info : equality_type -> string -> local_theory -> local_theory end structure Equality_Generator : EQUALITY_GENERATOR = struct open Generator_Aux datatype equality_type = BNF | EQ val debug = false fun debug_out s = if debug then writeln s else () val boolT = @{typ bool} fun compT T = T --> T --> boolT val orderify = map_atyps (fn T => T --> boolT) fun pcompT T = orderify T --> T --> boolT type info = {map : term, pequality : term, equality : term, equality_def : thm option, map_comp : thm option, partial_equality_thm : thm, equality_thm : thm, used_positions : bool list}; structure Data = Generic_Data ( type T = info Symtab.table; val empty = Symtab.empty; val extend = I; val merge = Symtab.merge (fn (info1 : info, info2 : info) => #equality info1 = #equality info2); ); fun add_info T info = Data.map (Symtab.update_new (T, info)) val get_info = Context.Proof #> Data.get #> Symtab.lookup fun the_info ctxt tyco = (case get_info ctxt tyco of SOME info => info | NONE => error ("no equality information available for type " ^ quote tyco)) fun declare_info tyco m p c c_def m_comp p_thm c_thm used_pos = Local_Theory.declaration {syntax = false, pervasive = false} (fn phi => add_info tyco {map = Morphism.term phi m, pequality = Morphism.term phi p, equality = Morphism.term phi c, equality_def = Option.map (Morphism.thm phi) c_def, map_comp = Option.map (Morphism.thm phi) m_comp, partial_equality_thm = Morphism.thm phi p_thm, equality_thm = Morphism.thm phi c_thm, used_positions = used_pos}) fun register_foreign_partial_and_full_equality tyco m p c c_def m_comp eq_thm c_thm = declare_info tyco m p c c_def m_comp eq_thm c_thm val mk_equality = mk_infer_const @{const_name equality} val mk_pequality = mk_infer_const @{const_name pequality} fun default_comp T = absdummy T (absdummy T @{term True}) (*%_ _. True*) fun register_foreign_equality T comp comp_thm lthy = let val tyco = (case T of Type (tyco, []) => tyco | _ => error "expected type constant with no arguments") val eq = @{thm equalityD2} OF [comp_thm] in register_foreign_partial_and_full_equality tyco (HOLogic.id_const T) comp comp NONE NONE eq comp_thm [] lthy end fun register_equality_of tyco lthy = let val (T,_) = typ_and_vs_of_typname (Proof_Context.theory_of lthy) tyco @{sort type} val comp = HOLogic.eq_const T val comp_thm = Thm.instantiate' [SOME (Thm.ctyp_of lthy T)] [] @{thm eq_equality} in register_foreign_equality T comp comp_thm lthy end fun generate_equalitys_from_bnf_fp tyco lthy = let val (tycos, Ts) = mutual_recursive_types tyco lthy val _ = map (fn tyco => "generating equality for type " ^ quote tyco) tycos |> cat_lines |> writeln val (tfrees, used_tfrees) = type_parameters (hd Ts) lthy val used_positions = map (member (op =) used_tfrees o TFree) tfrees val cs = map (subT "eq") used_tfrees val comp_Ts = map compT used_tfrees val arg_comps = map Free (cs ~~ comp_Ts) val dep_tycos = fold (add_used_tycos lthy) tycos [] val XTys = Bnf_Access.bnf_types lthy tycos val inst_types = typ_subst_atomic (XTys ~~ Ts) val cTys = map (map (map inst_types)) (Bnf_Access.constr_argument_types lthy tycos) val map_simps = Bnf_Access.map_simps lthy tycos val case_simps = Bnf_Access.case_simps lthy tycos val maps = Bnf_Access.map_terms lthy tycos val map_comp_thms = Bnf_Access.map_comps lthy tycos val t_ixs = 0 upto (length Ts - 1) val compNs = (*TODO: clashes in presence of same type names in different theories*) map (Long_Name.base_name) tycos |> map (fn s => "equality_" ^ s) fun gen_vars prefix = map (fn (i, pty) => Free (prefix ^ ints_to_subscript [i], pty)) (t_ixs ~~ Ts) (* primrec definitions of partial equalitys *) fun mk_pcomp (tyco, T) = ("partial_equality_" ^ Long_Name.base_name tyco, pcompT T) fun constr_terms lthy = Bnf_Access.constr_terms lthy #> map (apsnd (map freeify_tvars o fst o strip_type) o dest_Const) fun generate_pcomp_eqs lthy (tyco, T) = let val constrs = constr_terms lthy tyco fun comp_arg T x y = let val m = Generator_Aux.create_map default_comp (K o Free o mk_pcomp) () (K false) (#used_positions oo the_info) (#map oo the_info) (K o #pequality oo the_info) tycos ((K o K) ()) T lthy val p = Generator_Aux.create_partial () (K false) (#used_positions oo the_info) (#map oo the_info) (K o #pequality oo the_info) tycos ((K o K) ()) T lthy in p $ (m $ x) $ y |> infer_type lthy end fun generate_eq lthy (c_T as (cN, Ts)) = let val arg_Ts' = map orderify Ts val c = Const (cN, arg_Ts' ---> orderify T) val (y, (xs, ys)) = Name.variant "y" (Variable.names_of lthy) |>> Free o rpair T ||> (fn ctxt => Name.invent_names ctxt "x" (arg_Ts' @ Ts) |> map Free) ||> chop (length Ts) val k = find_index (curry (op =) c_T) constrs val cases = constrs |> map_index (fn (i, (_, Ts')) => if i < k then fold_rev absdummy Ts' @{term False} else if k < i then fold_rev absdummy Ts' @{term False} else @{term list_all_eq} $ HOLogic.mk_list boolT (@{map 3} comp_arg Ts xs ys) |> lambdas ys) val lhs = Free (mk_pcomp (tyco, T)) $ list_comb (c, xs) $ y val rhs = list_comb (singleton (Bnf_Access.case_consts lthy) tyco, cases) $ y in HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs)) |> infer_type lthy end in map (generate_eq lthy) constrs end val eqs = map (generate_pcomp_eqs lthy) (tycos ~~ Ts) |> flat val bindings = tycos ~~ Ts |> map mk_pcomp |> map (fn (name, T) => (Binding.name name, SOME T, NoSyn)) val ((pcomps, pcomp_simps), lthy) = lthy |> Local_Theory.begin_nested |> snd |> (BNF_LFP_Rec_Sugar.primrec false [] bindings (map (fn t => ((Binding.empty_atts, t), [], [])) eqs)) |> Local_Theory.end_nested_result (fn phi => fn (pcomps, _, pcomp_simps) => (map (Morphism.term phi) pcomps, map (Morphism.fact phi) pcomp_simps)) (* definitions of equalitys via partial equalitys and maps *) fun generate_comp_def tyco lthy = let val cs = map (subT "eq") used_tfrees val arg_Ts = map compT used_tfrees val args = map Free (cs ~~ arg_Ts) val (pcomp, m) = AList.lookup (op =) (tycos ~~ (pcomps ~~ maps)) tyco |> the val ts = tfrees |> map TFree |> map (fn T => AList.lookup (op =) (used_tfrees ~~ args) T |> the_default (default_comp T)) val rhs = HOLogic.mk_comp (pcomp, list_comb (m, ts)) |> infer_type lthy val abs_def = lambdas args rhs val name = "equality_" ^ Long_Name.base_name tyco val ((comp, (_, prethm)), lthy) = Local_Theory.define ((Binding.name name, NoSyn), (Binding.empty_atts, abs_def)) lthy val eq = Logic.mk_equals (list_comb (comp, args), rhs) - val thm = Goal.prove lthy (map (fst o dest_Free) args) [] eq (K (unfold_tac lthy [prethm])) + val thm = + Goal.prove lthy (map (fst o dest_Free) args) [] eq + (fn {context = ctxt, ...} => unfold_tac ctxt [prethm]) in Local_Theory.note ((Binding.name (name ^ "_def"), []), [thm]) lthy |>> the_single o snd |>> `(K comp) end val ((comps, comp_defs), lthy) = lthy |> Local_Theory.begin_nested |> snd |> fold_map generate_comp_def tycos |>> split_list |> Local_Theory.end_nested_result (fn phi => fn (comps, comp_defs) => (map (Morphism.term phi) comps, map (Morphism.thm phi) comp_defs)) (* alternative simp-rules for equalitys *) val full_comps = map (list_comb o rpair arg_comps) comps fun generate_comp_simps (tyco, T) lthy = let val constrs = constr_terms lthy tyco fun comp_arg T x y = let fun create_comp (T as TFree _) = AList.lookup (op =) (used_tfrees ~~ arg_comps) T |> the_default (HOLogic.id_const dummyT) | create_comp (Type (tyco, Ts)) = (case AList.lookup (op =) (tycos ~~ comps) tyco of SOME c => list_comb (c, arg_comps) | NONE => let val {equality = c, used_positions = up, ...} = the_info lthy tyco val ts = (up ~~ Ts) |> map_filter (fn (b, T) => if b then SOME (create_comp T) else NONE) in list_comb (c, ts) end) | create_comp T = error ("unexpected schematic variable " ^ quote (Syntax.string_of_typ lthy T)) val comp = create_comp T in comp $ x $ y |> infer_type lthy end fun generate_eq_thm lthy (c_T as (_, Ts)) = let val (xs, ctxt) = Variable.names_of lthy |> fold_map (fn T => Name.variant "x" #>> Free o rpair T) Ts fun mk_const (c, Ts) = Const (c, Ts ---> T) val comp_const = AList.lookup (op =) (tycos ~~ comps) tyco |> the val lhs = list_comb (comp_const, arg_comps) $ list_comb (mk_const c_T, xs) val k = find_index (curry (op =) c_T) constrs fun mk_eq c ys rhs = let val y = list_comb (mk_const c, ys) val eq = HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs $ y, rhs)) in (ys, eq |> infer_type lthy) end val ((ys, eqs), _) = fold_map (fn (i, c as (_, Ts')) => fn ctxt => let val (ys, ctxt) = fold_map (fn T => Name.variant "y" #>> Free o rpair T) Ts' ctxt in (if i < k then mk_eq c ys @{term False} else if k < i then mk_eq c ys @{term False} else @{term list_all_eq} $ HOLogic.mk_list boolT (@{map 3} comp_arg Ts xs ys) |> mk_eq c ys, ctxt) end) (tag_list 0 constrs) ctxt |> apfst (apfst flat o split_list) val dep_comp_defs = map_filter (#equality_def o the_info lthy) dep_tycos val dep_map_comps = map_filter (#map_comp o the_info lthy) dep_tycos val thms = prove_multi_future lthy (map (fst o dest_Free) (xs @ ys) @ cs) [] eqs (fn {context = ctxt, ...} => Goal.conjunction_tac 1 THEN unfold_tac ctxt (@{thms id_apply o_def} @ flat case_simps @ flat pcomp_simps @ dep_map_comps @ comp_defs @ dep_comp_defs @ flat map_simps)) in thms end val thms = map (generate_eq_thm lthy) constrs |> flat val simp_thms = map (Local_Defs.unfold lthy @{thms list_all_eq_unfold}) thms val name = "equality_" ^ Long_Name.base_name tyco in lthy |> Local_Theory.note ((Binding.name (name ^ "_simps"), @{attributes [simp, code]}), simp_thms) |> snd |> (fn lthy => (thms, lthy)) end val (comp_simps, lthy) = lthy |> Local_Theory.begin_nested |> snd |> fold_map generate_comp_simps (tycos ~~ Ts) |> Local_Theory.end_nested_result (fn phi => map (Morphism.fact phi)) (* partial theorems *) val set_funs = Bnf_Access.set_terms lthy tycos val x_vars = gen_vars "x" val free_names = map (fst o dest_Free) (x_vars @ arg_comps) val xi_vars = map_index (fn (i, _) => map_index (fn (j, pty) => Free ("x" ^ ints_to_subscript [i, j], pty)) used_tfrees) Ts fun mk_eq_thm' mk_eq' = map_index (fn (i, ((set_funs, x), xis)) => let fun create_cond ((set_t, xi), c) = let val rhs = mk_eq' lthy c $ xi |> HOLogic.mk_Trueprop val lhs = HOLogic.mk_mem (xi, set_t $ x) |> HOLogic.mk_Trueprop in Logic.all xi (Logic.mk_implies (lhs, rhs)) end val used_sets = map (the o AList.lookup (op =) (map TFree tfrees ~~ set_funs)) used_tfrees val conds = map create_cond (used_sets ~~ xis ~~ arg_comps) val concl = mk_eq' lthy (nth full_comps i) $ x |> HOLogic.mk_Trueprop in Logic.list_implies (conds, concl) |> infer_type lthy end) (set_funs ~~ x_vars ~~ xi_vars) val induct_thms = Bnf_Access.induct_thms lthy tycos val set_simps = Bnf_Access.set_simps lthy tycos val case_thms = Bnf_Access.case_thms lthy tycos val distinct_thms = Bnf_Access.distinct_thms lthy tycos val inject_thms = Bnf_Access.inject_thms lthy tycos val rec_info = (the_info lthy, #used_positions, tycos) val split_IHs = split_IHs rec_info val unknown_value = false (* effect of choosing false / true not yet visible *) fun induct_tac ctxt f = ((DETERM o Induction.induction_tac ctxt false (map (fn x => [SOME (NONE, (x, unknown_value))]) x_vars) [] [] (SOME induct_thms) []) THEN_ALL_NEW (fn i => Subgoal.SUBPROOF (fn {context = ctxt, prems = prems, params = iparams, ...} => f (i - 1) ctxt prems iparams) ctxt i)) 1 val recursor_tac = std_recursor_tac rec_info used_tfrees (fn info => #partial_equality_thm info) fun instantiate_IHs IHs pre_conds = map (fn IH => OF_option IH (replicate (Thm.nprems_of IH - length pre_conds) NONE @ map SOME pre_conds)) IHs (* partial eq-theorem *) val _ = debug_out "Partial equality" val eq_thms' = mk_eq_thm' mk_pequality fun eq_solve_tac i ctxt IH_prems xs = let val (i, j) = ind_case_to_idxs cTys i val k = length IH_prems - length arg_comps val pre_conds = drop k IH_prems val IH = take k IH_prems val comp_simps = nth comp_simps i val case_thm = nth case_thms i val distinct_thms = nth distinct_thms i val inject_thms = nth inject_thms i val set_thms = nth set_simps i in (* after induction *) resolve_tac ctxt @{thms pequalityI} 1 THEN Subgoal.FOCUS (fn focus => let val y = #params focus |> hd val yt = y |> snd |> Thm.term_of val ctxt = #context focus val pre_cond = map (fn pre_cond => Local_Defs.unfold ctxt set_thms pre_cond) pre_conds val IH = instantiate_IHs IH pre_cond val xs_tys = map (fastype_of o Thm.term_of o snd) xs val IHs = split_IHs xs_tys IH fun sub_case_tac j' (ctxt, y_simp, _) = if j = j' then unfold_tac ctxt (y_simp @ comp_simps) THEN unfold_tac ctxt @{thms list_all_eq} THEN unfold_tac ctxt (@{thms in_set_simps} @ inject_thms @ @{thms refl_True}) THEN conjI_tac @{thms conj_weak_cong} ctxt xs (fn ctxt' => fn k => resolve_tac ctxt @{thms pequalityD} 1 THEN recursor_tac pre_cond (nth xs_tys k) (nth IHs k) ctxt') else (* different constructors *) unfold_tac ctxt (y_simp @ distinct_thms @ comp_simps @ @{thms bool.simps}) in mk_case_tac ctxt [[SOME yt]] case_thm sub_case_tac end ) ctxt 1 end val eq_thms' = prove_multi_future lthy free_names [] eq_thms' (fn {context = ctxt, ...} => induct_tac ctxt eq_solve_tac) val _ = debug_out (@{make_string} eq_thms') (* total theorems *) fun mk_eq_sym_trans_thm mk_eq_sym_trans compI2 compE2 thms' = let val conds = map (fn c => mk_eq_sym_trans lthy c |> HOLogic.mk_Trueprop) arg_comps val thms = map (fn i => mk_eq_sym_trans lthy (nth full_comps i) |> HOLogic.mk_Trueprop |> (fn concl => Logic.list_implies (conds,concl))) t_ixs val thms = prove_multi_future lthy free_names [] thms (fn {context = ctxt, ...} => ALLGOALS Goal.conjunction_tac THEN Method.intros_tac ctxt (@{thm conjI} :: compI2 :: thms') [] THEN ALLGOALS (eresolve_tac ctxt [compE2])) in thms end val eq_thms = mk_eq_sym_trans_thm mk_equality @{thm equalityI2} @{thm equalityD2} eq_thms' val _ = debug_out "full equality thms" fun mk_comp_thm (i, e) = let val conds = map (fn c => mk_equality lthy c |> HOLogic.mk_Trueprop) arg_comps val nearly_thm = e val thm = mk_equality lthy (nth full_comps i) |> HOLogic.mk_Trueprop |> (fn concl => Logic.list_implies (conds, concl)) in Goal.prove_future lthy free_names [] thm (K (resolve_tac lthy [nearly_thm] 1 THEN ALLGOALS (assume_tac lthy))) end val comp_thms = map_index mk_comp_thm eq_thms val (_, lthy) = fold_map (fn (thm, cname) => Local_Theory.note ((Binding.name cname, []), [thm])) (comp_thms ~~ compNs) lthy val _ = debug_out (@{make_string} comp_thms) val (_, lthy) = fold_map (fn (thm, cname) => Local_Theory.note ((Binding.name (cname ^ "_pointwise"), []), [thm])) (eq_thms' ~~ compNs) lthy in ((pcomps ~~ pcomp_simps, comps ~~ comp_defs), lthy) ||> fold (fn (((((((tyco, map), pcomp), comp), comp_def), map_comp) , peq_thm), comp_thm) => declare_info tyco map pcomp comp (SOME comp_def) (SOME map_comp) peq_thm comp_thm used_positions) (tycos ~~ maps ~~ pcomps ~~ comps ~~ comp_defs ~~ map_comp_thms ~~ eq_thms' ~~ comp_thms) end fun generate_equality gen_type tyco lthy = let val _ = is_some (get_info lthy tyco) andalso error ("type " ^ quote tyco ^ " does already have a equality") in case gen_type of BNF => generate_equalitys_from_bnf_fp tyco lthy |> snd | EQ => register_equality_of tyco lthy end fun ensure_info gen_type tyco lthy = (case get_info lthy tyco of SOME _ => lthy | NONE => generate_equality gen_type tyco lthy) fun generate_equality_cmd tyco param = Named_Target.theory_map ( if param = "eq" then generate_equality EQ tyco else if param = "" then generate_equality BNF tyco else error ("unknown parameter, expecting no parameter for BNF-datatypes, " ^ "or \"eq\" for types where the built-in equality \"=\" should be used.")) val _ = Theory.setup (Derive_Manager.register_derive "equality" "generate an equality function, options are () and (eq)" generate_equality_cmd) end diff --git a/thys/Deriving/Hash_Generator/hash_generator.ML b/thys/Deriving/Hash_Generator/hash_generator.ML --- a/thys/Deriving/Hash_Generator/hash_generator.ML +++ b/thys/Deriving/Hash_Generator/hash_generator.ML @@ -1,411 +1,413 @@ (* Title: Deriving class instances for datatypes Author: Christian Sternagel and René Thiemann Maintainer: Christian Sternagel and René Thiemann License: LGPL *) signature HASHCODE_GENERATOR = sig type info = {map : term, (* take % x. x, if there is no map *) phash : term, (* partial hash *) hash : term, (* full hash *) hash_def : thm option, (* definition of hash, important for nesting *) map_comp : thm option, (* hashositionality of map, important for nesting *) used_positions : bool list} (* registers some type which is already instance of hashcode class in hash generator where some type must just be a type without type-arguments *) val register_hash_of : string -> local_theory -> local_theory val register_foreign_hash : typ -> (* type-constant without type-variables *) term -> (* hash-function for type *) local_theory -> local_theory val register_foreign_partial_and_full_hash : string -> (* long type name *) term -> (* map function, should be \x. x, if there is no map *) term -> (* partial hash-function of type (hashcode, 'b)ty => hashcode, where 'a is used, 'b is unused *) term -> (* (full) hash-function of type ('a \ hashcode) \ ('a,'b)ty \ hashcode, where 'a is used, 'b is unused *) thm option -> (* hash_def, should be full_hash = phash o map ahash ..., important for nesting *) thm option -> (* map compositionality, important for nesting *) bool list -> (*used positions*) local_theory -> local_theory datatype hash_type = HASHCODE | BNF val generate_hashs_from_bnf_fp : string -> (* name of type *) local_theory -> ((term * thm list) list * (* partial hashs + simp-rules *) (term * thm) list) * (* non-partial hash + def_rule *) local_theory val generate_hash : hash_type -> string -> (* name of type *) local_theory -> local_theory (* construct hashcode instance for datatype *) val hashable_instance : string -> theory -> theory val get_info : Proof.context -> string -> info option (* ensures that the info will be available on later requests *) val ensure_info : hash_type -> string -> local_theory -> local_theory end structure Hashcode_Generator : HASHCODE_GENERATOR = struct open Generator_Aux datatype hash_type = BNF | HASHCODE val hash_name = @{const_name "hashcode"} val hashS = @{sort hashable} val hashT = @{typ hashcode} fun hashfunT T = T --> hashT val hashify = map_atyps (fn _ => hashT) fun phashfunT T = hashify T --> hashT val max_int = 2147483648 (* 2 ^^ 31 *) fun int_of_string s = fold (fn c => fn i => (1792318057 * i + Char.ord c) mod max_int) (String.explode s) 0 (* all numbers in int_of_string and create_factors are primes (31-bit) *) fun create_factor ty_name con_name i = (1444315237 * int_of_string ty_name + 1336760419 * int_of_string con_name + 2044890737 * (i + 1) ) mod max_int fun create_hashes ty_name con_name Ts = map (fn i => HOLogic.mk_number hashT (create_factor ty_name con_name i)) (0 upto length Ts) |> HOLogic.mk_list hashT fun create_def_size _ = 10 type info = {map : term, phash : term, hash : term, hash_def : thm option, map_comp : thm option, used_positions : bool list}; structure Data = Generic_Data ( type T = info Symtab.table; val empty = Symtab.empty; val extend = I; val merge = Symtab.merge (fn (info1 : info, info2 : info) => #hash info1 = #hash info2); ); fun add_info T info = Data.map (Symtab.update_new (T, info)) val get_info = Context.Proof #> Data.get #> Symtab.lookup fun the_info ctxt tyco = (case get_info ctxt tyco of SOME info => info | NONE => error ("no hash_code information available for type " ^ quote tyco)) fun declare_info tyco m p c c_def m_hash used_pos = Local_Theory.declaration {syntax = false, pervasive = false} (fn phi => add_info tyco {map = Morphism.term phi m, phash = Morphism.term phi p, hash = Morphism.term phi c, hash_def = Option.map (Morphism.thm phi) c_def, map_comp = Option.map (Morphism.thm phi) m_hash, used_positions = used_pos}) fun register_foreign_partial_and_full_hash tyco m p c c_def m_hash eq_thm c_thm = declare_info tyco m p c c_def m_hash eq_thm c_thm fun default_hash T = absdummy T @{term "0 :: hashcode"} (*%_. 0*) fun register_foreign_hash T hash lthy = let val tyco = (case T of Type (tyco, []) => tyco | _ => error "expected type constant with no arguments") in register_foreign_partial_and_full_hash tyco (HOLogic.id_const T) hash hash NONE NONE [] lthy end fun register_hash_of tyco lthy = let val _ = is_class_instance (Proof_Context.theory_of lthy) tyco hashS orelse error ("type " ^ quote tyco ^ " is not an instance of class \"hashable\"") val (T,_) = typ_and_vs_of_typname (Proof_Context.theory_of lthy) tyco @{sort type} val hash = Const (hash_name, hashfunT T) in register_foreign_hash T hash lthy end fun generate_hashs_from_bnf_fp tyco lthy = let val (tycos, Ts) = mutual_recursive_types tyco lthy val _ = map (fn tyco => "generating hash-function for type " ^ quote tyco) tycos |> cat_lines |> writeln val (tfrees, used_tfrees) = type_parameters (hd Ts) lthy val used_positions = map (member (op =) used_tfrees o TFree) tfrees val cs = map (subT "h") used_tfrees val hash_Ts = map hashfunT used_tfrees val arg_hashs = map Free (cs ~~ hash_Ts) val dep_tycos = fold (add_used_tycos lthy) tycos [] val map_simps = Bnf_Access.map_simps lthy tycos val case_simps = Bnf_Access.case_simps lthy tycos val maps = Bnf_Access.map_terms lthy tycos val map_comp_thms = Bnf_Access.map_comps lthy tycos (* primrec definitions of partial hashs *) fun mk_phash (tyco, T) = ("partial_hash_code_" ^ Long_Name.base_name tyco, phashfunT T) fun constr_terms lthy = Bnf_Access.constr_terms lthy #> map (apsnd (map freeify_tvars o fst o strip_type) o dest_Const) fun generate_phash_eqs lthy (tyco, T) = let val constrs = constr_terms lthy tyco fun hash_arg T x = let val m = Generator_Aux.create_map default_hash (K o Free o mk_phash) () (K false) (#used_positions oo the_info) (#map oo the_info) (K o #phash oo the_info) tycos ((K o K) ()) T lthy val p = Generator_Aux.create_partial () (K false) (#used_positions oo the_info) (#map oo the_info) (K o #phash oo the_info) tycos ((K o K) ()) T lthy in p $ (m $ x) |> infer_type lthy end fun generate_eq lthy (cN, Ts) = let val arg_Ts' = map hashify Ts val c = Const (cN, arg_Ts' ---> hashify T) val xs = Name.invent_names (Variable.names_of lthy) "x" (arg_Ts') |> map Free val lhs = Free (mk_phash (tyco, T)) $ list_comb (c, xs) val rhs = @{term hash_combine} $ HOLogic.mk_list hashT (@{map 2} hash_arg Ts xs) $ create_hashes tyco cN Ts in HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs)) |> infer_type lthy end in map (generate_eq lthy) constrs end val eqs = map (generate_phash_eqs lthy) (tycos ~~ Ts) |> flat val bindings = tycos ~~ Ts |> map mk_phash |> map (fn (name, T) => (Binding.name name, SOME T, NoSyn)) val ((phashs, phash_simps), lthy) = lthy |> Local_Theory.begin_nested |> snd |> BNF_LFP_Rec_Sugar.primrec false [] bindings (map (fn t => ((Binding.empty_atts, t), [], [])) eqs) |> Local_Theory.end_nested_result (fn phi => fn (phashs, _, phash_simps) => (map (Morphism.term phi) phashs, map (Morphism.fact phi) phash_simps)) (* definitions of hashs via partial hashs and maps *) fun generate_hash_def tyco lthy = let val cs = map (subT "h") used_tfrees val arg_Ts = map hashfunT used_tfrees val args = map Free (cs ~~ arg_Ts) val (phash, m) = AList.lookup (op =) (tycos ~~ (phashs ~~ maps)) tyco |> the val ts = tfrees |> map TFree |> map (fn T => AList.lookup (op =) (used_tfrees ~~ args) T |> the_default (default_hash T)) val rhs = HOLogic.mk_comp (phash, list_comb (m, ts)) |> infer_type lthy val abs_def = lambdas args rhs val name = "hash_code_" ^ Long_Name.base_name tyco val ((hash, (_, prethm)), lthy) = Local_Theory.define ((Binding.name name, NoSyn), (Binding.empty_atts, abs_def)) lthy val eq = Logic.mk_equals (list_comb (hash, args), rhs) - val thm = Goal.prove lthy (map (fst o dest_Free) args) [] eq (K (unfold_tac lthy [prethm])) + val thm = + Goal.prove lthy (map (fst o dest_Free) args) [] eq + (fn {context = ctxt, ...} => unfold_tac ctxt [prethm]) in Local_Theory.note ((Binding.name (name ^ "_def"), []), [thm]) lthy |>> the_single o snd |>> `(K hash) end val ((hashs, hash_defs), lthy) = lthy |> Local_Theory.begin_nested |> snd |> fold_map generate_hash_def tycos |>> split_list |> Local_Theory.end_nested_result (fn phi => fn (hashs, hash_defs) => (map (Morphism.term phi) hashs, map (Morphism.thm phi) hash_defs)) (* alternative simp-rules for hashs *) fun generate_hash_simps (tyco, T) lthy = let val constrs = constr_terms lthy tyco fun hash_arg T x = let fun create_hash (T as TFree _) = AList.lookup (op =) (used_tfrees ~~ arg_hashs) T |> the_default (HOLogic.id_const dummyT) | create_hash (Type (tyco, Ts)) = (case AList.lookup (op =) (tycos ~~ hashs) tyco of SOME c => list_comb (c, arg_hashs) | NONE => let val {hash = c, used_positions = up, ...} = the_info lthy tyco val ts = (up ~~ Ts) |> map_filter (fn (b, T) => if b then SOME (create_hash T) else NONE) in list_comb (c, ts) end) | create_hash T = error ("unexpected schematic variable " ^ quote (Syntax.string_of_typ lthy T)) val hash = create_hash T in hash $ x |> infer_type lthy end fun generate_eq_thm lthy (c_T as (cN, Ts)) = let val xs = Variable.names_of lthy |> fold_map (fn T => Name.variant "x" #>> Free o rpair T) Ts |> fst fun mk_const (c, Ts) = Const (c, Ts ---> T) val hash_const = AList.lookup (op =) (tycos ~~ hashs) tyco |> the val lhs = list_comb (hash_const, arg_hashs) $ list_comb (mk_const c_T, xs) val rhs = @{term hash_combine} $ HOLogic.mk_list hashT (@{map 2} hash_arg Ts xs) $ create_hashes tyco cN Ts val eq = HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs)) |> infer_type lthy val dep_hash_defs = map_filter (#hash_def o the_info lthy) dep_tycos val dep_map_comps = map_filter (#map_comp o the_info lthy) dep_tycos val thms = prove_multi_future lthy (map (fst o dest_Free) xs @ cs) [] [eq] (fn {context = ctxt, ...} => Goal.conjunction_tac 1 THEN unfold_tac ctxt (@{thms id_apply o_def} @ flat case_simps @ flat phash_simps @ dep_map_comps @ hash_defs @ dep_hash_defs @ flat map_simps)) in thms end val thms = map (generate_eq_thm lthy) constrs |> flat val simp_thms = map (Local_Defs.unfold lthy @{thms hash_combine_unfold}) thms val name = "hash_code_" ^ Long_Name.base_name tyco in lthy |> Local_Theory.note ((Binding.name (name ^ "_simps"), @{attributes [simp, code]}), simp_thms) |> snd |> (fn lthy => (thms, lthy)) end val lthy = lthy |> Local_Theory.begin_nested |> snd |> fold_map generate_hash_simps (tycos ~~ Ts) |> snd |> Local_Theory.end_nested in ((phashs ~~ phash_simps, hashs ~~ hash_defs), lthy) ||> fold (fn (((((tyco, map), phash), hash), hash_def), map_comp) => declare_info tyco map phash hash (SOME hash_def) (SOME map_comp) used_positions) (tycos ~~ maps ~~ phashs ~~ hashs ~~ hash_defs ~~ map_comp_thms) end fun generate_hash gen_type tyco lthy = let val _ = is_some (get_info lthy tyco) andalso error ("type " ^ quote tyco ^ " does already have a hash") in case gen_type of BNF => generate_hashs_from_bnf_fp tyco lthy |> snd | HASHCODE => register_hash_of tyco lthy end fun ensure_info gen_type tyco lthy = (case get_info lthy tyco of SOME _ => lthy | NONE => generate_hash gen_type tyco lthy) fun dest_hash ctxt tname = (case get_info ctxt tname of SOME {hash = c, ...} => let val Ts = fastype_of c |> strip_type |> fst |> `((fn x => x - 1) o length) |> uncurry take in (c, Ts) end | NONE => error ("no hash info for type " ^ quote tname)) fun all_tys hash free_types = let val Ts = fastype_of hash |> strip_type |> fst |> List.last |> dest_Type |> snd in rename_types (Ts ~~ free_types) end fun mk_hash_rhs c Ts = list_comb (c, map (fn T => Const (hash_name, T)) Ts) fun mk_hash_def T rhs = Logic.mk_equals (Const (hash_name, hashfunT T), rhs) fun hashable_instance tname thy = let val _ = is_class_instance thy tname hashS andalso error ("type " ^ quote tname ^ " is already an instance of class \"hashcode\"") val _ = writeln ("deriving \"hashable\" instance for type " ^ quote tname) val thy = Named_Target.theory_map (ensure_info BNF tname) thy val {used_positions = us, ...} = the (get_info (Named_Target.theory_init thy) tname) val (_, xs) = typ_and_vs_of_used_typname tname us hashS val (_, (hashs_thm,lthy)) = Class.instantiation ([tname], xs, hashS) thy |> (fn ctxt => let val (c, Ts) = dest_hash ctxt tname val typ_mapping = all_tys c (map TFree xs) val hash_rhs = mk_hash_rhs c Ts val hash_def = mk_hash_def dummyT hash_rhs |> typ_mapping |> infer_type ctxt val ty = Term.fastype_of (snd (Logic.dest_equals hash_def)) |> Term.dest_Type |> snd |> hd val ty_it = Type (@{type_name itself}, [ty]) val hashs_rhs = lambda (Free ("x",ty_it)) (HOLogic.mk_number @{typ nat} (create_def_size tname)) val hashs_def = mk_def (ty_it --> @{typ nat}) @{const_name def_hashmap_size} hashs_rhs val basename = Long_Name.base_name tname in Generator_Aux.define_overloaded_generic ((Binding.name ("hashcode_" ^ basename ^ "_def"), @{attributes [code]}), hash_def) ctxt ||> define_overloaded ("def_hashmap_size_" ^ basename ^ "_def", hashs_def) end) in Class.prove_instantiation_exit (fn ctxt => Class.intro_classes_tac ctxt [] THEN unfold_tac ctxt [hashs_thm] THEN simp_tac ctxt 1 ) lthy end fun generate_hash_cmd tyco param = Named_Target.theory_map ( if param = "hashcode" then generate_hash HASHCODE tyco else if param = "" then generate_hash BNF tyco else error ("unknown parameter, expecting no parameter for BNF-datatypes, " ^ "or \"hashcode\" for types where the class-instance hashcode should be used.")) val _ = Theory.setup (Derive_Manager.register_derive "hash_code" "generate a hash function, options are () and (hashcode)" generate_hash_cmd #> Derive_Manager.register_derive "hashable" "register types in class hashable" (fn tyname => K (hashable_instance tyname))) end diff --git a/thys/Forcing/Synthetic_Definition.thy b/thys/Forcing/Synthetic_Definition.thy --- a/thys/Forcing/Synthetic_Definition.thy +++ b/thys/Forcing/Synthetic_Definition.thy @@ -1,128 +1,129 @@ section\Automatic synthesis of formulas\ theory Synthetic_Definition imports Utils keywords "synthesize" :: thy_decl % "ML" and "synthesize_notc" :: thy_decl % "ML" and "from_schematic" begin ML\ val $` = curry ((op $) o swap) infix $` fun pair f g x = (f x, g x) fun display kind pos (thms,thy) = let val _ = Proof_Display.print_results true pos thy ((kind,""),[thms]) in thy end fun prove_tc_form goal thms ctxt = Goal.prove ctxt [] [] goal - (fn _ => rewrite_goal_tac ctxt thms 1 - THEN TypeCheck.typecheck_tac ctxt) + (fn {context = ctxt', ...} => + rewrite_goal_tac ctxt' thms 1 + THEN TypeCheck.typecheck_tac ctxt') fun prove_sats goal thms thm_auto ctxt = - let val ctxt' = ctxt |> Simplifier.add_simp (thm_auto |> hd) - in Goal.prove ctxt [] [] goal - (fn _ => rewrite_goal_tac ctxt thms 1 - THEN PARALLEL_ALLGOALS (asm_simp_tac ctxt') - THEN TypeCheck.typecheck_tac ctxt') - end + (fn {context = ctxt', ...} => + let val ctxt'' = ctxt' |> Simplifier.add_simp (thm_auto |> hd) in + rewrite_goal_tac ctxt'' thms 1 + THEN PARALLEL_ALLGOALS (asm_simp_tac ctxt'') + THEN TypeCheck.typecheck_tac ctxt'' + end) fun is_mem \<^Const_>\mem for _ _\ = true | is_mem _ = false fun synth_thm_sats def_name term lhs set env hyps vars vs pos thm_auto lthy = let val (_,tm,ctxt1) = Utils.thm_concl_tm lthy term val (thm_refs,ctxt2) = Variable.import true [Proof_Context.get_thm lthy term] ctxt1 |>> #2 val vs' = map (Thm.term_of o #2) vs val vars' = map (Thm.term_of o #2) vars val r_tm = tm |> Utils.dest_lhs_def |> fold (op $`) vs' val sats = \<^Const>\apply for \<^Const>\satisfies for set r_tm\ env\ val rhs = \<^Const>\IFOL.eq \<^Type>\i\ for sats \<^Const>\succ for \<^Const>\zero\\\ val concl = \<^Const>\iff for lhs rhs\ val g_iff = Logic.list_implies(hyps, Utils.tp concl) val thm = prove_sats g_iff thm_refs thm_auto ctxt2 val name = Binding.name (def_name ^ "_iff_sats") val thm = Utils.fix_vars thm (map (#1 o dest_Free) vars') lthy in Local_Theory.note ((name, []), [thm]) lthy |> display "theorem" pos end fun synth_thm_tc def_name term hyps vars pos lthy = let val (_,tm,ctxt1) = Utils.thm_concl_tm lthy term val (thm_refs,ctxt2) = Variable.import true [Proof_Context.get_thm lthy term] ctxt1 |>> #2 val vars' = map (Thm.term_of o #2) vars val tc_attrib = @{attributes [TC]} val r_tm = tm |> Utils.dest_lhs_def |> fold (op $`) vars' val concl = \<^Const>\mem for r_tm \<^Const>\formula\\ val g = Logic.list_implies(hyps, Utils.tp concl) val thm = prove_tc_form g thm_refs ctxt2 val name = Binding.name (def_name ^ "_type") val thm = Utils.fix_vars thm (map (#1 o dest_Free) vars') ctxt2 in Local_Theory.note ((name, tc_attrib), [thm]) lthy |> display "theorem" pos end fun synthetic_def def_name thmref pos tc auto thy = let val (thm_ref,_) = thmref |>> Facts.ref_name val thm = Proof_Context.get_thm thy thm_ref; val thm_vars = rev (Term.add_vars (Thm.full_prop_of thm) []); val (((_,inst),thm_tms),_) = Variable.import true [thm] thy val vars = map (fn v => (v, the (Vars.lookup inst v))) thm_vars; val (tm,hyps) = thm_tms |> hd |> pair Thm.concl_of Thm.prems_of val (lhs,rhs) = tm |> Utils.dest_iff_tms o Utils.dest_trueprop val ((set,t),env) = rhs |> Utils.dest_sats_frm fun relevant ts \<^Const_>\mem for t _\ = not (Term.is_Free t) orelse member (op =) ts (t |> Term.dest_Free |> #1) | relevant _ _ = false val t_vars = sort_strings (Term.add_free_names t []) val vs = filter (member (op =) t_vars o #1 o #1 o #1) vars val at = fold_rev (lambda o Thm.term_of o #2) vs t val hyps' = filter (relevant t_vars o Utils.dest_trueprop) hyps in Local_Theory.define ((Binding.name def_name, NoSyn), ((Binding.name (def_name ^ "_def"), []), at)) thy |> #2 |> (if tc then synth_thm_tc def_name (def_name ^ "_def") hyps' vs pos else I) |> (if auto then synth_thm_sats def_name (def_name ^ "_def") lhs set env hyps vars vs pos thm_tms else I) end \ ML\ local val synth_constdecl = Parse.position (Parse.string -- ((Parse.$$$ "from_schematic" |-- Parse.thm))); val _ = Outer_Syntax.local_theory \<^command_keyword>\synthesize\ "ML setup for synthetic definitions" (synth_constdecl >> (fn ((bndg,thm),p) => synthetic_def bndg thm p true true)) val _ = Outer_Syntax.local_theory \<^command_keyword>\synthesize_notc\ "ML setup for synthetic definitions" (synth_constdecl >> (fn ((bndg,thm),p) => synthetic_def bndg thm p false false)) in end \ text\The \<^ML>\synthetic_def\ function extracts definitions from schematic goals. A new definition is added to the context. \ (* example of use *) (* schematic_goal mem_formula_ex : assumes "m\nat" "n\ nat" "env \ list(M)" shows "nth(m,env) \ nth(n,env) \ sats(M,?frm,env)" by (insert assms ; (rule sep_rules empty_iff_sats cartprod_iff_sats | simp del:sats_cartprod_fm)+) synthesize "\" from_schematic mem_formula_ex *) end diff --git a/thys/Forcing/renaming.ML b/thys/Forcing/renaming.ML --- a/thys/Forcing/renaming.ML +++ b/thys/Forcing/renaming.ML @@ -1,178 +1,184 @@ (* Builds the finite mapping. *) structure Renaming = struct open Utils fun sum_ f g m n p = \<^Const>\Renaming.sum for f g m n p\ fun mk_ren rho rho' ctxt = let val rs = to_ML_list rho val rs' = to_ML_list rho' val ixs = 0 upto (length rs-1) fun err t = "The element " ^ Syntax.string_of_term ctxt t ^ " is missing in the target environment" fun mkp i = case find_index (fn x => x = nth rs i) rs' of ~1 => nth rs i |> err |> error | j => mk_Pair (mk_ZFnat i) (mk_ZFnat j) in map mkp ixs |> mk_FinSet end fun mk_dom_lemma ren rho = let val n = rho |> to_ML_list |> length |> mk_ZFnat in eq_ n \<^Const>\domain for ren\ |> tp end fun ren_tc_goal fin ren rho rho' = let val n = rho |> to_ML_list |> length val m = rho' |> to_ML_list |> length val fun_ty = if fin then @{const_name "FiniteFun"} else @{const_abbrev "function_space"} val ty = Const (fun_ty,@{typ "i \ i \ i"}) $ mk_ZFnat n $ mk_ZFnat m in mem_ ren ty |> tp end fun ren_action_goal ren rho rho' ctxt = let val setV = Variable.variant_frees ctxt [] [("A",@{typ i})] |> hd |> Free val j = Variable.variant_frees ctxt [] [("j",@{typ i})] |> hd |> Free val vs = rho |> to_ML_list val ws = rho' |> to_ML_list |> filter isFree val h1 = subset_ (vs|> mk_FinSet) setV val h2 = lt_ j (mk_ZFnat (length vs)) val fvs = ([j,setV ] @ ws |> filter isFree) |> map freeName val lhs = nth_ j rho val rhs = nth_ (app_ ren j) rho' val concl = eq_ lhs rhs in (Logic.list_implies([tp h1,tp h2],tp concl),fvs) end fun sum_tc_goal f m n p = let val m_length = m |> to_ML_list |> length |> mk_ZFnat val n_length = n |> to_ML_list |> length |> mk_ZFnat val p_length = p |> length_ val id_fun = \<^Const>\id for p_length\ val sum_fun = sum_ f id_fun m_length n_length p_length val dom = add_ m_length p_length val codom = add_ n_length p_length val fun_ty = @{const_abbrev "function_space"} val ty = Const (fun_ty,@{typ "i \ i \ i"}) $ dom $ codom in (sum_fun, mem_ sum_fun ty |> tp) end fun sum_action_goal ren rho rho' ctxt = let val setV = Variable.variant_frees ctxt [] [("A",@{typ i})] |> hd |> Free val envV = Variable.variant_frees ctxt [] [("env",@{typ i})] |> hd |> Free val j = Variable.variant_frees ctxt [] [("j",@{typ i})] |> hd |> Free val vs = rho |> to_ML_list val ws = rho' |> to_ML_list |> filter isFree val envL = envV |> length_ val rhoL = vs |> length |> mk_ZFnat val h1 = subset_ (append vs ws |> mk_FinSet) setV val h2 = lt_ j (add_ rhoL envL) val h3 = mem_ envV (list_ setV) val fvs = ([j,setV,envV] @ ws |> filter isFree) |> map freeName val lhs = nth_ j (concat_ rho envV) val rhs = nth_ (app_ ren j) (concat_ rho' envV) val concl = eq_ lhs rhs in (Logic.list_implies([tp h1,tp h2,tp h3],tp concl),fvs) end (* Tactics *) fun fin ctxt = REPEAT (resolve_tac ctxt [@{thm nat_succI}] 1) THEN resolve_tac ctxt [@{thm nat_0I}] 1 fun step ctxt thm = asm_full_simp_tac ctxt 1 THEN asm_full_simp_tac ctxt 1 THEN EqSubst.eqsubst_tac ctxt [1] [@{thm app_fun} OF [thm]] 1 THEN simp_tac ctxt 1 THEN simp_tac ctxt 1 fun fin_fun_tac ctxt = REPEAT ( resolve_tac ctxt [@{thm consI}] 1 THEN resolve_tac ctxt [@{thm ltD}] 1 THEN simp_tac ctxt 1 THEN resolve_tac ctxt [@{thm ltD}] 1 THEN simp_tac ctxt 1) THEN resolve_tac ctxt [@{thm emptyI}] 1 THEN REPEAT (simp_tac ctxt 1) fun ren_thm e e' ctxt = let val r = mk_ren e e' ctxt val fin_tc_goal = ren_tc_goal true r e e' val dom_goal = mk_dom_lemma r e val tc_goal = ren_tc_goal false r e e' val (action_goal,fvs) = ren_action_goal r e e' ctxt - val fin_tc_lemma = Goal.prove ctxt [] [] fin_tc_goal (fn _ => fin_fun_tac ctxt) - val dom_lemma = Goal.prove ctxt [] [] dom_goal (fn _ => blast_tac ctxt 1) - val tc_lemma = Goal.prove ctxt [] [] tc_goal - (fn _ => EqSubst.eqsubst_tac ctxt [1] [dom_lemma] 1 - THEN resolve_tac ctxt [@{thm FiniteFun_is_fun}] 1 - THEN resolve_tac ctxt [fin_tc_lemma] 1) - val action_lemma = Goal.prove ctxt [] [] action_goal - (fn _ => - forward_tac ctxt [@{thm le_natI}] 1 - THEN fin ctxt - THEN REPEAT (resolve_tac ctxt [@{thm natE}] 1 - THEN step ctxt tc_lemma) - THEN (step ctxt tc_lemma) - ) + val fin_tc_lemma = + Goal.prove ctxt [] [] fin_tc_goal (fin_fun_tac o #context) + val dom_lemma = + Goal.prove ctxt [] [] dom_goal + (fn {context = ctxt', ...} => blast_tac ctxt' 1) + val tc_lemma = + Goal.prove ctxt [] [] tc_goal + (fn {context = ctxt', ...} => + EqSubst.eqsubst_tac ctxt' [1] [dom_lemma] 1 + THEN resolve_tac ctxt' [@{thm FiniteFun_is_fun}] 1 + THEN resolve_tac ctxt' [fin_tc_lemma] 1) + val action_lemma = + Goal.prove ctxt [] [] action_goal + (fn {context = ctxt', ...} => + forward_tac ctxt' [@{thm le_natI}] 1 + THEN fin ctxt' + THEN REPEAT (resolve_tac ctxt' [@{thm natE}] 1 THEN step ctxt' tc_lemma) + THEN (step ctxt' tc_lemma)) in (action_lemma, tc_lemma, fvs, r) - end + end (* Returns the sum renaming, the goal for type_checking, and the actual lemmas for the left part of the sum. *) fun sum_ren_aux e e' ctxt = let val env = Variable.variant_frees ctxt [] [("env",@{typ i})] |> hd |> Free val (left_action_lemma,left_tc_lemma,_,r) = ren_thm e e' ctxt val (sum_ren,sum_goal_tc) = sum_tc_goal r e e' env val setV = Variable.variant_frees ctxt [] [("A",@{typ i})] |> hd |> Free fun hyp en = mem_ en (list_ setV) in (sum_ren, freeName env, Logic.list_implies (map (fn e => e |> hyp |> tp) [env], sum_goal_tc), left_tc_lemma, left_action_lemma) end fun sum_tc_lemma rho rho' ctxt = let val (sum_ren, envVar, tc_goal, left_tc_lemma, left_action_lemma) = sum_ren_aux rho rho' ctxt val (goal,fvs) = sum_action_goal sum_ren rho rho' ctxt val r = mk_ren rho rho' ctxt - in (sum_ren, goal,envVar, r,left_tc_lemma, left_action_lemma ,fvs, Goal.prove ctxt [] [] tc_goal - (fn _ => - resolve_tac ctxt [@{thm sum_type_id_aux2}] 1 - THEN asm_simp_tac ctxt 4 - THEN simp_tac ctxt 1 - THEN resolve_tac ctxt [left_tc_lemma] 1 - THEN (fin ctxt) - THEN (fin ctxt) - )) + in + (sum_ren, goal,envVar, r,left_tc_lemma, left_action_lemma, fvs, + Goal.prove ctxt [] [] tc_goal + (fn {context = ctxt', ...} => + resolve_tac ctxt' [@{thm sum_type_id_aux2}] 1 + THEN asm_simp_tac ctxt' 4 + THEN simp_tac ctxt' 1 + THEN resolve_tac ctxt' [left_tc_lemma] 1 + THEN (fin ctxt') + THEN (fin ctxt'))) end fun sum_rename rho rho' ctxt = let val (_, goal, _, left_rename, left_tc_lemma, left_action_lemma, fvs, sum_tc_lemma) = sum_tc_lemma rho rho' ctxt val action_lemma = fix_vars left_action_lemma fvs ctxt - in (sum_tc_lemma, Goal.prove ctxt [] [] goal - (fn _ => resolve_tac ctxt [@{thm sum_action_id_aux}] 1 - THEN (simp_tac ctxt 4) - THEN (simp_tac ctxt 1) - THEN (resolve_tac ctxt [left_tc_lemma] 1) - THEN (asm_full_simp_tac ctxt 1) - THEN (asm_full_simp_tac ctxt 1) - THEN (simp_tac ctxt 1) - THEN (simp_tac ctxt 1) - THEN (simp_tac ctxt 1) - THEN (full_simp_tac ctxt 1) - THEN (resolve_tac ctxt [action_lemma] 1) - THEN (blast_tac ctxt 1) - THEN (full_simp_tac ctxt 1) - THEN (full_simp_tac ctxt 1) - - ), fvs, left_rename - ) -end ; + in + (sum_tc_lemma, + Goal.prove ctxt [] [] goal + (fn {context = ctxt', ...} => + resolve_tac ctxt' [@{thm sum_action_id_aux}] 1 + THEN (simp_tac ctxt' 4) + THEN (simp_tac ctxt' 1) + THEN (resolve_tac ctxt' [left_tc_lemma] 1) + THEN (asm_full_simp_tac ctxt' 1) + THEN (asm_full_simp_tac ctxt' 1) + THEN (simp_tac ctxt' 1) + THEN (simp_tac ctxt' 1) + THEN (simp_tac ctxt' 1) + THEN (full_simp_tac ctxt' 1) + THEN (resolve_tac ctxt' [action_lemma] 1) + THEN (blast_tac ctxt' 1) + THEN (full_simp_tac ctxt' 1) + THEN (full_simp_tac ctxt' 1)), + fvs, left_rename) + end; end \ No newline at end of file diff --git a/thys/Generic_Deriving/derive_laws.ML b/thys/Generic_Deriving/derive_laws.ML --- a/thys/Generic_Deriving/derive_laws.ML +++ b/thys/Generic_Deriving/derive_laws.ML @@ -1,139 +1,144 @@ open Derive_Util signature DERIVE_LAWS = sig (* prove the iso theorem*) val prove_isomorphism : type_info -> Proof.context -> thm option * Proof.context val prove_instance_tac : typ -> class_info -> instance_info -> type_info -> Proof.context -> tactic val prove_equivalence_law : class_info -> instance_info -> Proof.context -> thm val prove_combinator_instance : (thm list list -> local_theory -> Proof.context) -> local_theory -> Proof.state end structure Derive_Laws : DERIVE_LAWS = struct fun get_class_info thy classname = Symtab.lookup (Class_Data.get thy) classname fun prove_isomorphism type_info lthy = let val tname_short = Long_Name.base_name (#tname type_info) val from_info = the (#from_info (#rep_info type_info)) val to_info = the (#to_info (#rep_info type_info)) val from_f = hd (#fs from_info) val to_f = hd (#fs to_info) val from_induct = hd (the (#inducts from_info)) val to_induct = hd (the (#inducts to_info)) val iso_thm = HOLogic.mk_Trueprop (Const (\<^const_name>\Derive.iso\, dummyT) $ from_f $ to_f) |> Syntax.check_term lthy val induct_tac_to = (Induct_Tacs.induct_tac lthy [[SOME "b"]] (SOME [to_induct]) 2) val induct_tac_from = (Induct_Tacs.induct_tac lthy [[SOME "a"]] (SOME [from_induct]) 1) - val iso_thm_proved = Goal.prove lthy [] [] iso_thm - (fn _ => (resolve_tac lthy [@{thm Derive.iso_intro}] 1) THEN - induct_tac_to THEN induct_tac_from THEN - (ALLGOALS (asm_full_simp_tac lthy))) + val iso_thm_proved = + Goal.prove lthy [] [] iso_thm + (fn {context = ctxt, ...} => + resolve_tac ctxt [@{thm Derive.iso_intro}] 1 THEN + induct_tac_to THEN induct_tac_from THEN + ALLGOALS (asm_full_simp_tac ctxt)) val ((_,thms),lthy') = Local_Theory.note ((Binding.name ("conversion_iso_" ^ tname_short),[]), [iso_thm_proved]) lthy val ctxt_thy = Proof_Context.init_global (Proof_Context.theory_of lthy') val thm = singleton (Proof_Context.export lthy' ctxt_thy) (hd thms) in (SOME thm, lthy') end fun prove_equivalence_law cl_info inst_info ctxt = let val class = #class cl_info val classname = #classname cl_info val class_law = the (#class_law cl_info) val op_defs = #defs inst_info val axioms = the (#axioms cl_info) val axioms_def = the_list (#axioms_def cl_info) val class_def = the (#class_def cl_info) val ops_raw = map (Thm.full_prop_of #> HOLogic.dest_Trueprop #> HOLogic.dest_eq #> fst #> strip_comb #> fst) op_defs val ops = map (dest_Const #> apsnd (K dummyT) #> Const) ops_raw val class_law_const = Thm.full_prop_of class_law |> HOLogic.dest_Trueprop |> HOLogic.dest_eq |> fst |> strip_comb |> fst val class_law_const_dummy = dest_Const class_law_const |> apsnd (K dummyT) |> Const val axioms_thms = map (Local_Defs.unfold ctxt (class_def :: axioms_def)) axioms val superclasses = get_superclasses class classname (Proof_Context.theory_of ctxt) val superclass_laws = map (get_class_info (Proof_Context.theory_of ctxt) #> the #> #equivalence_thm #> the) superclasses val t = list_comb (class_law_const_dummy,ops) |> HOLogic.mk_Trueprop |> Syntax.check_term ctxt in Goal.prove ctxt [] [] t - (fn _ => (Local_Defs.unfold_tac ctxt [class_law]) - THEN (HEADGOAL (Method.insert_tac ctxt (axioms_thms@superclass_laws))) - THEN (HEADGOAL (asm_full_simp_tac ctxt))) + (fn {context = ctxt', ...} => + Local_Defs.unfold_tac ctxt' [class_law] + THEN (HEADGOAL (Method.insert_tac ctxt' (axioms_thms @ superclass_laws))) + THEN (HEADGOAL (asm_full_simp_tac ctxt'))) end fun prove_instance_tac T cl_info inst_info ty_info ctxt = let val transfer_thms = the (#transfer_law cl_info) |> map snd |> flat val iso_thm = the (#iso_thm ty_info) val ops = map (Thm.full_prop_of #> HOLogic.dest_Trueprop #> HOLogic.dest_eq #> fst #> strip_comb #> fst) (#defs inst_info) val op_defs = #defs inst_info val class_law = the (#class_law cl_info) val equivalence_thm = the (#equivalence_thm cl_info) val class_law_const = Thm.full_prop_of class_law |> HOLogic.dest_Trueprop |> HOLogic.dest_eq |> fst |> strip_comb |> fst val ops = map (dest_Const #> apsnd (K dummyT) #> Const) ops val class_law_const_dummy = dest_Const class_law_const |> apsnd (K dummyT) |> Const val class_law_inst = HOLogic.mk_Trueprop (list_comb (class_law_const_dummy, ops)) |> singleton (Type_Infer_Context.infer_types ctxt) |> (fn t => subst_TVars ([(Term.add_tvar_names t [] |> hd,T)]) t ) val transfer_thm_inst = (hd transfer_thms) OF [iso_thm,equivalence_thm] - val class_law_inst_proved = Goal.prove ctxt [] [] class_law_inst - (fn _ => (Local_Defs.unfold_tac ctxt op_defs) - THEN (Proof_Context.fact_tac ctxt [transfer_thm_inst] 1)) + val class_law_inst_proved = + Goal.prove ctxt [] [] class_law_inst + (fn {context = ctxt', ...} => + Local_Defs.unfold_tac ctxt' op_defs + THEN (Proof_Context.fact_tac ctxt' [transfer_thm_inst] 1)) val ctxt_thy = Proof_Context.init_global (Proof_Context.theory_of ctxt) val class_law_inst_def = singleton (Proof_Context.export ctxt ctxt_thy) class_law_inst_proved val class_law_unfolded = Local_Defs.unfold ctxt [class_law] class_law_inst_def val class_tac = Class.intro_classes_tac ctxt [] THEN (ALLGOALS (Method.insert_tac ctxt [class_law_unfolded])) THEN (ALLGOALS (blast_tac ctxt)) in class_tac end fun prove_combinator_instance after_qed lthy = let fun class_tac thms ctxt = Class.intro_classes_tac ctxt [] THEN (ALLGOALS (Method.insert_tac ctxt thms)) THEN (ALLGOALS (blast_tac ctxt)) fun prove_class_laws_manually st ctxt = let val thm = #goal (Proof.simple_goal st) val goal = (Class.intro_classes_tac ctxt []) thm |> Seq.hd val goals = Thm.prems_of goal val goals_formatted = (map single (goals ~~ (replicate (length goals) []))) fun prove_class thms = let val ctxt_thy = Proof_Context.init_global (Proof_Context.theory_of ctxt) val thms' = (Proof_Context.export ctxt ctxt_thy) (flat thms) in Class.prove_instantiation_exit (class_tac thms') ctxt end fun after_qed' thms _ = prove_class thms |> Named_Target.theory_init |> after_qed [] val st' = Proof.theorem NONE after_qed' goals_formatted ctxt in st' end val st = Class.instantiation_instance I lthy in prove_class_laws_manually st lthy end end \ No newline at end of file diff --git a/thys/Goedel_HFSet_Semanticless/SyntaxN.thy b/thys/Goedel_HFSet_Semanticless/SyntaxN.thy --- a/thys/Goedel_HFSet_Semanticless/SyntaxN.thy +++ b/thys/Goedel_HFSet_Semanticless/SyntaxN.thy @@ -1,1409 +1,1409 @@ chapter\Syntax of Terms and Formulas using Nominal Logic\ theory SyntaxN imports Nominal2.Nominal2 HereditarilyFinite.OrdArith begin section\Terms and Formulas\ subsection\Hf is a pure permutation type\ instantiation hf :: pt begin definition "p \ (s::hf) = s" instance by standard (simp_all add: permute_hf_def) end instance hf :: pure proof qed (rule permute_hf_def) atom_decl name declare fresh_set_empty [simp] lemma supp_name [simp]: fixes i::name shows "supp i = {atom i}" by (rule supp_at_base) subsection\The datatypes\ nominal_datatype tm = Zero | Var name | Eats tm tm nominal_datatype fm = Mem tm tm (infixr "IN" 150) | Eq tm tm (infixr "EQ" 150) | Disj fm fm (infixr "OR" 130) | Neg fm | Ex x::name f::fm binds x in f text \Mem, Eq are atomic formulas; Disj, Neg, Ex are non-atomic\ declare tm.supp [simp] fm.supp [simp] subsection\Substitution\ nominal_function subst :: "name \ tm \ tm \ tm" where "subst i x Zero = Zero" | "subst i x (Var k) = (if i=k then x else Var k)" | "subst i x (Eats t u) = Eats (subst i x t) (subst i x u)" by (auto simp: eqvt_def subst_graph_aux_def) (metis tm.strong_exhaust) nominal_termination (eqvt) by lexicographic_order lemma fresh_subst_if [simp]: "j \ subst i x t \ (atom i \ t \ j \ t) \ (j \ x \ (j \ t \ j = atom i))" by (induct t rule: tm.induct) (auto simp: fresh_at_base) lemma forget_subst_tm [simp]: "atom a \ tm \ subst a x tm = tm" by (induct tm rule: tm.induct) (simp_all add: fresh_at_base) lemma subst_tm_id [simp]: "subst a (Var a) tm = tm" by (induct tm rule: tm.induct) simp_all lemma subst_tm_commute [simp]: "atom j \ tm \ subst j u (subst i t tm) = subst i (subst j u t) tm" by (induct tm rule: tm.induct) (auto simp: fresh_Pair) lemma subst_tm_commute2 [simp]: "atom j \ t \ atom i \ u \ i \ j \ subst j u (subst i t tm) = subst i t (subst j u tm)" by (induct tm rule: tm.induct) auto lemma repeat_subst_tm [simp]: "subst i u (subst i t tm) = subst i (subst i u t) tm" by (induct tm rule: tm.induct) auto nominal_function subst_fm :: "fm \ name \ tm \ fm" ("_'(_::=_')" [1000, 0, 0] 200) where Mem: "(Mem t u)(i::=x) = Mem (subst i x t) (subst i x u)" | Eq: "(Eq t u)(i::=x) = Eq (subst i x t) (subst i x u)" | Disj: "(Disj A B)(i::=x) = Disj (A(i::=x)) (B(i::=x))" | Neg: "(Neg A)(i::=x) = Neg (A(i::=x))" | Ex: "atom j \ (i, x) \ (Ex j A)(i::=x) = Ex j (A(i::=x))" apply (simp add: eqvt_def subst_fm_graph_aux_def) apply auto [16] apply (rule_tac y=a and c="(aa, b)" in fm.strong_exhaust) apply (auto simp: eqvt_at_def fresh_star_def fresh_Pair fresh_at_base) apply (metis flip_at_base_simps(3) flip_fresh_fresh) done nominal_termination (eqvt) by lexicographic_order lemma size_subst_fm [simp]: "size (A(i::=x)) = size A" by (nominal_induct A avoiding: i x rule: fm.strong_induct) auto lemma forget_subst_fm [simp]: "atom a \ A \ A(a::=x) = A" by (nominal_induct A avoiding: a x rule: fm.strong_induct) (auto simp: fresh_at_base) lemma subst_fm_id [simp]: "A(a::=Var a) = A" by (nominal_induct A avoiding: a rule: fm.strong_induct) (auto simp: fresh_at_base) lemma fresh_subst_fm_if [simp]: "j \ (A(i::=x)) \ (atom i \ A \ j \ A) \ (j \ x \ (j \ A \ j = atom i))" by (nominal_induct A avoiding: i x rule: fm.strong_induct) (auto simp: fresh_at_base) lemma subst_fm_commute [simp]: "atom j \ A \ (A(i::=t))(j::=u) = A(i ::= subst j u t)" by (nominal_induct A avoiding: i j t u rule: fm.strong_induct) (auto simp: fresh_at_base) lemma repeat_subst_fm [simp]: "(A(i::=t))(i::=u) = A(i ::= subst i u t)" by (nominal_induct A avoiding: i t u rule: fm.strong_induct) auto lemma subst_fm_Ex_with_renaming: "atom i' \ (A, i, j, t) \ (Ex i A)(j ::= t) = Ex i' (((i \ i') \ A)(j ::= t))" by (rule subst [of "Ex i' ((i \ i') \ A)" "Ex i A"]) (auto simp: Abs1_eq_iff flip_def swap_commute) text \the simplifier cannot apply the rule above, because it introduces a new variable at the right hand side.\ simproc_setup subst_fm_renaming ("(Ex i A)(j ::= t)") = \fn _ => fn ctxt => fn ctrm => let val _ $ (_ $ i $ A) $ j $ t = Thm.term_of ctrm val atoms = Simplifier.prems_of ctxt |> map_filter (fn thm => case Thm.prop_of thm of _ $ (Const (@{const_name fresh}, _) $ atm $ _) => SOME (atm) | _ => NONE) |> distinct ((=)) fun get_thm atm = let val goal = HOLogic.mk_Trueprop (mk_fresh atm (HOLogic.mk_tuple [A, i, j, t])) in - SOME ((Goal.prove ctxt [] [] goal (K (asm_full_simp_tac ctxt 1))) + SOME ((Goal.prove ctxt [] [] goal (fn {context = ctxt', ...} => asm_full_simp_tac ctxt' 1)) RS @{thm subst_fm_Ex_with_renaming} RS eq_reflection) handle ERROR _ => NONE end in get_first get_thm atoms end \ subsection\Derived syntax\ subsubsection\Ordered pairs\ definition HPair :: "tm \ tm \ tm" where "HPair a b = Eats (Eats Zero (Eats (Eats Zero b) a)) (Eats (Eats Zero a) a)" lemma HPair_eqvt [eqvt]: "(p \ HPair a b) = HPair (p \ a) (p \ b)" by (auto simp: HPair_def) lemma fresh_HPair [simp]: "x \ HPair a b \ (x \ a \ x \ b)" by (auto simp: HPair_def) lemma HPair_injective_iff [iff]: "HPair a b = HPair a' b' \ (a = a' \ b = b')" by (auto simp: HPair_def) lemma subst_tm_HPair [simp]: "subst i x (HPair a b) = HPair (subst i x a) (subst i x b)" by (auto simp: HPair_def) subsubsection\Ordinals\ definition SUCC :: "tm \ tm" where "SUCC x \ Eats x x" fun ORD_OF :: "nat \ tm" where "ORD_OF 0 = Zero" | "ORD_OF (Suc k) = SUCC (ORD_OF k)" lemma SUCC_fresh_iff [simp]: "a \ SUCC t \ a \ t" by (simp add: SUCC_def) lemma SUCC_eqvt [eqvt]: "(p \ SUCC a) = SUCC (p \ a)" by (simp add: SUCC_def) lemma SUCC_subst [simp]: "subst i t (SUCC k) = SUCC (subst i t k)" by (simp add: SUCC_def) lemma ORD_OF_fresh [simp]: "a \ ORD_OF n" by (induct n) (auto simp: SUCC_def) lemma ORD_OF_eqvt [eqvt]: "(p \ ORD_OF n) = ORD_OF (p \ n)" by (induct n) (auto simp: permute_pure SUCC_eqvt) subsection\Derived logical connectives\ abbreviation Imp :: "fm \ fm \ fm" (infixr "IMP" 125) where "Imp A B \ Disj (Neg A) B" abbreviation All :: "name \ fm \ fm" where "All i A \ Neg (Ex i (Neg A))" abbreviation All2 :: "name \ tm \ fm \ fm" \ \bounded universal quantifier, for Sigma formulas\ where "All2 i t A \ All i ((Var i IN t) IMP A)" subsubsection\Conjunction\ definition Conj :: "fm \ fm \ fm" (infixr "AND" 135) where "Conj A B \ Neg (Disj (Neg A) (Neg B))" lemma Conj_eqvt [eqvt]: "p \ (A AND B) = (p \ A) AND (p \ B)" by (simp add: Conj_def) lemma fresh_Conj [simp]: "a \ A AND B \ (a \ A \ a \ B)" by (auto simp: Conj_def) lemma supp_Conj [simp]: "supp (A AND B) = supp A \ supp B" by (auto simp: Conj_def) lemma size_Conj [simp]: "size (A AND B) = size A + size B + 4" by (simp add: Conj_def) lemma Conj_injective_iff [iff]: "(A AND B) = (A' AND B') \ (A = A' \ B = B')" by (auto simp: Conj_def) lemma subst_fm_Conj [simp]: "(A AND B)(i::=x) = (A(i::=x)) AND (B(i::=x))" by (auto simp: Conj_def) subsubsection\If and only if\ definition Iff :: "fm \ fm \ fm" (infixr "IFF" 125) where "Iff A B = Conj (Imp A B) (Imp B A)" lemma Iff_eqvt [eqvt]: "p \ (A IFF B) = (p \ A) IFF (p \ B)" by (simp add: Iff_def) lemma fresh_Iff [simp]: "a \ A IFF B \ (a \ A \ a \ B)" by (auto simp: Conj_def Iff_def) lemma size_Iff [simp]: "size (A IFF B) = 2*(size A + size B) + 8" by (simp add: Iff_def) lemma Iff_injective_iff [iff]: "(A IFF B) = (A' IFF B') \ (A = A' \ B = B')" by (auto simp: Iff_def) lemma subst_fm_Iff [simp]: "(A IFF B)(i::=x) = (A(i::=x)) IFF (B(i::=x))" by (auto simp: Iff_def) section\Axioms and Theorems\ subsection\Logical axioms\ inductive_set boolean_axioms :: "fm set" where Ident: "A IMP A \ boolean_axioms" | DisjI1: "A IMP (A OR B) \ boolean_axioms" | DisjCont: "(A OR A) IMP A \ boolean_axioms" | DisjAssoc: "(A OR (B OR C)) IMP ((A OR B) OR C) \ boolean_axioms" | DisjConj: "(C OR A) IMP (((Neg C) OR B) IMP (A OR B)) \ boolean_axioms" inductive_set special_axioms :: "fm set" where I: "A(i::=x) IMP (Ex i A) \ special_axioms" inductive_set induction_axioms :: "fm set" where ind: "atom (j::name) \ (i,A) \ A(i::=Zero) IMP ((All i (All j (A IMP (A(i::= Var j) IMP A(i::= Eats(Var i)(Var j)))))) IMP (All i A)) \ induction_axioms" subsection \Concrete variables\ declare Abs_name_inject[simp] abbreviation "X0 \ Abs_name (Atom (Sort ''SyntaxN.name'' []) 0)" abbreviation "X1 \ Abs_name (Atom (Sort ''SyntaxN.name'' []) (Suc 0))" \ \We prefer @{term "Suc 0"} because simplification will transform 1 to that form anyway.\ abbreviation "X2 \ Abs_name (Atom (Sort ''SyntaxN.name'' []) 2)" abbreviation "X3 \ Abs_name (Atom (Sort ''SyntaxN.name'' []) 3)" abbreviation "X4 \ Abs_name (Atom (Sort ''SyntaxN.name'' []) 4)" subsection\The HF axioms\ definition HF1 :: fm where \ \the axiom @{term"z=0 \ (\x. \ x \<^bold>\ z)"}\ "HF1 = (Var X0 EQ Zero) IFF (All X1 (Neg (Var X1 IN Var X0)))" definition HF2 :: fm where \ \the axiom @{term"z = x \ y \ (\u. u \<^bold>\ z \ u \<^bold>\ x | u=y)"}\ "HF2 \ Var X0 EQ Eats (Var X1) (Var X2) IFF All X3 (Var X3 IN Var X0 IFF Var X3 IN Var X1 OR Var X3 EQ Var X2)" definition HF_axioms where "HF_axioms = {HF1, HF2}" subsection\Equality axioms\ definition refl_ax :: fm where "refl_ax = Var X1 EQ Var X1" definition eq_cong_ax :: fm where "eq_cong_ax = ((Var X1 EQ Var X2) AND (Var X3 EQ Var X4)) IMP ((Var X1 EQ Var X3) IMP (Var X2 EQ Var X4))" definition mem_cong_ax :: fm where "mem_cong_ax = ((Var X1 EQ Var X2) AND (Var X3 EQ Var X4)) IMP ((Var X1 IN Var X3) IMP (Var X2 IN Var X4))" definition eats_cong_ax :: fm where "eats_cong_ax = ((Var X1 EQ Var X2) AND (Var X3 EQ Var X4)) IMP ((Eats (Var X1) (Var X3)) EQ (Eats (Var X2) (Var X4)))" definition equality_axioms :: "fm set" where "equality_axioms = {refl_ax, eq_cong_ax, mem_cong_ax, eats_cong_ax}" subsection\The proof system\ text\This arbitrary additional axiom generalises the statements of the incompleteness theorems and other results to any formal system stronger than the HF theory. The additional axiom could be the conjunction of any finite number of assertions. Any more general extension must be a form that can be formalised for the proof predicate.\ consts extra_axiom :: fm inductive hfthm :: "fm set \ fm \ bool" (infixl "\" 55) where Hyp: "A \ H \ H \ A" | Extra: "H \ extra_axiom" | Bool: "A \ boolean_axioms \ H \ A" | Eq: "A \ equality_axioms \ H \ A" | Spec: "A \ special_axioms \ H \ A" | HF: "A \ HF_axioms \ H \ A" | Ind: "A \ induction_axioms \ H \ A" | MP: "H \ A IMP B \ H' \ A \ H \ H' \ B" | Exists: "H \ A IMP B \ atom i \ B \ \C \ H. atom i \ C \ H \ (Ex i A) IMP B" subsection\Derived rules of inference\ lemma contraction: "insert A (insert A H) \ B \ insert A H \ B" by (metis insert_absorb2) lemma thin_Un: "H \ A \ H \ H' \ A" by (metis Bool MP boolean_axioms.Ident sup_commute) lemma thin: "H \ A \ H \ H' \ H' \ A" by (metis Un_absorb1 thin_Un) lemma thin0: "{} \ A \ H \ A" by (metis sup_bot_left thin_Un) lemma thin1: "H \ B \ insert A H \ B" by (metis subset_insertI thin) lemma thin2: "insert A1 H \ B \ insert A1 (insert A2 H) \ B" by (blast intro: thin) lemma thin3: "insert A1 (insert A2 H) \ B \ insert A1 (insert A2 (insert A3 H)) \ B" by (blast intro: thin) lemma thin4: "insert A1 (insert A2 (insert A3 H)) \ B \ insert A1 (insert A2 (insert A3 (insert A4 H))) \ B" by (blast intro: thin) lemma rotate2: "insert A2 (insert A1 H) \ B \ insert A1 (insert A2 H) \ B" by (blast intro: thin) lemma rotate3: "insert A3 (insert A1 (insert A2 H)) \ B \ insert A1 (insert A2 (insert A3 H)) \ B" by (blast intro: thin) lemma rotate4: "insert A4 (insert A1 (insert A2 (insert A3 H))) \ B \ insert A1 (insert A2 (insert A3 (insert A4 H))) \ B" by (blast intro: thin) lemma rotate5: "insert A5 (insert A1 (insert A2 (insert A3 (insert A4 H)))) \ B \ insert A1 (insert A2 (insert A3 (insert A4 (insert A5 H)))) \ B" by (blast intro: thin) lemma rotate6: "insert A6 (insert A1 (insert A2 (insert A3 (insert A4 (insert A5 H))))) \ B \ insert A1 (insert A2 (insert A3 (insert A4 (insert A5 (insert A6 H))))) \ B" by (blast intro: thin) lemma rotate7: "insert A7 (insert A1 (insert A2 (insert A3 (insert A4 (insert A5 (insert A6 H)))))) \ B \ insert A1 (insert A2 (insert A3 (insert A4 (insert A5 (insert A6 (insert A7 H)))))) \ B" by (blast intro: thin) lemma rotate8: "insert A8 (insert A1 (insert A2 (insert A3 (insert A4 (insert A5 (insert A6 (insert A7 H))))))) \ B \ insert A1 (insert A2 (insert A3 (insert A4 (insert A5 (insert A6 (insert A7 (insert A8 H))))))) \ B" by (blast intro: thin) lemma rotate9: "insert A9 (insert A1 (insert A2 (insert A3 (insert A4 (insert A5 (insert A6 (insert A7 (insert A8 H)))))))) \ B \ insert A1 (insert A2 (insert A3 (insert A4 (insert A5 (insert A6 (insert A7 (insert A8 (insert A9 H)))))))) \ B" by (blast intro: thin) lemma rotate10: "insert A10 (insert A1 (insert A2 (insert A3 (insert A4 (insert A5 (insert A6 (insert A7 (insert A8 (insert A9 H))))))))) \ B \ insert A1 (insert A2 (insert A3 (insert A4 (insert A5 (insert A6 (insert A7 (insert A8 (insert A9 (insert A10 H))))))))) \ B" by (blast intro: thin) lemma rotate11: "insert A11 (insert A1 (insert A2 (insert A3 (insert A4 (insert A5 (insert A6 (insert A7 (insert A8 (insert A9 (insert A10 H)))))))))) \ B \ insert A1 (insert A2 (insert A3 (insert A4 (insert A5 (insert A6 (insert A7 (insert A8 (insert A9 (insert A10 (insert A11 H)))))))))) \ B" by (blast intro: thin) lemma rotate12: "insert A12 (insert A1 (insert A2 (insert A3 (insert A4 (insert A5 (insert A6 (insert A7 (insert A8 (insert A9 (insert A10 (insert A11 H))))))))))) \ B \ insert A1 (insert A2 (insert A3 (insert A4 (insert A5 (insert A6 (insert A7 (insert A8 (insert A9 (insert A10 (insert A11 (insert A12 H))))))))))) \ B" by (blast intro: thin) lemma rotate13: "insert A13 (insert A1 (insert A2 (insert A3 (insert A4 (insert A5 (insert A6 (insert A7 (insert A8 (insert A9 (insert A10 (insert A11 (insert A12 H)))))))))))) \ B \ insert A1 (insert A2 (insert A3 (insert A4 (insert A5 (insert A6 (insert A7 (insert A8 (insert A9 (insert A10 (insert A11 (insert A12 (insert A13 H)))))))))))) \ B" by (blast intro: thin) lemma rotate14: "insert A14 (insert A1 (insert A2 (insert A3 (insert A4 (insert A5 (insert A6 (insert A7 (insert A8 (insert A9 (insert A10 (insert A11 (insert A12 (insert A13 H))))))))))))) \ B \ insert A1 (insert A2 (insert A3 (insert A4 (insert A5 (insert A6 (insert A7 (insert A8 (insert A9 (insert A10 (insert A11 (insert A12 (insert A13 (insert A14 H))))))))))))) \ B" by (blast intro: thin) lemma rotate15: "insert A15 (insert A1 (insert A2 (insert A3 (insert A4 (insert A5 (insert A6 (insert A7 (insert A8 (insert A9 (insert A10 (insert A11 (insert A12 (insert A13 (insert A14 H)))))))))))))) \ B \ insert A1 (insert A2 (insert A3 (insert A4 (insert A5 (insert A6 (insert A7 (insert A8 (insert A9 (insert A10 (insert A11 (insert A12 (insert A13 (insert A14 (insert A15 H)))))))))))))) \ B" by (blast intro: thin) lemma MP_same: "H \ A IMP B \ H \ A \ H \ B" by (metis MP Un_absorb) lemma MP_thin: "HA \ A IMP B \ HB \ A \ HA \ HB \ H \ H \ B" by (metis MP_same le_sup_iff thin) lemma MP_null: "{} \ A IMP B \ H \ A \ H \ B" by (metis MP_same thin0) lemma Disj_commute: "H \ B OR A \ H \ A OR B" using DisjConj [of B A B] Ident [of B] by (metis Bool MP_same) lemma S: assumes "H \ A IMP (B IMP C)" "H' \ A IMP B" shows "H \ H' \ A IMP C" proof - have "H' \ H \ (Neg A) OR (C OR (Neg A))" by (metis Bool MP MP_same boolean_axioms.DisjConj Disj_commute DisjAssoc assms) thus ?thesis by (metis Bool Disj_commute Un_commute MP_same DisjAssoc DisjCont DisjI1) qed lemma Assume: "insert A H \ A" by (metis Hyp insertI1) lemmas AssumeH = Assume Assume [THEN rotate2] Assume [THEN rotate3] Assume [THEN rotate4] Assume [THEN rotate5] Assume [THEN rotate6] Assume [THEN rotate7] Assume [THEN rotate8] Assume [THEN rotate9] Assume [THEN rotate10] Assume [THEN rotate11] Assume [THEN rotate12] declare AssumeH [intro!] lemma Imp_triv_I: "H \ B \ H \ A IMP B" by (metis Bool Disj_commute MP_same boolean_axioms.DisjI1) lemma DisjAssoc1: "H \ A OR (B OR C) \ H \ (A OR B) OR C" by (metis Bool MP_same boolean_axioms.DisjAssoc) lemma DisjAssoc2: "H \ (A OR B) OR C \ H \ A OR (B OR C)" by (metis DisjAssoc1 Disj_commute) lemma Disj_commute_Imp: "H \ (B OR A) IMP (A OR B)" using DisjConj [of B A B] Ident [of B] by (metis Bool DisjAssoc2 Disj_commute MP_same) lemma Disj_Semicong_1: "H \ A OR C \ H \ A IMP B \ H \ B OR C" using DisjConj [of A C B] by (metis Bool Disj_commute MP_same) lemma Imp_Imp_commute: "H \ B IMP (A IMP C) \ H \ A IMP (B IMP C)" by (metis DisjAssoc1 DisjAssoc2 Disj_Semicong_1 Disj_commute_Imp) subsection\The Deduction Theorem\ lemma deduction_Diff: assumes "H \ B" shows "H - {C} \ C IMP B" using assms proof (induct) case (Hyp A H) thus ?case by (metis Bool Imp_triv_I boolean_axioms.Ident hfthm.Hyp member_remove remove_def) next case (Extra H) thus ?case by (metis Imp_triv_I hfthm.Extra) next case (Bool A H) thus ?case by (metis Imp_triv_I hfthm.Bool) next case (Eq A H) thus ?case by (metis Imp_triv_I hfthm.Eq) next case (Spec A H) thus ?case by (metis Imp_triv_I hfthm.Spec) next case (HF A H) thus ?case by (metis Imp_triv_I hfthm.HF) next case (Ind A H) thus ?case by (metis Imp_triv_I hfthm.Ind) next case (MP H A B H') hence "(H-{C}) \ (H'-{C}) \ Imp C B" by (simp add: S) thus ?case by (metis Un_Diff) next case (Exists H A B i) show ?case proof (cases "C \ H") case True hence "atom i \ C" using Exists by auto moreover have "H - {C} \ A IMP C IMP B" using Exists by (metis Imp_Imp_commute) ultimately have "H - {C} \ (Ex i A) IMP C IMP B" using Exists by (metis fm.fresh(3) fm.fresh(4) hfthm.Exists member_remove remove_def) thus ?thesis by (metis Imp_Imp_commute) next case False hence "H - {C} = H" by auto thus ?thesis using Exists by (metis Imp_triv_I hfthm.Exists) qed qed theorem Imp_I [intro!]: "insert A H \ B \ H \ A IMP B" by (metis Diff_insert_absorb Imp_triv_I deduction_Diff insert_absorb) lemma anti_deduction: "H \ A IMP B \ insert A H \ B" by (metis Assume MP_same thin1) subsection\Cut rules\ lemma cut: "H \ A \ insert A H' \ B \ H \ H' \ B" by (metis MP Un_commute Imp_I) lemma cut_same: "H \ A \ insert A H \ B \ H \ B" by (metis Un_absorb cut) lemma cut_thin: "HA \ A \ insert A HB \ B \ HA \ HB \ H \ H \ B" by (metis thin cut) lemma cut0: "{} \ A \ insert A H \ B \ H \ B" by (metis cut_same thin0) lemma cut1: "{A} \ B \ H \ A \ H \ B" by (metis cut sup_bot_right) lemma rcut1: "{A} \ B \ insert B H \ C \ insert A H \ C" by (metis Assume cut1 cut_same rotate2 thin1) lemma cut2: "\{A,B} \ C; H \ A; H \ B\ \ H \ C" by (metis Un_empty_right Un_insert_right cut cut_same) lemma rcut2: "{A,B} \ C \ insert C H \ D \ H \ B \ insert A H \ D" by (metis Assume cut2 cut_same insert_commute thin1) lemma cut3: "\{A,B,C} \ D; H \ A; H \ B; H \ C\ \ H \ D" by (metis MP_same cut2 Imp_I) lemma cut4: "\{A,B,C,D} \ E; H \ A; H \ B; H \ C; H \ D\ \ H \ E" by (metis MP_same cut3 [of B C D] Imp_I) section\Miscellaneous logical rules\ lemma Disj_I1: "H \ A \ H \ A OR B" by (metis Bool MP_same boolean_axioms.DisjI1) lemma Disj_I2: "H \ B \ H \ A OR B" by (metis Disj_commute Disj_I1) lemma Peirce: "H \ (Neg A) IMP A \ H \ A" using DisjConj [of "Neg A" A A] DisjCont [of A] by (metis Bool MP_same boolean_axioms.Ident) lemma Contra: "insert (Neg A) H \ A \ H \ A" by (metis Peirce Imp_I) lemma Imp_Neg_I: "H \ A IMP B \ H \ A IMP (Neg B) \ H \ Neg A" by (metis DisjConj [of B "Neg A" "Neg A"] DisjCont Bool Disj_commute MP_same) lemma NegNeg_I: "H \ A \ H \ Neg (Neg A)" using DisjConj [of "Neg (Neg A)" "Neg A" "Neg (Neg A)"] by (metis Bool Ident MP_same) lemma NegNeg_D: "H \ Neg (Neg A) \ H \ A" by (metis Disj_I1 Peirce) lemma Neg_D: "H \ Neg A \ H \ A \ H \ B" by (metis Imp_Neg_I Imp_triv_I NegNeg_D) lemma Disj_Neg_1: "H \ A OR B \ H \ Neg B \ H \ A" by (metis Disj_I1 Disj_Semicong_1 Disj_commute Peirce) lemma Disj_Neg_2: "H \ A OR B \ H \ Neg A \ H \ B" by (metis Disj_Neg_1 Disj_commute) lemma Neg_Disj_I: "H \ Neg A \ H \ Neg B \ H \ Neg (A OR B)" by (metis Bool Disj_Neg_1 MP_same boolean_axioms.Ident DisjAssoc) lemma Conj_I [intro!]: "H \ A \ H \ B \ H \ A AND B" by (metis Conj_def NegNeg_I Neg_Disj_I) lemma Conj_E1: "H \ A AND B \ H \ A" by (metis Conj_def Bool Disj_Neg_1 NegNeg_D boolean_axioms.DisjI1) lemma Conj_E2: "H \ A AND B \ H \ B" by (metis Conj_def Bool Disj_I2 Disj_Neg_2 MP_same DisjAssoc Ident) lemma Conj_commute: "H \ B AND A \ H \ A AND B" by (metis Conj_E1 Conj_E2 Conj_I) lemma Conj_E: assumes "insert A (insert B H) \ C" shows "insert (A AND B) H \ C" apply (rule cut_same [where A=A], metis Conj_E1 Hyp insertI1) by (metis (full_types) AssumeH(2) Conj_E2 assms cut_same [where A=B] insert_commute thin2) lemmas Conj_EH = Conj_E Conj_E [THEN rotate2] Conj_E [THEN rotate3] Conj_E [THEN rotate4] Conj_E [THEN rotate5] Conj_E [THEN rotate6] Conj_E [THEN rotate7] Conj_E [THEN rotate8] Conj_E [THEN rotate9] Conj_E [THEN rotate10] declare Conj_EH [intro!] lemma Neg_I0: assumes "(\B. atom i \ B \ insert A H \ B)" shows "H \ Neg A" by (rule Imp_Neg_I [where B = "Zero IN Zero"]) (auto simp: assms) lemma Neg_mono: "insert A H \ B \ insert (Neg B) H \ Neg A" by (rule Neg_I0) (metis Hyp Neg_D insert_commute insertI1 thin1) lemma Conj_mono: "insert A H \ B \ insert C H \ D \ insert (A AND C) H \ B AND D" by (metis Conj_E1 Conj_E2 Conj_I Hyp Un_absorb2 cut insertI1 subset_insertI) lemma Disj_mono: assumes "insert A H \ B" "insert C H \ D" shows "insert (A OR C) H \ B OR D" proof - { fix A B C H have "insert (A OR C) H \ (A IMP B) IMP C OR B" by (metis Bool Hyp MP_same boolean_axioms.DisjConj insertI1) hence "insert A H \ B \ insert (A OR C) H \ C OR B" by (metis MP_same Un_absorb Un_insert_right Imp_I thin_Un) } thus ?thesis by (metis cut_same assms thin2) qed lemma Disj_E: assumes A: "insert A H \ C" and B: "insert B H \ C" shows "insert (A OR B) H \ C" by (metis A B Disj_mono NegNeg_I Peirce) lemmas Disj_EH = Disj_E Disj_E [THEN rotate2] Disj_E [THEN rotate3] Disj_E [THEN rotate4] Disj_E [THEN rotate5] Disj_E [THEN rotate6] Disj_E [THEN rotate7] Disj_E [THEN rotate8] Disj_E [THEN rotate9] Disj_E [THEN rotate10] declare Disj_EH [intro!] lemma Contra': "insert A H \ Neg A \ H \ Neg A" by (metis Contra Neg_mono) lemma NegNeg_E [intro!]: "insert A H \ B \ insert (Neg (Neg A)) H \ B" by (metis NegNeg_D Neg_mono) declare NegNeg_E [THEN rotate2, intro!] declare NegNeg_E [THEN rotate3, intro!] declare NegNeg_E [THEN rotate4, intro!] declare NegNeg_E [THEN rotate5, intro!] declare NegNeg_E [THEN rotate6, intro!] declare NegNeg_E [THEN rotate7, intro!] declare NegNeg_E [THEN rotate8, intro!] lemma Imp_E: assumes A: "H \ A" and B: "insert B H \ C" shows "insert (A IMP B) H \ C" proof - have "insert (A IMP B) H \ B" by (metis Hyp A thin1 MP_same insertI1) thus ?thesis by (metis cut [where B=C] Un_insert_right sup_commute sup_idem B) qed lemma Imp_cut: assumes "insert C H \ A IMP B" "{A} \ C" shows "H \ A IMP B" by (metis Contra Disj_I1 Neg_mono assms rcut1) lemma Iff_I [intro!]: "insert A H \ B \ insert B H \ A \ H \ A IFF B" by (metis Iff_def Conj_I Imp_I) lemma Iff_MP_same: "H \ A IFF B \ H \ A \ H \ B" by (metis Iff_def Conj_E1 MP_same) lemma Iff_MP2_same: "H \ A IFF B \ H \ B \ H \ A" by (metis Iff_def Conj_E2 MP_same) lemma Iff_refl [intro!]: "H \ A IFF A" by (metis Hyp Iff_I insertI1) lemma Iff_sym: "H \ A IFF B \ H \ B IFF A" by (metis Iff_def Conj_commute) lemma Iff_trans: "H \ A IFF B \ H \ B IFF C \ H \ A IFF C" unfolding Iff_def by (metis Conj_E1 Conj_E2 Conj_I Disj_Semicong_1 Disj_commute) lemma Iff_E: "insert A (insert B H) \ C \ insert (Neg A) (insert (Neg B) H) \ C \ insert (A IFF B) H \ C" apply (auto simp: Iff_def insert_commute) apply (metis Disj_I1 Hyp anti_deduction insertCI) apply (metis Assume Disj_I1 anti_deduction) done lemma Iff_E1: assumes A: "H \ A" and B: "insert B H \ C" shows "insert (A IFF B) H \ C" by (metis Iff_def A B Conj_E Imp_E insert_commute thin1) lemma Iff_E2: assumes A: "H \ A" and B: "insert B H \ C" shows "insert (B IFF A) H \ C" by (metis Iff_def A B Bool Conj_E2 Conj_mono Imp_E boolean_axioms.Ident) lemma Iff_MP_left: "H \ A IFF B \ insert A H \ C \ insert B H \ C" by (metis Hyp Iff_E2 cut_same insertI1 insert_commute thin1) lemma Iff_MP_left': "H \ A IFF B \ insert B H \ C \ insert A H \ C" by (metis Iff_MP_left Iff_sym) lemma Swap: "insert (Neg B) H \ A \ insert (Neg A) H \ B" by (metis NegNeg_D Neg_mono) lemma Cases: "insert A H \ B \ insert (Neg A) H \ B \ H \ B" by (metis Contra Neg_D Neg_mono) lemma Neg_Conj_E: "H \ B \ insert (Neg A) H \ C \ insert (Neg (A AND B)) H \ C" by (metis Conj_I Swap thin1) lemma Disj_CI: "insert (Neg B) H \ A \ H \ A OR B" by (metis Contra Disj_I1 Disj_I2 Swap) lemma Disj_3I: "insert (Neg A) (insert (Neg C) H) \ B \ H \ A OR B OR C" by (metis Disj_CI Disj_commute insert_commute) lemma Contrapos1: "H \ A IMP B \ H \ Neg B IMP Neg A" by (metis Bool MP_same boolean_axioms.DisjConj boolean_axioms.Ident) lemma Contrapos2: "H \ (Neg B) IMP (Neg A) \ H \ A IMP B" by (metis Bool MP_same boolean_axioms.DisjConj boolean_axioms.Ident) lemma ContraAssumeN [intro]: "B \ H \ insert (Neg B) H \ A" by (metis Hyp Swap thin1) lemma ContraAssume: "Neg B \ H \ insert B H \ A" by (metis Disj_I1 Hyp anti_deduction) lemma ContraProve: "H \ B \ insert (Neg B) H \ A" by (metis Swap thin1) lemma Disj_IE1: "insert B H \ C \ insert (A OR B) H \ A OR C" by (metis Assume Disj_mono) lemmas Disj_IE1H = Disj_IE1 Disj_IE1 [THEN rotate2] Disj_IE1 [THEN rotate3] Disj_IE1 [THEN rotate4] Disj_IE1 [THEN rotate5] Disj_IE1 [THEN rotate6] Disj_IE1 [THEN rotate7] Disj_IE1 [THEN rotate8] declare Disj_IE1H [intro!] subsection\Quantifier reasoning\ lemma Ex_I: "H \ A(i::=x) \ H \ Ex i A" by (metis MP_same Spec special_axioms.intros) lemma Ex_E: assumes "insert A H \ B" "atom i \ B" "\C \ H. atom i \ C" shows "insert (Ex i A) H \ B" by (metis Exists Imp_I anti_deduction assms) lemma Ex_E_with_renaming: assumes "insert ((i \ i') \ A) H \ B" "atom i' \ (A,i,B)" "\C \ H. atom i' \ C" shows "insert (Ex i A) H \ B" proof - have "Ex i A = Ex i' ((i \ i') \ A)" using assms apply (auto simp: Abs1_eq_iff fresh_Pair) apply (metis flip_at_simps(2) fresh_at_base_permute_iff)+ done thus ?thesis by (metis Ex_E assms fresh_Pair) qed lemmas Ex_EH = Ex_E Ex_E [THEN rotate2] Ex_E [THEN rotate3] Ex_E [THEN rotate4] Ex_E [THEN rotate5] Ex_E [THEN rotate6] Ex_E [THEN rotate7] Ex_E [THEN rotate8] Ex_E [THEN rotate9] Ex_E [THEN rotate10] declare Ex_EH [intro!] lemma Ex_mono: "insert A H \ B \ \C \ H. atom i \ C \ insert (Ex i A) H \ (Ex i B)" by (auto simp add: intro: Ex_I [where x="Var i"]) lemma All_I [intro!]: "H \ A \ \C \ H. atom i \ C \ H \ All i A" by (auto intro: ContraProve Neg_I0) lemma All_D: "H \ All i A \ H \ A(i::=x)" by (metis Assume Ex_I NegNeg_D Neg_mono SyntaxN.Neg cut_same) lemma All_E: "insert (A(i::=x)) H \ B \ insert (All i A) H \ B" by (metis Ex_I NegNeg_D Neg_mono SyntaxN.Neg) lemma All_E': "H \ All i A \ insert (A(i::=x)) H \ B \ H \ B" by (metis All_D cut_same) lemma All2_E: "\atom i \ t; H \ x IN t; insert (A(i::=x)) H \ B\ \ insert (All2 i t A) H \ B" apply (rule All_E [where x=x], auto) by (metis Swap thin1) lemma All2_E': "\H \ All2 i t A; H \ x IN t; insert (A(i::=x)) H \ B; atom i \ t\ \ H \ B" by (metis All2_E cut_same) subsection\Congruence rules\ lemma Neg_cong: "H \ A IFF A' \ H \ Neg A IFF Neg A'" by (metis Iff_def Conj_E1 Conj_E2 Conj_I Contrapos1) lemma Disj_cong: "H \ A IFF A' \ H \ B IFF B' \ H \ A OR B IFF A' OR B'" by (metis Conj_E1 Conj_E2 Disj_mono Iff_I Iff_def anti_deduction) lemma Conj_cong: "H \ A IFF A' \ H \ B IFF B' \ H \ A AND B IFF A' AND B'" by (metis Conj_def Disj_cong Neg_cong) lemma Imp_cong: "H \ A IFF A' \ H \ B IFF B' \ H \ (A IMP B) IFF (A' IMP B')" by (metis Disj_cong Neg_cong) lemma Iff_cong: "H \ A IFF A' \ H \ B IFF B' \ H \ (A IFF B) IFF (A' IFF B')" by (metis Iff_def Conj_cong Imp_cong) lemma Ex_cong: "H \ A IFF A' \ \C \ H. atom i \ C \ H \ (Ex i A) IFF (Ex i A')" apply (rule Iff_I) apply (metis Ex_mono Hyp Iff_MP_same Un_absorb Un_insert_right insertI1 thin_Un) apply (metis Ex_mono Hyp Iff_MP2_same Un_absorb Un_insert_right insertI1 thin_Un) done lemma All_cong: "H \ A IFF A' \ \C \ H. atom i \ C \ H \ (All i A) IFF (All i A')" by (metis Ex_cong Neg_cong) lemma Subst: "H \ A \ \B \ H. atom i \ B \ H \ A (i::=x)" by (metis All_D All_I) section\Equality reasoning\ subsection\The congruence property for @{term Eq}, and other basic properties of equality\ lemma Eq_cong1: "{} \ (t EQ t' AND u EQ u') IMP (t EQ u IMP t' EQ u')" proof - obtain v2::name and v3::name and v4::name where v2: "atom v2 \ (t,X1,X3,X4)" and v3: "atom v3 \ (t,t',X1,v2,X4)" and v4: "atom v4 \ (t,t',u,X1,v2,v3)" by (metis obtain_fresh) have "{} \ (Var X1 EQ Var X2 AND Var X3 EQ Var X4) IMP (Var X1 EQ Var X3 IMP Var X2 EQ Var X4)" by (rule Eq) (simp add: eq_cong_ax_def equality_axioms_def) hence "{} \ (Var X1 EQ Var X2 AND Var X3 EQ Var X4) IMP (Var X1 EQ Var X3 IMP Var X2 EQ Var X4)" by (drule_tac i=X1 and x="Var X1" in Subst) simp_all hence "{} \ (Var X1 EQ Var v2 AND Var X3 EQ Var X4) IMP (Var X1 EQ Var X3 IMP Var v2 EQ Var X4)" by (drule_tac i=X2 and x="Var v2" in Subst) simp_all hence "{} \ (Var X1 EQ Var v2 AND Var v3 EQ Var X4) IMP (Var X1 EQ Var v3 IMP Var v2 EQ Var X4)" using v2 by (drule_tac i=X3 and x="Var v3" in Subst) simp_all hence "{} \ (Var X1 EQ Var v2 AND Var v3 EQ Var v4) IMP (Var X1 EQ Var v3 IMP Var v2 EQ Var v4)" using v2 v3 by (drule_tac i=X4 and x="Var v4" in Subst) simp_all hence "{} \ (t EQ Var v2 AND Var v3 EQ Var v4) IMP (t EQ Var v3 IMP Var v2 EQ Var v4)" using v2 v3 v4 by (drule_tac i=X1 and x=t in Subst) simp_all hence "{} \ (t EQ t' AND Var v3 EQ Var v4) IMP (t EQ Var v3 IMP t' EQ Var v4)" using v2 v3 v4 by (drule_tac i=v2 and x="t'" in Subst) simp_all hence "{} \ (t EQ t' AND u EQ Var v4) IMP (t EQ u IMP t' EQ Var v4)" using v3 v4 by (drule_tac i=v3 and x=u in Subst) simp_all thus ?thesis using v4 by (drule_tac i=v4 and x="u'" in Subst) simp_all qed lemma Refl [iff]: "H \ t EQ t" proof - have "{} \ Var X1 EQ Var X1" by (rule Eq) (simp add: equality_axioms_def refl_ax_def) hence "{} \ t EQ t" by (drule_tac i=X1 and x=t in Subst) simp_all thus ?thesis by (metis empty_subsetI thin) qed text\Apparently necessary in order to prove the congruence property.\ lemma Sym: assumes "H \ t EQ u" shows "H \ u EQ t" proof - have "{} \ (t EQ u AND t EQ t) IMP (t EQ t IMP u EQ t)" by (rule Eq_cong1) moreover have "{t EQ u} \ t EQ u AND t EQ t" by (metis Assume Conj_I Refl) ultimately have "{t EQ u} \ u EQ t" by (metis MP_same MP Refl sup_bot_left) thus "H \ u EQ t" by (metis assms cut1) qed lemma Sym_L: "insert (t EQ u) H \ A \ insert (u EQ t) H \ A" by (metis Assume Sym Un_empty_left Un_insert_left cut) lemma Trans: assumes "H \ x EQ y" "H \ y EQ z" shows "H \ x EQ z" proof - have "\H. H \ (x EQ x AND y EQ z) IMP (x EQ y IMP x EQ z)" by (metis Eq_cong1 bot_least thin) moreover have "{x EQ y, y EQ z} \ x EQ x AND y EQ z" by (metis Assume Conj_I Refl thin1) ultimately have "{x EQ y, y EQ z} \ x EQ z" by (metis Hyp MP_same insertI1) thus ?thesis by (metis assms cut2) qed lemma Eq_cong: assumes "H \ t EQ t'" "H \ u EQ u'" shows "H \ t EQ u IFF t' EQ u'" proof - { fix t t' u u' assume "H \ t EQ t'" "H \ u EQ u'" moreover have "{t EQ t', u EQ u'} \ t EQ u IMP t' EQ u'" using Eq_cong1 by (metis Assume Conj_I MP_null insert_commute) ultimately have "H \ t EQ u IMP t' EQ u'" by (metis cut2) } thus ?thesis by (metis Iff_def Conj_I assms Sym) qed lemma Eq_Trans_E: "H \ x EQ u \ insert (t EQ u) H \ A \ insert (x EQ t) H \ A" by (metis Assume Sym_L Trans cut_same thin1 thin2) subsection\The congruence property for @{term Mem}\ lemma Mem_cong1: "{} \ (t EQ t' AND u EQ u') IMP (t IN u IMP t' IN u')" proof - obtain v2::name and v3::name and v4::name where v2: "atom v2 \ (t,X1,X3,X4)" and v3: "atom v3 \ (t,t',X1,v2,X4)" and v4: "atom v4 \ (t,t',u,X1,v2,v3)" by (metis obtain_fresh) have "{} \ (Var X1 EQ Var X2 AND Var X3 EQ Var X4) IMP (Var X1 IN Var X3 IMP Var X2 IN Var X4)" by (metis mem_cong_ax_def equality_axioms_def insert_iff Eq) hence "{} \ (Var X1 EQ Var v2 AND Var X3 EQ Var X4) IMP (Var X1 IN Var X3 IMP Var v2 IN Var X4)" by (drule_tac i=X2 and x="Var v2" in Subst) simp_all hence "{} \ (Var X1 EQ Var v2 AND Var v3 EQ Var X4) IMP (Var X1 IN Var v3 IMP Var v2 IN Var X4)" using v2 by (drule_tac i=X3 and x="Var v3" in Subst) simp_all hence "{} \ (Var X1 EQ Var v2 AND Var v3 EQ Var v4) IMP (Var X1 IN Var v3 IMP Var v2 IN Var v4)" using v2 v3 by (drule_tac i=X4 and x="Var v4" in Subst) simp_all hence "{} \ (t EQ Var v2 AND Var v3 EQ Var v4) IMP (t IN Var v3 IMP Var v2 IN Var v4)" using v2 v3 v4 by (drule_tac i=X1 and x=t in Subst) simp_all hence "{} \ (t EQ t' AND Var v3 EQ Var v4) IMP (t IN Var v3 IMP t' IN Var v4)" using v2 v3 v4 by (drule_tac i=v2 and x=t' in Subst) simp_all hence "{} \ (t EQ t' AND u EQ Var v4) IMP (t IN u IMP t' IN Var v4)" using v3 v4 by (drule_tac i=v3 and x=u in Subst) simp_all thus ?thesis using v4 by (drule_tac i=v4 and x=u' in Subst) simp_all qed lemma Mem_cong: assumes "H \ t EQ t'" "H \ u EQ u'" shows "H \ t IN u IFF t' IN u'" proof - { fix t t' u u' have cong: "{t EQ t', u EQ u'} \ t IN u IMP t' IN u'" by (metis AssumeH(2) Conj_I MP_null Mem_cong1 insert_commute) } thus ?thesis by (metis Iff_def Conj_I cut2 assms Sym) qed subsection\The congruence properties for @{term Eats} and @{term HPair}\ lemma Eats_cong1: "{} \ (t EQ t' AND u EQ u') IMP (Eats t u EQ Eats t' u')" proof - obtain v2::name and v3::name and v4::name where v2: "atom v2 \ (t,X1,X3,X4)" and v3: "atom v3 \ (t,t',X1,v2,X4)" and v4: "atom v4 \ (t,t',u,X1,v2,v3)" by (metis obtain_fresh) have "{} \ (Var X1 EQ Var X2 AND Var X3 EQ Var X4) IMP (Eats (Var X1) (Var X3) EQ Eats (Var X2) (Var X4))" by (metis eats_cong_ax_def equality_axioms_def insert_iff Eq) hence "{} \ (Var X1 EQ Var v2 AND Var X3 EQ Var X4) IMP (Eats (Var X1) (Var X3) EQ Eats (Var v2) (Var X4))" by (drule_tac i=X2 and x="Var v2" in Subst) simp_all hence "{} \ (Var X1 EQ Var v2 AND Var v3 EQ Var X4) IMP (Eats (Var X1) (Var v3) EQ Eats (Var v2) (Var X4))" using v2 by (drule_tac i=X3 and x="Var v3" in Subst) simp_all hence "{} \ (Var X1 EQ Var v2 AND Var v3 EQ Var v4) IMP (Eats (Var X1) (Var v3) EQ Eats (Var v2) (Var v4))" using v2 v3 by (drule_tac i=X4 and x="Var v4" in Subst) simp_all hence "{} \ (t EQ Var v2 AND Var v3 EQ Var v4) IMP (Eats t (Var v3) EQ Eats (Var v2) (Var v4))" using v2 v3 v4 by (drule_tac i=X1 and x=t in Subst) simp_all hence "{} \ (t EQ t' AND Var v3 EQ Var v4) IMP (Eats t (Var v3) EQ Eats t' (Var v4))" using v2 v3 v4 by (drule_tac i=v2 and x=t' in Subst) simp_all hence "{} \ (t EQ t' AND u EQ Var v4) IMP (Eats t u EQ Eats t' (Var v4))" using v3 v4 by (drule_tac i=v3 and x=u in Subst) simp_all thus ?thesis using v4 by (drule_tac i=v4 and x=u' in Subst) simp_all qed lemma Eats_cong: "\H \ t EQ t'; H \ u EQ u'\ \ H \ Eats t u EQ Eats t' u'" by (metis Conj_I anti_deduction Eats_cong1 cut1) lemma HPair_cong: "\H \ t EQ t'; H \ u EQ u'\ \ H \ HPair t u EQ HPair t' u'" by (metis HPair_def Eats_cong Refl) lemma SUCC_cong: "H \ t EQ t' \ H \ SUCC t EQ SUCC t'" by (metis Eats_cong SUCC_def) subsection\Substitution for Equalities\ lemma Eq_subst_tm_Iff: "{t EQ u} \ subst i t tm EQ subst i u tm" by (induct tm rule: tm.induct) (auto simp: Eats_cong) lemma Eq_subst_fm_Iff: "insert (t EQ u) H \ A(i::=t) IFF A(i::=u)" proof - have "{ t EQ u } \ A(i::=t) IFF A(i::=u)" by (nominal_induct A avoiding: i t u rule: fm.strong_induct) (auto simp: Disj_cong Neg_cong Ex_cong Mem_cong Eq_cong Eq_subst_tm_Iff) thus ?thesis by (metis Assume cut1) qed lemma Var_Eq_subst_Iff: "insert (Var i EQ t) H \ A(i::=t) IFF A" by (metis Eq_subst_fm_Iff Iff_sym subst_fm_id) lemma Var_Eq_imp_subst_Iff: "H \ Var i EQ t \ H \ A(i::=t) IFF A" by (metis Var_Eq_subst_Iff cut_same) subsection\Congruence Rules for Predicates\ lemma P1_cong: fixes tms :: "tm list" assumes "\i t x. atom i \ tms \ (P t)(i::=x) = P (subst i x t)" and "H \ x EQ x'" shows "H \ P x IFF P x'" proof - obtain i::name where i: "atom i \ tms" by (metis obtain_fresh) have "insert (x EQ x') H \ (P (Var i))(i::=x) IFF (P(Var i))(i::=x')" by (rule Eq_subst_fm_Iff) thus ?thesis using assms i by (metis cut_same subst.simps(2)) qed lemma P2_cong: fixes tms :: "tm list" assumes sub: "\i t u x. atom i \ tms \ (P t u)(i::=x) = P (subst i x t) (subst i x u)" and eq: "H \ x EQ x'" "H \ y EQ y'" shows "H \ P x y IFF P x' y'" proof - have yy': "{ y EQ y' } \ P x' y IFF P x' y'" by (rule P1_cong [where tms="[y,x']@tms"]) (auto simp: fresh_Cons sub) have "{ x EQ x' } \ P x y IFF P x' y" by (rule P1_cong [where tms="[y,x']@tms"]) (auto simp: fresh_Cons sub) hence "{x EQ x', y EQ y'} \ P x y IFF P x' y'" by (metis Assume Iff_trans cut1 rotate2 yy') thus ?thesis by (metis cut2 eq) qed lemma P3_cong: fixes tms :: "tm list" assumes sub: "\i t u v x. atom i \ tms \ (P t u v)(i::=x) = P (subst i x t) (subst i x u) (subst i x v)" and eq: "H \ x EQ x'" "H \ y EQ y'" "H \ z EQ z'" shows "H \ P x y z IFF P x' y' z'" proof - obtain i::name where i: "atom i \ (z,z',y,y',x,x')" by (metis obtain_fresh) have tl: "{ y EQ y', z EQ z' } \ P x' y z IFF P x' y' z'" by (rule P2_cong [where tms="[z,z',y,y',x,x']@tms"]) (auto simp: fresh_Cons sub) have hd: "{ x EQ x' } \ P x y z IFF P x' y z" by (rule P1_cong [where tms="[z,y,x']@tms"]) (auto simp: fresh_Cons sub) have "{x EQ x', y EQ y', z EQ z'} \ P x y z IFF P x' y' z'" by (metis Assume thin1 hd [THEN cut1] tl Iff_trans) thus ?thesis by (rule cut3) (rule eq)+ qed lemma P4_cong: fixes tms :: "tm list" assumes sub: "\i t1 t2 t3 t4 x. atom i \ tms \ (P t1 t2 t3 t4)(i::=x) = P (subst i x t1) (subst i x t2) (subst i x t3) (subst i x t4)" and eq: "H \ x1 EQ x1'" "H \ x2 EQ x2'" "H \ x3 EQ x3'" "H \ x4 EQ x4'" shows "H \ P x1 x2 x3 x4 IFF P x1' x2' x3' x4'" proof - obtain i::name where i: "atom i \ (x4,x4',x3,x3',x2,x2',x1,x1')" by (metis obtain_fresh) have tl: "{ x2 EQ x2', x3 EQ x3', x4 EQ x4' } \ P x1' x2 x3 x4 IFF P x1' x2' x3' x4'" by (rule P3_cong [where tms="[x4,x4',x3,x3',x2,x2',x1,x1']@tms"]) (auto simp: fresh_Cons sub) have hd: "{ x1 EQ x1' } \ P x1 x2 x3 x4 IFF P x1' x2 x3 x4" by (auto simp: fresh_Cons sub intro!: P1_cong [where tms="[x4,x3,x2,x1']@tms"]) have "{x1 EQ x1', x2 EQ x2', x3 EQ x3', x4 EQ x4'} \ P x1 x2 x3 x4 IFF P x1' x2' x3' x4'" by (metis Assume thin1 hd [THEN cut1] tl Iff_trans) thus ?thesis by (rule cut4) (rule eq)+ qed section\Zero and Falsity\ lemma Mem_Zero_iff: assumes "atom i \ t" shows "H \ (t EQ Zero) IFF (All i (Neg ((Var i) IN t)))" proof - obtain i'::name where i': "atom i' \ (t, X0, X1, i)" by (rule obtain_fresh) have "{} \ ((Var X0) EQ Zero) IFF (All X1 (Neg ((Var X1) IN (Var X0))))" by (simp add: HF HF_axioms_def HF1_def) then have "{} \ (((Var X0) EQ Zero) IFF (All X1 (Neg ((Var X1) IN (Var X0)))))(X0 ::= t)" by (rule Subst) simp hence "{} \ (t EQ Zero) IFF (All i' (Neg ((Var i') IN t)))" using i' by simp also have "... = (FRESH i'. (t EQ Zero) IFF (All i' (Neg ((Var i') IN t))))" using i' by simp also have "... = (t EQ Zero) IFF (All i (Neg ((Var i) IN t)))" using assms by simp finally show ?thesis by (metis empty_subsetI thin) qed lemma Mem_Zero_E [intro!]: "insert (x IN Zero) H \ A" proof - obtain i::name where "atom i \ Zero" by (rule obtain_fresh) hence "{} \ All i (Neg ((Var i) IN Zero))" by (metis Mem_Zero_iff Iff_MP_same Refl) hence "{} \ Neg (x IN Zero)" by (drule_tac x=x in All_D) simp thus ?thesis by (metis Contrapos2 Hyp Imp_triv_I MP_same empty_subsetI insertI1 thin) qed declare Mem_Zero_E [THEN rotate2, intro!] declare Mem_Zero_E [THEN rotate3, intro!] declare Mem_Zero_E [THEN rotate4, intro!] declare Mem_Zero_E [THEN rotate5, intro!] declare Mem_Zero_E [THEN rotate6, intro!] declare Mem_Zero_E [THEN rotate7, intro!] declare Mem_Zero_E [THEN rotate8, intro!] subsection\The Formula @{term Fls}\ definition Fls where "Fls \ Zero IN Zero" lemma Fls_eqvt [eqvt]: "(p \ Fls) = Fls" by (simp add: Fls_def) lemma Fls_fresh [simp]: "a \ Fls" by (simp add: Fls_def) lemma Neg_I [intro!]: "insert A H \ Fls \ H \ Neg A" unfolding Fls_def by (rule Neg_I0) (metis Mem_Zero_E cut_same) lemma Neg_E [intro!]: "H \ A \ insert (Neg A) H \ Fls" by (rule ContraProve) declare Neg_E [THEN rotate2, intro!] declare Neg_E [THEN rotate3, intro!] declare Neg_E [THEN rotate4, intro!] declare Neg_E [THEN rotate5, intro!] declare Neg_E [THEN rotate6, intro!] declare Neg_E [THEN rotate7, intro!] declare Neg_E [THEN rotate8, intro!] text\We need these because Neg (A IMP B) doesn't have to be syntactically a conjunction.\ lemma Neg_Imp_I [intro!]: "H \ A \ insert B H \ Fls \ H \ Neg (A IMP B)" by (metis NegNeg_I Neg_Disj_I Neg_I) lemma Neg_Imp_E [intro!]: "insert (Neg B) (insert A H) \ C \ insert (Neg (A IMP B)) H \ C" apply (rule cut_same [where A=A]) apply (metis Assume Disj_I1 NegNeg_D Neg_mono) apply (metis Swap Imp_I rotate2 thin1) done declare Neg_Imp_E [THEN rotate2, intro!] declare Neg_Imp_E [THEN rotate3, intro!] declare Neg_Imp_E [THEN rotate4, intro!] declare Neg_Imp_E [THEN rotate5, intro!] declare Neg_Imp_E [THEN rotate6, intro!] declare Neg_Imp_E [THEN rotate7, intro!] declare Neg_Imp_E [THEN rotate8, intro!] lemma Fls_E [intro!]: "insert Fls H \ A" by (metis Mem_Zero_E Fls_def) declare Fls_E [THEN rotate2, intro!] declare Fls_E [THEN rotate3, intro!] declare Fls_E [THEN rotate4, intro!] declare Fls_E [THEN rotate5, intro!] declare Fls_E [THEN rotate6, intro!] declare Fls_E [THEN rotate7, intro!] declare Fls_E [THEN rotate8, intro!] lemma truth_provable: "H \ (Neg Fls)" by (metis Fls_E Neg_I) lemma ExFalso: "H \ Fls \ H \ A" by (metis Neg_D truth_provable) subsection\More properties of @{term Zero}\ lemma Eq_Zero_D: assumes "H \ t EQ Zero" "H \ u IN t" shows "H \ A" proof - obtain i::name where i: "atom i \ t" by (rule obtain_fresh) with assms have an: "H \ (All i (Neg ((Var i) IN t)))" by (metis Iff_MP_same Mem_Zero_iff) have "H \ Neg (u IN t)" using All_D [OF an, of u] i by simp thus ?thesis using assms by (metis Neg_D) qed lemma Eq_Zero_thm: assumes "atom i \ t" shows "{All i (Neg ((Var i) IN t))} \ t EQ Zero" by (metis Assume Iff_MP2_same Mem_Zero_iff assms) lemma Eq_Zero_I: assumes insi: "insert ((Var i) IN t) H \ Fls" and i1: "atom i \ t" and i2: "\B \ H. atom i \ B" shows "H \ t EQ Zero" proof - have "H \ All i (Neg ((Var i) IN t))" by (metis All_I Neg_I i2 insi) thus ?thesis by (metis cut_same cut [OF Eq_Zero_thm [OF i1] Hyp] insertCI insert_is_Un) qed subsection\Basic properties of @{term Eats}\ lemma Eq_Eats_iff: assumes "atom i \ (z,t,u)" shows "H \ (z EQ Eats t u) IFF (All i (Var i IN z IFF Var i IN t OR Var i EQ u))" proof - obtain v1::name and v2::name and i'::name where v1: "atom v1 \ (z,X0,X2,X3)" and v2: "atom v2 \ (t,z,X0,v1,X3)" and i': "atom i' \ (t,u,z,X0,v1,v2,X3)" by (metis obtain_fresh) have "{} \ ((Var X0) EQ (Eats (Var X1) (Var X2))) IFF (All X3 (Var X3 IN Var X0 IFF Var X3 IN Var X1 OR Var X3 EQ Var X2))" by (simp add: HF HF_axioms_def HF2_def) hence "{} \ ((Var X0) EQ (Eats (Var X1) (Var X2))) IFF (All X3 (Var X3 IN Var X0 IFF Var X3 IN Var X1 OR Var X3 EQ Var X2))" by (drule_tac i=X0 and x="Var X0" in Subst) simp_all hence "{} \ ((Var X0) EQ (Eats (Var v1) (Var X2))) IFF (All X3 (Var X3 IN Var X0 IFF Var X3 IN Var v1 OR Var X3 EQ Var X2))" using v1 by (drule_tac i=X1 and x="Var v1" in Subst) simp_all hence "{} \ ((Var X0) EQ (Eats (Var v1) (Var v2))) IFF (All X3 (Var X3 IN Var X0 IFF Var X3 IN Var v1 OR Var X3 EQ Var v2))" using v1 v2 by (drule_tac i=X2 and x="Var v2" in Subst) simp_all hence "{} \ (((Var X0) EQ (Eats (Var v1) (Var v2))) IFF (All X3 (Var X3 IN Var X0 IFF Var X3 IN Var v1 OR Var X3 EQ Var v2)))(X0 ::= z)" by (rule Subst) simp hence "{} \ ((z EQ (Eats (Var v1) (Var v2))) IFF (All i' (Var i' IN z IFF Var i' IN Var v1 OR Var i' EQ Var v2)))" using v1 v2 i' by (simp add: Conj_def Iff_def) hence "{} \ (z EQ (Eats t (Var v2))) IFF (All i' (Var i' IN z IFF Var i' IN t OR Var i' EQ Var v2))" using v1 v2 i' by (drule_tac i=v1 and x=t in Subst) simp_all hence "{} \ (z EQ Eats t u) IFF (All i' (Var i' IN z IFF Var i' IN t OR Var i' EQ u))" using v1 v2 i' by (drule_tac i=v2 and x=u in Subst) simp_all also have "... = (FRESH i'. (z EQ Eats t u) IFF (All i' (Var i' IN z IFF Var i' IN t OR Var i' EQ u)))" using i' by simp also have "... = (z EQ Eats t u) IFF (All i (Var i IN z IFF Var i IN t OR Var i EQ u))" using assms i' by simp finally show ?thesis by (rule thin0) qed lemma Eq_Eats_I: "H \ All i (Var i IN z IFF Var i IN t OR Var i EQ u) \ atom i \ (z,t,u) \ H \ z EQ Eats t u" by (metis Iff_MP2_same Eq_Eats_iff) lemma Mem_Eats_Iff: "H \ x IN (Eats t u) IFF x IN t OR x EQ u" proof - obtain i::name where "atom i \ (Eats t u, t, u)" by (rule obtain_fresh) thus ?thesis using Iff_MP_same [OF Eq_Eats_iff, THEN All_D] by auto qed lemma Mem_Eats_I1: "H \ u IN t \ H \ u IN Eats t z" by (metis Disj_I1 Iff_MP2_same Mem_Eats_Iff) lemma Mem_Eats_I2: "H \ u EQ z \ H \ u IN Eats t z" by (metis Disj_I2 Iff_MP2_same Mem_Eats_Iff) lemma Mem_Eats_E: assumes A: "insert (u IN t) H \ C" and B: "insert (u EQ z) H \ C" shows "insert (u IN Eats t z) H \ C" by (rule Mem_Eats_Iff [of _ u t z, THEN Iff_MP_left']) (metis A B Disj_E) lemmas Mem_Eats_EH = Mem_Eats_E Mem_Eats_E [THEN rotate2] Mem_Eats_E [THEN rotate3] Mem_Eats_E [THEN rotate4] Mem_Eats_E [THEN rotate5] Mem_Eats_E [THEN rotate6] Mem_Eats_E [THEN rotate7] Mem_Eats_E [THEN rotate8] declare Mem_Eats_EH [intro!] lemma Mem_SUCC_I1: "H \ u IN t \ H \ u IN SUCC t" by (metis Mem_Eats_I1 SUCC_def) lemma Mem_SUCC_I2: "H \ u EQ t \ H \ u IN SUCC t" by (metis Mem_Eats_I2 SUCC_def) lemma Mem_SUCC_Refl [simp]: "H \ k IN SUCC k" by (metis Mem_SUCC_I2 Refl) lemma Mem_SUCC_E: assumes "insert (u IN t) H \ C" "insert (u EQ t) H \ C" shows "insert (u IN SUCC t) H \ C" by (metis assms Mem_Eats_E SUCC_def) lemmas Mem_SUCC_EH = Mem_SUCC_E Mem_SUCC_E [THEN rotate2] Mem_SUCC_E [THEN rotate3] Mem_SUCC_E [THEN rotate4] Mem_SUCC_E [THEN rotate5] Mem_SUCC_E [THEN rotate6] Mem_SUCC_E [THEN rotate7] Mem_SUCC_E [THEN rotate8] lemma Eats_EQ_Zero_E: "insert (Eats t u EQ Zero) H \ A" by (metis Assume Eq_Zero_D Mem_Eats_I2 Refl) lemmas Eats_EQ_Zero_EH = Eats_EQ_Zero_E Eats_EQ_Zero_E [THEN rotate2] Eats_EQ_Zero_E [THEN rotate3] Eats_EQ_Zero_E [THEN rotate4] Eats_EQ_Zero_E [THEN rotate5] Eats_EQ_Zero_E [THEN rotate6] Eats_EQ_Zero_E [THEN rotate7] Eats_EQ_Zero_E [THEN rotate8] declare Eats_EQ_Zero_EH [intro!] lemma Eats_EQ_Zero_E2: "insert (Zero EQ Eats t u) H \ A" by (metis Eats_EQ_Zero_E Sym_L) lemmas Eats_EQ_Zero_E2H = Eats_EQ_Zero_E2 Eats_EQ_Zero_E2 [THEN rotate2] Eats_EQ_Zero_E2 [THEN rotate3] Eats_EQ_Zero_E2 [THEN rotate4] Eats_EQ_Zero_E2 [THEN rotate5] Eats_EQ_Zero_E2 [THEN rotate6] Eats_EQ_Zero_E2 [THEN rotate7] Eats_EQ_Zero_E2 [THEN rotate8] declare Eats_EQ_Zero_E2H [intro!] section\Bounded Quantification involving @{term Eats}\ lemma All2_cong: "H \ t EQ t' \ H \ A IFF A' \ \C \ H. atom i \ C \ H \ (All2 i t A) IFF (All2 i t' A')" by (metis All_cong Imp_cong Mem_cong Refl) lemma All2_Zero_E [intro!]: "H \ B \ insert (All2 i Zero A) H \ B" by (rule thin1) lemma All2_Eats_I_D: "atom i \ (t,u) \ { All2 i t A, A(i::=u)} \ (All2 i (Eats t u) A)" apply (auto, auto intro!: Ex_I [where x="Var i"]) apply (metis Assume thin1 Var_Eq_subst_Iff [THEN Iff_MP_same]) done lemma All2_Eats_I: "\atom i \ (t,u); H \ All2 i t A; H \ A(i::=u)\ \ H \ (All2 i (Eats t u) A)" by (rule cut2 [OF All2_Eats_I_D], auto) lemma All2_Eats_E1: "\atom i \ (t,u); \C \ H. atom i \ C\ \ insert (All2 i (Eats t u) A) H \ All2 i t A" by auto (metis Assume Ex_I Imp_E Mem_Eats_I1 Neg_mono subst_fm_id) lemma All2_Eats_E2: "\atom i \ (t,u); \C \ H. atom i \ C\ \ insert (All2 i (Eats t u) A) H \ A(i::=u)" by (rule All_E [where x=u]) (auto intro: ContraProve Mem_Eats_I2) lemma All2_Eats_E: assumes i: "atom i \ (t,u)" and B: "insert (All2 i t A) (insert (A(i::=u)) H) \ B" shows "insert (All2 i (Eats t u) A) H \ B" using i apply (rule cut_thin [OF All2_Eats_E2, where HB = "insert (All2 i (Eats t u) A) H"], auto) apply (rule cut_thin [OF All2_Eats_E1 B], auto) done lemma All2_SUCC_I: "atom i \ t \ H \ All2 i t A \ H \ A(i::=t) \ H \ (All2 i (SUCC t) A)" by (simp add: SUCC_def All2_Eats_I) lemma All2_SUCC_E: assumes "atom i \ t" and "insert (All2 i t A) (insert (A(i::=t)) H) \ B" shows "insert (All2 i (SUCC t) A) H \ B" by (simp add: SUCC_def All2_Eats_E assms) lemma All2_SUCC_E': assumes "H \ u EQ SUCC t" and "atom i \ t" "\C \ H. atom i \ C" and "insert (All2 i t A) (insert (A(i::=t)) H) \ B" shows "insert (All2 i u A) H \ B" by (metis All2_SUCC_E Iff_MP_left' Iff_refl All2_cong assms) section\Induction\ lemma Ind: assumes j: "atom (j::name) \ (i,A)" and prems: "H \ A(i::=Zero)" "H \ All i (All j (A IMP (A(i::= Var j) IMP A(i::= Eats(Var i)(Var j)))))" shows "H \ A" proof - have "{A(i::=Zero), All i (All j (A IMP (A(i::= Var j) IMP A(i::= Eats(Var i)(Var j)))))} \ All i A" by (metis j hfthm.Ind ind anti_deduction insert_commute) hence "H \ (All i A)" by (metis cut2 prems) thus ?thesis by (metis All_E' Assume subst_fm_id) qed end diff --git a/thys/Refine_Imperative_HOL/Sepref_Frame.thy b/thys/Refine_Imperative_HOL/Sepref_Frame.thy --- a/thys/Refine_Imperative_HOL/Sepref_Frame.thy +++ b/thys/Refine_Imperative_HOL/Sepref_Frame.thy @@ -1,541 +1,541 @@ section \Frame Inference\ theory Sepref_Frame imports Sepref_Basic Sepref_Constraints begin text \ In this theory, we provide a specific frame inference tactic for Sepref. The first tactic, \frame_tac\, is a standard frame inference tactic, based on the assumption that only @{const hn_ctxt}-assertions need to be matched. The second tactic, \merge_tac\, resolves entailments of the form \F1 \\<^sub>A F2 \\<^sub>t ?F\ that occur during translation of if and case statements. It synthesizes a new frame ?F, where refinements of variables with equal refinements in \F1\ and \F2\ are preserved, and the others are set to @{const hn_invalid}. \ definition mismatch_assn :: "('a \ 'c \ assn) \ ('a \ 'c \ assn) \ 'a \ 'c \ assn" where "mismatch_assn R1 R2 x y \ R1 x y \\<^sub>A R2 x y" abbreviation "hn_mismatch R1 R2 \ hn_ctxt (mismatch_assn R1 R2)" lemma recover_pure_aux: "CONSTRAINT is_pure R \ hn_invalid R x y \\<^sub>t hn_ctxt R x y" by (auto simp: is_pure_conv invalid_pure_recover hn_ctxt_def) lemma frame_thms: "P \\<^sub>t P" "P\\<^sub>tP' \ F\\<^sub>tF' \ F*P \\<^sub>t F'*P'" "hn_ctxt R x y \\<^sub>t hn_invalid R x y" "hn_ctxt R x y \\<^sub>t hn_ctxt (\_ _. true) x y" "CONSTRAINT is_pure R \ hn_invalid R x y \\<^sub>t hn_ctxt R x y" apply - applyS simp applyS (rule entt_star_mono; assumption) subgoal apply (simp add: hn_ctxt_def) apply (rule enttI) apply (rule ent_trans[OF invalidate[of R]]) by solve_entails applyS (sep_auto simp: hn_ctxt_def) applyS (erule recover_pure_aux) done named_theorems_rev sepref_frame_match_rules \Sepref: Additional frame rules\ text \Rules to discharge unmatched stuff\ (*lemma frame_rem_thms: "P \\<^sub>t P" "P \\<^sub>t emp" by sep_auto+ *) lemma frame_rem1: "P\\<^sub>tP" by simp lemma frame_rem2: "F \\<^sub>t F' \ F * hn_ctxt A x y \\<^sub>t F' * hn_ctxt A x y" apply (rule entt_star_mono) by auto lemma frame_rem3: "F \\<^sub>t F' \ F * hn_ctxt A x y \\<^sub>t F'" using frame_thms(2) by fastforce lemma frame_rem4: "P \\<^sub>t emp" by simp lemmas frame_rem_thms = frame_rem1 frame_rem2 frame_rem3 frame_rem4 named_theorems_rev sepref_frame_rem_rules \Sepref: Additional rules to resolve remainder of frame-pairing\ lemma ent_disj_star_mono: "\ A \\<^sub>A C \\<^sub>A E; B \\<^sub>A D \\<^sub>A F \ \ A*B \\<^sub>A C*D \\<^sub>A E*F" by (metis ent_disjI1 ent_disjI2 ent_disjE ent_star_mono) lemma entt_disj_star_mono: "\ A \\<^sub>A C \\<^sub>t E; B \\<^sub>A D \\<^sub>t F \ \ A*B \\<^sub>A C*D \\<^sub>t E*F" proof - assume a1: "A \\<^sub>A C \\<^sub>t E" assume "B \\<^sub>A D \\<^sub>t F" then have "A * B \\<^sub>A C * D \\<^sub>A true * E * (true * F)" using a1 by (simp add: ent_disj_star_mono enttD) then show ?thesis by (metis (no_types) assn_times_comm enttI merge_true_star_ctx star_aci(3)) qed lemma hn_merge1: (*"emp \\<^sub>A emp \\<^sub>A emp"*) "F \\<^sub>A F \\<^sub>t F" "\ hn_ctxt R1 x x' \\<^sub>A hn_ctxt R2 x x' \\<^sub>t hn_ctxt R x x'; Fl \\<^sub>A Fr \\<^sub>t F \ \ Fl * hn_ctxt R1 x x' \\<^sub>A Fr * hn_ctxt R2 x x' \\<^sub>t F * hn_ctxt R x x'" apply simp by (rule entt_disj_star_mono; simp) lemma hn_merge2: "hn_invalid R x x' \\<^sub>A hn_ctxt R x x' \\<^sub>t hn_invalid R x x'" "hn_ctxt R x x' \\<^sub>A hn_invalid R x x' \\<^sub>t hn_invalid R x x'" by (sep_auto eintros: invalidate ent_disjE intro!: ent_imp_entt simp: hn_ctxt_def)+ lemma invalid_assn_mono: "hn_ctxt A x y \\<^sub>t hn_ctxt B x y \ hn_invalid A x y \\<^sub>t hn_invalid B x y" by (clarsimp simp: invalid_assn_def entailst_def entails_def hn_ctxt_def) (force simp: mod_star_conv) lemma hn_merge3: (* Not used *) "\NO_MATCH (hn_invalid XX) R2; hn_ctxt R1 x x' \\<^sub>A hn_ctxt R2 x x' \\<^sub>t hn_ctxt Rm x x'\ \ hn_invalid R1 x x' \\<^sub>A hn_ctxt R2 x x' \\<^sub>t hn_invalid Rm x x'" "\NO_MATCH (hn_invalid XX) R1; hn_ctxt R1 x x' \\<^sub>A hn_ctxt R2 x x' \\<^sub>t hn_ctxt Rm x x'\ \ hn_ctxt R1 x x' \\<^sub>A hn_invalid R2 x x' \\<^sub>t hn_invalid Rm x x'" apply (meson entt_disjD1 entt_disjD2 entt_disjE entt_trans frame_thms(3) invalid_assn_mono) apply (meson entt_disjD1 entt_disjD2 entt_disjE entt_trans frame_thms(3) invalid_assn_mono) done lemmas merge_thms = hn_merge1 hn_merge2 named_theorems sepref_frame_merge_rules \Sepref: Additional merge rules\ lemma hn_merge_mismatch: "hn_ctxt R1 x x' \\<^sub>A hn_ctxt R2 x x' \\<^sub>t hn_mismatch R1 R2 x x'" by (sep_auto simp: hn_ctxt_def mismatch_assn_def) lemma is_merge: "P1\\<^sub>AP2\\<^sub>tP \ P1\\<^sub>AP2\\<^sub>tP" . lemma merge_mono: "\A\\<^sub>tA'; B\\<^sub>tB'; A'\\<^sub>AB' \\<^sub>t C\ \ A\\<^sub>AB \\<^sub>t C" by (meson entt_disjE entt_disjI1_direct entt_disjI2_direct entt_trans) text \Apply forward rule on left or right side of merge\ lemma gen_merge_cons1: "\A\\<^sub>tA'; A'\\<^sub>AB \\<^sub>t C\ \ A\\<^sub>AB \\<^sub>t C" by (meson merge_mono entt_refl) lemma gen_merge_cons2: "\B\\<^sub>tB'; A\\<^sub>AB' \\<^sub>t C\ \ A\\<^sub>AB \\<^sub>t C" by (meson merge_mono entt_refl) lemmas gen_merge_cons = gen_merge_cons1 gen_merge_cons2 text \These rules are applied to recover pure values that have been destroyed by rule application\ definition "RECOVER_PURE P Q \ P \\<^sub>t Q" lemma recover_pure: "RECOVER_PURE emp emp" "\RECOVER_PURE P2 Q2; RECOVER_PURE P1 Q1\ \ RECOVER_PURE (P1*P2) (Q1*Q2)" "CONSTRAINT is_pure R \ RECOVER_PURE (hn_invalid R x y) (hn_ctxt R x y)" "RECOVER_PURE (hn_ctxt R x y) (hn_ctxt R x y)" unfolding RECOVER_PURE_def subgoal by sep_auto subgoal by (drule (1) entt_star_mono) subgoal by (rule recover_pure_aux) subgoal by sep_auto done lemma recover_pure_triv: "RECOVER_PURE P P" unfolding RECOVER_PURE_def by sep_auto text \Weakening the postcondition by converting @{const invalid_assn} to @{term "\_ _. true"}\ definition "WEAKEN_HNR_POST \ \' \'' \ (\h. h\\) \ (\'' \\<^sub>t \')" lemma weaken_hnr_postI: assumes "WEAKEN_HNR_POST \ \'' \'" assumes "hn_refine \ c \' R a" shows "hn_refine \ c \'' R a" apply (rule hn_refine_preI) apply (rule hn_refine_cons_post) apply (rule assms) using assms(1) unfolding WEAKEN_HNR_POST_def by blast lemma weaken_hnr_post_triv: "WEAKEN_HNR_POST \ P P" unfolding WEAKEN_HNR_POST_def by sep_auto lemma weaken_hnr_post: "\WEAKEN_HNR_POST \ P P'; WEAKEN_HNR_POST \' Q Q'\ \ WEAKEN_HNR_POST (\*\') (P*Q) (P'*Q')" "WEAKEN_HNR_POST (hn_ctxt R x y) (hn_ctxt R x y) (hn_ctxt R x y)" "WEAKEN_HNR_POST (hn_ctxt R x y) (hn_invalid R x y) (hn_ctxt (\_ _. true) x y)" proof (goal_cases) case 1 thus ?case unfolding WEAKEN_HNR_POST_def apply clarsimp apply (rule entt_star_mono) by (auto simp: mod_star_conv) next case 2 thus ?case by (rule weaken_hnr_post_triv) next case 3 thus ?case unfolding WEAKEN_HNR_POST_def by (sep_auto simp: invalid_assn_def hn_ctxt_def) qed lemma reorder_enttI: assumes "A*true = C*true" assumes "B*true = D*true" shows "(A\\<^sub>tB) \ (C\\<^sub>tD)" apply (intro eq_reflection) unfolding entt_def_true by (simp add: assms) lemma merge_sat1: "(A\\<^sub>AA' \\<^sub>t Am) \ (A\\<^sub>AAm \\<^sub>t Am)" using entt_disjD1 entt_disjE by blast lemma merge_sat2: "(A\\<^sub>AA' \\<^sub>t Am) \ (Am\\<^sub>AA' \\<^sub>t Am)" using entt_disjD2 entt_disjE by blast ML \ signature SEPREF_FRAME = sig (* Check if subgoal is a frame obligation *) (*val is_frame : term -> bool *) (* Check if subgoal is a merge obligation *) val is_merge: term -> bool (* Perform frame inference *) val frame_tac: (Proof.context -> tactic') -> Proof.context -> tactic' (* Perform merging *) val merge_tac: (Proof.context -> tactic') -> Proof.context -> tactic' val frame_step_tac: (Proof.context -> tactic') -> bool -> Proof.context -> tactic' (* Reorder frame *) val prepare_frame_tac : Proof.context -> tactic' (* Solve a RECOVER_PURE goal, inserting constraints as necessary *) val recover_pure_tac: Proof.context -> tactic' (* Split precondition of hnr-goal into frame and arguments *) val align_goal_tac: Proof.context -> tactic' (* Normalize goal's precondition *) val norm_goal_pre_tac: Proof.context -> tactic' (* Rearrange precondition of hnr-term according to parameter order, normalize all relations *) val align_rl_conv: Proof.context -> conv (* Convert hn_invalid to \_ _. true in postcondition of hnr-goal. Makes proving the goal easier.*) val weaken_post_tac: Proof.context -> tactic' val add_normrel_eq : thm -> Context.generic -> Context.generic val del_normrel_eq : thm -> Context.generic -> Context.generic val get_normrel_eqs : Proof.context -> thm list val cfg_debug: bool Config.T val setup: theory -> theory end structure Sepref_Frame : SEPREF_FRAME = struct val cfg_debug = Attrib.setup_config_bool @{binding sepref_debug_frame} (K false) val DCONVERSION = Sepref_Debugging.DBG_CONVERSION cfg_debug val dbg_msg_tac = Sepref_Debugging.dbg_msg_tac cfg_debug structure normrel_eqs = Named_Thms ( val name = @{binding sepref_frame_normrel_eqs} val description = "Equations to normalize relations for frame matching" ) val add_normrel_eq = normrel_eqs.add_thm val del_normrel_eq = normrel_eqs.del_thm val get_normrel_eqs = normrel_eqs.get val mk_entailst = HOLogic.mk_binrel @{const_name "entailst"} local open Sepref_Basic Refine_Util Conv fun assn_ord p = case apply2 dest_hn_ctxt_opt p of (NONE,NONE) => EQUAL | (SOME _, NONE) => LESS | (NONE, SOME _) => GREATER | (SOME (_,a,_), SOME (_,a',_)) => Term_Ord.fast_term_ord (a,a') in fun reorder_ctxt_conv ctxt ct = let val cert = Thm.cterm_of ctxt val new_ct = Thm.term_of ct |> strip_star |> sort assn_ord |> list_star |> cert val thm = Goal.prove_internal ctxt [] (mk_cequals (ct,new_ct)) (fn _ => simp_tac (put_simpset HOL_basic_ss ctxt addsimps @{thms star_aci}) 1) in thm end fun prepare_fi_conv ctxt ct = case Thm.term_of ct of @{mpat "?P \\<^sub>t ?Q"} => let val cert = Thm.cterm_of ctxt (* Build table from abs-vars to ctxt *) val (Qm, Qum) = strip_star Q |> filter_out is_true |> List.partition is_hn_ctxt val Qtab = ( Qm |> map (fn x => (#2 (dest_hn_ctxt x),(NONE,x))) |> Termtab.make ) handle e as (Termtab.DUP _) => ( tracing ("Dup heap: " ^ @{make_string} ct); raise e) (* Go over entries in P and try to find a partner *) val (Qtab,Pum) = fold (fn a => fn (Qtab,Pum) => case dest_hn_ctxt_opt a of NONE => (Qtab,a::Pum) | SOME (_,p,_) => ( case Termtab.lookup Qtab p of SOME (NONE,tg) => (Termtab.update (p,(SOME a,tg)) Qtab, Pum) | _ => (Qtab,a::Pum) ) ) (strip_star P) (Qtab,[]) val Pum = filter_out is_true Pum (* Read out information from Qtab *) val (pairs,Qum2) = Termtab.dest Qtab |> map #2 |> List.partition (is_some o #1) |> apfst (map (apfst the)) |> apsnd (map #2) (* Build reordered terms: P' = fst pairs * Pum, Q' = snd pairs * (Qum2*Qum) *) val P' = mk_star (list_star (map fst pairs), list_star Pum) val Q' = mk_star (list_star (map snd pairs), list_star (Qum2@Qum)) val new_ct = mk_entailst (P', Q') |> cert val msg_tac = dbg_msg_tac (Sepref_Debugging.msg_allgoals "Solving frame permutation") ctxt 1 val tac = msg_tac THEN ALLGOALS (resolve_tac ctxt @{thms reorder_enttI}) THEN star_permute_tac ctxt val thm = Goal.prove_internal ctxt [] (mk_cequals (ct,new_ct)) (fn _ => tac) in thm end | _ => no_conv ct end fun is_merge @{mpat "Trueprop (_ \\<^sub>A _ \\<^sub>t _)"} = true | is_merge _ = false fun is_gen_frame @{mpat "Trueprop (_ \\<^sub>t _)"} = true | is_gen_frame _ = false fun prepare_frame_tac ctxt = let open Refine_Util Conv val frame_ss = put_simpset HOL_basic_ss ctxt addsimps @{thms mult_1_right[where 'a=assn] mult_1_left[where 'a=assn]} in CONVERSION Thm.eta_conversion THEN' (*CONCL_COND' is_frame THEN'*) simp_tac frame_ss THEN' CONVERSION (HOL_concl_conv (fn _ => prepare_fi_conv ctxt) ctxt) end local fun wrap_side_tac side_tac dbg tac = tac THEN_ALL_NEW_FWD ( CONCL_COND' is_gen_frame ORELSE' (if dbg then TRY_SOLVED' else SOLVED') side_tac ) in fun frame_step_tac side_tac dbg ctxt = let open Refine_Util Conv (* Constraint solving is built-in *) val side_tac = Sepref_Constraints.constraint_tac ctxt ORELSE' side_tac ctxt val frame_thms = @{thms frame_thms} @ Named_Theorems_Rev.get ctxt @{named_theorems_rev sepref_frame_match_rules} val merge_thms = @{thms merge_thms} @ Named_Theorems.get ctxt @{named_theorems sepref_frame_merge_rules} val ss = put_simpset HOL_basic_ss ctxt addsimps normrel_eqs.get ctxt fun frame_thm_tac dbg = wrap_side_tac side_tac dbg (resolve_tac ctxt frame_thms) fun merge_thm_tac dbg = wrap_side_tac side_tac dbg (resolve_tac ctxt merge_thms) fun thm_tac dbg = CONCL_COND' is_merge THEN_ELSE' (merge_thm_tac dbg, frame_thm_tac dbg) in full_simp_tac ss THEN' thm_tac dbg end end fun frame_loop_tac side_tac ctxt = let in TRY o ( REPEAT_ALL_NEW (DETERM o frame_step_tac side_tac false ctxt) ) end fun frame_tac side_tac ctxt = let open Refine_Util Conv val frame_rem_thms = @{thms frame_rem_thms} @ Named_Theorems_Rev.get ctxt @{named_theorems_rev sepref_frame_rem_rules} val solve_remainder_tac = TRY o REPEAT_ALL_NEW (DETERM o resolve_tac ctxt frame_rem_thms) in (prepare_frame_tac ctxt THEN' resolve_tac ctxt @{thms ent_star_mono entt_star_mono}) THEN_ALL_NEW_LIST [ frame_loop_tac side_tac ctxt, solve_remainder_tac ] end fun merge_tac side_tac ctxt = let open Refine_Util Conv val merge_conv = arg1_conv (binop_conv (reorder_ctxt_conv ctxt)) in CONVERSION Thm.eta_conversion THEN' CONCL_COND' is_merge THEN' simp_tac (put_simpset HOL_basic_ss ctxt addsimps @{thms star_aci}) THEN' CONVERSION (HOL_concl_conv (fn _ => merge_conv) ctxt) THEN' frame_loop_tac side_tac ctxt end val setup = normrel_eqs.setup local open Sepref_Basic fun is_invalid @{mpat "hn_invalid _ _ _ :: assn"} = true | is_invalid _ = false fun contains_invalid @{mpat "Trueprop (RECOVER_PURE ?Q _)"} = exists is_invalid (strip_star Q) | contains_invalid _ = false in fun recover_pure_tac ctxt = CONCL_COND' contains_invalid THEN_ELSE' ( REPEAT_ALL_NEW (DETERM o (resolve_tac ctxt @{thms recover_pure} ORELSE' Sepref_Constraints.constraint_tac ctxt)), resolve_tac ctxt @{thms recover_pure_triv} ) end local open Sepref_Basic Refine_Util datatype cte = Other of term | Hn of term * term * term fun dest_ctxt_elem @{mpat "hn_ctxt ?R ?a ?c"} = Hn (R,a,c) | dest_ctxt_elem t = Other t fun mk_ctxt_elem (Other t) = t | mk_ctxt_elem (Hn (R,a,c)) = @{mk_term "hn_ctxt ?R ?a ?c"} fun match x (Hn (_,y,_)) = x aconv y | match _ _ = false fun dest_with_frame (*ctxt*) _ t = let val (P,c,Q,R,a) = dest_hn_refine t val (_,(_,args)) = dest_hnr_absfun a val pre_ctes = strip_star P |> map dest_ctxt_elem val (pre_args,frame) = (case split_matching match args pre_ctes of NONE => raise TERM("align_conv: Could not match all arguments",[P,a]) | SOME x => x) in ((frame,pre_args),c,Q,R,a) end fun align_goal_conv_aux ctxt t = let val ((frame,pre_args),c,Q,R,a) = dest_with_frame ctxt t val P' = apply2 (list_star o map mk_ctxt_elem) (frame,pre_args) |> mk_star val t' = mk_hn_refine (P',c,Q,R,a) in t' end fun align_rl_conv_aux ctxt t = let val ((frame,pre_args),c,Q,R,a) = dest_with_frame ctxt t val _ = frame = [] orelse raise TERM ("align_rl_conv: Extra preconditions in rule",[t,list_star (map mk_ctxt_elem frame)]) val P' = list_star (map mk_ctxt_elem pre_args) val t' = mk_hn_refine (P',c,Q,R,a) in t' end fun normrel_conv ctxt = let val ss = put_simpset HOL_basic_ss ctxt addsimps normrel_eqs.get ctxt in Simplifier.rewrite ss end in - fun align_goal_conv ctxt = f_tac_conv ctxt (align_goal_conv_aux ctxt) (star_permute_tac ctxt) + fun align_goal_conv ctxt = f_tac_conv ctxt (align_goal_conv_aux ctxt) star_permute_tac fun norm_goal_pre_conv ctxt = let open Conv val nr_conv = normrel_conv ctxt in HOL_concl_conv (fn _ => hn_refine_conv nr_conv all_conv all_conv all_conv all_conv) ctxt end fun norm_goal_pre_tac ctxt = CONVERSION (norm_goal_pre_conv ctxt) fun align_rl_conv ctxt = let open Conv val nr_conv = normrel_conv ctxt in - HOL_concl_conv (fn ctxt => f_tac_conv ctxt (align_rl_conv_aux ctxt) (star_permute_tac ctxt)) ctxt + HOL_concl_conv (fn ctxt => f_tac_conv ctxt (align_rl_conv_aux ctxt) star_permute_tac) ctxt then_conv HOL_concl_conv (K (hn_refine_conv nr_conv all_conv nr_conv nr_conv all_conv)) ctxt end fun align_goal_tac ctxt = CONCL_COND' is_hn_refine_concl THEN' DCONVERSION ctxt (HOL_concl_conv align_goal_conv ctxt) end fun weaken_post_tac ctxt = TRADE (fn ctxt => resolve_tac ctxt @{thms weaken_hnr_postI} THEN' SOLVED' (REPEAT_ALL_NEW (DETERM o resolve_tac ctxt @{thms weaken_hnr_post weaken_hnr_post_triv})) ) ctxt end \ setup Sepref_Frame.setup method_setup weaken_hnr_post = \Scan.succeed (fn ctxt => SIMPLE_METHOD' (Sepref_Frame.weaken_post_tac ctxt))\ \Convert "hn_invalid" to "hn_ctxt (\_ _. true)" in postcondition of hn_refine goal\ (* TODO: Improper, modifies all h\_ premises that happen to be there. Use tagging to protect! *) method extract_hnr_invalids = ( rule hn_refine_preI, ((drule mod_starD hn_invalidI | elim conjE exE)+)? ) \ \Extract \hn_invalid _ _ _ = true\ preconditions from \hn_refine\ goal.\ lemmas [sepref_frame_normrel_eqs] = the_pure_pure pure_the_pure end diff --git a/thys/Refine_Imperative_HOL/Sepref_Id_Op.thy b/thys/Refine_Imperative_HOL/Sepref_Id_Op.thy --- a/thys/Refine_Imperative_HOL/Sepref_Id_Op.thy +++ b/thys/Refine_Imperative_HOL/Sepref_Id_Op.thy @@ -1,346 +1,346 @@ section \Operation Identification Phase\ theory Sepref_Id_Op imports Main Automatic_Refinement.Refine_Lib Automatic_Refinement.Autoref_Tagging "Lib/Named_Theorems_Rev" begin text \ The operation identification phase is adapted from the Autoref tool. The basic idea is to have a type system, which works on so called interface types (also called conceptual types). Each conceptual type denotes an abstract data type, e.g., set, map, priority queue. Each abstract operation, which must be a constant applied to its arguments, is assigned a conceptual type. Additionally, there is a set of {\emph pattern rewrite rules}, which are applied to subterms before type inference takes place, and which may be backtracked over. This way, encodings of abstract operations in Isabelle/HOL, like @{term [source] "\_. None"} for the empty map, or @{term [source] "fun_upd m k (Some v)"} for map update, can be rewritten to abstract operations, and get properly typed. \ subsection "Proper Protection of Term" text \ The following constants are meant to encode abstraction and application as proper HOL-constants, and thus avoid strange effects with HOL's higher-order unification heuristics and automatic beta and eta-contraction. The first step of operation identification is to protect the term by replacing all function applications and abstractions be the constants defined below. \ definition [simp]: "PROTECT2 x (y::prop) \ x" consts DUMMY :: "prop" abbreviation PROTECT2_syn ("'(#_#')") where "PROTECT2_syn t \ PROTECT2 t DUMMY" abbreviation (input)ABS2 :: "('a\'b)\'a\'b" (binder "\\<^sub>2" 10) where "ABS2 f \ (\x. PROTECT2 (f x) DUMMY)" lemma beta: "(\\<^sub>2x. f x)$x \ f x" by simp text \ Another version of @{const "APP"}. Treated like @{const APP} by our tool. Required to avoid infinite pattern rewriting in some cases, e.g., map-lookup. \ definition APP' (infixl "$''" 900) where [simp, autoref_tag_defs]: "f$'a \ f a" text \ Sometimes, whole terms should be protected from being processed by our tool. For example, our tool should not look into numerals. For this reason, the \PR_CONST\ tag indicates terms that our tool shall handle as atomic constants, an never look into them. The special form \UNPROTECT\ can be used inside pattern rewrite rules. It has the effect to revert the protection from its argument, and then wrap it into a \PR_CONST\. \ definition [simp, autoref_tag_defs]: "PR_CONST x \ x" \ \Tag to protect constant\ definition [simp, autoref_tag_defs]: "UNPROTECT x \ x" \ \Gets converted to @{term PR_CONST}, after unprotecting its content\ subsection \Operation Identification\ text \ Indicator predicate for conceptual typing of a constant \ definition intf_type :: "'a \ 'b itself \ bool" (infix "::\<^sub>i" 10) where [simp]: "c::\<^sub>iI \ True" lemma itypeI: "c::\<^sub>iI" by simp lemma itypeI': "intf_type c TYPE('T)" by (rule itypeI) lemma itype_self: "(c::'a) ::\<^sub>i TYPE('a)" by simp definition CTYPE_ANNOT :: "'b \ 'a itself \ 'b" (infix ":::\<^sub>i" 10) where [simp]: "c:::\<^sub>iI \ c" text \ Wrapper predicate for an conceptual type inference \ definition ID :: "'a \ 'a \ 'c itself \ bool" where [simp]: "ID t t' T \ t=t'" subsubsection \Conceptual Typing Rules\ lemma ID_unfold_vars: "ID x y T \ x\y" by simp lemma ID_PR_CONST_trigger: "ID (PR_CONST x) y T \ ID (PR_CONST x) y T" . lemma pat_rule: "\ p\p'; ID p' t' T \ \ ID p t' T" by simp lemma app_rule: "\ ID f f' TYPE('a\'b); ID x x' TYPE('a)\ \ ID (f$x) (f'$x') TYPE('b)" by simp lemma app'_rule: "\ ID f f' TYPE('a\'b); ID x x' TYPE('a)\ \ ID (f$'x) (f'$x') TYPE('b)" by simp lemma abs_rule: "\ \x x'. ID x x' TYPE('a) \ ID (t x) (t' x x') TYPE('b) \ \ ID (\\<^sub>2x. t x) (\\<^sub>2x'. t' x' x') TYPE('a\'b)" by simp lemma id_rule: "c::\<^sub>iI \ ID c c I" by simp lemma annot_rule: "ID t t' I \ ID (t:::\<^sub>iI) t' I" by simp lemma fallback_rule: "ID (c::'a) c TYPE('c)" by simp lemma unprotect_rl1: "ID (PR_CONST x) t T \ ID (UNPROTECT x) t T" by simp subsection \ ML-Level code \ ML \ infix 0 THEN_ELSE_COMB' signature ID_OP_TACTICAL = sig val SOLVE_FWD: tactic' -> tactic' val DF_SOLVE_FWD: bool -> tactic' -> tactic' end structure Id_Op_Tactical :ID_OP_TACTICAL = struct fun SOLVE_FWD tac i st = SOLVED' ( tac THEN_ALL_NEW_FWD (SOLVE_FWD tac)) i st (* Search for solution with DFS-strategy. If dbg-flag is given, return sequence of stuck states if no solution is found. *) fun DF_SOLVE_FWD dbg tac = let val stuck_list_ref = Unsynchronized.ref [] fun stuck_tac _ st = if dbg then ( stuck_list_ref := st :: !stuck_list_ref; Seq.empty ) else Seq.empty fun rec_tac i st = ( (tac THEN_ALL_NEW_FWD (SOLVED' rec_tac)) ORELSE' stuck_tac ) i st fun fail_tac _ _ = if dbg then Seq.of_list (rev (!stuck_list_ref)) else Seq.empty in rec_tac ORELSE' fail_tac end end \ named_theorems_rev id_rules "Operation identification rules" named_theorems_rev pat_rules "Operation pattern rules" named_theorems_rev def_pat_rules "Definite operation pattern rules (not backtracked over)" ML \ structure Id_Op = struct fun id_a_conv cnv ct = case Thm.term_of ct of @{mpat "ID _ _ _"} => Conv.fun_conv (Conv.fun_conv (Conv.arg_conv cnv)) ct | _ => raise CTERM("id_a_conv",[ct]) fun protect env (@{mpat "?t:::\<^sub>i?I"}) = let val t = protect env t in @{mk_term env: "?t:::\<^sub>i?I"} end | protect _ (t as @{mpat "PR_CONST _"}) = t | protect env (t1$t2) = let val t1 = protect env t1 val t2 = protect env t2 in @{mk_term env: "?t1.0 $ ?t2.0"} end | protect env (Abs (x,T,t)) = let val t = protect (T::env) t in @{mk_term env: "\v_x::?'v_T. PROTECT2 ?t DUMMY"} end | protect _ t = t fun protect_conv ctxt = Refine_Util.f_tac_conv ctxt (protect []) - (simp_tac - (put_simpset HOL_basic_ss ctxt addsimps @{thms PROTECT2_def APP_def}) 1) + (fn goal_ctxt => + simp_tac (put_simpset HOL_basic_ss goal_ctxt addsimps @{thms PROTECT2_def APP_def}) 1) fun unprotect_conv ctxt = Simplifier.rewrite (put_simpset HOL_basic_ss ctxt addsimps @{thms PROTECT2_def APP_def}) fun do_unprotect_tac ctxt = resolve_tac ctxt @{thms unprotect_rl1} THEN' CONVERSION (Refine_Util.HOL_concl_conv (fn ctxt => id_a_conv (unprotect_conv ctxt)) ctxt) val cfg_id_debug = Attrib.setup_config_bool @{binding id_debug} (K false) val cfg_id_trace_fallback = Attrib.setup_config_bool @{binding id_trace_fallback} (K false) fun dest_id_rl thm = case Thm.concl_of thm of @{mpat (typs) "Trueprop (?c::\<^sub>iTYPE(?'v_T))"} => (c,T) | _ => raise THM("dest_id_rl",~1,[thm]) val add_id_rule = snd oo Thm.proof_attributes [Named_Theorems_Rev.add @{named_theorems_rev id_rules}] datatype id_tac_mode = Init | Step | Normal | Solve fun id_tac ss ctxt = let open Id_Op_Tactical val certT = Thm.ctyp_of ctxt val cert = Thm.cterm_of ctxt val thy = Proof_Context.theory_of ctxt val id_rules = Named_Theorems_Rev.get ctxt @{named_theorems_rev id_rules} val pat_rules = Named_Theorems_Rev.get ctxt @{named_theorems_rev pat_rules} val def_pat_rules = Named_Theorems_Rev.get ctxt @{named_theorems_rev def_pat_rules} val rl_net = Tactic.build_net ( (pat_rules |> map (fn thm => thm RS @{thm pat_rule})) @ @{thms annot_rule app_rule app'_rule abs_rule} @ (id_rules |> map (fn thm => thm RS @{thm id_rule})) ) val def_rl_net = Tactic.build_net ( (def_pat_rules |> map (fn thm => thm RS @{thm pat_rule})) ) val id_pr_const_rename_tac = resolve_tac ctxt @{thms ID_PR_CONST_trigger} THEN' Subgoal.FOCUS (fn { context=ctxt, prems, ... } => let fun is_ID @{mpat "Trueprop (ID _ _ _)"} = true | is_ID _ = false val prems = filter (Thm.prop_of #> is_ID) prems val eqs = map (fn thm => thm RS @{thm ID_unfold_vars}) prems val conv = Conv.rewrs_conv eqs val conv = fn ctxt => (Conv.top_sweep_conv (K conv) ctxt) val conv = fn ctxt => Conv.fun2_conv (Conv.arg_conv (conv ctxt)) val conv = Refine_Util.HOL_concl_conv conv ctxt in CONVERSION conv 1 end ) ctxt THEN' resolve_tac ctxt @{thms id_rule} THEN' resolve_tac ctxt id_rules val ityping = id_rules |> map dest_id_rl |> filter (is_Const o #1) |> map (apfst (#1 o dest_Const)) |> Symtab.make_list val has_type = Symtab.defined ityping fun mk_fallback name cT = case try (Sign.the_const_constraint thy) name of SOME T => try (Thm.instantiate' [SOME (certT cT), SOME (certT T)] [SOME (cert (Const (name,cT)))]) @{thm fallback_rule} | NONE => NONE fun trace_fallback thm = Config.get ctxt cfg_id_trace_fallback andalso let open Pretty val p = block [str "ID_OP: Applying fallback rule: ", Thm.pretty_thm ctxt thm] in string_of p |> tracing; false end val fallback_tac = CONVERSION Thm.eta_conversion THEN' IF_EXGOAL (fn i => fn st => case Logic.concl_of_goal (Thm.prop_of st) i of @{mpat "Trueprop (ID (mpaq_STRUCT (mpaq_Const ?name ?cT)) _ _)"} => ( if not (has_type name) then case mk_fallback name cT of SOME thm => (trace_fallback thm; resolve_tac ctxt [thm] i st) | NONE => Seq.empty else Seq.empty ) | _ => Seq.empty) val init_tac = CONVERSION ( Refine_Util.HOL_concl_conv (fn ctxt => (id_a_conv (protect_conv ctxt))) ctxt ) val step_tac = (FIRST' [ assume_tac ctxt, eresolve_tac ctxt @{thms id_rule}, resolve_from_net_tac ctxt def_rl_net, resolve_from_net_tac ctxt rl_net, id_pr_const_rename_tac, do_unprotect_tac ctxt, fallback_tac]) val solve_tac = DF_SOLVE_FWD (Config.get ctxt cfg_id_debug) step_tac in case ss of Init => init_tac | Step => step_tac | Normal => init_tac THEN' solve_tac | Solve => solve_tac end end \ subsection \Default Setup\ subsubsection \Numerals\ lemma pat_numeral[def_pat_rules]: "numeral$x \ UNPROTECT (numeral$x)" by simp lemma id_nat_const[id_rules]: "(PR_CONST (a::nat)) ::\<^sub>i TYPE(nat)" by simp lemma id_int_const[id_rules]: "(PR_CONST (a::int)) ::\<^sub>i TYPE(int)" by simp (*subsection \Example\ schematic_lemma "ID (\a b. (b(1::int\2::nat) |`(-{3})) a, Map.empty, \a. case a of None \ Some a | Some _ \ None) (?c) (?T::?'d itself)" (*"TERM (?c,?T)"*) using [[id_debug]] apply (tactic {* Id_Op.id_tac Id_Op.Normal @{context} 1 *}) done *) end diff --git a/thys/Refine_Imperative_HOL/Sepref_Monadify.thy b/thys/Refine_Imperative_HOL/Sepref_Monadify.thy --- a/thys/Refine_Imperative_HOL/Sepref_Monadify.thy +++ b/thys/Refine_Imperative_HOL/Sepref_Monadify.thy @@ -1,308 +1,308 @@ section \Monadify\ theory Sepref_Monadify imports Sepref_Basic Sepref_Id_Op begin text \ In this phase, a monadic program is converted to complete monadic form, that is, computation of compound expressions are made visible as top-level operations in the monad. The monadify process is separated into 2 steps. \begin{enumerate} \item In a first step, eta-expansion is used to add missing operands to operations and combinators. This way, operators and combinators always occur with the same arity, which simplifies further processing. \item In a second step, computation of compound operands is flattened, introducing new bindings for the intermediate values. \end{enumerate} \ definition SP \ \Tag to protect content from further application of arity and combinator equations\ where [simp]: "SP x \ x" lemma SP_cong[cong]: "SP x \ SP x" by simp lemma PR_CONST_cong[cong]: "PR_CONST x \ PR_CONST x" by simp definition RCALL \ \Tag that marks recursive call\ where [simp]: "RCALL D \ D" definition EVAL \ \Tag that marks evaluation of plain expression for monadify phase\ where [simp]: "EVAL x \ RETURN x" text \ Internally, the package first applies rewriting rules from \sepref_monadify_arity\, which use eta-expansion to ensure that every combinator has enough actual parameters. Moreover, this phase will mark recursive calls by the tag @{const RCALL}. Next, rewriting rules from \sepref_monadify_comb\ are used to add @{const EVAL}-tags to plain expressions that should be evaluated in the monad. The @{const EVAL} tags are flattened using a default simproc that generates left-to-right argument order. \ lemma monadify_simps: "Refine_Basic.bind$(RETURN$x)$(\\<^sub>2x. f x) = f x" "EVAL$x \ RETURN$x" by simp_all definition [simp]: "PASS \ RETURN" \ \Pass on value, invalidating old one\ lemma remove_pass_simps: "Refine_Basic.bind$(PASS$x)$(\\<^sub>2x. f x) \ f x" "Refine_Basic.bind$m$(\\<^sub>2x. PASS$x) \ m" by simp_all definition COPY :: "'a \ 'a" \ \Marks required copying of parameter\ where [simp]: "COPY x \ x" lemma RET_COPY_PASS_eq: "RETURN$(COPY$p) = PASS$p" by simp named_theorems_rev sepref_monadify_arity "Sepref.Monadify: Arity alignment equations" named_theorems_rev sepref_monadify_comb "Sepref.Monadify: Combinator equations" ML \ structure Sepref_Monadify = struct local fun cr_var (i,T) = ("v"^string_of_int i, Free ("__v"^string_of_int i,T)) fun lambda2_name n t = let val t = @{mk_term "PROTECT2 ?t DUMMY"} in Term.lambda_name n t end fun bind_args exp0 [] = exp0 | bind_args exp0 ((x,m)::xms) = let val lr = bind_args exp0 xms |> incr_boundvars 1 |> lambda2_name x in @{mk_term "Refine_Basic.bind$?m$?lr"} end fun monadify t = let val (f,args) = Autoref_Tagging.strip_app t val _ = not (is_Abs f) orelse raise TERM ("monadify: higher-order",[t]) val argTs = map fastype_of args (*val args = map monadify args*) val args = map (fn a => @{mk_term "EVAL$?a"}) args (*val fT = fastype_of f val argTs = binder_types fT*) val argVs = tag_list 0 argTs |> map cr_var val res0 = let val x = Autoref_Tagging.list_APP (f,map #2 argVs) in @{mk_term "SP (RETURN$?x)"} end val res = bind_args res0 (argVs ~~ args) in res end fun monadify_conv_aux ctxt ct = case Thm.term_of ct of @{mpat "EVAL$_"} => let - val ss = put_simpset HOL_basic_ss ctxt - val ss = (ss addsimps @{thms monadify_simps SP_def}) - val tac = (simp_tac ss 1) + fun tac goal_ctxt = + simp_tac (put_simpset HOL_basic_ss goal_ctxt addsimps @{thms monadify_simps SP_def}) 1 in (*Refine_Util.monitor_conv "monadify"*) ( Refine_Util.f_tac_conv ctxt (dest_comb #> #2 #> monadify) tac) ct end | t => raise TERM ("monadify_conv",[t]) (*fun extract_comb_conv ctxt = Conv.rewrs_conv (Named_Theorems_Rev.get ctxt @{named_theorems_rev sepref_monadify_evalcomb}) *) in (* val monadify_conv = Conv.top_conv (fn ctxt => Conv.try_conv ( extract_comb_conv ctxt else_conv monadify_conv_aux ctxt ) ) *) val monadify_simproc = Simplifier.make_simproc @{context} "monadify_simproc" {lhss = [Logic.varify_global @{term "EVAL$a"}], proc = K (try o monadify_conv_aux)}; end local open Sepref_Basic fun mark_params t = let val (P,c,Q,R,a) = dest_hn_refine t val pps = strip_star P |> map_filter (dest_hn_ctxt_opt #> map_option #2) fun tr env (t as @{mpat "RETURN$?x"}) = if is_Bound x orelse member (aconv) pps x then @{mk_term env: "PASS$?x"} else t | tr env (t1$t2) = tr env t1 $ tr env t2 | tr env (Abs (x,T,t)) = Abs (x,T,tr (T::env) t) | tr _ t = t val a = tr [] a in mk_hn_refine (P,c,Q,R,a) end in fun mark_params_conv ctxt = Refine_Util.f_tac_conv ctxt (mark_params) - (simp_tac (put_simpset HOL_basic_ss ctxt addsimps @{thms PASS_def}) 1) + (fn goal_ctxt => simp_tac (put_simpset HOL_basic_ss goal_ctxt addsimps @{thms PASS_def}) 1) end local open Sepref_Basic fun dp ctxt (@{mpat "Refine_Basic.bind$(PASS$?p)$(?t' AS\<^sub>p (\_. PROTECT2 _ DUMMY))"}) = let val (t',ps) = let val ((t',rc),ctxt) = dest_lambda_rc ctxt t' val f = case t' of @{mpat "PROTECT2 ?f _"} => f | _ => raise Match val (f,ps) = dp ctxt f val t' = @{mk_term "PROTECT2 ?f DUMMY"} val t' = rc t' in (t',ps) end val dup = member (aconv) ps p val t = if dup then @{mk_term "Refine_Basic.bind$(RETURN$(COPY$?p))$?t'"} else @{mk_term "Refine_Basic.bind$(PASS$?p)$?t'"} in (t,p::ps) end | dp ctxt (t1$t2) = (#1 (dp ctxt t1) $ #1 (dp ctxt t2),[]) | dp ctxt (t as (Abs _)) = (apply_under_lambda (#1 oo dp) ctxt t,[]) | dp _ t = (t,[]) fun dp_conv ctxt = Refine_Util.f_tac_conv ctxt (#1 o dp ctxt) - (ALLGOALS (simp_tac (put_simpset HOL_basic_ss ctxt addsimps @{thms RET_COPY_PASS_eq}))) + (fn goal_ctxt => + ALLGOALS (simp_tac (put_simpset HOL_basic_ss goal_ctxt addsimps @{thms RET_COPY_PASS_eq}))) in fun dup_tac ctxt = CONVERSION (Sepref_Basic.hn_refine_concl_conv_a dp_conv ctxt) end fun arity_tac ctxt = let val arity1_ss = put_simpset HOL_basic_ss ctxt addsimps ((Named_Theorems_Rev.get ctxt @{named_theorems_rev sepref_monadify_arity})) |> Simplifier.add_cong @{thm SP_cong} |> Simplifier.add_cong @{thm PR_CONST_cong} val arity2_ss = put_simpset HOL_basic_ss ctxt addsimps @{thms beta SP_def} in simp_tac arity1_ss THEN' simp_tac arity2_ss end fun comb_tac ctxt = let val comb1_ss = put_simpset HOL_basic_ss ctxt addsimps (Named_Theorems_Rev.get ctxt @{named_theorems_rev sepref_monadify_comb}) (*addsimps (Named_Theorems_Rev.get ctxt @{named_theorems_rev sepref_monadify_evalcomb})*) addsimprocs [monadify_simproc] |> Simplifier.add_cong @{thm SP_cong} |> Simplifier.add_cong @{thm PR_CONST_cong} val comb2_ss = put_simpset HOL_basic_ss ctxt addsimps @{thms SP_def} in simp_tac comb1_ss THEN' simp_tac comb2_ss end (*fun ops_tac ctxt = CONVERSION ( Sepref_Basic.hn_refine_concl_conv_a monadify_conv ctxt)*) fun mark_params_tac ctxt = CONVERSION ( Refine_Util.HOL_concl_conv (K (mark_params_conv ctxt)) ctxt) fun contains_eval @{mpat "Trueprop (hn_refine _ _ _ _ ?a)"} = Term.exists_subterm (fn @{mpat EVAL} => true | _ => false) a | contains_eval t = raise TERM("contains_eval",[t]); fun remove_pass_tac ctxt = simp_tac (put_simpset HOL_basic_ss ctxt addsimps @{thms remove_pass_simps}) fun monadify_tac dbg ctxt = let open Sepref_Basic in PHASES' [ ("arity", arity_tac, 0), ("comb", comb_tac, 0), (*("ops", ops_tac, 0),*) ("check_EVAL", K (CONCL_COND' (not o contains_eval)), 0), ("mark_params", mark_params_tac, 0), ("dup", dup_tac, 0), ("remove_pass", remove_pass_tac, 0) ] (flag_phases_ctrl dbg) ctxt end end \ lemma dflt_arity[sepref_monadify_arity]: "RETURN \ \\<^sub>2x. SP RETURN$x" "RECT \ \\<^sub>2B x. SP RECT$(\\<^sub>2D x. B$(\\<^sub>2x. RCALL$D$x)$x)$x" "case_list \ \\<^sub>2fn fc l. SP case_list$fn$(\\<^sub>2x xs. fc$x$xs)$l" "case_prod \ \\<^sub>2fp p. SP case_prod$(\\<^sub>2a b. fp$a$b)$p" "case_option \ \\<^sub>2fn fs ov. SP case_option$fn$(\\<^sub>2x. fs$x)$ov" "If \ \\<^sub>2b t e. SP If$b$t$e" "Let \ \\<^sub>2x f. SP Let$x$(\\<^sub>2x. f$x)" by (simp_all only: SP_def APP_def PROTECT2_def RCALL_def) lemma dflt_comb[sepref_monadify_comb]: "\B x. RECT$B$x \ Refine_Basic.bind$(EVAL$x)$(\\<^sub>2x. SP (RECT$B$x))" "\D x. RCALL$D$x \ Refine_Basic.bind$(EVAL$x)$(\\<^sub>2x. SP (RCALL$D$x))" "\fn fc l. case_list$fn$fc$l \ Refine_Basic.bind$(EVAL$l)$(\\<^sub>2l. (SP case_list$fn$fc$l))" "\fp p. case_prod$fp$p \ Refine_Basic.bind$(EVAL$p)$(\\<^sub>2p. (SP case_prod$fp$p))" "\fn fs ov. case_option$fn$fs$ov \ Refine_Basic.bind$(EVAL$ov)$(\\<^sub>2ov. (SP case_option$fn$fs$ov))" "\b t e. If$b$t$e \ Refine_Basic.bind$(EVAL$b)$(\\<^sub>2b. (SP If$b$t$e))" "\x. RETURN$x \ Refine_Basic.bind$(EVAL$x)$(\\<^sub>2x. SP (RETURN$x))" "\x f. Let$x$f \ Refine_Basic.bind$(EVAL$x)$(\\<^sub>2x. (SP Let$x$f))" by (simp_all) lemma dflt_plain_comb[sepref_monadify_comb]: "EVAL$(If$b$t$e) \ Refine_Basic.bind$(EVAL$b)$(\\<^sub>2b. If$b$(EVAL$t)$(EVAL$e))" "EVAL$(case_list$fn$(\\<^sub>2x xs. fc x xs)$l) \ Refine_Basic.bind$(EVAL$l)$(\\<^sub>2l. case_list$(EVAL$fn)$(\\<^sub>2x xs. EVAL$(fc x xs))$l)" "EVAL$(case_prod$(\\<^sub>2a b. fp a b)$p) \ Refine_Basic.bind$(EVAL$p)$(\\<^sub>2p. case_prod$(\\<^sub>2a b. EVAL$(fp a b))$p)" "EVAL$(case_option$fn$(\\<^sub>2x. fs x)$ov) \ Refine_Basic.bind$(EVAL$ov)$(\\<^sub>2ov. case_option$(EVAL$fn)$(\\<^sub>2x. EVAL$(fs x))$ov)" "EVAL $ (Let $ v $ (\\<^sub>2x. f x)) \ (\) $ (EVAL $ v) $ (\\<^sub>2x. EVAL $ (f x))" apply (rule eq_reflection, simp split: list.split prod.split option.split)+ done lemma evalcomb_PR_CONST[sepref_monadify_comb]: "EVAL$(PR_CONST x) \ SP (RETURN$(PR_CONST x))" by simp end diff --git a/thys/Refine_Monadic/Generic/RefineG_Transfer.thy b/thys/Refine_Monadic/Generic/RefineG_Transfer.thy --- a/thys/Refine_Monadic/Generic/RefineG_Transfer.thy +++ b/thys/Refine_Monadic/Generic/RefineG_Transfer.thy @@ -1,290 +1,290 @@ section \Transfer between Domains\ theory RefineG_Transfer imports "../Refine_Misc" begin text \Currently, this theory is specialized to transfers that include no data refinement. \ definition "REFINEG_TRANSFER_POST_SIMP x y \ x=y" definition [simp]: "REFINEG_TRANSFER_ALIGN x y == True" lemma REFINEG_TRANSFER_ALIGNI: "REFINEG_TRANSFER_ALIGN x y" by simp lemma START_REFINEG_TRANSFER: assumes "REFINEG_TRANSFER_ALIGN d c" assumes "c\a" assumes "REFINEG_TRANSFER_POST_SIMP c d" shows "d\a" using assms by (simp add: REFINEG_TRANSFER_POST_SIMP_def) lemma STOP_REFINEG_TRANSFER: "REFINEG_TRANSFER_POST_SIMP c c" unfolding REFINEG_TRANSFER_POST_SIMP_def .. ML \ structure RefineG_Transfer = struct structure Post_Processors = Theory_Data ( type T = (Proof.context -> tactic') Symtab.table val empty = Symtab.empty val extend = I val merge = Symtab.join (K snd) ) fun add_post_processor name tac = Post_Processors.map (Symtab.update_new (name,tac)) fun delete_post_processor name = Post_Processors.map (Symtab.delete name) val get_post_processors = Post_Processors.get #> Symtab.dest fun post_process_tac ctxt = let val tacs = get_post_processors (Proof_Context.theory_of ctxt) |> map (fn (_,tac) => tac ctxt) val tac = REPEAT_DETERM' (CHANGED o EVERY' (map (fn t => TRY o t) tacs)) in tac end structure Post_Simp = Generic_Data ( type T = simpset val empty = HOL_basic_ss val extend = I val merge = Raw_Simplifier.merge_ss ) fun post_simps_op f a context = let val ctxt = Context.proof_of context fun do_it ss = simpset_of (f (put_simpset ss ctxt, a)) in Post_Simp.map do_it context end val add_post_simps = post_simps_op (op addsimps) val del_post_simps = post_simps_op (op delsimps) fun get_post_ss ctxt = let val ss = Post_Simp.get (Context.Proof ctxt) val ctxt = put_simpset ss ctxt in ctxt end structure post_subst = Named_Thms ( val name = @{binding refine_transfer_post_subst} val description = "Refinement Framework: " ^ "Transfer postprocessing substitutions" ); fun post_subst_tac ctxt = let val s_thms = post_subst.get ctxt - val dis_tac = (ALLGOALS (Tagged_Solver.solve_tac ctxt)) + fun dis_tac goal_ctxt = ALLGOALS (Tagged_Solver.solve_tac goal_ctxt) val cnv = Cond_Rewr_Conv.cond_rewrs_conv dis_tac s_thms val ts_conv = Conv.top_sweep_conv cnv ctxt val ss = get_post_ss ctxt in REPEAT o CHANGED o (Simplifier.simp_tac ss THEN' CONVERSION ts_conv) end structure transfer = Named_Thms ( val name = @{binding refine_transfer} val description = "Refinement Framework: " ^ "Transfer rules" ); fun transfer_tac thms ctxt i st = let val thms = thms @ transfer.get ctxt; val ss = put_simpset HOL_basic_ss ctxt addsimps @{thms nested_case_prod_simp} in REPEAT_DETERM1 ( COND (has_fewer_prems (Thm.nprems_of st)) no_tac ( FIRST [ Method.assm_tac ctxt i, resolve_tac ctxt thms i, Tagged_Solver.solve_tac ctxt i, CHANGED_PROP (simp_tac ss i)] )) st end (* Adjust right term to have same structure as left one *) fun align_tac ctxt = IF_EXGOAL (fn i => fn st => case Logic.concl_of_goal (Thm.prop_of st) i of @{mpat "Trueprop (REFINEG_TRANSFER_ALIGN ?c _)"} => let val c = Thm.cterm_of ctxt c val cT = Thm.ctyp_of_cterm c val rl = @{thm REFINEG_TRANSFER_ALIGNI} |> Thm.incr_indexes (Thm.maxidx_of st + 1) |> Thm.instantiate' [NONE,SOME cT] [NONE,SOME c] (*val _ = tracing (@{make_string} rl)*) in resolve_tac ctxt [rl] i st end | _ => Seq.empty ) fun post_transfer_tac thms ctxt = let open Autoref_Tacticals in resolve_tac ctxt @{thms START_REFINEG_TRANSFER} THEN' align_tac ctxt THEN' IF_SOLVED (transfer_tac thms ctxt) (post_process_tac ctxt THEN' resolve_tac ctxt @{thms STOP_REFINEG_TRANSFER}) (K all_tac) end fun get_post_simp_rules context = Context.proof_of context |> get_post_ss |> simpset_of |> Raw_Simplifier.dest_ss |> #simps |> map snd local val add_ps = Thm.declaration_attribute (add_post_simps o single) val del_ps = Thm.declaration_attribute (del_post_simps o single) in val setup = I #> add_post_processor "RefineG_Transfer.post_subst" post_subst_tac #> post_subst.setup #> transfer.setup #> Attrib.setup @{binding refine_transfer_post_simp} (Attrib.add_del add_ps del_ps) ("declaration of transfer post simplification rules") #> Global_Theory.add_thms_dynamic ( @{binding refine_transfer_post_simps}, get_post_simp_rules) end end \ setup \RefineG_Transfer.setup\ method_setup refine_transfer = \Scan.lift (Args.mode "post") -- Attrib.thms >> (fn (post,thms) => fn ctxt => SIMPLE_METHOD' ( if post then RefineG_Transfer.post_transfer_tac thms ctxt else RefineG_Transfer.transfer_tac thms ctxt)) \ "Invoke transfer rules" locale transfer = fixes \ :: "'c \ 'a::complete_lattice" begin text \ In the following, we define some transfer lemmas for general HOL - constructs. \ lemma transfer_if[refine_transfer]: assumes "b \ \ s1 \ S1" assumes "\b \ \ s2 \ S2" shows "\ (if b then s1 else s2) \ (if b then S1 else S2)" using assms by auto lemma transfer_prod[refine_transfer]: assumes "\a b. \ (f a b) \ F a b" shows "\ (case_prod f x) \ (case_prod F x)" using assms by (auto split: prod.split) lemma transfer_Let[refine_transfer]: assumes "\x. \ (f x) \ F x" shows "\ (Let x f) \ Let x F" using assms by auto lemma transfer_option[refine_transfer]: assumes "\ fa \ Fa" assumes "\x. \ (fb x) \ Fb x" shows "\ (case_option fa fb x) \ case_option Fa Fb x" using assms by (auto split: option.split) lemma transfer_sum[refine_transfer]: assumes "\l. \ (fl l) \ Fl l" assumes "\r. \ (fr r) \ Fr r" shows "\ (case_sum fl fr x) \ (case_sum Fl Fr x)" using assms by (auto split: sum.split) lemma transfer_list[refine_transfer]: assumes "\ fn \ Fn" assumes "\x xs. \ (fc x xs) \ Fc x xs" shows "\ (case_list fn fc l) \ case_list Fn Fc l" using assms by (auto split: list.split) lemma transfer_rec_list[refine_transfer]: assumes FN: "\s. \ (fn s) \ fn' s" assumes FC: "\x l rec rec' s. \ \s. \ (rec s) \ (rec' s) \ \ \ (fc x l rec s) \ fc' x l rec' s" shows "\ (rec_list fn fc l s) \ rec_list fn' fc' l s" apply (induct l arbitrary: s) apply (simp add: FN) apply (simp add: FC) done lemma transfer_rec_nat[refine_transfer]: assumes FN: "\s. \ (fn s) \ fn' s" assumes FC: "\n rec rec' s. \ \s. \ (rec s) \ rec' s \ \ \ (fs n rec s) \ fs' n rec' s" shows "\ (rec_nat fn fs n s) \ rec_nat fn' fs' n s" apply (induct n arbitrary: s) apply (simp add: FN) apply (simp add: FC) done end text \Transfer into complete lattice structure\ locale ordered_transfer = transfer + constrains \ :: "'c::complete_lattice \ 'a::complete_lattice" text \Transfer into complete lattice structure with distributive transfer function.\ locale dist_transfer = ordered_transfer + constrains \ :: "'c::complete_lattice \ 'a::complete_lattice" assumes \_dist: "\A. is_chain A \ \ (Sup A) = Sup (\`A)" begin lemma \_mono[simp, intro!]: "mono \" apply rule apply (subgoal_tac "is_chain {x,y}") apply (drule \_dist) apply (auto simp: le_iff_sup) [] apply (rule chainI) apply auto done lemma \_strict[simp]: "\ bot = bot" using \_dist[of "{}"] by simp end text \Transfer into ccpo\ locale ccpo_transfer = transfer \ for \ :: "'c::ccpo \ 'a::complete_lattice" text \Transfer into ccpo with distributive transfer function.\ locale dist_ccpo_transfer = ccpo_transfer \ for \ :: "'c::ccpo \ 'a::complete_lattice" + assumes \_dist: "\A. is_chain A \ \ (Sup A) = Sup (\`A)" begin lemma \_mono[simp, intro!]: "mono \" proof fix x y :: 'c assume LE: "x\y" hence C[simp, intro!]: "is_chain {x,y}" by (auto intro: chainI) from LE have "\ x \ sup (\ x) (\ y)" by simp also have "\ = Sup (\`{x,y})" by simp also have "\ = \ (Sup {x,y})" by (rule \_dist[symmetric]) simp also have "Sup {x,y} = y" apply (rule antisym) apply (rule ccpo_Sup_least[OF C]) using LE apply auto [] apply (rule ccpo_Sup_upper[OF C]) by auto finally show "\ x \ \ y" . qed lemma \_strict[simp]: "\ (Sup {}) = bot" using \_dist[of "{}"] by simp end end