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/Auto2_HOL/logic_steps.ML b/thys/Auto2_HOL/logic_steps.ML --- a/thys/Auto2_HOL/logic_steps.ML +++ b/thys/Auto2_HOL/logic_steps.ML @@ -1,1081 +1,1081 @@ (* File: logic_steps.ML Author: Bohua Zhan Core (logic) proofsteps. *) signature LOGIC_PROOFSTEPS = sig (* General logic *) val shadow_prop_item: proofstep val shadow_term_item: proofstep val exists_elim_prfstep: proofstep val add_logic_proofsteps: theory -> theory val mk_all_disj: term list * term list -> term val strip_all_disj: term -> term list * term list val norm_all_disj: Proof.context -> conv val replace_disj_vars: Proof.context -> term list * term list -> term list * term list val disj_prop_match: Proof.context -> id_inst -> term * (term list * term list) * ((indexname * typ) list * cterm list) -> id_inst_th list val norm_conj: conv (* DISJ items. *) val TY_DISJ: string val disj_to_ritems: bool -> term -> thm -> raw_item list val disj_to_update: bool -> term -> box_id * int option * thm -> raw_update val dest_tname_of_disj: cterm list -> term * cterm list val is_match_prem_only: box_item -> bool val analyze_disj_th: Proof.context -> thm -> term * thm val disj_rewr_terms: term list -> term list val output_disj_fn: item_output val disj_prop_matcher: item_matcher val reduce_disj_True: conv val match_update_prfstep: proofstep val match_one_sch_prfstep: proofstep val disj_match_iff_prfstep: proofstep val disj_create_case_prfstep: proofstep val disj_shadow_prfstep: proofstep val add_disj_proofsteps: theory -> theory (* Normalizers *) val split_not_imp_th: thm -> thm list val split_conj_gen_th: Proof.context -> thm -> thm list val eq_normalizer: normalizer val property_normalizer: normalizer val disj_normalizer: normalizer val logic_thm_update: Proof.context -> box_id * thm -> raw_update val add_disj_normalizers: theory -> theory end; structure Logic_ProofSteps : LOGIC_PROOFSTEPS = struct fun boolVar s = Var ((s, 0), boolT) (* Shadowing based on equivalence. For both PROP and TERM items, shadowing is based on subterm equivalence, skipping any Not (~) at head. *) fun shadow_item_fn ctxt item1 item2 = if #sc item1 = 0 andalso #sc item2 = 0 then [] else let val id = BoxItem.merged_id ctxt [item1, item2] val _ = assert (forall (BoxItem.match_ty_strs [TY_TERM, TY_PROP]) [item1, item2]) "shadow_item_fn" val (tname1, tname2) = (the_single (#tname item1), the_single (#tname item2)) handle List.Empty => raise Fail "shadow_item_fn" val (t1, t2) = (Thm.term_of tname1, Thm.term_of tname2) val (lhs, rhs) = if fastype_of t1 = boolT andalso is_neg t1 andalso fastype_of t2 = boolT andalso is_neg t2 then (UtilLogic.get_cneg tname1, UtilLogic.get_cneg tname2) else (tname1, tname2) val equiv_ids = (RewriteTable.subequiv_info ctxt id (lhs, rhs)) |> map fst |> filter BoxID.has_incr_id |> Util.max_partial (BoxID.is_eq_ancestor ctxt) val item_to_shadow = if #uid item1 > #uid item2 then item1 else item2 fun process_id id' = ShadowItem {id = id', item = item_to_shadow} in map process_id equiv_ids end val shadow_prop_item = {name = "shadow_prop", args = [TypedUniv TY_PROP, TypedUniv TY_PROP], func = TwoStep shadow_item_fn} val shadow_term_item = {name = "shadow_term", args = [TypedUniv TY_TERM, TypedUniv TY_TERM], func = TwoStep shadow_item_fn} fun eq_abs_fn ctxt item1 item2 = let val {id = id1, tname = tname1, ...} = item1 val {id = id2, tname = tname2, ...} = item2 val id = BoxID.merge_boxes ctxt (id1, id2) val (ct1, ct2) = apply2 the_single (tname1, tname2) val (t1, t2) = apply2 Thm.term_of (ct1, ct2) in if not (Util.is_abs t1) orelse not (Util.is_abs t2) then [] else if RewriteTable.is_equiv id ctxt (ct1, ct2) then [] else if Term_Ord.term_ord (t2, t1) = LESS then [] else let fun process_equiv (id', eq_th) = let val (lhs, rhs) = (Util.lhs_of eq_th, Util.rhs_of eq_th) in AddItems {id = id', sc = SOME 1, raw_items = [Fact (TY_EQ, [lhs, rhs], eq_th)]} end in (Matcher.rewrite_match ctxt (t1, ct2) (id, fo_init)) |> map (fn ((id, _), eq_th) => (id, eq_th)) |> filter (BoxID.has_incr_id o fst) |> Util.max_partial (BoxID.id_is_eq_ancestor ctxt) |> map process_equiv end end val eq_abs_prfstep = {name = "eq_abs", args = [TypedUniv TY_TERM, TypedUniv TY_TERM], func = TwoStep eq_abs_fn} (* Given an assumption of the form EX x. A, we produce an assumption A with x in A replaced by a free variable. To avoid name collisions, when the update is produced x is replaced by an "internal" free variable, with suffix '_'. When the update is applied, that internal free variable is replaced by a fresh variable as determined by the context. We produce at most two variables at a time. *) fun exists_elim_fn ctxt {id, prop, ...} = if not (BoxID.has_incr_id id) then [] else let val t = prop_of' prop in if is_ex t orelse is_bex t orelse (is_neg t andalso is_obj_all (dest_not t)) orelse (is_neg t andalso is_ball (dest_not t)) then let val (ritems, th') = prop |> apply_to_thm' (UtilLogic.normalize_exists ctxt) |> Update.apply_exists_ritems ctxt in [AddItems {id = id, sc = NONE, raw_items = ritems @ [Update.thm_to_ritem th']}] end else [] end val exists_elim_prfstep = {name = "exists_elim", args = [TypedUniv TY_PROP], func = OneStep exists_elim_fn} val add_logic_proofsteps = fold add_prfstep [ shadow_prop_item, shadow_term_item, exists_elim_prfstep, eq_abs_prfstep ] (* Given (x_1 ... x_n, A_1 ... A_n), create the corresponding forall-disj term. *) fun mk_all_disj (vars, terms) = case vars of [] => list_disj terms | var :: vars' => mk_obj_all var (mk_all_disj (vars', terms)) (* Normalize t into a disjunction of terms. *) fun strip_disj t = if is_disj t then maps strip_disj [dest_arg1 t, dest_arg t] else if is_imp t then maps strip_disj [get_neg (dest_arg1 t), dest_arg t] else if is_neg t then let val t' = dest_not t in if is_neg t' then strip_disj (dest_not t') else if is_conj t' then maps strip_disj [get_neg (dest_arg1 t'), get_neg (dest_arg t')] else [t] end else [t] (* Normalize a term into the form !x_1 ... x_n. A_1 | ... | A_n *) fun strip_all_disj t = if is_obj_all t then case t of - _ $ Abs (abs as (_, T, _)) => + _ $ (u as Abs _) => let - val (x, body) = Term.dest_abs abs - val var = Free (x, T) + val (v, body) = Term.dest_abs_global u + val var = Free v val (vars, disjs) = strip_all_disj body in if exists (Util.occurs_free var) disjs then (var :: vars, disjs) else (vars, disjs) end | f $ arg => strip_all_disj (f $ UtilLogic.force_abs_form arg) | _ => raise Fail "strip_all_disj" else if is_ball t then case t of - _ $ S $ Abs (abs as (_, T, _)) => + _ $ S $ (u as Abs _) => let - val (x, body) = Term.dest_abs abs - val var = Free (x, T) + val (v, body) = Term.dest_abs_global u + val var = Free v val mem = mk_mem (var, S) val (vars, disjs) = strip_all_disj body in (var :: vars, get_neg mem :: disjs) end | f $ S $ arg => strip_all_disj (f $ S $ UtilLogic.force_abs_form arg) | _ => raise Fail "strip_all_disj" else if is_neg t andalso is_ex (dest_not t) then case dest_not t of - _ $ Abs (abs as (_, T, _)) => + _ $ (u as Abs _) => let - val (x, body) = Term.dest_abs abs - val var = Free (x, T) + val (v, body) = Term.dest_abs_global u + val var = Free v val (vars, disjs) = strip_all_disj (get_neg body) in if exists (Util.occurs_free var) disjs then (var :: vars, disjs) else (vars, disjs) end | f $ arg => strip_all_disj (get_neg (f $ UtilLogic.force_abs_form arg)) | _ => raise Fail "strip_all_disj" else if is_neg t andalso is_bex (dest_not t) then case dest_not t of - _ $ S $ Abs (abs as (_, T, _)) => + _ $ S $ (u as Abs _) => let - val (x, body) = Term.dest_abs abs - val var = Free (x, T) + val (v, body) = Term.dest_abs_global u + val var = Free v val mem = mk_mem (var, S) val (vars, disjs) = strip_all_disj (get_neg body) in (var :: vars, get_neg mem :: disjs) end | f $ S $ arg => strip_all_disj (get_neg (f $ S $ UtilLogic.force_abs_form arg)) | _ => raise Fail "strip_all_disj" else if is_disj t then let val (v1, ts1) = strip_all_disj (dest_arg1 t) val (v2, ts2) = strip_all_disj (dest_arg t) in (v1 @ v2, ts1 @ ts2) end else if is_imp t then let val (v1, ts1) = strip_all_disj (get_neg (dest_arg1 t)) val (v2, ts2) = strip_all_disj (dest_arg t) in (v1 @ v2, ts1 @ ts2) end else if is_neg t then let val t' = dest_not t in if is_neg t' then strip_all_disj (dest_not t') else if is_conj t' then let val (v1, ts1) = strip_all_disj (get_neg (dest_arg1 t')) val (v2, ts2) = strip_all_disj (get_neg (dest_arg t')) in (v1 @ v2, ts1 @ ts2) end else ([], [t]) end else ([], [t]) (* Normalize (A_1 | A_2 | ... | A_m) | (B_1 | B_2 | ... | B_n) *) fun norm_disj_clauses ct = let val (arg1, _) = Util.dest_binop_args (Thm.term_of ct) in if is_disj arg1 then Conv.every_conv [rewr_obj_eq UtilBase.disj_assoc_th, Conv.arg_conv norm_disj_clauses] ct else Conv.all_conv ct end (* Normalize ct. *) fun norm_disj ct = let val t = Thm.term_of ct val _ = assert (fastype_of t = boolT) "norm_disj: wrong type" in if is_disj t then Conv.every_conv [Conv.binop_conv norm_disj, norm_disj_clauses] ct else if is_imp t then Conv.every_conv [rewr_obj_eq UtilBase.imp_conv_disj_th, norm_disj] ct else if is_neg t andalso is_neg (dest_not t) then Conv.every_conv [rewr_obj_eq UtilBase.nn_cancel_th, norm_disj] ct else if is_neg t andalso is_conj (dest_not t) then Conv.every_conv [rewr_obj_eq UtilBase.de_Morgan_conj_th, norm_disj] ct else Conv.all_conv ct end (* Normalize to forall at the top-level *) fun norm_all_disj ctxt ct = let val t = Thm.term_of ct in if is_obj_all t then if not (Util.is_subterm (Bound 0) (dest_arg t)) then Conv.every_conv [rewr_obj_eq UtilBase.all_trivial_th, norm_all_disj ctxt] ct else Conv.every_conv [Conv.binder_conv (norm_all_disj o snd) ctxt] ct else if is_ball t then Conv.every_conv [rewr_obj_eq UtilBase.Ball_def_th, norm_all_disj ctxt] ct else if is_neg t andalso is_ex (dest_not t) then Conv.every_conv [rewr_obj_eq UtilBase.not_ex_th, norm_all_disj ctxt] ct else if is_neg t andalso is_bex (dest_not t) then Conv.every_conv [ Conv.arg_conv (rewr_obj_eq UtilBase.Bex_def_th), norm_all_disj ctxt] ct else if is_disj t then let val eq1 = Conv.binop_conv (norm_all_disj ctxt) ct val rhs = Util.rhs_of eq1 in if is_obj_all (dest_arg1 rhs) then Conv.every_conv [Conv.rewr_conv eq1, rewr_obj_eq UtilBase.disj_commute_th, rewr_obj_eq UtilBase.swap_all_disj_th, norm_all_disj ctxt] ct else if is_obj_all (dest_arg rhs) then Conv.every_conv [Conv.rewr_conv eq1, rewr_obj_eq UtilBase.swap_all_disj_th, norm_all_disj ctxt] ct else Conv.every_conv [Conv.rewr_conv eq1, norm_disj_clauses] ct end else if is_imp t then Conv.every_conv [rewr_obj_eq UtilBase.imp_conv_disj_th, norm_all_disj ctxt] ct else if is_neg t andalso is_neg (dest_not t) then Conv.every_conv [rewr_obj_eq UtilBase.nn_cancel_th, norm_all_disj ctxt] ct else if is_neg t andalso is_conj (dest_not t) then Conv.every_conv [rewr_obj_eq UtilBase.de_Morgan_conj_th, norm_all_disj ctxt] ct else Conv.all_conv ct end fun mk_disj_eq eq_ths = case eq_ths of [] => raise Fail "mk_disj_eq" | [eq] => eq | eq :: eqs' => Drule.binop_cong_rule UtilBase.cDisj eq (mk_disj_eq eqs') (* Sort A | (B_1 | B_2 | ... | B_n), assuming the right side is sorted. *) fun sort_disj_clause_aux ct = if not (is_disj (Thm.term_of ct)) then Conv.all_conv ct else let val (arg1, arg2) = Util.dest_binop_args (Thm.term_of ct) in if is_disj arg2 then if Term_Ord.term_ord (dest_arg1 arg2, arg1) = LESS then Conv.every_conv [rewr_obj_eq (obj_sym UtilBase.disj_assoc_th), Conv.arg1_conv (rewr_obj_eq UtilBase.disj_commute_th), rewr_obj_eq UtilBase.disj_assoc_th, Conv.arg_conv sort_disj_clause_aux] ct else Conv.all_conv ct else if Term_Ord.term_ord (arg2, arg1) = LESS then rewr_obj_eq (obj_sym UtilBase.disj_commute_th) ct else Conv.all_conv ct end (* Sort A_1 | ... | A_n. *) fun sort_disj_clause ct = if not (is_disj (Thm.term_of ct)) then Conv.all_conv ct else Conv.every_conv [Conv.arg_conv sort_disj_clause, sort_disj_clause_aux] ct (* Apply cv on the body of !x y z. ... .*) fun forall_body_conv cv ctxt ct = if is_obj_all (Thm.term_of ct) then Conv.binder_conv (fn (_, ctxt) => forall_body_conv cv ctxt) ctxt ct else cv ct fun norm_all_disj_sorted ctxt ct = Conv.every_conv [norm_all_disj ctxt, forall_body_conv sort_disj_clause ctxt] ct fun abstract_eq ctxt var eq = let val (x, T) = Term.dest_Free var val all_const = Const (UtilBase.All_name, (T --> boolT) --> boolT) in Drule.arg_cong_rule (Thm.cterm_of ctxt all_const) (Thm.abstract_rule x (Thm.cterm_of ctxt var) eq) end fun replace_disj_vars ctxt (vars, disjs) = let val vars' = vars |> map Term.dest_Free |> Variable.variant_frees ctxt [] |> map Free val subst = vars ~~ vars' in (vars', map (subst_atomic subst) disjs) end (* Matching for all-disj propositions *) fun disj_prop_match ctxt (id, (tyinst, inst)) (t, (var_t, ts), (var_u, cus)) = let val thy = Proof_Context.theory_of ctxt val us = map Thm.term_of cus in if length var_t <> length var_u orelse length ts <> length us then [] else let (* First match the types (return [] if no match). *) val tys_t = map fastype_of var_t val tys_u = map snd var_u val tyinst' = fold (Sign.typ_match thy) (tys_t ~~ tys_u) tyinst val var_t' = map (Envir.subst_term_types tyinst') var_t val ts' = ts |> map (Envir.subst_term_types tyinst') val var_ct' = map (Thm.cterm_of ctxt) var_t' val cus' = cus |> map (Thm.instantiate_cterm (TVars.empty, Vars.make (var_u ~~ var_ct'))) (* Match the type-instantiated pattern with term. *) val insts = Matcher.rewrite_match_subset ctxt var_t' (ts', cus') (id, (tyinst', inst)) fun process_inst ((id', instsp'), ths) = let (* Equality between normalized t and u *) val eq_th = ths |> mk_disj_eq |> fold (abstract_eq ctxt) (rev var_t') |> apply_to_lhs (norm_all_disj_sorted ctxt) |> apply_to_rhs (norm_all_disj_sorted ctxt) (* Equality between un-normalized t and u *) val t' = Util.subst_term_norm instsp' t val norm1 = norm_all_disj_sorted ctxt (Thm.cterm_of ctxt t') val cu = Thm.cterm_of ctxt (mk_all_disj (var_t', map Thm.term_of cus')) val norm2 = norm_all_disj_sorted ctxt cu val eq_th' = Util.transitive_list [norm1, eq_th, meta_sym norm2] in ((id', instsp'), eq_th') end in map process_inst insts end handle Type.TYPE_MATCH => [] end fun norm_conj_de_Morgan ct = let val t = Thm.term_of ct in if is_neg t andalso is_disj (dest_not t) then Conv.every_conv [rewr_obj_eq UtilBase.de_Morgan_disj_th, Conv.arg_conv norm_conj_de_Morgan] ct else Conv.all_conv ct end fun norm_conj_not_imp ct = let val t = Thm.term_of ct in if is_neg t andalso is_imp (dest_not t) then Conv.every_conv [rewr_obj_eq UtilBase.not_imp_th, Conv.arg_conv norm_conj_not_imp] ct else Conv.all_conv ct end fun norm_nn_cancel ct = let val t = Thm.term_of ct in if is_neg t andalso is_neg (dest_not t) then Conv.every_conv [rewr_obj_eq UtilBase.nn_cancel_th, norm_nn_cancel] ct else Conv.all_conv ct end fun norm_conj ct = let val t = Thm.term_of ct val _ = assert (fastype_of t = boolT) "norm_conj: wrong type" in if is_neg t andalso is_neg (dest_not t) then norm_nn_cancel ct else if is_neg t andalso is_imp (dest_not t) then norm_conj_not_imp ct else if is_neg t andalso is_disj (dest_not t) then norm_conj_de_Morgan ct else Conv.all_conv ct end (* DISJ items. *) val TY_DISJ = "DISJ" (* Given a theorem in the form of a disjunction, possibly containing schematic variables, return the corresponding DISJ item. *) fun disj_to_ritems prem_only disj_head th = let val subs = strip_disj (prop_of' th) in if length subs = 1 then if Util.has_vars (Thm.prop_of th) then let fun th_to_ritem th = Fact (TY_DISJ, UtilLogic.term_of_bool prem_only :: disj :: strip_disj (prop_of' th), th) in map th_to_ritem (th |> apply_to_thm' norm_conj |> UtilLogic.split_conj_th) end else [Fact (TY_PROP, [prop_of' th], th)] else let val tname = UtilLogic.term_of_bool prem_only :: disj_head :: subs in [Fact (TY_DISJ, tname, th)] end end fun disj_to_update prem_only disj_head (id, sc, th) = if Thm.prop_of th aconv pFalse then ResolveBox {id = id, th = th} else AddItems {id = id, sc = sc, raw_items = disj_to_ritems prem_only disj_head th} fun get_disj_head t = if is_disj t then disj else if is_imp t then get_disj_head (dest_arg t) else if is_neg t andalso is_neg (dest_not t) then get_disj_head (dest_not (dest_not t)) else if is_neg t andalso is_conj (dest_not t) then conj else if is_obj_all t orelse is_ball t then case dest_arg t of - Abs (x, T, body) => get_disj_head (snd (Term.dest_abs (x, T, body))) + u as Abs _ => get_disj_head (snd (Term.dest_abs_global u)) | _ => imp else if is_neg t andalso is_ex (dest_not t) then conj else if is_neg t andalso is_bex (dest_not t) then conj else imp (* Given a theorem th, return equivalent theorem in disjunctive form, with possible schematic variables. Also return whether th is "active", that is, whether it is originally a conjunctive goal or disjunctive fact, as opposed to implications. *) fun analyze_disj_th ctxt th = let val head = get_disj_head (prop_of' th) val th' = th |> apply_to_thm' (norm_all_disj ctxt) |> apply_to_thm (UtilLogic.to_meta_conv ctxt) |> Util.forall_elim_sch in (head, th') end (* Deconstruct the tname of a DISJ item. *) fun dest_tname_of_disj tname = case tname of _ :: disj_head :: rest => (Thm.term_of disj_head, rest) | _ => raise Fail "dest_tname_of_disj: too few terms in tname." (* Determine whether the item is for matching premises only (from the first entry in tname. *) fun is_match_prem_only {tname, ...} = UtilLogic.bool_of_term (Thm.term_of (hd tname)) fun is_active_head disj_head = (not (disj_head aconv imp)) fun disj_rewr_terms ts = if UtilLogic.bool_of_term (hd ts) then [] else drop 2 ts fun output_disj_fn ctxt (ts, _) = let val (match_prem, disj_head, subs) = (hd ts, hd (tl ts), tl (tl ts)) val prefix = if UtilLogic.bool_of_term match_prem then "(match_prem) " else "" in if disj_head aconv disj then prefix ^ ((foldr1 mk_disj subs) |> Syntax.string_of_term ctxt) else if disj_head aconv conj then prefix ^ ((foldr1 mk_conj (map get_neg subs)) |> get_neg |> Syntax.string_of_term ctxt) else if disj_head aconv imp then prefix ^ ((foldr1 mk_imp (subs |> split_last |> apfst (map get_neg) |> apsnd single |> (op @))) |> Syntax.string_of_term ctxt) else raise Fail "output_disj_fn: unexpected disj_head." end val disj_prop_matcher = let fun pre_match pat {tname, ...} ctxt = let val (var_t, ts) = (strip_all_disj pat) |> replace_disj_vars ctxt val (_, cus) = dest_tname_of_disj tname val us = map Thm.term_of cus val var_u = fold Term.add_vars us [] in length ts = length us andalso length var_t = length var_u end fun match pat item ctxt (id, instsp) = let val {tname, prop = th, ...} = item val (_, cus) = dest_tname_of_disj tname val us = map Thm.term_of cus val var_u = fold Term.add_vars us [] val (var_t, ts) = (strip_all_disj pat) |> replace_disj_vars ctxt fun process_perm perm = map (pair (map Var perm)) (disj_prop_match ctxt (id, instsp) (pat, (var_t, ts), (perm, cus))) fun process_inst (var_u, ((id', instsp'), eq_th)) = let val eq_th' = make_trueprop_eq (meta_sym eq_th) val forall_th = th |> fold Thm.forall_intr (rev (map (Thm.cterm_of ctxt) var_u)) |> apply_to_thm (UtilLogic.to_obj_conv ctxt) in ((id', instsp'), Thm.equal_elim eq_th' forall_th) end in if length var_t <> length var_u orelse length ts <> length us then [] else var_u |> Util.all_permutes |> maps process_perm |> map process_inst end in {pre_match = pre_match, match = match} end (* Given ct in the form p_1 | ... | p_n, apply cv to each of p_i. *) fun ac_disj_conv cv ct = if is_disj (Thm.term_of ct) then Conv.every_conv [Conv.arg1_conv cv, Conv.arg_conv (ac_disj_conv cv)] ct else cv ct (* Assume ct is a disjunction, associating to the right. *) fun reduce_disj_True ct = if is_disj (Thm.term_of ct) then ((rewr_obj_eq UtilBase.disj_True1_th) else_conv ((Conv.arg_conv reduce_disj_True) then_conv (rewr_obj_eq UtilBase.disj_True2_th))) ct else Conv.all_conv ct (* Handles also the case where pat is in not-conj or imp form. *) fun match_prop ctxt (id, item2) pat = let val disj_pats = strip_disj pat (* th is pat'(inst), where pat' is one of the disjunctive terms of pat. *) fun process_inst ((id, inst), th) = let (* Construct the theorem pat'(inst) == True. *) val to_eqT_cv = (th RS UtilBase.eq_True_th) |> rewr_obj_eq |> Conv.try_conv (* Rewrite pat(inst) using the above, then rewrite to True. *) val pat_eqT = pat |> Thm.cterm_of ctxt |> norm_disj |> Util.subst_thm ctxt inst |> apply_to_rhs (ac_disj_conv to_eqT_cv) |> apply_to_rhs reduce_disj_True |> to_obj_eq val patT = pat_eqT RS UtilBase.eq_True_inv_th in ((id, inst), patT) end val insts1 = (ItemIO.match_arg ctxt (PropMatch pat) item2 (id, fo_init)) |> filter (BoxID.has_incr_id o fst o fst) val insts2 = if length disj_pats > 1 then map process_inst (maps (match_prop ctxt (id, item2)) disj_pats) else [] in insts1 @ insts2 end (* Given theorem ~P, cancel any disjunct that is aconv to P. It is possible to leave one disjunct P un-cancelled. *) fun disj_cancel_cv ctxt th ct = if is_disj (Thm.term_of ct) then Conv.every_conv [Conv.arg_conv (disj_cancel_cv ctxt th), Conv.try_conv (rewr_obj_eq (th RS UtilBase.or_cancel1_th)), Conv.try_conv (rewr_obj_eq (th RS UtilBase.or_cancel2_th))] ct else Conv.all_conv ct (* Given theorem P and a disjuncion theorem, return new disjunction theorem with ~P cancelled. If all disjuncts can be cancelled, return False. *) fun disj_cancel_prop ctxt th prop = let val th' = if is_neg (prop_of' th) then th else th RS UtilBase.nn_create_th val prop' = prop |> apply_to_thm' (disj_cancel_cv ctxt th') val P = th' |> prop_of' |> dest_not in if prop_of' prop' aconv P then [th', prop'] MRS UtilBase.contra_triv_th else prop' end (* Reduce a disjunction p_1 | ... | t | ... | p_n by matching ~t with the second item. If the disjunction contains schematic variables, t must have either zero or the largest number of schematic variables. *) fun match_update_fn ctxt item1 item2 = if is_match_prem_only item1 andalso Util.has_vars (Thm.prop_of (#prop item1)) then [] else let val {id, prop, tname, ...} = item1 val thy = Proof_Context.theory_of ctxt val (disj_head, csubs) = dest_tname_of_disj tname val subs = map Thm.term_of csubs fun count_var t = length (Term.add_vars t []) val max_nvar = fold (curry Int.max) (map count_var subs) 0 fun is_priority_term t = if is_neg t then exists (is_ex orf is_bex) (strip_conj (dest_not t)) else exists (is_obj_all orf is_ball) (strip_disj t) val has_priority_term = exists is_priority_term (map get_neg subs) val (NO_MATCH, SLOW_MATCH, YES_MATCH) = (0, 1, 2) (* Test whether to perform matching on pattern. *) fun test_do_match t = let val nvar = count_var t val neg_t = get_neg t in if not (Util.is_pattern t) then NO_MATCH else if length subs > 1 andalso Property.is_property_prem thy neg_t then NO_MATCH else if has_priority_term andalso not (is_priority_term neg_t) then SLOW_MATCH else if is_mem neg_t andalso Term.is_Var (dest_arg1 neg_t) andalso null (Term.add_frees (dest_arg neg_t) []) then SLOW_MATCH else if nvar = 0 orelse nvar = max_nvar then YES_MATCH else NO_MATCH end (* Match the negation of subs[i] with th2. For each match, instantiate in prop all schematic variables in t, so that t becomes ~th2. Then remove t from prop in the instantiated version. *) fun get_matches i = let val t = nth subs i val do_match = test_do_match t fun process_inst ((id', inst), th) = let val prop' = prop |> Util.subst_thm_thy thy inst |> disj_cancel_prop ctxt th val sc = if do_match = SLOW_MATCH then 200 else 10 in disj_to_update false disj_head (id', SOME sc, prop') :: (if count_var t > 0 then [] else [ShadowItem {id = id', item = item1}]) end in if do_match = NO_MATCH then [] else t |> get_neg |> match_prop ctxt (id, item2) |> maps process_inst end fun get_matches_no_var () = let fun process_inst (id', ths) = let val prop' = prop |> fold (disj_cancel_prop ctxt) ths in disj_to_update false disj_head (id', SOME 1, prop') :: (if is_match_prem_only item1 then [] else [ShadowItem {id = id', item = item1}]) end fun get_match_at_id id' insts = insts |> filter (fn ((id, _), _) => BoxID.is_eq_ancestor ctxt id id') |> map snd |> take 1 fun get_matches_at_id all_insts id' = (id', maps (get_match_at_id id') all_insts) fun merge_matches all_insts = let val ids = distinct (op =) (maps (map (fst o fst)) all_insts) in map (get_matches_at_id all_insts) ids end val _ = assert (length subs >= 2) val ts = [hd subs, List.last subs] in ts |> map get_neg |> map (match_prop ctxt (id, item2)) |> merge_matches |> maps process_inst end in if max_nvar = 0 then get_matches_no_var () else maps get_matches (0 upto (length subs - 1)) end val match_update_prfstep = {name = "disj_match_update", args = [TypedUniv TY_DISJ, PropMatch (boolVar "A")], func = TwoStep match_update_fn} (* For DISJ items with a single term, of form f p1 ... pn, match t against each of p_i. *) fun match_one_sch_fn ctxt item1 item2 = if is_match_prem_only item1 then [] else let val {id, tname, prop = th1, ...} = item1 val thy = Proof_Context.theory_of ctxt val subs = (dest_tname_of_disj tname) |> snd |> map Thm.term_of in if length subs > 1 then [] else let val t = the_single subs val args = Util.dest_args t fun count_var t = length (Term.add_vars t []) val nvar = count_var t fun get_matches i = if count_var (nth args i) < nvar then [] else let val arg = nth args i val targ = TypedMatch (TY_TERM, arg) val insts = (ItemIO.match_arg ctxt targ item2 (id, fo_init)) |> filter (BoxID.has_incr_id o fst o fst) fun inst_to_updt ((id', inst), _) = let val th1' = Util.subst_thm_thy thy inst th1 val prop' = prop_of' th1' in if is_eq_term prop' andalso RewriteTable.is_equiv_t id' ctxt (dest_eq prop') then [] else [Update.thm_update (id', th1')] end in maps inst_to_updt insts end in maps get_matches (0 upto (length args - 1)) end end val match_one_sch_prfstep = {name = "disj_match_one_sch", args = [TypedUniv TY_DISJ, TypedUniv TY_TERM], func = TwoStep match_one_sch_fn} fun disj_match_iff_fn ctxt {id, tname, prop, ...} = if not (BoxID.has_incr_id id) then [] else let val (_, csubs) = dest_tname_of_disj tname val subs = map Thm.term_of csubs in if length subs > 1 then [] else if not (is_eq_term (the_single subs) andalso fastype_of (dest_arg (the_single subs)) = boolT) then [] else let val cv = (UtilLogic.to_obj_conv ctxt) then_conv (Trueprop_conv norm_disj) val forward = prop |> UtilLogic.equiv_forward_th |> apply_to_thm cv val backward = prop |> UtilLogic.equiv_backward_th |> apply_to_thm cv in [disj_to_update false imp (id, NONE, forward), disj_to_update false imp (id, NONE, backward)] end end val disj_match_iff_prfstep = {name = "disj_match_iff", args = [TypedUniv TY_DISJ], func = OneStep disj_match_iff_fn} (* For active case, create box checking the next case. *) fun disj_create_case_fn _ {id, tname, ...} = if not (BoxID.has_incr_id id) then [] else if exists Util.has_vars (map Thm.term_of tname) then [] else let val (disj_head, csubs) = dest_tname_of_disj tname in if not (is_active_head disj_head) then [] else if length csubs = 1 then [] else let val subs = map Thm.term_of csubs in [AddBoxes {id = id, sc = NONE, init_assum = mk_Trueprop (hd subs)}] end end val disj_create_case_prfstep = {name = "disj_create_case", args = [TypedUniv TY_DISJ], func = OneStep disj_create_case_fn} (* item1 dominates item2 if the disjunctive terms in item1 is a subset of that for item2. *) fun disj_shadow_fn ctxt (item1 as {tname = tname1, ...}) (item2 as {tname = tname2, ...}) = let val id = BoxItem.merged_id ctxt [item1, item2] val (disj_head1, subs1) = dest_tname_of_disj tname1 val (disj_head2, subs2) = dest_tname_of_disj tname2 in if not (BoxID.has_incr_id id) then [] else if not (is_active_head disj_head1) andalso is_active_head disj_head2 then [] else if is_match_prem_only item1 andalso not (is_match_prem_only item2) then [] else if subset (op aconvc) (subs1, subs2) then [ShadowItem {id = id, item = item2}] else [] end val disj_shadow_prfstep = {name = "disj_shadow", args = [TypedUniv TY_DISJ, TypedUniv TY_DISJ], func = TwoStep disj_shadow_fn} val add_disj_proofsteps = fold ItemIO.add_item_type [ (TY_DISJ, SOME disj_rewr_terms, SOME output_disj_fn, NONE) ] #> fold ItemIO.add_prop_matcher [ (TY_DISJ, disj_prop_matcher) ] #> fold add_prfstep [ match_update_prfstep, match_one_sch_prfstep, disj_match_iff_prfstep, disj_create_case_prfstep, disj_shadow_prfstep ] (* Normalizers *) fun split_not_imp_th th = th |> apply_to_thm' norm_conj_not_imp |> UtilLogic.split_conj_th (* Generalized form of splitting A & B. Also deal with cases ~(A | B) and ~(A --> B). *) fun split_conj_gen_th _ th = th |> apply_to_thm' norm_conj |> UtilLogic.split_conj_th fun eq_normalizer _ ritem = case ritem of Handler _ => [ritem] | Fact (ty_str, _, th) => if ty_str <> TY_PROP then [ritem] else if is_eq_term (prop_of' th) then let val (lhs, rhs) = dest_eq (prop_of' th) in if fastype_of lhs = boolT then map Update.thm_to_ritem (UtilLogic.split_conj_th (th RS UtilBase.iffD_th)) else [Fact (TY_EQ, [lhs, rhs], th)] end else [ritem] fun property_normalizer _ ritem = case ritem of Handler _ => [ritem] | Fact (ty_str, _, th) => if ty_str <> TY_PROP then [ritem] else if Property.is_property (prop_of' th) then [Fact (TY_PROPERTY, [prop_of' th], th)] else [ritem] fun disj_normalizer ctxt ritem = case ritem of Handler _ => [ritem] | Fact (ty_str, _, th) => if ty_str <> TY_PROP then [ritem] else let val t = prop_of' th in if is_neg t andalso is_conj (dest_not t) orelse is_disj t orelse is_imp t orelse is_obj_all t orelse is_ball t orelse is_neg t andalso is_ex (dest_not t) orelse is_neg t andalso is_bex (dest_not t) then let val (disj_head, disj_th) = analyze_disj_th ctxt th val prem_only = Auto2_State.lookup_prem_only ctxt t val disj_th = if prem_only then disj_th else snd (Normalizer.meta_use_vardefs disj_th) in disj_to_ritems prem_only disj_head disj_th end else [ritem] end fun logic_thm_update ctxt (id, th) = let val t = prop_of' th in if is_obj_all t orelse is_ball t orelse is_neg t andalso is_ex (dest_not t) orelse is_neg t andalso is_bex (dest_not t) orelse is_disj t orelse is_imp t orelse is_neg t andalso is_conj (dest_not t) then let val (disj_head, disj_th) = analyze_disj_th ctxt th val raw_items = disj_to_ritems true disj_head disj_th in AddItems {id = id, sc = NONE, raw_items = raw_items} end else Update.thm_update (id, th) end val add_disj_normalizers = Normalizer.add_th_normalizer ( "split_conj_gen", split_conj_gen_th ) #> fold Normalizer.add_normalizer [ ("eq", eq_normalizer), ("property", property_normalizer), ("disj", disj_normalizer) ] end (* structure Logic_ProofSteps. *) val _ = Theory.setup Logic_ProofSteps.add_logic_proofsteps val _ = Theory.setup Logic_ProofSteps.add_disj_proofsteps val _ = Theory.setup Logic_ProofSteps.add_disj_normalizers val TY_DISJ = Logic_ProofSteps.TY_DISJ diff --git a/thys/Auto2_HOL/matcher.ML b/thys/Auto2_HOL/matcher.ML --- a/thys/Auto2_HOL/matcher.ML +++ b/thys/Auto2_HOL/matcher.ML @@ -1,409 +1,409 @@ (* File: matcher.ML Author: Bohua Zhan Matching up to equivalence (E-matching) using a rewrite table. *) signature MATCHER = sig (* Internal *) val check_type_term: theory -> term * term -> id_inst -> (id_inst * term) option val check_type: theory -> typ * typ -> id_inst -> id_inst option val update_inst: term list -> indexname -> cterm -> id_inst -> id_inst_th list (* THe actual matching function. These are defined together. *) val match_list: Proof.context -> term list -> term list * cterm list -> id_inst -> id_inst_ths list val match_comb: Proof.context -> term list -> term * cterm -> id_inst -> id_inst_th list val match_head: Proof.context -> term list -> term * cterm -> id_inst -> id_inst_th list val match_all_head: Proof.context -> term list -> term * cterm -> id_inst -> id_inst_th list val match: Proof.context -> term list -> term * cterm -> id_inst -> id_inst_th list (* Defined in terms of match. *) val rewrite_match: Proof.context -> term * cterm -> id_inst -> id_inst_th list val rewrite_match_head: Proof.context -> term * cterm -> id_inst -> id_inst_th list val rewrite_match_list: Proof.context -> (bool * (term * cterm)) list -> id_inst -> id_inst_ths list val rewrite_match_subset: Proof.context -> term list -> term list * cterm list -> id_inst -> id_inst_ths list (* Prematching. *) val pre_match_type: Proof.context -> typ * typ -> bool val pre_match_comb: Proof.context -> term * cterm -> bool val pre_match_head': Proof.context -> term * cterm -> bool val pre_match_all_head: Proof.context -> term * cterm -> bool val pre_match: Proof.context -> term * cterm -> bool val pre_match_head: Proof.context -> term * cterm -> bool end; structure Matcher : MATCHER = struct fun compare_inst (((id1, inst1), _), ((id2, inst2), _)) = eq_list (op =) (id1, id2) andalso Util.eq_env (inst1, inst2) (* Match type at the top level for t and u. If there is no match, return NONE. Otherwise, return the updated instsp as well as t instantiated with the new type. *) fun check_type_term thy (t, u) (id, (tyinst, inst)) = let val (T, U) = (fastype_of t, fastype_of u) in if T = U then SOME ((id, (tyinst, inst)), t) else let val tyinst' = Sign.typ_match thy (T, U) tyinst val t' = Envir.subst_term_types tyinst' t in SOME ((id, (tyinst', inst)), t') end end handle Type.TYPE_MATCH => NONE (* Match two types. *) fun check_type thy (T, U) (id, (tyinst, inst)) = let val tyinst' = if T = U then tyinst else Sign.typ_match thy (T, U) tyinst in SOME (id, (tyinst', inst)) end handle Type.TYPE_MATCH => NONE (* Starting here, bd_vars is the list of free variables substituted for bound variables, when matching goes inside abstractions. *) fun is_open bd_vars u = case bd_vars of [] => false | _ => length (inter (op aconv) bd_vars (map Free (Term.add_frees u []))) > 0 (* Assign schematic variable with indexname ixn to u. Type of the schematic variable is determined by the type of u. *) fun update_inst bd_vars ixn cu (id, inst) = let val u = Thm.term_of cu in if is_open bd_vars u then [] else [((id, Util.update_env (ixn, u) inst), Thm.reflexive cu)] end (* Matching an order list of patterns against terms. *) fun match_list ctxt bd_vars (ts, cus) instsp = if null ts andalso null cus then [(instsp, [])] else if null ts orelse null cus then [] else let (* Two choices, one of which should always work (encounter no illegal higher-order patterns. *) fun hd_first () = let val insts_t = match ctxt bd_vars (hd ts, hd cus) instsp fun process_inst_t (instsp', th) = let val insts_ts' = match_list ctxt bd_vars (tl ts, tl cus) instsp' in map (apsnd (cons th)) insts_ts' end in maps process_inst_t insts_t end fun tl_first () = let val insts_ts' = match_list ctxt bd_vars (tl ts, tl cus) instsp fun process_inst_ts' (instsp', ths) = let val insts_t = match ctxt bd_vars (hd ts, hd cus) instsp' in map (apsnd (fn th => th :: ths)) insts_t end in maps process_inst_ts' insts_ts' end in hd_first () handle Fail "invalid pattern" => tl_first () end (* Match a non-AC function. *) and match_comb ctxt bd_vars (t, cu) (instsp as (_, (_, inst))) = let val (tf, targs) = Term.strip_comb t val (cuf, cuargs) = Drule.strip_comb cu val uf = Thm.term_of cuf in if Term.aconv_untyped (tf, uf) then let val instsps' = match_list ctxt bd_vars (targs, cuargs) instsp fun process_inst (instsp', ths) = (instsp', Util.comb_equiv (cuf, ths)) in map process_inst instsps' end else if is_Var tf then let val (ixn, _) = Term.dest_Var tf in case Vartab.lookup inst ixn of NONE => if subset (op aconv) (targs, bd_vars) andalso not (has_duplicates (op aconv) targs) then let val cu' = cu |> Thm.term_of |> fold Util.lambda_abstract (rev targs) |> Thm.cterm_of ctxt in map (fn (instsp', _) => (instsp', Thm.reflexive cu)) (update_inst bd_vars ixn cu' instsp) end else raise Fail "invalid pattern" | SOME (_, tf') => let val t' = Term.list_comb (tf', targs) |> Envir.beta_norm in match ctxt bd_vars (t', cu) instsp end end else [] end (* Match t and u at head. *) and match_head ctxt bd_vars (t, cu) (instsp as (_, (_, inst))) = let val u = Thm.term_of cu in if fastype_of t <> fastype_of u then [] else case (t, u) of (Var ((x, i), _), _) => (case Vartab.lookup inst (x, i) of NONE => if x = "NUMC" andalso not (Consts.is_const_ctxt ctxt u) then [] else if x = "FREE" andalso not (Term.is_Free u) then [] else update_inst bd_vars (x, i) cu instsp | SOME (_, u') => match_head ctxt bd_vars (u', cu) instsp) | (Free (a, _), Free (b, _)) => if a = b then [(instsp, Thm.reflexive cu)] else [] | (Const (a, _), Const (b, _)) => if a = b then [(instsp, Thm.reflexive cu)] else [] | (_ $ _, _) => match_comb ctxt bd_vars (t, cu) instsp | _ => [] end (* With fixed t, match with all equivalences of u. *) and match_all_head ctxt bd_vars (t, cu) (id, env) = let val u = Thm.term_of cu val u_equivs = if is_open bd_vars u then [(id, Thm.reflexive cu)] else RewriteTable.get_head_equiv_with_t ctxt (id, cu) t fun process_equiv (id', eq_u) = let val cu' = Thm.rhs_of eq_u val insts_u' = match_head ctxt bd_vars (t, cu') (id', env) fun process_inst ((id', env'), eq_th) = (* eq_th: t(env') == u', eq_u: u == u'. *) ((id', env'), Util.transitive_list [eq_th, meta_sym eq_u]) in map process_inst insts_u' end in maps process_equiv u_equivs end (* Match t and u, possibly by rewriting u at head. *) and match ctxt bd_vars (t, cu) (instsp as (_, (_, inst))) = case check_type_term (Proof_Context.theory_of ctxt) (t, Thm.term_of cu) instsp of NONE => [] | SOME (instsp', t') => case (t', Thm.term_of cu) of (Var ((x, i), _), _) => (case Vartab.lookup inst (x, i) of NONE => if member (op =) ["NUMC", "FREE"] x then match_all_head ctxt bd_vars (t', cu) instsp' else update_inst bd_vars (x, i) cu instsp' | SOME (_, u') => match ctxt bd_vars (u', cu) instsp') | (Abs (_, T, t'), Abs (x, U, _)) => ( case check_type (Proof_Context.theory_of ctxt) (T, U) instsp' of NONE => [] | SOME (instsp'' as (_, (tyinst', _))) => let - val (cv, cu') = Thm.dest_abs (SOME Name.uu) cu + val (cv, cu') = Thm.dest_abs_global cu val v = Thm.term_of cv val t'' = Envir.subst_term_types tyinst' t' val t_free = Term.subst_bound (v, t'') val insts' = match ctxt (v :: bd_vars) (t_free, cu') instsp'' fun process_inst (inst', th') = (inst', Thm.abstract_rule x cv th') in map process_inst insts' end) | _ => (* Free, Const, and comb case *) (match_all_head ctxt bd_vars (t', cu) instsp') |> distinct compare_inst (* Function for export *) fun rewrite_match_gen at_head ctxt (t, cu) (id, env) = if at_head then case check_type_term (Proof_Context.theory_of ctxt) (t, Thm.term_of cu) (id, env) of NONE => [] | SOME ((id, env), t) => match_head ctxt [] (t, cu) (id, env) else match ctxt [] (t, cu) (id, env) val rewrite_match = rewrite_match_gen false val rewrite_match_head = rewrite_match_gen true (* pairs is a list of (at_head, (t, u)). Match the pairs in sequence, and return a list of ((id, inst), ths), where ths is the list of equalities t(env) == u. *) fun rewrite_match_list ctxt pairs instsp = case pairs of [] => [(instsp, [])] | (at_head, (t, cu)) :: pairs' => let val insts_t = rewrite_match_gen at_head ctxt (t, cu) instsp fun process_inst_t (instsp', th) = let val insts_ts' = rewrite_match_list ctxt pairs' instsp' in map (apsnd (cons th)) insts_ts' end in maps process_inst_t insts_t end (* Given two lists of terms (ts, us), match ts with a subset of us. Return a list of ((id, inst), ths), where ths is the list of equalities t_i(env) == u_j. *) fun rewrite_match_subset ctxt bd_vars (ts, cus) instsp = case ts of [] => [(instsp, [])] | t :: ts' => let fun match_i i = map (pair i) (match ctxt bd_vars (t, nth cus i) instsp) fun process_match_i (i, (instsp', th)) = let val insts_ts' = rewrite_match_subset ctxt bd_vars (ts', nth_drop i cus) instsp' in map (apsnd (cons th)) insts_ts' end in (0 upto (length cus - 1)) |> maps match_i |> maps process_match_i end handle Fail "invalid pattern" => if length ts' > 0 andalso Term_Ord.term_ord (t, hd ts') = LESS then rewrite_match_subset ctxt bd_vars (ts' @ [t], cus) instsp else raise Fail "rewrite_match_subset: invalid pattern" (* Fast function for determining whether there can be a match between t and u. *) fun pre_match_type ctxt (T, U) = let val thy = Proof_Context.theory_of ctxt val _ = Sign.typ_match thy (T, U) Vartab.empty in true end handle Type.TYPE_MATCH => false fun pre_match_comb ctxt (t, cu) = let val (tf, targs) = Term.strip_comb t val (cuf, cuargs) = Drule.strip_comb cu val uf = Thm.term_of cuf in is_Var tf orelse (Term.aconv_untyped (tf, uf) andalso length targs = length cuargs andalso forall (pre_match ctxt) (targs ~~ cuargs)) end and pre_match_head' ctxt (t, cu) = let val u = Thm.term_of cu in case (t, u) of (Var ((x, _), T), _) => if Term.is_open u then false else if not (pre_match_type ctxt (T, fastype_of u)) then false else if x = "NUMC" then Consts.is_const_ctxt ctxt u else if x = "FREE" then Term.is_Free u else true | (Free (a, T), Free (b, U)) => (a = b andalso pre_match_type ctxt (T, U)) | (Const (a, T), Const (b, U)) => (a = b andalso pre_match_type ctxt (T, U)) | (_ $ _, _) => pre_match_comb ctxt (t, cu) | _ => false end and pre_match_all_head ctxt (t, cu) = let val thy = Proof_Context.theory_of ctxt val u = Thm.term_of cu val tyinst = Sign.typ_match thy (fastype_of t, fastype_of u) Vartab.empty val t' = Envir.subst_term_types tyinst t in if Term.is_open u then pre_match_head' ctxt (t', cu) else let val u_equivs = (RewriteTable.get_head_equiv_with_t ctxt ([], cu) t') |> map snd |> map Thm.rhs_of in exists (fn cu' => pre_match_head' ctxt (t', cu')) u_equivs end end handle Type.TYPE_MATCH => false and pre_match ctxt (t, cu) = let val u = Thm.term_of cu in case (t, u) of (Var ((x, _), T), _) => if Term.is_open u then false else if not (pre_match_type ctxt (T, fastype_of u)) then false else if member (op =) ["NUMC", "FREE"] x then pre_match_all_head ctxt (t, cu) else true | (Abs (_, T, t'), Abs (_, U, _)) => if not (pre_match_type ctxt (T, U)) then false else let - val (cv, cu') = Thm.dest_abs (SOME Name.uu) cu + val (cv, cu') = Thm.dest_abs_global cu val t'' = subst_bound (Thm.term_of cv, t') in pre_match ctxt (t'', cu') end | (Bound i, Bound j) => (i = j) | (_, _) => pre_match_all_head ctxt (t, cu) end fun pre_match_head ctxt (t, cu) = let val thy = Proof_Context.theory_of ctxt val u = Thm.term_of cu val tyinst = Sign.typ_match thy (fastype_of t, fastype_of u) Vartab.empty val t' = Envir.subst_term_types tyinst t in pre_match_head' ctxt (t', cu) end handle Type.TYPE_MATCH => false end (* structure Matcher. *) diff --git a/thys/Auto2_HOL/util.ML b/thys/Auto2_HOL/util.ML --- a/thys/Auto2_HOL/util.ML +++ b/thys/Auto2_HOL/util.ML @@ -1,926 +1,925 @@ (* File: util.ML Author: Bohua Zhan Utility functions. *) signature BASIC_UTIL = sig (* Exceptions *) val assert: bool -> string -> unit (* Types *) val propT: typ (* Lists *) val the_pair: 'a list -> 'a * 'a val the_triple: 'a list -> 'a * 'a * 'a val filter_split: ('a -> bool) -> 'a list -> 'a list * 'a list (* Managing matching environments. *) val fo_init: Type.tyenv * Envir.tenv val lookup_instn: Type.tyenv * Envir.tenv -> string * int -> term val lookup_inst: Type.tyenv * Envir.tenv -> string -> term (* Tracing functions. *) val trace_t: Proof.context -> string -> term -> unit val trace_tlist: Proof.context -> string -> term list -> unit val trace_thm: Proof.context -> string -> thm -> unit val trace_fullthm: Proof.context -> string -> thm -> unit val trace_thm_global: string -> thm -> unit val trace_fullthm_global: string -> thm -> unit (* Terms *) val dest_arg: term -> term val dest_arg1: term -> term (* Theorems *) val apply_to_thm: conv -> thm -> thm val meta_sym: thm -> thm val apply_to_lhs: conv -> thm -> thm val apply_to_rhs: conv -> thm -> thm end; signature UTIL = sig include BASIC_UTIL (* Lists *) val max: ('a * 'a -> order) -> 'a list -> 'a val max_partial: ('a -> 'a -> bool) -> 'a list -> 'a list val subsets: 'a list -> 'a list list val all_permutes: 'a list -> 'a list list val all_pairs: 'a list * 'b list -> ('a * 'b) list val remove_dup_lists: ('a * 'a -> order) -> 'a list * 'a list -> 'a list * 'a list val is_subseq: ('a * 'a -> bool) -> 'a list * 'a list -> bool (* Strings. *) val is_prefix_str: string -> string -> bool val is_just_internal: string -> bool (* Managing matching environments. *) val defined_instn: Type.tyenv * Envir.tenv -> string * int -> bool val lookup_tyinst: Type.tyenv * Envir.tenv -> string -> typ val update_env: indexname * term -> Type.tyenv * Envir.tenv -> Type.tyenv * Envir.tenv val eq_env: (Type.tyenv * Envir.tenv) * (Type.tyenv * Envir.tenv) -> bool (* Matching. *) val first_order_match_list: theory -> (term * term) list -> Type.tyenv * Envir.tenv -> Type.tyenv * Envir.tenv (* Printing functions, mostly from Isabelle Cookbook. *) val string_of_terms: Proof.context -> term list -> string val string_of_terms_global: theory -> term list -> string val string_of_tyenv: Proof.context -> Type.tyenv -> string val string_of_env: Proof.context -> Envir.tenv -> string val string_of_list: ('a -> string) -> 'a list -> string val string_of_list': ('a -> string) -> 'a list -> string val string_of_bool: bool -> string (* Managing context. *) val declare_free_term: term -> Proof.context -> Proof.context (* Terms. *) val is_abs: term -> bool val is_implies: term -> bool val dest_binop: term -> term * (term * term) val dest_binop_head: term -> term val dest_binop_args: term -> term * term val dest_args: term -> term list val dest_argn: int -> term -> term val get_head_name: term -> string val is_meta_eq: term -> bool val occurs_frees: term list -> term -> bool val occurs_free: term -> term -> bool val has_vars: term -> bool val is_subterm: term -> term -> bool val has_subterm: term list -> term -> bool val is_head: term -> term -> bool val lambda_abstract: term -> term -> term val is_pattern_list: term list -> bool val is_pattern: term -> bool val normalize_meta_all_imp: Proof.context -> conv val swap_meta_imp_alls: Proof.context -> conv val normalize_meta_horn: Proof.context -> conv val strip_meta_horn: term -> term list * (term list * term) val list_meta_horn: term list * (term list * term) -> term val to_internal_vars: Proof.context -> term list * term -> term list * term val rename_abs_term: term list -> term -> term val print_term_detail: Proof.context -> term -> string (* cterms. *) val dest_cargs: cterm -> cterm list val dest_binop_cargs: cterm -> cterm * cterm (* Theorems. *) val arg_backn_conv: int -> conv -> conv val argn_conv: int -> conv -> conv val comb_equiv: cterm * thm list -> thm val name_of_thm: thm -> string val update_name_of_thm: thm -> string -> thm -> thm val lhs_of: thm -> term val rhs_of: thm -> term val assume_meta_eq: theory -> term * term -> thm val assume_thm: Proof.context -> term -> thm val subst_thm_thy: theory -> Type.tyenv * Envir.tenv -> thm -> thm val subst_thm: Proof.context -> Type.tyenv * Envir.tenv -> thm -> thm val send_first_to_hyps: thm -> thm val send_all_to_hyps: thm -> thm val subst_thm_atomic: (cterm * cterm) list -> thm -> thm val subst_term_norm: Type.tyenv * Envir.tenv -> term -> term val concl_conv_n: int -> conv -> conv val concl_conv: conv -> conv val transitive_list: thm list -> thm val skip_n_conv: int -> conv -> conv val pattern_rewr_conv: term -> (term * thm) list -> conv val eq_cong_th: int -> term -> term -> Proof.context -> thm val forall_elim_sch: thm -> thm (* Conversions *) val reverse_eta_conv: Proof.context -> conv val repeat_n_conv: int -> conv -> conv (* Miscellaneous. *) val test_conv: Proof.context -> conv -> string -> string * string -> unit val term_pat_setup: theory -> theory val cterm_pat_setup: theory -> theory val type_pat_setup: theory -> theory val timer: string * (unit -> 'a) -> 'a val exn_trace: (unit -> 'a) -> 'a end; structure Util : UTIL = struct fun assert b exn_str = if b then () else raise Fail exn_str val propT = @{typ prop} fun max comp lst = let fun max2 t1 t2 = if comp (t1, t2) = LESS then t2 else t1 in case lst of [] => raise Fail "max: empty list" | l :: ls => fold max2 ls l end (* Given a function comp, remove y for each pair (x, y) such that comp x y = true (if x dominates y). *) fun max_partial comp lst = let fun helper taken remains = case remains of [] => taken | x :: xs => if exists (fn y => comp y x) taken then helper taken xs else helper (x :: filter_out (fn y => comp x y) taken) xs in helper [] lst end (* Return all subsets of lst. *) fun subsets [] = [[]] | subsets (l::ls) = let val prev = subsets ls in prev @ map (cons l) prev end (* List of all permutations of xs *) fun all_permutes xs = case xs of [] => [[]] | [x] => [[x]] | [x, y] => [[x, y], [y, x]] | _ => maps (fn i => map (cons (nth xs i)) (all_permutes (nth_drop i xs))) (0 upto (length xs - 1)) (* Convert list to pair. List must consist of exactly two items. *) fun the_pair lst = case lst of [i1, i2] => (i1, i2) | _ => raise Fail "the_pair" (* Convert list to triple. List must consist of exactly three items. *) fun the_triple lst = case lst of [i1, i2, i3] => (i1, i2, i3) | _ => raise Fail "the_triple" (* Split list into (ins, outs), where ins satisfy f, and outs don't. *) fun filter_split _ [] = ([], []) | filter_split f (x :: xs) = let val (ins, outs) = filter_split f xs in if f x then (x :: ins, outs) else (ins, x :: outs) end (* Form the Cartesian product of two lists. *) fun all_pairs (l1, l2) = maps (fn y => (map (fn x => (x, y)) l1)) l2 (* Given two sorted lists, remove all pairs of terms that appear in both lists, counting multiplicity. *) fun remove_dup_lists ord (xs, ys) = case xs of [] => ([], ys) | x :: xs' => case ys of [] => (xs, []) | y :: ys' => case ord (x, y) of LESS => apfst (cons x) (remove_dup_lists ord (xs', ys)) | EQUAL => remove_dup_lists ord (xs', ys') | GREATER => apsnd (cons y) (remove_dup_lists ord (xs, ys')) (* Whether l1 is a subsequence of l2. *) fun is_subseq eq (l1, l2) = case l1 of [] => true | x :: l1' => case l2 of [] => false | y :: l2' => if eq (x, y) then is_subseq eq (l1', l2') else is_subseq eq (l1, l2') (* Whether pre is a prefix of str. *) fun is_prefix_str pre str = is_prefix (op =) (String.explode pre) (String.explode str) (* Test whether x is followed by exactly one _. *) fun is_just_internal x = Name.is_internal x andalso not (Name.is_skolem x) val fo_init = (Vartab.empty, Vartab.empty) (* Lookup a Vartab inst with string and integer specifying indexname. *) fun defined_instn (_, inst) (str, n) = Vartab.defined inst (str, n) fun lookup_instn (_, inst) (str, n) = case Vartab.lookup inst (str, n) of NONE => raise Fail ("lookup_inst: not found " ^ str ^ (if n = 0 then "" else string_of_int n)) | SOME (_, u) => u fun lookup_inst (tyinst, inst) str = lookup_instn (tyinst, inst) (str, 0) fun lookup_tyinst (tyinst, _) str = case Vartab.lookup tyinst (str, 0) of NONE => raise Fail ("lookup_tyinst: not found " ^ str) | SOME (_, T) => T fun update_env (idx, t) (tyenv, tenv) = (tyenv, tenv |> Vartab.update_new (idx, (type_of t, t))) (* A rough comparison, simply compare the corresponding terms. *) fun eq_env ((_, inst1), (_, inst2)) = let val data1 = Vartab.dest inst1 val data2 = Vartab.dest inst2 fun compare_data (((x, i), (ty, t)), ((x', i'), (ty', t'))) = x = x' andalso i = i' andalso ty = ty' andalso t aconv t' in eq_set compare_data (data1, data2) end fun first_order_match_list thy pairs inst = case pairs of [] => inst | (t, u) :: rest => let val inst' = Pattern.first_order_match thy (t, u) inst in first_order_match_list thy rest inst' end fun string_of_terms ctxt ts = ts |> map (Syntax.pretty_term ctxt) |> Pretty.commas |> Pretty.block |> Pretty.string_of fun string_of_terms_global thy ts = ts |> map (Syntax.pretty_term_global thy) |> Pretty.commas |> Pretty.block |> Pretty.string_of fun pretty_helper aux env = env |> Vartab.dest |> map aux |> map (fn (s1, s2) => Pretty.block [s1, Pretty.str " := ", s2]) |> Pretty.enum "," "[" "]" |> Pretty.string_of fun string_of_tyenv ctxt tyenv = let fun get_typs (v, (s, T)) = (TVar (v, s), T) val print = apply2 (Syntax.pretty_typ ctxt) in pretty_helper (print o get_typs) tyenv end fun string_of_env ctxt env = let fun get_ts (v, (T, t)) = (Var (v, T), t) val print = apply2 (Syntax.pretty_term ctxt) in pretty_helper (print o get_ts) env end fun string_of_list func lst = Pretty.str_list "[" "]" (map func lst) |> Pretty.string_of fun string_of_list' func lst = if length lst = 1 then func (the_single lst) else string_of_list func lst fun string_of_bool b = if b then "true" else "false" fun trace_t ctxt s t = tracing (s ^ " " ^ (Syntax.string_of_term ctxt t)) fun trace_tlist ctxt s ts = tracing (s ^ " " ^ (string_of_terms ctxt ts)) fun trace_thm ctxt s th = tracing (s ^ " " ^ (th |> Thm.prop_of |> Syntax.string_of_term ctxt)) fun trace_fullthm ctxt s th = tracing (s ^ " [" ^ (Thm.hyps_of th |> string_of_terms ctxt) ^ "] ==> " ^ (Thm.prop_of th |> Syntax.string_of_term ctxt)) fun trace_thm_global s th = let val thy = Thm.theory_of_thm th in tracing (s ^ " " ^ (th |> Thm.prop_of |> Syntax.string_of_term_global thy)) end fun trace_fullthm_global s th = let val thy = Thm.theory_of_thm th in tracing (s ^ " [" ^ (Thm.hyps_of th |> string_of_terms_global thy) ^ "] ==> " ^ (Thm.prop_of th |> Syntax.string_of_term_global thy)) end fun declare_free_term t ctxt = if not (is_Free t) then raise Fail "declare_free_term: t not free." else ctxt |> Variable.add_fixes_direct [t |> Term.dest_Free |> fst] |> Variable.declare_term t fun is_abs t = case t of Abs _ => true | _ => false (* Whether a given term is of the form A ==> B. *) fun is_implies t = let val _ = assert (fastype_of t = propT) "is_implies: wrong type" in case t of @{const Pure.imp} $ _ $ _ => true | _ => false end (* Extract the last two arguments on t, collecting the rest into f. *) fun dest_binop t = case t of f $ a $ b => (f, (a, b)) | _ => raise Fail "dest_binop" fun dest_binop_head t = fst (dest_binop t) fun dest_binop_args t = snd (dest_binop t) (* Return the argument of t. If t is f applied to multiple arguments, return the last argument. *) fun dest_arg t = case t of _ $ arg => arg | _ => raise Fail "dest_arg" (* Return the first of two arguments of t. *) fun dest_arg1 t = case t of _ $ arg1 $ _ => arg1 | _ => raise Fail "dest_arg1" (* Return the list of all arguments of t. *) fun dest_args t = t |> Term.strip_comb |> snd (* Return the nth argument of t, counting from left and starting at zero. *) fun dest_argn n t = nth (dest_args t) n (* Return the name of the head function. *) fun get_head_name t = case Term.head_of t of Const (nm, _) => nm | _ => raise Fail "get_head_name" (* Whether the term is of the form A == B. *) fun is_meta_eq t = let val _ = assert (fastype_of t = propT) "is_meta_eq_term: wrong type" in case t of Const (@{const_name Pure.eq}, _) $ _ $ _ => true | _ => false end (* Given free variable freevar (or a list of free variables freevars), determine whether any of the inputs appears in t. *) fun occurs_frees freevars t = inter (op aconv) (map Free (Term.add_frees t [])) freevars <> [] fun occurs_free freevar t = occurs_frees [freevar] t (* Whether the given term contains schematic variables. *) fun has_vars t = length (Term.add_vars t []) > 0 (* Whether subt is a subterm of t. *) fun is_subterm subt t = exists_subterm (fn t' => t' aconv subt) t (* Whether any of subts is a subterm of t. *) fun has_subterm subts t = exists_subterm (fn t' => member (op aconv) subts t') t (* Whether s is a head term of t. *) fun is_head s t = let val (sf, sargs) = Term.strip_comb s val (tf, targs) = Term.strip_comb t in sf aconv tf andalso is_prefix (op aconv) sargs targs end (* If stmt is P(t), return %t. P(t). *) fun lambda_abstract t stmt = Term.lambda t (Term.abstract_over (t, stmt)) (* A more general criterion for patterns. In combinations, any argument that is a pattern (in the more general sense) frees up checking for any functional schematic variables in that argument. *) fun is_pattern_list ts = let fun is_funT T = case T of Type ("fun", _) => true | _ => false fun get_fun_vars t = (Term.add_vars t []) |> filter (is_funT o snd) |> map Var fun test_list exclude_vars ts = case ts of [] => true | [t] => test_term exclude_vars t | t :: ts' => if test_term exclude_vars t then test_list (merge (op aconv) (exclude_vars, get_fun_vars t)) ts' else if test_list exclude_vars ts' then test_term (distinct (op aconv) (exclude_vars @ maps get_fun_vars ts')) t else false and test_term exclude_vars t = case t of Abs (_, _, t') => test_term exclude_vars t' | _ => let val (head, args) = strip_comb t in if is_Var head then if member (op aconv) exclude_vars head then test_list exclude_vars args else forall is_Bound args andalso not (has_duplicates (op aconv) args) else test_list exclude_vars args end in test_list [] ts end fun is_pattern t = is_pattern_list [t] (* Push !!x to the right as much as possible. *) fun normalize_meta_all_imp_once ct = Conv.try_conv ( Conv.every_conv [Conv.rewr_conv (Thm.symmetric @{thm Pure.norm_hhf_eq}), Conv.arg_conv normalize_meta_all_imp_once]) ct fun normalize_meta_all_imp ctxt ct = let val t = Thm.term_of ct in if Logic.is_all t then Conv.every_conv [Conv.binder_conv (normalize_meta_all_imp o snd) ctxt, normalize_meta_all_imp_once] ct else if is_implies t then Conv.arg_conv (normalize_meta_all_imp ctxt) ct else Conv.all_conv ct end (* Rewrite A ==> !!v_i. B to !!v_i. A ==> B. *) fun swap_meta_imp_alls ctxt ct = let val t = Thm.term_of ct in if is_implies t andalso Logic.is_all (dest_arg t) then Conv.every_conv [Conv.rewr_conv @{thm Pure.norm_hhf_eq}, Conv.binder_conv (swap_meta_imp_alls o snd) ctxt] ct else Conv.all_conv ct end (* Normalize a horn clause into standard form !!v_i. A_i ==> B. *) fun normalize_meta_horn ctxt ct = let val t = Thm.term_of ct in if Logic.is_all t then Conv.binder_conv (normalize_meta_horn o snd) ctxt ct else if is_implies t then Conv.every_conv [Conv.arg_conv (normalize_meta_horn ctxt), swap_meta_imp_alls ctxt] ct else Conv.all_conv ct end (* Deconstruct a horn clause !!v_i. A_i ==> B into (v_i, (A_i, B)). *) fun strip_meta_horn t = case t of - Const (@{const_name Pure.all}, _) $ Abs (abs as (_, T, _)) => + Const (@{const_name Pure.all}, _) $ (u as Abs _) => let - val (x, body) = Term.dest_abs abs - val var = Free (x, T) + val (v, body) = Term.dest_abs_global u val (vars, (assums, concl)) = strip_meta_horn body in - (var :: vars, (assums, concl)) + (Free v :: vars, (assums, concl)) end | @{const Pure.imp} $ P $ Q => let val (vars, (assums, concl)) = strip_meta_horn Q in (vars, (P :: assums, concl)) end | _ => ([], ([], t)) fun list_meta_horn (vars, (As, B)) = (Logic.list_implies (As, B)) |> fold Logic.all (rev vars) fun to_internal_vars ctxt (vars, body) = let val vars' = vars |> map Term.dest_Free |> Variable.variant_frees ctxt [] |> map (apfst Name.internal) |> map Free val subst = vars ~~ vars' in (vars', subst_atomic subst body) end fun rename_abs_term vars t = case vars of [] => t | var :: rest => let val (x, _) = Term.dest_Free var in case t of A $ Abs (_, T1, body) => A $ Abs (x, T1, rename_abs_term rest body) | _ => error "rename_abs_term" end fun print_term_detail ctxt t = case t of Const (s, ty) => "Const (" ^ s ^ ", " ^ (Syntax.string_of_typ ctxt ty) | Free (s, ty) => "Free (" ^ s ^ ", " ^ (Syntax.string_of_typ ctxt ty) ^ ")" | Var ((x, i), ty) => "Var ((" ^ x ^ ", " ^ (string_of_int i) ^ "), " ^ (Syntax.string_of_typ ctxt ty) ^ ")" | Bound n => "Bound " ^ (string_of_int n) | Abs (s, ty, b) => "Abs (" ^ s ^ ", " ^ (Syntax.string_of_typ ctxt ty) ^ ", " ^ (print_term_detail ctxt b) ^ ")" | u $ v => "(" ^ print_term_detail ctxt u ^ ") $ (" ^ print_term_detail ctxt v ^ ")" (* Version for ct. *) fun dest_cargs ct = ct |> Drule.strip_comb |> snd fun dest_binop_cargs ct = (Thm.dest_arg1 ct, Thm.dest_arg ct) (* Apply cv to nth argument of t, counting from right and starting at 0. *) fun arg_backn_conv n cv ct = if n = 0 then Conv.arg_conv cv ct else Conv.fun_conv (arg_backn_conv (n-1) cv) ct (* Apply cv to nth argument of t, counting from left and starting at 0. *) fun argn_conv n cv ct = let val args_count = ct |> Thm.term_of |> dest_args |> length val _ = assert (n >= 0 andalso n < args_count) in arg_backn_conv (args_count - n - 1) cv ct end (* Given a head cterm f (function to be applied), and a list of equivalence theorems of arguments, produce an equivalent theorem for the overall term. *) fun comb_equiv (cf, arg_equivs) = Library.foldl (uncurry Thm.combination) (Thm.reflexive cf, arg_equivs) (* Retrive name of theorem. *) fun name_of_thm th = if Thm.has_name_hint th then Thm.get_name_hint th else raise Fail "name_of_thm: not found" (* Set the name of th to the name of ori_th, followed by suffix. *) fun update_name_of_thm ori_th suffix th = if Thm.has_name_hint ori_th then th |> Thm.put_name_hint (Thm.get_name_hint ori_th ^ suffix) else th val lhs_of = Thm.term_of o Thm.lhs_of val rhs_of = Thm.term_of o Thm.rhs_of fun assume_meta_eq thy (t1, t2) = Thm.assume (Thm.global_cterm_of thy (Logic.mk_equals (t1, t2))) fun assume_thm ctxt t = if type_of t <> propT then raise Fail "assume_thm: t is not of type prop" else Thm.assume (Thm.cterm_of ctxt t) (* Similar to Envir.subst_term. Apply an instantiation to a theorem. *) fun subst_thm_thy thy (tyinsts, insts) th = let fun process_tyenv (v, (S, T)) = ((v, S), Thm.global_ctyp_of thy T) val tys = map process_tyenv (Vartab.dest tyinsts) fun process_tenv (v, (T, u)) = ((v, Envir.subst_type tyinsts T), Thm.global_cterm_of thy u) val ts = map process_tenv (Vartab.dest insts) in th |> Drule.instantiate_normalize (TVars.make tys, Vars.make ts) end fun subst_thm ctxt (tyinsts, insts) th = subst_thm_thy (Proof_Context.theory_of ctxt) (tyinsts, insts) th fun send_first_to_hyps th = let val cprem = Thm.cprem_of th 1 in Thm.implies_elim th (Thm.assume cprem) end fun send_all_to_hyps th = let val _ = assert (forall (not o has_vars) (Thm.prems_of th)) "send_all_to_hyps: schematic variables in hyps." in funpow (Thm.nprems_of th) send_first_to_hyps th end (* Replace using subst the internal variables in th. This proceeds in several steps: first, pull any hypotheses of the theorem involving the replaced variables into statement of the theorem, perform the replacement (using forall_intr then forall_elim), finally return the hypotheses to their original place. *) fun subst_thm_atomic subst th = let val old_cts = map fst subst val old_ts = map Thm.term_of old_cts val new_cts = map snd subst val chyps = filter (fn ct => has_subterm old_ts (Thm.term_of ct)) (Thm.chyps_of th) in th |> fold Thm.implies_intr chyps |> fold Thm.forall_intr old_cts |> fold Thm.forall_elim (rev new_cts) |> funpow (length chyps) send_first_to_hyps end (* Substitution into terms used in auto2. Substitute types first and instantiate the types in the table of term instantiations. Also perform beta_norm at the end. *) fun subst_term_norm (tyinsts, insts) t = let fun inst_tenv tenv = tenv |> Vartab.dest |> map (fn (ixn, (T, v)) => (ixn, (Envir.subst_type tyinsts T, v))) |> Vartab.make in t |> Envir.subst_term_types tyinsts |> Envir.subst_term (tyinsts, inst_tenv insts) |> Envir.beta_norm end (* Apply the conversion cv to the statement of th, yielding the equivalent theorem. *) fun apply_to_thm cv th = let val eq = cv (Thm.cprop_of th) in if Thm.is_reflexive eq then th else Thm.equal_elim eq th end (* Given th of form A == B, get th' of form B == A. *) val meta_sym = Thm.symmetric (* Apply conv to rewrite the left hand side of th. *) fun apply_to_lhs cv th = let val eq = cv (Thm.lhs_of th) in if Thm.is_reflexive eq then th else Thm.transitive (meta_sym eq) th end (* Apply conv to rewrite the right hand side of th. *) fun apply_to_rhs cv th = let val eq = cv (Thm.rhs_of th) in if Thm.is_reflexive eq then th else Thm.transitive th eq end (* Using cv, rewrite the part of ct after stripping i premises. *) fun concl_conv_n i cv ct = if i = 0 then cv ct else (Conv.arg_conv (concl_conv_n (i-1) cv)) ct (* Rewrite part of ct after stripping all premises. *) fun concl_conv cv ct = case Thm.term_of ct of @{const Pure.imp} $ _ $ _ => Conv.arg_conv (concl_conv cv) ct | _ => cv ct (* Given a list of theorems A = B, B = C, etc., apply Thm.transitive to get equality between start and end. *) fun transitive_list ths = let fun rev_transitive btoc atob = let val (b, c) = btoc |> Thm.cprop_of |> Thm.dest_equals val (a, b') = atob |> Thm.cprop_of |> Thm.dest_equals in if b aconvc b' then if a aconvc b then btoc else if b aconvc c then atob else Thm.transitive atob btoc else let val _ = map (trace_thm_global "ths:") ths in raise Fail "transitive_list: intermediate does not agree" end end in fold rev_transitive (tl ths) (hd ths) end (* Skip to argument n times. For example, if applied to rewrite a proposition in implication form (==> or -->), it will skip the first n assumptions. *) fun skip_n_conv n cv = if n <= 0 then cv else Conv.arg_conv (skip_n_conv (n-1) cv) (* Given a term t, and pairs (a_i, eq_i), where a_i's are atomic subterms of t. Suppose the input is obtained by replacing each a_i by the left side of eq_i in t, obtain equality from t to the term obtained by replacing each a_i by the right side of eq_i in t. *) fun pattern_rewr_conv t eq_ths ct = case t of Bound _ => raise Fail "pattern_rewr_conv: bound variable." | Abs _ => raise Fail "pattern_rewr_conv: abs not supported." | _ $ _ => let val (f, arg) = Term.strip_comb t val (f', arg') = Drule.strip_comb ct val _ = assert (f aconv Thm.term_of f') "pattern_rewr_conv: input does not match pattern." val ths = map (fn (t, ct) => pattern_rewr_conv t eq_ths ct) (arg ~~ arg') in comb_equiv (f', ths) end | Const _ => let val _ = assert (t aconv Thm.term_of ct) "pattern_rewr_conv: input does not match pattern." in Conv.all_conv ct end | _ => (* Free and Var cases *) (case AList.lookup (op aconv) eq_ths t of NONE => Conv.all_conv ct | SOME eq_th => let val _ = assert (lhs_of eq_th aconv (Thm.term_of ct)) "pattern_rewr_conv: wrong substitution." in eq_th end) (* Given integer i, term b_i, and a term A = f(a_1, ..., a_n), produce the theorem a_i = b_i ==> A = f(a_1, ..., b_i, ..., a_n). *) fun eq_cong_th i bi A ctxt = let val thy = Proof_Context.theory_of ctxt val (cf, cargs) = Drule.strip_comb (Thm.cterm_of ctxt A) val _ = assert (i < length cargs) "eq_cong_th: i out of bounds." val ai = Thm.term_of (nth cargs i) val eq_i = assume_meta_eq thy (ai, bi) val eqs = map Thm.reflexive (take i cargs) @ [eq_i] @ map Thm.reflexive (drop (i + 1) cargs) val eq_A = comb_equiv (cf, eqs) in Thm.implies_intr (Thm.cprop_of eq_i) eq_A end (* Convert theorems of form !!x y. P x y into P ?x ?y (arbitrary number of quantifiers). *) fun forall_elim_sch th = case Thm.prop_of th of Const (@{const_name Pure.all}, _) $ Abs (x, T, _) => let val var_xs = map fst (Term.add_var_names (Thm.prop_of th) []) val x' = if member (op =) var_xs x then singleton (Name.variant_list var_xs) x else x val thy = Thm.theory_of_thm th in th |> Thm.forall_elim (Thm.global_cterm_of thy (Var ((x', 0), T))) |> forall_elim_sch end | _ => th (* Given P of function type, produce P == %x. P x. *) fun reverse_eta_conv ctxt ct = let val t = Thm.term_of ct val argT = Term.domain_type (fastype_of t) handle Match => raise CTERM ("reverse_eta_conv", [ct]) val rhs = Abs ("x", argT, t $ Bound 0) val eq_th = Thm.eta_conversion (Thm.cterm_of ctxt rhs) in meta_sym eq_th end (* Repeat cv exactly n times. *) fun repeat_n_conv n cv t = if n = 0 then Conv.all_conv t else (cv then_conv (repeat_n_conv (n-1) cv)) t (* Generic function for testing a conv. *) fun test_conv ctxt cv err_str (str1, str2) = let val (t1, t2) = (Proof_Context.read_term_pattern ctxt str1, Proof_Context.read_term_pattern ctxt str2) val th = cv (Thm.cterm_of ctxt t1) in if t1 aconv (lhs_of th) andalso t2 aconv (rhs_of th) then () else let val _ = trace_t ctxt "Input:" t1 val _ = trace_t ctxt "Expected:" t2 val _ = trace_t ctxt "Actual:" (Thm.prop_of th) in raise Fail err_str end end (* term_pat and typ_pat, from Isabelle Cookbook. *) val term_pat_setup = let val parser = Args.context -- Scan.lift Args.embedded_inner_syntax fun term_pat (ctxt, str) = str |> Proof_Context.read_term_pattern ctxt |> ML_Syntax.print_term |> ML_Syntax.atomic in ML_Antiquotation.inline @{binding "term_pat"} (parser >> term_pat) end val cterm_pat_setup = let val parser = Args.context -- Scan.lift Args.embedded_inner_syntax fun cterm_pat (ctxt, str) = str |> Proof_Context.read_term_pattern ctxt |> ML_Syntax.print_term |> ML_Syntax.atomic |> prefix "Thm.cterm_of ML_context" in ML_Antiquotation.value @{binding "cterm_pat"} (parser >> cterm_pat) end val type_pat_setup = let val parser = Args.context -- Scan.lift Args.embedded_inner_syntax fun typ_pat (ctxt, str) = let val ctxt' = Proof_Context.set_mode Proof_Context.mode_schematic ctxt in str |> Syntax.read_typ ctxt' |> ML_Syntax.print_typ |> ML_Syntax.atomic end in ML_Antiquotation.inline @{binding "typ_pat"} (parser >> typ_pat) end (* Time the given function f : unit -> 'a. *) fun timer (msg, f) = let val t_start = Timing.start () val res = f () val t_end = Timing.result t_start in (writeln (msg ^ (Timing.message t_end)); res) end (* When exception is shown when running function f : unit -> 'a, print stack trace. *) fun exn_trace f = Runtime.exn_trace f end (* structure Util. *) structure Basic_Util: BASIC_UTIL = Util open Basic_Util val _ = Theory.setup (Util.term_pat_setup) val _ = Theory.setup (Util.cterm_pat_setup) val _ = Theory.setup (Util.type_pat_setup) diff --git a/thys/Auto2_HOL/util_logic.ML b/thys/Auto2_HOL/util_logic.ML --- a/thys/Auto2_HOL/util_logic.ML +++ b/thys/Auto2_HOL/util_logic.ML @@ -1,630 +1,628 @@ (* File: util_logic.ML Author: Bohua Zhan Utility functions, after fixing object logic. *) signature BASIC_UTIL_LOGIC = sig (* Terms *) val Trueprop: term val is_Trueprop: term -> bool val mk_Trueprop: term -> term val dest_Trueprop: term -> term val Trueprop_conv: conv -> conv val pFalse: term val Not: term val mk_not: term -> term val dest_not: term -> term val is_neg: term -> bool val get_neg: term -> term val get_neg': term -> term val conj: term val is_conj: term -> bool val mk_conj: term * term -> term val strip_conj: term -> term list val list_conj: term list -> term val disj: term val is_disj: term -> bool val mk_disj: term * term -> term val strip_disj: term -> term list val list_disj: term list -> term val imp: term val is_imp: term -> bool val mk_imp: term * term -> term val dest_imp: term -> term * term val strip_obj_imp: term -> term list * term val list_obj_imp: term list * term -> term val is_obj_all: term -> bool val is_ball: term -> bool val mk_obj_all: term -> term -> term val is_ex: term -> bool val is_bex: term -> bool val mk_exists: term -> term -> term val is_mem: term -> bool val mk_mem: term * term -> term (* Theorems *) val is_true_th: thm -> bool val prop_of': thm -> term val cprop_of': thm -> cterm val concl_of': thm -> term val make_trueprop_eq: thm -> thm val assume_eq: theory -> term * term -> thm val apply_to_thm': conv -> thm -> thm val to_meta_eq: thm -> thm val to_obj_eq: thm -> thm val obj_sym: thm -> thm val rewr_obj_eq: thm -> conv val conj_left_th: thm -> thm val conj_right_th: thm -> thm val equiv_forward_th: thm -> thm val equiv_backward_th: thm -> thm val inv_backward_th: thm -> thm val to_obj_eq_th: thm -> thm val to_obj_eq_iff_th: thm -> thm val obj_sym_th: thm -> thm end; signature UTIL_LOGIC = sig include BASIC_UTIL_LOGIC (* Terms *) val term_of_bool: bool -> term val bool_of_term: term -> bool (* Finding subterms *) val list_subterms: term -> term list val get_all_subterms: term -> term list val get_all_subterms_skip_if: term -> term list (* cterms *) val get_cneg: cterm -> cterm (* Theorems *) val make_neg_eq: thm -> thm val rewrite_to_contra_form: conv val rewrite_from_contra_form: conv val to_obj_conv: Proof.context -> conv val to_obj_conv_on_horn: Proof.context -> conv val to_meta_conv: Proof.context -> conv val split_conj_th: thm -> thm list val split_conj_gen_th: thm -> thm list val split_not_disj_th: thm -> thm list val strip_horn': thm -> term list * term val mk_conjs_th: thm list -> thm val ex_elim: Proof.context -> term -> thm -> thm val force_abs_form: term -> term val strip_obj_horn: term -> term list * (term list * term) val list_obj_horn: term list * (term list * term) -> term val is_ex_form_gen: term -> bool val normalize_exists: Proof.context -> conv val strip_exists: term -> term list * term (* Wrapper around common tactics. *) val prove_by_tac: (Proof.context -> int -> tactic) -> Proof.context -> thm list -> term -> thm val contra_by_tac: (Proof.context -> int -> tactic) -> Proof.context -> thm list -> thm end; structure UtilLogic : UTIL_LOGIC = struct (* Booleans *) fun term_of_bool b = (if b then bTrue else bFalse) fun bool_of_term t = if t aconv bTrue then true else if t aconv bFalse then false else raise Fail "bool_of_term: unexpected t." (* Trueprop *) val Trueprop = Const (UtilBase.Trueprop_name, boolT --> propT) (* Returns whether the given term is Trueprop. *) fun is_Trueprop t = let val _ = assert (fastype_of t = propT) "is_Trueprop: wrong type" in case t of Const (c, _) $ _ => c = UtilBase.Trueprop_name | _ => false end fun mk_Trueprop P = Trueprop $ P fun dest_Trueprop t = if is_Trueprop t then dest_arg t else raise Fail "dest_Trueprop" fun Trueprop_conv cv ct = if is_Trueprop (Thm.term_of ct) then Conv.arg_conv cv ct else raise CTERM ("Trueprop_conv", [ct]) val pFalse = Trueprop $ bFalse (* Not *) val Not = Const (UtilBase.Not_name, boolT --> boolT) fun mk_not P = Not $ P (* Returns whether the given term is in neg form. *) fun is_neg t = let val _ = assert (fastype_of t = boolT) "is_neg: wrong type" in case t of Const (c, _) $ _ => c = UtilBase.Not_name | _ => false end fun dest_not t = if is_neg t then dest_arg t else raise Fail "dest_not" (* Returns the negation of the given term. Avoids double negatives. *) fun get_neg t = if is_neg t then dest_not t else Not $ t (* Version of get_neg for Trueprop terms. *) fun get_neg' t = let val _ = assert (is_Trueprop t) "get_neg': input should be a Trueprop." in t |> dest_Trueprop |> get_neg |> mk_Trueprop end (* Conjunction and disjunction *) val conj = Const (UtilBase.Conj_name, boolT --> boolT --> boolT) fun is_conj t = case t of Const (c, _) $ _ $ _ => c = UtilBase.Conj_name | _ => false fun mk_conj (t, u) = conj $ t $ u fun strip_conj t = if is_conj t then (dest_arg1 t) :: strip_conj (dest_arg t) else [t] fun list_conj ts = case ts of [] => error "list_conj" | [t] => t | t :: rest => mk_conj (t, list_conj rest) val disj = Const (UtilBase.Disj_name, boolT --> boolT --> boolT) fun is_disj t = case t of Const (c, _) $ _ $ _ => c = UtilBase.Disj_name | _ => false fun mk_disj (t, u) = disj $ t $ u fun strip_disj t = if is_disj t then (dest_arg1 t) :: strip_disj (dest_arg t) else [t] fun list_disj ts = case ts of [] => raise Fail "list_disj" | [t] => t | t :: ts' => mk_disj (t, list_disj ts') (* Object implication *) val imp = Const (UtilBase.Imp_name, boolT --> boolT --> boolT) fun is_imp t = case t of Const (c, _) $ _ $ _ => c = UtilBase.Imp_name | _ => false fun mk_imp (t, u) = imp $ t $ u fun dest_imp t = if is_imp t then (dest_arg1 t, dest_arg t) else raise Fail "dest_imp" (* Given t of form A1 --> ... --> An, return ([A1, ..., A(n-1)], An). *) fun strip_obj_imp t = if is_imp t then let val (As, B') = strip_obj_imp (dest_arg t) in (dest_arg1 t :: As, B') end else ([], t) fun list_obj_imp (As, C) = case As of [] => C | A :: rest => mk_imp (A, list_obj_imp (rest, C)) fun is_obj_all t = case t of Const (c, _) $ Abs _ => c = UtilBase.All_name | _ => false fun is_ball t = case t of Const (c, _) $ _ $ _ => c = UtilBase.Ball_name | _ => false fun mk_obj_all t body = let val (x, T) = Term.dest_Free t in Const (UtilBase.All_name, (T --> boolT) --> boolT) $ Term.absfree (x, T) body end fun is_ex t = case t of Const (c, _) $ Abs _ => c = UtilBase.Ex_name | _ => false fun is_bex t = case t of Const (c, _) $ _ $ _ => c = UtilBase.Bex_name | _ => false fun mk_exists t body = let val (x, T) = Term.dest_Free t in Const (UtilBase.Ex_name, (T --> boolT) --> boolT) $ Term.absfree (x, T) body end fun is_mem t = case t of Const (c, _) $ _ $ _ => c = UtilBase.Mem_name | _ => false fun mk_mem (x, A) = let val T = fastype_of x in Const (UtilBase.Mem_name, T --> UtilBase.mk_setT T --> boolT) $ x $ A end (* Finding subterms *) (* Get all subterms of t. *) fun list_subterms t = let fun dest_at_head t = case t of Abs (_, _, body) => dest_at_head body | _ => if Term.is_open t then t |> Term.strip_comb |> snd |> maps dest_at_head else [t] and dest t = case t of Abs _ => dest_at_head t | _ $ _ => t |> Term.strip_comb |> snd | _ => [] in dest t end (* List all (closed) subterms. Larger ones will be listed first. *) fun get_all_subterms t = t |> list_subterms |> maps get_all_subterms |> distinct (op aconv) |> cons t (* List all (closed) subterms. For terms of form if cond then yes or no, list only subterms of cond. *) fun get_all_subterms_skip_if t = if UtilBase.is_if t then t |> Util.dest_args |> hd |> get_all_subterms_skip_if |> cons t else t |> list_subterms |> maps get_all_subterms_skip_if |> distinct (op aconv) |> cons t (* cterms *) fun get_cneg ct = let val t = Thm.term_of ct val _ = assert (fastype_of t = boolT) "get_neg: wrong type" in if is_neg t then Thm.dest_arg ct else Thm.apply UtilBase.cNot ct end (* Theorems *) fun is_true_th th = Thm.eq_thm_prop (th, true_th) fun prop_of' th = dest_Trueprop (Thm.prop_of th) fun cprop_of' th = Thm.dest_arg (Thm.cprop_of th) fun concl_of' th = dest_Trueprop (Thm.concl_of th) (* Given an equality between bools, make it an equality between props, by applying the function Trueprop to both sides. *) fun make_trueprop_eq th = Thm.combination (Thm.reflexive UtilBase.cTrueprop) th (* Assumed theorems. *) fun assume_eq thy (t1, t2) = Thm.assume (Thm.global_cterm_of thy (mk_Trueprop (mk_eq (t1, t2)))) (* Apply cv to the statement of th, skipping Trueprop. *) fun apply_to_thm' cv th = apply_to_thm (Trueprop_conv cv) th val to_meta_eq = apply_to_thm (Util.concl_conv UtilBase.to_meta_eq_cv) val to_obj_eq = apply_to_thm (Util.concl_conv UtilBase.to_obj_eq_cv) val obj_sym = apply_to_thm (Util.concl_conv UtilBase.obj_sym_cv) (* Obtain rewriting conv from obj equality. *) fun rewr_obj_eq eq_th = Conv.rewr_conv (to_meta_eq eq_th) (* Given an equality A == B, make the equality ~A == ~B. Cancel ~~ on both sides if exists. *) fun make_neg_eq th = th |> Thm.combination (Thm.reflexive UtilBase.cNot) |> apply_to_lhs (Conv.try_conv (rewr_obj_eq UtilBase.nn_cancel_th)) |> apply_to_rhs (Conv.try_conv (rewr_obj_eq UtilBase.nn_cancel_th)) (* If ct is of the form [...] ==> False, leave it unchanged. Otherwise, change [...] ==> B to [..., ~ B] ==> False and change [...] ==> ~ B to [..., B] ==> False. *) fun rewrite_to_contra_form ct = let val (_, concl) = Logic.strip_horn (Thm.term_of ct) val concl' = dest_Trueprop concl in if concl' aconv bFalse then Conv.all_conv ct else if is_neg concl' then Util.concl_conv (Conv.rewr_conv UtilBase.to_contra_form_th') ct else Util.concl_conv (Conv.rewr_conv UtilBase.to_contra_form_th) ct end (* Rewrite ct from [...] ==> A ==> False to [...] ==> ~ A and from [...] ==> ~ A ==> False to [...] ==> A. *) fun rewrite_from_contra_form ct = let val (prems, concl) = Logic.strip_horn (Thm.term_of ct) val _ = assert (concl aconv pFalse) "rewrite_from_contra_form: concl should be false." val num_prems = length prems val last_prem' = dest_Trueprop (nth prems (num_prems-1)) val to_contra_form = if is_neg last_prem' then UtilBase.to_contra_form_th else UtilBase.to_contra_form_th' in Util.concl_conv_n (num_prems-1) (Conv.rewr_conv (meta_sym to_contra_form)) ct end (* Converts ==> to --> and !! to !. *) fun to_obj_conv ctxt ct = case Thm.term_of ct of @{const Pure.imp} $ _ $ _ => Conv.every_conv [Conv.binop_conv (to_obj_conv ctxt), Conv.rewr_conv UtilBase.atomize_imp_th] ct | Const (@{const_name Pure.all}, _) $ Abs _ => Conv.every_conv [Conv.binder_conv (to_obj_conv o snd) ctxt, Conv.rewr_conv UtilBase.atomize_all_th] ct | _ => Conv.all_conv ct (* When ct is of form A1 ==> ... ==> An, apply to_obj_conv to each Ai. *) fun to_obj_conv_on_horn ctxt ct = case Thm.term_of ct of @{const Pure.imp} $ _ $ _ => Conv.every_conv [Conv.arg1_conv (to_obj_conv ctxt), Conv.arg_conv (to_obj_conv_on_horn ctxt)] ct | Const (@{const_name Pure.all}, _) $ Abs _ => Conv.binder_conv (to_obj_conv_on_horn o snd) ctxt ct | _ => Conv.all_conv ct (* Convert !! and ==> on the outermost level. Pull all !! to the front. *) fun to_meta_conv ctxt ct = let val t = Thm.term_of ct in if is_Trueprop t andalso is_obj_all (dest_Trueprop t) then Conv.every_conv [Conv.rewr_conv (meta_sym UtilBase.atomize_all_th), to_meta_conv ctxt] ct else if is_Trueprop t andalso is_ball (dest_Trueprop t) then Conv.every_conv [Conv.arg_conv (rewr_obj_eq UtilBase.Ball_def_th), to_meta_conv ctxt] ct else if is_Trueprop t andalso is_imp (dest_Trueprop t) then Conv.every_conv [Conv.rewr_conv (meta_sym UtilBase.atomize_imp_th), to_meta_conv ctxt] ct else if Logic.is_all t then Conv.binder_conv (to_meta_conv o snd) ctxt ct else if Util.is_implies t then Conv.every_conv [Conv.arg_conv (to_meta_conv ctxt), Util.swap_meta_imp_alls ctxt] ct else Conv.all_conv ct end (* Modify th using imp_th, and add postfix to name (if available). *) fun thm_RS_mod imp_th suffix th = (th RS imp_th) |> Drule.zero_var_indexes |> Util.update_name_of_thm th suffix (* From A & B, obtain A. *) val conj_left_th = thm_RS_mod UtilBase.conjunct1_th "@left" (* From A & B, obtain B. *) val conj_right_th = thm_RS_mod UtilBase.conjunct2_th "@right" (* From (A::bool) = B, obtain A ==> B. *) val equiv_forward_th = thm_RS_mod UtilBase.iffD1_th "@eqforward" (* From (A::bool) = B, obtain B ==> A. *) val equiv_backward_th = thm_RS_mod UtilBase.iffD2_th "@eqbackward" (* From (A::bool) = B, obtain ~A ==> ~B. *) val inv_backward_th = thm_RS_mod UtilBase.inv_back_th "@invbackward" (* Same as to_obj_eq, except keeping names and indices. *) fun to_obj_eq_th th = th |> to_obj_eq |> Util.update_name_of_thm th "@obj_eq" (* Same as to_obj_eq_iff, except keeping names and indices. *) fun to_obj_eq_iff_th th = th |> UtilBase.to_obj_eq_iff |> Util.update_name_of_thm th "@iff" (* Same as obj_sym, except keeping names and indices. *) fun obj_sym_th th = th |> obj_sym |> Util.update_name_of_thm th "@sym" (* Given th of form (P ==>) A1 & ... & An, return theorems (P ==>) A1, ..., (P ==>) An, where there can be zero or more premises in front. *) fun split_conj_th th = if is_conj (prop_of' th) then (th RS UtilBase.conjunct1_th) :: (split_conj_th (th RS UtilBase.conjunct2_th)) else [th] (* More general form. *) fun split_conj_gen_th th = if is_conj (prop_of' th) then maps split_conj_gen_th [th RS UtilBase.conjunct1_th, th RS UtilBase.conjunct2_th] else [th] (* Given th of form ~ (A1 | ... | An), return theorems ~ A1, ... ~ An. *) fun split_not_disj_th th = let val t = prop_of' th in if is_neg t andalso is_disj (dest_not t) then (th RS UtilBase.or_intro1_th) :: (split_not_disj_th (th RS UtilBase.or_intro2_th)) else [th] end (* Similar to Logic.strip_horn, except remove Trueprop. *) fun strip_horn' th = (Logic.strip_horn (Thm.prop_of th)) |> apfst (map dest_Trueprop) |> apsnd dest_Trueprop fun mk_conjs_th ths = case ths of [] => raise Fail "mk_conjs_th" | [th] => th | th :: rest => [th, mk_conjs_th rest] MRS UtilBase.conjI_th (* Given th of form P x ==> False, where x is the given free variable, obtain new theorem of form (EX x. P x) ==> False. The function is written so it can be applied to multiple variables with fold. For example, "fold (ex_elim ctxt) [x, y] (P x y ==> False) will give (EX y x. P x y) ==> False. *) fun ex_elim ctxt freevar th = let val thy = Proof_Context.theory_of ctxt val th' = th |> Thm.forall_intr (Thm.cterm_of ctxt freevar) val head_prem = hd (Thm.prems_of UtilBase.exE_th') val inst = Pattern.match thy (head_prem, Thm.prop_of th') fo_init val exE_inst = Util.subst_thm ctxt inst UtilBase.exE_th' in Thm.elim_implies th' exE_inst end fun force_abs_form t = case t of Abs _ => t | _ => Abs ("x", domain_type (fastype_of t), t $ Bound 0) (* Normalization of object forall expressions in horn-clause form. *) fun strip_obj_horn t = if is_obj_all t then case t of - _ $ Abs (abs as (_, T, _)) => + _ $ (u as Abs _) => let - val (x, body) = Term.dest_abs abs - val var = Free (x, T) + val (v, body) = Term.dest_abs_global u val (vars, (assum, concl)) = strip_obj_horn body in - (var :: vars, (assum, concl)) + (Free v :: vars, (assum, concl)) end | f $ arg => strip_obj_horn (f $ force_abs_form arg) | _ => raise Fail "strip_obj_horn" else if is_ball t then case t of - _ $ S $ Abs (abs as (_, T, _)) => + _ $ S $ (u as Abs _) => let - val (x, body) = Term.dest_abs abs - val var = Free (x, T) + val (v, body) = Term.dest_abs_global u + val var = Free v val mem = mk_mem (var, S) val (vars, (assum, concl)) = strip_obj_horn body in (var :: vars, (mem :: assum, concl)) end | f $ S $ arg => strip_obj_horn (f $ S $ force_abs_form arg) | _ => raise Fail "strip_obj_horn" else if is_imp t then let val (vars, (assum, concl)) = strip_obj_horn (dest_arg t) in (vars, (dest_arg1 t :: assum, concl)) end else ([], ([], t)) fun list_obj_horn (vars, (As, B)) = (list_obj_imp (As, B)) |> fold mk_obj_all (rev vars) (* Whether t can be immediately rewritten into EX form. *) fun is_ex_form_gen t = is_ex t orelse is_bex t orelse (is_neg t andalso is_obj_all (dest_not t)) orelse (is_neg t andalso is_ball (dest_not t)) orelse (is_conj t andalso is_ex_form_gen (dest_arg t)) (* Rewrite A & EX v_i. B to EX v_i. A & B. *) fun swap_conj_exists ctxt ct = let val t = Thm.term_of ct in if is_conj t andalso is_ex (dest_arg t) then Conv.every_conv [rewr_obj_eq UtilBase.swap_ex_conj_th, Conv.binder_conv (swap_conj_exists o snd) ctxt] ct else Conv.all_conv ct end (* Normalize existence statement into EX v_i. A_1 & ... & A_n. This includes moving as many existence quantifiers to the left as possible. *) fun normalize_exists ctxt ct = let val t = Thm.term_of ct in if is_ex t then Conv.binder_conv (normalize_exists o snd) ctxt ct else if is_bex t then Conv.every_conv [rewr_obj_eq UtilBase.Bex_def_th, normalize_exists ctxt] ct else if is_neg t andalso is_obj_all (dest_not t) then Conv.every_conv [rewr_obj_eq UtilBase.not_all_th, normalize_exists ctxt] ct else if is_neg t andalso is_ball (dest_not t) then Conv.every_conv [Conv.arg_conv (rewr_obj_eq UtilBase.Ball_def_th), rewr_obj_eq UtilBase.not_all_th, Conv.binder_conv (K (rewr_obj_eq UtilBase.not_imp_th)) ctxt, normalize_exists ctxt] ct else if is_conj t then Conv.every_conv [Conv.arg_conv (normalize_exists ctxt), swap_conj_exists ctxt] ct else Conv.all_conv ct end (* Assume t is in normal form. *) fun strip_exists t = if is_ex t then case t of - _ $ Abs (abs as (_, T, _)) => + _ $ (u as Abs _) => let - val (x, body) = Term.dest_abs abs - val var = Free (x, T) + val (v, body) = Term.dest_abs_global u val (vars, body') = strip_exists body in - (var :: vars, body') + (Free v :: vars, body') end | _ => raise Fail "strip_exists" else ([], t) (* Generic wrapper function. tac can be arith_tac, simp_tac, etc. *) fun prove_by_tac tac ctxt ths goal = let val goal' = Logic.list_implies (map Thm.prop_of ths, mk_Trueprop goal) in - ths MRS (Goal.prove ctxt [] [] goal' (K (tac ctxt 1))) + ths MRS (Goal.prove ctxt [] [] goal' (HEADGOAL o tac o #context)) end fun contra_by_tac tac ctxt ths = prove_by_tac tac ctxt ths bFalse end (* structure UtilLogic *) structure Basic_UtilLogic: BASIC_UTIL_LOGIC = UtilLogic open Basic_UtilLogic diff --git a/thys/Auto2_Imperative_HOL/Functional/Indexed_PQueue.thy b/thys/Auto2_Imperative_HOL/Functional/Indexed_PQueue.thy --- a/thys/Auto2_Imperative_HOL/Functional/Indexed_PQueue.thy +++ b/thys/Auto2_Imperative_HOL/Functional/Indexed_PQueue.thy @@ -1,380 +1,380 @@ (* File: Indexed_PQueue.thy Author: Bohua Zhan *) section \Indexed priority queues\ theory Indexed_PQueue imports Arrays_Ex Mapping_Str begin text \ Verification of indexed priority queue: functional part. The data structure is also verified by Lammich in \cite{Refine_Imperative_HOL-AFP}. \ subsection \Successor functions, eq-pred predicate\ fun s1 :: "nat \ nat" where "s1 m = 2 * m + 1" fun s2 :: "nat \ nat" where "s2 m = 2 * m + 2" lemma s_inj [forward]: "s1 m = s1 m' \ m = m'" "s2 m = s2 m' \ m = m'" by auto lemma s_neq [resolve]: "s1 m \ s2 m'" "s1 m > m" "s2 m > m" "s2 m > s1 m" using s1.simps s2.simps by presburger+ setup \add_forward_prfstep_cond @{thm s_neq(2)} [with_term "s1 ?m"]\ setup \add_forward_prfstep_cond @{thm s_neq(3)} [with_term "s2 ?m"]\ setup \add_forward_prfstep_cond @{thm s_neq(4)} [with_term "s2 ?m", with_term "s1 ?m"]\ inductive eq_pred :: "nat \ nat \ bool" where "eq_pred n n" | "eq_pred n m \ eq_pred n (s1 m)" | "eq_pred n m \ eq_pred n (s2 m)" setup \add_case_induct_rule @{thm eq_pred.cases}\ setup \add_prop_induct_rule @{thm eq_pred.induct}\ setup \add_resolve_prfstep @{thm eq_pred.intros(1)}\ setup \fold add_backward_prfstep @{thms eq_pred.intros(2,3)}\ lemma eq_pred_parent1 [forward]: "eq_pred i (s1 k) \ i \ s1 k \ eq_pred i k" @proof @let "v = s1 k" @prop_induct "eq_pred i v" @qed lemma eq_pred_parent2 [forward]: "eq_pred i (s2 k) \ i \ s2 k \ eq_pred i k" @proof @let "v = s2 k" @prop_induct "eq_pred i v" @qed lemma eq_pred_cases: "eq_pred i j \ eq_pred (s1 i) j \ eq_pred (s2 i) j \ j = i \ j = s1 i \ j = s2 i" @proof @prop_induct "eq_pred i j" @qed setup \add_forward_prfstep_cond @{thm eq_pred_cases} [with_cond "?i \ s1 ?k", with_cond "?i \ s2 ?k"]\ lemma eq_pred_le [forward]: "eq_pred i j \ i \ j" @proof @prop_induct "eq_pred i j" @qed subsection \Heap property\ text \The corresponding tree is a heap\ definition is_heap :: "('a \ 'b::linorder) list \ bool" where [rewrite]: "is_heap xs = (\i j. eq_pred i j \ j < length xs \ snd (xs ! i) \ snd (xs ! j))" lemma is_heapD: "is_heap xs \ j < length xs \ eq_pred i j \ snd (xs ! i) \ snd (xs ! j)" by auto2 setup \add_forward_prfstep_cond @{thm is_heapD} [with_term "?xs ! ?j"]\ setup \del_prfstep_thm_eqforward @{thm is_heap_def}\ subsection \Bubble-down\ text \The corresponding tree is a heap, except k is not necessarily smaller than its descendents.\ definition is_heap_partial1 :: "('a \ 'b::linorder) list \ nat \ bool" where [rewrite]: "is_heap_partial1 xs k = (\i j. eq_pred i j \ i \ k \ j < length xs \ snd (xs ! i) \ snd (xs ! j))" text \Two cases of switching with s1 k.\ lemma bubble_down1: "s1 k < length xs \ is_heap_partial1 xs k \ snd (xs ! k) > snd (xs ! s1 k) \ snd (xs ! s1 k) \ snd (xs ! s2 k) \ is_heap_partial1 (list_swap xs k (s1 k)) (s1 k)" by auto2 setup \add_forward_prfstep_cond @{thm bubble_down1} [with_term "list_swap ?xs ?k (s1 ?k)"]\ lemma bubble_down2: "s1 k < length xs \ is_heap_partial1 xs k \ snd (xs ! k) > snd (xs ! s1 k) \ s2 k \ length xs \ is_heap_partial1 (list_swap xs k (s1 k)) (s1 k)" by auto2 setup \add_forward_prfstep_cond @{thm bubble_down2} [with_term "list_swap ?xs ?k (s1 ?k)"]\ text \One case of switching with s2 k.\ lemma bubble_down3: "s2 k < length xs \ is_heap_partial1 xs k \ snd (xs ! s1 k) > snd (xs ! s2 k) \ snd (xs ! k) > snd (xs ! s2 k) \ xs' = list_swap xs k (s2 k) \ is_heap_partial1 xs' (s2 k)" by auto2 setup \add_forward_prfstep_cond @{thm bubble_down3} [with_term "?xs'"]\ subsection \Bubble-up\ fun par :: "nat \ nat" where "par m = (m - 1) div 2" setup \register_wellform_data ("par m", ["m \ 0"])\ lemma ps_inverse [rewrite]: "par (s1 k) = k" "par (s2 k) = k" by simp+ lemma p_basic: "m \ 0 \ par m < m" by auto setup \add_forward_prfstep_cond @{thm p_basic} [with_term "par ?m"]\ lemma p_cases: "m \ 0 \ m = s1 (par m) \ m = s2 (par m)" by auto setup \add_forward_prfstep_cond @{thm p_cases} [with_term "par ?m"]\ lemma eq_pred_p_next: "i \ 0 \ eq_pred i j \ eq_pred (par i) j" @proof @prop_induct "eq_pred i j" @qed setup \add_forward_prfstep_cond @{thm eq_pred_p_next} [with_term "par ?i"]\ lemma heap_implies_hd_min [resolve]: "is_heap xs \ i < length xs \ xs \ [] \ snd (hd xs) \ snd (xs ! i)" @proof @strong_induct i @case "i = 0" @apply_induct_hyp "par i" @have "eq_pred (par i) i" @qed text \The corresponding tree is a heap, except k is not necessarily greater than its ancestors.\ definition is_heap_partial2 :: "('a \ 'b::linorder) list \ nat \ bool" where [rewrite]: "is_heap_partial2 xs k = (\i j. eq_pred i j \ j < length xs \ j \ k \ snd (xs ! i) \ snd (xs ! j))" lemma bubble_up1 [forward]: "k < length xs \ is_heap_partial2 xs k \ snd (xs ! k) < snd (xs ! par k) \ k \ 0 \ is_heap_partial2 (list_swap xs k (par k)) (par k)" by auto2 lemma bubble_up2 [forward]: "k < length xs \ is_heap_partial2 xs k \ snd (xs ! k) \ snd (xs ! par k) \ k \ 0 \ is_heap xs" by auto2 setup \del_prfstep_thm @{thm p_cases}\ subsection \Indexed priority queue\ type_synonym 'a idx_pqueue = "(nat \ 'a) list \ nat option list" fun index_of_pqueue :: "'a idx_pqueue \ bool" where "index_of_pqueue (xs, m) = ( (\i m ! (fst (xs ! i)) = Some i) \ (\i. \k i < length xs \ fst (xs ! i) = k))" setup \add_rewrite_rule @{thm index_of_pqueue.simps}\ lemma index_of_pqueueD1: "i < length xs \ index_of_pqueue (xs, m) \ fst (xs ! i) < length m \ m ! (fst (xs ! i)) = Some i" by auto2 setup \add_forward_prfstep_cond @{thm index_of_pqueueD1} [with_term "?xs ! ?i"]\ lemma index_of_pqueueD2 [forward]: "k < length m \ index_of_pqueue (xs, m) \ m ! k = Some i \ i < length xs \ fst (xs ! i) = k" by auto2 lemma index_of_pqueueD3 [forward]: "index_of_pqueue (xs, m) \ p \ set xs \ fst p < length m" @proof @obtain i where "i < length xs" "xs ! i = p" @qed setup \del_prfstep_thm_eqforward @{thm index_of_pqueue.simps}\ lemma has_index_unique_key [forward]: "index_of_pqueue (xs, m) \ unique_keys_set (set xs)" @proof @have "\a x y. (a, x) \ set xs \ (a, y) \ set xs \ x = y" @with @obtain i where "i < length xs" "xs ! i = (a, x)" @obtain j where "j < length xs" "xs ! j = (a, y)" @end @qed lemma has_index_keys_of [rewrite]: "index_of_pqueue (xs, m) \ has_key_alist xs k \ (k < length m \ m ! k \ None)" @proof @case "has_key_alist xs k" @with @obtain v' where "(k, v') \ set xs" @obtain i where "i < length xs \ xs ! i = (k, v')" @end @qed lemma has_index_distinct [forward]: "index_of_pqueue (xs, m) \ distinct xs" @proof @have "\ij j \ xs ! i \ xs ! j" @qed subsection \Basic operations on indexed\_queue\ fun idx_pqueue_swap_fun :: "(nat \ 'a) list \ nat option list \ nat \ nat \ (nat \ 'a) list \ nat option list" where "idx_pqueue_swap_fun (xs, m) i j = ( list_swap xs i j, ((m [fst (xs ! i) := Some j]) [fst (xs ! j) := Some i]))" lemma index_of_pqueue_swap [forward_arg]: "i < length xs \ j < length xs \ index_of_pqueue (xs, m) \ index_of_pqueue (idx_pqueue_swap_fun (xs, m) i j)" @proof @unfold "idx_pqueue_swap_fun (xs, m) i j" @qed lemma fst_idx_pqueue_swap [rewrite]: "fst (idx_pqueue_swap_fun (xs, m) i j) = list_swap xs i j" @proof @unfold "idx_pqueue_swap_fun (xs, m) i j" @qed lemma snd_idx_pqueue_swap [rewrite]: "length (snd (idx_pqueue_swap_fun (xs, m) i j)) = length m" @proof @unfold "idx_pqueue_swap_fun (xs, m) i j" @qed fun idx_pqueue_push_fun :: "nat \ 'a \ 'a idx_pqueue \ 'a idx_pqueue" where "idx_pqueue_push_fun k v (xs, m) = (xs @ [(k, v)], list_update m k (Some (length xs)))" lemma idx_pqueue_push_correct [forward_arg]: "index_of_pqueue (xs, m) \ k < length m \ \has_key_alist xs k \ r = idx_pqueue_push_fun k v (xs, m) \ index_of_pqueue r \ fst r = xs @ [(k, v)] \ length (snd r) = length m" @proof @unfold "idx_pqueue_push_fun k v (xs, m)" @qed fun idx_pqueue_pop_fun :: "'a idx_pqueue \ 'a idx_pqueue" where "idx_pqueue_pop_fun (xs, m) = (butlast xs, list_update m (fst (last xs)) None)" lemma idx_pqueue_pop_correct [forward_arg]: "index_of_pqueue (xs, m) \ xs \ [] \ r = idx_pqueue_pop_fun (xs, m) \ index_of_pqueue r \ fst r = butlast xs \ length (snd r) = length m" @proof @unfold "idx_pqueue_pop_fun (xs, m)" @have "length xs = length (butlast xs) + 1" @have "fst (xs ! length (butlast xs)) < length m" (* TODO: remove? *) @qed subsection \Bubble up and down\ function idx_bubble_down_fun :: "'a::linorder idx_pqueue \ nat \ 'a idx_pqueue" where "idx_bubble_down_fun (xs, m) k = ( if s2 k < length xs then if snd (xs ! s1 k) \ snd (xs ! s2 k) then if snd (xs ! k) > snd (xs ! s1 k) then idx_bubble_down_fun (idx_pqueue_swap_fun (xs, m) k (s1 k)) (s1 k) else (xs, m) else if snd (xs ! k) > snd (xs ! s2 k) then idx_bubble_down_fun (idx_pqueue_swap_fun (xs, m) k (s2 k)) (s2 k) else (xs, m) else if s1 k < length xs then if snd (xs ! k) > snd (xs ! s1 k) then idx_bubble_down_fun (idx_pqueue_swap_fun (xs, m) k (s1 k)) (s1 k) else (xs, m) else (xs, m))" by pat_completeness auto termination by (relation "measure (\((xs,_),k). (length xs - k))") (simp_all, auto2+) lemma idx_bubble_down_fun_correct: "r = idx_bubble_down_fun x k \ is_heap_partial1 (fst x) k \ is_heap (fst r) \ mset (fst r) = mset (fst x) \ length (snd r) = length (snd x)" @proof @fun_induct "idx_bubble_down_fun x k" @with @subgoal "(x = (xs, m), k = k)" @unfold "idx_bubble_down_fun (xs, m) k" @case "s2 k < length xs" @with @case "snd (xs ! s1 k) \ snd (xs ! s2 k)" @end @case "s1 k < length xs" @end @qed setup \add_forward_prfstep_cond @{thm idx_bubble_down_fun_correct} [with_term "?r"]\ lemma idx_bubble_down_fun_correct2 [forward]: "index_of_pqueue x \ index_of_pqueue (idx_bubble_down_fun x k)" @proof @fun_induct "idx_bubble_down_fun x k" @with @subgoal "(x = (xs, m), k = k)" @unfold "idx_bubble_down_fun (xs, m) k" @case "s2 k < length xs" @with @case "snd (xs ! s1 k) \ snd (xs ! s2 k)" @end @case "s1 k < length xs" @end @qed fun idx_bubble_up_fun :: "'a::linorder idx_pqueue \ nat \ 'a idx_pqueue" where "idx_bubble_up_fun (xs, m) k = ( if k = 0 then (xs, m) else if k < length xs then if snd (xs ! k) < snd (xs ! par k) then idx_bubble_up_fun (idx_pqueue_swap_fun (xs, m) k (par k)) (par k) else (xs, m) else (xs, m))" lemma idx_bubble_up_fun_correct: "r = idx_bubble_up_fun x k \ is_heap_partial2 (fst x) k \ is_heap (fst r) \ mset (fst r) = mset (fst x) \ length (snd r) = length (snd x)" @proof @fun_induct "idx_bubble_up_fun x k" @with @subgoal "(x = (xs, m), k = k)" @unfold "idx_bubble_up_fun (xs, m) k" @end @qed setup \add_forward_prfstep_cond @{thm idx_bubble_up_fun_correct} [with_term "?r"]\ lemma idx_bubble_up_fun_correct2 [forward]: "index_of_pqueue x \ index_of_pqueue (idx_bubble_up_fun x k)" @proof @fun_induct "idx_bubble_up_fun x k" @with @subgoal "(x = (xs, m), k = k)" @unfold "idx_bubble_up_fun (xs, m) k" @end @qed subsection \Main operations\ fun delete_min_idx_pqueue_fun :: "'a::linorder idx_pqueue \ (nat \ 'a) \ 'a idx_pqueue" where "delete_min_idx_pqueue_fun (xs, m) = ( let (xs', m') = idx_pqueue_swap_fun (xs, m) 0 (length xs - 1); a'' = idx_pqueue_pop_fun (xs', m') in (last xs', idx_bubble_down_fun a'' 0))" lemma delete_min_idx_pqueue_correct: "index_of_pqueue (xs, m) \ xs \ [] \ res = delete_min_idx_pqueue_fun (xs, m) \ index_of_pqueue (snd res)" @proof @unfold "delete_min_idx_pqueue_fun (xs, m)" @qed setup \add_forward_prfstep_cond @{thm delete_min_idx_pqueue_correct} [with_term "?res"]\ lemma hd_last_swap_eval_last [rewrite]: "xs \ [] \ last (list_swap xs 0 (length xs - 1)) = hd xs" @proof @let "xs' = list_swap xs 0 (length xs - 1)" @have "last xs' = xs' ! (length xs - 1)" @qed text \Correctness of delete-min.\ theorem delete_min_idx_pqueue_correct2: "is_heap xs \ xs \ [] \ res = delete_min_idx_pqueue_fun (xs, m) \ index_of_pqueue (xs, m) \ is_heap (fst (snd res)) \ fst res = hd xs \ length (snd (snd res)) = length m \ map_of_alist (fst (snd res)) = delete_map (fst (fst res)) (map_of_alist xs)" @proof @unfold "delete_min_idx_pqueue_fun (xs, m)" @let "xs' = list_swap xs 0 (length xs - 1)" @have "is_heap_partial1 (butlast xs') 0" @qed setup \add_forward_prfstep_cond @{thm delete_min_idx_pqueue_correct2} [with_term "?res"]\ fun insert_idx_pqueue_fun :: "nat \ 'a::linorder \ 'a idx_pqueue \ 'a idx_pqueue" where "insert_idx_pqueue_fun k v x = ( let x' = idx_pqueue_push_fun k v x in idx_bubble_up_fun x' (length (fst x') - 1))" lemma insert_idx_pqueue_correct [forward_arg]: "index_of_pqueue (xs, m) \ k < length m \ \has_key_alist xs k \ index_of_pqueue (insert_idx_pqueue_fun k v (xs, m))" @proof @unfold "insert_idx_pqueue_fun k v (xs, m)" @qed text \Correctness of insertion.\ theorem insert_idx_pqueue_correct2: "index_of_pqueue (xs, m) \ is_heap xs \ k < length m \ \has_key_alist xs k \ r = insert_idx_pqueue_fun k v (xs, m) \ is_heap (fst r) \ length (snd r) = length m \ map_of_alist (fst r) = map_of_alist xs { k \ v }" @proof @unfold "insert_idx_pqueue_fun k v (xs, m)" @have "is_heap_partial2 (xs @ [(k, v)]) (length xs)" @qed setup \add_forward_prfstep_cond @{thm insert_idx_pqueue_correct2} [with_term "?r"]\ fun update_idx_pqueue_fun :: "nat \ 'a::linorder \ 'a idx_pqueue \ 'a idx_pqueue" where "update_idx_pqueue_fun k v (xs, m) = ( if m ! k = None then insert_idx_pqueue_fun k v (xs, m) else let i = the (m ! k); xs' = list_update xs i (k, v) in if snd (xs ! i) \ v then idx_bubble_down_fun (xs', m) i else idx_bubble_up_fun (xs', m) i)" lemma update_idx_pqueue_correct [forward_arg]: "index_of_pqueue (xs, m) \ k < length m \ index_of_pqueue (update_idx_pqueue_fun k v (xs, m))" @proof @unfold "update_idx_pqueue_fun k v (xs, m)" - @let "i = the (m ! k)" - @let "xs' = list_update xs i (k, v)" + @let "i' = the (m ! k)" + @let "xs' = list_update xs i' (k, v)" @case "m ! k = None" @have "index_of_pqueue (xs', m)" @qed text \Correctness of update.\ theorem update_idx_pqueue_correct2: "index_of_pqueue (xs, m) \ is_heap xs \ k < length m \ r = update_idx_pqueue_fun k v (xs, m) \ is_heap (fst r) \ length (snd r) = length m \ map_of_alist (fst r) = map_of_alist xs { k \ v }" @proof @unfold "update_idx_pqueue_fun k v (xs, m)" @let "i = the (m ! k)" @let "xs' = list_update xs i (k, v)" @have "xs' = fst (xs', m)" (* TODO: remove *) @case "m ! k = None" @case "snd (xs ! the (m ! k)) \ v" @with @have "is_heap_partial1 xs' i" @end @have "is_heap_partial2 xs' i" @qed setup \add_forward_prfstep_cond @{thm update_idx_pqueue_correct2} [with_term "?r"]\ 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/Complx/OG_Syntax.thy b/thys/Complx/OG_Syntax.thy --- a/thys/Complx/OG_Syntax.thy +++ b/thys/Complx/OG_Syntax.thy @@ -1,420 +1,420 @@ (* * Copyright 2016, Data61, CSIRO * * This software may be distributed and modified according to the terms of * the BSD 2-Clause license. Note that NO WARRANTY is provided. * See "LICENSE_BSD2.txt" for details. * * @TAG(DATA61_BSD) *) section \Shallowly-embedded syntax for COMPLX programs\ theory OG_Syntax imports OG_Hoare OG_Tactics begin datatype ('s, 'p, 'f) ann_com = AnnCom "('s, 'p, 'f) ann" "('s,'p,'f) com" fun ann where "ann (AnnCom p q) = p" fun com where "com (AnnCom p q) = q" lemmas ann.simps[oghoare_simps] com.simps[oghoare_simps] syntax "_quote" :: "'b \ ('a \ 'b)" ("(\_\)" [0] 1000) "_antiquote" :: "('a \ 'b) \ 'b" ("\_" [1000] 1000) "_Assert" :: "'a \ 'a set" ("(\_\)" [0] 1000) translations "\b\" \ "CONST Collect \b\" parse_translation \ let fun quote_tr [t] = Syntax_Trans.quote_tr @{syntax_const "_antiquote"} t | quote_tr ts = raise TERM ("quote_tr", ts); in [(@{syntax_const "_quote"}, K quote_tr)] end \ syntax "_fst" :: "'a \ 'b \ 'a" ("_\<^sub>," [60] 61) "_snd" :: "'a \ 'b \ 'b" ("_\<^sub>." [60] 61) parse_translation \ let fun fst_tr ((Const (@{const_syntax Pair}, _) $ p $ c) :: ts) = p; fun snd_tr ((Const (@{const_syntax Pair}, _) $ p $ c) :: ts) = c; in [(@{syntax_const "_fst"}, K fst_tr), (@{syntax_const "_snd"}, K snd_tr)] end \ text\Syntax for commands and for assertions and boolean expressions in commands \com\ and annotated commands \ann_com\.\ syntax "_Annotation" :: "('s,'p,'f) ann_com \ ('s, 'p, 'f) ann" ("_\<^sub>?" [60] 61) "_Command" :: "('s,'p,'f) ann_com \ ('s,'p,'f) com" ("_\<^sub>!" [60] 61) parse_translation \ let fun ann_tr ((Const (@{const_syntax AnnCom}, _) $ p $ c) :: ts) = p | ann_tr (p :: ts) = Const (@{const_syntax ann}, dummyT) $ p | ann_tr x = raise TERM ("ann_tr", x); fun com_tr ((Const (@{const_syntax AnnCom}, _) $ p $ c) :: ts) = c | com_tr (c :: ts) = Const (@{const_syntax com}, dummyT) $ c | com_tr x = raise TERM ("com_tr", x); in [(@{syntax_const "_Annotation"}, K ann_tr), (@{syntax_const "_Command"}, K com_tr)] end \ syntax "_Seq" :: "('s,'p,'f) ann_com \ ('s,'p,'f) ann_com \ ('s,'p,'f) ann_com" ("(_,,/ _)" [55, 56] 55) "_AnnSeq" :: "('s,'p,'f) ann_com \ ('s,'p,'f) ann_com \ ('s,'p,'f) ann_com" ("(_;;//_)" [55, 56] 55) translations "_Seq c1 c2" \ "CONST AnnCom (CONST AnnComp (c1\<^sub>?) (c2\<^sub>?)) (CONST Seq (c1\<^sub>!) (c2\<^sub>!))" "_AnnSeq c1 c2" \ "CONST AnnCom (CONST AnnComp (c1\<^sub>?) (c2\<^sub>?)) (CONST Seq (c1\<^sub>!) (c2\<^sub>!))" syntax "_Assign" :: "idt \ 'b \ ('s,'p,'f) ann_com" ("(\_ :=/ _)" [70, 65] 61) "_AnnAssign" :: "'s assn \ idt \ 'b \ ('s,'p,'f) ann_com" ("(_//\_ :=/ _)" [90,70,65] 61) definition "FAKE_ANN \ UNIV" translations "r \x := a" \ "CONST AnnCom (CONST AnnExpr r) (CONST Basic \\(_update_name x (\_. a))\)" "\x := a" \ "CONST FAKE_ANN \x := a" abbreviation "update_var f S s \ (\v. f (\_. v) s) ` S" abbreviation "fun_to_rel f \ \ ((\s. (\v. (s, v)) ` f s) ` UNIV)" syntax "_Spec" :: "idt \ 'b \ ('s,'p,'f) ann_com" ("(\_ :\/ _)" [70, 65] 61) "_AnnSpec" :: "'a assn \ idt \ 'b \ ('s,'p,'f) ann_com" ("(_//\_ :\/ _)" [90,70,65] 61) translations "r \x :\ S" \ "CONST AnnCom (CONST AnnExpr r) (CONST Spec (CONST fun_to_rel \\(CONST update_var (_update_name x) S)\))" "\x :\ S" \ "CONST FAKE_ANN \x :\ S" nonterminal grds and grd syntax "_AnnCond1" :: "'s assn \ 's bexp \ ('s,'p,'f) ann_com \ ('s,'p,'f) ann_com \ ('s,'p,'f) ann_com" ("(_//IF _//(2THEN/ (_))//(2ELSE/ (_))//FI)" [90,0,0,0] 61) "_AnnCond2" :: "'s assn \ 's bexp \ ('s,'p,'f) ann_com \ ('s,'p,'f) ann_com" ("(_//IF _//(2THEN/ (_))//FI)" [90,0,0] 61) "_AnnWhile" :: "'s assn \ 's bexp \ 's assn \ ('s,'p,'f) ann_com \ ('s,'p,'f) ann_com" ("(_//WHILE _/ INV _//(2DO/ (_))//OD)" [90,0,0,0] 61) "_AnnAwait" :: "'s assn \ 's bexp \ ('s,'p,'f) ann_com \ ('s,'p,'f) ann_com" ("(_//AWAIT _/ (2THEN/ (_))/ END)" [90,0,0] 61) "_AnnAtom" :: "'s assn \ ('s,'p,'f) ann_com \ ('s,'p,'f) ann_com" ("(_//\_\)" [90,0] 61) "_AnnWait" :: "'s assn \ 's bexp \ ('s,'p,'f) ann_com" ("(_//WAIT _/ END)" [90,0] 61) "_Cond" :: "'s bexp \ ('s,'p,'f) ann_com \ ('s,'p,'f) ann_com \ ('s,'p,'f) ann_com" ("(IF _//(2THEN/ (_))//(2ELSE/ (_))//FI)" [0, 0, 0] 61) "_Cond2" :: "'s bexp \ ('s,'p,'f) ann_com \ ('s,'p,'f) ann_com" ("(IF _//(2THEN/ (_))//FI)" [0,0] 56) "_While_inv" :: "'s bexp \ 's assn \ ('s,'p,'f) ann_com \ ('s,'p,'f) ann_com" ("(WHILE _/ INV _//(2DO/ (_))//OD)" [0, 0, 0] 61) "_While" :: "'s bexp \ ('s,'p,'f) ann_com \ ('s,'p,'f) ann_com" ("(WHILE _//(2DO/ (_))//OD)" [0, 0] 61) "_Await" :: "'s bexp \ ('s,'p,'f) ann_com \ ('s,'p,'f) ann_com" ("(AWAIT _/ (2THEN/ (_))/ END)" [0,0] 61) "_Atom" :: "('s,'p,'f) ann_com \ ('s,'p,'f) ann_com" ("(\_\)" [0] 61) "_Wait" :: "'s bexp \ ('s,'p,'f) ann_com" ("(WAIT _/ END)" [0] 61) "_grd" :: "'f \ 's bexp \ grd" ("'(_, _')" [1000] 1000) "_last_grd" :: "grd \ grds" ("_" 1000) "_grds" :: "[grd, grds] \ grds" ("(_,/ _)" [999,1000] 1000) "_guards" :: "'s assn \ grds \ ('s,'p,'f) ann_com \ ('s,'p,'f) ann_com" ("(_//(2_ \/ (_)))" [90, 0, 56] 61) "_Throw" :: "('s,'p,'f) ann_com" ("THROW" 61) "_AnnThrow" :: "'s assn \ ('s,'p,'f) ann_com" ("(_/ THROW)" [90] 61) "_Try_Catch" :: "('s,'p,'f) ann_com \('s,'p,'f) ann_com \ ('s,'p,'f) ann_com" ("((2TRY/ (_))//(2CATCH/ (_))/ END)" [0,0] 71) "_AnnCallX" :: "'s assn \ ('s \ 's) \ 's assn \ 'p \ nat \ ('s \ 's \ 's) \ ('s\'s\('s,'p,'f) com) \ 's assn \ 's assn \ 's assn \ 's assn \ ('s,'p,'f) ann_com" ("(_//(2CALLX/ (_)//_/ _/ _//_/ _//_/ _//_/ _))" [90,1000,0,1000,0,1000,1000,0,0,0,0] 61) "_AnnSCall" :: "'s assn \ 'p \ nat \ ('s,'p,'f) ann_com" ("(_//SCALL _/ _)" [90,0,0] 61) "_Skip" :: "('s,'p,'f) ann_com" ("SKIP" 61) "_AnnSkip" :: "'s assn \ ('s,'p,'f) ann_com" ("(_/ SKIP)" [90] 61) translations "r IF b THEN c1 ELSE c2 FI" \ "CONST AnnCom (CONST AnnBin r (c1\<^sub>?) (c2\<^sub>?)) (CONST Cond \b\ (c1\<^sub>!) (c2\<^sub>!))" "r IF b THEN c FI" \ "r IF b THEN c ELSE SKIP FI" "r WHILE b INV i DO c OD" \ "CONST AnnCom (CONST AnnWhile r i (c\<^sub>?)) (CONST While \b\ (c\<^sub>!))" "r AWAIT b THEN c END" \ "CONST AnnCom (CONST AnnRec r (c\<^sub>?)) (CONST Await \b\ (c\<^sub>!))" "r \c\" \ "r AWAIT CONST True THEN c END" "r WAIT b END" \ "r AWAIT b THEN SKIP END" "IF b THEN c1 ELSE c2 FI" \ "CONST FAKE_ANN IF b THEN c1 ELSE c2 FI" "IF b THEN c FI" \ "CONST FAKE_ANN IF b THEN c ELSE SKIP FI" "WHILE b DO c OD" \ "CONST FAKE_ANN WHILE b INV CONST FAKE_ANN DO c OD" "WHILE b INV i DO c OD" \ "CONST FAKE_ANN WHILE b INV i DO c OD" "AWAIT b THEN c END" \ "CONST FAKE_ANN AWAIT b THEN c END" "\c\" \ "CONST FAKE_ANN AWAIT CONST True THEN c END" "WAIT b END" \ "AWAIT b THEN SKIP END" "_grd f g" \ "(f, g)" "_grds g gs" \ "g#gs" "_last_grd g" \ "[g]" "_guards r gs c" \ "CONST AnnCom (CONST ann_guards r gs (c\<^sub>?)) (CONST guards gs (c\<^sub>!))" "ai CALLX init r p n restore return arestoreq areturn arestorea A" \ "CONST AnnCom (CONST ann_call ai r n arestoreq areturn arestorea A) (CONST call init p restore return)" "r SCALL p n" \ "CONST AnnCom (CONST AnnCall r n) (CONST Call p)" "r THROW" \ "CONST AnnCom (CONST AnnExpr r) (CONST Throw)" "THROW" \ "CONST FAKE_ANN THROW" "TRY c1 CATCH c2 END" \ "CONST AnnCom (CONST AnnComp (c1\<^sub>?) (c2\<^sub>?)) (CONST Catch (c1\<^sub>!) (c2\<^sub>!))" "r SKIP" \ "CONST AnnCom (CONST AnnExpr r) (CONST Skip)" "SKIP" \ "CONST FAKE_ANN SKIP" nonterminal prgs syntax "_PAR" :: "prgs \ 'a" ("(COBEGIN//_//COEND)" [57] 56) "_prg" :: "['a, 'a, 'a] \ prgs" ("(2 _//_,/ _)" [60, 90, 90] 57) "_prgs" :: "['a, 'a, 'a, prgs] \ prgs" ("(2 _//_,/ _)//\//_" [60,90,90,57] 57) "_prg_scheme" :: "['a, 'a, 'a, 'a, 'a, 'a] \ prgs" (" (2SCHEME [_ \ _ < _]//_//_,/ _)" [0,0,0,60, 90,90] 57) translations "_prg c q a" \ "([((c\<^sub>?), q, a)], [(c\<^sub>!)])" "_prgs c q a ps" \ "(((c\<^sub>?), q, a) # (ps\<^sub>,), (c\<^sub>!) # (ps\<^sub>.))" "_PAR ps" \ "CONST AnnCom (CONST AnnPar (ps\<^sub>,)) (CONST Parallel (ps\<^sub>.))" "_prg_scheme j i k c q a" \ "(CONST map (\i. ((c\<^sub>?), q, a)) [j..i. (c\<^sub>!)) [j.. ('s,'p,'f) proc_assns \ 'f set \ ('s,'p,'f) ann_com \ 's assn \ 's assn \ bool" ("(4_),/ (4_)/ (|\\<^bsub>'/_\<^esub> (_//_, _))" [60,60,60,20,1000,1000]60) "_oghoare_seq" :: "('s,'p,'f) body \ ('s,'p,'f) proc_assns \ 'f set \'s assn \ ('s,'p,'f) ann_com \ 's assn \ 's assn \ bool" ("(4_),/ (4_)/ (|\\<^bsub>'/_\<^esub> (_//_//_, _))" [60,60,60,1000,20,1000,1000]60) translations "_oghoare \ \ F c Q A" \ "\, \ \\<^bsub>/F\<^esub> (c\<^sub>?) (c\<^sub>!) Q, A" "_oghoare_seq \ \ F P c Q A" \ "\, \ \\<^bsub>/F\<^esub> P (c\<^sub>?) (c\<^sub>!) Q, A" ML \val syntax_debug = false\ print_translation \ let fun quote_tr' f (t :: ts) = Term.list_comb (f $ Syntax_Trans.quote_tr' @{syntax_const "_antiquote"} t, ts) | quote_tr' _ _ = raise Match; fun annquote_tr' f (r :: t :: ts) = Term.list_comb (f $ r $ Syntax_Trans.quote_tr' @{syntax_const "_antiquote"} t, ts) | annquote_tr' _ _ = raise Match; val assert_tr' = quote_tr' (Syntax.const @{syntax_const "_Assert"}); fun annbexp_tr' name (r :: (Const (@{const_syntax Collect}, _) $ t) :: ts) = annquote_tr' (Syntax.const name) (r :: t :: ts) | annbexp_tr' name (r :: Const (@{const_syntax UNIV}, _) :: ts) = annquote_tr' (Syntax.const name) (r :: Abs ("s", dummyT, Const (@{const_syntax True}, dummyT)) :: ts) | annbexp_tr' name (r :: Const (@{const_syntax Set.empty}, _) :: ts) = annquote_tr' (Syntax.const name) (r :: Abs ("s", dummyT, Const (@{const_syntax False}, dummyT)) :: ts) | annbexp_tr' name x = let val _ = if syntax_debug then writeln ("annbexp_tr'\n " ^ @{make_string} x) else () in raise Match end; fun annassign_tr' (r :: Abs (x, _, f $ k $ Bound 0) :: ts) = quote_tr' (Syntax.const @{syntax_const "_AnnAssign"} $ r $ Syntax_Trans.update_name_tr' f) (Abs (x, dummyT, Syntax_Trans.const_abs_tr' k) :: ts) | annassign_tr' r = let val _ = writeln ("annassign_tr'\n " ^ @{make_string} r) in raise Match end; fun dest_list (Const (@{const_syntax Nil}, _)) = [] | dest_list (Const (@{const_syntax Cons}, _) $ x $ xs) = x :: dest_list xs | dest_list _ = raise Match; fun guards_lst_tr' [fg] = fg | guards_lst_tr' (t :: ts) = Syntax.const @{syntax_const "_grds"} $ t $ guards_lst_tr' ts | guards_lst_tr' [] = raise Match; fun new_AnnCom r c = (Const (@{const_syntax AnnCom}, dummyT) $ r $ c) fun new_Pair a b = (Const (@{const_syntax Pair}, dummyT) $ a $ b) fun dest_prod (Const (@{const_syntax Pair}, _) $ a $ b) = (a, b) | dest_prod _ = raise Match; fun prgs_tr' f pqa c = let val (p, qa) = dest_prod pqa val (q, a) = dest_prod qa in [new_AnnCom (f p) (f c), f q, f a] end fun prgs_lst_tr' [p] [c] = list_comb (Syntax.const @{syntax_const "_prg"}, prgs_tr' I p c) | prgs_lst_tr' (p :: ps) (c :: cs) = list_comb (Syntax.const @{syntax_const "_prgs"}, prgs_tr' I p c) $ prgs_lst_tr' ps cs | prgs_lst_tr' _ _= raise Match; fun AnnCom_tr (Const (@{const_syntax AnnPar}, _) $ (Const (@{const_syntax map}, _) $ Abs (i, T, p) $ (Const _ $ j $ k)) :: Const (@{const_syntax Parallel}, _) $ (Const (@{const_syntax map}, _) $ Abs (_, _, c) $ _) :: ts) = let val _ = if syntax_debug then writeln "prg_scheme" else () - fun dest_abs body = snd (Term.dest_abs (i, T, body)) in + fun dest_abs body = snd (Term.dest_abs_global (Abs (i, T, body))) in Syntax.const @{syntax_const "_PAR"} $ list_comb (Syntax.const @{syntax_const "_prg_scheme"} $ j $ Free (i, T) $ k, prgs_tr' dest_abs p c) end | AnnCom_tr (Const (@{const_syntax AnnPar}, _) $ ps :: Const (@{const_syntax Parallel}, _) $ cs :: ts) = let val _ = if syntax_debug then writeln "Par" else () val (ps', cs') = (dest_list ps, dest_list cs) in Syntax.const @{syntax_const "_PAR"} $ prgs_lst_tr' ps' cs' end | AnnCom_tr (Const (@{const_syntax AnnExpr}, _) $ r :: Const (@{const_syntax Basic}, _) $ Abs (x, _, f $ k $ Bound 0) :: ts) = let val _ = if syntax_debug then writeln "Basic'" else () in quote_tr' (Syntax.const @{syntax_const "_AnnAssign"} $ r $ Syntax_Trans.update_name_tr' f) (Abs (x, dummyT, Syntax_Trans.const_abs_tr' k) :: ts) end | AnnCom_tr (Const (@{const_syntax AnnExpr}, _) $ r :: Const (@{const_syntax Basic}, _) $ (f $ k) :: ts) = let val _ = if syntax_debug then writeln "Basic" else () in quote_tr' (Syntax.const @{syntax_const "_AnnAssign"} $ r $ Syntax_Trans.update_name_tr' f) (k :: ts) end | AnnCom_tr (Const (@{const_syntax AnnExpr}, _) $ r :: Const (@{const_syntax Spec}, _) $ (_ $ _ $ Abs (_,_, _ $ _ $ ((_ $ f) $ S $ _))) :: ts) = let val _ = if syntax_debug then writeln ("Spec") else () in (Syntax.const @{syntax_const "_AnnSpec"} $ r $ Syntax_Trans.update_name_tr' f $ Syntax_Trans.antiquote_tr' @{syntax_const "_antiquote"} S) end | AnnCom_tr (Const (@{const_syntax AnnComp}, _) $ r $ r' :: Const (@{const_syntax Seq}, _) $ c $ c' :: ts) = let val _ = if syntax_debug then writeln "Seq" else () in Syntax.const @{syntax_const "_AnnSeq"} $ new_AnnCom r c $ new_AnnCom r' c' end | AnnCom_tr (Const (@{const_syntax AnnRec}, _) $ r $ p :: Const (@{const_syntax Await}, _) $ b $ c :: ts) = let val _ = if syntax_debug then writeln "Await" else () in annbexp_tr' @{syntax_const "_AnnAwait"} (r :: b :: new_AnnCom p c :: ts) end | AnnCom_tr (Const (@{const_syntax AnnWhile}, _) $ r $ i $ p :: Const (@{const_syntax While}, _) $ b $ c :: ts) = let val _ = if syntax_debug then writeln "While" else () in annbexp_tr' @{syntax_const "_AnnWhile"} (r :: b :: i :: new_AnnCom p c :: ts) end | AnnCom_tr (Const (@{const_syntax AnnBin}, _) $ r $ p $ p' :: Const (@{const_syntax Cond}, _) $ b $ c $ c':: ts) = let val _ = if syntax_debug then writeln "Cond" else () in annbexp_tr' @{syntax_const "_AnnCond1"} (r :: b :: new_AnnCom p c :: new_AnnCom p' c' :: ts) end | AnnCom_tr (Const (@{const_syntax ann_guards}, _) $ r $ gs $ p :: Const (@{const_syntax guards}, _) $ _ $ c :: ts) = let val _ = if syntax_debug then writeln "guards" else () in Syntax.const @{syntax_const "_guards"} $ r $ guards_lst_tr' (dest_list gs) $ new_AnnCom p c end | AnnCom_tr (Const (@{const_syntax AnnRec}, _) $ r $ p :: Const (@{const_syntax Guard}, _) $ f $ g $ c :: ts) = let val _ = if syntax_debug then writeln "guards'" else () in Syntax.const @{syntax_const "_guards"} $ r $ new_Pair f g $ new_AnnCom p c end | AnnCom_tr (Const (@{const_syntax AnnCall}, _) $ r $ n :: Const (@{const_syntax Call}, _) $ p :: ts) = let val _ = if syntax_debug then writeln "SCall" else () in Syntax.const @{syntax_const "_AnnSCall"} $ r $ p $ n end | AnnCom_tr (Const (@{const_syntax ann_call}, _) $ ai $ r $ n $ arestoreq $ areturn $ arestorea $ A :: Const (@{const_syntax call}, _) $ init $ p $ restore $ return :: ts) = let val _ = if syntax_debug then writeln "CallX" else () in Syntax.const @{syntax_const "_AnnCallX"} $ ai $ init $ r $ p $ n $ restore $ return $ arestoreq $ areturn $ arestorea $ A end | AnnCom_tr (Const (@{const_syntax AnnComp}, _) $ r $ r' :: Const (@{const_syntax Catch}, _) $ c $ c' :: ts) = let val _ = if syntax_debug then writeln "Catch" else () in Syntax.const @{syntax_const "_Try_Catch"} $ new_AnnCom r c $ new_AnnCom r' c' end | AnnCom_tr (Const (@{const_syntax ann}, _) $ p :: Const (@{const_syntax com}, _) $ p' :: ts) = let val _ = if syntax_debug then writeln "ann_com" else () in if p = p' then p else raise Match end | AnnCom_tr x = let val _ = if syntax_debug then writeln ("AnnCom_tr\n " ^ @{make_string} x) else () in raise Match end; fun oghoare_tr (gamma :: sigma :: F :: r :: c :: Q :: A :: ts) = let val _ = if syntax_debug then writeln "oghoare" else () in Syntax.const @{syntax_const "_oghoare"} $ gamma $ sigma $ F $ new_AnnCom r c $ Q $ A end | oghoare_tr x = let val _ = writeln ("oghoare_tr\n " ^ @{make_string} x) in raise Match end; fun oghoare_seq_tr (gamma :: sigma :: F :: P :: r :: c :: Q :: A :: ts) = let val _ = if syntax_debug then writeln "oghoare_seq" else () in Syntax.const @{syntax_const "_oghoare_seq"} $ gamma $ sigma $ F $ P $ new_AnnCom r c $ Q $ A end | oghoare_seq_tr x = let val _ = writeln ("oghoare_seq_tr\n " ^ @{make_string} x) in raise Match end; in [(@{const_syntax Collect}, K assert_tr'), (@{const_syntax AnnCom}, K AnnCom_tr), (@{const_syntax oghoare}, K oghoare_tr), (@{const_syntax oghoare_seq}, K oghoare_seq_tr)] 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/Monad_Memo_DP/example/Bellman_Ford.thy b/thys/Monad_Memo_DP/example/Bellman_Ford.thy --- a/thys/Monad_Memo_DP/example/Bellman_Ford.thy +++ b/thys/Monad_Memo_DP/example/Bellman_Ford.thy @@ -1,1138 +1,1137 @@ subsection \The Bellman-Ford Algorithm\ theory Bellman_Ford imports "HOL-Library.IArray" "HOL-Library.Code_Target_Numeral" "HOL-Library.Product_Lexorder" "HOL-Library.RBT_Mapping" "../heap_monad/Heap_Main" Example_Misc "../util/Tracing" "../util/Ground_Function" begin subsubsection \Misc\ lemma nat_le_cases: fixes n :: nat assumes "i \ n" obtains "i < n" | "i = n" using assms by (cases "i = n") auto context dp_consistency_iterator begin lemma crel_vs_iterate_state: "crel_vs (=) () (iter_state f x)" if "((=) ===>\<^sub>T R) g f" by (metis crel_vs_iterate_state iter_state_iterate_state that) lemma consistent_crel_vs_iterate_state: "crel_vs (=) () (iter_state f x)" if "consistentDP f" using consistentDP_def crel_vs_iterate_state that by simp end instance extended :: (countable) countable proof standard obtain to_nat :: "'a \ nat" where "inj to_nat" by auto let ?f = "\ x. case x of Fin n \ to_nat n + 2 | Pinf \ 0 | Minf \ 1" from \inj _ \ have "inj ?f" by (auto simp: inj_def split: extended.split) then show "\to_nat :: 'a extended \ nat. inj to_nat" by auto qed instance extended :: (heap) heap .. instantiation "extended" :: (conditionally_complete_lattice) complete_lattice begin definition "Inf A = ( if A = {} \ A = {\} then \ else if -\ \ A \ \ bdd_below (Fin -` A) then -\ else Fin (Inf (Fin -` A)))" definition "Sup A = ( if A = {} \ A = {-\} then -\ else if \ \ A \ \ bdd_above (Fin -` A) then \ else Fin (Sup (Fin -` A)))" instance proof standard have [dest]: "Inf (Fin -` A) \ x" if "Fin x \ A" "bdd_below (Fin -` A)" for A and x :: 'a using that by (intro cInf_lower) auto have *: False if "\ z \ Inf (Fin -` A)" "\x. x \ A \ Fin z \ x" "Fin x \ A" for A and x z :: 'a using cInf_greatest[of "Fin -` A" z] that vimage_eq by force show "Inf A \ x" if "x \ A" for x :: "'a extended" and A using that unfolding Inf_extended_def by (cases x) auto show "z \ Inf A" if "\x. x \ A \ z \ x" for z :: "'a extended" and A using that unfolding Inf_extended_def apply (clarsimp; safe) apply force apply force subgoal by (cases z; force simp: bdd_below_def) subgoal by (cases z; force simp: bdd_below_def) subgoal for x y by (cases z; cases y) (auto elim: *) subgoal for x y by (cases z; cases y; simp; metis * less_eq_extended.elims(2)) done have [dest]: "x \ Sup (Fin -` A)" if "Fin x \ A" "bdd_above (Fin -` A)" for A and x :: 'a using that by (intro cSup_upper) auto have *: False if "\ Sup (Fin -` A) \ z" "\x. x \ A \ x \ Fin z" "Fin x \ A" for A and x z :: 'a using cSup_least[of "Fin -` A" z] that vimage_eq by force show "x \ Sup A" if "x \ A" for x :: "'a extended" and A using that unfolding Sup_extended_def by (cases x) auto show "Sup A \ z" if "\x. x \ A \ x \ z" for z :: "'a extended" and A using that unfolding Sup_extended_def apply (clarsimp; safe) apply force apply force subgoal by (cases z; force) subgoal by (cases z; force) subgoal for x y by (cases z; cases y) (auto elim: *) subgoal for x y by (cases z; cases y; simp; metis * extended.exhaust) done show "Inf {} = (top::'a extended)" unfolding Inf_extended_def top_extended_def by simp show "Sup {} = (bot::'a extended)" unfolding Sup_extended_def bot_extended_def by simp qed end instance "extended" :: ("{conditionally_complete_lattice,linorder}") complete_linorder .. lemma Minf_eq_zero[simp]: "-\ = 0 \ False" and Pinf_eq_zero[simp]: "\ = 0 \ False" unfolding zero_extended_def by auto lemma Sup_int: fixes x :: int and X :: "int set" assumes "X \ {}" "bdd_above X" shows "Sup X \ X \ (\y\X. y \ Sup X)" proof - from assms obtain x y where "X \ {..y}" "x \ X" by (auto simp: bdd_above_def) then have *: "finite (X \ {x..y})" "X \ {x..y} \ {}" and "x \ y" by (auto simp: subset_eq) have "\!x\X. (\y\X. y \ x)" proof { fix z assume "z \ X" have "z \ Max (X \ {x..y})" proof cases assume "x \ z" with \z \ X\ \X \ {..y}\ *(1) show ?thesis by (auto intro!: Max_ge) next assume "\ x \ z" then have "z < x" by simp also have "x \ Max (X \ {x..y})" using \x \ X\ *(1) \x \ y\ by (intro Max_ge) auto finally show ?thesis by simp qed } note le = this with Max_in[OF *] show ex: "Max (X \ {x..y}) \ X \ (\z\X. z \ Max (X \ {x..y}))" by auto fix z assume *: "z \ X \ (\y\X. y \ z)" with le have "z \ Max (X \ {x..y})" by auto moreover have "Max (X \ {x..y}) \ z" using * ex by auto ultimately show "z = Max (X \ {x..y})" by auto qed then show "Sup X \ X \ (\y\X. y \ Sup X)" unfolding Sup_int_def by (rule theI') qed lemmas Sup_int_in = Sup_int[THEN conjunct1] lemma Inf_int_in: fixes S :: "int set" assumes "S \ {}" "bdd_below S" shows "Inf S \ S" using assms unfolding Inf_int_def by (smt Sup_int_in bdd_above_uminus image_iff image_is_empty) lemma finite_setcompr_eq_image: "finite {f x |x. P x} \ finite (f ` {x. P x})" by (simp add: setcompr_eq_image) lemma finite_lists_length_le1: "finite {xs. length xs \ i \ set xs \ {0..(n::nat)}}" for i by (auto intro: finite_subset[OF _ finite_lists_length_le[OF finite_atLeastAtMost]]) lemma finite_lists_length_le2: "finite {xs. length xs + 1 \ i \ set xs \ {0..(n::nat)}}" for i by (auto intro: finite_subset[OF _ finite_lists_length_le1[of "i"]]) lemmas [simp] = finite_setcompr_eq_image finite_lists_length_le2[simplified] finite_lists_length_le1 lemma get_return: "run_state (State_Monad.bind State_Monad.get (\ m. State_Monad.return (f m))) m = (f m, m)" by (simp add: State_Monad.bind_def State_Monad.get_def) lemma list_pidgeonhole: assumes "set xs \ S" "card S < length xs" "finite S" obtains as a bs cs where "xs = as @ a # bs @ a # cs" proof - from assms have "\ distinct xs" by (metis card_mono distinct_card not_le) then show ?thesis by (metis append.assoc append_Cons not_distinct_conv_prefix split_list that) qed lemma path_eq_cycleE: assumes "v # ys @ [t] = as @ a # bs @ a # cs" obtains (Nil_Nil) "as = []" "cs = []" "v = a" "a = t" "ys = bs" | (Nil_Cons) cs' where "as = []" "v = a" "ys = bs @ a # cs'" "cs = cs' @ [t]" | (Cons_Nil) as' where "as = v # as'" "cs = []" "a = t" "ys = as' @ a # bs" | (Cons_Cons) as' cs' where "as = v # as'" "cs = cs' @ [t]" "ys = as' @ a # bs @ a # cs'" using assms by (auto simp: Cons_eq_append_conv append_eq_Cons_conv append_eq_append_conv2) lemma le_add_same_cancel1: "a + b \ a \ b \ 0" if "a < \" "-\ < a" for a b :: "int extended" using that by (cases a; cases b) (auto simp add: zero_extended_def) lemma add_gt_minfI: assumes "-\ < a" "-\ < b" shows "-\ < a + b" using assms by (cases a; cases b) auto lemma add_lt_infI: assumes "a < \" "b < \" shows "a + b < \" using assms by (cases a; cases b) auto lemma sum_list_not_infI: "sum_list xs < \" if "\ x \ set xs. x < \" for xs :: "int extended list" using that apply (induction xs) apply (simp add: zero_extended_def)+ by (smt less_extended_simps(2) plus_extended.elims) lemma sum_list_not_minfI: "sum_list xs > -\" if "\ x \ set xs. x > -\" for xs :: "int extended list" using that by (induction xs) (auto intro: add_gt_minfI simp: zero_extended_def) subsubsection \Single-Sink Shortest Path Problem\ datatype bf_result = Path "nat list" int | No_Path | Computation_Error context fixes n :: nat and W :: "nat \ nat \ int extended" begin context fixes t :: nat \ \Final node\ begin text \ The correctness proof closely follows Kleinberg \&\ Tardos: "Algorithm Design", chapter "Dynamic Programming" @{cite "Kleinberg-Tardos"} \ fun weight :: "nat list \ int extended" where "weight [v] = 0" | "weight (v # w # xs) = W v w + weight (w # xs)" definition "OPT i v = ( Min ( {weight (v # xs @ [t]) | xs. length xs + 1 \ i \ set xs \ {0..n}} \ {if t = v then 0 else \} ) )" lemma weight_alt_def': "weight (s # xs) + w = snd (fold (\j (i, x). (j, W i j + x)) xs (s, w))" by (induction xs arbitrary: s w; simp; smt add.commute add.left_commute) lemma weight_alt_def: "weight (s # xs) = snd (fold (\j (i, x). (j, W i j + x)) xs (s, 0))" by (rule weight_alt_def'[of s xs 0, simplified]) lemma weight_append: "weight (xs @ a # ys) = weight (xs @ [a]) + weight (a # ys)" by (induction xs rule: weight.induct; simp add: add.assoc) lemma OPT_0: "OPT 0 v = (if t = v then 0 else \)" unfolding OPT_def by simp subsubsection \Functional Correctness\ lemma OPT_cases: obtains (path) xs where "OPT i v = weight (v # xs @ [t])" "length xs + 1 \ i" "set xs \ {0..n}" | (sink) "v = t" "OPT i v = 0" | (unreachable) "v \ t" "OPT i v = \" unfolding OPT_def using Min_in[of "{weight (v # xs @ [t]) |xs. length xs + 1 \ i \ set xs \ {0..n}} \ {if t = v then 0 else \}"] by (auto simp: finite_lists_length_le2[simplified] split: if_split_asm) lemma OPT_Suc: "OPT (Suc i) v = min (OPT i v) (Min {OPT i w + W v w | w. w \ n})" (is "?lhs = ?rhs") if "t \ n" proof - have "OPT i w + W v w \ OPT (Suc i) v" if "w \ n" for w using OPT_cases[of i w] proof cases case (path xs) with \w \ n\ show ?thesis by (subst OPT_def) (auto intro!: Min_le exI[where x = "w # xs"] simp: add.commute) next case sink then show ?thesis by (subst OPT_def) (auto intro!: Min_le exI[where x = "[]"]) next case unreachable then show ?thesis by simp qed then have "Min {OPT i w + W v w |w. w \ n} \ OPT (Suc i) v" by (auto intro!: Min.boundedI) moreover have "OPT i v \ OPT (Suc i) v" unfolding OPT_def by (rule Min_antimono) auto ultimately have "?lhs \ ?rhs" by simp from OPT_cases[of "Suc i" v] have "?lhs \ ?rhs" proof cases case (path xs) note [simp] = path(1) from path consider (zero) "i = 0" "length xs = 0" | (new) "i > 0" "length xs = i" | (old) "length xs < i" by (cases "length xs = i") auto then show ?thesis proof cases case zero with path have "OPT (Suc i) v = W v t" by simp also have "W v t = OPT i t + W v t" unfolding OPT_def using \i = 0\ by auto also have "\ \ Min {OPT i w + W v w |w. w \ n}" using \t \ n\ by (auto intro: Min_le) finally show ?thesis by (rule min.coboundedI2) next case new with \_ = i\ obtain u ys where [simp]: "xs = u # ys" by (cases xs) auto from path have "OPT i u \ weight (u # ys @ [t])" unfolding OPT_def by (intro Min_le) auto from path have "Min {OPT i w + W v w |w. w \ n} \ W v u + OPT i u" by (intro Min_le) (auto simp: add.commute) also from \OPT i u \ _\ have "\ \ OPT (Suc i) v" by (simp add: add_left_mono) finally show ?thesis by (rule min.coboundedI2) next case old with path have "OPT i v \ OPT (Suc i) v" by (auto 4 3 intro: Min_le simp: OPT_def) then show ?thesis by (rule min.coboundedI1) qed next case unreachable then show ?thesis by simp next case sink then have "OPT i v \ OPT (Suc i) v" unfolding OPT_def by simp then show ?thesis by (rule min.coboundedI1) qed with \?lhs \ ?rhs\ show ?thesis by (rule order.antisym) qed fun bf :: "nat \ nat \ int extended" where "bf 0 v = (if t = v then 0 else \)" | "bf (Suc i) v = min_list (bf i v # [W v w + bf i w . w \ [0 ..< Suc n]])" lemmas [simp del] = bf.simps lemmas bf_simps[simp] = bf.simps[unfolded min_list_fold] lemma bf_correct: "OPT i j = bf i j" if \t \ n\ proof (induction i arbitrary: j) case 0 then show ?case by (simp add: OPT_0) next case (Suc i) have *: "{bf i w + W j w |w. w \ n} = set (map (\w. W j w + bf i w) [0..t \ n\ show ?case by (simp add: OPT_Suc del: upt_Suc, subst Min.set_eq_fold[symmetric], auto simp: *) qed subsubsection \Functional Memoization\ memoize_fun bf\<^sub>m: bf with_memory dp_consistency_mapping monadifies (state) bf.simps text \Generated Definitions\ context includes state_monad_syntax begin thm bf\<^sub>m'.simps bf\<^sub>m_def end text \Correspondence Proof\ memoize_correct by memoize_prover print_theorems lemmas [code] = bf\<^sub>m.memoized_correct interpretation iterator "\ (x, y). x \ n \ y \ n" "\ (x, y). if y < n then (x, y + 1) else (x + 1, 0)" "\ (x, y). x * (n + 1) + y" by (rule table_iterator_up) interpretation bottom_up: dp_consistency_iterator_empty "\ (_::(nat \ nat, int extended) mapping). True" "\ (x, y). bf x y" "\ k. do {m \ State_Monad.get; State_Monad.return (Mapping.lookup m k :: int extended option)}" "\ k v. do {m \ State_Monad.get; State_Monad.set (Mapping.update k v m)}" "\ (x, y). x \ n \ y \ n" "\ (x, y). if y < n then (x, y + 1) else (x + 1, 0)" "\ (x, y). x * (n + 1) + y" Mapping.empty .. definition "iter_bf = iter_state (\ (x, y). bf\<^sub>m' x y)" lemma iter_bf_unfold[code]: "iter_bf = (\ (i, j). (if i \ n \ j \ n then do { bf\<^sub>m' i j; iter_bf (if j < n then (i, j + 1) else (i + 1, 0)) } else State_Monad.return ()))" unfolding iter_bf_def by (rule ext) (safe, clarsimp simp: iter_state_unfold) lemmas bf_memoized = bf\<^sub>m.memoized[OF bf\<^sub>m.crel] lemmas bf_bottom_up = bottom_up.memoized[OF bf\<^sub>m.crel, folded iter_bf_def] text \ This will be our final implementation, which includes detection of negative cycles. See the corresponding section below for the correctness proof. \ definition "bellman_ford \ do { _ \ iter_bf (n, n); xs \ State_Main.map\<^sub>T' (\i. bf\<^sub>m' n i) [0.. State_Main.map\<^sub>T' (\i. bf\<^sub>m' (n + 1) i) [0.. do { _ \ iter_bf (n, n); (\\xs. \\ys. State_Monad.return (if xs = ys then Some xs else None)\ . (State_Main.map\<^sub>T . \\i. bf\<^sub>m' (n + 1) i\ . \[0..)\) . (State_Main.map\<^sub>T . \\i. bf\<^sub>m' n i\ . \[0..) }" unfolding State_Monad_Ext.fun_app_lifted_def bellman_ford_def State_Main.map\<^sub>T_def bind_left_identity . end subsubsection \Imperative Memoization\ context fixes mem :: "nat ref \ nat ref \ int extended option array ref \ int extended option array ref" assumes mem_is_init: "mem = result_of (init_state (n + 1) 1 0) Heap.empty" begin lemma [intro]: "dp_consistency_heap_array_pair' (n + 1) fst snd id 1 0 mem" by (standard; simp add: mem_is_init injective_def) interpretation iterator "\ (x, y). x \ n \ y \ n" "\ (x, y). if y < n then (x, y + 1) else (x + 1, 0)" "\ (x, y). x * (n + 1) + y" by (rule table_iterator_up) lemma [intro]: "dp_consistency_heap_array_pair_iterator (n + 1) fst snd id 1 0 mem (\ (x, y). if y < n then (x, y + 1) else (x + 1, 0)) (\ (x, y). x * (n + 1) + y) (\ (x, y). x \ n \ y \ n)" by (standard; simp add: mem_is_init injective_def) memoize_fun bf\<^sub>h: bf with_memory (default_proof) dp_consistency_heap_array_pair_iterator where size = "n + 1" and key1 = "fst :: nat \ nat \ nat" and key2 = "snd :: nat \ nat \ nat" and k1 = "1 :: nat" and k2 = "0 :: nat" and to_index = "id :: nat \ nat" and mem = mem and cnt = "\ (x, y). x \ n \ y \ n" and nxt = "\ (x :: nat, y). if y < n then (x, y + 1) else (x + 1, 0)" and sizef = "\ (x, y). x * (n + 1) + y" monadifies (heap) bf.simps memoize_correct by memoize_prover lemmas memoized_empty = bf\<^sub>h.memoized_empty[OF bf\<^sub>h.consistent_DP_iter_and_compute[OF bf\<^sub>h.crel]] lemmas iter_heap_unfold = iter_heap_unfold end (* Fixed Memory *) subsubsection \Detecting Negative Cycles\ definition "shortest v = ( Inf ( {weight (v # xs @ [t]) | xs. set xs \ {0..n}} \ {if t = v then 0 else \} ) )" definition "is_path xs \ weight (xs @ [t]) < \" definition "has_negative_cycle \ \xs a ys. set (a # xs @ ys) \ {0..n} \ weight (a # xs @ [a]) < 0 \ is_path (a # ys)" definition "reaches a \ \xs. is_path (a # xs) \ a \ n \ set xs \ {0..n}" lemma fold_sum_aux': assumes "\u \ set (a # xs). \v \ set (xs @ [b]). f v + W u v \ f u" shows "sum_list (map f (a # xs)) \ sum_list (map f (xs @ [b])) + weight (a # xs @ [b])" using assms by (induction xs arbitrary: a; simp) (smt ab_semigroup_add_class.add_ac(1) add.left_commute add_mono) lemma fold_sum_aux: assumes "\u \ set (a # xs). \v \ set (a # xs). f v + W u v \ f u" shows "sum_list (map f (a # xs @ [a])) \ sum_list (map f (a # xs @ [a])) + weight (a # xs @ [a])" using fold_sum_aux'[of a xs a f] assms by auto (metis (no_types, opaque_lifting) add.assoc add.commute add_left_mono) context begin private definition "is_path2 xs \ weight xs < \" private lemma is_path2_remove_cycle: assumes "is_path2 (as @ a # bs @ a # cs)" shows "is_path2 (as @ a # cs)" proof - have "weight (as @ a # bs @ a # cs) = weight (as @ [a]) + weight (a # bs @ [a]) + weight (a # cs)" by (metis Bellman_Ford.weight_append append_Cons append_assoc) with assms have "weight (as @ [a]) < \" "weight (a # cs) < \" unfolding is_path2_def by (simp, metis Pinf_add_right antisym less_extended_simps(4) not_less add.commute)+ then show ?thesis unfolding is_path2_def by (subst weight_append) (rule add_lt_infI) qed private lemma is_path_eq: "is_path xs \ is_path2 (xs @ [t])" unfolding is_path_def is_path2_def .. lemma is_path_remove_cycle: assumes "is_path (as @ a # bs @ a # cs)" shows "is_path (as @ a # cs)" using assms unfolding is_path_eq by (simp add: is_path2_remove_cycle) lemma is_path_remove_cycle2: assumes "is_path (as @ t # cs)" shows "is_path as" using assms unfolding is_path_eq by (simp add: is_path2_remove_cycle) end (* private lemmas *) lemma is_path_shorten: assumes "is_path (i # xs)" "i \ n" "set xs \ {0..n}" "t \ n" "t \ i" obtains xs where "is_path (i # xs)" "i \ n" "set xs \ {0..n}" "length xs < n" proof (cases "length xs < n") case True with assms show ?thesis by (auto intro: that) next case False then have "length xs \ n" by auto with assms(1,3) show ?thesis proof (induction "length xs" arbitrary: xs rule: less_induct) case less then have "length (i # xs @ [t]) > card ({0..n})" by auto moreover from less.prems \i \ n\ \t \ n\ have "set (i # xs @ [t]) \ {0..n}" by auto ultimately obtain a as bs cs where *: "i # xs @ [t] = as @ a # bs @ a # cs" by (elim list_pidgeonhole) auto obtain ys where ys: "is_path (i # ys)" "length ys < length xs" "set (i # ys) \ {0..n}" apply atomize_elim using * proof (cases rule: path_eq_cycleE) case Nil_Nil with \t \ i\ show "\ys. is_path (i # ys) \ length ys < length xs \ set (i # ys) \ {0..n}" by auto next case (Nil_Cons cs') then show "\ys. is_path (i # ys) \ length ys < length xs \ set (i # ys) \ {0..n}" using \set (i # xs @ [t]) \ {0..n}\ \is_path (i # xs)\ is_path_remove_cycle[of "[]"] by - (rule exI[where x = cs'], simp) next case (Cons_Nil as') then show "\ys. is_path (i # ys) \ length ys < length xs \ set (i # ys) \ {0..n}" using \set (i # xs @ [t]) \ {0..n}\ \is_path (i # xs)\ by - (rule exI[where x = as'], auto intro: is_path_remove_cycle2) next case (Cons_Cons as' cs') then show "\ys. is_path (i # ys) \ length ys < length xs \ set (i # ys) \ {0..n}" using \set (i # xs @ [t]) \ {0..n}\ \is_path (i # xs)\ is_path_remove_cycle[of "i # as'"] by - (rule exI[where x = "as' @ a # cs'"], auto) qed then show ?thesis by (cases "n \ length ys") (auto intro: that less) qed qed lemma reaches_non_inf_path: assumes "reaches i" "i \ n" "t \ n" shows "OPT n i < \" proof (cases "t = i") case True with \i \ n\ \t \ n\ have "OPT n i \ 0" unfolding OPT_def by (auto intro: Min_le simp: finite_lists_length_le2[simplified]) then show ?thesis using less_linear by (fastforce simp: zero_extended_def) next case False from assms(1) obtain xs where "is_path (i # xs)" "i \ n" "set xs \ {0..n}" unfolding reaches_def by safe then obtain xs where xs: "is_path (i # xs)" "i \ n" "set xs \ {0..n}" "length xs < n" using \t \ i\ \t \ n\ by (auto intro: is_path_shorten) then have "weight (i # xs @ [t]) < \" unfolding is_path_def by auto with xs(2-) show ?thesis unfolding OPT_def by (elim order.strict_trans1[rotated]) (auto simp: setcompr_eq_image finite_lists_length_le2[simplified]) qed lemma OPT_sink_le_0: "OPT i t \ 0" unfolding OPT_def by (auto simp: finite_lists_length_le2[simplified]) lemma is_path_appendD: assumes "is_path (as @ a # bs)" shows "is_path (a # bs)" using assms weight_append[of as a "bs @ [t]"] unfolding is_path_def by simp (metis Pinf_add_right add.commute less_extended_simps(4) not_less_iff_gr_or_eq) lemma has_negative_cycleI: assumes "set (a # xs @ ys) \ {0..n}" "weight (a # xs @ [a]) < 0" "is_path (a # ys)" shows has_negative_cycle using assms unfolding has_negative_cycle_def by auto lemma OPT_cases2: obtains (path) xs where "v \ t" "OPT i v \ \" "OPT i v = weight (v # xs @ [t])" "length xs + 1 \ i" "set xs \ {0..n}" | (unreachable) "v \ t" "OPT i v = \" | (sink) "v = t" "OPT i v \ 0" unfolding OPT_def using Min_in[of "{weight (v # xs @ [t]) |xs. length xs + 1 \ i \ set xs \ {0..n}} \ {if t = v then 0 else \}"] by (cases "v = t"; force simp: finite_lists_length_le2[simplified] split: if_split_asm) lemma shortest_le_OPT: assumes "v \ n" shows "shortest v \ OPT i v" unfolding OPT_def shortest_def apply (subst Min_Inf) apply (simp add: setcompr_eq_image finite_lists_length_le2[simplified]; fail)+ apply (rule Inf_superset_mono) apply auto done context assumes W_wellformed: "\i \ n. \j \ n. W i j > -\" assumes "t \ n" begin lemma weight_not_minfI: "-\ < weight xs" if "set xs \ {0..n}" "xs \ []" using that using W_wellformed \t \ n\ by (induction xs rule: induct_list012) (auto intro: add_gt_minfI simp: zero_extended_def) lemma OPT_not_minfI: "OPT n i > -\" if "i \ n" proof - have "OPT n i \ {weight (i # xs @ [t]) |xs. length xs + 1 \ n \ set xs \ {0..n}} \ {if t = i then 0 else \}" unfolding OPT_def by (rule Min_in) (auto simp: setcompr_eq_image finite_lists_length_le2[simplified]) with that \t \ n\ show ?thesis by (auto 4 3 intro!: weight_not_minfI simp: zero_extended_def) qed theorem detects_cycle: assumes has_negative_cycle shows "\i \ n. OPT (n + 1) i < OPT n i" proof - from assms \t \ n\ obtain xs a ys where cycle: "a \ n" "set xs \ {0..n}" "set ys \ {0..n}" "weight (a # xs @ [a]) < 0" "is_path (a # ys)" unfolding has_negative_cycle_def by clarsimp then have "reaches a" unfolding reaches_def by auto have reaches: "reaches x" if "x \ set xs" for x proof - from that obtain as bs where "xs = as @ x # bs" by atomize_elim (rule split_list) with cycle have "weight (x # bs @ [a]) < \" using weight_append[of "a # as" x "bs @ [a]"] by simp (metis Pinf_add_right Pinf_le add.commute less_eq_extended.simps(2) not_less) moreover from \reaches a\ obtain cs where "local.weight (a # cs @ [t]) < \" "set cs \ {0..n}" unfolding reaches_def is_path_def by auto ultimately show ?thesis unfolding reaches_def is_path_def using \a \ n\ weight_append[of "x # bs" a "cs @ [t]"] cycle(2) \xs = _\ by - (rule exI[where x = "bs @ [a] @ cs"], auto intro: add_lt_infI) qed let ?S = "sum_list (map (OPT n) (a # xs @ [a]))" obtain u v where "u \ n" "v \ n" "OPT n v + W u v < OPT n u" proof (atomize_elim, rule ccontr) assume "\u v. u \ n \ v \ n \ OPT n v + W u v < OPT n u" then have "?S \ ?S + weight (a # xs @ [a])" using cycle(1-3) by (subst fold_sum_aux; fastforce simp: subset_eq) moreover have "?S > -\" using cycle(1-4) by (intro sum_list_not_minfI, auto intro!: OPT_not_minfI) moreover have "?S < \" using reaches \t \ n\ cycle(1,2) by (intro sum_list_not_infI) (auto intro: reaches_non_inf_path \reaches a\ simp: subset_eq) ultimately have "weight (a # xs @ [a]) \ 0" by (simp add: le_add_same_cancel1) with \weight _ < 0\ show False by simp qed then show ?thesis by - (rule exI[where x = u], auto 4 4 intro: Min.coboundedI min.strict_coboundedI2 elim: order.strict_trans1[rotated] simp: OPT_Suc[OF \t \ n\]) qed corollary bf_detects_cycle: assumes has_negative_cycle shows "\i \ n. bf (n + 1) i < bf n i" using detects_cycle[OF assms] unfolding bf_correct[OF \t \ n\] . lemma shortest_cases: assumes "v \ n" obtains (path) xs where "shortest v = weight (v # xs @ [t])" "set xs \ {0..n}" | (sink) "v = t" "shortest v = 0" | (unreachable) "v \ t" "shortest v = \" | (negative_cycle) "shortest v = -\" "\x. \xs. set xs \ {0..n} \ weight (v # xs @ [t]) < Fin x" proof - let ?S = "{weight (v # xs @ [t]) | xs. set xs \ {0..n}} \ {if t = v then 0 else \}" have "?S \ {}" by auto have Minf_lowest: False if "-\ < a" "-\ = a" for a :: "int extended" using that by auto show ?thesis proof (cases "shortest v") case (Fin x) then have "-\ \ ?S" "bdd_below (Fin -` ?S)" "?S \ {\}" "x = Inf (Fin -` ?S)" unfolding shortest_def Inf_extended_def by (auto split: if_split_asm) from this(1-3) have "x \ Fin -` ?S" unfolding \x = _\ by (intro Inf_int_in, auto simp: zero_extended_def) (smt empty_iff extended.exhaust insertI2 mem_Collect_eq vimage_eq) with \shortest v = _\ show ?thesis unfolding vimage_eq by (auto split: if_split_asm intro: that) next case Pinf with \?S \ {}\ have "t \ v" unfolding shortest_def Inf_extended_def by (auto split: if_split_asm) with \_ = \\ show ?thesis by (auto intro: that) next case Minf then have "?S \ {}" "?S \ {\}" "-\ \ ?S \ \ bdd_below (Fin -` ?S)" unfolding shortest_def Inf_extended_def by (auto split: if_split_asm) from this(3) have "\x. \xs. set xs \ {0..n} \ weight (v # xs @ [t]) < Fin x" proof assume "-\ \ ?S" with weight_not_minfI have False using \v \ n\ \t \ n\ by (auto split: if_split_asm elim: Minf_lowest[rotated]) then show ?thesis .. next assume "\ bdd_below (Fin -` ?S)" show ?thesis proof fix x :: int let ?m = "min x (-1)" from \\ bdd_below _\ obtain m where "Fin m \ ?S" "m < ?m" unfolding bdd_below_def by - (simp, drule spec[of _ "?m"], force) then show "\xs. set xs \ {0..n} \ weight (v # xs @ [t]) < Fin x" by (auto split: if_split_asm simp: zero_extended_def) (metis less_extended_simps(1))+ qed qed with \shortest v = _\ show ?thesis by (auto intro: that) qed qed lemma simple_paths: assumes "\ has_negative_cycle" "weight (v # xs @ [t]) < \" "set xs \ {0..n}" "v \ n" obtains ys where "weight (v # ys @ [t]) \ weight (v # xs @ [t])" "set ys \ {0..n}" "length ys < n" | "v = t" using assms(2-) proof (atomize_elim, induction "length xs" arbitrary: xs rule: less_induct) case (less ys) note ys = less.prems(1,2) note IH = less.hyps have path: "is_path (v # ys)" using is_path_def not_less_iff_gr_or_eq ys(1) by fastforce show ?case proof (cases "length ys \ n") case True with ys \v \ n\ \t \ n\ obtain a as bs cs where "v # ys @ [t] = as @ a # bs @ a # cs" by - (rule list_pidgeonhole[of "v # ys @ [t]" "{0..n}"], auto) then show ?thesis proof (cases rule: path_eq_cycleE) case Nil_Nil then show ?thesis by simp next case (Nil_Cons cs') then have *: "weight (v # ys @ [t]) = weight (a # bs @ [a]) + weight (a # cs' @ [t])" by (simp add: weight_append[of "a # bs" a "cs' @ [t]", simplified]) show ?thesis proof (cases "weight (a # bs @ [a]) < 0") case True with Nil_Cons \set ys \ _\ path show ?thesis using assms(1) by (force intro: has_negative_cycleI[of a bs ys]) next case False then have "weight (a # bs @ [a]) \ 0" by auto with * ys have "weight (a # cs' @ [t]) \ weight (v # ys @ [t])" using add_mono not_le by fastforce with Nil_Cons \length ys \ n\ ys show ?thesis using IH[of cs'] by simp (meson le_less_trans order_trans) qed next case (Cons_Nil as') with ys have *: "weight (v # ys @ [t]) = weight (v # as' @ [t]) + weight (a # bs @ [a])" using weight_append[of "v # as'" t "bs @ [t]"] by simp show ?thesis proof (cases "weight (a # bs @ [a]) < 0") case True with Cons_Nil \set ys \ _\ path assms(1) show ?thesis using is_path_appendD[of "v # as'"] by (force intro: has_negative_cycleI[of a bs bs]) next case False then have "weight (a # bs @ [a]) \ 0" by auto with * ys(1) have "weight (v # as' @ [t]) \ weight (v # ys @ [t])" using add_left_mono by fastforce with Cons_Nil \length ys \ n\ \v \ n\ ys show ?thesis using IH[of as'] by simp (meson le_less_trans order_trans) qed next case (Cons_Cons as' cs') with ys have *: "weight (v # ys @ [t]) = weight (v # as' @ a # cs' @ [t]) + weight (a # bs @ [a])" using weight_append[of "v # as'" a "bs @ a # cs' @ [t]"] weight_append[of "a # bs" a "cs' @ [t]"] weight_append[of "v # as'" a "cs' @ [t]"] by (simp add: algebra_simps) show ?thesis proof (cases "weight (a # bs @ [a]) < 0") case True with Cons_Cons \set ys \ _\ path assms(1) show ?thesis using is_path_appendD[of "v # as'"] by (force intro: has_negative_cycleI[of a bs "bs @ a # cs'"]) next case False then have "weight (a # bs @ [a]) \ 0" by auto with * ys have "weight (v # as' @ a # cs' @ [t]) \ weight (v # ys @ [t])" using add_left_mono by fastforce with Cons_Cons \v \ n\ ys show ?thesis using is_path_remove_cycle2 IH[of "as' @ a # cs'"] by simp (meson le_less_trans order_trans) qed qed next case False with \set ys \ _\ show ?thesis by auto qed qed theorem shorter_than_OPT_n_has_negative_cycle: assumes "shortest v < OPT n v" "v \ n" shows has_negative_cycle proof - from assms obtain ys where ys: "weight (v # ys @ [t]) < OPT n v" "set ys \ {0..n}" apply (cases rule: OPT_cases2[of v n]; cases rule: shortest_cases[OF \v \ n\]; simp) apply (metis uminus_extended.cases) using less_extended_simps(2) less_trans apply blast apply (metis less_eq_extended.elims(2) less_extended_def zero_extended_def) done show ?thesis proof (cases "v = t") case True with ys \t \ n\ show ?thesis using OPT_sink_le_0[of n] unfolding has_negative_cycle_def is_path_def using less_extended_def by force next case False show ?thesis proof (rule ccontr) assume "\ has_negative_cycle" with False False ys \v \ n\ obtain xs where "weight (v # xs @ [t]) \ weight (v # ys @ [t])" "set xs \ {0..n}" "length xs < n" using less_extended_def by (fastforce elim!: simple_paths[of v ys]) then have "OPT n v \ weight (v # xs @ [t])" unfolding OPT_def by (intro Min_le) auto with \_ \ weight (v # ys @ [t])\ \weight (v # ys @ [t]) < OPT n v\ show False by simp qed qed qed corollary detects_cycle_has_negative_cycle: assumes "OPT (n + 1) v < OPT n v" "v \ n" shows has_negative_cycle using assms shortest_le_OPT[of v "n + 1"] shorter_than_OPT_n_has_negative_cycle[of v] by auto corollary bellman_ford_detects_cycle: "has_negative_cycle \ (\v \ n. OPT (n + 1) v < OPT n v)" using detects_cycle_has_negative_cycle detects_cycle by blast corollary bellman_ford_shortest_paths: assumes "\ has_negative_cycle" shows "\v \ n. bf n v = shortest v" proof - have "OPT n v \ shortest v" if "v \ n" for v using that assms shorter_than_OPT_n_has_negative_cycle[of v] by force then show ?thesis unfolding bf_correct[OF \t \ n\, symmetric] by (safe, rule order.antisym) (auto elim: shortest_le_OPT) qed lemma OPT_mono: "OPT m v \ OPT n v" if \v \ n\ \n \ m\ using that unfolding OPT_def by (intro Min_antimono) auto corollary bf_fix: assumes "\ has_negative_cycle" "m \ n" shows "\v \ n. bf m v = bf n v" proof (intro allI impI) fix v assume "v \ n" from \v \ n\ \n \ m\ have "shortest v \ OPT m v" by (simp add: shortest_le_OPT) moreover from \v \ n\ \n \ m\ have "OPT m v \ OPT n v" by (rule OPT_mono) moreover from \v \ n\ assms have "OPT n v \ shortest v" using shorter_than_OPT_n_has_negative_cycle[of v] by force ultimately show "bf m v = bf n v" unfolding bf_correct[OF \t \ n\, symmetric] by simp qed lemma bellman_ford_correct': "bf\<^sub>m.crel_vs (=) (if has_negative_cycle then None else Some (map shortest [0..m' = bf\<^sub>m.crel[unfolded bf\<^sub>m.consistentDP_def, THEN rel_funD, of "(m, x)" "(m, y)" for m x y, unfolded prod.case] have "?l = ?r" supply [simp del] = bf_simps supply [simp add] = bf_fix[rule_format, symmetric] bellman_ford_shortest_paths[rule_format, symmetric] unfolding Wrap_def App_def using bf_detects_cycle by (fastforce elim: nat_le_cases) \ \Slightly transform the goal, then apply parametric reasoning like usual.\ show ?thesis \ \Roughly \ unfolding bellman_ford_alt_def \?l = ?r\ \ \Obtain parametric form.\ apply (rule bf\<^sub>m.crel_vs_bind_ignore[rotated]) \ \Drop bind.\ apply (rule bottom_up.consistent_crel_vs_iterate_state[OF bf\<^sub>m.crel, folded iter_bf_def]) apply (subst Transfer.Rel_def[symmetric]) \ \Setup typical goal for automated reasoner.\ \ \We need to reason manually because we are not in the context where \bf\<^sub>m\ was defined.\ \ \This is roughly what @{method "memoize_prover_match_step"}/\Transform_Tactic.step_tac\ does.\ - ML_prf \val ctxt = @{context}\ - apply (tactic \Transform_Tactic.solve_relator_tac ctxt 1\ + apply (tactic \Transform_Tactic.solve_relator_tac \<^context> 1\ | rule HOL.refl | rule bf\<^sub>m.dp_match_rule | rule bf\<^sub>m.crel_vs_return_ext | (subst Rel_def, rule crel_bf\<^sub>m') - | tactic \Transform_Tactic.transfer_raw_tac ctxt 1\)+ + | tactic \Transform_Tactic.transfer_raw_tac \<^context> 1\)+ done qed theorem bellman_ford_correct: "fst (run_state bellman_ford Mapping.empty) = (if has_negative_cycle then None else Some (map shortest [0..m.cmem_empty bellman_ford_correct'[unfolded bf\<^sub>m.crel_vs_def, rule_format, of Mapping.empty] unfolding bf\<^sub>m.crel_vs_def by auto end (* Wellformedness *) end (* Final Node *) end (* Bellman Ford *) subsubsection \Extracting an Executable Constant for the Imperative Implementation\ ground_function (prove_termination) bf\<^sub>h'_impl: bf\<^sub>h'.simps lemma bf\<^sub>h'_impl_def: fixes n :: nat fixes mem :: "nat ref \ nat ref \ int extended option array ref \ int extended option array ref" assumes mem_is_init: "mem = result_of (init_state (n + 1) 1 0) Heap.empty" shows "bf\<^sub>h'_impl n w t mem = bf\<^sub>h' n w t mem" proof - have "bf\<^sub>h'_impl n w t mem i j = bf\<^sub>h' n w t mem i j" for i j by (induction rule: bf\<^sub>h'.induct[OF mem_is_init]; simp add: bf\<^sub>h'.simps[OF mem_is_init]; solve_cong simp ) then show ?thesis by auto qed definition "iter_bf_heap n w t mem = iterator_defs.iter_heap (\(x, y). x \ n \ y \ n) (\(x, y). if y < n then (x, y + 1) else (x + 1, 0)) (\(x, y). bf\<^sub>h'_impl n w t mem x y)" lemma iter_bf_heap_unfold[code]: "iter_bf_heap n w t mem = (\ (i, j). (if i \ n \ j \ n then do { bf\<^sub>h'_impl n w t mem i j; iter_bf_heap n w t mem (if j < n then (i, j + 1) else (i + 1, 0)) } else Heap_Monad.return ()))" unfolding iter_bf_heap_def by (rule ext) (safe, simp add: iter_heap_unfold) definition "bf_impl n w t i j = do { mem \ (init_state (n + 1) (1::nat) (0::nat) :: (nat ref \ nat ref \ int extended option array ref \ int extended option array ref) Heap); iter_bf_heap n w t mem (0, 0); bf\<^sub>h'_impl n w t mem i j }" lemma bf_impl_correct: "bf n w t i j = result_of (bf_impl n w t i j) Heap.empty" using memoized_empty[OF HOL.refl, of n w t "(i, j)"] by (simp add: execute_bind_success[OF succes_init_state] bf_impl_def bf\<^sub>h'_impl_def iter_bf_heap_def ) subsubsection \Test Cases\ definition "G\<^sub>1_list = [[(1 :: nat,-6 :: int), (2,4), (3,5)], [(3,10)], [(3,2)], []]" definition "G\<^sub>2_list = [[(1 :: nat,-6 :: int), (2,4), (3,5)], [(3,10)], [(3,2)], [(0, -5)]]" definition "G\<^sub>3_list = [[(1 :: nat,-1 :: int), (2,2)], [(2,5), (3,4)], [(3,2), (4,3)], [(2,-2), (4,2)], []]" definition "G\<^sub>4_list = [[(1 :: nat,-1 :: int), (2,2)], [(2,5), (3,4)], [(3,2), (4,3)], [(2,-3), (4,2)], []]" definition "graph_of a i j = case_option \ (Fin o snd) (List.find (\ p. fst p = j) (a !! i))" definition "test_bf = bf_impl 3 (graph_of (IArray G\<^sub>1_list)) 3 3 0" code_reflect Test functions test_bf text \One can see a trace of the calls to the memory in the output\ ML \Test.test_bf ()\ lemma bottom_up_alt[code]: "bf n W t i j = fst (run_state (iter_bf n W t (0, 0) \ (\_. bf\<^sub>m' n W t i j)) Mapping.empty)" using bf_bottom_up by auto definition "bf_ia n W t i j = (let W' = graph_of (IArray W) in fst (run_state (iter_bf n W' t (i, j) \ (\_. bf\<^sub>m' n W' t i j)) Mapping.empty) )" \ \Component tests.\ lemma "fst (run_state (bf\<^sub>m' 3 (graph_of (IArray G\<^sub>1_list)) 3 3 0) Mapping.empty) = 4" "bf 3 (graph_of (IArray G\<^sub>1_list)) 3 3 0 = 4" by eval+ \ \Regular test cases.\ lemma "fst (run_state (bellman_ford 3 (graph_of (IArray G\<^sub>1_list)) 3) Mapping.empty) = Some [4, 10, 2, 0]" "fst (run_state (bellman_ford 4 (graph_of (IArray G\<^sub>3_list)) 4) Mapping.empty) = Some [4, 5, 3, 1, 0]" by eval+ \ \Test detection of negative cycles.\ lemma "fst (run_state (bellman_ford 3 (graph_of (IArray G\<^sub>2_list)) 3) Mapping.empty) = None" "fst (run_state (bellman_ford 4 (graph_of (IArray G\<^sub>4_list)) 4) Mapping.empty) = None" by eval+ end (* Theory *) \ No newline at end of file diff --git a/thys/Monad_Memo_DP/state_monad/DP_CRelVS_Ext.thy b/thys/Monad_Memo_DP/state_monad/DP_CRelVS_Ext.thy --- a/thys/Monad_Memo_DP/state_monad/DP_CRelVS_Ext.thy +++ b/thys/Monad_Memo_DP/state_monad/DP_CRelVS_Ext.thy @@ -1,241 +1,243 @@ +(* FIXME dead code!? *) + theory DP_CRelVS_Ext - imports "transform/Transform_Cmd" + imports "Transform_Cmd" begin notation fun_app_lifted (infixl "." 999) notation Transfer.Rel ("Rel") thm map_cong lemma mapT_cong: assumes "xs = ys" "\x. x\set ys \ f x = g x" shows "map\<^sub>T . \f\ . \xs\ = map\<^sub>T . \g\ . \ys\" unfolding map\<^sub>T_def unfolding assms(1) using assms(2) by (induction ys) (auto simp: return_app_return) thm fold_cong lemma foldT_cong: assumes "xs = ys" "\x. x\set ys \ f x = g x" shows "fold\<^sub>T . \f\ . \xs\ = fold\<^sub>T . \g\ . \ys\" unfolding fold\<^sub>T_def unfolding assms(1) using assms(2) by (induction ys) (auto simp: return_app_return) lemma abs_unit_cong: (* for lazy checkmem *) assumes "x = y" shows "(\_::unit. x) = (\_. y)" using assms .. lemma App_cong: assumes "f x = g y" shows "App f x = App g y" unfolding App_def using assms . lemmas [fundef_cong] = return_app_return_cong ifT_cong mapT_cong foldT_cong abs_unit_cong App_cong context dp_consistency begin context includes lifting_syntax begin named_theorems dp_match_rule lemma refl2: "is_equality R \ Rel R x x" unfolding is_equality_def Rel_def by simp lemma rel_fun2: assumes "is_equality R0" "\x. Rel R1 (f x) (g x)" shows "Rel (rel_fun R0 R1) f g" using assms unfolding is_equality_def Rel_def by auto thm if_cong lemma if\<^sub>T_cong2: assumes "Rel (=) b c" "c \ Rel (crel_vs R) x x\<^sub>T" "\c \ Rel (crel_vs R) y y\<^sub>T" shows "Rel (crel_vs R) (if (Wrap b) then x else y) (if\<^sub>T \c\ x\<^sub>T y\<^sub>T)" using assms unfolding if\<^sub>T_def left_identity Rel_def Wrap_def by (auto split: if_split) lemma if\<^sub>T_cong2': assumes "c \ Rel (crel_vs R) x x\<^sub>T" "\c \ Rel (crel_vs R) y y\<^sub>T" shows "Rel (crel_vs R) (if Wrap c then x else y) (if\<^sub>T \c\ x\<^sub>T y\<^sub>T)" apply (rule if\<^sub>T_cong2[OF _ assms]) apply (auto simp: Rel_def) done (* thm map_cong lemma map1T_cong: assumes "xs = ys" "\x. x\set ys \ f\<^sub>T . \x\ = g\<^sub>T . \x\" shows "map1\<^sub>T . f\<^sub>T . \xs\ = map1\<^sub>T . g\<^sub>T . \ys\" unfolding map1\<^sub>T_def unfolding assms(1) using assms(2) apply (induction ys) apply (auto simp: return_app_return) subgoal apply (rule state.expand) apply (rule ext) unfolding fun_app_lifted_def unfolding bind_def return_def apply (auto split: prod.splits) *) lemma map\<^sub>T_transfer2: "Rel (crel_vs ((R0 ===>\<^sub>T R1) ===>\<^sub>T list_all2 R0 ===>\<^sub>T list_all2 R1)) map map\<^sub>T" unfolding Rel_def by (fact map\<^sub>T_transfer) lemma map\<^sub>T_cong2: assumes "is_equality R" "Rel R xs ys" "\x. x\set ys \ Rel (crel_vs S) (f x) (f\<^sub>T' x)" shows "Rel (crel_vs (list_all2 S)) (App (App map (Wrap f)) (Wrap xs)) (map\<^sub>T . \f\<^sub>T'\ . \ys\)" unfolding map\<^sub>T_def unfolding return_app_return unfolding assms(2)[unfolded Rel_def assms(1)[unfolded is_equality_def]] using assms(3) unfolding Rel_def Wrap_def App_def apply (induction ys) apply (tactic \Transform_Tactic.unfold_dp_defs_tac @{context} true true 1\) subgoal premises by transfer_prover subgoal premises prems for a ys apply (tactic \Transform_Tactic.unfold_dp_defs_tac @{context} true true 1\) apply (unfold return_app_return Wrap_App_Wrap) supply [transfer_rule] = prems(2)[OF list.set_intros(1)] prems(1)[OF prems(2)[OF list.set_intros(2)], simplified] by transfer_prover done lemma map\<^sub>T_cong2': assumes "\x. x\set xs \ Rel (crel_vs S) (f x) (f\<^sub>T' x)" shows "Rel (crel_vs (list_all2 S)) (map f xs) (map\<^sub>T . \f\<^sub>T'\ . \xs\)" apply (rule map\<^sub>T_cong2[OF is_equality_eq _ assms, unfolded Wrap_def App_def]) apply (auto simp: Rel_def) done lemma fold\<^sub>T_transfer2: "Rel (crel_vs ((R0 ===>\<^sub>T R1 ===>\<^sub>T R1) ===>\<^sub>T list_all2 R0 ===>\<^sub>T R1 ===>\<^sub>T R1)) fold fold\<^sub>T" unfolding Rel_def by (fact fold\<^sub>T_transfer) lemma fold\<^sub>T_cong2: assumes "is_equality R" "Rel R xs ys" "\x. x\set ys \ Rel (crel_vs (S ===> crel_vs S)) (f x) (f\<^sub>T' x)" shows "Rel (crel_vs (S ===> crel_vs S)) (fold f xs) (fold\<^sub>T . \f\<^sub>T'\ . \ys\)" unfolding fold\<^sub>T_def unfolding return_app_return unfolding assms(2)[unfolded Rel_def assms(1)[unfolded is_equality_def]] using assms(3) unfolding Rel_def apply (induction ys) apply (tactic \Transform_Tactic.unfold_dp_defs_tac @{context} true true 1\) subgoal premises by transfer_prover subgoal premises prems for a ys apply (tactic \Transform_Tactic.unfold_dp_defs_tac @{context} true true 1\) apply (unfold return_app_return Wrap_App_Wrap) supply [transfer_rule] = prems(2)[OF list.set_intros(1)] prems(1)[OF prems(2)[OF list.set_intros(2)], simplified] by transfer_prover done thm BNF_Fixpoint_Base.case_prod_transfer lemma prod_case_transfer2: assumes "Rel (R0 ===> R1 ===> crel_vs S) f g" shows "Rel (rel_prod R0 R1 ===> crel_vs S) (case_prod f) (case_prod g)" supply [transfer_rule] = assms[unfolded Rel_def] unfolding Rel_def by transfer_prover lemma prod_case_cong2: assumes "is_equality R" "Rel R x y" "\u v. y = (u, v) \ Rel (crel_vs S) (f u v) (g u v)" shows "Rel (crel_vs S) (case_prod f x) (\case_prod g\ . \y\)" unfolding return_app_return using assms by (auto simp: is_equality_def Rel_def split: prod.split) thm option.case_transfer lemma option_case_transfer2: assumes "Rel (crel_vs S) f0 g0" "Rel (R ===> crel_vs S) f1 g1" shows "Rel (rel_option R ===> crel_vs S) (case_option f0 f1) (case_option g0 g1)" supply [transfer_rule] = assms[unfolded Rel_def] unfolding Rel_def by transfer_prover thm option.case_cong lemma option_case_cong2: assumes "is_equality R" "Rel R x y" "y = None \ Rel (crel_vs S) f0 g0" "\x2. y = Some x2 \ Rel (crel_vs S) (f1 x2) (g1 x2)" shows "Rel (crel_vs S) (case_option f0 f1 x) (\case_option g0 g1\ . \y\)" unfolding return_app_return using assms by (auto simp: is_equality_def Rel_def split: option.split) thm list.case_transfer lemma list_case_transfer2: assumes "Rel (crel_vs S) f0 g0" "Rel (R ===> list_all2 R ===> crel_vs S) f1 g1" shows "Rel (list_all2 R ===> crel_vs S) (case_list f0 f1) (case_list g0 g1)" supply [transfer_rule] = assms[unfolded Rel_def] unfolding Rel_def by transfer_prover thm list.case_cong lemma list_case_cong2: assumes "is_equality R" "Rel R x y" "x = Nil \ Rel (crel_vs S) f0 g0" "\x2 x2s. x = x2#x2s \ Rel (crel_vs S) (f1 x2 x2s) (g1 x2 x2s)" shows "Rel (crel_vs S) (case_list f0 f1 x) (\case_list g0 g1\ . \y\)" unfolding return_app_return using assms by (auto simp: is_equality_def Rel_def split: list.split) lemmas [dp_match_rule] = map\<^sub>T_cong2 fold\<^sub>T_cong2 if\<^sub>T_cong2 lemmas [dp_match_rule] = prod_case_cong2 prod_case_transfer2 option_case_cong2 option_case_transfer2 list_case_cong2 list_case_transfer2 lemmas [dp_match_rule] = crel_vs_return crel_vs_fun_app rel_fun2 refl2 lemmas [dp_match_rule] = map\<^sub>T_transfer2 fold\<^sub>T_transfer2 thm dp_match_rule end (* lifting_syntax *) end (* dp_consistency *) no_notation fun_app_lifted (infixl "." 999) no_notation Transfer.Rel ("Rel") end (* theory *) diff --git a/thys/Monad_Memo_DP/transform/Transform.ML b/thys/Monad_Memo_DP/transform/Transform.ML --- a/thys/Monad_Memo_DP/transform/Transform.ML +++ b/thys/Monad_Memo_DP/transform/Transform.ML @@ -1,238 +1,239 @@ -structure Transform_DP = struct +signature TRANSFORM_DP = +sig + val dp_fun_part1_cmd: + (binding * string) * ((bool * (xstring * Position.T)) * (string * string) list) option + -> local_theory -> local_theory + val dp_fun_part2_cmd: string * (Facts.ref * Token.src list) list -> local_theory -> local_theory + val dp_correct_cmd: local_theory -> Proof.state +end + +structure Transform_DP : TRANSFORM_DP = +struct fun dp_interpretation standard_proof locale_name instance qualifier dp_term lthy = lthy |> Interpretation.isar_interpretation ([(locale_name, ((qualifier, true), (Expression.Named (("dp", dp_term) :: instance), [])))], []) |> (if standard_proof then Proof.global_default_proof else Proof.global_immediate_proof) -fun prep_params (((scope, tm_str), def_thms_opt), mem_locale_opt) lthy = +fun prep_params (((scope, tm_str), def_thms_opt), mem_locale_opt) ctxt = let - val tm = Syntax.read_term lthy tm_str + val tm = Syntax.read_term ctxt tm_str val scope' = (Binding.is_empty scope? Binding.map_name (fn _ => Transform_Misc.term_name tm ^ "\<^sub>T")) scope - val def_thms_opt' = Option.map (Attrib.eval_thms lthy) def_thms_opt - val mem_locale_opt' = Option.map (Locale.check (Proof_Context.theory_of lthy)) mem_locale_opt + val def_thms_opt' = Option.map (Attrib.eval_thms ctxt) def_thms_opt + val mem_locale_opt' = Option.map (Locale.check (Proof_Context.theory_of ctxt)) mem_locale_opt in (scope', tm, def_thms_opt', mem_locale_opt') end (* fun dp_interpretation_cmd args lthy = let val (scope, tm, _, mem_locale_opt) = prep_params args lthy val scope_name = Binding.name_of scope in case mem_locale_opt of NONE => lthy | SOME x => dp_interpretation x scope_name (Transform_Misc.uncurry tm) lthy end *) fun do_monadify heap_name scope tm mem_locale_opt def_thms_opt lthy = let val monad_consts = Transform_Const.get_monad_const heap_name val scope_name = Binding.name_of scope val memoizer_opt = if is_none mem_locale_opt then NONE else SOME (Transform_Misc.locale_term lthy scope_name "checkmem") val old_info_opt = Function_Common.import_function_data tm lthy val old_defs_opt = [ K def_thms_opt, K (Option.mapPartial #simps old_info_opt) ] |> Library.get_first (fn x => x ()) val old_defs = case old_defs_opt of SOME defs => defs | NONE => raise TERM("no definition", [tm]) - val ((_, old_defs_imported), _) = Variable.import true old_defs lthy + val ((_, old_defs_imported), ctxt') = lthy + |> fold Variable.declare_thm old_defs + |> Variable.import true old_defs (* val new_bind = Binding.suffix_name "\<^sub>T'" scope val new_bindT = Binding.suffix_name "\<^sub>T" scope *) val new_bind = Binding.suffix_name "'" scope val new_bindT = scope fun dest_def (def, def_imported) = let - val def_imported_meta = def_imported |> Local_Defs.meta_rewrite_rule lthy + val def_imported_meta = def_imported |> Local_Defs.meta_rewrite_rule ctxt' val eqs = def_imported_meta |> Thm.full_prop_of val (head, _) = Logic.dest_equals eqs |> fst |> Transform_Misc.behead tm (*val _ = if Term.aconv_untyped (head, tm) then () else raise THM("invalid definition", 0, [def])*) - val Abs t = Term.lambda_name (Binding.name_of new_bind, head) eqs - val (t_name, eqs') = Term.dest_abs t + val t = Term.lambda_name (Binding.name_of new_bind, head) eqs + val ((t_name, _), eqs') = Term.dest_abs_global t val _ = @{assert} (t_name = Binding.name_of new_bind) (*val eqs' = Term.subst_atomic [(head, Free (Binding.name_of new_bind, fastype_of head))] eqs*) - val (rhs_conv, eqsT, n_args) = Transform_Term.lift_equation monad_consts lthy (Logic.dest_equals eqs') memoizer_opt - val def_meta' = def |> Local_Defs.meta_rewrite_rule lthy |> Conv.fconv_rule (Conv.arg_conv (rhs_conv lthy)) + val (rhs_conv, eqsT, n_args) = Transform_Term.lift_equation monad_consts ctxt' (Logic.dest_equals eqs') memoizer_opt + val def_meta' = def |> Local_Defs.meta_rewrite_rule ctxt' |> Conv.fconv_rule (Conv.arg_conv (rhs_conv ctxt')) val def_meta_simped = def_meta' |> Conv.fconv_rule ( - repeat_sweep_conv (K Transform_Term.rewrite_pureapp_beta_conv) lthy + Transform_Term.repeat_sweep_conv (K Transform_Term.rewrite_pureapp_beta_conv) ctxt' ) (* val eqsT_simped = eqsT - |> Syntax.check_term lthy - |> Thm.cterm_of lthy - |> repeat_sweep_conv (K Transform_Term.rewrite_app_beta_conv) lthy + |> Syntax.check_term ctxt' + |> Thm.cterm_of ctxt' + |> Transform_Term.repeat_sweep_conv (K Transform_Term.rewrite_app_beta_conv) ctxt' |> Thm.full_prop_of |> Logic.dest_equals |> snd *) in ((def_meta_simped, eqsT), n_args) end val ((old_defs', new_defs_raw), n_args) = map dest_def (old_defs ~~ old_defs_imported) |> split_list |>> split_list ||> Transform_Misc.the_element val new_defs = Syntax.check_props lthy new_defs_raw |> map (fn eqsT => eqsT |> Thm.cterm_of lthy - |> repeat_sweep_conv (K (#rewrite_app_beta_conv monad_consts)) lthy + |> Transform_Term.repeat_sweep_conv (K (#rewrite_app_beta_conv monad_consts)) lthy |> Thm.full_prop_of |> Logic.dest_equals |> snd) (*val _ = map (Pretty.writeln o Syntax.pretty_term @{context} o Thm.full_prop_of) old_defs'*) (*val (new_defs, lthy) = Variable.importT_terms new_defs lthy*) - val (new_info, lthy) = Transform_Misc.add_function new_bind new_defs lthy + val (new_info, lthy1) = Transform_Misc.add_function new_bind new_defs lthy val replay_tac = case old_info_opt of NONE => no_tac - | SOME info => Transform_Tactic.totality_replay_tac info new_info lthy + | SOME info => Transform_Tactic.totality_replay_tac info new_info lthy1 val totality_tac = replay_tac - ORELSE (Function_Common.termination_prover_tac false lthy + ORELSE (Function_Common.termination_prover_tac false lthy1 THEN Transform_Tactic.my_print_tac "termination by default prover") - val (new_info, lthy) = Function.prove_termination NONE totality_tac lthy + val (new_info, lthy2) = Function.prove_termination NONE totality_tac lthy1 val new_def' = new_info |> #simps |> the val head' = new_info |> #fs |> the_single - val headT = Transform_Term.wrap_head monad_consts head' n_args |> Syntax.check_term lthy - val ((headTC, (_, new_defT)), lthy) = Local_Theory.define ((new_bindT, NoSyn), ((Thm.def_binding new_bindT,[]), headT)) lthy + val headT = Transform_Term.wrap_head monad_consts head' n_args |> Syntax.check_term lthy2 + val ((headTC, (_, new_defT)), lthy) = Local_Theory.define ((new_bindT, NoSyn), ((Thm.def_binding new_bindT,[]), headT)) lthy2 - val lthy = Transform_Data.commit_dp_info (#monad_name monad_consts) ({ + val lthy3 = Transform_Data.commit_dp_info (#monad_name monad_consts) ({ old_head = tm, new_head' = head', new_headT = headTC, old_defs = old_defs', new_def' = new_def', new_defT = new_defT }) lthy - val _ = Proof_Display.print_consts true (Position.thread_data ()) lthy (K false) [ + val _ = Proof_Display.print_consts true (Position.thread_data ()) lthy3 (K false) [ (Binding.name_of new_bind, Term.type_of head'), (Binding.name_of new_bindT, Term.type_of headTC)] - in lthy end + in lthy3 end fun gen_dp_monadify prep_term args lthy = let val (scope, tm, def_thms_opt, mem_locale_opt) = prep_params args lthy (* val memoizer_opt = memoizer_scope_opt |> Option.map (fn memoizer_scope => Syntax.read_term lthy (Long_Name.qualify memoizer_scope Transform_Const.checkmemVN)) val _ = memoizer_opt |> Option.map (fn memoizer => if Term.aconv_untyped (head_of memoizer, @{term mem_defs.checkmem}) then () else raise TERM("invalid memoizer", [the memoizer_opt])) *) in do_monadify "state" scope tm mem_locale_opt def_thms_opt lthy end val dp_monadify_cmd = gen_dp_monadify Syntax.read_term fun dp_fun_part1_cmd ((scope, tm_str), (mem_locale_instance_opt)) lthy = let val scope_name = Binding.name_of scope val tm = Syntax.read_term lthy tm_str val _ = if is_Free tm then warning ("Free term: " ^ (Syntax.pretty_term lthy tm |> Pretty.string_of)) else () val mem_locale_opt' = Option.map (Locale.check (Proof_Context.theory_of lthy) o (snd o fst)) mem_locale_instance_opt - val lthy_f = case mem_locale_instance_opt of - NONE => I + val lthy1 = case mem_locale_instance_opt of + NONE => lthy | SOME ((standard_proof, locale_name), instance) => let val locale_name = Locale.check (Proof_Context.theory_of lthy) locale_name val instance = map (apsnd (Syntax.read_term lthy)) instance in - dp_interpretation standard_proof locale_name instance scope_name (Transform_Misc.uncurry tm) + dp_interpretation standard_proof locale_name instance scope_name (Transform_Misc.uncurry tm) lthy end - - val lthy = lthy_f lthy - val lthy = Transform_Data.add_tmp_cmd_info (Binding.reset_pos scope, tm, mem_locale_opt') lthy - in - lthy + Transform_Data.add_tmp_cmd_info (Binding.reset_pos scope, tm, mem_locale_opt') lthy1 end fun dp_fun_part2_cmd (heap_name, def_thms_str) lthy = let val {scope, head=tm, locale=locale_opt, dp_info=dp_info_opt} = Transform_Data.get_last_cmd_info lthy val _ = if is_none dp_info_opt then () else raise TERM("already monadified", [tm]) val def_thms = Attrib.eval_thms lthy def_thms_str - val heap_typ = Syntax.read_typ - val lthy = do_monadify heap_name scope tm locale_opt (SOME def_thms) lthy in - lthy + do_monadify heap_name scope tm locale_opt (SOME def_thms) lthy end fun dp_correct_cmd lthy = let val {scope, head=tm, locale=locale_opt, dp_info=dp_info_opt} = Transform_Data.get_last_cmd_info lthy val dp_info = case dp_info_opt of SOME x => x | NONE => raise TERM("not yet monadified", [tm]) val _ = if is_some locale_opt then () else raise TERM("not interpreted yet", [tm]) val scope_name = Binding.name_of scope val consistentDP = Transform_Misc.locale_term lthy scope_name "consistentDP" val dpT' = #new_head' dp_info val dpT'_curried = dpT' |> Transform_Misc.uncurry val goal_pat = consistentDP $ dpT'_curried val goal_prop = Syntax.check_term lthy (HOLogic.mk_Trueprop goal_pat) val tuple_pat = type_of dpT' |> strip_type |> fst |> length |> Name.invent_list [] "a" |> map (fn s => Var ((s, 0), TVar ((s, 0), @{sort type}))) |> HOLogic.mk_tuple |> Thm.cterm_of lthy val memoized_thm_opt = Transform_Misc.locale_thms lthy scope_name "memoized" |> the_single |> SOME handle ERROR msg => (warning msg; NONE) val memoized_thm'_opt = memoized_thm_opt |> Option.map (Drule.infer_instantiate' lthy [NONE, SOME tuple_pat]) fun display_thms thm_binds ctxt = - Proof_Display.print_results true (Position.thread_data ()) ctxt((Thm.theoremK, ""), [thm_binds]) + Proof_Display.print_results true (Position.thread_data ()) ctxt ((Thm.theoremK, ""), [thm_binds]) val crel_thm_name = "crel" val memoized_thm_name = "memoized_correct" - fun afterqed thmss ctxt = + fun afterqed result lthy1 = let - val [[crel_thm]] = thmss - - val (crel_thm_binds, ctxt) = Local_Theory.note ( - (Binding.qualify_name true scope crel_thm_name, []), - [crel_thm] - ) ctxt - - val _ = display_thms crel_thm_binds ctxt + val [[crel_thm]] = result - val ctxt = case memoized_thm'_opt of NONE => ctxt | SOME memoized_thm' => let - val (memoized_thm_binds, ctxt) = Local_Theory.note ( - (Binding.qualify_name true scope memoized_thm_name, []), - [(crel_thm RS memoized_thm') |> Local_Defs.unfold lthy @{thms prod.case}] - ) ctxt + val (crel_thm_binds, lthy2) = lthy1 + |> Local_Theory.note ((Binding.qualify_name true scope crel_thm_name, []), [crel_thm]) - val _ = display_thms memoized_thm_binds ctxt - in ctxt end + val _ = display_thms crel_thm_binds lthy2 in - ctxt + case memoized_thm'_opt of + NONE => lthy2 + | SOME memoized_thm' => + let + val (memoized_thm_binds, lthy3) = lthy2 + |> Local_Theory.note + ((Binding.qualify_name true scope memoized_thm_name, []), + [(crel_thm RS memoized_thm') |> Local_Defs.unfold lthy @{thms prod.case}]) + val _ = display_thms memoized_thm_binds lthy3 + in lthy3 end end - - val goal = Proof.theorem NONE afterqed [[(goal_prop, [])]] lthy in - goal + Proof.theorem NONE afterqed [[(goal_prop, [])]] lthy end - end diff --git a/thys/Monad_Memo_DP/transform/Transform_Cmd.thy b/thys/Monad_Memo_DP/transform/Transform_Cmd.thy --- a/thys/Monad_Memo_DP/transform/Transform_Cmd.thy +++ b/thys/Monad_Memo_DP/transform/Transform_Cmd.thy @@ -1,75 +1,76 @@ subsection \Tool Setup\ theory Transform_Cmd imports "../Pure_Monad" "../state_monad/DP_CRelVS" "../heap_monad/DP_CRelVH" keywords "memoize_fun" :: thy_decl and "monadifies" :: thy_decl and "memoize_correct" :: thy_goal and "with_memory" :: quasi_command and "default_proof" :: quasi_command begin ML_file \../transform/Transform_Misc.ML\ ML_file \../transform/Transform_Const.ML\ ML_file \../transform/Transform_Data.ML\ ML_file \../transform/Transform_Tactic.ML\ ML_file \../transform/Transform_Term.ML\ ML_file \../transform/Transform.ML\ ML_file \../transform/Transform_Parser.ML\ ML \ val _ = - Outer_Syntax.local_theory @{command_keyword memoize_fun} "whatever" + Outer_Syntax.local_theory @{command_keyword memoize_fun} "" (Transform_Parser.dp_fun_part1_parser >> Transform_DP.dp_fun_part1_cmd) val _ = - Outer_Syntax.local_theory @{command_keyword monadifies} "whatever" + Outer_Syntax.local_theory @{command_keyword monadifies} "" (Transform_Parser.dp_fun_part2_parser >> Transform_DP.dp_fun_part2_cmd) -\ -ML \ val _ = - Outer_Syntax.local_theory_to_proof @{command_keyword memoize_correct} "whatever" + Outer_Syntax.local_theory_to_proof @{command_keyword memoize_correct} "" (Scan.succeed Transform_DP.dp_correct_cmd) \ method_setup memoize_prover = \ -Scan.succeed (fn ctxt => (SIMPLE_METHOD' ( - Transform_Data.get_last_cmd_info ctxt - |> Transform_Tactic.solve_consistentDP_tac ctxt))) + Scan.succeed (fn ctxt => SIMPLE_METHOD' ( + Transform_Data.get_last_cmd_info ctxt + |> Transform_Tactic.solve_consistentDP_tac ctxt)) \ method_setup memoize_prover_init = \ -Scan.succeed (fn ctxt => (SIMPLE_METHOD' ( - Transform_Data.get_last_cmd_info ctxt - |> Transform_Tactic.prepare_consistentDP_tac ctxt))) + Scan.succeed (fn ctxt => SIMPLE_METHOD' ( + Transform_Data.get_last_cmd_info ctxt + |> Transform_Tactic.prepare_consistentDP_tac ctxt)) \ method_setup memoize_prover_case_init = \ -Scan.succeed (fn ctxt => (SIMPLE_METHOD' ( - Transform_Data.get_last_cmd_info ctxt - |> Transform_Tactic.prepare_case_tac ctxt))) + Scan.succeed (fn ctxt => SIMPLE_METHOD' ( + Transform_Data.get_last_cmd_info ctxt + |> Transform_Tactic.prepare_case_tac ctxt)) \ method_setup memoize_prover_match_step = \ -Scan.succeed (fn ctxt => (SIMPLE_METHOD' ( - Transform_Data.get_last_cmd_info ctxt - |> Transform_Tactic.step_tac ctxt))) + Scan.succeed (fn ctxt => SIMPLE_METHOD' ( + Transform_Data.get_last_cmd_info ctxt + |> Transform_Tactic.step_tac ctxt)) \ method_setup memoize_unfold_defs = \ -Scan.option (Scan.lift (Args.parens Args.name) -- Args.term) >> (fn tm_opt => fn ctxt => (SIMPLE_METHOD' ( - Transform_Data.get_or_last_cmd_info ctxt tm_opt - |> Transform_Tactic.dp_unfold_defs_tac ctxt))) + Scan.option (Scan.lift (Args.parens Args.name) -- Args.term) >> + (fn tm_opt => fn ctxt => SIMPLE_METHOD' + (Transform_Data.get_or_last_cmd_info ctxt tm_opt + |> Transform_Tactic.dp_unfold_defs_tac ctxt)) \ method_setup memoize_combinator_init = \ -Scan.option (Scan.lift (Args.parens Args.name) -- Args.term) >> (fn tm_opt => fn ctxt => (SIMPLE_METHOD' ( - Transform_Data.get_or_last_cmd_info ctxt tm_opt - |> Transform_Tactic.prepare_combinator_tac ctxt))) + Scan.option (Scan.lift (Args.parens Args.name) -- Args.term) >> + (fn tm_opt => fn ctxt => SIMPLE_METHOD' + (Transform_Data.get_or_last_cmd_info ctxt tm_opt + |> Transform_Tactic.prepare_combinator_tac ctxt)) \ -end (* theory *) + +end diff --git a/thys/Monad_Memo_DP/transform/Transform_Const.ML b/thys/Monad_Memo_DP/transform/Transform_Const.ML --- a/thys/Monad_Memo_DP/transform/Transform_Const.ML +++ b/thys/Monad_Memo_DP/transform/Transform_Const.ML @@ -1,98 +1,112 @@ -structure Transform_Const = struct +signature TRANSFORM_CONST = +sig + type MONAD_CONSTS = { + monad_name: string, + mk_stateT: typ -> typ, + return: term -> term, + app: (term * term) -> term, + if_termN: string, + checkmemVN: string, + rewrite_app_beta_conv: conv + } + val get_monad_const: string -> MONAD_CONSTS +end + +structure Transform_Const : TRANSFORM_CONST = +struct val pureappN = @{const_name Pure_Monad.App} fun pureapp tm = Const (pureappN, dummyT) $ tm type MONAD_CONSTS = { monad_name: string, mk_stateT: typ -> typ, return: term -> term, app: (term * term) -> term, if_termN: string, checkmemVN: string, rewrite_app_beta_conv: conv } - + val state_monad: MONAD_CONSTS = let val memT = TFree ("MemoryType", @{sort type}) val memT = dummyT - + fun mk_stateT tp = Type (@{type_name State_Monad.state}, [memT, tp]) - + val returnN = @{const_name State_Monad.return} fun return tm = Const (returnN, dummyT --> mk_stateT dummyT) $ tm - + val appN = @{const_name State_Monad_Ext.fun_app_lifted} fun app (tm0, tm1) = Const (appN, dummyT) $ tm0 $ tm1 - + fun checkmem'C ctxt = Transform_Misc.get_const_pat ctxt "checkmem'" fun checkmem' ctxt param body = checkmem'C ctxt $ param $ body - + val checkmemVN = "checkmem" val checkmemC = @{const_name "state_mem_defs.checkmem"} fun rewrite_app_beta_conv ctm = case Thm.term_of ctm of Const (@{const_name State_Monad_Ext.fun_app_lifted}, _) $ (Const (@{const_name State_Monad.return}, _) $ Abs _) $ (Const (@{const_name State_Monad.return}, _) $ _) => Conv.rewr_conv @{thm State_Monad_Ext.return_app_return_meta} ctm | _ => Conv.no_conv ctm in { monad_name = "state", mk_stateT = mk_stateT, return = return, app = app, if_termN = @{const_name State_Monad_Ext.if\<^sub>T}, checkmemVN = checkmemVN, rewrite_app_beta_conv = rewrite_app_beta_conv } end val heap_monad: MONAD_CONSTS = let fun mk_stateT tp = Type (@{type_name Heap_Monad.Heap}, [tp]) - + val returnN = @{const_name Heap_Monad.return} fun return tm = Const (returnN, dummyT --> mk_stateT dummyT) $ tm - + val appN = @{const_name Heap_Monad_Ext.fun_app_lifted} fun app (tm0, tm1) = Const (appN, dummyT) $ tm0 $ tm1 - + fun checkmem'C ctxt = Transform_Misc.get_const_pat ctxt "checkmem'" fun checkmem' ctxt param body = checkmem'C ctxt $ param $ body - + val checkmemVN = "checkmem" val checkmemC = @{const_name "heap_mem_defs.checkmem"} fun rewrite_app_beta_conv ctm = case Thm.term_of ctm of Const (@{const_name Heap_Monad_Ext.fun_app_lifted}, _) $ (Const (@{const_name Heap_Monad.return}, _) $ Abs _) $ (Const (@{const_name Heap_Monad.return}, _) $ _) => Conv.rewr_conv @{thm Heap_Monad_Ext.return_app_return_meta} ctm | _ => Conv.no_conv ctm in { monad_name = "heap", mk_stateT = mk_stateT, return = return, app = app, if_termN = @{const_name Heap_Monad_Ext.if\<^sub>T}, checkmemVN = checkmemVN, rewrite_app_beta_conv = rewrite_app_beta_conv } end val monad_consts_dict = [ ("state", state_monad), ("heap", heap_monad) ] fun get_monad_const name = case AList.lookup op= monad_consts_dict name of SOME consts => consts | NONE => error("unrecognized monad: " ^ name ^ " , choices: " ^ commas (map fst monad_consts_dict)); end - diff --git a/thys/Monad_Memo_DP/transform/Transform_Data.ML b/thys/Monad_Memo_DP/transform/Transform_Data.ML --- a/thys/Monad_Memo_DP/transform/Transform_Data.ML +++ b/thys/Monad_Memo_DP/transform/Transform_Data.ML @@ -1,142 +1,167 @@ -structure Transform_Data = struct +signature TRANSFORM_DATA = +sig + type dp_info = { + old_head: term, + new_head': term, + new_headT: term, + + old_defs: thm list, + new_defT: thm, + new_def': thm list + } + type cmd_info = { + scope: binding, + head: term, + locale: string option, + dp_info: dp_info option + } + val get_dp_info: string -> Proof.context -> term -> dp_info option + val get_last_cmd_info: Proof.context -> cmd_info + val commit_dp_info: string -> dp_info -> local_theory -> local_theory + val add_tmp_cmd_info: binding * term * string option -> local_theory -> local_theory + val get_or_last_cmd_info: Proof.context -> (string * term) option -> cmd_info +end + +structure Transform_Data : TRANSFORM_DATA = +struct type dp_info = { old_head: term, new_head': term, new_headT: term, old_defs: thm list, new_defT: thm, new_def': thm list } type cmd_info = { scope: binding, head: term, locale: string option, dp_info: dp_info option } fun map_cmd_info f0 f1 f2 f3 {scope, head, locale, dp_info} = {scope = f0 scope, head = f1 head, locale = f2 locale, dp_info = f3 dp_info} fun map_cmd_dp_info f = map_cmd_info I I I f structure Data = Generic_Data ( type T = { monadified_terms: (string * cmd_info Item_Net.T) list, last_cmd_info: cmd_info option } val empty = { monadified_terms = ["state", "heap"] ~~ replicate 2 (Item_Net.init (op aconv o apply2 #head) (single o #head)), last_cmd_info = NONE } val extend = I fun merge ( {monadified_terms = m0, ...}, {monadified_terms = m1, ...} ) = let val keys0 = map fst m0 val keys1 = map fst m1 val _ = @{assert} (keys0 = keys1) val vals = map Item_Net.merge (map snd m0 ~~ map snd m1) val ms = keys0 ~~ vals in {monadified_terms = ms, last_cmd_info = NONE} end ) fun transform_dp_info phi {old_head, new_head', new_headT, old_defs, new_defT, new_def'} = { old_head = Morphism.term phi old_head, new_head' = Morphism.term phi new_head', new_headT = Morphism.term phi new_headT, - + old_defs = Morphism.fact phi old_defs, new_def' = Morphism.fact phi new_def', new_defT = Morphism.thm phi new_defT } -fun get_monadified_terms_generic monad_name ctxt = - Data.get ctxt +fun get_monadified_terms_generic monad_name context = + Data.get context |> #monadified_terms |> (fn l => AList.lookup op= l monad_name) |> the -fun get_monadified_terms monad_name lthy = - get_monadified_terms_generic monad_name (Context.Proof lthy) +fun get_monadified_terms monad_name ctxt = + get_monadified_terms_generic monad_name (Context.Proof ctxt) fun map_data f0 f1 = Data.map (fn {monadified_terms, last_cmd_info} => {monadified_terms = f0 monadified_terms, last_cmd_info = f1 last_cmd_info}) fun map_monadified_terms f = map_data f I fun map_last_cmd_info f = map_data I f -fun put_monadified_terms_generic monad_name new_terms ctxt = - ctxt |> map_monadified_terms (AList.update op= (monad_name, new_terms)) +fun put_monadified_terms_generic monad_name new_terms context = + context |> map_monadified_terms (AList.update op= (monad_name, new_terms)) -fun map_monadified_terms_generic monad_name f ctxt = - ctxt |> map_monadified_terms (AList.map_entry op= monad_name f) +fun map_monadified_terms_generic monad_name f context = + context |> map_monadified_terms (AList.map_entry op= monad_name f) -fun put_last_cmd_info cmd_info_opt ctxt = - map_last_cmd_info (K cmd_info_opt) ctxt +fun put_last_cmd_info cmd_info_opt context = + map_last_cmd_info (K cmd_info_opt) context -fun get_cmd_info monad_name lthy tm = - get_monadified_terms monad_name lthy +fun get_cmd_info monad_name ctxt tm = + get_monadified_terms monad_name ctxt |> (fn net => Item_Net.retrieve net tm) -fun get_dp_info monad_name lthy tm = - get_cmd_info monad_name lthy tm +fun get_dp_info monad_name ctxt tm = + get_cmd_info monad_name ctxt tm |> (fn result => case result of {dp_info = SOME dp_info', ...} :: _ => SOME dp_info' | _ => NONE) -fun get_last_cmd_info_generic ctxt = - Data.get ctxt +fun get_last_cmd_info_generic context = + Data.get context |> #last_cmd_info |> the -fun get_last_cmd_info lthy = - get_last_cmd_info_generic (Context.Proof lthy) +fun get_last_cmd_info ctxt = + get_last_cmd_info_generic (Context.Proof ctxt) fun commit_dp_info monad_name dp_info = Local_Theory.declaration {pervasive = false, syntax = false} - (fn phi => fn ctxt => + (fn phi => fn context => let - val old_cmd_info = get_last_cmd_info_generic ctxt + val old_cmd_info = get_last_cmd_info_generic context val new_dp_info = transform_dp_info phi dp_info val new_cmd_info = old_cmd_info |> map_cmd_dp_info (K (SOME new_dp_info)) in - ctxt + context |> map_monadified_terms_generic monad_name (Item_Net.update new_cmd_info) |> put_last_cmd_info (SOME new_cmd_info) end) fun add_tmp_cmd_info (scope, head, locale_opt) = Local_Theory.declaration {pervasive = false, syntax = false} - (fn phi => fn ctxt => + (fn phi => fn context => let val new_cmd_info = { scope = Morphism.binding phi scope, head = Morphism.term phi head, locale = locale_opt, dp_info = NONE } in - ctxt |> put_last_cmd_info (SOME new_cmd_info) + context |> put_last_cmd_info (SOME new_cmd_info) end ) -fun get_or_last_cmd_info lthy monad_name_tm_opt = +fun get_or_last_cmd_info ctxt monad_name_tm_opt = case monad_name_tm_opt of - NONE => get_last_cmd_info lthy - | SOME (monad_name, tm) => get_cmd_info monad_name lthy tm |> the_single + NONE => get_last_cmd_info ctxt + | SOME (monad_name, tm) => get_cmd_info monad_name ctxt tm |> the_single end diff --git a/thys/Monad_Memo_DP/transform/Transform_Misc.ML b/thys/Monad_Memo_DP/transform/Transform_Misc.ML --- a/thys/Monad_Memo_DP/transform/Transform_Misc.ML +++ b/thys/Monad_Memo_DP/transform/Transform_Misc.ML @@ -1,73 +1,86 @@ -structure Transform_Misc = struct +signature TRANSFORM_MISC = +sig + val get_const_pat: Proof.context -> string -> term + val totality_of: Function.info -> thm + val rel_of: Function.info -> Proof.context -> thm + val the_element: int list -> int + val add_function: binding -> term list -> local_theory -> Function.info * local_theory + val behead: term -> term -> term * term list + val term_name: term -> string + val locale_term: Proof.context -> string -> string -> term + val locale_thms: Proof.context -> string -> string -> thm list + val uncurry: term -> term +end + +structure Transform_Misc : TRANSFORM_MISC = +struct fun import_function_info term_opt ctxt = case term_opt of SOME tm => (case Function_Common.import_function_data tm ctxt of SOME info => info | NONE => raise TERM("not a function", [tm])) | NONE => (case Function_Common.import_last_function ctxt of SOME info => info | NONE => error "no function defined yet") fun get_const_pat ctxt tm_pat = let val (Const (name, _)) = Proof_Context.read_const {proper=false, strict=false} ctxt tm_pat in Const (name, dummyT) end fun head_of (func_info: Function.info) = #fs func_info |> the_single fun bind_of (func_info: Function.info) = #fnames func_info |> the_single fun totality_of (func_info: Function.info) = func_info |> #totality |> the; fun rel_of (func_info: Function.info) ctxt = Inductive.the_inductive ctxt (#R func_info) |> snd |> #eqs |> the_single; fun the_element l = if tl l |> find_first (not o equal (hd l)) |> is_none then hd l else (@{print} l; error "inconsistent n_args") fun add_function bind defs = let val fixes = [(bind, NONE, NoSyn)]; val specs = map (fn def => (((Binding.empty, []), def), [], [])) defs - val pat_completeness_auto = fn ctxt => - Pat_Completeness.pat_completeness_tac ctxt 1 - THEN auto_tac ctxt in - Function.add_function fixes specs Function_Fun.fun_config pat_completeness_auto + Function.add_function fixes specs Function_Fun.fun_config + (fn ctxt => Pat_Completeness.pat_completeness_tac ctxt 1 THEN auto_tac ctxt) end fun behead head tm = let val head_nargs = strip_comb head |> snd |> length val (tm_head, tm_args) = strip_comb tm val (tm_args0, tm_args1) = chop head_nargs tm_args val tm_head' = list_comb (tm_head, tm_args0) val _ = if Term.aconv_untyped (head, tm_head') then () else raise TERM("head does not match", [head, tm_head']) in (tm_head', tm_args1) end fun term_name tm = if is_Free tm orelse is_Const tm then Term.term_name tm else raise TERM("not an atom, explicit name required", [tm]) - fun locale_term lthy locale_name term_name = - Syntax.read_term lthy (Long_Name.qualify locale_name term_name) + fun locale_term ctxt locale_name term_name = + Syntax.read_term ctxt (Long_Name.qualify locale_name term_name) - fun locale_thms lthy locale_name thms_name = - Proof_Context.get_thms lthy (Long_Name.qualify locale_name thms_name) + fun locale_thms ctxt locale_name thms_name = + Proof_Context.get_thms ctxt (Long_Name.qualify locale_name thms_name) fun uncurry tm = let val arg_typs = fastype_of tm |> binder_types val arg_names = Name.invent_list [] "a" (length arg_typs) val args = map Free (arg_names ~~ arg_typs) val args_tuple = HOLogic.mk_tuple args val tm' = list_comb (tm, args) |> HOLogic.tupled_lambda args_tuple in tm' end end diff --git a/thys/Monad_Memo_DP/transform/Transform_Parser.ML b/thys/Monad_Memo_DP/transform/Transform_Parser.ML --- a/thys/Monad_Memo_DP/transform/Transform_Parser.ML +++ b/thys/Monad_Memo_DP/transform/Transform_Parser.ML @@ -1,44 +1,52 @@ -structure Transform_Parser = struct +signature TRANSFORM_PARSER = +sig + val dp_fun_part1_parser: ((binding * string) * ((bool * (string * Position.T)) * + (string * string) list) option) parser + val dp_fun_part2_parser: (string * (Facts.ref * Token.src list) list) parser +end + +structure Transform_Parser : TRANSFORM_PARSER = +struct val dp_fun_parser = Parse.binding (* name of instantiation and monadified term *) (* fun dp_fun binding = Transform_Data.update_last_binding binding *) val memoizes_parser = Parse.name_position (* name of locale, e.g. dp_consistency_rbt *) val monadifies_parser = Parse.term (* term to be monadified *) -- Scan.option ( @{keyword "("} |-- Parse.thms1 --| (* optional definitions, ".simps" as default *) @{keyword ")"}) val dp_monadify_cmd_parser = Scan.optional (Parse.binding --| Parse.$$$ ":") Binding.empty (* optional scope *) -- Parse.term (* term to be monadified *) -- Scan.option (@{keyword "("} |-- (* optional definitions, ".simps" as default *) Parse.thms1 --| @{keyword ")"}) -- Scan.option (@{keyword with_memory} |-- Parse.name_position) (* e.g. dp_consistency_rbt *) val instance = (Parse.where_ |-- Parse.and_list1 (Parse.name -- (Parse.$$$ "=" |-- Parse.term)) || Scan.succeed []) val dp_fun_part1_parser = (Parse.binding --| Parse.$$$ ":") (* scope, e.g., bf\<^sub>T *) -- Parse.term (* term to be monadified, e.g., bf *) -- Scan.option (@{keyword with_memory} |-- Parse.opt_keyword "default_proof" -- Parse.name_position -- instance ) (* e.g. dp_consistency_rbt *) val dp_fun_part2_parser = (* monadifies *) (@{keyword "("} |-- Parse.name --| @{keyword ")"}) -- Parse.thms1 end diff --git a/thys/Monad_Memo_DP/transform/Transform_Tactic.ML b/thys/Monad_Memo_DP/transform/Transform_Tactic.ML --- a/thys/Monad_Memo_DP/transform/Transform_Tactic.ML +++ b/thys/Monad_Memo_DP/transform/Transform_Tactic.ML @@ -1,141 +1,157 @@ -structure Transform_Tactic = struct +signature TRANSFORM_TACTIC = +sig + val my_print_tac: string -> tactic + val totality_resolve_tac: thm -> thm -> thm -> Proof.context -> tactic + val totality_replay_tac: Function.info -> Function.info -> Proof.context -> tactic + val solve_relator_tac: Proof.context -> int -> tactic + val transfer_raw_tac: Proof.context -> int -> tactic + val step_tac: Proof.context -> Transform_Data.cmd_info -> int -> tactic + val prepare_case_tac: Proof.context -> Transform_Data.cmd_info -> int -> tactic + val prepare_consistentDP_tac: Proof.context -> Transform_Data.cmd_info -> int -> tactic + val solve_consistentDP_tac: Proof.context -> Transform_Data.cmd_info -> int -> tactic + val prepare_combinator_tac: Proof.context -> Transform_Data.cmd_info -> int -> tactic + val dp_unfold_defs_tac: Proof.context -> Transform_Data.cmd_info -> int -> tactic +end + +structure Transform_Tactic : TRANSFORM_TACTIC = +struct fun my_print_tac msg st = (tracing msg; all_tac st) - + fun totality_resolve_tac totality0 def0 def1 ctxt = let val totality0_unfolded = totality0 |> Local_Defs.unfold ctxt [def0] val totality1 = totality0_unfolded |> Local_Defs.fold ctxt [def1] in if Thm.full_prop_of totality0_unfolded aconv Thm.full_prop_of totality1 then let val msg = Pretty.string_of (Pretty.block [ Pretty.str "Failed to transform totality from", Pretty.brk 1, Pretty.quote (Syntax.pretty_term ctxt (Thm.full_prop_of def0)), Pretty.brk 1, Pretty.str "to", Pretty.brk 1, Pretty.quote (Syntax.pretty_term ctxt (Thm.full_prop_of def1)), Pretty.brk 1]) in (*print_tac ctxt msg THEN*) no_tac end else HEADGOAL (resolve_tac ctxt [totality1]) end - + fun totality_blast_tac totality0 def0 def1 ctxt = HEADGOAL ( (resolve_tac ctxt [totality0 RS @{thm rev_iffD1}]) THEN' (resolve_tac ctxt [@{thm arg_cong[where f=HOL.All]}]) THEN' SELECT_GOAL (unfold_tac ctxt (map (Local_Defs.abs_def_rule ctxt) [def0, def1])) THEN' (resolve_tac ctxt [@{thm arg_cong[where f=Wellfounded.accp]}]) THEN' (Blast.depth_tac ctxt 2) ) - + fun totality_replay_tac old_info new_info ctxt = let val totality0 = Transform_Misc.totality_of old_info val def0 = Transform_Misc.rel_of old_info ctxt val def1 = Transform_Misc.rel_of new_info ctxt fun my_print_tac msg st = (tracing msg; all_tac st) in no_tac ORELSE (totality_resolve_tac totality0 def0 def1 ctxt THEN my_print_tac "termination by replaying") ORELSE (totality_blast_tac totality0 def0 def1 ctxt THEN my_print_tac "termination by blast") end fun dp_intro_tac ctxt (cmd_info: Transform_Data.cmd_info) = let val scope_name = Binding.name_of (#scope cmd_info) val consistentDP_rule = Transform_Misc.locale_thms ctxt scope_name "consistentDP_intro" in resolve_tac ctxt consistentDP_rule end fun expand_relator_tac ctxt = SELECT_GOAL (Local_Defs.fold_tac ctxt (Transfer.get_relator_eq ctxt)) fun solve_relator_tac ctxt = SOLVED' (Transfer.eq_tac ctxt) fun split_params_tac ctxt = clarify_tac ctxt fun dp_induct_tac ctxt (cmd_info: Transform_Data.cmd_info) = let val dpT' = cmd_info |> #dp_info |> the |> #new_head' val dpT'_info = Function.get_info ctxt dpT' val induct_rule = dpT'_info |> #inducts |> the in resolve_tac ctxt induct_rule end fun dp_unfold_def_tac ctxt (cmd_info: Transform_Data.cmd_info) sel = cmd_info |> #dp_info |> the |> sel |> map (Local_Defs.meta_rewrite_rule ctxt) - |> Conv.rewrs_conv + |> Conv.rewrs_conv |> Conv.try_conv |> Conv.binop_conv - |> HOLogic.Trueprop_conv + |> HOLogic.Trueprop_conv |> Conv.concl_conv ~1 |> (fn cv => Conv.params_conv ~1 (K cv) ctxt) |> CONVERSION (* |> EqSubst.eqsubst_tac ctxt [0] : may rewrite locale parameters in certain situations *) fun dp_match_rule_tac ctxt (cmd_info: Transform_Data.cmd_info) = let val scope_name = Binding.name_of (#scope cmd_info) val dp_match_rules = Transform_Misc.locale_thms ctxt scope_name "dp_match_rule" in resolve_tac ctxt dp_match_rules end fun checkmem_tac ctxt (cmd_info: Transform_Data.cmd_info) = let val scope_name = Binding.name_of (#scope cmd_info) val dp_match_rules = Transform_Misc.locale_thms ctxt scope_name "crel_vs_checkmem_tupled" in resolve_tac ctxt dp_match_rules THEN' SOLVED' (clarify_tac ctxt) THEN' Transfer.eq_tac ctxt end fun solve_IH_tac ctxt = Method.assm_tac ctxt fun transfer_raw_tac ctxt = resolve_tac ctxt (Transfer.get_transfer_raw ctxt) fun step_tac ctxt (cmd_info: Transform_Data.cmd_info) = solve_IH_tac ctxt ORELSE' solve_relator_tac ctxt ORELSE' dp_match_rule_tac ctxt cmd_info ORELSE' transfer_raw_tac ctxt fun prepare_case_tac ctxt (cmd_info: Transform_Data.cmd_info) = dp_unfold_def_tac ctxt cmd_info #new_def' THEN' checkmem_tac ctxt cmd_info THEN' dp_unfold_def_tac ctxt cmd_info #old_defs fun solve_case_tac ctxt (cmd_info: Transform_Data.cmd_info) = prepare_case_tac ctxt cmd_info THEN' REPEAT_ALL_NEW (step_tac ctxt cmd_info) fun prepare_consistentDP_tac ctxt (cmd_info: Transform_Data.cmd_info) = dp_intro_tac ctxt cmd_info THEN' expand_relator_tac ctxt THEN' split_params_tac ctxt THEN' dp_induct_tac ctxt cmd_info fun solve_consistentDP_tac ctxt (cmd_info: Transform_Data.cmd_info) = prepare_consistentDP_tac ctxt cmd_info THEN_ALL_NEW SOLVED' (solve_case_tac ctxt cmd_info) fun prepare_combinator_tac ctxt (cmd_info: Transform_Data.cmd_info) = EqSubst.eqsubst_tac ctxt [0] @{thms Rel_def[symmetric]} THEN' dp_unfold_def_tac ctxt cmd_info (single o #new_defT) THEN' REPEAT_ALL_NEW (resolve_tac ctxt (@{thm Rel_abs} :: Transform_Misc.locale_thms ctxt "local" "crel_vs_return_ext")) THEN' (SELECT_GOAL (unfold_tac ctxt @{thms Rel_def})) fun dp_unfold_defs_tac ctxt (cmd_info: Transform_Data.cmd_info) = dp_unfold_def_tac ctxt cmd_info #new_def' THEN' dp_unfold_def_tac ctxt cmd_info #old_defs end diff --git a/thys/Monad_Memo_DP/transform/Transform_Term.ML b/thys/Monad_Memo_DP/transform/Transform_Term.ML --- a/thys/Monad_Memo_DP/transform/Transform_Term.ML +++ b/thys/Monad_Memo_DP/transform/Transform_Term.ML @@ -1,340 +1,344 @@ -fun list_conv (head_conv, arg_convs) lthy = - Library.foldl (uncurry Conv.combination_conv) (head_conv lthy, map (fn conv => conv lthy) arg_convs) +signature TRANSFORM_TERM = +sig + val repeat_sweep_conv: (Proof.context -> conv) -> Proof.context -> conv + val rewrite_pureapp_beta_conv: conv + val wrap_head: Transform_Const.MONAD_CONSTS -> term -> int -> term + val lift_equation: Transform_Const.MONAD_CONSTS -> Proof.context -> + term * term -> term option -> (Proof.context -> conv) * term * int +end + +structure Transform_Term : TRANSFORM_TERM = +struct + +fun list_conv (head_conv, arg_convs) ctxt = + Library.foldl (uncurry Conv.combination_conv) + (head_conv ctxt, map (fn conv => conv ctxt) arg_convs) fun eta_conv1 ctxt = (Conv.abs_conv (K Conv.all_conv) ctxt) else_conv (Thm.eta_long_conversion then_conv Conv.abs_conv (K Thm.eta_conversion) ctxt) fun eta_conv_n n = - funpow n (fn conv => fn ctxt => eta_conv1 ctxt then_conv Conv.abs_conv (fn (_, ctxt) => conv ctxt) ctxt) (K Conv.all_conv) + funpow n (fn conv => fn ctxt => eta_conv1 ctxt then_conv Conv.abs_conv (conv o #2) ctxt) (K Conv.all_conv) fun conv_changed conv ctm = let val eq = conv ctm in if Thm.is_reflexive eq then Conv.no_conv ctm else eq end fun repeat_sweep_conv conv = Conv.repeat_conv o conv_changed o Conv.top_sweep_conv conv val app_mark_conv = Conv.rewr_conv @{thm App_def[symmetric]} val app_unmark_conv = Conv.rewr_conv @{thm App_def} val wrap_mark_conv = Conv.rewr_conv @{thm Wrap_def[symmetric]} -structure Transform_Term = struct - fun eta_expand tm = let val n_args = Integer.min 1 (length (binder_types (fastype_of tm))) val (args, body) = Term.strip_abs_eta n_args tm in Library.foldr (uncurry Term.absfree) (args, body) end fun is_ctr_sugar ctxt tp_name = is_some (Ctr_Sugar.ctr_sugar_of ctxt tp_name) fun type_nargs tp = tp |> strip_type |> fst |> length fun term_nargs tm = type_nargs (fastype_of tm) -fun - lift_type (monad_consts: Transform_Const.MONAD_CONSTS) ctxt tp = #mk_stateT monad_consts (lift_type' monad_consts ctxt tp) -and - lift_type' monad_consts ctxt (tp as Type (@{type_name fun}, _)) - = lift_type' monad_consts ctxt (domain_type tp) --> lift_type monad_consts ctxt (range_type tp) -| lift_type' monad_consts ctxt (tp as Type (tp_name, tp_args)) - = if is_ctr_sugar ctxt tp_name then Type (tp_name, map (lift_type' monad_consts ctxt) tp_args) +fun lift_type (monad_consts: Transform_Const.MONAD_CONSTS) ctxt tp = + #mk_stateT monad_consts (lift_type' monad_consts ctxt tp) +and lift_type' monad_consts ctxt (tp as Type (@{type_name fun}, _)) = + lift_type' monad_consts ctxt (domain_type tp) --> lift_type monad_consts ctxt (range_type tp) + | lift_type' monad_consts ctxt (tp as Type (tp_name, tp_args)) = + if is_ctr_sugar ctxt tp_name then Type (tp_name, map (lift_type' monad_consts ctxt) tp_args) else if null tp_args then tp (* int, nat, \ *) else raise TYPE("not a ctr_sugar", [tp], []) -| lift_type' _ _ tp = tp + | lift_type' _ _ tp = tp fun is_atom_type monad_consts ctxt tp = tp = lift_type' monad_consts ctxt tp fun is_1st_type monad_consts ctxt tp = body_type tp :: binder_types tp |> forall (is_atom_type monad_consts ctxt) fun orig_atom ctxt atom_name = Proof_Context.read_term_pattern ctxt atom_name fun is_1st_term monad_consts ctxt tm = is_1st_type monad_consts ctxt (fastype_of tm) fun is_1st_atom monad_consts ctxt atom_name = is_1st_term monad_consts ctxt (orig_atom ctxt atom_name) -fun wrap_1st_term monad_consts ctxt tm n_args_opt inner_wrap = +fun wrap_1st_term monad_consts tm n_args_opt inner_wrap = let val n_args = the_default (term_nargs tm) n_args_opt val (vars_name_typ, body) = Term.strip_abs_eta n_args tm - fun wrap (name_typ, (conv, tm)) = ( - eta_conv1 ctxt then_conv Conv.abs_conv (K conv) ctxt then_conv wrap_mark_conv, - #return monad_consts (Term.absfree name_typ tm) - ) + fun wrap (name_typ, (conv, tm)) = + (fn ctxt => eta_conv1 ctxt then_conv Conv.abs_conv (conv o #2) ctxt then_conv wrap_mark_conv, + #return monad_consts (Term.absfree name_typ tm)) val (conv, result) = Library.foldr wrap (vars_name_typ, ( - if inner_wrap then (wrap_mark_conv, #return monad_consts body) else (Conv.all_conv, body) + if inner_wrap then (K wrap_mark_conv, #return monad_consts body) else (K Conv.all_conv, body) )) - in - (K conv, result) + (conv, result) end fun lift_1st_atom monad_consts ctxt atom (name, tp) = let val (arg_typs, body_typ) = strip_type tp val n_args = term_nargs (orig_atom ctxt name) val (arg_typs, body_arg_typs) = chop n_args arg_typs val arg_typs' = map (lift_type' monad_consts ctxt) arg_typs val body_typ' = lift_type' monad_consts ctxt (body_arg_typs ---> body_typ) val tm' = atom (name, arg_typs' ---> body_typ') (* " *) in - wrap_1st_term monad_consts ctxt tm' (SOME n_args) true + wrap_1st_term monad_consts tm' (SOME n_args) true end fun fixed_args head_n_args tm = let val (tm_head, tm_args) = strip_comb tm val n_tm_args = length tm_args in head_n_args tm_head |> Option.mapPartial (fn n_args => if n_tm_args > n_args then NONE else if n_tm_args < n_args then raise TERM("need " ^ string_of_int n_args ^ " args", [tm]) else SOME (tm_head, tm_args)) end fun lift_abs' monad_consts ctxt (name, typ) cont lift_dict body = let val free = Free (name, typ) val typ' = lift_type' monad_consts ctxt typ val freeT' = Free (name, typ') val freeT = #return monad_consts (freeT') val lift_dict' = if is_atom_type monad_consts ctxt typ then lift_dict else (free, (K wrap_mark_conv, freeT))::lift_dict val (conv_free, body_free) = (cont (lift_dict') body) val body' = lambda freeT' body_free - fun conv lthy = - eta_conv1 ctxt then_conv Conv.abs_conv (fn (_, lthy') => conv_free lthy') lthy - in - (conv, body') - end + fun conv ctxt' = eta_conv1 ctxt' then_conv Conv.abs_conv (conv_free o #2) ctxt' + in (conv, body') end fun lift_arg monad_consts ctxt lift_dict tm = (* let val (conv, tm') = lift_term ctxt lift_dict (eta_expand tm) fun conv' ctxt = Conv.try_conv (eta_conv1 ctxt) then_conv (conv ctxt) in (conv', tm') end eta_expand AFTER lifting *) lift_term monad_consts ctxt lift_dict tm and lift_term monad_consts ctxt lift_dict tm = let val case_terms = Ctr_Sugar.ctr_sugars_of ctxt |> map #casex fun lookup_case_term tm = find_first (fn x => Term.aconv_untyped (x, tm)) case_terms val check_cont = lift_term monad_consts ctxt val check_cont_arg = lift_arg monad_consts ctxt fun check_const tm = case tm of Const (_, typ) => ( case Transform_Data.get_dp_info (#monad_name monad_consts) ctxt tm of SOME {new_headT=Const (name, _), ...} => SOME (K Conv.all_conv, Const (name, lift_type monad_consts ctxt typ)) | SOME {new_headT=tm', ...} => raise TERM("not a constant", [tm']) | NONE => NONE) | _ => NONE fun check_1st_atom tm = case tm of Const (name, typ) => if is_1st_atom monad_consts ctxt name then SOME (lift_1st_atom monad_consts ctxt Const (name, typ)) else NONE | Free (name, typ) => if is_1st_atom monad_consts ctxt name then SOME (lift_1st_atom monad_consts ctxt Free (name, typ)) else NONE | _ => (* if is_1st_term ctxt tm andalso exists_subterm (AList.defined (op aconv) lift_dict) tm then SOME (wrap_1st_term tm NONE) else *) NONE (* fun check_dict tm = (* TODO: map -> mapT, dummyT *) AList.lookup Term.aconv_untyped lift_dict tm |> Option.map (fn tm' => if is_Const tm then (@{assert} (is_Const tm'); map_types (K (lift_type ctxt (type_of tm))) tm') else tm') *) fun check_dict tm = AList.lookup Term.aconv_untyped lift_dict tm fun check_if tm = fixed_args (fn head => if Term.aconv_untyped (head, @{term If}) then SOME 3 else NONE) tm |> Option.map (fn (_, args) => let val (arg_convs, args') = map (check_cont lift_dict) args |> split_list val conv = list_conv (K Conv.all_conv, arg_convs) val tm' = list_comb (Const (#if_termN monad_consts, dummyT), args') in (conv, tm') end) fun check_abs tm = case tm of - Abs (name, typ, body) => + Abs _ => let - val (name', body') = Term.dest_abs (name, typ, body) + val ((name', typ), body') = Term.dest_abs_global tm val (conv, tm') = lift_abs' monad_consts ctxt (name', typ) check_cont lift_dict body' - fun convT lthy = conv lthy then_conv wrap_mark_conv + fun convT ctxt' = conv ctxt' then_conv wrap_mark_conv val tmT = #return monad_consts tm' in SOME (convT, tmT) end | _ => NONE fun check_case tm = fixed_args (lookup_case_term #> Option.map (fn cs => term_nargs cs - 1)) tm |> Option.map (fn (head, args) => let val (case_name, case_type) = lookup_case_term head |> the |> dest_Const val ((clause_typs, _), _) = strip_type case_type |>> split_last - + val clase_nparams = clause_typs |> map type_nargs (* ('a\'b) \ ('a\'b) |> type_nargs = 1*) - + fun lift_clause n_param clause = let val (vars_name_typ, body) = Term.strip_abs_eta n_param clause val abs_lift_wraps = map (lift_abs' monad_consts ctxt) vars_name_typ val lift_wrap = Library.foldr (op o) (abs_lift_wraps, I) check_cont val (conv, clauseT) = lift_wrap lift_dict body in (conv, clauseT) end - + val head' = Const (case_name, dummyT) (* clauses are sufficient for type inference *) val (convs, clauses') = map2 lift_clause clase_nparams args |> split_list - fun conv lthy = list_conv (K Conv.all_conv, convs) lthy then_conv wrap_mark_conv + fun conv ctxt' = list_conv (K Conv.all_conv, convs) ctxt' then_conv wrap_mark_conv val tm' = #return monad_consts (list_comb (head', clauses')) in (conv, tm') end) fun check_app tm = case tm of f $ x => let val (f_conv, tmf) = check_cont lift_dict f val (x_conv, tmx) = check_cont_arg lift_dict x val tm' = #app monad_consts (tmf, tmx) - fun conv lthy = Conv.combination_conv (f_conv lthy then_conv app_mark_conv) (x_conv lthy) + fun conv ctxt' = Conv.combination_conv (f_conv ctxt' then_conv app_mark_conv) (x_conv ctxt') in SOME (conv, tm') end | _ => NONE fun check_pure tm = if tm |> exists_subterm (AList.defined (op aconv) lift_dict) orelse not (is_1st_term monad_consts ctxt tm) then NONE - else SOME (wrap_1st_term monad_consts ctxt tm NONE true) + else SOME (wrap_1st_term monad_consts tm NONE true) fun choke tm = raise TERM("cannot process term", [tm]) val checks = [ check_pure, check_const, check_case, check_if, check_abs, check_app, check_dict, check_1st_atom, choke ] in get_first (fn check => check tm) checks |> the end fun rewrite_pureapp_beta_conv ctm = case Thm.term_of ctm of Const (@{const_name Pure_Monad.App}, _) $ (Const (@{const_name Pure_Monad.Wrap}, _) $ Abs _) $ (Const (@{const_name Pure_Monad.Wrap}, _) $ _) => Conv.rewr_conv @{thm Wrap_App_Wrap} ctm | _ => Conv.no_conv ctm fun monadify monad_consts ctxt tm = let val (_, tm) = lift_term monad_consts ctxt [] tm (*val tm = rewrite_return_app_return tm*) val tm = Syntax.check_term ctxt tm in tm end fun wrap_head (monad_consts: Transform_Const.MONAD_CONSTS) head n_args = Library.foldr (fn (typ, tm) => #return monad_consts (absdummy typ tm)) (replicate n_args dummyT, list_comb (head, rev (map_range Bound n_args))) fun lift_head monad_consts ctxt head n_args = let val dest_head = if is_Const head then dest_Const else dest_Free val (head_name, head_typ) = dest_head head val (arg_typs, body_typ) = strip_type head_typ val (arg_typs0, arg_typs1) = chop n_args arg_typs val arg_typs0' = map (lift_type' monad_consts ctxt) arg_typs0 val arg_typs1T = lift_type monad_consts ctxt (arg_typs1 ---> body_typ) val head_typ' = arg_typs0' ---> arg_typs1T val head' = Free (head_name, head_typ') - val (head_conv, headT) = wrap_1st_term monad_consts ctxt head' (SOME n_args) false + val (head_conv, headT) = wrap_1st_term monad_consts head' (SOME n_args) false in (head', (head_conv, headT)) end fun lift_equation monad_consts ctxt (lhs, rhs) memoizer_opt = let val (head, args) = strip_comb lhs val n_args = length args val (head', (head_conv, headT)) = lift_head monad_consts ctxt head n_args val args' = args |> map (map_aterms (fn tm => tm |> map_types (if is_Const tm then K dummyT else lift_type' monad_consts ctxt))) val lhs' = list_comb (head', args') val frees = fold Term.add_frees args [] |> filter_out (is_atom_type monad_consts ctxt o snd) - + val lift_dict_args = frees |> map (fn (name, typ) => ( - Free (name, typ), + Free (name, typ), (K wrap_mark_conv, #return monad_consts (Free (name, lift_type' monad_consts ctxt typ))) )) val lift_dict = (head, (head_conv, headT)) :: lift_dict_args val (rhs_conv, rhsT) = lift_term monad_consts ctxt lift_dict rhs val rhsT_memoized = case memoizer_opt of SOME memoizer => memoizer $ HOLogic.mk_tuple args $ rhsT | NONE => rhsT val eqs' = (lhs', rhsT_memoized) |> HOLogic.mk_eq |> HOLogic.mk_Trueprop in (rhs_conv, eqs', n_args) end - end diff --git a/thys/Monad_Memo_DP/util/Ground_Function.ML b/thys/Monad_Memo_DP/util/Ground_Function.ML --- a/thys/Monad_Memo_DP/util/Ground_Function.ML +++ b/thys/Monad_Memo_DP/util/Ground_Function.ML @@ -1,39 +1,47 @@ (** Define a new ground constant from an existing function definition **) -structure Ground_Function = struct + +signature GROUND_FUNCTION = +sig + val mk_fun: bool -> thm list -> binding -> local_theory -> local_theory +end + +structure Ground_Function : GROUND_FUNCTION = +struct fun add_function bind defs = let val fixes = [(bind, NONE, NoSyn)]; val specs = map (fn def => (((Binding.empty, []), def), [], [])) defs - val pat_completeness_auto = fn ctxt => - Pat_Completeness.pat_completeness_tac ctxt 1 - THEN auto_tac ctxt in - Function.add_function fixes specs Function_Fun.fun_config pat_completeness_auto + Function.add_function fixes specs Function_Fun.fun_config + (fn ctxt => Pat_Completeness.pat_completeness_tac ctxt 1 THEN auto_tac ctxt) end fun dest_hol_eq_prop t = - let - val Const ("HOL.Trueprop", _) $ (Const ("HOL.eq", _) $ a $ b) = t + let val \<^Const_>\Trueprop for \<^Const_>\HOL.eq _ for a b\\ = t in (a, b) end + fun get_fun_head t = let val (t, _) = dest_hol_eq_prop t val t = Term.head_of t val Const (fun_name, fun_ty) = t in (fun_name, fun_ty) end -fun mk_fun termination simps binding ctxt = +fun mk_fun termination simps binding lthy = let val eqns = map Thm.concl_of simps - val (eqns, _) = Variable.import_terms true eqns ctxt + val (eqns, _) = Variable.import_terms true eqns lthy val (f_name, f_ty) = get_fun_head (hd eqns) val s = Binding.name_of binding val replacement = (Const (f_name, f_ty), Free (s, f_ty)) val eqns = map (subst_free [replacement]) eqns - val (_, ctxt) = add_function binding eqns ctxt - fun prove_termination lthy = - Function.prove_termination NONE (Function_Common.termination_prover_tac false lthy) lthy - in ctxt |> (if termination then snd o prove_termination else I) end + fun prove_termination lthy' = lthy' + |> Function.prove_termination NONE (Function_Common.termination_prover_tac false lthy') + in + lthy + |> add_function binding eqns |> #2 + |> termination ? (prove_termination #> #2) + end end diff --git a/thys/Nominal2/Nominal2_Base.thy b/thys/Nominal2/Nominal2_Base.thy --- a/thys/Nominal2/Nominal2_Base.thy +++ b/thys/Nominal2/Nominal2_Base.thy @@ -1,3435 +1,3435 @@ (* Title: Nominal2_Base Authors: Christian Urban, Brian Huffman, Cezary Kaliszyk Basic definitions and lemma infrastructure for Nominal Isabelle. *) theory Nominal2_Base imports "HOL-Library.Infinite_Set" "HOL-Library.Multiset" "HOL-Library.FSet" FinFun.FinFun keywords "atom_decl" "equivariance" :: thy_decl begin declare [[typedef_overloaded]] section \Atoms and Sorts\ text \A simple implementation for \atom_sorts\ is strings.\ (* types atom_sort = string *) text \To deal with Church-like binding we use trees of strings as sorts.\ datatype atom_sort = Sort "string" "atom_sort list" datatype atom = Atom atom_sort nat text \Basic projection function.\ primrec sort_of :: "atom \ atom_sort" where "sort_of (Atom s n) = s" primrec nat_of :: "atom \ nat" where "nat_of (Atom s n) = n" text \There are infinitely many atoms of each sort.\ lemma INFM_sort_of_eq: shows "INFM a. sort_of a = s" proof - have "INFM i. sort_of (Atom s i) = s" by simp moreover have "inj (Atom s)" by (simp add: inj_on_def) ultimately show "INFM a. sort_of a = s" by (rule INFM_inj) qed lemma infinite_sort_of_eq: shows "infinite {a. sort_of a = s}" using INFM_sort_of_eq unfolding INFM_iff_infinite . lemma atom_infinite [simp]: shows "infinite (UNIV :: atom set)" using subset_UNIV infinite_sort_of_eq by (rule infinite_super) lemma obtain_atom: fixes X :: "atom set" assumes X: "finite X" obtains a where "a \ X" "sort_of a = s" proof - from X have "MOST a. a \ X" unfolding MOST_iff_cofinite by simp with INFM_sort_of_eq have "INFM a. sort_of a = s \ a \ X" by (rule INFM_conjI) then obtain a where "a \ X" "sort_of a = s" by (auto elim: INFM_E) then show ?thesis .. qed lemma atom_components_eq_iff: fixes a b :: atom shows "a = b \ sort_of a = sort_of b \ nat_of a = nat_of b" by (induct a, induct b, simp) section \Sort-Respecting Permutations\ definition "perm \ {f. bij f \ finite {a. f a \ a} \ (\a. sort_of (f a) = sort_of a)}" typedef perm = "perm" proof show "id \ perm" unfolding perm_def by simp qed lemma permI: assumes "bij f" and "MOST x. f x = x" and "\a. sort_of (f a) = sort_of a" shows "f \ perm" using assms unfolding perm_def MOST_iff_cofinite by simp lemma perm_is_bij: "f \ perm \ bij f" unfolding perm_def by simp lemma perm_is_finite: "f \ perm \ finite {a. f a \ a}" unfolding perm_def by simp lemma perm_is_sort_respecting: "f \ perm \ sort_of (f a) = sort_of a" unfolding perm_def by simp lemma perm_MOST: "f \ perm \ MOST x. f x = x" unfolding perm_def MOST_iff_cofinite by simp lemma perm_id: "id \ perm" unfolding perm_def by simp lemma perm_comp: assumes f: "f \ perm" and g: "g \ perm" shows "(f \ g) \ perm" apply (rule permI) apply (rule bij_comp) apply (rule perm_is_bij [OF g]) apply (rule perm_is_bij [OF f]) apply (rule MOST_rev_mp [OF perm_MOST [OF g]]) apply (rule MOST_rev_mp [OF perm_MOST [OF f]]) apply (simp) apply (simp add: perm_is_sort_respecting [OF f]) apply (simp add: perm_is_sort_respecting [OF g]) done lemma perm_inv: assumes f: "f \ perm" shows "(inv f) \ perm" apply (rule permI) apply (rule bij_imp_bij_inv) apply (rule perm_is_bij [OF f]) apply (rule MOST_mono [OF perm_MOST [OF f]]) apply (erule subst, rule inv_f_f) apply (rule bij_is_inj [OF perm_is_bij [OF f]]) apply (rule perm_is_sort_respecting [OF f, THEN sym, THEN trans]) apply (simp add: surj_f_inv_f [OF bij_is_surj [OF perm_is_bij [OF f]]]) done lemma bij_Rep_perm: "bij (Rep_perm p)" using Rep_perm [of p] unfolding perm_def by simp lemma finite_Rep_perm: "finite {a. Rep_perm p a \ a}" using Rep_perm [of p] unfolding perm_def by simp lemma sort_of_Rep_perm: "sort_of (Rep_perm p a) = sort_of a" using Rep_perm [of p] unfolding perm_def by simp lemma Rep_perm_ext: "Rep_perm p1 = Rep_perm p2 \ p1 = p2" by (simp add: fun_eq_iff Rep_perm_inject [symmetric]) instance perm :: size .. subsection \Permutations form a (multiplicative) group\ instantiation perm :: group_add begin definition "0 = Abs_perm id" definition "- p = Abs_perm (inv (Rep_perm p))" definition "p + q = Abs_perm (Rep_perm p \ Rep_perm q)" definition "(p1::perm) - p2 = p1 + - p2" lemma Rep_perm_0: "Rep_perm 0 = id" unfolding zero_perm_def by (simp add: Abs_perm_inverse perm_id) lemma Rep_perm_add: "Rep_perm (p1 + p2) = Rep_perm p1 \ Rep_perm p2" unfolding plus_perm_def by (simp add: Abs_perm_inverse perm_comp Rep_perm) lemma Rep_perm_uminus: "Rep_perm (- p) = inv (Rep_perm p)" unfolding uminus_perm_def by (simp add: Abs_perm_inverse perm_inv Rep_perm) instance apply standard unfolding Rep_perm_inject [symmetric] unfolding minus_perm_def unfolding Rep_perm_add unfolding Rep_perm_uminus unfolding Rep_perm_0 by (simp_all add: o_assoc inv_o_cancel [OF bij_is_inj [OF bij_Rep_perm]]) end section \Implementation of swappings\ definition swap :: "atom \ atom \ perm" ("'(_ \ _')") where "(a \ b) = Abs_perm (if sort_of a = sort_of b then (\c. if a = c then b else if b = c then a else c) else id)" lemma Rep_perm_swap: "Rep_perm (a \ b) = (if sort_of a = sort_of b then (\c. if a = c then b else if b = c then a else c) else id)" unfolding swap_def apply (rule Abs_perm_inverse) apply (rule permI) apply (auto simp: bij_def inj_on_def surj_def)[1] apply (rule MOST_rev_mp [OF MOST_neq(1) [of a]]) apply (rule MOST_rev_mp [OF MOST_neq(1) [of b]]) apply (simp) apply (simp) done lemmas Rep_perm_simps = Rep_perm_0 Rep_perm_add Rep_perm_uminus Rep_perm_swap lemma swap_different_sorts [simp]: "sort_of a \ sort_of b \ (a \ b) = 0" by (rule Rep_perm_ext) (simp add: Rep_perm_simps) lemma swap_cancel: shows "(a \ b) + (a \ b) = 0" and "(a \ b) + (b \ a) = 0" by (rule_tac [!] Rep_perm_ext) (simp_all add: Rep_perm_simps fun_eq_iff) lemma swap_self [simp]: "(a \ a) = 0" by (rule Rep_perm_ext, simp add: Rep_perm_simps fun_eq_iff) lemma minus_swap [simp]: "- (a \ b) = (a \ b)" by (rule minus_unique [OF swap_cancel(1)]) lemma swap_commute: "(a \ b) = (b \ a)" by (rule Rep_perm_ext) (simp add: Rep_perm_swap fun_eq_iff) lemma swap_triple: assumes "a \ b" and "c \ b" assumes "sort_of a = sort_of b" "sort_of b = sort_of c" shows "(a \ c) + (b \ c) + (a \ c) = (a \ b)" using assms by (rule_tac Rep_perm_ext) (auto simp: Rep_perm_simps fun_eq_iff) section \Permutation Types\ text \ Infix syntax for \permute\ has higher precedence than addition, but lower than unary minus. \ class pt = fixes permute :: "perm \ 'a \ 'a" ("_ \ _" [76, 75] 75) assumes permute_zero [simp]: "0 \ x = x" assumes permute_plus [simp]: "(p + q) \ x = p \ (q \ x)" begin lemma permute_diff [simp]: shows "(p - q) \ x = p \ - q \ x" using permute_plus [of p "- q" x] by simp lemma permute_minus_cancel [simp]: shows "p \ - p \ x = x" and "- p \ p \ x = x" unfolding permute_plus [symmetric] by simp_all lemma permute_swap_cancel [simp]: shows "(a \ b) \ (a \ b) \ x = x" unfolding permute_plus [symmetric] by (simp add: swap_cancel) lemma permute_swap_cancel2 [simp]: shows "(a \ b) \ (b \ a) \ x = x" unfolding permute_plus [symmetric] by (simp add: swap_commute) lemma inj_permute [simp]: shows "inj (permute p)" by (rule inj_on_inverseI) (rule permute_minus_cancel) lemma surj_permute [simp]: shows "surj (permute p)" by (rule surjI, rule permute_minus_cancel) lemma bij_permute [simp]: shows "bij (permute p)" by (rule bijI [OF inj_permute surj_permute]) lemma inv_permute: shows "inv (permute p) = permute (- p)" by (rule inv_equality) (simp_all) lemma permute_minus: shows "permute (- p) = inv (permute p)" by (simp add: inv_permute) lemma permute_eq_iff [simp]: shows "p \ x = p \ y \ x = y" by (rule inj_permute [THEN inj_eq]) end subsection \Permutations for atoms\ instantiation atom :: pt begin definition "p \ a = (Rep_perm p) a" instance apply standard apply(simp_all add: permute_atom_def Rep_perm_simps) done end lemma sort_of_permute [simp]: shows "sort_of (p \ a) = sort_of a" unfolding permute_atom_def by (rule sort_of_Rep_perm) lemma swap_atom: shows "(a \ b) \ c = (if sort_of a = sort_of b then (if c = a then b else if c = b then a else c) else c)" unfolding permute_atom_def by (simp add: Rep_perm_swap) lemma swap_atom_simps [simp]: "sort_of a = sort_of b \ (a \ b) \ a = b" "sort_of a = sort_of b \ (a \ b) \ b = a" "c \ a \ c \ b \ (a \ b) \ c = c" unfolding swap_atom by simp_all lemma perm_eq_iff: fixes p q :: "perm" shows "p = q \ (\a::atom. p \ a = q \ a)" unfolding permute_atom_def by (metis Rep_perm_ext ext) subsection \Permutations for permutations\ instantiation perm :: pt begin definition "p \ q = p + q - p" instance apply standard apply (simp add: permute_perm_def) apply (simp add: permute_perm_def algebra_simps) done end lemma permute_self: shows "p \ p = p" unfolding permute_perm_def by (simp add: add.assoc) lemma pemute_minus_self: shows "- p \ p = p" unfolding permute_perm_def by (simp add: add.assoc) subsection \Permutations for functions\ instantiation "fun" :: (pt, pt) pt begin definition "p \ f = (\x. p \ (f (- p \ x)))" instance apply standard apply (simp add: permute_fun_def) apply (simp add: permute_fun_def minus_add) done end lemma permute_fun_app_eq: shows "p \ (f x) = (p \ f) (p \ x)" unfolding permute_fun_def by simp lemma permute_fun_comp: shows "p \ f = (permute p) o f o (permute (-p))" by (simp add: comp_def permute_fun_def) subsection \Permutations for booleans\ instantiation bool :: pt begin definition "p \ (b::bool) = b" instance apply standard apply(simp_all add: permute_bool_def) done end lemma permute_boolE: fixes P::"bool" shows "p \ P \ P" by (simp add: permute_bool_def) lemma permute_boolI: fixes P::"bool" shows "P \ p \ P" by(simp add: permute_bool_def) subsection \Permutations for sets\ instantiation "set" :: (pt) pt begin definition "p \ X = {p \ x | x. x \ X}" instance apply standard apply (auto simp: permute_set_def) done end lemma permute_set_eq: shows "p \ X = {x. - p \ x \ X}" unfolding permute_set_def by (auto) (metis permute_minus_cancel(1)) lemma permute_set_eq_image: shows "p \ X = permute p ` X" unfolding permute_set_def by auto lemma permute_set_eq_vimage: shows "p \ X = permute (- p) -` X" unfolding permute_set_eq vimage_def by simp lemma permute_finite [simp]: shows "finite (p \ X) = finite X" unfolding permute_set_eq_vimage using bij_permute by (rule finite_vimage_iff) lemma swap_set_not_in: assumes a: "a \ S" "b \ S" shows "(a \ b) \ S = S" unfolding permute_set_def using a by (auto simp: swap_atom) lemma swap_set_in: assumes a: "a \ S" "b \ S" "sort_of a = sort_of b" shows "(a \ b) \ S \ S" unfolding permute_set_def using a by (auto simp: swap_atom) lemma swap_set_in_eq: assumes a: "a \ S" "b \ S" "sort_of a = sort_of b" shows "(a \ b) \ S = (S - {a}) \ {b}" unfolding permute_set_def using a by (auto simp: swap_atom) lemma swap_set_both_in: assumes a: "a \ S" "b \ S" shows "(a \ b) \ S = S" unfolding permute_set_def using a by (auto simp: swap_atom) lemma mem_permute_iff: shows "(p \ x) \ (p \ X) \ x \ X" unfolding permute_set_def by auto lemma empty_eqvt: shows "p \ {} = {}" unfolding permute_set_def by (simp) lemma insert_eqvt: shows "p \ (insert x A) = insert (p \ x) (p \ A)" unfolding permute_set_eq_image image_insert .. subsection \Permutations for @{typ unit}\ instantiation unit :: pt begin definition "p \ (u::unit) = u" instance by standard (simp_all add: permute_unit_def) end subsection \Permutations for products\ instantiation prod :: (pt, pt) pt begin primrec permute_prod where Pair_eqvt: "p \ (x, y) = (p \ x, p \ y)" instance by standard auto end subsection \Permutations for sums\ instantiation sum :: (pt, pt) pt begin primrec permute_sum where Inl_eqvt: "p \ (Inl x) = Inl (p \ x)" | Inr_eqvt: "p \ (Inr y) = Inr (p \ y)" instance by standard (case_tac [!] x, simp_all) end subsection \Permutations for @{typ "'a list"}\ instantiation list :: (pt) pt begin primrec permute_list where Nil_eqvt: "p \ [] = []" | Cons_eqvt: "p \ (x # xs) = p \ x # p \ xs" instance by standard (induct_tac [!] x, simp_all) end lemma set_eqvt: shows "p \ (set xs) = set (p \ xs)" by (induct xs) (simp_all add: empty_eqvt insert_eqvt) subsection \Permutations for @{typ "'a option"}\ instantiation option :: (pt) pt begin primrec permute_option where None_eqvt: "p \ None = None" | Some_eqvt: "p \ (Some x) = Some (p \ x)" instance by standard (induct_tac [!] x, simp_all) end subsection \Permutations for @{typ "'a multiset"}\ instantiation multiset :: (pt) pt begin definition "p \ M = {# p \ x. x :# M #}" instance proof fix M :: "'a multiset" and p q :: "perm" show "0 \ M = M" unfolding permute_multiset_def by (induct_tac M) (simp_all) show "(p + q) \ M = p \ q \ M" unfolding permute_multiset_def by (induct_tac M) (simp_all) qed end lemma permute_multiset [simp]: fixes M N::"('a::pt) multiset" shows "(p \ {#}) = ({#} ::('a::pt) multiset)" and "(p \ add_mset x M) = add_mset (p \ x) (p \ M)" and "(p \ (M + N)) = (p \ M) + (p \ N)" unfolding permute_multiset_def by (simp_all) subsection \Permutations for @{typ "'a fset"}\ instantiation fset :: (pt) pt begin context includes fset.lifting begin lift_definition "permute_fset" :: "perm \ 'a fset \ 'a fset" is "permute :: perm \ 'a set \ 'a set" by simp end context includes fset.lifting begin instance proof fix x :: "'a fset" and p q :: "perm" show "0 \ x = x" by transfer simp show "(p + q) \ x = p \ q \ x" by transfer simp qed end end context includes fset.lifting begin lemma permute_fset [simp]: fixes S::"('a::pt) fset" shows "(p \ {||}) = ({||} ::('a::pt) fset)" and "(p \ finsert x S) = finsert (p \ x) (p \ S)" apply (transfer, simp add: empty_eqvt) apply (transfer, simp add: insert_eqvt) done lemma fset_eqvt: shows "p \ (fset S) = fset (p \ S)" by transfer simp end subsection \Permutations for @{typ "('a, 'b) finfun"}\ instantiation finfun :: (pt, pt) pt begin lift_definition permute_finfun :: "perm \ ('a, 'b) finfun \ ('a, 'b) finfun" is "permute :: perm \ ('a \ 'b) \ ('a \ 'b)" apply(simp add: permute_fun_comp) apply(rule finfun_right_compose) apply(rule finfun_left_compose) apply(assumption) apply(simp) done instance apply standard apply(transfer) apply(simp) apply(transfer) apply(simp) done end subsection \Permutations for @{typ char}, @{typ nat}, and @{typ int}\ instantiation char :: pt begin definition "p \ (c::char) = c" instance by standard (simp_all add: permute_char_def) end instantiation nat :: pt begin definition "p \ (n::nat) = n" instance by standard (simp_all add: permute_nat_def) end instantiation int :: pt begin definition "p \ (i::int) = i" instance by standard (simp_all add: permute_int_def) end section \Pure types\ text \Pure types will have always empty support.\ class pure = pt + assumes permute_pure: "p \ x = x" text \Types @{typ unit} and @{typ bool} are pure.\ instance unit :: pure proof qed (rule permute_unit_def) instance bool :: pure proof qed (rule permute_bool_def) text \Other type constructors preserve purity.\ instance "fun" :: (pure, pure) pure by standard (simp add: permute_fun_def permute_pure) instance set :: (pure) pure by standard (simp add: permute_set_def permute_pure) instance prod :: (pure, pure) pure by standard (induct_tac x, simp add: permute_pure) instance sum :: (pure, pure) pure by standard (induct_tac x, simp_all add: permute_pure) instance list :: (pure) pure by standard (induct_tac x, simp_all add: permute_pure) instance option :: (pure) pure by standard (induct_tac x, simp_all add: permute_pure) subsection \Types @{typ char}, @{typ nat}, and @{typ int}\ instance char :: pure proof qed (rule permute_char_def) instance nat :: pure proof qed (rule permute_nat_def) instance int :: pure proof qed (rule permute_int_def) section \Infrastructure for Equivariance and \Perm_simp\\ subsection \Basic functions about permutations\ ML_file \nominal_basics.ML\ subsection \Eqvt infrastructure\ text \Setup of the theorem attributes \eqvt\ and \eqvt_raw\.\ ML_file \nominal_thmdecls.ML\ lemmas [eqvt] = (* pt types *) permute_prod.simps permute_list.simps permute_option.simps permute_sum.simps (* sets *) empty_eqvt insert_eqvt set_eqvt (* fsets *) permute_fset fset_eqvt (* multisets *) permute_multiset subsection \\perm_simp\ infrastructure\ definition "unpermute p = permute (- p)" lemma eqvt_apply: fixes f :: "'a::pt \ 'b::pt" and x :: "'a::pt" shows "p \ (f x) \ (p \ f) (p \ x)" unfolding permute_fun_def by simp lemma eqvt_lambda: fixes f :: "'a::pt \ 'b::pt" shows "p \ f \ (\x. p \ (f (unpermute p x)))" unfolding permute_fun_def unpermute_def by simp lemma eqvt_bound: shows "p \ unpermute p x \ x" unfolding unpermute_def by simp text \provides \perm_simp\ methods\ ML_file \nominal_permeq.ML\ method_setup perm_simp = \Nominal_Permeq.args_parser >> Nominal_Permeq.perm_simp_meth\ \pushes permutations inside.\ method_setup perm_strict_simp = \Nominal_Permeq.args_parser >> Nominal_Permeq.perm_strict_simp_meth\ \pushes permutations inside, raises an error if it cannot solve all permutations.\ simproc_setup perm_simproc ("p \ t") = \fn _ => fn ctxt => fn ctrm => case Thm.term_of (Thm.dest_arg ctrm) of Free _ => NONE | Var _ => NONE | \<^Const_>\permute _ for _ _\ => NONE | _ => let val thm = Nominal_Permeq.eqvt_conv ctxt Nominal_Permeq.eqvt_strict_config ctrm handle ERROR _ => Thm.reflexive ctrm in if Thm.is_reflexive thm then NONE else SOME(thm) end \ subsubsection \Equivariance for permutations and swapping\ lemma permute_eqvt: shows "p \ (q \ x) = (p \ q) \ (p \ x)" unfolding permute_perm_def by simp (* the normal version of this lemma would cause loops *) lemma permute_eqvt_raw [eqvt_raw]: shows "p \ permute \ permute" apply(simp add: fun_eq_iff permute_fun_def) apply(subst permute_eqvt) apply(simp) done lemma zero_perm_eqvt [eqvt]: shows "p \ (0::perm) = 0" unfolding permute_perm_def by simp lemma add_perm_eqvt [eqvt]: fixes p p1 p2 :: perm shows "p \ (p1 + p2) = p \ p1 + p \ p2" unfolding permute_perm_def by (simp add: perm_eq_iff) lemma swap_eqvt [eqvt]: shows "p \ (a \ b) = (p \ a \ p \ b)" unfolding permute_perm_def by (auto simp: swap_atom perm_eq_iff) lemma uminus_eqvt [eqvt]: fixes p q::"perm" shows "p \ (- q) = - (p \ q)" unfolding permute_perm_def by (simp add: diff_add_eq_diff_diff_swap) subsubsection \Equivariance of Logical Operators\ lemma eq_eqvt [eqvt]: shows "p \ (x = y) \ (p \ x) = (p \ y)" unfolding permute_eq_iff permute_bool_def .. lemma Not_eqvt [eqvt]: shows "p \ (\ A) \ \ (p \ A)" by (simp add: permute_bool_def) lemma conj_eqvt [eqvt]: shows "p \ (A \ B) \ (p \ A) \ (p \ B)" by (simp add: permute_bool_def) lemma imp_eqvt [eqvt]: shows "p \ (A \ B) \ (p \ A) \ (p \ B)" by (simp add: permute_bool_def) declare imp_eqvt[folded HOL.induct_implies_def, eqvt] lemma all_eqvt [eqvt]: shows "p \ (\x. P x) = (\x. (p \ P) x)" unfolding All_def by (perm_simp) (rule refl) declare all_eqvt[folded HOL.induct_forall_def, eqvt] lemma ex_eqvt [eqvt]: shows "p \ (\x. P x) = (\x. (p \ P) x)" unfolding Ex_def by (perm_simp) (rule refl) lemma ex1_eqvt [eqvt]: shows "p \ (\!x. P x) = (\!x. (p \ P) x)" unfolding Ex1_def by (perm_simp) (rule refl) lemma if_eqvt [eqvt]: shows "p \ (if b then x else y) = (if p \ b then p \ x else p \ y)" by (simp add: permute_fun_def permute_bool_def) lemma True_eqvt [eqvt]: shows "p \ True = True" unfolding permute_bool_def .. lemma False_eqvt [eqvt]: shows "p \ False = False" unfolding permute_bool_def .. lemma disj_eqvt [eqvt]: shows "p \ (A \ B) \ (p \ A) \ (p \ B)" by (simp add: permute_bool_def) lemma all_eqvt2: shows "p \ (\x. P x) = (\x. p \ P (- p \ x))" by (perm_simp add: permute_minus_cancel) (rule refl) lemma ex_eqvt2: shows "p \ (\x. P x) = (\x. p \ P (- p \ x))" by (perm_simp add: permute_minus_cancel) (rule refl) lemma ex1_eqvt2: shows "p \ (\!x. P x) = (\!x. p \ P (- p \ x))" by (perm_simp add: permute_minus_cancel) (rule refl) lemma the_eqvt: assumes unique: "\!x. P x" shows "(p \ (THE x. P x)) = (THE x. (p \ P) x)" apply(rule the1_equality [symmetric]) apply(rule_tac p="-p" in permute_boolE) apply(perm_simp add: permute_minus_cancel) apply(rule unique) apply(rule_tac p="-p" in permute_boolE) apply(perm_simp add: permute_minus_cancel) apply(rule theI'[OF unique]) done lemma the_eqvt2: assumes unique: "\!x. P x" shows "(p \ (THE x. P x)) = (THE x. p \ P (- p \ x))" apply(rule the1_equality [symmetric]) apply(simp only: ex1_eqvt2[symmetric]) apply(simp add: permute_bool_def unique) apply(simp add: permute_bool_def) apply(rule theI'[OF unique]) done subsubsection \Equivariance of Set operators\ lemma mem_eqvt [eqvt]: shows "p \ (x \ A) \ (p \ x) \ (p \ A)" unfolding permute_bool_def permute_set_def by (auto) lemma Collect_eqvt [eqvt]: shows "p \ {x. P x} = {x. (p \ P) x}" unfolding permute_set_eq permute_fun_def by (auto simp: permute_bool_def) lemma Bex_eqvt [eqvt]: shows "p \ (\x \ S. P x) = (\x \ (p \ S). (p \ P) x)" unfolding Bex_def by simp lemma Ball_eqvt [eqvt]: shows "p \ (\x \ S. P x) = (\x \ (p \ S). (p \ P) x)" unfolding Ball_def by simp lemma image_eqvt [eqvt]: shows "p \ (f ` A) = (p \ f) ` (p \ A)" unfolding image_def by simp lemma Image_eqvt [eqvt]: shows "p \ (R `` A) = (p \ R) `` (p \ A)" unfolding Image_def by simp lemma UNIV_eqvt [eqvt]: shows "p \ UNIV = UNIV" unfolding UNIV_def by (perm_simp) (rule refl) lemma inter_eqvt [eqvt]: shows "p \ (A \ B) = (p \ A) \ (p \ B)" unfolding Int_def by simp lemma Inter_eqvt [eqvt]: shows "p \ \S = \(p \ S)" unfolding Inter_eq by simp lemma union_eqvt [eqvt]: shows "p \ (A \ B) = (p \ A) \ (p \ B)" unfolding Un_def by simp lemma Union_eqvt [eqvt]: shows "p \ \A = \(p \ A)" unfolding Union_eq by perm_simp rule lemma Diff_eqvt [eqvt]: fixes A B :: "'a::pt set" shows "p \ (A - B) = (p \ A) - (p \ B)" unfolding set_diff_eq by simp lemma Compl_eqvt [eqvt]: fixes A :: "'a::pt set" shows "p \ (- A) = - (p \ A)" unfolding Compl_eq_Diff_UNIV by simp lemma subset_eqvt [eqvt]: shows "p \ (S \ T) \ (p \ S) \ (p \ T)" unfolding subset_eq by simp lemma psubset_eqvt [eqvt]: shows "p \ (S \ T) \ (p \ S) \ (p \ T)" unfolding psubset_eq by simp lemma vimage_eqvt [eqvt]: shows "p \ (f -` A) = (p \ f) -` (p \ A)" unfolding vimage_def by simp lemma foldr_eqvt[eqvt]: "p \ foldr f xs = foldr (p \ f) (p \ xs)" apply(induct xs) apply(simp_all) apply(perm_simp exclude: foldr) apply(simp) done (* FIXME: eqvt attribute *) lemma Sigma_eqvt: shows "(p \ (X \ Y)) = (p \ X) \ (p \ Y)" unfolding Sigma_def by (perm_simp) (rule refl) text \ In order to prove that lfp is equivariant we need two auxiliary classes which specify that (<=) and Inf are equivariant. Instances for bool and fun are given. \ class le_eqvt = pt + assumes le_eqvt [eqvt]: "p \ (x \ y) = ((p \ x) \ (p \ (y :: 'a :: {order, pt})))" class inf_eqvt = pt + assumes inf_eqvt [eqvt]: "p \ (Inf X) = Inf (p \ (X :: 'a :: {complete_lattice, pt} set))" instantiation bool :: le_eqvt begin instance apply standard unfolding le_bool_def apply(perm_simp) apply(rule refl) done end instantiation "fun" :: (pt, le_eqvt) le_eqvt begin instance apply standard unfolding le_fun_def apply(perm_simp) apply(rule refl) done end instantiation bool :: inf_eqvt begin instance apply standard unfolding Inf_bool_def apply(perm_simp) apply(rule refl) done end instantiation "fun" :: (pt, inf_eqvt) inf_eqvt begin instance apply standard unfolding Inf_fun_def apply(perm_simp) apply(rule refl) done end lemma lfp_eqvt [eqvt]: fixes F::"('a \ 'b) \ ('a::pt \ 'b::{inf_eqvt, le_eqvt})" shows "p \ (lfp F) = lfp (p \ F)" unfolding lfp_def by simp lemma finite_eqvt [eqvt]: shows "p \ finite A = finite (p \ A)" unfolding finite_def by simp lemma fun_upd_eqvt[eqvt]: shows "p \ (f(x := y)) = (p \ f)((p \ x) := (p \ y))" unfolding fun_upd_def by simp lemma comp_eqvt [eqvt]: shows "p \ (f \ g) = (p \ f) \ (p \ g)" unfolding comp_def by simp subsubsection \Equivariance for product operations\ lemma fst_eqvt [eqvt]: shows "p \ (fst x) = fst (p \ x)" by (cases x) simp lemma snd_eqvt [eqvt]: shows "p \ (snd x) = snd (p \ x)" by (cases x) simp lemma split_eqvt [eqvt]: shows "p \ (case_prod P x) = case_prod (p \ P) (p \ x)" unfolding split_def by simp subsubsection \Equivariance for list operations\ lemma append_eqvt [eqvt]: shows "p \ (xs @ ys) = (p \ xs) @ (p \ ys)" by (induct xs) auto lemma rev_eqvt [eqvt]: shows "p \ (rev xs) = rev (p \ xs)" by (induct xs) (simp_all add: append_eqvt) lemma map_eqvt [eqvt]: shows "p \ (map f xs) = map (p \ f) (p \ xs)" by (induct xs) (simp_all) lemma removeAll_eqvt [eqvt]: shows "p \ (removeAll x xs) = removeAll (p \ x) (p \ xs)" by (induct xs) (auto) lemma filter_eqvt [eqvt]: shows "p \ (filter f xs) = filter (p \ f) (p \ xs)" apply(induct xs) apply(simp) apply(simp only: filter.simps permute_list.simps if_eqvt) apply(simp only: permute_fun_app_eq) done lemma distinct_eqvt [eqvt]: shows "p \ (distinct xs) = distinct (p \ xs)" apply(induct xs) apply(simp add: permute_bool_def) apply(simp add: conj_eqvt Not_eqvt mem_eqvt set_eqvt) done lemma length_eqvt [eqvt]: shows "p \ (length xs) = length (p \ xs)" by (induct xs) (simp_all add: permute_pure) subsubsection \Equivariance for @{typ "'a option"}\ lemma map_option_eqvt[eqvt]: shows "p \ (map_option f x) = map_option (p \ f) (p \ x)" by (cases x) (simp_all) subsubsection \Equivariance for @{typ "'a fset"}\ context includes fset.lifting begin lemma in_fset_eqvt [eqvt]: shows "(p \ (x |\| S)) = ((p \ x) |\| (p \ S))" by transfer simp lemma union_fset_eqvt [eqvt]: shows "(p \ (S |\| T)) = ((p \ S) |\| (p \ T))" by (induct S) (simp_all) lemma inter_fset_eqvt [eqvt]: shows "(p \ (S |\| T)) = ((p \ S) |\| (p \ T))" by transfer simp lemma subset_fset_eqvt [eqvt]: shows "(p \ (S |\| T)) = ((p \ S) |\| (p \ T))" by transfer simp lemma map_fset_eqvt [eqvt]: shows "p \ (f |`| S) = (p \ f) |`| (p \ S)" by transfer simp end subsubsection \Equivariance for @{typ "('a, 'b) finfun"}\ lemma finfun_update_eqvt [eqvt]: shows "(p \ (finfun_update f a b)) = finfun_update (p \ f) (p \ a) (p \ b)" by (transfer) (simp) lemma finfun_const_eqvt [eqvt]: shows "(p \ (finfun_const b)) = finfun_const (p \ b)" by (transfer) (simp) lemma finfun_apply_eqvt [eqvt]: shows "(p \ (finfun_apply f b)) = finfun_apply (p \ f) (p \ b)" by (transfer) (simp) section \Supp, Freshness and Supports\ context pt begin definition supp :: "'a \ atom set" where "supp x = {a. infinite {b. (a \ b) \ x \ x}}" definition fresh :: "atom \ 'a \ bool" ("_ \ _" [55, 55] 55) where "a \ x \ a \ supp x" end lemma supp_conv_fresh: shows "supp x = {a. \ a \ x}" unfolding fresh_def by simp lemma swap_rel_trans: assumes "sort_of a = sort_of b" assumes "sort_of b = sort_of c" assumes "(a \ c) \ x = x" assumes "(b \ c) \ x = x" shows "(a \ b) \ x = x" proof (cases) assume "a = b \ c = b" with assms show "(a \ b) \ x = x" by auto next assume *: "\ (a = b \ c = b)" have "((a \ c) + (b \ c) + (a \ c)) \ x = x" using assms by simp also have "(a \ c) + (b \ c) + (a \ c) = (a \ b)" using assms * by (simp add: swap_triple) finally show "(a \ b) \ x = x" . qed lemma swap_fresh_fresh: assumes a: "a \ x" and b: "b \ x" shows "(a \ b) \ x = x" proof (cases) assume asm: "sort_of a = sort_of b" have "finite {c. (a \ c) \ x \ x}" "finite {c. (b \ c) \ x \ x}" using a b unfolding fresh_def supp_def by simp_all then have "finite ({c. (a \ c) \ x \ x} \ {c. (b \ c) \ x \ x})" by simp then obtain c where "(a \ c) \ x = x" "(b \ c) \ x = x" "sort_of c = sort_of b" by (rule obtain_atom) (auto) then show "(a \ b) \ x = x" using asm by (rule_tac swap_rel_trans) (simp_all) next assume "sort_of a \ sort_of b" then show "(a \ b) \ x = x" by simp qed subsection \supp and fresh are equivariant\ lemma supp_eqvt [eqvt]: shows "p \ (supp x) = supp (p \ x)" unfolding supp_def by simp lemma fresh_eqvt [eqvt]: shows "p \ (a \ x) = (p \ a) \ (p \ x)" unfolding fresh_def by simp lemma fresh_permute_iff: shows "(p \ a) \ (p \ x) \ a \ x" by (simp only: fresh_eqvt[symmetric] permute_bool_def) lemma fresh_permute_left: shows "a \ p \ x \ - p \ a \ x" proof assume "a \ p \ x" then have "- p \ a \ - p \ p \ x" by (simp only: fresh_permute_iff) then show "- p \ a \ x" by simp next assume "- p \ a \ x" then have "p \ - p \ a \ p \ x" by (simp only: fresh_permute_iff) then show "a \ p \ x" by simp qed section \supports\ definition supports :: "atom set \ 'a::pt \ bool" (infixl "supports" 80) where "S supports x \ \a b. (a \ S \ b \ S \ (a \ b) \ x = x)" lemma supp_is_subset: fixes S :: "atom set" and x :: "'a::pt" assumes a1: "S supports x" and a2: "finite S" shows "(supp x) \ S" proof (rule ccontr) assume "\ (supp x \ S)" then obtain a where b1: "a \ supp x" and b2: "a \ S" by auto from a1 b2 have "\b. b \ S \ (a \ b) \ x = x" unfolding supports_def by auto then have "{b. (a \ b) \ x \ x} \ S" by auto with a2 have "finite {b. (a \ b) \ x \ x}" by (simp add: finite_subset) then have "a \ (supp x)" unfolding supp_def by simp with b1 show False by simp qed lemma supports_finite: fixes S :: "atom set" and x :: "'a::pt" assumes a1: "S supports x" and a2: "finite S" shows "finite (supp x)" proof - have "(supp x) \ S" using a1 a2 by (rule supp_is_subset) then show "finite (supp x)" using a2 by (simp add: finite_subset) qed lemma supp_supports: fixes x :: "'a::pt" shows "(supp x) supports x" unfolding supports_def proof (intro strip) fix a b assume "a \ (supp x) \ b \ (supp x)" then have "a \ x" and "b \ x" by (simp_all add: fresh_def) then show "(a \ b) \ x = x" by (simp add: swap_fresh_fresh) qed lemma supports_fresh: fixes x :: "'a::pt" assumes a1: "S supports x" and a2: "finite S" and a3: "a \ S" shows "a \ x" unfolding fresh_def proof - have "(supp x) \ S" using a1 a2 by (rule supp_is_subset) then show "a \ (supp x)" using a3 by auto qed lemma supp_is_least_supports: fixes S :: "atom set" and x :: "'a::pt" assumes a1: "S supports x" and a2: "finite S" and a3: "\S'. finite S' \ (S' supports x) \ S \ S'" shows "(supp x) = S" proof (rule equalityI) show "(supp x) \ S" using a1 a2 by (rule supp_is_subset) with a2 have fin: "finite (supp x)" by (rule rev_finite_subset) have "(supp x) supports x" by (rule supp_supports) with fin a3 show "S \ supp x" by blast qed lemma subsetCI: shows "(\x. x \ A \ x \ B \ False) \ A \ B" by auto lemma finite_supp_unique: assumes a1: "S supports x" assumes a2: "finite S" assumes a3: "\a b. \a \ S; b \ S; sort_of a = sort_of b\ \ (a \ b) \ x \ x" shows "(supp x) = S" using a1 a2 proof (rule supp_is_least_supports) fix S' assume "finite S'" and "S' supports x" show "S \ S'" proof (rule subsetCI) fix a assume "a \ S" and "a \ S'" have "finite (S \ S')" using \finite S\ \finite S'\ by simp then obtain b where "b \ S \ S'" and "sort_of b = sort_of a" by (rule obtain_atom) then have "b \ S" and "b \ S'" and "sort_of a = sort_of b" by simp_all then have "(a \ b) \ x = x" using \a \ S'\ \S' supports x\ by (simp add: supports_def) moreover have "(a \ b) \ x \ x" using \a \ S\ \b \ S\ \sort_of a = sort_of b\ by (rule a3) ultimately show "False" by simp qed qed section \Support w.r.t. relations\ text \ This definition is used for unquotient types, where alpha-equivalence does not coincide with equality. \ definition "supp_rel R x = {a. infinite {b. \(R ((a \ b) \ x) x)}}" section \Finitely-supported types\ class fs = pt + assumes finite_supp: "finite (supp x)" lemma pure_supp: fixes x::"'a::pure" shows "supp x = {}" unfolding supp_def by (simp add: permute_pure) lemma pure_fresh: fixes x::"'a::pure" shows "a \ x" unfolding fresh_def by (simp add: pure_supp) instance pure < fs by standard (simp add: pure_supp) subsection \Type \<^typ>\atom\ is finitely-supported.\ lemma supp_atom: shows "supp a = {a}" apply (rule finite_supp_unique) apply (clarsimp simp add: supports_def) apply simp apply simp done lemma fresh_atom: shows "a \ b \ a \ b" unfolding fresh_def supp_atom by simp instance atom :: fs by standard (simp add: supp_atom) section \Type \<^typ>\perm\ is finitely-supported.\ lemma perm_swap_eq: shows "(a \ b) \ p = p \ (p \ (a \ b)) = (a \ b)" unfolding permute_perm_def by (metis add_diff_cancel minus_perm_def) lemma supports_perm: shows "{a. p \ a \ a} supports p" unfolding supports_def unfolding perm_swap_eq by (simp add: swap_eqvt) lemma finite_perm_lemma: shows "finite {a::atom. p \ a \ a}" using finite_Rep_perm [of p] unfolding permute_atom_def . lemma supp_perm: shows "supp p = {a. p \ a \ a}" apply (rule finite_supp_unique) apply (rule supports_perm) apply (rule finite_perm_lemma) apply (simp add: perm_swap_eq swap_eqvt) apply (auto simp: perm_eq_iff swap_atom) done lemma fresh_perm: shows "a \ p \ p \ a = a" unfolding fresh_def by (simp add: supp_perm) lemma supp_swap: shows "supp (a \ b) = (if a = b \ sort_of a \ sort_of b then {} else {a, b})" by (auto simp: supp_perm swap_atom) lemma fresh_swap: shows "a \ (b \ c) \ (sort_of b \ sort_of c) \ b = c \ (a \ b \ a \ c)" by (simp add: fresh_def supp_swap supp_atom) lemma fresh_zero_perm: shows "a \ (0::perm)" unfolding fresh_perm by simp lemma supp_zero_perm: shows "supp (0::perm) = {}" unfolding supp_perm by simp lemma fresh_plus_perm: fixes p q::perm assumes "a \ p" "a \ q" shows "a \ (p + q)" using assms unfolding fresh_def by (auto simp: supp_perm) lemma supp_plus_perm: fixes p q::perm shows "supp (p + q) \ supp p \ supp q" by (auto simp: supp_perm) lemma fresh_minus_perm: fixes p::perm shows "a \ (- p) \ a \ p" unfolding fresh_def unfolding supp_perm apply(simp) apply(metis permute_minus_cancel) done lemma supp_minus_perm: fixes p::perm shows "supp (- p) = supp p" unfolding supp_conv_fresh by (simp add: fresh_minus_perm) lemma plus_perm_eq: fixes p q::"perm" assumes asm: "supp p \ supp q = {}" shows "p + q = q + p" unfolding perm_eq_iff proof fix a::"atom" show "(p + q) \ a = (q + p) \ a" proof - { assume "a \ supp p" "a \ supp q" then have "(p + q) \ a = (q + p) \ a" by (simp add: supp_perm) } moreover { assume a: "a \ supp p" "a \ supp q" then have "p \ a \ supp p" by (simp add: supp_perm) then have "p \ a \ supp q" using asm by auto with a have "(p + q) \ a = (q + p) \ a" by (simp add: supp_perm) } moreover { assume a: "a \ supp p" "a \ supp q" then have "q \ a \ supp q" by (simp add: supp_perm) then have "q \ a \ supp p" using asm by auto with a have "(p + q) \ a = (q + p) \ a" by (simp add: supp_perm) } ultimately show "(p + q) \ a = (q + p) \ a" using asm by blast qed qed lemma supp_plus_perm_eq: fixes p q::perm assumes asm: "supp p \ supp q = {}" shows "supp (p + q) = supp p \ supp q" proof - { fix a::"atom" assume "a \ supp p" then have "a \ supp q" using asm by auto then have "a \ supp (p + q)" using \a \ supp p\ by (simp add: supp_perm) } moreover { fix a::"atom" assume "a \ supp q" then have "a \ supp p" using asm by auto then have "a \ supp (q + p)" using \a \ supp q\ by (simp add: supp_perm) then have "a \ supp (p + q)" using asm plus_perm_eq by metis } ultimately have "supp p \ supp q \ supp (p + q)" by blast then show "supp (p + q) = supp p \ supp q" using supp_plus_perm by blast qed lemma perm_eq_iff2: fixes p q :: "perm" shows "p = q \ (\a::atom \ supp p \ supp q. p \ a = q \ a)" unfolding perm_eq_iff apply(auto) apply(case_tac "a \ p \ a \ q") apply(simp add: fresh_perm) apply(simp add: fresh_def) done instance perm :: fs by standard (simp add: supp_perm finite_perm_lemma) section \Finite Support instances for other types\ subsection \Type @{typ "'a \ 'b"} is finitely-supported.\ lemma supp_Pair: shows "supp (x, y) = supp x \ supp y" by (simp add: supp_def Collect_imp_eq Collect_neg_eq) lemma fresh_Pair: shows "a \ (x, y) \ a \ x \ a \ y" by (simp add: fresh_def supp_Pair) lemma supp_Unit: shows "supp () = {}" by (simp add: supp_def) lemma fresh_Unit: shows "a \ ()" by (simp add: fresh_def supp_Unit) instance prod :: (fs, fs) fs apply standard apply (case_tac x) apply (simp add: supp_Pair finite_supp) done subsection \Type @{typ "'a + 'b"} is finitely supported\ lemma supp_Inl: shows "supp (Inl x) = supp x" by (simp add: supp_def) lemma supp_Inr: shows "supp (Inr x) = supp x" by (simp add: supp_def) lemma fresh_Inl: shows "a \ Inl x \ a \ x" by (simp add: fresh_def supp_Inl) lemma fresh_Inr: shows "a \ Inr y \ a \ y" by (simp add: fresh_def supp_Inr) instance sum :: (fs, fs) fs apply standard apply (case_tac x) apply (simp_all add: supp_Inl supp_Inr finite_supp) done subsection \Type @{typ "'a option"} is finitely supported\ lemma supp_None: shows "supp None = {}" by (simp add: supp_def) lemma supp_Some: shows "supp (Some x) = supp x" by (simp add: supp_def) lemma fresh_None: shows "a \ None" by (simp add: fresh_def supp_None) lemma fresh_Some: shows "a \ Some x \ a \ x" by (simp add: fresh_def supp_Some) instance option :: (fs) fs apply standard apply (induct_tac x) apply (simp_all add: supp_None supp_Some finite_supp) done subsubsection \Type @{typ "'a list"} is finitely supported\ lemma supp_Nil: shows "supp [] = {}" by (simp add: supp_def) lemma fresh_Nil: shows "a \ []" by (simp add: fresh_def supp_Nil) lemma supp_Cons: shows "supp (x # xs) = supp x \ supp xs" by (simp add: supp_def Collect_imp_eq Collect_neg_eq) lemma fresh_Cons: shows "a \ (x # xs) \ a \ x \ a \ xs" by (simp add: fresh_def supp_Cons) lemma supp_append: shows "supp (xs @ ys) = supp xs \ supp ys" by (induct xs) (auto simp: supp_Nil supp_Cons) lemma fresh_append: shows "a \ (xs @ ys) \ a \ xs \ a \ ys" by (induct xs) (simp_all add: fresh_Nil fresh_Cons) lemma supp_rev: shows "supp (rev xs) = supp xs" by (induct xs) (auto simp: supp_append supp_Cons supp_Nil) lemma fresh_rev: shows "a \ rev xs \ a \ xs" by (induct xs) (auto simp: fresh_append fresh_Cons fresh_Nil) lemma supp_removeAll: fixes x::"atom" shows "supp (removeAll x xs) = supp xs - {x}" by (induct xs) (auto simp: supp_Nil supp_Cons supp_atom) lemma supp_of_atom_list: fixes as::"atom list" shows "supp as = set as" by (induct as) (simp_all add: supp_Nil supp_Cons supp_atom) instance list :: (fs) fs apply standard apply (induct_tac x) apply (simp_all add: supp_Nil supp_Cons finite_supp) done section \Support and Freshness for Applications\ lemma fresh_conv_MOST: shows "a \ x \ (MOST b. (a \ b) \ x = x)" unfolding fresh_def supp_def unfolding MOST_iff_cofinite by simp lemma fresh_fun_app: assumes "a \ f" and "a \ x" shows "a \ f x" using assms unfolding fresh_conv_MOST unfolding permute_fun_app_eq by (elim MOST_rev_mp) (simp) lemma supp_fun_app: shows "supp (f x) \ (supp f) \ (supp x)" using fresh_fun_app unfolding fresh_def by auto subsection \Equivariance Predicate \eqvt\ and \eqvt_at\\ definition "eqvt f \ \p. p \ f = f" lemma eqvt_boolI: fixes f::"bool" shows "eqvt f" unfolding eqvt_def by (simp add: permute_bool_def) text \equivariance of a function at a given argument\ definition "eqvt_at f x \ \p. p \ (f x) = f (p \ x)" lemma eqvtI: shows "(\p. p \ f \ f) \ eqvt f" unfolding eqvt_def by simp lemma eqvt_at_perm: assumes "eqvt_at f x" shows "eqvt_at f (q \ x)" proof - { fix p::"perm" have "p \ (f (q \ x)) = p \ q \ (f x)" using assms by (simp add: eqvt_at_def) also have "\ = (p + q) \ (f x)" by simp also have "\ = f ((p + q) \ x)" using assms by (simp only: eqvt_at_def) finally have "p \ (f (q \ x)) = f (p \ q \ x)" by simp } then show "eqvt_at f (q \ x)" unfolding eqvt_at_def by simp qed lemma supp_fun_eqvt: assumes a: "eqvt f" shows "supp f = {}" using a unfolding eqvt_def unfolding supp_def by simp lemma fresh_fun_eqvt: assumes a: "eqvt f" shows "a \ f" using a unfolding fresh_def by (simp add: supp_fun_eqvt) lemma fresh_fun_eqvt_app: assumes a: "eqvt f" shows "a \ x \ a \ f x" proof - from a have "supp f = {}" by (simp add: supp_fun_eqvt) then show "a \ x \ a \ f x" unfolding fresh_def using supp_fun_app by auto qed lemma supp_fun_app_eqvt: assumes a: "eqvt f" shows "supp (f x) \ supp x" using fresh_fun_eqvt_app[OF a] unfolding fresh_def by auto lemma supp_eqvt_at: assumes asm: "eqvt_at f x" and fin: "finite (supp x)" shows "supp (f x) \ supp x" apply(rule supp_is_subset) unfolding supports_def unfolding fresh_def[symmetric] using asm apply(simp add: eqvt_at_def) apply(simp add: swap_fresh_fresh) apply(rule fin) done lemma finite_supp_eqvt_at: assumes asm: "eqvt_at f x" and fin: "finite (supp x)" shows "finite (supp (f x))" apply(rule finite_subset) apply(rule supp_eqvt_at[OF asm fin]) apply(rule fin) done lemma fresh_eqvt_at: assumes asm: "eqvt_at f x" and fin: "finite (supp x)" and fresh: "a \ x" shows "a \ f x" using fresh unfolding fresh_def using supp_eqvt_at[OF asm fin] by auto text \for handling of freshness of functions\ simproc_setup fresh_fun_simproc ("a \ (f::'a::pt \'b::pt)") = \fn _ => fn ctxt => fn ctrm => let val _ $ _ $ f = Thm.term_of ctrm in case (Term.add_frees f [], Term.add_vars f []) of ([], []) => SOME(@{thm fresh_fun_eqvt[simplified eqvt_def, THEN Eq_TrueI]}) | (x::_, []) => let val argx = Free x val absf = absfree x f val cty_inst = [SOME (Thm.ctyp_of ctxt (fastype_of argx)), SOME (Thm.ctyp_of ctxt (fastype_of f))] val ctrm_inst = [NONE, SOME (Thm.cterm_of ctxt absf), SOME (Thm.cterm_of ctxt argx)] val thm = Thm.instantiate' cty_inst ctrm_inst @{thm fresh_fun_app} in SOME(thm RS @{thm Eq_TrueI}) end | (_, _) => NONE end \ subsection \helper functions for \nominal_functions\\ lemma THE_defaultI2: assumes "\!x. P x" "\x. P x \ Q x" shows "Q (THE_default d P)" by (iprover intro: assms THE_defaultI') lemma the_default_eqvt: assumes unique: "\!x. P x" shows "(p \ (THE_default d P)) = (THE_default (p \ d) (p \ P))" apply(rule THE_default1_equality [symmetric]) apply(rule_tac p="-p" in permute_boolE) apply(simp add: ex1_eqvt) apply(rule unique) apply(rule_tac p="-p" in permute_boolE) apply(rule subst[OF permute_fun_app_eq]) apply(simp) apply(rule THE_defaultI'[OF unique]) done lemma fundef_ex1_eqvt: fixes x::"'a::pt" assumes f_def: "f == (\x::'a. THE_default (d x) (G x))" assumes eqvt: "eqvt G" assumes ex1: "\!y. G x y" shows "(p \ (f x)) = f (p \ x)" apply(simp only: f_def) apply(subst the_default_eqvt) apply(rule ex1) apply(rule THE_default1_equality [symmetric]) apply(rule_tac p="-p" in permute_boolE) apply(perm_simp add: permute_minus_cancel) using eqvt[simplified eqvt_def] apply(simp) apply(rule ex1) apply(rule THE_defaultI2) apply(rule_tac p="-p" in permute_boolE) apply(perm_simp add: permute_minus_cancel) apply(rule ex1) apply(perm_simp) using eqvt[simplified eqvt_def] apply(simp) done lemma fundef_ex1_eqvt_at: fixes x::"'a::pt" assumes f_def: "f == (\x::'a. THE_default (d x) (G x))" assumes eqvt: "eqvt G" assumes ex1: "\!y. G x y" shows "eqvt_at f x" unfolding eqvt_at_def using assms by (auto intro: fundef_ex1_eqvt) lemma fundef_ex1_prop: fixes x::"'a::pt" assumes f_def: "f == (\x::'a. THE_default (d x) (G x))" assumes P_all: "\x y. G x y \ P x y" assumes ex1: "\!y. G x y" shows "P x (f x)" unfolding f_def using ex1 apply(erule_tac ex1E) apply(rule THE_defaultI2) apply(blast) apply(rule P_all) apply(assumption) done section \Support of Finite Sets of Finitely Supported Elements\ text \support and freshness for atom sets\ lemma supp_finite_atom_set: fixes S::"atom set" assumes "finite S" shows "supp S = S" apply(rule finite_supp_unique) apply(simp add: supports_def) apply(simp add: swap_set_not_in) apply(rule assms) apply(simp add: swap_set_in) done lemma supp_cofinite_atom_set: fixes S::"atom set" assumes "finite (UNIV - S)" shows "supp S = (UNIV - S)" apply(rule finite_supp_unique) apply(simp add: supports_def) apply(simp add: swap_set_both_in) apply(rule assms) apply(subst swap_commute) apply(simp add: swap_set_in) done lemma fresh_finite_atom_set: fixes S::"atom set" assumes "finite S" shows "a \ S \ a \ S" unfolding fresh_def by (simp add: supp_finite_atom_set[OF assms]) lemma fresh_minus_atom_set: fixes S::"atom set" assumes "finite S" shows "a \ S - T \ (a \ T \ a \ S)" unfolding fresh_def by (auto simp: supp_finite_atom_set assms) lemma Union_supports_set: shows "(\x \ S. supp x) supports S" proof - { fix a b have "\x \ S. (a \ b) \ x = x \ (a \ b) \ S = S" unfolding permute_set_def by force } then show "(\x \ S. supp x) supports S" unfolding supports_def by (simp add: fresh_def[symmetric] swap_fresh_fresh) qed lemma Union_of_finite_supp_sets: fixes S::"('a::fs set)" assumes fin: "finite S" shows "finite (\x\S. supp x)" using fin by (induct) (auto simp: finite_supp) lemma Union_included_in_supp: fixes S::"('a::fs set)" assumes fin: "finite S" shows "(\x\S. supp x) \ supp S" proof - have eqvt: "eqvt (\S. \x \ S. supp x)" unfolding eqvt_def by simp have "(\x\S. supp x) = supp (\x\S. supp x)" by (rule supp_finite_atom_set[symmetric]) (rule Union_of_finite_supp_sets[OF fin]) also have "\ \ supp S" using eqvt by (rule supp_fun_app_eqvt) finally show "(\x\S. supp x) \ supp S" . qed lemma supp_of_finite_sets: fixes S::"('a::fs set)" assumes fin: "finite S" shows "(supp S) = (\x\S. supp x)" apply(rule subset_antisym) apply(rule supp_is_subset) apply(rule Union_supports_set) apply(rule Union_of_finite_supp_sets[OF fin]) apply(rule Union_included_in_supp[OF fin]) done lemma finite_sets_supp: fixes S::"('a::fs set)" assumes "finite S" shows "finite (supp S)" using assms by (simp only: supp_of_finite_sets Union_of_finite_supp_sets) lemma supp_of_finite_union: fixes S T::"('a::fs) set" assumes fin1: "finite S" and fin2: "finite T" shows "supp (S \ T) = supp S \ supp T" using fin1 fin2 by (simp add: supp_of_finite_sets) lemma fresh_finite_union: fixes S T::"('a::fs) set" assumes fin1: "finite S" and fin2: "finite T" shows "a \ (S \ T) \ a \ S \ a \ T" unfolding fresh_def by (simp add: supp_of_finite_union[OF fin1 fin2]) lemma supp_of_finite_insert: fixes S::"('a::fs) set" assumes fin: "finite S" shows "supp (insert x S) = supp x \ supp S" using fin by (simp add: supp_of_finite_sets) lemma fresh_finite_insert: fixes S::"('a::fs) set" assumes fin: "finite S" shows "a \ (insert x S) \ a \ x \ a \ S" using fin unfolding fresh_def by (simp add: supp_of_finite_insert) lemma supp_set_empty: shows "supp {} = {}" unfolding supp_def by (simp add: empty_eqvt) lemma fresh_set_empty: shows "a \ {}" by (simp add: fresh_def supp_set_empty) lemma supp_set: fixes xs :: "('a::fs) list" shows "supp (set xs) = supp xs" apply(induct xs) apply(simp add: supp_set_empty supp_Nil) apply(simp add: supp_Cons supp_of_finite_insert) done lemma fresh_set: fixes xs :: "('a::fs) list" shows "a \ (set xs) \ a \ xs" unfolding fresh_def by (simp add: supp_set) subsection \Type @{typ "'a multiset"} is finitely supported\ lemma set_mset_eqvt [eqvt]: shows "p \ (set_mset M) = set_mset (p \ M)" by (induct M) (simp_all add: insert_eqvt empty_eqvt) lemma supp_set_mset: shows "supp (set_mset M) \ supp M" apply (rule supp_fun_app_eqvt) unfolding eqvt_def apply(perm_simp) apply(simp) done lemma Union_finite_multiset: fixes M::"'a::fs multiset" shows "finite (\{supp x | x. x \# M})" proof - have "finite (\(supp ` {x. x \# M}))" by (induct M) (simp_all add: Collect_imp_eq Collect_neg_eq finite_supp) then show "finite (\{supp x | x. x \# M})" by (simp only: image_Collect) qed lemma Union_supports_multiset: shows "\{supp x | x. x \# M} supports M" proof - have sw: "\a b. ((\x. x \# M \ (a \ b) \ x = x) \ (a \ b) \ M = M)" unfolding permute_multiset_def by (induct M) simp_all have "(\x\set_mset M. supp x) supports M" by (auto intro!: sw swap_fresh_fresh simp add: fresh_def supports_def) also have "(\x\set_mset M. supp x) = (\{supp x | x. x \# M})" by auto finally show "(\{supp x | x. x \# M}) supports M" . qed lemma Union_included_multiset: fixes M::"('a::fs multiset)" shows "(\{supp x | x. x \# M}) \ supp M" proof - have "(\{supp x | x. x \# M}) = (\x \ set_mset M. supp x)" by auto also have "... = supp (set_mset M)" by (simp add: supp_of_finite_sets) also have " ... \ supp M" by (rule supp_set_mset) finally show "(\{supp x | x. x \# M}) \ supp M" . qed lemma supp_of_multisets: fixes M::"('a::fs multiset)" shows "(supp M) = (\{supp x | x. x \# M})" apply(rule subset_antisym) apply(rule supp_is_subset) apply(rule Union_supports_multiset) apply(rule Union_finite_multiset) apply(rule Union_included_multiset) done lemma multisets_supp_finite: fixes M::"('a::fs multiset)" shows "finite (supp M)" by (simp only: supp_of_multisets Union_finite_multiset) lemma supp_of_multiset_union: fixes M N::"('a::fs) multiset" shows "supp (M + N) = supp M \ supp N" by (auto simp: supp_of_multisets) lemma supp_empty_mset [simp]: shows "supp {#} = {}" unfolding supp_def by simp instance multiset :: (fs) fs by standard (rule multisets_supp_finite) subsection \Type @{typ "'a fset"} is finitely supported\ lemma supp_fset [simp]: shows "supp (fset S) = supp S" unfolding supp_def by (simp add: fset_eqvt fset_cong) lemma supp_empty_fset [simp]: shows "supp {||} = {}" unfolding supp_def by simp lemma fresh_empty_fset: shows "a \ {||}" unfolding fresh_def by (simp) lemma supp_finsert [simp]: fixes x::"'a::fs" and S::"'a fset" shows "supp (finsert x S) = supp x \ supp S" apply(subst supp_fset[symmetric]) apply(simp add: supp_of_finite_insert) done lemma fresh_finsert: fixes x::"'a::fs" and S::"'a fset" shows "a \ finsert x S \ a \ x \ a \ S" unfolding fresh_def by simp lemma fset_finite_supp: fixes S::"('a::fs) fset" shows "finite (supp S)" by (induct S) (simp_all add: finite_supp) lemma supp_union_fset: fixes S T::"'a::fs fset" shows "supp (S |\| T) = supp S \ supp T" by (induct S) (auto) lemma fresh_union_fset: fixes S T::"'a::fs fset" shows "a \ S |\| T \ a \ S \ a \ T" unfolding fresh_def by (simp add: supp_union_fset) instance fset :: (fs) fs by standard (rule fset_finite_supp) subsection \Type @{typ "('a, 'b) finfun"} is finitely supported\ lemma fresh_finfun_const: shows "a \ (finfun_const b) \ a \ b" by (simp add: fresh_def supp_def) lemma fresh_finfun_update: shows "\a \ f; a \ x; a \ y\ \ a \ finfun_update f x y" unfolding fresh_conv_MOST unfolding finfun_update_eqvt by (elim MOST_rev_mp) (simp) lemma supp_finfun_const: shows "supp (finfun_const b) = supp(b)" by (simp add: supp_def) lemma supp_finfun_update: shows "supp (finfun_update f x y) \ supp(f, x, y)" using fresh_finfun_update by (auto simp: fresh_def supp_Pair) instance finfun :: (fs, fs) fs apply standard apply(induct_tac x rule: finfun_weak_induct) apply(simp add: supp_finfun_const finite_supp) apply(rule finite_subset) apply(rule supp_finfun_update) apply(simp add: supp_Pair finite_supp) done section \Freshness and Fresh-Star\ lemma fresh_Unit_elim: shows "(a \ () \ PROP C) \ PROP C" by (simp add: fresh_Unit) lemma fresh_Pair_elim: shows "(a \ (x, y) \ PROP C) \ (a \ x \ a \ y \ PROP C)" by rule (simp_all add: fresh_Pair) (* this rule needs to be added before the fresh_prodD is *) (* added to the simplifier with mksimps *) lemma [simp]: shows "a \ x1 \ a \ x2 \ a \ (x1, x2)" by (simp add: fresh_Pair) lemma fresh_PairD: shows "a \ (x, y) \ a \ x" and "a \ (x, y) \ a \ y" by (simp_all add: fresh_Pair) declaration \fn _ => let val mksimps_pairs = (@{const_name Nominal2_Base.fresh}, @{thms fresh_PairD}) :: mksimps_pairs in Simplifier.map_ss (fn ss => Simplifier.set_mksimps (mksimps mksimps_pairs) ss) end \ text \The fresh-star generalisation of fresh is used in strong induction principles.\ definition fresh_star :: "atom set \ 'a::pt \ bool" ("_ \* _" [80,80] 80) where "as \* x \ \a \ as. a \ x" lemma fresh_star_supp_conv: shows "supp x \* y \ supp y \* x" by (auto simp: fresh_star_def fresh_def) lemma fresh_star_perm_set_conv: fixes p::"perm" assumes fresh: "as \* p" and fin: "finite as" shows "supp p \* as" apply(rule fresh_star_supp_conv) apply(simp add: supp_finite_atom_set fin fresh) done lemma fresh_star_atom_set_conv: assumes fresh: "as \* bs" and fin: "finite as" "finite bs" shows "bs \* as" using fresh unfolding fresh_star_def fresh_def by (auto simp: supp_finite_atom_set fin) lemma atom_fresh_star_disjoint: assumes fin: "finite bs" shows "as \* bs \ (as \ bs = {})" unfolding fresh_star_def fresh_def by (auto simp: supp_finite_atom_set fin) lemma fresh_star_Pair: shows "as \* (x, y) = (as \* x \ as \* y)" by (auto simp: fresh_star_def fresh_Pair) lemma fresh_star_list: shows "as \* (xs @ ys) \ as \* xs \ as \* ys" and "as \* (x # xs) \ as \* x \ as \* xs" and "as \* []" by (auto simp: fresh_star_def fresh_Nil fresh_Cons fresh_append) lemma fresh_star_set: fixes xs::"('a::fs) list" shows "as \* set xs \ as \* xs" unfolding fresh_star_def by (simp add: fresh_set) lemma fresh_star_singleton: fixes a::"atom" shows "as \* {a} \ as \* a" by (simp add: fresh_star_def fresh_finite_insert fresh_set_empty) lemma fresh_star_fset: fixes xs::"('a::fs) list" shows "as \* fset S \ as \* S" by (simp add: fresh_star_def fresh_def) lemma fresh_star_Un: shows "(as \ bs) \* x = (as \* x \ bs \* x)" by (auto simp: fresh_star_def) lemma fresh_star_insert: shows "(insert a as) \* x = (a \ x \ as \* x)" by (auto simp: fresh_star_def) lemma fresh_star_Un_elim: "((as \ bs) \* x \ PROP C) \ (as \* x \ bs \* x \ PROP C)" unfolding fresh_star_def apply(rule) apply(erule meta_mp) apply(auto) done lemma fresh_star_insert_elim: "(insert a as \* x \ PROP C) \ (a \ x \ as \* x \ PROP C)" unfolding fresh_star_def by rule (simp_all add: fresh_star_def) lemma fresh_star_empty_elim: "({} \* x \ PROP C) \ PROP C" by (simp add: fresh_star_def) lemma fresh_star_Unit_elim: shows "(a \* () \ PROP C) \ PROP C" by (simp add: fresh_star_def fresh_Unit) lemma fresh_star_Pair_elim: shows "(a \* (x, y) \ PROP C) \ (a \* x \ a \* y \ PROP C)" by (rule, simp_all add: fresh_star_Pair) lemma fresh_star_zero: shows "as \* (0::perm)" unfolding fresh_star_def by (simp add: fresh_zero_perm) lemma fresh_star_plus: fixes p q::perm shows "\a \* p; a \* q\ \ a \* (p + q)" unfolding fresh_star_def by (simp add: fresh_plus_perm) lemma fresh_star_permute_iff: shows "(p \ a) \* (p \ x) \ a \* x" unfolding fresh_star_def by (metis mem_permute_iff permute_minus_cancel(1) fresh_permute_iff) lemma fresh_star_eqvt [eqvt]: shows "p \ (as \* x) \ (p \ as) \* (p \ x)" unfolding fresh_star_def by simp section \Induction principle for permutations\ lemma smaller_supp: assumes a: "a \ supp p" shows "supp ((p \ a \ a) + p) \ supp p" proof - have "supp ((p \ a \ a) + p) \ supp p" unfolding supp_perm by (auto simp: swap_atom) moreover have "a \ supp ((p \ a \ a) + p)" by (simp add: supp_perm) then have "supp ((p \ a \ a) + p) \ supp p" using a by auto ultimately show "supp ((p \ a \ a) + p) \ supp p" by auto qed lemma perm_struct_induct[consumes 1, case_names zero swap]: assumes S: "supp p \ S" and zero: "P 0" and swap: "\p a b. \P p; supp p \ S; a \ S; b \ S; a \ b; sort_of a = sort_of b\ \ P ((a \ b) + p)" shows "P p" proof - have "finite (supp p)" by (simp add: finite_supp) then show "P p" using S proof(induct A\"supp p" arbitrary: p rule: finite_psubset_induct) case (psubset p) then have ih: "\q. supp q \ supp p \ P q" by auto have as: "supp p \ S" by fact { assume "supp p = {}" then have "p = 0" by (simp add: supp_perm perm_eq_iff) then have "P p" using zero by simp } moreover { assume "supp p \ {}" then obtain a where a0: "a \ supp p" by blast then have a1: "p \ a \ S" "a \ S" "sort_of (p \ a) = sort_of a" "p \ a \ a" using as by (auto simp: supp_atom supp_perm swap_atom) let ?q = "(p \ a \ a) + p" have a2: "supp ?q \ supp p" using a0 smaller_supp by simp then have "P ?q" using ih by simp moreover have "supp ?q \ S" using as a2 by simp ultimately have "P ((p \ a \ a) + ?q)" using as a1 swap by simp moreover have "p = (p \ a \ a) + ?q" by (simp add: perm_eq_iff) ultimately have "P p" by simp } ultimately show "P p" by blast qed qed lemma perm_simple_struct_induct[case_names zero swap]: assumes zero: "P 0" and swap: "\p a b. \P p; a \ b; sort_of a = sort_of b\ \ P ((a \ b) + p)" shows "P p" by (rule_tac S="supp p" in perm_struct_induct) (auto intro: zero swap) lemma perm_struct_induct2[consumes 1, case_names zero swap plus]: assumes S: "supp p \ S" assumes zero: "P 0" assumes swap: "\a b. \sort_of a = sort_of b; a \ b; a \ S; b \ S\ \ P (a \ b)" assumes plus: "\p1 p2. \P p1; P p2; supp p1 \ S; supp p2 \ S\ \ P (p1 + p2)" shows "P p" using S by (induct p rule: perm_struct_induct) (auto intro: zero plus swap simp add: supp_swap) lemma perm_simple_struct_induct2[case_names zero swap plus]: assumes zero: "P 0" assumes swap: "\a b. \sort_of a = sort_of b; a \ b\ \ P (a \ b)" assumes plus: "\p1 p2. \P p1; P p2\ \ P (p1 + p2)" shows "P p" by (rule_tac S="supp p" in perm_struct_induct2) (auto intro: zero swap plus) lemma supp_perm_singleton: fixes p::"perm" shows "supp p \ {b} \ p = 0" proof - { assume "supp p \ {b}" then have "p = 0" by (induct p rule: perm_struct_induct) (simp_all) } then show "supp p \ {b} \ p = 0" by (auto simp: supp_zero_perm) qed lemma supp_perm_pair: fixes p::"perm" shows "supp p \ {a, b} \ p = 0 \ p = (b \ a)" proof - { assume "supp p \ {a, b}" then have "p = 0 \ p = (b \ a)" apply (induct p rule: perm_struct_induct) apply (auto simp: swap_cancel supp_zero_perm supp_swap) apply (simp add: swap_commute) done } then show "supp p \ {a, b} \ p = 0 \ p = (b \ a)" by (auto simp: supp_zero_perm supp_swap split: if_splits) qed lemma supp_perm_eq: assumes "(supp x) \* p" shows "p \ x = x" proof - from assms have "supp p \ {a. a \ x}" unfolding supp_perm fresh_star_def fresh_def by auto then show "p \ x = x" proof (induct p rule: perm_struct_induct) case zero show "0 \ x = x" by simp next case (swap p a b) then have "a \ x" "b \ x" "p \ x = x" by simp_all then show "((a \ b) + p) \ x = x" by (simp add: swap_fresh_fresh) qed qed text \same lemma as above, but proved with a different induction principle\ lemma supp_perm_eq_test: assumes "(supp x) \* p" shows "p \ x = x" proof - from assms have "supp p \ {a. a \ x}" unfolding supp_perm fresh_star_def fresh_def by auto then show "p \ x = x" proof (induct p rule: perm_struct_induct2) case zero show "0 \ x = x" by simp next case (swap a b) then have "a \ x" "b \ x" by simp_all then show "(a \ b) \ x = x" by (simp add: swap_fresh_fresh) next case (plus p1 p2) have "p1 \ x = x" "p2 \ x = x" by fact+ then show "(p1 + p2) \ x = x" by simp qed qed lemma perm_supp_eq: assumes a: "(supp p) \* x" shows "p \ x = x" proof - from assms have "supp p \ {a. a \ x}" unfolding supp_perm fresh_star_def fresh_def by auto then show "p \ x = x" proof (induct p rule: perm_struct_induct2) case zero show "0 \ x = x" by simp next case (swap a b) then have "a \ x" "b \ x" by simp_all then show "(a \ b) \ x = x" by (simp add: swap_fresh_fresh) next case (plus p1 p2) have "p1 \ x = x" "p2 \ x = x" by fact+ then show "(p1 + p2) \ x = x" by simp qed qed lemma supp_perm_perm_eq: assumes a: "\a \ supp x. p \ a = q \ a" shows "p \ x = q \ x" proof - from a have "\a \ supp x. (-q + p) \ a = a" by simp then have "\a \ supp x. a \ supp (-q + p)" unfolding supp_perm by simp then have "supp x \* (-q + p)" unfolding fresh_star_def fresh_def by simp then have "(-q + p) \ x = x" by (simp only: supp_perm_eq) then show "p \ x = q \ x" by (metis permute_minus_cancel permute_plus) qed text \disagreement set\ definition dset :: "perm \ perm \ atom set" where "dset p q = {a::atom. p \ a \ q \ a}" lemma ds_fresh: assumes "dset p q \* x" shows "p \ x = q \ x" using assms unfolding dset_def fresh_star_def fresh_def by (auto intro: supp_perm_perm_eq) lemma atom_set_perm_eq: assumes a: "as \* p" shows "p \ as = as" proof - from a have "supp p \ {a. a \ as}" unfolding supp_perm fresh_star_def fresh_def by auto then show "p \ as = as" proof (induct p rule: perm_struct_induct) case zero show "0 \ as = as" by simp next case (swap p a b) then have "a \ as" "b \ as" "p \ as = as" by simp_all then show "((a \ b) + p) \ as = as" by (simp add: swap_set_not_in) qed qed section \Avoiding of atom sets\ text \ For every set of atoms, there is another set of atoms avoiding a finitely supported c and there is a permutation which 'translates' between both sets. \ lemma at_set_avoiding_aux: fixes Xs::"atom set" and As::"atom set" assumes b: "Xs \ As" and c: "finite As" shows "\p. (p \ Xs) \ As = {} \ (supp p) = (Xs \ (p \ Xs))" proof - from b c have "finite Xs" by (rule finite_subset) then show ?thesis using b proof (induct rule: finite_subset_induct) case empty have "0 \ {} \ As = {}" by simp moreover have "supp (0::perm) = {} \ 0 \ {}" by (simp add: supp_zero_perm) ultimately show ?case by blast next case (insert x Xs) then obtain p where p1: "(p \ Xs) \ As = {}" and p2: "supp p = (Xs \ (p \ Xs))" by blast from \x \ As\ p1 have "x \ p \ Xs" by fast with \x \ Xs\ p2 have "x \ supp p" by fast hence px: "p \ x = x" unfolding supp_perm by simp have "finite (As \ p \ Xs \ supp p)" using \finite As\ \finite Xs\ by (simp add: permute_set_eq_image finite_supp) then obtain y where "y \ (As \ p \ Xs \ supp p)" "sort_of y = sort_of x" by (rule obtain_atom) hence y: "y \ As" "y \ p \ Xs" "y \ supp p" "sort_of y = sort_of x" by simp_all hence py: "p \ y = y" "x \ y" using \x \ As\ by (auto simp: supp_perm) let ?q = "(x \ y) + p" have q: "?q \ insert x Xs = insert y (p \ Xs)" unfolding insert_eqvt using \p \ x = x\ \sort_of y = sort_of x\ using \x \ p \ Xs\ \y \ p \ Xs\ by (simp add: swap_atom swap_set_not_in) have "?q \ insert x Xs \ As = {}" using \y \ As\ \p \ Xs \ As = {}\ unfolding q by simp moreover have "supp (x \ y) \ supp p = {}" using px py \sort_of y = sort_of x\ unfolding supp_swap by (simp add: supp_perm) then have "supp ?q = (supp (x \ y) \ supp p)" by (simp add: supp_plus_perm_eq) then have "supp ?q = insert x Xs \ ?q \ insert x Xs" using p2 \sort_of y = sort_of x\ \x \ y\ unfolding q supp_swap by auto ultimately show ?case by blast qed qed lemma at_set_avoiding: assumes a: "finite Xs" and b: "finite (supp c)" obtains p::"perm" where "(p \ Xs)\*c" and "(supp p) = (Xs \ (p \ Xs))" using a b at_set_avoiding_aux [where Xs="Xs" and As="Xs \ supp c"] unfolding fresh_star_def fresh_def by blast lemma at_set_avoiding1: assumes "finite xs" and "finite (supp c)" shows "\p. (p \ xs) \* c" using assms apply(erule_tac c="c" in at_set_avoiding) apply(auto) done lemma at_set_avoiding2: assumes "finite xs" and "finite (supp c)" "finite (supp x)" and "xs \* x" shows "\p. (p \ xs) \* c \ supp x \* p" using assms apply(erule_tac c="(c, x)" in at_set_avoiding) apply(simp add: supp_Pair) apply(rule_tac x="p" in exI) apply(simp add: fresh_star_Pair) apply(rule fresh_star_supp_conv) apply(auto simp: fresh_star_def) done lemma at_set_avoiding3: assumes "finite xs" and "finite (supp c)" "finite (supp x)" and "xs \* x" shows "\p. (p \ xs) \* c \ supp x \* p \ supp p = xs \ (p \ xs)" using assms apply(erule_tac c="(c, x)" in at_set_avoiding) apply(simp add: supp_Pair) apply(rule_tac x="p" in exI) apply(simp add: fresh_star_Pair) apply(rule fresh_star_supp_conv) apply(auto simp: fresh_star_def) done lemma at_set_avoiding2_atom: assumes "finite (supp c)" "finite (supp x)" and b: "a \ x" shows "\p. (p \ a) \ c \ supp x \* p" proof - have a: "{a} \* x" unfolding fresh_star_def by (simp add: b) obtain p where p1: "(p \ {a}) \* c" and p2: "supp x \* p" using at_set_avoiding2[of "{a}" "c" "x"] assms a by blast have c: "(p \ a) \ c" using p1 unfolding fresh_star_def Ball_def by(erule_tac x="p \ a" in allE) (simp add: permute_set_def) hence "p \ a \ c \ supp x \* p" using p2 by blast then show "\p. (p \ a) \ c \ supp x \* p" by blast qed section \Renaming permutations\ lemma set_renaming_perm: assumes b: "finite bs" shows "\q. (\b \ bs. q \ b = p \ b) \ supp q \ bs \ (p \ bs)" using b proof (induct) case empty have "(\b \ {}. 0 \ b = p \ b) \ supp (0::perm) \ {} \ p \ {}" by (simp add: permute_set_def supp_perm) then show "\q. (\b \ {}. q \ b = p \ b) \ supp q \ {} \ p \ {}" by blast next case (insert a bs) then have " \q. (\b \ bs. q \ b = p \ b) \ supp q \ bs \ p \ bs" by simp then obtain q where *: "\b \ bs. q \ b = p \ b" and **: "supp q \ bs \ p \ bs" by (metis empty_subsetI insert(3) supp_swap) { assume 1: "q \ a = p \ a" have "\b \ (insert a bs). q \ b = p \ b" using 1 * by simp moreover have "supp q \ insert a bs \ p \ insert a bs" using ** by (auto simp: insert_eqvt) ultimately have "\q. (\b \ insert a bs. q \ b = p \ b) \ supp q \ insert a bs \ p \ insert a bs" by blast } moreover { assume 2: "q \ a \ p \ a" define q' where "q' = ((q \ a) \ (p \ a)) + q" have "\b \ insert a bs. q' \ b = p \ b" using 2 * \a \ bs\ unfolding q'_def by (auto simp: swap_atom) moreover { have "{q \ a, p \ a} \ insert a bs \ p \ insert a bs" using ** apply (auto simp: supp_perm insert_eqvt) apply (subgoal_tac "q \ a \ bs \ p \ bs") apply(auto)[1] apply(subgoal_tac "q \ a \ {a. q \ a \ a}") apply(blast) apply(simp) done then have "supp (q \ a \ p \ a) \ insert a bs \ p \ insert a bs" unfolding supp_swap by auto moreover have "supp q \ insert a bs \ p \ insert a bs" using ** by (auto simp: insert_eqvt) ultimately have "supp q' \ insert a bs \ p \ insert a bs" unfolding q'_def using supp_plus_perm by blast } ultimately have "\q. (\b \ insert a bs. q \ b = p \ b) \ supp q \ insert a bs \ p \ insert a bs" by blast } ultimately show "\q. (\b \ insert a bs. q \ b = p \ b) \ supp q \ insert a bs \ p \ insert a bs" by blast qed lemma set_renaming_perm2: shows "\q. (\b \ bs. q \ b = p \ b) \ supp q \ bs \ (p \ bs)" proof - have "finite (bs \ supp p)" by (simp add: finite_supp) then obtain q where *: "\b \ bs \ supp p. q \ b = p \ b" and **: "supp q \ (bs \ supp p) \ (p \ (bs \ supp p))" using set_renaming_perm by blast from ** have "supp q \ bs \ (p \ bs)" by (auto simp: inter_eqvt) moreover have "\b \ bs - supp p. q \ b = p \ b" apply(auto) apply(subgoal_tac "b \ supp q") apply(simp add: fresh_def[symmetric]) apply(simp add: fresh_perm) apply(clarify) apply(rotate_tac 2) apply(drule subsetD[OF **]) apply(simp add: inter_eqvt supp_eqvt permute_self) done ultimately have "(\b \ bs. q \ b = p \ b) \ supp q \ bs \ (p \ bs)" using * by auto then show "\q. (\b \ bs. q \ b = p \ b) \ supp q \ bs \ (p \ bs)" by blast qed lemma list_renaming_perm: shows "\q. (\b \ set bs. q \ b = p \ b) \ supp q \ set bs \ (p \ set bs)" proof (induct bs) case (Cons a bs) then have " \q. (\b \ set bs. q \ b = p \ b) \ supp q \ set bs \ p \ (set bs)" by simp then obtain q where *: "\b \ set bs. q \ b = p \ b" and **: "supp q \ set bs \ p \ (set bs)" by (blast) { assume 1: "a \ set bs" have "q \ a = p \ a" using * 1 by (induct bs) (auto) then have "\b \ set (a # bs). q \ b = p \ b" using * by simp moreover have "supp q \ set (a # bs) \ p \ (set (a # bs))" using ** by (auto simp: insert_eqvt) ultimately have "\q. (\b \ set (a # bs). q \ b = p \ b) \ supp q \ set (a # bs) \ p \ (set (a # bs))" by blast } moreover { assume 2: "a \ set bs" define q' where "q' = ((q \ a) \ (p \ a)) + q" have "\b \ set (a # bs). q' \ b = p \ b" unfolding q'_def using 2 * \a \ set bs\ by (auto simp: swap_atom) moreover { have "{q \ a, p \ a} \ set (a # bs) \ p \ (set (a # bs))" using ** apply (auto simp: supp_perm insert_eqvt) apply (subgoal_tac "q \ a \ set bs \ p \ set bs") apply(auto)[1] apply(subgoal_tac "q \ a \ {a. q \ a \ a}") apply(blast) apply(simp) done then have "supp (q \ a \ p \ a) \ set (a # bs) \ p \ set (a # bs)" unfolding supp_swap by auto moreover have "supp q \ set (a # bs) \ p \ (set (a # bs))" using ** by (auto simp: insert_eqvt) ultimately have "supp q' \ set (a # bs) \ p \ (set (a # bs))" unfolding q'_def using supp_plus_perm by blast } ultimately have "\q. (\b \ set (a # bs). q \ b = p \ b) \ supp q \ set (a # bs) \ p \ (set (a # bs))" by blast } ultimately show "\q. (\b \ set (a # bs). q \ b = p \ b) \ supp q \ set (a # bs) \ p \ (set (a # bs))" by blast next case Nil have "(\b \ set []. 0 \ b = p \ b) \ supp (0::perm) \ set [] \ p \ set []" by (simp add: supp_zero_perm) then show "\q. (\b \ set []. q \ b = p \ b) \ supp q \ set [] \ p \ (set [])" by blast qed section \Concrete Atoms Types\ text \ Class \at_base\ allows types containing multiple sorts of atoms. Class \at\ only allows types with a single sort. \ class at_base = pt + fixes atom :: "'a \ atom" assumes atom_eq_iff [simp]: "atom a = atom b \ a = b" assumes atom_eqvt: "p \ (atom a) = atom (p \ a)" declare atom_eqvt [eqvt] class at = at_base + assumes sort_of_atom_eq [simp]: "sort_of (atom a) = sort_of (atom b)" lemma sort_ineq [simp]: assumes "sort_of (atom a) \ sort_of (atom b)" shows "atom a \ atom b" using assms by metis lemma supp_at_base: fixes a::"'a::at_base" shows "supp a = {atom a}" by (simp add: supp_atom [symmetric] supp_def atom_eqvt) lemma fresh_at_base: shows "sort_of a \ sort_of (atom b) \ a \ b" and "a \ b \ a \ atom b" unfolding fresh_def apply(simp_all add: supp_at_base) apply(metis) done (* solves the freshness only if the inequality can be shown by the simproc below *) lemma fresh_ineq_at_base [simp]: shows "a \ atom b \ a \ b" by (simp add: fresh_at_base) lemma fresh_atom_at_base [simp]: fixes b::"'a::at_base" shows "a \ atom b \ a \ b" by (simp add: fresh_def supp_at_base supp_atom) lemma fresh_star_atom_at_base: fixes b::"'a::at_base" shows "as \* atom b \ as \* b" by (simp add: fresh_star_def fresh_atom_at_base) lemma if_fresh_at_base [simp]: shows "atom a \ x \ P (if a = x then t else s) = P s" and "atom a \ x \ P (if x = a then t else s) = P s" by (simp_all add: fresh_at_base) simproc_setup fresh_ineq ("x \ (y::'a::at_base)") = \fn _ => fn ctxt => fn ctrm => case Thm.term_of ctrm of \<^Const_>\Not for \<^Const_>\HOL.eq _ for lhs rhs\\ => let fun first_is_neg lhs rhs [] = NONE | first_is_neg lhs rhs (thm::thms) = (case Thm.prop_of thm of _ $ \<^Const_>\Not for \<^Const_>\HOL.eq _ for l r\\ => (if l = lhs andalso r = rhs then SOME(thm) else if r = lhs andalso l = rhs then SOME(thm RS @{thm not_sym}) else first_is_neg lhs rhs thms) | _ => first_is_neg lhs rhs thms) val simp_thms = @{thms fresh_Pair fresh_at_base atom_eq_iff} val prems = Simplifier.prems_of ctxt |> filter (fn thm => case Thm.prop_of thm of _ $ \<^Const_>\fresh _ for \_ $ a\ b\ => (let val atms = a :: HOLogic.strip_tuple b in member (op =) atms lhs andalso member (op =) atms rhs end) | _ => false) |> map (simplify (put_simpset HOL_basic_ss ctxt addsimps simp_thms)) |> map (HOLogic.conj_elims ctxt) |> flat in case first_is_neg lhs rhs prems of SOME(thm) => SOME(thm RS @{thm Eq_TrueI}) | NONE => NONE end | _ => NONE \ instance at_base < fs proof qed (simp add: supp_at_base) lemma at_base_infinite [simp]: shows "infinite (UNIV :: 'a::at_base set)" (is "infinite ?U") proof obtain a :: 'a where "True" by auto assume "finite ?U" hence "finite (atom ` ?U)" by (rule finite_imageI) then obtain b where b: "b \ atom ` ?U" "sort_of b = sort_of (atom a)" by (rule obtain_atom) from b(2) have "b = atom ((atom a \ b) \ a)" unfolding atom_eqvt [symmetric] by (simp add: swap_atom) hence "b \ atom ` ?U" by simp with b(1) show "False" by simp qed lemma swap_at_base_simps [simp]: fixes x y::"'a::at_base" shows "sort_of (atom x) = sort_of (atom y) \ (atom x \ atom y) \ x = y" and "sort_of (atom x) = sort_of (atom y) \ (atom x \ atom y) \ y = x" and "atom x \ a \ atom x \ b \ (a \ b) \ x = x" unfolding atom_eq_iff [symmetric] unfolding atom_eqvt [symmetric] by simp_all lemma obtain_at_base: assumes X: "finite X" obtains a::"'a::at_base" where "atom a \ X" proof - have "inj (atom :: 'a \ atom)" by (simp add: inj_on_def) with X have "finite (atom -` X :: 'a set)" by (rule finite_vimageI) with at_base_infinite have "atom -` X \ (UNIV :: 'a set)" by auto then obtain a :: 'a where "atom a \ X" by auto thus ?thesis .. qed lemma obtain_fresh': assumes fin: "finite (supp x)" obtains a::"'a::at_base" where "atom a \ x" using obtain_at_base[where X="supp x"] by (auto simp: fresh_def fin) lemma obtain_fresh: fixes x::"'b::fs" obtains a::"'a::at_base" where "atom a \ x" by (rule obtain_fresh') (auto simp: finite_supp) lemma supp_finite_set_at_base: assumes a: "finite S" shows "supp S = atom ` S" apply(simp add: supp_of_finite_sets[OF a]) apply(simp add: supp_at_base) apply(auto) done (* FIXME lemma supp_cofinite_set_at_base: assumes a: "finite (UNIV - S)" shows "supp S = atom ` (UNIV - S)" apply(rule finite_supp_unique) *) lemma fresh_finite_set_at_base: fixes a::"'a::at_base" assumes a: "finite S" shows "atom a \ S \ a \ S" unfolding fresh_def apply(simp add: supp_finite_set_at_base[OF a]) apply(subst inj_image_mem_iff) apply(simp add: inj_on_def) apply(simp) done lemma fresh_at_base_permute_iff [simp]: fixes a::"'a::at_base" shows "atom (p \ a) \ p \ x \ atom a \ x" unfolding atom_eqvt[symmetric] by (simp only: fresh_permute_iff) lemma fresh_at_base_permI: shows "atom a \ p \ p \ a = a" by (simp add: fresh_def supp_perm) section \Infrastructure for concrete atom types\ definition flip :: "'a::at_base \ 'a \ perm" ("'(_ \ _')") where "(a \ b) = (atom a \ atom b)" lemma flip_fresh_fresh: assumes "atom a \ x" "atom b \ x" shows "(a \ b) \ x = x" using assms by (simp add: flip_def swap_fresh_fresh) lemma flip_self [simp]: "(a \ a) = 0" unfolding flip_def by (rule swap_self) lemma flip_commute: "(a \ b) = (b \ a)" unfolding flip_def by (rule swap_commute) lemma minus_flip [simp]: "- (a \ b) = (a \ b)" unfolding flip_def by (rule minus_swap) lemma add_flip_cancel: "(a \ b) + (a \ b) = 0" unfolding flip_def by (rule swap_cancel) lemma permute_flip_cancel [simp]: "(a \ b) \ (a \ b) \ x = x" unfolding permute_plus [symmetric] add_flip_cancel by simp lemma permute_flip_cancel2 [simp]: "(a \ b) \ (b \ a) \ x = x" by (simp add: flip_commute) lemma flip_eqvt [eqvt]: shows "p \ (a \ b) = (p \ a \ p \ b)" unfolding flip_def by (simp add: swap_eqvt atom_eqvt) lemma flip_at_base_simps [simp]: shows "sort_of (atom a) = sort_of (atom b) \ (a \ b) \ a = b" and "sort_of (atom a) = sort_of (atom b) \ (a \ b) \ b = a" and "\a \ c; b \ c\ \ (a \ b) \ c = c" and "sort_of (atom a) \ sort_of (atom b) \ (a \ b) \ x = x" unfolding flip_def unfolding atom_eq_iff [symmetric] unfolding atom_eqvt [symmetric] by simp_all text \the following two lemmas do not hold for \at_base\, only for single sort atoms from at\ lemma flip_triple: fixes a b c::"'a::at" assumes "a \ b" and "c \ b" shows "(a \ c) + (b \ c) + (a \ c) = (a \ b)" unfolding flip_def by (rule swap_triple) (simp_all add: assms) lemma permute_flip_at: fixes a b c::"'a::at" shows "(a \ b) \ c = (if c = a then b else if c = b then a else c)" unfolding flip_def apply (rule atom_eq_iff [THEN iffD1]) apply (subst atom_eqvt [symmetric]) apply (simp add: swap_atom) done lemma flip_at_simps [simp]: fixes a b::"'a::at" shows "(a \ b) \ a = b" and "(a \ b) \ b = a" unfolding permute_flip_at by simp_all subsection \Syntax for coercing at-elements to the atom-type\ syntax "_atom_constrain" :: "logic \ type \ logic" ("_:::_" [4, 0] 3) translations "_atom_constrain a t" => "CONST atom (_constrain a t)" subsection \A lemma for proving instances of class \at\.\ setup \Sign.add_const_constraint (@{const_name "permute"}, NONE)\ setup \Sign.add_const_constraint (@{const_name "atom"}, NONE)\ text \ New atom types are defined as subtypes of \<^typ>\atom\. \ lemma exists_eq_simple_sort: shows "\a. a \ {a. sort_of a = s}" by (rule_tac x="Atom s 0" in exI, simp) lemma exists_eq_sort: shows "\a. a \ {a. sort_of a \ range sort_fun}" by (rule_tac x="Atom (sort_fun x) y" in exI, simp) lemma at_base_class: fixes sort_fun :: "'b \ atom_sort" fixes Rep :: "'a \ atom" and Abs :: "atom \ 'a" assumes type: "type_definition Rep Abs {a. sort_of a \ range sort_fun}" assumes atom_def: "\a. atom a = Rep a" assumes permute_def: "\p a. p \ a = Abs (p \ Rep a)" shows "OFCLASS('a, at_base_class)" proof interpret type_definition Rep Abs "{a. sort_of a \ range sort_fun}" by (rule type) have sort_of_Rep: "\a. sort_of (Rep a) \ range sort_fun" using Rep by simp fix a b :: 'a and p p1 p2 :: perm show "0 \ a = a" unfolding permute_def by (simp add: Rep_inverse) show "(p1 + p2) \ a = p1 \ p2 \ a" unfolding permute_def by (simp add: Abs_inverse sort_of_Rep) show "atom a = atom b \ a = b" unfolding atom_def by (simp add: Rep_inject) show "p \ atom a = atom (p \ a)" unfolding permute_def atom_def by (simp add: Abs_inverse sort_of_Rep) qed (* lemma at_class: fixes s :: atom_sort fixes Rep :: "'a \ atom" and Abs :: "atom \ 'a" assumes type: "type_definition Rep Abs {a. sort_of a \ range (\x::unit. s)}" assumes atom_def: "\a. atom a = Rep a" assumes permute_def: "\p a. p \ a = Abs (p \ Rep a)" shows "OFCLASS('a, at_class)" proof interpret type_definition Rep Abs "{a. sort_of a \ range (\x::unit. s)}" by (rule type) have sort_of_Rep: "\a. sort_of (Rep a) = s" using Rep by (simp add: image_def) fix a b :: 'a and p p1 p2 :: perm show "0 \ a = a" unfolding permute_def by (simp add: Rep_inverse) show "(p1 + p2) \ a = p1 \ p2 \ a" unfolding permute_def by (simp add: Abs_inverse sort_of_Rep) show "sort_of (atom a) = sort_of (atom b)" unfolding atom_def by (simp add: sort_of_Rep) show "atom a = atom b \ a = b" unfolding atom_def by (simp add: Rep_inject) show "p \ atom a = atom (p \ a)" unfolding permute_def atom_def by (simp add: Abs_inverse sort_of_Rep) qed *) lemma at_class: fixes s :: atom_sort fixes Rep :: "'a \ atom" and Abs :: "atom \ 'a" assumes type: "type_definition Rep Abs {a. sort_of a = s}" assumes atom_def: "\a. atom a = Rep a" assumes permute_def: "\p a. p \ a = Abs (p \ Rep a)" shows "OFCLASS('a, at_class)" proof interpret type_definition Rep Abs "{a. sort_of a = s}" by (rule type) have sort_of_Rep: "\a. sort_of (Rep a) = s" using Rep by (simp add: image_def) fix a b :: 'a and p p1 p2 :: perm show "0 \ a = a" unfolding permute_def by (simp add: Rep_inverse) show "(p1 + p2) \ a = p1 \ p2 \ a" unfolding permute_def by (simp add: Abs_inverse sort_of_Rep) show "sort_of (atom a) = sort_of (atom b)" unfolding atom_def by (simp add: sort_of_Rep) show "atom a = atom b \ a = b" unfolding atom_def by (simp add: Rep_inject) show "p \ atom a = atom (p \ a)" unfolding permute_def atom_def by (simp add: Abs_inverse sort_of_Rep) qed lemma at_class_sort: fixes s :: atom_sort fixes Rep :: "'a \ atom" and Abs :: "atom \ 'a" fixes a::"'a" assumes type: "type_definition Rep Abs {a. sort_of a = s}" assumes atom_def: "\a. atom a = Rep a" shows "sort_of (atom a) = s" using atom_def type unfolding type_definition_def by simp setup \Sign.add_const_constraint (@{const_name "permute"}, SOME @{typ "perm \ 'a::pt \ 'a"})\ setup \Sign.add_const_constraint (@{const_name "atom"}, SOME @{typ "'a::at_base \ atom"})\ section \Library functions for the nominal infrastructure\ ML_file \nominal_library.ML\ section \The freshness lemma according to Andy Pitts\ lemma freshness_lemma: fixes h :: "'a::at \ 'b::pt" assumes a: "\a. atom a \ (h, h a)" shows "\x. \a. atom a \ h \ h a = x" proof - from a obtain b where a1: "atom b \ h" and a2: "atom b \ h b" by (auto simp: fresh_Pair) show "\x. \a. atom a \ h \ h a = x" proof (intro exI allI impI) fix a :: 'a assume a3: "atom a \ h" show "h a = h b" proof (cases "a = b") assume "a = b" thus "h a = h b" by simp next assume "a \ b" hence "atom a \ b" by (simp add: fresh_at_base) with a3 have "atom a \ h b" by (rule fresh_fun_app) with a2 have d1: "(atom b \ atom a) \ (h b) = (h b)" by (rule swap_fresh_fresh) from a1 a3 have d2: "(atom b \ atom a) \ h = h" by (rule swap_fresh_fresh) from d1 have "h b = (atom b \ atom a) \ (h b)" by simp also have "\ = ((atom b \ atom a) \ h) ((atom b \ atom a) \ b)" by (rule permute_fun_app_eq) also have "\ = h a" using d2 by simp finally show "h a = h b" by simp qed qed qed lemma freshness_lemma_unique: fixes h :: "'a::at \ 'b::pt" assumes a: "\a. atom a \ (h, h a)" shows "\!x. \a. atom a \ h \ h a = x" proof (rule ex_ex1I) from a show "\x. \a. atom a \ h \ h a = x" by (rule freshness_lemma) next fix x y assume x: "\a. atom a \ h \ h a = x" assume y: "\a. atom a \ h \ h a = y" from a x y show "x = y" by (auto simp: fresh_Pair) qed text \packaging the freshness lemma into a function\ definition Fresh :: "('a::at \ 'b::pt) \ 'b" where "Fresh h = (THE x. \a. atom a \ h \ h a = x)" lemma Fresh_apply: fixes h :: "'a::at \ 'b::pt" assumes a: "\a. atom a \ (h, h a)" assumes b: "atom a \ h" shows "Fresh h = h a" unfolding Fresh_def proof (rule the_equality) show "\a'. atom a' \ h \ h a' = h a" proof (intro strip) fix a':: 'a assume c: "atom a' \ h" from a have "\x. \a. atom a \ h \ h a = x" by (rule freshness_lemma) with b c show "h a' = h a" by auto qed next fix fr :: 'b assume "\a. atom a \ h \ h a = fr" with b show "fr = h a" by auto qed lemma Fresh_apply': fixes h :: "'a::at \ 'b::pt" assumes a: "atom a \ h" "atom a \ h a" shows "Fresh h = h a" apply (rule Fresh_apply) apply (auto simp: fresh_Pair intro: a) done simproc_setup Fresh_simproc ("Fresh (h::'a::at \ 'b::pt)") = \fn _ => fn ctxt => fn ctrm => let val _ $ h = Thm.term_of ctrm val atoms = Simplifier.prems_of ctxt |> map_filter (fn thm => case Thm.prop_of thm of _ $ \<^Const_>\fresh _ for \<^Const_>\atom _ for atm\ _\ => SOME atm | _ => NONE) |> distinct (op =) fun get_thm atm = let val goal1 = HOLogic.mk_Trueprop (mk_fresh (mk_atom atm) h) val goal2 = HOLogic.mk_Trueprop (mk_fresh (mk_atom atm) (h $ atm)) - val thm1 = Goal.prove ctxt [] [] goal1 (K (asm_simp_tac ctxt 1)) - val thm2 = Goal.prove ctxt [] [] goal2 (K (asm_simp_tac ctxt 1)) + val thm1 = Goal.prove ctxt [] [] goal1 (fn {context = ctxt', ...} => asm_simp_tac ctxt' 1) + val thm2 = Goal.prove ctxt [] [] goal2 (fn {context = ctxt', ...} => asm_simp_tac ctxt' 1) in SOME (@{thm Fresh_apply'} OF [thm1, thm2] RS eq_reflection) end handle ERROR _ => NONE in get_first get_thm atoms end \ lemma Fresh_eqvt: fixes h :: "'a::at \ 'b::pt" assumes a: "\a. atom a \ (h, h a)" shows "p \ (Fresh h) = Fresh (p \ h)" proof - from a obtain a::"'a::at" where fr: "atom a \ h" "atom a \ h a" by (metis fresh_Pair) then have fr_p: "atom (p \ a) \ (p \ h)" "atom (p \ a) \ (p \ h) (p \ a)" by (metis atom_eqvt fresh_permute_iff eqvt_apply)+ have "p \ (Fresh h) = p \ (h a)" using fr by simp also have "... = (p \ h) (p \ a)" by simp also have "... = Fresh (p \ h)" using fr_p by simp finally show "p \ (Fresh h) = Fresh (p \ h)" . qed lemma Fresh_supports: fixes h :: "'a::at \ 'b::pt" assumes a: "\a. atom a \ (h, h a)" shows "(supp h) supports (Fresh h)" apply (simp add: supports_def fresh_def [symmetric]) apply (simp add: Fresh_eqvt [OF a] swap_fresh_fresh) done notation Fresh (binder "FRESH " 10) lemma FRESH_f_iff: fixes P :: "'a::at \ 'b::pure" fixes f :: "'b \ 'c::pure" assumes P: "finite (supp P)" shows "(FRESH x. f (P x)) = f (FRESH x. P x)" proof - obtain a::'a where "atom a \ P" using P by (rule obtain_fresh') then show "(FRESH x. f (P x)) = f (FRESH x. P x)" by (simp add: pure_fresh) qed lemma FRESH_binop_iff: fixes P :: "'a::at \ 'b::pure" fixes Q :: "'a::at \ 'c::pure" fixes binop :: "'b \ 'c \ 'd::pure" assumes P: "finite (supp P)" and Q: "finite (supp Q)" shows "(FRESH x. binop (P x) (Q x)) = binop (FRESH x. P x) (FRESH x. Q x)" proof - from assms have "finite (supp (P, Q))" by (simp add: supp_Pair) then obtain a::'a where "atom a \ (P, Q)" by (rule obtain_fresh') then show ?thesis by (simp add: pure_fresh) qed lemma FRESH_conj_iff: fixes P Q :: "'a::at \ bool" assumes P: "finite (supp P)" and Q: "finite (supp Q)" shows "(FRESH x. P x \ Q x) \ (FRESH x. P x) \ (FRESH x. Q x)" using P Q by (rule FRESH_binop_iff) lemma FRESH_disj_iff: fixes P Q :: "'a::at \ bool" assumes P: "finite (supp P)" and Q: "finite (supp Q)" shows "(FRESH x. P x \ Q x) \ (FRESH x. P x) \ (FRESH x. Q x)" using P Q by (rule FRESH_binop_iff) section \Automation for creating concrete atom types\ text \At the moment only single-sort concrete atoms are supported.\ ML_file \nominal_atoms.ML\ section \Automatic equivariance procedure for inductive definitions\ ML_file \nominal_eqvt.ML\ end diff --git a/thys/Nominal2/nominal_dt_alpha.ML b/thys/Nominal2/nominal_dt_alpha.ML --- a/thys/Nominal2/nominal_dt_alpha.ML +++ b/thys/Nominal2/nominal_dt_alpha.ML @@ -1,870 +1,873 @@ (* Title: nominal_dt_alpha.ML Author: Cezary Kaliszyk Author: Christian Urban Definitions and proofs for the alpha-relations. *) signature NOMINAL_DT_ALPHA = sig val comb_binders: Proof.context -> bmode -> term list -> (term option * int) list -> term val define_raw_alpha: raw_dt_info -> bn_info list -> bclause list list list -> term list -> Proof.context -> alpha_result * local_theory val induct_prove: typ list -> (typ * (term -> term)) list -> thm -> (Proof.context -> int -> tactic) -> Proof.context -> thm list val alpha_prove: term list -> (term * ((term * term) -> term)) list -> thm -> (Proof.context -> int -> tactic) -> Proof.context -> thm list val raw_prove_alpha_distincts: Proof.context -> alpha_result -> raw_dt_info -> thm list val raw_prove_alpha_eq_iff: Proof.context -> alpha_result -> raw_dt_info -> thm list val raw_prove_refl: Proof.context -> alpha_result -> thm -> thm list val raw_prove_sym: Proof.context -> alpha_result -> thm list -> thm list val raw_prove_trans: Proof.context -> alpha_result -> thm list -> thm list -> thm list val raw_prove_equivp: Proof.context -> alpha_result -> thm list -> thm list -> thm list -> thm list * thm list val raw_prove_bn_imp: Proof.context -> alpha_result -> thm list val raw_fv_bn_rsp_aux: Proof.context -> alpha_result -> term list -> term list -> term list -> thm list -> thm list val raw_size_rsp_aux: Proof.context -> alpha_result -> thm list -> thm list val raw_constrs_rsp: Proof.context -> alpha_result -> term list list -> thm list -> thm list list val raw_alpha_bn_rsp: alpha_result -> thm list -> thm list -> thm list val raw_perm_bn_rsp: Proof.context -> alpha_result -> term list -> thm list -> thm list val mk_funs_rsp: Proof.context -> thm -> thm val mk_alpha_permute_rsp: Proof.context -> thm -> thm end structure Nominal_Dt_Alpha: NOMINAL_DT_ALPHA = struct open Nominal_Permeq open Nominal_Dt_Data fun lookup xs x = the (AList.lookup (op=) xs x) fun group xs = AList.group (op=) xs (** definition of the inductive rules for alpha and alpha_bn **) (* construct the compound terms for prod_fv and prod_alpha *) fun mk_prod_fv (t1, t2) = let val A = domain_type (fastype_of t1) val B = domain_type (fastype_of t2) in \<^Const>\prod_fv A B for t1 t2\ end fun mk_prod_alpha (t1, t2) = let val A = domain_type (fastype_of t1) val B = domain_type (fastype_of t2) in \<^Const>\prod_alpha A B for t1 t2\ end (* generates the compound binder terms *) fun comb_binders lthy bmode args binders = let fun bind_set lthy args (NONE, i) = setify lthy (nth args i) | bind_set _ args (SOME bn, i) = bn $ (nth args i) fun bind_lst lthy args (NONE, i) = listify lthy (nth args i) | bind_lst _ args (SOME bn, i) = bn $ (nth args i) val (combine_fn, bind_fn) = case bmode of Lst => (mk_append, bind_lst) | Set => (mk_union, bind_set) | Res => (mk_union, bind_set) in binders |> map (bind_fn lthy args) |> foldl1 combine_fn end (* produces the term for an alpha with abstraction *) fun mk_alpha_term bmode fv alpha args args' binders binders' = let val (alpha_name, binder_ty) = case bmode of Lst => (@{const_name "alpha_lst"}, @{typ "atom list"}) | Set => (@{const_name "alpha_set"}, @{typ "atom set"}) | Res => (@{const_name "alpha_res"}, @{typ "atom set"}) val ty = fastype_of args val pair_ty = HOLogic.mk_prodT (binder_ty, ty) val alpha_ty = [ty, ty] ---> @{typ "bool"} val fv_ty = ty --> @{typ "atom set"} val pair_lhs = HOLogic.mk_prod (binders, args) val pair_rhs = HOLogic.mk_prod (binders', args') in HOLogic.exists_const \<^Type>\perm\ $ Abs ("p", \<^Type>\perm\, Const (alpha_name, [pair_ty, alpha_ty, fv_ty, @{typ "perm"}, pair_ty] ---> @{typ bool}) $ pair_lhs $ alpha $ fv $ (Bound 0) $ pair_rhs) end (* for non-recursive binders we have to produce alpha_bn premises *) fun mk_alpha_bn_prem alpha_bn_map args args' bodies binder = case binder of (NONE, _) => [] | (SOME bn, i) => if member (op=) bodies i then [] else [lookup alpha_bn_map bn $ nth args i $ nth args' i] (* generate the premises for an alpha rule; mk_frees is used if no binders are present *) fun mk_alpha_prems lthy alpha_map alpha_bn_map is_rec (args, args') bclause = let fun mk_frees i = let val arg = nth args i val arg' = nth args' i val ty = fastype_of arg in if nth is_rec i then fst (lookup alpha_map ty) $ arg $ arg' else HOLogic.mk_eq (arg, arg') end fun mk_alpha_fv i = let val ty = fastype_of (nth args i) in case AList.lookup (op=) alpha_map ty of NONE => (HOLogic.eq_const ty, supp_const ty) | SOME (alpha, fv) => (alpha, fv) end in case bclause of BC (_, [], bodies) => map (HOLogic.mk_Trueprop o mk_frees) bodies | BC (bmode, binders, bodies) => let val (alphas, fvs) = split_list (map mk_alpha_fv bodies) val comp_fv = foldl1 mk_prod_fv fvs val comp_alpha = foldl1 mk_prod_alpha alphas val comp_args = foldl1 HOLogic.mk_prod (map (nth args) bodies) val comp_args' = foldl1 HOLogic.mk_prod (map (nth args') bodies) val comp_binders = comb_binders lthy bmode args binders val comp_binders' = comb_binders lthy bmode args' binders val alpha_prem = mk_alpha_term bmode comp_fv comp_alpha comp_args comp_args' comp_binders comp_binders' val alpha_bn_prems = flat (map (mk_alpha_bn_prem alpha_bn_map args args' bodies) binders) in map HOLogic.mk_Trueprop (alpha_prem::alpha_bn_prems) end end (* produces the introduction rule for an alpha rule *) fun mk_alpha_intros lthy alpha_map alpha_bn_map (constr, ty, arg_tys, is_rec) bclauses = let val arg_names = Old_Datatype_Prop.make_tnames arg_tys val arg_names' = Name.variant_list arg_names arg_names val args = map Free (arg_names ~~ arg_tys) val args' = map Free (arg_names' ~~ arg_tys) val alpha = fst (lookup alpha_map ty) val concl = HOLogic.mk_Trueprop (alpha $ list_comb (constr, args) $ list_comb (constr, args')) val prems = map (mk_alpha_prems lthy alpha_map alpha_bn_map is_rec (args, args')) bclauses in Library.foldr Logic.mk_implies (flat prems, concl) end (* produces the premise of an alpha-bn rule; we only need to treat the case special where the binding clause is empty; - if the body is not included in the bn_info, then we either produce an equation or an alpha-premise - if the body is included in the bn_info, then we create either a recursive call to alpha-bn, or no premise *) fun mk_alpha_bn lthy alpha_map alpha_bn_map bn_args is_rec (args, args') bclause = let fun mk_alpha_bn_prem i = let val arg = nth args i val arg' = nth args' i val ty = fastype_of arg in case AList.lookup (op=) bn_args i of NONE => (case (AList.lookup (op=) alpha_map ty) of NONE => [HOLogic.mk_eq (arg, arg')] | SOME (alpha, _) => [alpha $ arg $ arg']) | SOME (NONE) => [] | SOME (SOME bn) => [lookup alpha_bn_map bn $ arg $ arg'] end in case bclause of BC (_, [], bodies) => map HOLogic.mk_Trueprop (flat (map mk_alpha_bn_prem bodies)) | _ => mk_alpha_prems lthy alpha_map alpha_bn_map is_rec (args, args') bclause end fun mk_alpha_bn_intro lthy bn_trm alpha_map alpha_bn_map (bn_args, (constr, _, arg_tys, is_rec)) bclauses = let val arg_names = Old_Datatype_Prop.make_tnames arg_tys val arg_names' = Name.variant_list arg_names arg_names val args = map Free (arg_names ~~ arg_tys) val args' = map Free (arg_names' ~~ arg_tys) val alpha_bn = lookup alpha_bn_map bn_trm val concl = HOLogic.mk_Trueprop (alpha_bn $ list_comb (constr, args) $ list_comb (constr, args')) val prems = map (mk_alpha_bn lthy alpha_map alpha_bn_map bn_args is_rec (args, args')) bclauses in Library.foldr Logic.mk_implies (flat prems, concl) end fun mk_alpha_bn_intros lthy alpha_map alpha_bn_map constrs_info bclausesss (bn_trm, bn_n, bn_argss) = let val nth_constrs_info = nth constrs_info bn_n val nth_bclausess = nth bclausesss bn_n in map2 (mk_alpha_bn_intro lthy bn_trm alpha_map alpha_bn_map) (bn_argss ~~ nth_constrs_info) nth_bclausess end fun define_raw_alpha raw_dt_info bn_info bclausesss fvs lthy = let val RawDtInfo {raw_dt_names, raw_tys, raw_cns_info, ...} = raw_dt_info val alpha_names = map (prefix "alpha_" o Long_Name.base_name) raw_dt_names val alpha_tys = map (fn ty => [ty, ty] ---> @{typ bool}) raw_tys val alpha_frees = map Free (alpha_names ~~ alpha_tys) val alpha_map = raw_tys ~~ (alpha_frees ~~ fvs) val (bns, bn_tys) = split_list (map (fn (bn, i, _) => (bn, i)) bn_info) val bn_names = map (fn bn => Long_Name.base_name (fst (dest_Const bn))) bns val alpha_bn_names = map (prefix "alpha_") bn_names val alpha_bn_arg_tys = map (nth raw_tys) bn_tys val alpha_bn_tys = map (fn ty => [ty, ty] ---> @{typ "bool"}) alpha_bn_arg_tys val alpha_bn_frees = map Free (alpha_bn_names ~~ alpha_bn_tys) val alpha_bn_map = bns ~~ alpha_bn_frees val alpha_intros = map2 (map2 (mk_alpha_intros lthy alpha_map alpha_bn_map)) raw_cns_info bclausesss val alpha_bn_intros = map (mk_alpha_bn_intros lthy alpha_map alpha_bn_map raw_cns_info bclausesss) bn_info val all_alpha_names = map (fn (a, ty) => ((Binding.name a, ty), NoSyn)) (alpha_names @ alpha_bn_names ~~ alpha_tys @ alpha_bn_tys) val all_alpha_intros = map (pair Binding.empty_atts) (flat alpha_intros @ flat alpha_bn_intros) val (alpha_info, lthy') = lthy |> Local_Theory.begin_nested |> snd |> Inductive.add_inductive {quiet_mode = true, verbose = false, alt_name = Binding.empty, coind = false, no_elim = false, no_ind = false, skip_mono = false} all_alpha_names [] all_alpha_intros [] |> Local_Theory.end_nested_result Inductive.transform_result; val phi = Proof_Context.export_morphism lthy' (Proof_Context.transfer (Proof_Context.theory_of lthy') lthy); val alpha_info_global = Inductive.transform_result phi alpha_info; val all_alpha_trms = #preds alpha_info_global |> map Type.legacy_freeze (*FIXME*) val alpha_raw_induct = #raw_induct alpha_info_global val alpha_intros = #intrs alpha_info_global; val alpha_cases = #elims alpha_info_global; val (alpha_trms, alpha_bn_trms) = chop (length fvs) all_alpha_trms val alpha_tys = map (domain_type o fastype_of) alpha_trms val alpha_bn_tys = map (domain_type o fastype_of) alpha_bn_trms val alpha_names = map (fst o dest_Const) alpha_trms val alpha_bn_names = map (fst o dest_Const) alpha_bn_trms in (AlphaResult {alpha_names = alpha_names, alpha_trms = alpha_trms, alpha_tys = alpha_tys, alpha_bn_names = alpha_bn_names, alpha_bn_trms = alpha_bn_trms, alpha_bn_tys = alpha_bn_tys, alpha_intros = alpha_intros, alpha_cases = alpha_cases, alpha_raw_induct = alpha_raw_induct}, lthy') end (** induction proofs **) (* proof by structural induction over data types *) fun induct_prove tys props induct_thm cases_tac ctxt = let val (arg_names, ctxt') = Variable.variant_fixes (replicate (length tys) "x") ctxt val args = map2 (curry Free) arg_names tys val true_trms = replicate (length tys) (K \<^Const>\True\) fun apply_all x fs = map (fn f => f x) fs val goals = group (props @ (tys ~~ true_trms)) |> map snd |> map2 apply_all args |> map fold_conj |> foldr1 HOLogic.mk_conj |> HOLogic.mk_Trueprop fun tac ctxt = HEADGOAL (DETERM o (resolve_tac ctxt [induct_thm]) THEN_ALL_NEW (REPEAT_ALL_NEW (FIRST' [resolve_tac ctxt @{thms TrueI conjI}, cases_tac ctxt]))) in - Goal.prove ctxt' [] [] goals (fn {context, ...} => tac context) + Goal.prove ctxt' [] [] goals (tac o #context) |> singleton (Proof_Context.export ctxt' ctxt) |> Old_Datatype_Aux.split_conj_thm |> map Old_Datatype_Aux.split_conj_thm |> flat |> filter_out (is_true o Thm.concl_of) |> map zero_var_indexes end (* proof by rule induction over the alpha-definitions *) fun alpha_prove alphas props alpha_induct_thm cases_tac ctxt = let val arg_tys = map (domain_type o fastype_of) alphas val ((arg_names1, arg_names2), ctxt') = ctxt |> Variable.variant_fixes (replicate (length alphas) "x") ||>> Variable.variant_fixes (replicate (length alphas) "y") val args1 = map2 (curry Free) arg_names1 arg_tys val args2 = map2 (curry Free) arg_names2 arg_tys val true_trms = replicate (length alphas) (K \<^Const>\True\) fun apply_all x fs = map (fn f => f x) fs val goals_rhs = group (props @ (alphas ~~ true_trms)) |> map snd |> map2 apply_all (args1 ~~ args2) |> map fold_conj fun apply_trm_pair t (ar1, ar2) = t $ ar1 $ ar2 val goals_lhs = map2 apply_trm_pair alphas (args1 ~~ args2) val goals = (map2 (curry HOLogic.mk_imp) goals_lhs goals_rhs) |> foldr1 HOLogic.mk_conj |> HOLogic.mk_Trueprop fun tac ctxt = HEADGOAL (DETERM o (resolve_tac ctxt [alpha_induct_thm]) THEN_ALL_NEW FIRST' [resolve_tac ctxt @{thms TrueI}, cases_tac ctxt]) in - Goal.prove ctxt' [] [] goals (fn {context, ...} => tac context) + Goal.prove ctxt' [] [] goals (tac o #context) |> singleton (Proof_Context.export ctxt' ctxt) |> Old_Datatype_Aux.split_conj_thm |> map (fn th => th RS mp) |> map Old_Datatype_Aux.split_conj_thm |> flat |> filter_out (is_true o Thm.concl_of) |> map zero_var_indexes end (** produces the distinctness theorems **) (* transforms the distinctness theorems of the constructors into "not-alphas" of the constructors *) fun mk_distinct_goal ty_trm_assoc neq = let val (Const (@{const_name "HOL.eq"}, ty_eq) $ lhs $ rhs) = HOLogic.dest_not (HOLogic.dest_Trueprop neq) val ty_str = fst (dest_Type (domain_type ty_eq)) in Const (lookup ty_trm_assoc ty_str, ty_eq) $ lhs $ rhs |> HOLogic.mk_not |> HOLogic.mk_Trueprop end fun distinct_tac ctxt cases_thms distinct_thms = resolve_tac ctxt [notI] THEN' eresolve_tac ctxt cases_thms THEN_ALL_NEW asm_full_simp_tac (put_simpset HOL_ss ctxt addsimps distinct_thms) fun raw_prove_alpha_distincts ctxt alpha_result raw_dt_info = let val AlphaResult {alpha_names, alpha_cases, ...} = alpha_result val RawDtInfo {raw_dt_names, raw_distinct_thms, ...} = raw_dt_info val ty_trm_assoc = raw_dt_names ~~ alpha_names fun mk_alpha_distinct raw_distinct_trm = let val goal = mk_distinct_goal ty_trm_assoc raw_distinct_trm in Goal.prove ctxt [] [] goal - (K (distinct_tac ctxt alpha_cases raw_distinct_thms 1)) + (fn {context = ctxt', ...} => distinct_tac ctxt' alpha_cases raw_distinct_thms 1) end in map (mk_alpha_distinct o Thm.prop_of) raw_distinct_thms end (** produces the alpha_eq_iff simplification rules **) (* in case a theorem is of the form (Rel Const Const), it will be rewritten to ((Rel Const = Const) = True) *) fun mk_simp_rule thm = case Thm.prop_of thm of \<^Const_>\Trueprop for \_ $ Const _ $ Const _\\ => thm RS @{thm eqTrueI} | _ => thm fun alpha_eq_iff_tac ctxt dist_inj intros elims = SOLVED' (asm_full_simp_tac (put_simpset HOL_ss ctxt addsimps intros)) ORELSE' (resolve_tac ctxt @{thms iffI} THEN' RANGE [eresolve_tac ctxt elims THEN_ALL_NEW asm_full_simp_tac (put_simpset HOL_ss ctxt addsimps dist_inj), asm_full_simp_tac (put_simpset HOL_ss ctxt addsimps intros)]) fun mk_alpha_eq_iff_goal thm = let val prop = Thm.prop_of thm; val concl = HOLogic.dest_Trueprop (Logic.strip_imp_concl prop); val hyps = map HOLogic.dest_Trueprop (Logic.strip_imp_prems prop); fun list_conj l = foldr1 HOLogic.mk_conj l; in if hyps = [] then HOLogic.mk_Trueprop concl else HOLogic.mk_Trueprop (HOLogic.mk_eq (concl, list_conj hyps)) end; fun raw_prove_alpha_eq_iff ctxt alpha_result raw_dt_info = let val AlphaResult {alpha_intros, alpha_cases, ...} = alpha_result val RawDtInfo {raw_distinct_thms, raw_inject_thms, ...} = raw_dt_info val ((_, thms_imp), ctxt') = Variable.import false alpha_intros ctxt; val goals = map mk_alpha_eq_iff_goal thms_imp; - val tac = alpha_eq_iff_tac ctxt (raw_distinct_thms @ raw_inject_thms) alpha_intros alpha_cases 1; - val thms = map (fn goal => Goal.prove ctxt' [] [] goal (K tac)) goals; + fun tac goal_ctxt = + alpha_eq_iff_tac goal_ctxt (raw_distinct_thms @ raw_inject_thms) alpha_intros alpha_cases 1; + val thms = map (fn goal => Goal.prove ctxt' [] [] goal (tac o #context)) goals; in Variable.export ctxt' ctxt thms |> map mk_simp_rule end (** reflexivity proof for the alphas **) val exi_zero = @{lemma "P (0::perm) \ (\x. P x)" by auto} fun cases_tac intros ctxt = let val prod_simps = @{thms split_conv prod_alpha_def rel_prod_conv} val unbound_tac = REPEAT o (eresolve_tac ctxt @{thms conjE}) THEN' assume_tac ctxt val bound_tac = EVERY' [ resolve_tac ctxt [exi_zero], resolve_tac ctxt @{thms alpha_refl}, asm_full_simp_tac (put_simpset HOL_ss ctxt addsimps prod_simps) ] in resolve_tac ctxt intros THEN_ALL_NEW FIRST' [resolve_tac ctxt @{thms refl}, unbound_tac, bound_tac] end fun raw_prove_refl ctxt alpha_result raw_dt_induct = let val AlphaResult {alpha_trms, alpha_tys, alpha_bn_trms, alpha_bn_tys, alpha_intros, ...} = alpha_result val props = map (fn (ty, c) => (ty, fn x => c $ x $ x)) ((alpha_tys ~~ alpha_trms) @ (alpha_bn_tys ~~ alpha_bn_trms)) in induct_prove alpha_tys props raw_dt_induct (cases_tac alpha_intros) ctxt end (** symmetry proof for the alphas **) val exi_neg = @{lemma "(\(p::perm). P p) \ (\q. P q \ Q (- q)) \ \p. Q p" by (erule exE, rule_tac x="-p" in exI, auto)} (* for premises that contain binders *) fun prem_bound_tac pred_names alpha_eqvt ctxt = let fun trans_prem_tac pred_names ctxt = SUBPROOF (fn {prems, context = ctxt', ...} => let val prems' = map (transform_prem1 ctxt' pred_names) prems in resolve_tac ctxt' prems' 1 end) ctxt val prod_simps = @{thms split_conv permute_prod.simps prod_alpha_def rel_prod_conv alphas} in EVERY' [ eresolve_tac ctxt [exi_neg], resolve_tac ctxt @{thms alpha_sym_eqvt}, asm_full_simp_tac (put_simpset HOL_ss ctxt addsimps prod_simps), eqvt_tac ctxt (eqvt_relaxed_config addpres alpha_eqvt) THEN' resolve_tac ctxt @{thms refl}, trans_prem_tac pred_names ctxt] end fun raw_prove_sym ctxt alpha_result alpha_eqvt = let val AlphaResult {alpha_names, alpha_trms, alpha_bn_names, alpha_bn_trms, alpha_intros, alpha_raw_induct, ...} = alpha_result val alpha_trms = alpha_trms @ alpha_bn_trms val alpha_names = alpha_names @ alpha_bn_names val props = map (fn t => fn (x, y) => t $ y $ x) alpha_trms fun tac ctxt = resolve_tac ctxt alpha_intros THEN_ALL_NEW FIRST' [assume_tac ctxt, resolve_tac ctxt @{thms sym} THEN' assume_tac ctxt, prem_bound_tac alpha_names alpha_eqvt ctxt] in alpha_prove alpha_trms (alpha_trms ~~ props) alpha_raw_induct tac ctxt end (** transitivity proof for alphas **) (* applies cases rules and resolves them with the last premise *) fun ecases_tac cases = Subgoal.FOCUS (fn {context = ctxt, prems, ...} => HEADGOAL (resolve_tac ctxt cases THEN' resolve_tac ctxt [List.last prems])) fun aatac pred_names = SUBPROOF (fn {context = ctxt, prems, ...} => HEADGOAL (resolve_tac ctxt (map (transform_prem1 ctxt pred_names) prems))) (* instantiates exI with the permutation p + q *) val perm_inst_tac = Subgoal.FOCUS (fn {context = ctxt, params, ...} => let val (p, q) = apply2 snd (last2 params) val pq_inst = foldl1 (uncurry Thm.apply) [@{cterm "plus::perm => perm => perm"}, p, q] val exi_inst = Thm.instantiate' [SOME (@{ctyp "perm"})] [NONE, SOME pq_inst] @{thm exI} in HEADGOAL (resolve_tac ctxt [exi_inst]) end) fun non_trivial_cases_tac pred_names intros alpha_eqvt ctxt = let val prod_simps = @{thms split_conv alphas permute_prod.simps prod_alpha_def rel_prod_conv} in resolve_tac ctxt intros THEN_ALL_NEW (asm_simp_tac (put_simpset HOL_basic_ss ctxt) THEN' TRY o EVERY' (* if binders are present *) [ eresolve_tac ctxt @{thms exE}, eresolve_tac ctxt @{thms exE}, perm_inst_tac ctxt, resolve_tac ctxt @{thms alpha_trans_eqvt}, assume_tac ctxt, aatac pred_names ctxt, eqvt_tac ctxt (eqvt_relaxed_config addpres alpha_eqvt) THEN' resolve_tac ctxt @{thms refl}, asm_full_simp_tac (put_simpset HOL_ss ctxt addsimps prod_simps) ]) end fun prove_trans_tac alpha_result raw_dt_thms alpha_eqvt ctxt = let val AlphaResult {alpha_names, alpha_bn_names, alpha_intros, alpha_cases, ...} = alpha_result val alpha_names = alpha_names @ alpha_bn_names fun all_cases ctxt = asm_full_simp_tac (put_simpset HOL_basic_ss ctxt addsimps raw_dt_thms) THEN' TRY o non_trivial_cases_tac alpha_names alpha_intros alpha_eqvt ctxt in EVERY' [ resolve_tac ctxt @{thms allI}, resolve_tac ctxt @{thms impI}, ecases_tac alpha_cases ctxt THEN_ALL_NEW all_cases ctxt ] end fun prep_trans_goal alpha_trm (arg1, arg2) = let val arg_ty = fastype_of arg1 val mid = alpha_trm $ arg2 $ (Bound 0) val rhs = alpha_trm $ arg1 $ (Bound 0) in HOLogic.all_const arg_ty $ Abs ("z", arg_ty, HOLogic.mk_imp (mid, rhs)) end fun raw_prove_trans ctxt alpha_result raw_dt_thms alpha_eqvt = let val AlphaResult {alpha_trms, alpha_bn_trms, alpha_raw_induct, ...} = alpha_result val alpha_trms = alpha_trms @ alpha_bn_trms val props = map prep_trans_goal alpha_trms in alpha_prove alpha_trms (alpha_trms ~~ props) alpha_raw_induct (prove_trans_tac alpha_result raw_dt_thms alpha_eqvt) ctxt end (** proves the equivp predicate for all alphas **) val reflp_def' = @{lemma "reflp R == \x. R x x" by (simp add: reflp_def refl_on_def)} val symp_def' = @{lemma "symp R \ \x y . R x y --> R y x" by (simp add: symp_def sym_def)} val transp_def' = @{lemma "transp R \ \x y. R x y \ (\z. R y z \ R x z)" by (rule eq_reflection, auto simp add: trans_def transp_def)} fun raw_prove_equivp ctxt alpha_result refl symm trans = let val AlphaResult {alpha_trms, alpha_bn_trms, ...} = alpha_result val refl' = map (fold_rule ctxt [reflp_def'] o atomize ctxt) refl val symm' = map (fold_rule ctxt [symp_def'] o atomize ctxt) symm val trans' = map (fold_rule ctxt [transp_def'] o atomize ctxt) trans fun prep_goal t = \<^Const>\Trueprop for \<^Const>\equivp \domain_type (fastype_of t)\ for t\\ in Goal.prove_common ctxt NONE [] [] (map prep_goal (alpha_trms @ alpha_bn_trms)) - (K (HEADGOAL (Goal.conjunction_tac THEN_ALL_NEW (resolve_tac ctxt @{thms equivpI} THEN' - RANGE [resolve_tac ctxt refl', resolve_tac ctxt symm', resolve_tac ctxt trans'])))) + (fn {context = ctxt', ...} => + HEADGOAL (Goal.conjunction_tac THEN_ALL_NEW (resolve_tac ctxt' @{thms equivpI} THEN' + RANGE [resolve_tac ctxt' refl', resolve_tac ctxt' symm', resolve_tac ctxt' trans']))) |> chop (length alpha_trms) end (* proves that alpha_raw implies alpha_bn *) fun raw_prove_bn_imp_tac alpha_result ctxt = SUBPROOF (fn {prems, context = ctxt', ...} => let val AlphaResult {alpha_names, alpha_intros, ...} = alpha_result val prems' = flat (map Old_Datatype_Aux.split_conj_thm prems) val prems'' = map (transform_prem1 ctxt' alpha_names) prems' in HEADGOAL (REPEAT_ALL_NEW (FIRST' [ resolve_tac ctxt' @{thms TrueI}, resolve_tac ctxt' @{thms conjI}, resolve_tac ctxt' prems', resolve_tac ctxt' prems'', resolve_tac ctxt' alpha_intros ])) end) ctxt fun raw_prove_bn_imp ctxt alpha_result = let val AlphaResult {alpha_trms, alpha_tys, alpha_bn_trms, alpha_raw_induct, ...} = alpha_result val arg_ty = domain_type o fastype_of val ty_assoc = alpha_tys ~~ alpha_trms val props = map (fn t => (lookup ty_assoc (arg_ty t), fn (x, y) => t $ x $ y)) alpha_bn_trms in alpha_prove (alpha_trms @ alpha_bn_trms) props alpha_raw_induct (raw_prove_bn_imp_tac alpha_result) ctxt end (* respectfulness for fv_raw / bn_raw *) fun raw_fv_bn_rsp_aux ctxt alpha_result fvs bns fv_bns simps = let val AlphaResult {alpha_trms, alpha_tys, alpha_bn_trms, alpha_raw_induct, ...} = alpha_result val arg_ty = domain_type o fastype_of val ty_assoc = alpha_tys ~~ alpha_trms fun mk_eq' t x y = HOLogic.mk_eq (t $ x, t $ y) val prop1 = map (fn t => (lookup ty_assoc (arg_ty t), fn (x, y) => mk_eq' t x y)) fvs val prop2 = map (fn t => (lookup ty_assoc (arg_ty t), fn (x, y) => mk_eq' t x y)) (bns @ fv_bns) val prop3 = map2 (fn t1 => fn t2 => (t1, fn (x, y) => mk_eq' t2 x y)) alpha_bn_trms fv_bns val simpset = put_simpset HOL_ss ctxt addsimps (simps @ @{thms alphas prod_fv.simps set_simps append.simps} @ @{thms Un_assoc Un_insert_left Un_empty_right Un_empty_left}) in alpha_prove (alpha_trms @ alpha_bn_trms) (prop1 @ prop2 @ prop3) alpha_raw_induct (K (asm_full_simp_tac simpset)) ctxt end (* respectfulness for size *) fun raw_size_rsp_aux ctxt alpha_result simps = let val AlphaResult {alpha_trms, alpha_tys, alpha_bn_trms, alpha_bn_tys, alpha_raw_induct, ...} = alpha_result fun mk_prop ty (x, y) = HOLogic.mk_eq (HOLogic.size_const ty $ x, HOLogic.size_const ty $ y) val props = (alpha_trms @ alpha_bn_trms) ~~ (map mk_prop (alpha_tys @ alpha_bn_tys)) val simpset = put_simpset HOL_ss ctxt addsimps (simps @ @{thms alphas prod_alpha_def rel_prod_conv permute_prod_def prod.case prod.rec}) val tac = (TRY o REPEAT o eresolve_tac ctxt @{thms exE}) THEN' asm_full_simp_tac simpset in alpha_prove (alpha_trms @ alpha_bn_trms) props alpha_raw_induct (K tac) ctxt end (* respectfulness for constructors *) fun raw_constr_rsp_tac ctxt alpha_intros simps = let val pre_simpset = put_simpset HOL_ss ctxt addsimps @{thms rel_fun_def} val post_simpset = put_simpset HOL_ss ctxt addsimps @{thms alphas prod_alpha_def rel_prod_conv prod_fv.simps fresh_star_zero permute_zero prod.case} @ simps in asm_full_simp_tac pre_simpset THEN' REPEAT o (resolve_tac ctxt @{thms allI impI}) THEN' resolve_tac ctxt alpha_intros THEN_ALL_NEW (TRY o (resolve_tac ctxt [exi_zero]) THEN' asm_full_simp_tac post_simpset) end fun raw_constrs_rsp ctxt alpha_result constrs simps = let val AlphaResult {alpha_trms, alpha_tys, alpha_intros, ...} = alpha_result fun lookup ty = case AList.lookup (op=) (alpha_tys ~~ alpha_trms) ty of NONE => HOLogic.eq_const ty | SOME alpha => alpha fun rel_fun_app (t1, t2) = Const (@{const_name "rel_fun"}, dummyT) $ t1 $ t2 fun prep_goal trm = trm |> strip_type o fastype_of |> (fn (tys, ty) => tys @ [ty]) |> map lookup |> foldr1 rel_fun_app |> (fn t => t $ trm $ trm) |> Syntax.check_term ctxt |> HOLogic.mk_Trueprop in map (fn constrs => Goal.prove_common ctxt NONE [] [] (map prep_goal constrs) - (K (HEADGOAL - (Goal.conjunction_tac THEN_ALL_NEW raw_constr_rsp_tac ctxt alpha_intros simps)))) constrs + (fn {context = ctxt', ...} => + HEADGOAL + (Goal.conjunction_tac THEN_ALL_NEW raw_constr_rsp_tac ctxt' alpha_intros simps))) constrs end (* rsp lemmas for alpha_bn relations *) val rsp_equivp = @{lemma "[|equivp Q; !!x y. R x y ==> Q x y|] ==> (R ===> R ===> (=)) Q Q" by (simp only: rel_fun_def equivp_def, metis)} (* we have to reorder the alpha_bn_imps theorems in order to be in order with alpha_bn_trms *) fun raw_alpha_bn_rsp alpha_result alpha_bn_equivp alpha_bn_imps = let val AlphaResult {alpha_bn_trms, ...} = alpha_result fun mk_map thm = thm |> `Thm.prop_of |>> List.last o snd o strip_comb |>> HOLogic.dest_Trueprop |>> head_of |>> fst o dest_Const val alpha_bn_imps' = map (lookup (map mk_map alpha_bn_imps) o fst o dest_Const) alpha_bn_trms fun mk_thm thm1 thm2 = (Thm.forall_intr_vars thm2) COMP (thm1 RS rsp_equivp) in map2 mk_thm alpha_bn_equivp alpha_bn_imps' end (* rsp for permute_bn functions *) val perm_bn_rsp = @{lemma "(\x y p. R x y \ R (f p x) (f p y)) \ ((=) ===> R ===> R) f f" by (simp add: rel_fun_def)} fun raw_prove_perm_bn_tac alpha_result simps ctxt = SUBPROOF (fn {prems, context = ctxt', ...} => let val AlphaResult {alpha_names, alpha_bn_names, alpha_intros, ...} = alpha_result val prems' = flat (map Old_Datatype_Aux.split_conj_thm prems) val prems'' = map (transform_prem1 ctxt' (alpha_names @ alpha_bn_names)) prems' in HEADGOAL (simp_tac (put_simpset HOL_basic_ss ctxt' addsimps (simps @ prems')) THEN' TRY o REPEAT_ALL_NEW (FIRST' [ resolve_tac ctxt' @{thms TrueI}, resolve_tac ctxt' @{thms conjI}, resolve_tac ctxt' @{thms refl}, resolve_tac ctxt' prems', resolve_tac ctxt' prems'', resolve_tac ctxt' alpha_intros ])) end) ctxt fun raw_perm_bn_rsp ctxt alpha_result perm_bns simps = let val AlphaResult {alpha_trms, alpha_tys, alpha_bn_trms, alpha_bn_tys, alpha_raw_induct, ...} = alpha_result val perm_bn_ty = range_type o range_type o fastype_of val ty_assoc = (alpha_tys @ alpha_bn_tys) ~~ (alpha_trms @ alpha_bn_trms) val ([p], ctxt') = Variable.variant_fixes ["p"] ctxt val p = Free (p, \<^Type>\perm\) fun mk_prop t = let val alpha_trm = lookup ty_assoc (perm_bn_ty t) in (alpha_trm, fn (x, y) => alpha_trm $ (t $ p $ x) $ (t $ p $ y)) end val goals = map mk_prop perm_bns in alpha_prove (alpha_trms @ alpha_bn_trms) goals alpha_raw_induct (raw_prove_perm_bn_tac alpha_result simps) ctxt |> Proof_Context.export ctxt' ctxt |> map (atomize ctxt) |> map single |> map (curry (op OF) perm_bn_rsp) end (* transformation of the natural rsp-lemmas into standard form *) val fun_rsp = @{lemma "(\x y. R x y \ f x = f y) \ (R ===> (=)) f f" by (simp add: rel_fun_def)} fun mk_funs_rsp ctxt thm = thm |> atomize ctxt |> single |> curry (op OF) fun_rsp val permute_rsp = @{lemma "(\x y p. R x y \ R (permute p x) (permute p y)) ==> ((=) ===> R ===> R) permute permute" by (simp add: rel_fun_def)} fun mk_alpha_permute_rsp ctxt thm = thm |> atomize ctxt |> single |> curry (op OF) permute_rsp end (* structure *) diff --git a/thys/Nominal2/nominal_dt_quot.ML b/thys/Nominal2/nominal_dt_quot.ML --- a/thys/Nominal2/nominal_dt_quot.ML +++ b/thys/Nominal2/nominal_dt_quot.ML @@ -1,751 +1,754 @@ (* Title: nominal_dt_alpha.ML Author: Christian Urban Author: Cezary Kaliszyk Performing quotient constructions, lifting theorems and deriving support properties for the quotient types. *) signature NOMINAL_DT_QUOT = sig val define_qtypes: (string list * binding * mixfix) list -> typ list -> term list -> thm list -> local_theory -> Quotient_Info.quotients list * local_theory val define_qconsts: typ list -> (string * term * mixfix * thm) list -> local_theory -> Quotient_Info.quotconsts list * local_theory val define_qperms: typ list -> string list -> (string * sort) list -> (string * term * mixfix * thm) list -> thm list -> local_theory -> local_theory val define_qsizes: typ list -> string list -> (string * sort) list -> (string * term * mixfix * thm) list -> local_theory -> local_theory val lift_thms: typ list -> thm list -> thm list -> Proof.context -> thm list * Proof.context val prove_supports: Proof.context -> thm list -> term list -> thm list val prove_fsupp: Proof.context -> typ list -> thm -> thm list -> thm list val fs_instance: typ list -> string list -> (string * sort) list -> thm list -> local_theory -> local_theory val prove_fv_supp: typ list -> term list -> term list -> term list -> term list -> thm list -> thm list -> thm list -> thm list -> thm -> bclause list list -> Proof.context -> thm list val prove_bns_finite: typ list -> term list -> thm -> thm list -> Proof.context -> thm list val prove_perm_bn_alpha_thms: typ list -> term list -> term list -> thm -> thm list -> thm list -> thm list -> Proof.context -> thm list val prove_permute_bn_thms: typ list -> term list -> term list -> thm -> thm list -> thm list -> thm list -> Proof.context -> thm list val prove_strong_exhausts: Proof.context -> thm list -> bclause list list list -> thm list -> thm list -> thm list -> thm list -> thm list -> thm list val prove_strong_induct: Proof.context -> thm -> thm list -> thm list -> bclause list list list -> thm list end structure Nominal_Dt_Quot: NOMINAL_DT_QUOT = struct open Nominal_Permeq fun lookup xs x = the (AList.lookup (op=) xs x) (* defines the quotient types *) fun define_qtypes qtys_descr alpha_tys alpha_trms alpha_equivp_thms lthy = let val qty_args1 = map2 (fn ty => fn trm => (ty, trm, false)) alpha_tys alpha_trms val qty_args2 = map2 (fn descr => fn args1 => (descr, args1, (NONE, NONE))) qtys_descr qty_args1 val qty_args3 = qty_args2 ~~ alpha_equivp_thms in fold_map (Quotient_Type.add_quotient_type {overloaded = false}) qty_args3 lthy end (* a wrapper for lifting a raw constant *) fun lift_raw_const qtys (qconst_name, rconst, mx', rsp_thm) lthy = let val rty = fastype_of rconst val qty = Quotient_Term.derive_qtyp lthy qtys rty val lhs_raw = Free (qconst_name, qty) val rhs_raw = rconst val raw_var = (Binding.name qconst_name, NONE, mx') val ((binding, _, mx), ctxt) = Proof_Context.cert_var raw_var lthy val lhs = Syntax.check_term ctxt lhs_raw val rhs = Syntax.check_term ctxt rhs_raw val (lhs_str, lhs_ty) = dest_Free lhs handle TERM _ => error "Constant already defined." val _ = if null (strip_abs_vars rhs) then () else error "The definiens cannot be an abstraction" val _ = if is_Const rhs then () else warning "The definiens is not a constant" in Quotient_Def.add_quotient_def (((binding, mx), Binding.empty_atts), (lhs, rhs)) rsp_thm ctxt end (* defines quotient constants *) fun define_qconsts qtys consts_specs lthy = let val (qconst_infos, lthy') = fold_map (lift_raw_const qtys) consts_specs lthy val phi = Proof_Context.export_morphism lthy' (Proof_Context.transfer (Proof_Context.theory_of lthy') lthy) in (map (Quotient_Info.transform_quotconsts phi) qconst_infos, lthy') end (* defines the quotient permutations and proves pt-class *) fun define_qperms qtys qfull_ty_names tvs perm_specs raw_perm_laws lthy = let val lthy1 = lthy |> Local_Theory.exit_global |> Class.instantiation (qfull_ty_names, tvs, @{sort pt}) val (_, lthy2) = define_qconsts qtys perm_specs lthy1 val ((_, raw_perm_laws'), lthy3) = Variable.importT raw_perm_laws lthy2 val lifted_perm_laws = map (Quotient_Tacs.lifted lthy3 qtys []) raw_perm_laws' |> Variable.exportT lthy3 lthy2 fun tac ctxt = Class.intro_classes_tac ctxt [] THEN (ALLGOALS (resolve_tac ctxt lifted_perm_laws)) in lthy2 |> Class.prove_instantiation_exit tac |> Named_Target.theory_init end (* defines the size functions and proves size-class *) fun define_qsizes qtys qfull_ty_names tvs size_specs lthy = let fun tac ctxt = Class.intro_classes_tac ctxt [] in lthy |> Local_Theory.exit_global |> Class.instantiation (qfull_ty_names, tvs, @{sort size}) |> snd o (define_qconsts qtys size_specs) |> Class.prove_instantiation_exit tac |> Named_Target.theory_init end (* lifts a theorem and cleans all "_raw" parts from variable names *) local val any = Scan.one (Symbol.not_eof) val raw = Scan.this_string "_raw" val exclude = Scan.repeat (Scan.unless raw any) --| raw >> implode val parser = Scan.repeat (exclude || any) in fun unraw_str s = s |> raw_explode |> Scan.finite Symbol.stopper parser >> implode |> fst end fun unraw_vars_thm ctxt thm = let fun unraw_var_str ((s, i), T) = ((unraw_str s, i), T) val vars = Term.add_vars (Thm.prop_of thm) [] val vars' = map (Thm.cterm_of ctxt o Var o unraw_var_str) vars in Thm.instantiate (TVars.empty, Vars.make (vars ~~ vars')) thm end fun unraw_bounds_thm th = let val trm = Thm.prop_of th val trm' = Term.map_abs_vars unraw_str trm in Thm.rename_boundvars trm trm' th end fun lift_thms qtys simps thms ctxt = (map (Quotient_Tacs.lifted ctxt qtys simps #> unraw_bounds_thm #> unraw_vars_thm ctxt #> Drule.zero_var_indexes) thms, ctxt) fun mk_supports_goal ctxt qtrm = let val vs = fresh_args ctxt qtrm val rhs = list_comb (qtrm, vs) val lhs = fold (curry HOLogic.mk_prod) vs @{term "()"} |> mk_supp in mk_supports lhs rhs |> HOLogic.mk_Trueprop end fun supports_tac ctxt perm_simps = let val simpset1 = put_simpset HOL_basic_ss ctxt addsimps @{thms supports_def fresh_def[symmetric]} val simpset2 = put_simpset HOL_ss ctxt addsimps @{thms swap_fresh_fresh fresh_Pair} in EVERY' [ simp_tac simpset1, eqvt_tac ctxt (eqvt_strict_config addpres perm_simps), simp_tac simpset2 ] end fun prove_supports_single ctxt perm_simps qtrm = let val goal = mk_supports_goal ctxt qtrm val ctxt' = Proof_Context.augment goal ctxt in Goal.prove ctxt' [] [] goal - (K (HEADGOAL (supports_tac ctxt perm_simps))) + (fn {context = goal_ctxt, ...} => HEADGOAL (supports_tac goal_ctxt perm_simps)) |> singleton (Proof_Context.export ctxt' ctxt) end fun prove_supports ctxt perm_simps qtrms = map (prove_supports_single ctxt perm_simps) qtrms (* finite supp lemmas for qtypes *) fun prove_fsupp ctxt qtys qinduct qsupports_thms = let val (vs, ctxt') = Variable.variant_fixes (replicate (length qtys) "x") ctxt val goals = vs ~~ qtys |> map Free |> map (mk_finite o mk_supp) |> foldr1 (HOLogic.mk_conj) |> HOLogic.mk_Trueprop val tac = EVERY' [ resolve_tac ctxt' @{thms supports_finite}, resolve_tac ctxt' qsupports_thms, asm_simp_tac (put_simpset HOL_ss ctxt' addsimps @{thms finite_supp supp_Pair finite_Un}) ] in Goal.prove ctxt' [] [] goals - (K (HEADGOAL (resolve_tac ctxt' [qinduct] THEN_ALL_NEW tac))) + (fn {context = goal_ctxt, ...} => + HEADGOAL (resolve_tac goal_ctxt [qinduct] THEN_ALL_NEW tac)) |> singleton (Proof_Context.export ctxt' ctxt) |> Old_Datatype_Aux.split_conj_thm |> map zero_var_indexes end (* finite supp instances *) fun fs_instance qtys qfull_ty_names tvs qfsupp_thms lthy = let val lthy1 = lthy |> Local_Theory.exit_global |> Class.instantiation (qfull_ty_names, tvs, @{sort fs}) fun tac ctxt = Class.intro_classes_tac ctxt [] THEN (ALLGOALS (resolve_tac ctxt qfsupp_thms)) in lthy1 |> Class.prove_instantiation_exit tac |> Named_Target.theory_init end (* proves that fv and fv_bn equals supp *) fun gen_mk_goals fv supp = let val arg_ty = fastype_of fv |> domain_type in (arg_ty, fn x => HOLogic.mk_eq (fv $ x, supp x)) end fun mk_fvs_goals fv = gen_mk_goals fv mk_supp fun mk_fv_bns_goals fv_bn alpha_bn = gen_mk_goals fv_bn (mk_supp_rel alpha_bn) fun add_simpset ctxt thms = put_simpset HOL_basic_ss ctxt addsimps thms fun symmetric thms = map (fn thm => thm RS @{thm sym}) thms val supp_Abs_set = @{thms supp_Abs(1)[symmetric]} val supp_Abs_res = @{thms supp_Abs(2)[symmetric]} val supp_Abs_lst = @{thms supp_Abs(3)[symmetric]} fun mk_supp_abs ctxt (BC (Set, _, _)) = EqSubst.eqsubst_tac ctxt [1] supp_Abs_set | mk_supp_abs ctxt (BC (Res, _, _)) = EqSubst.eqsubst_tac ctxt [1] supp_Abs_res | mk_supp_abs ctxt (BC (Lst, _, _)) = EqSubst.eqsubst_tac ctxt [1] supp_Abs_lst fun mk_supp_abs_tac ctxt [] = [] | mk_supp_abs_tac ctxt (BC (_, [], _)::xs) = mk_supp_abs_tac ctxt xs | mk_supp_abs_tac ctxt (bc::xs) = (DETERM o mk_supp_abs ctxt bc)::mk_supp_abs_tac ctxt xs fun mk_bn_supp_abs_tac ctxt trm = trm |> fastype_of |> body_type |> (fn ty => case ty of @{typ "atom set"} => simp_tac (add_simpset ctxt supp_Abs_set) | @{typ "atom list"} => simp_tac (add_simpset ctxt supp_Abs_lst) | _ => raise TERM ("mk_bn_supp_abs_tac", [trm])) val thms1 = @{thms supp_Pair supp_eqvt[symmetric] Un_assoc conj_assoc} val thms2 = @{thms de_Morgan_conj Collect_disj_eq finite_Un} val thms3 = @{thms alphas prod_alpha_def prod_fv.simps rel_prod_conv permute_prod_def prod.rec prod.case prod.inject not_True_eq_False empty_def[symmetric] finite.emptyI} fun prove_fv_supp qtys qtrms fvs fv_bns alpha_bns fv_simps eq_iffs perm_simps fv_bn_eqvts qinduct bclausess ctxt = let val goals1 = map mk_fvs_goals fvs val goals2 = map2 mk_fv_bns_goals fv_bns alpha_bns fun tac ctxt = SUBGOAL (fn (goal, i) => let val (fv_fun, arg) = goal |> Envir.eta_contract |> Logic.strip_assums_concl |> HOLogic.dest_Trueprop |> fst o HOLogic.dest_eq |> dest_comb val supp_abs_tac = case (AList.lookup (op=) (qtrms ~~ bclausess) (head_of arg)) of SOME bclauses => EVERY' (mk_supp_abs_tac ctxt bclauses) | NONE => mk_bn_supp_abs_tac ctxt fv_fun val eqvt_rconfig = eqvt_relaxed_config addpres (perm_simps @ fv_bn_eqvts) in EVERY' [ TRY o asm_full_simp_tac (add_simpset ctxt (@{thm supp_Pair[symmetric]}::fv_simps)), TRY o supp_abs_tac, TRY o simp_tac (add_simpset ctxt @{thms supp_def supp_rel_def}), TRY o eqvt_tac ctxt eqvt_rconfig, TRY o simp_tac (add_simpset ctxt (@{thms Abs_eq_iff} @ eq_iffs)), TRY o asm_full_simp_tac (add_simpset ctxt thms3), TRY o simp_tac (add_simpset ctxt thms2), TRY o asm_full_simp_tac (add_simpset ctxt (thms1 @ (symmetric fv_bn_eqvts)))] i end) in induct_prove qtys (goals1 @ goals2) qinduct tac ctxt |> map (atomize ctxt) |> map (simplify (put_simpset HOL_basic_ss ctxt addsimps @{thms fun_eq_iff[symmetric]})) end fun prove_bns_finite qtys qbns qinduct qbn_simps ctxt = let fun mk_goal qbn = let val arg_ty = domain_type (fastype_of qbn) val finite = @{term "finite :: atom set => bool"} in (arg_ty, fn x => finite $ (to_set (qbn $ x))) end val props = map mk_goal qbns val ss_tac = asm_full_simp_tac (put_simpset HOL_basic_ss ctxt addsimps (qbn_simps @ @{thms set_simps set_append finite_insert finite.emptyI finite_Un})) in induct_prove qtys props qinduct (K ss_tac) ctxt end fun prove_perm_bn_alpha_thms qtys qperm_bns alpha_bns qinduct qperm_bn_simps qeq_iffs qalpha_refls ctxt = let val ([p], ctxt') = Variable.variant_fixes ["p"] ctxt val p = Free (p, \<^Type>\perm\) fun mk_goal qperm_bn alpha_bn = let val arg_ty = domain_type (fastype_of alpha_bn) in (arg_ty, fn x => (mk_id (Abs ("", arg_ty, alpha_bn $ Bound 0 $ (qperm_bn $ p $ Bound 0)))) $ x) end val props = map2 mk_goal qperm_bns alpha_bns val ss = @{thm id_def}::qperm_bn_simps @ qeq_iffs @ qalpha_refls val ss_tac = asm_full_simp_tac (put_simpset HOL_ss ctxt addsimps ss) in induct_prove qtys props qinduct (K ss_tac) ctxt' |> Proof_Context.export ctxt' ctxt |> map (simplify (put_simpset HOL_basic_ss ctxt addsimps @{thms id_def})) end fun prove_permute_bn_thms qtys qbns qperm_bns qinduct qperm_bn_simps qbn_defs qbn_eqvts ctxt = let val ([p], ctxt') = Variable.variant_fixes ["p"] ctxt val p = Free (p, \<^Type>\perm\) fun mk_goal qbn qperm_bn = let val arg_ty = domain_type (fastype_of qbn) in (arg_ty, fn x => (mk_id (Abs ("", arg_ty, HOLogic.mk_eq (mk_perm p (qbn $ Bound 0), qbn $ (qperm_bn $ p $ Bound 0)))) $ x)) end val props = map2 mk_goal qbns qperm_bns val ss = @{thm id_def}::qperm_bn_simps @ qbn_defs val ss_tac = EVERY' [asm_full_simp_tac (put_simpset HOL_basic_ss ctxt' addsimps ss), TRY o eqvt_tac ctxt' (eqvt_strict_config addpres qbn_eqvts), TRY o asm_full_simp_tac (put_simpset HOL_basic_ss ctxt')] in induct_prove qtys props qinduct (K ss_tac) ctxt' |> Proof_Context.export ctxt' ctxt |> map (simplify (put_simpset HOL_basic_ss ctxt addsimps @{thms id_def})) end (*** proves strong exhauts theorems ***) (* fixme: move into nominal_library *) fun abs_const bmode ty = let val (const_name, binder_ty, abs_ty) = case bmode of Lst => (@{const_name "Abs_lst"}, @{typ "atom list"}, @{type_name abs_lst}) | Set => (@{const_name "Abs_set"}, @{typ "atom set"}, @{type_name abs_set}) | Res => (@{const_name "Abs_res"}, @{typ "atom set"}, @{type_name abs_res}) in Const (const_name, [binder_ty, ty] ---> Type (abs_ty, [ty])) end fun mk_abs bmode trm1 trm2 = abs_const bmode (fastype_of trm2) $ trm1 $ trm2 fun is_abs_eq thm = let fun is_abs trm = case head_of trm of \<^Const_>\Abs_set _\ => true | \<^Const_>\Abs_lst _\ => true | \<^Const_>\Abs_res _\ => true | _ => false in thm |> Thm.prop_of |> HOLogic.dest_Trueprop |> HOLogic.dest_eq |> fst |> is_abs end (* adds a freshness condition to the assumptions *) fun mk_ecase_prems lthy c (params, prems, concl) bclauses = let val tys = map snd params val binders = get_all_binders bclauses fun prep_binder (opt, i) = let val t = Bound (length tys - i - 1) in case opt of NONE => setify_ty lthy (nth tys i) t | SOME bn => to_set_ty (fastype_of1 (tys, bn $ t)) (bn $ t) end val prems' = case binders of [] => prems (* case: no binders *) | _ => binders (* case: binders *) |> map prep_binder |> fold_union_env tys |> (fn t => mk_fresh_star t c) |> (fn t => HOLogic.mk_Trueprop t :: prems) in mk_full_horn params prems' concl end (* derives the freshness theorem that there exists a p, such that (p o as) #* (c, t1,..., tn) *) fun fresh_thm ctxt c parms binders bn_finite_thms = let fun prep_binder (opt, i) = case opt of NONE => setify ctxt (nth parms i) | SOME bn => to_set (bn $ (nth parms i)) fun prep_binder2 (opt, i) = case opt of NONE => atomify ctxt (nth parms i) | SOME bn => bn $ (nth parms i) val rhs = HOLogic.mk_tuple ([c] @ parms @ (map prep_binder2 binders)) val lhs = binders |> map prep_binder |> fold_union |> mk_perm (Bound 0) val goal = mk_fresh_star lhs rhs |> mk_exists ("p", \<^Type>\perm\) |> HOLogic.mk_Trueprop val ss = bn_finite_thms @ @{thms supp_Pair finite_supp finite_sets_supp} @ @{thms finite.intros finite_Un finite_set finite_fset} in Goal.prove ctxt [] [] goal - (K (HEADGOAL (resolve_tac ctxt @{thms at_set_avoiding1} - THEN_ALL_NEW (simp_tac (put_simpset HOL_ss ctxt addsimps ss))))) + (fn {context = ctxt', ...} => + HEADGOAL (resolve_tac ctxt' @{thms at_set_avoiding1} + THEN_ALL_NEW (simp_tac (put_simpset HOL_ss ctxt' addsimps ss)))) end (* derives an abs_eq theorem of the form Exists q. [as].x = [p o as].(q o x) for non-recursive binders Exists q. [as].x = [q o as].(q o x) for recursive binders *) fun abs_eq_thm ctxt fprops p parms bn_eqvt permute_bns (bclause as (BC (bmode, binders, bodies))) = case binders of [] => [] | _ => let val rec_flag = is_recursive_binder bclause val binder_trm = comb_binders ctxt bmode parms binders val body_trm = foldl1 HOLogic.mk_prod (map (nth parms) bodies) val abs_lhs = mk_abs bmode binder_trm body_trm val abs_rhs = if rec_flag then mk_abs bmode (mk_perm (Bound 0) binder_trm) (mk_perm (Bound 0) body_trm) else mk_abs bmode (mk_perm p binder_trm) (mk_perm (Bound 0) body_trm) val abs_eq = HOLogic.mk_eq (abs_lhs, abs_rhs) val peq = HOLogic.mk_eq (mk_perm (Bound 0) binder_trm, mk_perm p binder_trm) val goal = HOLogic.mk_conj (abs_eq, peq) |> (fn t => HOLogic.mk_exists ("q", @{typ "perm"}, t)) |> HOLogic.mk_Trueprop val ss = fprops @ @{thms set_simps set_append union_eqvt} @ @{thms fresh_star_Un fresh_star_Pair fresh_star_list fresh_star_singleton fresh_star_fset fresh_star_set} - val tac1 = + fun tac1 ctxt' = if rec_flag - then resolve_tac ctxt @{thms Abs_rename_set' Abs_rename_res' Abs_rename_lst'} - else resolve_tac ctxt @{thms Abs_rename_set Abs_rename_res Abs_rename_lst} + then resolve_tac ctxt' @{thms Abs_rename_set' Abs_rename_res' Abs_rename_lst'} + else resolve_tac ctxt' @{thms Abs_rename_set Abs_rename_res Abs_rename_lst} - val tac2 = - EVERY' [simp_tac (put_simpset HOL_basic_ss ctxt addsimps ss), - TRY o simp_tac (put_simpset HOL_ss ctxt)] + fun tac2 ctxt' = + EVERY' [simp_tac (put_simpset HOL_basic_ss ctxt' addsimps ss), + TRY o simp_tac (put_simpset HOL_ss ctxt')] in - [ Goal.prove ctxt [] [] goal (K (HEADGOAL (tac1 THEN_ALL_NEW tac2))) + [ Goal.prove ctxt [] [] goal (fn {context = ctxt', ...} => + HEADGOAL (tac1 ctxt' THEN_ALL_NEW tac2 ctxt')) |> (if rec_flag then eqvt_rule ctxt (eqvt_strict_config addpres bn_eqvt) else eqvt_rule ctxt (eqvt_strict_config addpres permute_bns)) ] end val setify = @{lemma "xs = ys ==> set xs = set ys" by simp} -fun case_tac ctxt c bn_finite_thms eq_iff_thms bn_eqvt permute_bns perm_bn_alphas +fun case_tac goal_ctxt c bn_finite_thms eq_iff_thms bn_eqvt permute_bns perm_bn_alphas prems bclausess qexhaust_thm = let fun aux_tac prem bclauses = case (get_all_binders bclauses) of - [] => EVERY' [resolve_tac ctxt [prem], assume_tac ctxt] - | binders => Subgoal.SUBPROOF (fn {params, prems, concl, context = ctxt, ...} => + [] => EVERY' [resolve_tac goal_ctxt [prem], assume_tac goal_ctxt] + | binders => Subgoal.SUBPROOF (fn {params, prems, concl, context = goal_ctxt1, ...} => let val parms = map (Thm.term_of o snd) params - val fthm = fresh_thm ctxt c parms binders bn_finite_thms + val fthm = fresh_thm goal_ctxt1 c parms binders bn_finite_thms val ss = @{thms fresh_star_Pair union_eqvt fresh_star_Un} val (([(_, fperm)], fprops), ctxt') = Obtain.result - (fn ctxt' => EVERY1 [eresolve_tac ctxt [exE], - full_simp_tac (put_simpset HOL_basic_ss ctxt' addsimps ss), - REPEAT o (eresolve_tac ctxt @{thms conjE})]) [fthm] ctxt + (fn goal_ctxt2 => + EVERY1 [eresolve_tac goal_ctxt2 [exE], + full_simp_tac (put_simpset HOL_basic_ss goal_ctxt2 addsimps ss), + REPEAT o (eresolve_tac goal_ctxt2 @{thms conjE})]) [fthm] goal_ctxt1 val abs_eq_thms = flat (map (abs_eq_thm ctxt' fprops (Thm.term_of fperm) parms bn_eqvt permute_bns) bclauses) val ((_, eqs), ctxt'') = Obtain.result - (fn ctxt'' => EVERY1 - [ REPEAT o (eresolve_tac ctxt @{thms exE}), - REPEAT o (eresolve_tac ctxt @{thms conjE}), - REPEAT o (dresolve_tac ctxt [setify]), - full_simp_tac (put_simpset HOL_basic_ss ctxt'' + (fn goal_ctxt3 => EVERY1 + [ REPEAT o (eresolve_tac goal_ctxt3 @{thms exE}), + REPEAT o (eresolve_tac goal_ctxt3 @{thms conjE}), + REPEAT o (dresolve_tac goal_ctxt3 [setify]), + full_simp_tac (put_simpset HOL_basic_ss goal_ctxt3 addsimps @{thms set_append set_simps})]) abs_eq_thms ctxt' val (abs_eqs, peqs) = split_filter is_abs_eq eqs val fprops' = map (eqvt_rule ctxt'' (eqvt_strict_config addpres permute_bns)) fprops @ map (eqvt_rule ctxt'' (eqvt_strict_config addpres bn_eqvt)) fprops (* for freshness conditions *) - val tac1 = SOLVED' (EVERY' - [ simp_tac (put_simpset HOL_basic_ss ctxt'' addsimps peqs), - rewrite_goal_tac ctxt'' (@{thms fresh_star_Un[THEN eq_reflection]}), - conj_tac ctxt'' (DETERM o resolve_tac ctxt'' fprops') ]) + fun tac1 ctxt = SOLVED' (EVERY' + [ simp_tac (put_simpset HOL_basic_ss ctxt addsimps peqs), + rewrite_goal_tac ctxt (@{thms fresh_star_Un[THEN eq_reflection]}), + conj_tac ctxt (DETERM o resolve_tac ctxt fprops') ]) (* for equalities between constructors *) - val tac2 = SOLVED' (EVERY' + fun tac2 ctxt = SOLVED' (EVERY' [ resolve_tac ctxt [@{thm ssubst} OF prems], - rewrite_goal_tac ctxt'' (map safe_mk_equiv eq_iff_thms), - rewrite_goal_tac ctxt'' (map safe_mk_equiv abs_eqs), - conj_tac ctxt'' (DETERM o resolve_tac ctxt'' (@{thms refl} @ perm_bn_alphas)) ]) + rewrite_goal_tac ctxt (map safe_mk_equiv eq_iff_thms), + rewrite_goal_tac ctxt (map safe_mk_equiv abs_eqs), + conj_tac ctxt (DETERM o resolve_tac ctxt (@{thms refl} @ perm_bn_alphas)) ]) (* proves goal "P" *) val side_thm = Goal.prove ctxt'' [] [] (Thm.term_of concl) - (K (EVERY1 [ resolve_tac ctxt'' [prem], RANGE [tac1, tac2] ])) - |> singleton (Proof_Context.export ctxt'' ctxt) + (fn {context = goal_ctxt4, ...} => + EVERY1 [ resolve_tac goal_ctxt4 [prem], RANGE [tac1 goal_ctxt4, tac2 goal_ctxt4]]) + |> singleton (Proof_Context.export ctxt'' goal_ctxt1) in - resolve_tac ctxt [side_thm] 1 - end) ctxt + resolve_tac goal_ctxt1 [side_thm] 1 + end) goal_ctxt in - EVERY1 [resolve_tac ctxt [qexhaust_thm], RANGE (map2 aux_tac prems bclausess)] + EVERY1 [resolve_tac goal_ctxt [qexhaust_thm], RANGE (map2 aux_tac prems bclausess)] end fun prove_strong_exhausts lthy exhausts bclausesss bn_finite_thms eq_iff_thms bn_eqvt permute_bns perm_bn_alphas = let val ((_, exhausts'), lthy') = Variable.import true exhausts lthy val ([c, a], lthy'') = Variable.variant_fixes ["c", "'a"] lthy' val c = Free (c, TFree (a, @{sort fs})) val (ecases, main_concls) = exhausts' (* ecases are of the form (params, prems, concl) *) |> map Thm.prop_of |> map Logic.strip_horn |> split_list val ecases' = (map o map) strip_full_horn ecases val premss = (map2 o map2) (mk_ecase_prems lthy'' c) ecases' bclausesss - fun tac bclausess exhaust {prems, context} = - case_tac context c bn_finite_thms eq_iff_thms bn_eqvt permute_bns perm_bn_alphas - prems bclausess exhaust - fun prove prems bclausess exhaust concl = - Goal.prove lthy'' [] prems concl (tac bclausess exhaust) + Goal.prove lthy'' [] prems concl (fn {prems, context = goal_ctxt} => + case_tac goal_ctxt c bn_finite_thms eq_iff_thms bn_eqvt permute_bns perm_bn_alphas + prems bclausess exhaust) in map4 prove premss bclausesss exhausts' main_concls |> Proof_Context.export lthy'' lthy end (** strong induction theorems **) fun add_c_prop c c_ty trm = let val (P, arg) = dest_comb trm val (P_name, P_ty) = dest_Free P val (ty_args, bool) = strip_type P_ty in Free (P_name, (c_ty :: ty_args) ---> bool) $ c $ arg end fun add_qnt_c_prop c_name c_ty trm = trm |> HOLogic.dest_Trueprop |> incr_boundvars 1 |> add_c_prop (Bound 0) c_ty |> HOLogic.mk_Trueprop |> mk_all (c_name, c_ty) fun prep_prem lthy c c_name c_ty bclauses (params, prems, concl) = let val tys = map snd params val binders = get_all_binders bclauses fun prep_binder (opt, i) = let val t = Bound (length tys - i - 1) in case opt of NONE => setify_ty lthy (nth tys i) t | SOME bn => to_set_ty (fastype_of1 (tys, bn $ t)) (bn $ t) end val prems' = prems |> map (incr_boundvars 1) |> map (add_qnt_c_prop c_name c_ty) val prems'' = case binders of [] => prems' (* case: no binders *) | _ => binders (* case: binders *) |> map prep_binder |> fold_union_env tys |> incr_boundvars 1 |> (fn t => mk_fresh_star_ty c_ty t (Bound 0)) |> (fn t => HOLogic.mk_Trueprop t :: prems') val concl' = concl |> HOLogic.dest_Trueprop |> incr_boundvars 1 |> add_c_prop (Bound 0) c_ty |> HOLogic.mk_Trueprop in mk_full_horn (params @ [(c_name, c_ty)]) prems'' concl' end fun prove_strong_induct lthy induct exhausts size_thms bclausesss = let - val ((_, [induct']), lthy') = Variable.import true [induct] lthy + val ((_, [induct']), ctxt') = Variable.import true [induct] lthy - val ([c_name, a], lthy'') = Variable.variant_fixes ["c", "'a"] lthy' + val ([c_name, a], ctxt'') = Variable.variant_fixes ["c", "'a"] ctxt' val c_ty = TFree (a, @{sort fs}) val c = Free (c_name, c_ty) val (prems, concl) = induct' |> Thm.prop_of |> Logic.strip_horn val concls = concl |> HOLogic.dest_Trueprop |> HOLogic.dest_conj |> map (add_c_prop c c_ty) |> map HOLogic.mk_Trueprop val prems' = prems |> map strip_full_horn - |> map2 (prep_prem lthy'' c c_name c_ty) (flat bclausesss) + |> map2 (prep_prem ctxt'' c c_name c_ty) (flat bclausesss) - fun pat_tac ctxt thm = - Subgoal.FOCUS (fn {params, context = ctxt', ...} => + fun pat_tac goal_ctxt thm = + Subgoal.FOCUS (fn {params, context = goal_ctxt1, ...} => let val ty_parms = map (fn (_, ct) => (fastype_of (Thm.term_of ct), ct)) params val vs = Term.add_vars (Thm.prop_of thm) [] val vs_tys = map (Type.legacy_freeze_type o snd) vs val assigns = map (lookup ty_parms) vs_tys - val thm' = infer_instantiate ctxt' (map #1 vs ~~ assigns) thm + val thm' = infer_instantiate goal_ctxt1 (map #1 vs ~~ assigns) thm in - resolve_tac ctxt' [thm'] 1 - end) ctxt - THEN_ALL_NEW asm_full_simp_tac (put_simpset HOL_basic_ss ctxt) + resolve_tac goal_ctxt1 [thm'] 1 + end) goal_ctxt + THEN_ALL_NEW asm_full_simp_tac (put_simpset HOL_basic_ss goal_ctxt) fun size_simp_tac ctxt = simp_tac (put_simpset size_ss ctxt addsimps (@{thms comp_def snd_conv} @ size_thms)) in - Goal.prove_common lthy'' NONE [] prems' concls - (fn {prems, context = ctxt} => - Induction_Schema.induction_schema_tac ctxt prems - THEN RANGE (map (pat_tac ctxt) exhausts) 1 - THEN prove_termination_ind ctxt 1 - THEN ALLGOALS (size_simp_tac ctxt)) - |> Proof_Context.export lthy'' lthy + Goal.prove_common ctxt'' NONE [] prems' concls + (fn {prems, context = goal_ctxt} => + Induction_Schema.induction_schema_tac goal_ctxt prems + THEN RANGE (map (pat_tac goal_ctxt) exhausts) 1 + THEN prove_termination_ind goal_ctxt 1 + THEN ALLGOALS (size_simp_tac goal_ctxt)) + |> Proof_Context.export ctxt'' lthy end end (* structure *) diff --git a/thys/Nominal2/nominal_dt_rawfuns.ML b/thys/Nominal2/nominal_dt_rawfuns.ML --- a/thys/Nominal2/nominal_dt_rawfuns.ML +++ b/thys/Nominal2/nominal_dt_rawfuns.ML @@ -1,557 +1,557 @@ (* Title: nominal_dt_rawfuns.ML Author: Cezary Kaliszyk Author: Christian Urban Definitions of the raw fv, fv_bn and permute functions. *) signature NOMINAL_DT_RAWFUNS = sig val get_all_binders: bclause list -> (term option * int) list val is_recursive_binder: bclause -> bool val define_raw_bns: raw_dt_info -> (binding * typ option * mixfix) list -> Specification.multi_specs -> local_theory -> (term list * thm list * bn_info list * thm list * local_theory) val define_raw_fvs: raw_dt_info -> bn_info list -> bclause list list list -> Proof.context -> term list * term list * thm list * thm list * local_theory val define_raw_bn_perms: raw_dt_info -> bn_info list -> local_theory -> (term list * thm list * local_theory) val define_raw_perms: raw_dt_info -> local_theory -> (term list * thm list * thm list) * local_theory val raw_prove_eqvt: term list -> thm list -> thm list -> Proof.context -> thm list end structure Nominal_Dt_RawFuns: NOMINAL_DT_RAWFUNS = struct open Nominal_Permeq fun get_all_binders bclauses = bclauses |> map (fn (BC (_, binders, _)) => binders) |> flat |> remove_dups (op =) fun is_recursive_binder (BC (_, binders, bodies)) = case (inter (op =) (map snd binders) bodies) of nil => false | _ => true fun lookup xs x = the (AList.lookup (op=) xs x) (** functions that define the raw binding functions **) (* strip_bn_fun takes a rhs of a bn function: this can only contain unions or appends of elements; in case of recursive calls it returns also the applied bn function *) fun strip_bn_fun ctxt args t = let fun aux t = case t of \<^Const_>\sup _ for l r\ => aux l @ aux r | \<^Const_>\append _ for l r\ => aux l @ aux r | \<^Const_>\insert _ for \<^Const_>\atom _ for \x as Var _\\ y\ => (find_index (equal x) args, NONE) :: aux y | \<^Const>\Cons _ for \<^Const_>\atom _ for \x as Var _\\ y\ => (find_index (equal x) args, NONE) :: aux y | \<^Const_>\bot _\ => [] | \<^Const_>\Nil _\ => [] | (f as Const _) $ (x as Var _) => [(find_index (equal x) args, SOME f)] | _ => error ("Unsupported binding function: " ^ (Syntax.string_of_term ctxt t)) in aux t end (** definition of the raw binding functions **) fun prep_bn_info ctxt dt_names dts bn_funs eqs = let fun process_eq eq = let val (lhs, rhs) = eq |> HOLogic.dest_Trueprop |> HOLogic.dest_eq val (bn_fun, [cnstr]) = strip_comb lhs val (_, ty) = dest_Const bn_fun val (ty_name, _) = dest_Type (domain_type ty) val dt_index = find_index (fn x => x = ty_name) dt_names val (cnstr_head, cnstr_args) = strip_comb cnstr val cnstr_name = Long_Name.base_name (fst (dest_Const cnstr_head)) val rhs_elements = strip_bn_fun ctxt cnstr_args rhs in ((bn_fun, dt_index), (cnstr_name, rhs_elements)) end (* order according to constructor names *) fun cntrs_order ((bn, dt_index), data) = let val dt = nth dts dt_index val cts = snd dt val ct_names = map (Binding.name_of o (fn (x, _, _) => x)) cts in (bn, (bn, dt_index, order (op=) ct_names data)) end in eqs |> map process_eq |> AList.group (op=) (* eqs grouped according to bn_functions *) |> map cntrs_order (* inner data ordered according to constructors *) |> order (op=) bn_funs (* ordered according to bn_functions *) end fun define_raw_bns raw_dt_info raw_bn_funs raw_bn_eqs lthy = if null raw_bn_funs then ([], [], [], [], lthy) else let val RawDtInfo {raw_dt_names, raw_dts, raw_inject_thms, raw_distinct_thms, raw_size_thms, ...} = raw_dt_info val lthy1 = lthy |> Local_Theory.begin_nested |> snd |> Function.add_function raw_bn_funs raw_bn_eqs Function_Common.default_config (pat_completeness_simp (raw_inject_thms @ raw_distinct_thms)) |> snd |> Local_Theory.end_nested; val ({fs, simps, inducts, ...}, lthy2) = prove_termination_fun raw_size_thms lthy1 val raw_bn_induct = the inducts val raw_bn_eqs = the simps val raw_bn_info = prep_bn_info lthy raw_dt_names raw_dts fs (map Thm.prop_of raw_bn_eqs) in (fs, raw_bn_eqs, raw_bn_info, raw_bn_induct, lthy2) end (** functions that construct the equations for fv and fv_bn **) fun mk_fv_rhs lthy fv_map fv_bn_map args (BC (bmode, binders, bodies)) = let fun mk_fv_body i = let val arg = nth args i val ty = fastype_of arg in case AList.lookup (op=) fv_map ty of NONE => mk_supp arg | SOME fv => fv $ arg end fun bind_set (NONE, i) = (setify lthy (nth args i), @{term "{}::atom set"}) | bind_set (SOME bn, i) = (bn $ (nth args i), if member (op=) bodies i then @{term "{}::atom set"} else lookup fv_bn_map bn $ (nth args i)) fun bind_lst (NONE, i) = (listify lthy (nth args i), @{term "[]::atom list"}) | bind_lst (SOME bn, i) = (bn $ (nth args i), if member (op=) bodies i then @{term "[]::atom list"} else lookup fv_bn_map bn $ (nth args i)) val (combine_fn, bind_fn) = case bmode of Lst => (fold_append, bind_lst) | Set => (fold_union, bind_set) | Res => (fold_union, bind_set) val t1 = map mk_fv_body bodies val (t2, t3) = binders |> map bind_fn |> split_list |> apply2 combine_fn in mk_union (mk_diff (fold_union t1, to_set t2), to_set t3) end (* in case of fv_bn we have to treat the case special, where an "empty" binding clause is given *) fun mk_fv_bn_rhs lthy fv_map fv_bn_map bn_args args bclause = let fun mk_fv_bn_body i = let val arg = nth args i val ty = fastype_of arg in case AList.lookup (op=) bn_args i of NONE => (case (AList.lookup (op=) fv_map ty) of NONE => mk_supp arg | SOME fv => fv $ arg) | SOME (NONE) => @{term "{}::atom set"} | SOME (SOME bn) => lookup fv_bn_map bn $ arg end in case bclause of BC (_, [], bodies) => fold_union (map mk_fv_bn_body bodies) | _ => mk_fv_rhs lthy fv_map fv_bn_map args bclause end fun mk_fv_eq lthy fv_map fv_bn_map (constr, ty, arg_tys, _) bclauses = let val arg_names = Old_Datatype_Prop.make_tnames arg_tys val args = map Free (arg_names ~~ arg_tys) val fv = lookup fv_map ty val lhs = fv $ list_comb (constr, args) val rhs_trms = map (mk_fv_rhs lthy fv_map fv_bn_map args) bclauses val rhs = fold_union rhs_trms in HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs)) end fun mk_fv_bn_eq lthy bn_trm fv_map fv_bn_map (bn_args, (constr, _, arg_tys, _)) bclauses = let val arg_names = Old_Datatype_Prop.make_tnames arg_tys val args = map Free (arg_names ~~ arg_tys) val fv_bn = lookup fv_bn_map bn_trm val lhs = fv_bn $ list_comb (constr, args) val rhs_trms = map (mk_fv_bn_rhs lthy fv_map fv_bn_map bn_args args) bclauses val rhs = fold_union rhs_trms in HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs)) end fun mk_fv_bn_eqs lthy fv_map fv_bn_map constrs_info bclausesss (bn_trm, bn_n, bn_argss) = let val nth_constrs_info = nth constrs_info bn_n val nth_bclausess = nth bclausesss bn_n in map2 (mk_fv_bn_eq lthy bn_trm fv_map fv_bn_map) (bn_argss ~~ nth_constrs_info) nth_bclausess end fun define_raw_fvs raw_dt_info bn_info bclausesss lthy = let val RawDtInfo {raw_dt_names, raw_tys, raw_cns_info, raw_inject_thms, raw_distinct_thms, raw_size_thms, ...} = raw_dt_info val fv_names = map (prefix "fv_" o Long_Name.base_name) raw_dt_names val fv_tys = map (fn ty => ty --> @{typ "atom set"}) raw_tys val fv_frees = map Free (fv_names ~~ fv_tys); val fv_map = raw_tys ~~ fv_frees val (bns, bn_tys) = split_list (map (fn (bn, i, _) => (bn, i)) bn_info) val bn_names = map (fn bn => Long_Name.base_name (fst (dest_Const bn))) bns val fv_bn_names = map (prefix "fv_") bn_names val fv_bn_arg_tys = map (nth raw_tys) bn_tys val fv_bn_tys = map (fn ty => ty --> @{typ "atom set"}) fv_bn_arg_tys val fv_bn_frees = map Free (fv_bn_names ~~ fv_bn_tys) val fv_bn_map = bns ~~ fv_bn_frees val fv_eqs = map2 (map2 (mk_fv_eq lthy fv_map fv_bn_map)) raw_cns_info bclausesss val fv_bn_eqs = map (mk_fv_bn_eqs lthy fv_map fv_bn_map raw_cns_info bclausesss) bn_info val all_fun_names = map (fn s => (Binding.name s, NONE, NoSyn)) (fv_names @ fv_bn_names) val all_fun_eqs = map (fn t => ((Binding.empty_atts, t), [], [])) (flat fv_eqs @ flat fv_bn_eqs) val lthy' = lthy |> Local_Theory.begin_nested |> snd |> Function.add_function all_fun_names all_fun_eqs Function_Common.default_config (pat_completeness_simp (raw_inject_thms @ raw_distinct_thms)) |> snd |> Local_Theory.end_nested; val ({fs, simps, inducts, ...}, lthy'') = prove_termination_fun raw_size_thms lthy' val morphism = Proof_Context.export_morphism lthy'' (Proof_Context.transfer (Proof_Context.theory_of lthy'') lthy) val simps_exp = map (Morphism.thm morphism) (the simps) val inducts_exp = map (Morphism.thm morphism) (the inducts) val (fvs', fv_bns') = chop (length fv_frees) fs (* grafting the types so that they coincide with the input into the function package *) val fvs'' = map2 (fn t => fn ty => Const (fst (dest_Const t), ty) ) fvs' fv_tys val fv_bns'' = map2 (fn t => fn ty => Const (fst (dest_Const t), ty) ) fv_bns' fv_bn_tys in (fvs'', fv_bns'', simps_exp, inducts_exp, lthy'') end (** definition of raw permute_bn functions **) fun mk_perm_bn_eq_rhs p perm_bn_map bn_args (i, arg) = case AList.lookup (op=) bn_args i of NONE => arg | SOME (NONE) => mk_perm p arg | SOME (SOME bn) => (lookup perm_bn_map bn) $ p $ arg fun mk_perm_bn_eq lthy bn_trm perm_bn_map bn_args (constr, _, arg_tys, _) = let val p = Free ("p", \<^Type>\perm\) val arg_names = Old_Datatype_Prop.make_tnames arg_tys val args = map Free (arg_names ~~ arg_tys) val perm_bn = lookup perm_bn_map bn_trm val lhs = perm_bn $ p $ list_comb (constr, args) val rhs = list_comb (constr, map_index (mk_perm_bn_eq_rhs p perm_bn_map bn_args) args) in HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs)) end fun mk_perm_bn_eqs lthy perm_bn_map cns_info (bn_trm, bn_n, bn_argss) = let val nth_cns_info = nth cns_info bn_n in map2 (mk_perm_bn_eq lthy bn_trm perm_bn_map) bn_argss nth_cns_info end fun define_raw_bn_perms raw_dt_info bn_info lthy = if null bn_info then ([], [], lthy) else let val RawDtInfo {raw_tys, raw_cns_info, raw_inject_thms, raw_distinct_thms, raw_size_thms, ...} = raw_dt_info val (bns, bn_tys) = split_list (map (fn (bn, i, _) => (bn, i)) bn_info) val bn_names = map (fn bn => Long_Name.base_name (fst (dest_Const bn))) bns val perm_bn_names = map (prefix "permute_") bn_names val perm_bn_arg_tys = map (nth raw_tys) bn_tys val perm_bn_tys = map (fn ty => @{typ "perm"} --> ty --> ty) perm_bn_arg_tys val perm_bn_frees = map Free (perm_bn_names ~~ perm_bn_tys) val perm_bn_map = bns ~~ perm_bn_frees val perm_bn_eqs = map (mk_perm_bn_eqs lthy perm_bn_map raw_cns_info) bn_info val all_fun_names = map (fn s => (Binding.name s, NONE, NoSyn)) perm_bn_names val all_fun_eqs = map (fn t => ((Binding.empty_atts, t), [], [])) (flat perm_bn_eqs) val prod_simps = @{thms prod.inject HOL.simp_thms} val lthy' = lthy |> Local_Theory.begin_nested |> snd |> Function.add_function all_fun_names all_fun_eqs Function_Common.default_config (pat_completeness_simp (prod_simps @ raw_inject_thms @ raw_distinct_thms)) |> snd |> Local_Theory.end_nested; val ({fs, simps, ...}, lthy'') = prove_termination_fun raw_size_thms lthy' val morphism = Proof_Context.export_morphism lthy'' (Proof_Context.transfer (Proof_Context.theory_of lthy'') lthy) val simps_exp = map (Morphism.thm morphism) (the simps) in (fs, simps_exp, lthy'') end (*** raw permutation functions ***) (** proves the two pt-type class properties **) fun prove_permute_zero induct bnfs perm_defs perm_fns ctxt = let val perm_types = map (body_type o fastype_of) perm_fns val perm_indnames = Old_Datatype_Prop.make_tnames perm_types fun single_goal ((perm_fn, T), x) = HOLogic.mk_eq (perm_fn $ @{term "0::perm"} $ Free (x, T), Free (x, T)) val goals = HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj (map single_goal (perm_fns ~~ perm_types ~~ perm_indnames))) val map_ids = map BNF_Def.map_ident_of_bnf bnfs val rules = @{thm permute_zero} :: perm_defs @ map_ids val congs = map BNF_Def.map_cong0_of_bnf bnfs in Goal.prove ctxt perm_indnames [] goals (fn {context = ctxt', ...} => (Old_Datatype_Aux.ind_tac ctxt' induct perm_indnames THEN_ALL_NEW asm_simp_tac (fold Simplifier.add_cong congs (put_simpset HOL_basic_ss ctxt' addsimps rules))) 1) |> Old_Datatype_Aux.split_conj_thm end fun prove_permute_plus induct bnfs perm_defs perm_fns ctxt = let val p = Free ("p", \<^Type>\perm\) val q = Free ("q", \<^Type>\perm\) val perm_types = map (body_type o fastype_of) perm_fns val perm_indnames = Old_Datatype_Prop.make_tnames perm_types fun single_goal ((perm_fn, T), x) = HOLogic.mk_eq (perm_fn $ (mk_plus p q) $ Free (x, T), perm_fn $ p $ (perm_fn $ q $ Free (x, T))) val goals = HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj (map single_goal (perm_fns ~~ perm_types ~~ perm_indnames))) val map_comps = map BNF_Def.map_comp_of_bnf bnfs val rules = @{thms permute_plus o_def} @ perm_defs @ map_comps val congs = map BNF_Def.map_cong0_of_bnf bnfs in Goal.prove ctxt ("p" :: "q" :: perm_indnames) [] goals (fn {context = ctxt', ...} => (Old_Datatype_Aux.ind_tac ctxt' induct perm_indnames THEN_ALL_NEW asm_simp_tac (fold Simplifier.add_cong congs (put_simpset HOL_basic_ss ctxt' addsimps rules))) 1) |> Old_Datatype_Aux.split_conj_thm end (* Return the map operator for the given type, along with its list of argument types, if a map operator exists; otherwise return NONE *) fun mk_map_of_type lthy (Type (c, tys)) = let fun mk_map bnf = let val live = BNF_Def.live_of_bnf bnf val t = BNF_Def.mk_map live tys tys (BNF_Def.map_of_bnf bnf) val (map_arg_tys, _) = BNF_Util.strip_typeN live (fastype_of t) in (t, map domain_type map_arg_tys) end in Option.map mk_map (BNF_Def.bnf_of lthy c) end | mk_map_of_type _ _ = NONE fun mk_perm_eq lthy ty_perm_assoc cnstr = let (* permute function with boolean flag indicating recursion *) fun mk_perm p ty = case (AList.lookup (op=) ty_perm_assoc ty) of SOME perm => (perm $ p, true) | NONE => (case mk_map_of_type lthy ty of SOME (t, tys) => let val (ts, recs) = split_list (map (mk_perm p) tys) in if exists I recs then (list_comb (t, ts), true) else (perm_const ty $ p, false) end | NONE => (perm_const ty $ p, false)) fun lookup_perm p (ty, arg) = fst (mk_perm p ty) $ arg val p = Free ("p", \<^Type>\perm\) val (arg_tys, ty) = strip_type (fastype_of cnstr) val arg_names = Name.variant_list ["p"] (Old_Datatype_Prop.make_tnames arg_tys) val args = map Free (arg_names ~~ arg_tys) val lhs = lookup_perm p (ty, list_comb (cnstr, args)) val rhs = list_comb (cnstr, map (lookup_perm p) (arg_tys ~~ args)) val eq = HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs)) in ((Binding.empty_atts, eq), [], []) end fun define_raw_perms raw_dt_info lthy = let val RawDtInfo {raw_dt_names, raw_fp_sugars, raw_tys, raw_ty_args, raw_all_cns, raw_induct_thm, ...} = raw_dt_info val bnfs = (#fp_nesting_bnfs o hd) raw_fp_sugars val perm_fn_names = raw_dt_names |> map Long_Name.base_name |> map (prefix "permute_") val perm_fn_types = map perm_ty raw_tys val perm_fn_frees = map Free (perm_fn_names ~~ perm_fn_types) val perm_fn_binds = map (fn s => (Binding.name s, NONE, NoSyn)) perm_fn_names val perm_eqs = map (mk_perm_eq lthy (raw_tys ~~ perm_fn_frees)) (flat raw_all_cns) fun tac ctxt (_, _, simps) = Class.intro_classes_tac ctxt [] THEN ALLGOALS (resolve_tac ctxt simps) fun morphism phi (fvs, dfs, simps) = (map (Morphism.term phi) fvs, map (Morphism.thm phi) dfs, map (Morphism.thm phi) simps); val ((perm_funs, perm_eq_thms), lthy') = lthy |> Local_Theory.exit_global |> Class.instantiation (raw_dt_names, raw_ty_args, @{sort pt}) |> BNF_LFP_Compat.primrec perm_fn_binds perm_eqs val perm_zero_thms = prove_permute_zero raw_induct_thm bnfs perm_eq_thms perm_funs lthy' val perm_plus_thms = prove_permute_plus raw_induct_thm bnfs perm_eq_thms perm_funs lthy' in lthy' |> Class.prove_instantiation_exit_result morphism tac (perm_funs, perm_eq_thms, perm_zero_thms @ perm_plus_thms) ||> Named_Target.theory_init end (** equivarance proofs **) val eqvt_apply_sym = @{thm eqvt_apply[symmetric]} fun subproof_tac const_names simps = SUBPROOF (fn {prems, context = ctxt, ...} => HEADGOAL (simp_tac (put_simpset HOL_basic_ss ctxt addsimps simps) THEN' eqvt_tac ctxt (eqvt_relaxed_config addexcls const_names) THEN' simp_tac (put_simpset HOL_basic_ss ctxt addsimps (prems @ [eqvt_apply_sym])))) fun prove_eqvt_tac insts ind_thms const_names simps ctxt = HEADGOAL (Object_Logic.full_atomize_tac ctxt THEN' (DETERM o (Induct_Tacs.induct_tac ctxt insts (SOME ind_thms))) THEN_ALL_NEW subproof_tac const_names simps ctxt) fun mk_eqvt_goal pi const arg = let val lhs = mk_perm pi (const $ arg) val rhs = const $ (mk_perm pi arg) in HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs)) end fun raw_prove_eqvt consts ind_thms simps ctxt = if null consts then [] else let val ([p], ctxt') = Variable.variant_fixes ["p"] ctxt val p = Free (p, \<^Type>\perm\) val arg_tys = consts |> map fastype_of |> map domain_type val (arg_names, ctxt'') = Variable.variant_fixes (Old_Datatype_Prop.make_tnames arg_tys) ctxt' val args = map Free (arg_names ~~ arg_tys) val goals = map2 (mk_eqvt_goal p) consts args val insts = map (single o SOME) arg_names val const_names = map (fst o dest_Const) consts in - Goal.prove_common ctxt'' NONE [] [] goals (fn {context, ...} => - prove_eqvt_tac insts ind_thms const_names simps context) + Goal.prove_common ctxt'' NONE [] [] goals (fn {context = goal_ctxt, ...} => + prove_eqvt_tac insts ind_thms const_names simps goal_ctxt) |> Proof_Context.export ctxt'' ctxt end end (* structure *) diff --git a/thys/Nominal2/nominal_eqvt.ML b/thys/Nominal2/nominal_eqvt.ML --- a/thys/Nominal2/nominal_eqvt.ML +++ b/thys/Nominal2/nominal_eqvt.ML @@ -1,147 +1,147 @@ (* Title: nominal_eqvt.ML Author: Stefan Berghofer (original code) Author: Christian Urban Author: Tjark Weber Automatic proofs for equivariance of inductive predicates. *) signature NOMINAL_EQVT = sig val raw_equivariance: Proof.context -> term list -> thm -> thm list -> thm list val equivariance_cmd: string -> Proof.context -> local_theory end structure Nominal_Eqvt : NOMINAL_EQVT = struct open Nominal_Permeq; open Nominal_ThmDecls; fun atomize_conv ctxt = Raw_Simplifier.rewrite_cterm (true, false, false) (K (K NONE)) (put_simpset HOL_basic_ss ctxt addsimps @{thms induct_atomize}) fun atomize_intr ctxt = Conv.fconv_rule (Conv.prems_conv ~1 (atomize_conv ctxt)) fun atomize_induct ctxt = Conv.fconv_rule (Conv.prems_conv ~1 - (Conv.params_conv ~1 (K (Conv.prems_conv ~1 (atomize_conv ctxt))) ctxt)) + (Conv.params_conv ~1 (Conv.prems_conv ~1 o atomize_conv) ctxt)) (** equivariance tactics **) fun eqvt_rel_single_case_tac ctxt pred_names pi intro = let val cpi = Thm.cterm_of ctxt pi val pi_intro_rule = Thm.instantiate' [] [NONE, SOME cpi] @{thm permute_boolI} val eqvt_sconfig = eqvt_strict_config addexcls pred_names in eqvt_tac ctxt eqvt_sconfig THEN' - SUBPROOF (fn {prems, context as ctxt, ...} => + SUBPROOF (fn {prems, context = goal_ctxt, ...} => let val simps1 = - put_simpset HOL_basic_ss ctxt addsimps @{thms permute_fun_def permute_self split_paired_all} + put_simpset HOL_basic_ss goal_ctxt addsimps @{thms permute_fun_def permute_self split_paired_all} val simps2 = - put_simpset HOL_basic_ss ctxt addsimps @{thms permute_bool_def permute_minus_cancel(2)} - val prems' = map (transform_prem2 ctxt pred_names) prems - val prems'' = map (fn thm => eqvt_rule ctxt eqvt_sconfig (thm RS pi_intro_rule)) prems' + put_simpset HOL_basic_ss goal_ctxt addsimps @{thms permute_bool_def permute_minus_cancel(2)} + val prems' = map (transform_prem2 goal_ctxt pred_names) prems + val prems'' = map (fn thm => eqvt_rule goal_ctxt eqvt_sconfig (thm RS pi_intro_rule)) prems' val prems''' = map (simplify simps2 o simplify simps1) prems'' in - HEADGOAL (resolve_tac ctxt [intro] THEN_ALL_NEW - resolve_tac ctxt (prems' @ prems'' @ prems''')) + HEADGOAL (resolve_tac goal_ctxt [intro] THEN_ALL_NEW + resolve_tac goal_ctxt (prems' @ prems'' @ prems''')) end) ctxt end fun eqvt_rel_tac ctxt pred_names pi induct intros = let val cases = map (eqvt_rel_single_case_tac ctxt pred_names pi) intros in EVERY' ((DETERM o resolve_tac ctxt [induct]) :: cases) end (** equivariance procedure **) fun prepare_goal ctxt pi pred_with_args = let val (c, xs) = strip_comb pred_with_args fun is_nonfixed_Free (Free (s, _)) = not (Variable.is_fixed ctxt s) | is_nonfixed_Free _ = false fun mk_perm_nonfixed_Free t = if is_nonfixed_Free t then mk_perm pi t else t in HOLogic.mk_imp (pred_with_args, list_comb (c, map mk_perm_nonfixed_Free xs)) end fun name_of (Const (s, _)) = s fun raw_equivariance ctxt preds raw_induct intrs = let (* FIXME: polymorphic predicates should either be rejected or specialized to arguments of sort pt *) val is_already_eqvt = filter (is_eqvt ctxt) preds val _ = if null is_already_eqvt then () else error ("Already equivariant: " ^ commas (map (Syntax.string_of_term ctxt) is_already_eqvt)) val pred_names = map (name_of o head_of) preds val raw_induct' = atomize_induct ctxt raw_induct val intrs' = map (atomize_intr ctxt) intrs val (([raw_concl], [raw_pi]), ctxt') = ctxt |> Variable.import_terms false [Thm.concl_of raw_induct'] ||>> Variable.variant_fixes ["p"] val pi = Free (raw_pi, \<^Type>\perm\) val preds_with_args = raw_concl |> HOLogic.dest_Trueprop |> HOLogic.dest_conj |> map (fst o HOLogic.dest_imp) val goal = preds_with_args |> map (prepare_goal ctxt pi) |> foldr1 HOLogic.mk_conj |> HOLogic.mk_Trueprop in Goal.prove ctxt' [] [] goal - (fn {context, ...} => eqvt_rel_tac context pred_names pi raw_induct' intrs' 1) + (fn {context = goal_ctxt, ...} => eqvt_rel_tac goal_ctxt pred_names pi raw_induct' intrs' 1) |> Old_Datatype_Aux.split_conj_thm |> Proof_Context.export ctxt' ctxt |> map (fn th => th RS mp) |> map zero_var_indexes end (** stores thm under name.eqvt and adds [eqvt]-attribute **) fun note_named_thm (name, thm) ctxt = let val thm_name = Binding.qualified_name (Long_Name.qualify (Long_Name.base_name name) "eqvt") val attr = Attrib.internal (K eqvt_add) val ((_, [thm']), ctxt') = Local_Theory.note ((thm_name, [attr]), [thm]) ctxt in (thm', ctxt') end (** equivariance command **) fun equivariance_cmd pred_name ctxt = let val ({names, ...}, {preds, raw_induct, intrs, ...}) = Inductive.the_inductive_global ctxt (long_name ctxt pred_name) val thms = raw_equivariance ctxt preds raw_induct intrs in fold_map note_named_thm (names ~~ thms) ctxt |> snd end val _ = Outer_Syntax.local_theory @{command_keyword equivariance} "Proves equivariance for inductive predicate involving nominal datatypes." (Parse.const >> equivariance_cmd) end (* structure *) diff --git a/thys/Nominal2/nominal_inductive.ML b/thys/Nominal2/nominal_inductive.ML --- a/thys/Nominal2/nominal_inductive.ML +++ b/thys/Nominal2/nominal_inductive.ML @@ -1,440 +1,441 @@ (* Title: nominal_inductive.ML Author: Christian Urban Author: Tjark Weber Infrastructure for proving strong induction theorems for inductive predicates involving nominal datatypes. Code based on an earlier version by Stefan Berghofer. *) signature NOMINAL_INDUCTIVE = sig val prove_strong_inductive: string list -> string list -> term list list -> thm -> thm list -> Proof.context -> Proof.state val prove_strong_inductive_cmd: xstring * (string * string list) list -> Proof.context -> Proof.state end structure Nominal_Inductive : NOMINAL_INDUCTIVE = struct fun mk_cplus p q = Thm.apply (Thm.apply @{cterm "plus :: perm => perm => perm"} p) q fun mk_cminus p = Thm.apply @{cterm "uminus :: perm => perm"} p fun minus_permute_intro_tac ctxt p = resolve_tac ctxt [Thm.instantiate' [] [SOME (mk_cminus p)] @{thm permute_boolE}] fun minus_permute_elim p thm = thm RS (Thm.instantiate' [] [NONE, SOME (mk_cminus p)] @{thm permute_boolI}) (* fixme: move to nominal_library *) fun real_head_of \<^Const_>\Trueprop for t\ = real_head_of t | real_head_of \<^Const_>\Pure.imp for _ t\ = real_head_of t | real_head_of \<^Const_>\Pure.all _ for \Abs (_, _, t)\\ = real_head_of t | real_head_of \<^Const_>\All _ for \Abs (_, _, t)\\ = real_head_of t | real_head_of \<^Const_>\HOL.induct_forall _ for \Abs (_, _, t)\\ = real_head_of t | real_head_of t = head_of t fun mk_vc_compat (avoid, avoid_trm) prems concl_args params = if null avoid then [] else let val vc_goal = concl_args |> HOLogic.mk_tuple |> mk_fresh_star avoid_trm |> HOLogic.mk_Trueprop |> (curry Logic.list_implies) prems |> fold_rev (Logic.all o Free) params val finite_goal = avoid_trm |> mk_finite |> HOLogic.mk_Trueprop |> (curry Logic.list_implies) prems |> fold_rev (Logic.all o Free) params in [vc_goal, finite_goal] end (* fixme: move to nominal_library *) fun map_term prop f trm = if prop trm then f trm else case trm of (t1 $ t2) => map_term prop f t1 $ map_term prop f t2 | Abs (x, T, t) => Abs (x, T, map_term prop f t) | _ => trm fun add_p_c p (c, c_ty) trm = let val (P, args) = strip_comb trm val (P_name, P_ty) = dest_Free P val (ty_args, bool) = strip_type P_ty val args' = map (mk_perm p) args in list_comb (Free (P_name, (c_ty :: ty_args) ---> bool), c :: args') |> (fn t => HOLogic.all_const c_ty $ lambda c t ) |> (fn t => HOLogic.all_const \<^Type>\perm\ $ lambda p t) end fun induct_forall_const T = \<^Const>\HOL.induct_forall T\ fun mk_induct_forall (a, T) t = induct_forall_const T $ Abs (a, T, t) fun add_c_prop qnt Ps (c, c_name, c_ty) trm = let fun add t = let val (P, args) = strip_comb t val (P_name, P_ty) = dest_Free P val (ty_args, bool) = strip_type P_ty val args' = args |> qnt ? map (incr_boundvars 1) in list_comb (Free (P_name, (c_ty :: ty_args) ---> bool), c :: args') |> qnt ? mk_induct_forall (c_name, c_ty) end in map_term (member (op =) Ps o head_of) add trm end fun prep_prem Ps c_name c_ty (avoid, avoid_trm) (params, prems, concl) = let val prems' = prems |> map (incr_boundvars 1) |> map (add_c_prop true Ps (Bound 0, c_name, c_ty)) val avoid_trm' = avoid_trm |> fold_rev absfree (params @ [(c_name, c_ty)]) |> strip_abs_body |> (fn t => mk_fresh_star_ty c_ty t (Bound 0)) |> HOLogic.mk_Trueprop val prems'' = if null avoid then prems' else avoid_trm' :: prems' val concl' = concl |> incr_boundvars 1 |> add_c_prop false Ps (Bound 0, c_name, c_ty) in mk_full_horn (params @ [(c_name, c_ty)]) prems'' concl' end (* fixme: move to nominal_library *) fun same_name (Free (a1, _), Free (a2, _)) = (a1 = a2) | same_name (Var (a1, _), Var (a2, _)) = (a1 = a2) | same_name (Const (a1, _), Const (a2, _)) = (a1 = a2) | same_name _ = false (* local abbreviations *) local open Nominal_Permeq in (* by default eqvt_strict_config contains unwanted @{thm permute_pure} *) val eqvt_sconfig = eqvt_strict_config addpres @{thms permute_minus_cancel} fun eqvt_stac ctxt = eqvt_tac ctxt eqvt_sconfig fun eqvt_srule ctxt = eqvt_rule ctxt eqvt_sconfig end val all_elims = let fun spec' ct = Thm.instantiate' [SOME (Thm.ctyp_of_cterm ct)] [NONE, SOME ct] @{thm spec} in fold (fn ct => fn th => th RS spec' ct) end fun helper_tac flag prm p ctxt = Subgoal.SUBPROOF (fn {context = ctxt', prems, ...} => let val prems' = prems |> map (minus_permute_elim p) |> map (eqvt_srule ctxt') val prm' = (prems' MRS prm) |> flag ? (all_elims [p]) |> flag ? (eqvt_srule ctxt') in asm_full_simp_tac (put_simpset HOL_ss ctxt' addsimps (prm' :: @{thms HOL.induct_forall_def})) 1 end) ctxt fun non_binder_tac prem intr_cvars Ps ctxt = Subgoal.SUBPROOF (fn {context = ctxt', params, prems, ...} => let val (prms, p, _) = split_last2 (map snd params) val prm_tys = map (fastype_of o Thm.term_of) prms val cperms = map (Thm.cterm_of ctxt' o perm_const) prm_tys val p_prms = map2 (fn ct1 => fn ct2 => Thm.mk_binop ct1 p ct2) cperms prms val prem' = prem |> infer_instantiate ctxt' (map (#1 o dest_Var o Thm.term_of) intr_cvars ~~ p_prms) |> eqvt_srule ctxt' (* for inductive-premises*) fun tac1 prm = helper_tac true prm p ctxt' (* for non-inductive premises *) fun tac2 prm = EVERY' [ minus_permute_intro_tac ctxt' p, eqvt_stac ctxt', helper_tac false prm p ctxt' ] fun select prm (t, i) = (if member same_name Ps (real_head_of t) then tac1 prm else tac2 prm) i in EVERY1 [ eqvt_stac ctxt', resolve_tac ctxt' [prem'], RANGE (map (SUBGOAL o select) prems) ] end) ctxt fun fresh_thm ctxt user_thm p c concl_args avoid_trm = let val conj1 = mk_fresh_star (mk_perm (Bound 0) (mk_perm p avoid_trm)) c val conj2 = mk_fresh_star_ty \<^Type>\perm\ (mk_supp (HOLogic.mk_tuple (map (mk_perm p) concl_args))) (Bound 0) val fresh_goal = mk_exists ("q", \<^Type>\perm\) (HOLogic.mk_conj (conj1, conj2)) |> HOLogic.mk_Trueprop val ss = @{thms finite_supp supp_Pair finite_Un permute_finite} @ @{thms fresh_star_Pair fresh_star_permute_iff} val simp = asm_full_simp_tac (put_simpset HOL_ss ctxt addsimps ss) in Goal.prove ctxt [] [] fresh_goal - (K (HEADGOAL (resolve_tac ctxt @{thms at_set_avoiding2} + (fn {context = goal_ctxt, ...} => + HEADGOAL (resolve_tac goal_ctxt @{thms at_set_avoiding2} THEN_ALL_NEW EVERY' [cut_facts_tac user_thm, REPEAT o - eresolve_tac ctxt @{thms conjE}, simp]))) + eresolve_tac goal_ctxt @{thms conjE}, simp])) end val supp_perm_eq' = @{lemma "fresh_star (supp (permute p x)) q ==> permute p x == permute (q + p) x" by (simp add: supp_perm_eq)} val fresh_star_plus = @{lemma "fresh_star (permute q (permute p x)) c ==> fresh_star (permute (q + p) x) c" by (simp add: permute_plus)} fun binder_tac prem intr_cvars param_trms Ps user_thm avoid_trm concl_args ctxt = Subgoal.FOCUS (fn {context = ctxt, params, prems, concl, ...} => let val (prms, p, c) = split_last2 (map snd params) val prm_trms = map Thm.term_of prms val prm_tys = map fastype_of prm_trms val avoid_trm' = subst_free (param_trms ~~ prm_trms) avoid_trm val concl_args' = map (subst_free (param_trms ~~ prm_trms)) concl_args val user_thm' = map (infer_instantiate ctxt (map (#1 o dest_Var o Thm.term_of) intr_cvars ~~ prms)) user_thm |> map (full_simplify (put_simpset HOL_ss ctxt addsimps (@{thm fresh_star_Pair}::prems))) val fthm = fresh_thm ctxt user_thm' (Thm.term_of p) (Thm.term_of c) concl_args' avoid_trm' val (([(_, q)], fprop :: fresh_eqs), ctxt') = Obtain.result (K (EVERY1 [eresolve_tac ctxt @{thms exE}, full_simp_tac (put_simpset HOL_basic_ss ctxt addsimps @{thms supp_Pair fresh_star_Un}), REPEAT o eresolve_tac ctxt @{thms conjE}, dresolve_tac ctxt [fresh_star_plus], REPEAT o dresolve_tac ctxt [supp_perm_eq']])) [fthm] ctxt val expand_conv = Conv.try_conv (Conv.rewrs_conv fresh_eqs) fun expand_conv_bot ctxt = Conv.bottom_conv (K expand_conv) ctxt val cperms = map (Thm.cterm_of ctxt' o perm_const) prm_tys val qp_prms = map2 (fn ct1 => fn ct2 => Thm.mk_binop ct1 (mk_cplus q p) ct2) cperms prms val prem' = prem |> infer_instantiate ctxt' (map (#1 o dest_Var o Thm.term_of) intr_cvars ~~ qp_prms) |> eqvt_srule ctxt' val fprop' = eqvt_srule ctxt' fprop val tac_fresh = simp_tac (put_simpset HOL_basic_ss ctxt' addsimps [fprop']) (* for inductive-premises*) - fun tac1 prm = helper_tac true prm (mk_cplus q p) ctxt' + fun tac1 goal_ctxt prm = helper_tac true prm (mk_cplus q p) goal_ctxt (* for non-inductive premises *) - fun tac2 prm = - EVERY' [ minus_permute_intro_tac ctxt' (mk_cplus q p), - eqvt_stac ctxt', - helper_tac false prm (mk_cplus q p) ctxt' ] + fun tac2 goal_ctxt prm = + EVERY' [ minus_permute_intro_tac goal_ctxt (mk_cplus q p), + eqvt_stac goal_ctxt, + helper_tac false prm (mk_cplus q p) goal_ctxt ] - fun select prm (t, i) = - (if member same_name Ps (real_head_of t) then tac1 prm else tac2 prm) i + fun select goal_ctxt prm (t, i) = + (if member same_name Ps (real_head_of t) then tac1 goal_ctxt prm else tac2 goal_ctxt prm) i val side_thm = Goal.prove ctxt' [] [] (Thm.term_of concl) - (fn {context = ctxt'', ...} => - EVERY1 [ CONVERSION (expand_conv_bot ctxt''), - eqvt_stac ctxt'', - resolve_tac ctxt'' [prem'], - RANGE (tac_fresh :: map (SUBGOAL o select) prems) ]) + (fn {context = goal_ctxt, ...} => + EVERY1 [ CONVERSION (expand_conv_bot goal_ctxt), + eqvt_stac goal_ctxt, + resolve_tac goal_ctxt [prem'], + RANGE (tac_fresh :: map (SUBGOAL o select goal_ctxt) prems) ]) |> singleton (Proof_Context.export ctxt' ctxt) in resolve_tac ctxt [side_thm] 1 end) ctxt fun case_tac ctxt Ps avoid avoid_trm intr_cvars param_trms prem user_thm concl_args = let val tac1 = non_binder_tac prem intr_cvars Ps ctxt val tac2 = binder_tac prem intr_cvars param_trms Ps user_thm avoid_trm concl_args ctxt in EVERY' [ resolve_tac ctxt @{thms allI}, resolve_tac ctxt @{thms allI}, if null avoid then tac1 else tac2 ] end fun prove_sinduct_tac raw_induct user_thms Ps avoids avoid_trms intr_cvars param_trms concl_args {prems, context = ctxt} = let val cases_tac = @{map 7} (case_tac ctxt Ps) avoids avoid_trms intr_cvars param_trms prems user_thms concl_args in EVERY1 [ DETERM o resolve_tac ctxt [raw_induct], RANGE cases_tac ] end val normalise = @{lemma "(Q \ (\p c. P p c)) \ (\c. Q \ P (0::perm) c)" by simp} fun prove_strong_inductive pred_names rule_names avoids raw_induct intrs ctxt = let val ((_, [raw_induct']), ctxt') = Variable.import true [raw_induct] ctxt val (ind_prems, ind_concl) = raw_induct' |> Thm.prop_of |> Logic.strip_horn |>> map strip_full_horn val params = map (fn (x, _, _) => x) ind_prems val param_trms = (map o map) Free params val intr_vars_tys = map (fn t => rev (Term.add_vars (Thm.prop_of t) [])) intrs val intr_vars = (map o map) fst intr_vars_tys val intr_vars_substs = map2 (curry (op ~~)) intr_vars param_trms val intr_cvars = (map o map) (Thm.cterm_of ctxt o Var) intr_vars_tys val (intr_prems, intr_concls) = intrs |> map Thm.prop_of |> map2 subst_Vars intr_vars_substs |> map Logic.strip_horn |> split_list val intr_concls_args = map (snd o fixed_nonfixed_args ctxt' o HOLogic.dest_Trueprop) intr_concls val avoid_trms = avoids |> (map o map) (setify ctxt') |> map fold_union val vc_compat_goals = map4 mk_vc_compat (avoids ~~ avoid_trms) intr_prems intr_concls_args params val ([c_name, a, p], ctxt'') = Variable.variant_fixes ["c", "'a", "p"] ctxt' val c_ty = TFree (a, @{sort fs}) val c = Free (c_name, c_ty) val p = Free (p, \<^Type>\perm\) val (preconds, ind_concls) = ind_concl |> HOLogic.dest_Trueprop |> HOLogic.dest_conj |> map HOLogic.dest_imp |> split_list val Ps = map (fst o strip_comb) ind_concls val ind_concl' = ind_concls |> map (add_p_c p (c, c_ty)) |> (curry (op ~~)) preconds |> map HOLogic.mk_imp |> fold_conj |> HOLogic.mk_Trueprop val ind_prems' = ind_prems |> map2 (prep_prem Ps c_name c_ty) (avoids ~~ avoid_trms) fun after_qed ctxt_outside user_thms ctxt = let val strong_ind_thms = Goal.prove ctxt [] ind_prems' ind_concl' (prove_sinduct_tac raw_induct user_thms Ps avoids avoid_trms intr_cvars param_trms intr_concls_args) |> singleton (Proof_Context.export ctxt ctxt_outside) |> Old_Datatype_Aux.split_conj_thm |> map (fn thm => thm RS normalise) |> map (asm_full_simplify (put_simpset HOL_basic_ss ctxt addsimps @{thms permute_zero induct_rulify})) |> map (Drule.rotate_prems (length ind_prems')) |> map zero_var_indexes val qualified_thm_name = pred_names |> map Long_Name.base_name |> space_implode "_" |> (fn s => Binding.qualify false s (Binding.name "strong_induct")) val attrs = [ Attrib.internal (K (Rule_Cases.consumes 1)), Attrib.internal (K (Rule_Cases.case_names rule_names)) ] in ctxt |> Local_Theory.note ((qualified_thm_name, attrs), strong_ind_thms) |> snd end in Proof.theorem NONE (after_qed ctxt) ((map o map) (rpair []) vc_compat_goals) ctxt'' end fun prove_strong_inductive_cmd (pred_name, avoids) ctxt = let val ({names, ...}, {raw_induct, intrs, ...}) = Inductive.the_inductive_global ctxt (long_name ctxt pred_name) val rule_names = hd names |> the o Induct.lookup_inductP ctxt |> fst o Rule_Cases.get |> map (fst o fst) val case_names = map fst avoids val _ = case duplicates (op =) case_names of [] => () | xs => error ("Duplicate case names: " ^ commas_quote xs) val _ = case subtract (op =) rule_names case_names of [] => () | xs => error ("No such case(s) in inductive definition: " ^ commas_quote xs) val avoids_ordered = order_default (op =) [] rule_names avoids fun read_avoids avoid_trms intr = let (* fixme hack *) val (((_, inst), _), ctxt') = Variable.import true [intr] ctxt val trms = build (inst |> Vars.fold (cons o Thm.term_of o snd)) val ctxt'' = fold Variable.declare_term trms ctxt' in map (Syntax.read_term ctxt'') avoid_trms end val avoid_trms = map2 read_avoids avoids_ordered intrs in prove_strong_inductive names rule_names avoid_trms raw_induct intrs ctxt end (* outer syntax *) local val single_avoid_parser = Parse.name -- (@{keyword ":"} |-- Parse.and_list1 Parse.term) val avoids_parser = Scan.optional (@{keyword "avoids"} |-- Parse.enum1 "|" single_avoid_parser) [] val main_parser = Parse.name -- avoids_parser in val _ = Outer_Syntax.local_theory_to_proof @{command_keyword nominal_inductive} "prove strong induction theorem for inductive predicate involving nominal datatypes" (main_parser >> prove_strong_inductive_cmd) end end diff --git a/thys/Nominal2/nominal_library.ML b/thys/Nominal2/nominal_library.ML --- a/thys/Nominal2/nominal_library.ML +++ b/thys/Nominal2/nominal_library.ML @@ -1,488 +1,488 @@ (* Title: nominal_library.ML Author: Christian Urban Library functions for nominal. *) signature NOMINAL_LIBRARY = sig val mk_sort_of: term -> term val atom_ty: typ -> typ val atom_const: typ -> term val mk_atom_ty: typ -> term -> term val mk_atom: term -> term val mk_atom_set_ty: typ -> term -> term val mk_atom_set: term -> term val mk_atom_fset_ty: typ -> term -> term val mk_atom_fset: term -> term val mk_atom_list_ty: typ -> term -> term val mk_atom_list: term -> term val is_atom: Proof.context -> typ -> bool val is_atom_set: Proof.context -> typ -> bool val is_atom_fset: Proof.context -> typ -> bool val is_atom_list: Proof.context -> typ -> bool val to_set_ty: typ -> term -> term val to_set: term -> term val atomify_ty: Proof.context -> typ -> term -> term val atomify: Proof.context -> term -> term val setify_ty: Proof.context -> typ -> term -> term val setify: Proof.context -> term -> term val listify_ty: Proof.context -> typ -> term -> term val listify: Proof.context -> term -> term val fresh_const: typ -> term val mk_fresh_ty: typ -> term -> term -> term val mk_fresh: term -> term -> term val fresh_star_const: typ -> term val mk_fresh_star_ty: typ -> term -> term -> term val mk_fresh_star: term -> term -> term val supp_const: typ -> term val mk_supp_ty: typ -> term -> term val mk_supp: term -> term val supp_rel_const: typ -> term val mk_supp_rel_ty: typ -> term -> term -> term val mk_supp_rel: term -> term -> term val supports_const: typ -> term val mk_supports_ty: typ -> term -> term -> term val mk_supports: term -> term -> term val finite_const: typ -> term val mk_finite_ty: typ -> term -> term val mk_finite: term -> term val mk_diff: term * term -> term val mk_append: term * term -> term val mk_union: term * term -> term val fold_union: term list -> term val fold_append: term list -> term val mk_conj: term * term -> term val fold_conj: term list -> term val fold_conj_balanced: term list -> term (* functions for de-Bruijn open terms *) val mk_binop_env: typ list -> string -> term * term -> term val mk_union_env: typ list -> term * term -> term val fold_union_env: typ list -> term list -> term (* fresh arguments for a term *) val fresh_args: Proof.context -> term -> term list (* some logic operations *) val strip_full_horn: term -> (string * typ) list * term list * term val mk_full_horn: (string * typ) list -> term list -> term -> term (* datatype operations *) type cns_info = (term * typ * typ list * bool list) list val all_dtyp_constrs_types: Old_Datatype_Aux.descr -> cns_info list (* tactics for function package *) val size_ss: simpset val pat_completeness_simp: thm list -> Proof.context -> tactic val prove_termination_ind: Proof.context -> int -> tactic val prove_termination_fun: thm list -> Proof.context -> Function.info * local_theory (* transformations of premises in inductions *) val transform_prem1: Proof.context -> string list -> thm -> thm val transform_prem2: Proof.context -> string list -> thm -> thm (* transformation into the object logic *) val atomize: Proof.context -> thm -> thm val atomize_rule: Proof.context -> int -> thm -> thm val atomize_concl: Proof.context -> thm -> thm (* applies a tactic to a formula composed of conjunctions *) val conj_tac: Proof.context -> (int -> tactic) -> int -> tactic end structure Nominal_Library: NOMINAL_LIBRARY = struct fun mk_sort_of t = \<^Const>\sort_of for t\; fun atom_ty ty = ty --> @{typ "atom"}; fun atom_const ty = \<^Const>\atom ty\ fun mk_atom_ty ty t = atom_const ty $ t; fun mk_atom t = mk_atom_ty (fastype_of t) t; fun mk_atom_set_ty ty t = let val atom_ty = HOLogic.dest_setT ty in \<^Const>\image atom_ty \<^Type>\atom\ for \atom_const atom_ty\ t\ end fun mk_atom_fset_ty ty t = let val atom_ty = dest_fsetT ty in \<^Const>\fimage atom_ty \<^Type>\atom\ for \atom_const atom_ty\ t\ end fun mk_atom_list_ty ty t = let val atom_ty = dest_listT ty in \<^Const>\map atom_ty \<^Type>\atom\ for \atom_const atom_ty\ t\ end fun mk_atom_set t = mk_atom_set_ty (fastype_of t) t fun mk_atom_fset t = mk_atom_fset_ty (fastype_of t) t fun mk_atom_list t = mk_atom_list_ty (fastype_of t) t (* coerces a list into a set *) fun to_set_ty ty t = case ty of @{typ "atom list"} => @{term "set :: atom list => atom set"} $ t | @{typ "atom fset"} => @{term "fset :: atom fset => atom set"} $ t | _ => t fun to_set t = to_set_ty (fastype_of t) t (* testing for concrete atom types *) fun is_atom ctxt ty = Sign.of_sort (Proof_Context.theory_of ctxt) (ty, @{sort at_base}) fun is_atom_set ctxt \<^Type>\fun ty \<^Type>\bool\\ = is_atom ctxt ty | is_atom_set _ _ = false; fun is_atom_fset ctxt \<^Type>\fset ty\ = is_atom ctxt ty | is_atom_fset _ _ = false; fun is_atom_list ctxt \<^Type>\list ty\ = is_atom ctxt ty | is_atom_list _ _ = false (* functions that coerce singletons, sets, fsets and lists of concrete atoms into general atoms sets / lists *) fun atomify_ty ctxt ty t = if is_atom ctxt ty then mk_atom_ty ty t else if is_atom_set ctxt ty then mk_atom_set_ty ty t else if is_atom_fset ctxt ty then mk_atom_fset_ty ty t else if is_atom_list ctxt ty then mk_atom_list_ty ty t else raise TERM ("atomify: term is not an atom, set or list of atoms", [t]) fun setify_ty ctxt ty t = if is_atom ctxt ty then HOLogic.mk_set \<^Type>\atom\ [mk_atom_ty ty t] else if is_atom_set ctxt ty then mk_atom_set_ty ty t else if is_atom_fset ctxt ty then @{term "fset :: atom fset => atom set"} $ mk_atom_fset_ty ty t else if is_atom_list ctxt ty then @{term "set :: atom list => atom set"} $ mk_atom_list_ty ty t else raise TERM ("setify: term is not an atom, set or list of atoms", [t]) fun listify_ty ctxt ty t = if is_atom ctxt ty then HOLogic.mk_list \<^Type>\atom\ [mk_atom_ty ty t] else if is_atom_list ctxt ty then mk_atom_list_ty ty t else raise TERM ("listify: term is not an atom or list of atoms", [t]) fun atomify ctxt t = atomify_ty ctxt (fastype_of t) t fun setify ctxt t = setify_ty ctxt (fastype_of t) t fun listify ctxt t = listify_ty ctxt (fastype_of t) t fun fresh_const ty = \<^Const>\fresh ty\ fun mk_fresh_ty ty t1 t2 = fresh_const ty $ t1 $ t2 fun mk_fresh t1 t2 = mk_fresh_ty (fastype_of t2) t1 t2 fun fresh_star_const ty = \<^Const>\fresh_star ty\ fun mk_fresh_star_ty ty t1 t2 = fresh_star_const ty $ t1 $ t2 fun mk_fresh_star t1 t2 = mk_fresh_star_ty (fastype_of t2) t1 t2 fun supp_const ty = \<^Const>\supp ty\ fun mk_supp_ty ty t = supp_const ty $ t fun mk_supp t = mk_supp_ty (fastype_of t) t fun supp_rel_const ty = \<^Const>\supp_rel ty\ fun mk_supp_rel_ty ty r t = supp_rel_const ty $ r $ t fun mk_supp_rel r t = mk_supp_rel_ty (fastype_of t) r t fun supports_const ty = \<^Const>\supports ty\; fun mk_supports_ty ty t1 t2 = supports_const ty $ t1 $ t2; fun mk_supports t1 t2 = mk_supports_ty (fastype_of t2) t1 t2; fun finite_const ty = \<^Const>\finite ty\ fun mk_finite_ty ty t = finite_const ty $ t fun mk_finite t = mk_finite_ty (HOLogic.dest_setT (fastype_of t)) t (* functions that construct differences, appends and unions but avoid producing empty atom sets or empty atom lists *) fun mk_diff (@{term "{}::atom set"}, _) = @{term "{}::atom set"} | mk_diff (t1, @{term "{}::atom set"}) = t1 | mk_diff (@{term "set ([]::atom list)"}, _) = @{term "set ([]::atom list)"} | mk_diff (t1, @{term "set ([]::atom list)"}) = t1 | mk_diff (t1, t2) = HOLogic.mk_binop @{const_name minus} (t1, t2) fun mk_append (t1, @{term "[]::atom list"}) = t1 | mk_append (@{term "[]::atom list"}, t2) = t2 | mk_append (t1, t2) = HOLogic.mk_binop @{const_name "append"} (t1, t2) fun mk_union (t1, @{term "{}::atom set"}) = t1 | mk_union (@{term "{}::atom set"}, t2) = t2 | mk_union (t1, @{term "set ([]::atom list)"}) = t1 | mk_union (@{term "set ([]::atom list)"}, t2) = t2 | mk_union (t1, t2) = HOLogic.mk_binop @{const_name "sup"} (t1, t2) fun fold_union trms = fold_rev (curry mk_union) trms @{term "{}::atom set"} fun fold_append trms = fold_rev (curry mk_append) trms @{term "[]::atom list"} fun mk_conj (t1, @{term "True"}) = t1 | mk_conj (@{term "True"}, t2) = t2 | mk_conj (t1, t2) = HOLogic.mk_conj (t1, t2) fun fold_conj trms = fold_rev (curry mk_conj) trms @{term "True"} fun fold_conj_balanced ts = Balanced_Tree.make HOLogic.mk_conj ts (* functions for de-Bruijn open terms *) fun mk_binop_env tys c (t, u) = let val ty = fastype_of1 (tys, t) in Const (c, [ty, ty] ---> ty) $ t $ u end fun mk_union_env tys (t1, @{term "{}::atom set"}) = t1 | mk_union_env tys (@{term "{}::atom set"}, t2) = t2 | mk_union_env tys (t1, @{term "set ([]::atom list)"}) = t1 | mk_union_env tys (@{term "set ([]::atom list)"}, t2) = t2 | mk_union_env tys (t1, t2) = mk_binop_env tys @{const_name "sup"} (t1, t2) fun fold_union_env tys trms = fold_left (mk_union_env tys) trms @{term "{}::atom set"} (* produces fresh arguments for a term *) fun fresh_args ctxt f = f |> fastype_of |> binder_types |> map (pair "z") |> Variable.variant_frees ctxt [f] |> map Free (** some logic operations **) (* decompses a formula into params, premises and a conclusion *) fun strip_full_horn trm = let fun strip_outer_params \<^Const_>\Pure.all _ for \Abs (a, T, t)\\ = strip_outer_params t |>> cons (a, T) | strip_outer_params B = ([], B) val (params, body) = strip_outer_params trm val (prems, concl) = Logic.strip_horn body in (params, prems, concl) end (* composes a formula out of params, premises and a conclusion *) fun mk_full_horn params prems concl = Logic.list_implies (prems, concl) |> fold_rev mk_all params (** datatypes **) (* constructor infos *) type cns_info = (term * typ * typ list * bool list) list (* - term for constructor constant - type of the constructor - types of the arguments - flags indicating whether the argument is recursive *) (* returns info about constructors in a datatype *) fun all_dtyp_constrs_info descr = map (fn (_, (ty, vs, constrs)) => map (pair (ty, vs)) constrs) descr (* returns the constants of the constructors plus the corresponding type and types of arguments *) fun all_dtyp_constrs_types descr = let fun aux ((ty_name, vs), (cname, args)) = let val vs_tys = map (Old_Datatype_Aux.typ_of_dtyp descr) vs val ty = Type (ty_name, vs_tys) val arg_tys = map (Old_Datatype_Aux.typ_of_dtyp descr) args val is_rec = map Old_Datatype_Aux.is_rec_type args in (Const (cname, arg_tys ---> ty), ty, arg_tys, is_rec) end in map (map aux) (all_dtyp_constrs_info descr) end (** function package tactics **) fun pat_completeness_simp simps ctxt = let val simpset = put_simpset HOL_basic_ss ctxt addsimps (@{thms sum.inject sum.distinct} @ simps) in Pat_Completeness.pat_completeness_tac ctxt 1 THEN ALLGOALS (asm_full_simp_tac simpset) end (* simpset for size goals *) val size_ss = simpset_of (put_simpset HOL_ss @{context} addsimprocs [@{simproc natless_cancel_numerals}] addsimps @{thms in_measure wf_measure sum.case add_Suc_right add.right_neutral zero_less_Suc prod.size(1) mult_Suc_right}) val natT = @{typ nat} fun size_prod_const T1 T2 = \<^Const>\size_prod T1 T2\ fun snd_const T1 T2 = \<^Const>\Product_Type.snd T1 T2\ fun mk_measure_trm f ctxt T = HOLogic.dest_setT T |> fst o HOLogic.dest_prodT |> f |> curry (op $) \<^Const>\measure dummyT\ |> Syntax.check_term ctxt (* wf-goal arising in induction_schema *) fun prove_termination_ind ctxt = let fun mk_size_measure T = case T of \<^Type>\sum T1 T2\ => Sum_Tree.mk_sumcase T1 T2 \<^Type>\nat\ (mk_size_measure T1) (mk_size_measure T2) | \<^Type>\prod T1 T2\ => HOLogic.mk_comp (mk_size_measure T2, snd_const T1 T2) | _ => HOLogic.size_const T val measure_trm = mk_measure_trm (mk_size_measure) ctxt in Function_Relation.relation_tac ctxt measure_trm end (* wf-goal arising in function definitions *) fun prove_termination_fun size_simps ctxt = let fun mk_size_measure T = case T of \<^Type>\sum T1 T2\ => Sum_Tree.mk_sumcase T1 T2 natT (mk_size_measure T1) (mk_size_measure T2) | \<^Type>\prod T1 T2\ => size_prod_const T1 T2 $ (mk_size_measure T1) $ (mk_size_measure T2) | _ => HOLogic.size_const T val measure_trm = mk_measure_trm (mk_size_measure) ctxt val tac = Function_Relation.relation_tac ctxt measure_trm THEN_ALL_NEW simp_tac (put_simpset size_ss ctxt addsimps size_simps) in Function.prove_termination NONE (HEADGOAL tac) ctxt end (** transformations of premises (in inductive proofs) **) (* given the theorem F[t]; proves the theorem F[f t] - F needs to be monotone - f returns either SOME for a term it fires on and NONE elsewhere *) fun map_term f t = (case f t of NONE => map_term' f t | x => x) and map_term' f (t $ u) = (case (map_term f t, map_term f u) of (NONE, NONE) => NONE | (SOME t'', NONE) => SOME (t'' $ u) | (NONE, SOME u'') => SOME (t $ u'') | (SOME t'', SOME u'') => SOME (t'' $ u'')) | map_term' f (Abs (s, T, t)) = (case map_term f t of NONE => NONE | SOME t'' => SOME (Abs (s, T, t''))) | map_term' _ _ = NONE; fun map_thm_tac ctxt tac thm = let val monos = Inductive.get_monos ctxt val simpset = put_simpset HOL_basic_ss ctxt addsimps @{thms split_def} in EVERY [cut_facts_tac [thm] 1, eresolve_tac ctxt [rev_mp] 1, REPEAT_DETERM (FIRSTGOAL (simp_tac simpset THEN' resolve_tac ctxt monos)), - REPEAT_DETERM (resolve_tac ctxt [impI] 1 THEN (assume_tac ctxt 1 ORELSE tac))] + REPEAT_DETERM (resolve_tac ctxt [impI] 1 THEN (assume_tac ctxt 1 ORELSE tac ctxt))] end fun map_thm ctxt f tac thm = let val opt_goal_trm = map_term f (Thm.prop_of thm) in case opt_goal_trm of NONE => thm | SOME goal => - Goal.prove ctxt [] [] goal (fn _ => map_thm_tac ctxt tac thm) + Goal.prove ctxt [] [] goal (fn {context = ctxt', ...} => map_thm_tac ctxt' tac thm) end (* inductive premises can be of the form R ... /\ P ...; split_conj_i picks out the part R or P part *) fun split_conj1 names \<^Const_>\conj for f1 _\ = (case head_of f1 of Const (name, _) => if member (op =) names name then SOME f1 else NONE | _ => NONE) | split_conj1 _ _ = NONE; fun split_conj2 names \<^Const_>\conj for f1 f2\ = (case head_of f1 of Const (name, _) => if member (op =) names name then SOME f2 else NONE | _ => NONE) | split_conj2 _ _ = NONE; fun transform_prem1 ctxt names thm = - map_thm ctxt (split_conj1 names) (eresolve_tac ctxt [conjunct1] 1) thm + map_thm ctxt (split_conj1 names) (fn ctxt' => eresolve_tac ctxt' [conjunct1] 1) thm fun transform_prem2 ctxt names thm = - map_thm ctxt (split_conj2 names) (eresolve_tac ctxt [conjunct2] 1) thm + map_thm ctxt (split_conj2 names) (fn ctxt' => eresolve_tac ctxt' [conjunct2] 1) thm (* transforms a theorem into one of the object logic *) fun atomize ctxt = Conv.fconv_rule (Object_Logic.atomize ctxt) o Thm.forall_intr_vars; fun atomize_rule ctxt i thm = Conv.fconv_rule (Conv.concl_conv i (Object_Logic.atomize ctxt)) thm fun atomize_concl ctxt thm = atomize_rule ctxt (length (Thm.prems_of thm)) thm (* applies a tactic to a formula composed of conjunctions *) fun conj_tac ctxt tac i = let fun select (trm, i) = case trm of \<^Const_>\Trueprop for t'\ => select (t', i) | \<^Const_>\conj for _ _\ => EVERY' [resolve_tac ctxt @{thms conjI}, RANGE [conj_tac ctxt tac, conj_tac ctxt tac]] i | _ => tac i in SUBGOAL select i end end (* structure *) open Nominal_Library; diff --git a/thys/Nominal2/nominal_mutual.ML b/thys/Nominal2/nominal_mutual.ML --- a/thys/Nominal2/nominal_mutual.ML +++ b/thys/Nominal2/nominal_mutual.ML @@ -1,508 +1,514 @@ (* Nominal Mutual Functions Author: Christian Urban heavily based on the code of Alexander Krauss (code forked on 14 January 2011) Joachim Breitner helped with the auxiliary graph definitions (7 August 2012) Mutual recursive nominal function definitions. *) signature NOMINAL_FUNCTION_MUTUAL = sig val prepare_nominal_function_mutual : Nominal_Function_Common.nominal_function_config -> string (* defname *) -> ((string * typ) * mixfix) list -> term list -> local_theory -> ((thm (* goalstate *) * (Proof.context -> thm -> Nominal_Function_Common.nominal_function_result) (* proof continuation *) ) * local_theory) end structure Nominal_Function_Mutual: NOMINAL_FUNCTION_MUTUAL = struct open Function_Lib open Function_Common open Nominal_Function_Common type qgar = string * (string * typ) list * term list * term list * term datatype mutual_part = MutualPart of {i : int, i' : int, fvar : string * typ, cargTs: typ list, f_def: term, f: term option, f_defthm : thm option} datatype mutual_info = Mutual of {n : int, n' : int, fsum_var : string * typ, ST: typ, RST: typ, parts: mutual_part list, fqgars: qgar list, qglrs: ((string * typ) list * term list * term * term) list, fsum : term option} fun mutual_induct_Pnames n = if n < 5 then fst (chop n ["P","Q","R","S"]) else map (fn i => "P" ^ string_of_int i) (1 upto n) fun get_part fname = the o find_first (fn (MutualPart {fvar=(n,_), ...}) => n = fname) (* FIXME *) fun mk_prod_abs e (t1, t2) = let val bTs = rev (map snd e) val T1 = fastype_of1 (bTs, t1) val T2 = fastype_of1 (bTs, t2) in HOLogic.pair_const T1 T2 $ t1 $ t2 end fun analyze_eqs ctxt defname fs eqs = let val num = length fs val fqgars = map (split_def ctxt (K true)) eqs val arity_of = map (fn (fname,_,_,args,_) => (fname, length args)) fqgars |> AList.lookup (op =) #> the fun curried_types (fname, fT) = let val (caTs, uaTs) = chop (arity_of fname) (binder_types fT) in (caTs, uaTs ---> body_type fT) end val (caTss, resultTs) = split_list (map curried_types fs) val argTs = map (foldr1 HOLogic.mk_prodT) caTss val dresultTs = distinct (op =) resultTs val n' = length dresultTs val RST = Balanced_Tree.make (uncurry Sum_Tree.mk_sumT) dresultTs val ST = Balanced_Tree.make (uncurry Sum_Tree.mk_sumT) argTs val fsum_type = ST --> RST val ([fsum_var_name], _) = Variable.add_fixes [ defname ^ "_sum" ] ctxt val fsum_var = (fsum_var_name, fsum_type) fun define (fvar as (n, _)) caTs resultT i = let val vars = map_index (fn (j,T) => Free ("x" ^ string_of_int j, T)) caTs (* FIXME: Bind xs properly *) val i' = find_index (fn Ta => Ta = resultT) dresultTs + 1 val f_exp = Sum_Tree.mk_proj RST n' i' (Free fsum_var $ Sum_Tree.mk_inj ST num i (foldr1 HOLogic.mk_prod vars)) val def = Term.abstract_over (Free fsum_var, fold_rev lambda vars f_exp) val rew = (n, fold_rev lambda vars f_exp) in (MutualPart {i=i, i'=i', fvar=fvar,cargTs=caTs,f_def=def,f=NONE,f_defthm=NONE}, rew) end val (parts, rews) = split_list (map4 define fs caTss resultTs (1 upto num)) fun convert_eqs (f, qs, gs, args, rhs) = let val MutualPart {i, i', ...} = get_part f parts val rhs' = rhs |> map_aterms (fn t as Free (n, _) => the_default t (AList.lookup (op =) rews n) | t => t) in (qs, gs, Sum_Tree.mk_inj ST num i (foldr1 (mk_prod_abs qs) args), Envir.beta_norm (Sum_Tree.mk_inj RST n' i' rhs')) end val qglrs = map convert_eqs fqgars in Mutual {n=num, n'=n', fsum_var=fsum_var, ST=ST, RST=RST, parts=parts, fqgars=fqgars, qglrs=qglrs, fsum=NONE} end fun define_projections fixes mutual fsum lthy = let fun def ((MutualPart {i=i, i'=i', fvar=(fname, fT), cargTs, f_def, ...}), (_, mixfix)) lthy = let val ((f, (_, f_defthm)), lthy') = Local_Theory.define ((Binding.name fname, mixfix), ((Binding.concealed (Binding.name (fname ^ "_def")), []), Term.subst_bound (fsum, f_def))) lthy in (MutualPart {i=i, i'=i', fvar=(fname, fT), cargTs=cargTs, f_def=f_def, f=SOME f, f_defthm=SOME f_defthm }, lthy') end val Mutual { n, n', fsum_var, ST, RST, parts, fqgars, qglrs, ... } = mutual val (parts', lthy') = fold_map def (parts ~~ fixes) lthy in (Mutual { n=n, n'=n', fsum_var=fsum_var, ST=ST, RST=RST, parts=parts', fqgars=fqgars, qglrs=qglrs, fsum=SOME fsum }, lthy') end fun in_context ctxt (f, pre_qs, pre_gs, pre_args, pre_rhs) F = let val oqnames = map fst pre_qs val (qs, _) = Variable.variant_fixes oqnames ctxt |>> map2 (fn (_, T) => fn n => Free (n, T)) pre_qs fun inst t = subst_bounds (rev qs, t) val gs = map inst pre_gs val args = map inst pre_args val rhs = inst pre_rhs val cqs = map (Thm.cterm_of ctxt) qs val (ags, ctxt') = fold_map Thm.assume_hyps (map (Thm.cterm_of ctxt) gs) ctxt val import = fold Thm.forall_elim cqs #> fold Thm.elim_implies ags val export = fold_rev (Thm.implies_intr o Thm.cprop_of) ags #> fold_rev forall_intr_rename (oqnames ~~ cqs) in F ctxt' (f, qs, gs, args, rhs) import export end fun recover_mutual_psimp all_orig_fdefs parts ctxt (fname, _, _, args, rhs) import (export : thm -> thm) sum_psimp_eq = let val (MutualPart {f=SOME f, ...}) = get_part fname parts val psimp = import sum_psimp_eq val ((simp, restore_cond), ctxt') = case cprems_of psimp of [] => ((psimp, I), ctxt) | [cond] => let val (asm, ctxt') = Thm.assume_hyps cond ctxt in ((Thm.implies_elim psimp asm, Thm.implies_intr cond), ctxt') end | _ => raise General.Fail "Too many conditions" in Goal.prove ctxt' [] [] (HOLogic.Trueprop $ HOLogic.mk_eq (list_comb (f, args), rhs)) - (fn _ => (Local_Defs.unfold_tac ctxt' all_orig_fdefs) - THEN EqSubst.eqsubst_tac ctxt' [0] [simp] 1 - THEN (simp_tac ctxt') 1) + (fn {context = goal_ctxt, ...} => + (Local_Defs.unfold_tac goal_ctxt all_orig_fdefs) + THEN EqSubst.eqsubst_tac goal_ctxt [0] [simp] 1 + THEN (simp_tac goal_ctxt) 1) |> restore_cond |> export end val inl_perm = @{lemma "x = Inl y ==> projl (permute p x) = permute p (projl x)" by simp} val inr_perm = @{lemma "x = Inr y ==> projr (permute p x) = permute p (projr x)" by simp} fun recover_mutual_eqvt eqvt_thm all_orig_fdefs parts ctxt (fname, _, _, args, _) import (export : thm -> thm) sum_psimp_eq = let val (MutualPart {f=SOME f, ...}) = get_part fname parts val psimp = import sum_psimp_eq val ((cond, simp, restore_cond), ctxt') = case cprems_of psimp of [] => (([], psimp, I), ctxt) | [cond] => let val (asm, ctxt') = Thm.assume_hyps cond ctxt in (([asm], Thm.implies_elim psimp asm, Thm.implies_intr cond), ctxt') end | _ => raise General.Fail "Too many conditions" val ([p], ctxt'') = ctxt' |> fold Variable.declare_term args |> Variable.variant_fixes ["p"] val p = Free (p, \<^Type>\perm\) val simpset = put_simpset HOL_basic_ss ctxt'' addsimps @{thms permute_sum.simps[symmetric] Pair_eqvt[symmetric] sum.sel} @ [(cond MRS eqvt_thm) RS @{thm sym}] @ [inl_perm, inr_perm, simp] val goal_lhs = mk_perm p (list_comb (f, args)) val goal_rhs = list_comb (f, map (mk_perm p) args) in Goal.prove ctxt'' [] [] (HOLogic.Trueprop $ HOLogic.mk_eq (goal_lhs, goal_rhs)) - (fn _ => Local_Defs.unfold_tac ctxt'' all_orig_fdefs + (fn {context = goal_ctxt, ...} => + Local_Defs.unfold_tac goal_ctxt all_orig_fdefs THEN (asm_full_simp_tac simpset 1)) |> singleton (Proof_Context.export ctxt'' ctxt) |> restore_cond |> export end fun mk_applied_form ctxt caTs thm = let val xs = map_index (fn (i,T) => Thm.cterm_of ctxt (Free ("x" ^ string_of_int i, T))) caTs (* FIXME: Bind xs properly *) in fold (fn x => fn thm => Thm.combination thm (Thm.reflexive x)) xs thm |> Conv.fconv_rule (Thm.beta_conversion true) |> fold_rev Thm.forall_intr xs |> Thm.forall_elim_vars 0 end fun mutual_induct_rules ctxt induct all_f_defs (Mutual {n, ST, parts, ...}) = let val cert = Thm.cterm_of ctxt val newPs = map2 (fn Pname => fn MutualPart {cargTs, ...} => Free (Pname, cargTs ---> HOLogic.boolT)) (mutual_induct_Pnames (length parts)) parts fun mk_P (MutualPart {cargTs, ...}) P = let val avars = map_index (fn (i,T) => Var (("a", i), T)) cargTs val atup = foldr1 HOLogic.mk_prod avars in HOLogic.tupled_lambda atup (list_comb (P, avars)) end val Ps = map2 mk_P parts newPs val case_exp = Sum_Tree.mk_sumcases HOLogic.boolT Ps val induct_inst = Thm.forall_elim (cert case_exp) induct |> full_simplify (put_simpset Sum_Tree.sumcase_split_ss ctxt) |> full_simplify (put_simpset HOL_basic_ss ctxt addsimps all_f_defs) fun project rule (MutualPart {cargTs, i, ...}) k = let val afs = map_index (fn (j,T) => Free ("a" ^ string_of_int (j + k), T)) cargTs (* FIXME! *) val inj = Sum_Tree.mk_inj ST n i (foldr1 HOLogic.mk_prod afs) in (rule |> Thm.forall_elim (cert inj) |> full_simplify (put_simpset Sum_Tree.sumcase_split_ss ctxt) |> fold_rev (Thm.forall_intr o cert) (afs @ newPs), k + length cargTs) end in fst (fold_map (project induct_inst) parts 0) end fun forall_elim s \<^Const_>\Pure.all _ for \Abs (_, _, t)\\ = subst_bound (s, t) | forall_elim _ t = t val forall_elim_list = fold forall_elim fun split_conj_thm th = (split_conj_thm (th RS conjunct1)) @ (split_conj_thm (th RS conjunct2)) handle THM _ => [th]; fun prove_eqvt ctxt fs argTss eqvts_thms induct_thms = let fun aux argTs s = argTs |> map (pair s) |> Variable.variant_frees ctxt fs val argss' = map2 aux argTss (Name.invent (Variable.names_of ctxt) "" (length fs)) val argss = (map o map) Free argss' val arg_namess = (map o map) fst argss' val insts = (map o map) SOME arg_namess val ([p_name], ctxt') = Variable.variant_fixes ["p"] ctxt val p = Free (p_name, \<^Type>\perm\) (* extracting the acc-premises from the induction theorems *) val acc_prems = map Thm.prop_of induct_thms |> map2 forall_elim_list argss |> map (strip_qnt_body @{const_name Pure.all}) |> map (curry Logic.nth_prem 1) |> map HOLogic.dest_Trueprop fun mk_goal acc_prem (f, args) = let val goal_lhs = mk_perm p (list_comb (f, args)) val goal_rhs = list_comb (f, map (mk_perm p) args) in HOLogic.mk_imp (acc_prem, HOLogic.mk_eq (goal_lhs, goal_rhs)) end val goal = fold_conj_balanced (map2 mk_goal acc_prems (fs ~~ argss)) |> HOLogic.mk_Trueprop val induct_thm = case induct_thms of [thm] => thm |> Variable.gen_all ctxt' |> Thm.permute_prems 0 1 |> (fn thm => atomize_rule ctxt' (length (Thm.prems_of thm) - 1) thm) | thms => thms |> map (Variable.gen_all ctxt') |> map (Rule_Cases.add_consumes 1) |> snd o Rule_Cases.strict_mutual_rule ctxt' |> atomize_concl ctxt' fun tac ctxt thm = resolve_tac ctxt [Variable.gen_all ctxt thm] THEN_ALL_NEW assume_tac ctxt in Goal.prove ctxt' (flat arg_namess) [] goal - (fn {context = ctxt'', ...} => - HEADGOAL (DETERM o (resolve_tac ctxt'' [induct_thm]) THEN' - RANGE (map (tac ctxt'') eqvts_thms))) + (fn {context = goal_ctxt, ...} => + HEADGOAL (DETERM o (resolve_tac goal_ctxt [induct_thm]) THEN' + RANGE (map (tac goal_ctxt) eqvts_thms))) |> singleton (Proof_Context.export ctxt' ctxt) |> split_conj_thm |> map (fn th => th RS mp) end fun mk_partial_rules_mutual ctxt inner_cont (m as Mutual {parts, fqgars, ...}) proof = let val result = inner_cont proof val NominalFunctionResult {G, R, cases, psimps, simple_pinducts=[simple_pinduct], termination, domintros, eqvts=[eqvt],...} = result val (all_f_defs, fs) = map (fn MutualPart {f_defthm = SOME f_def, f = SOME f, cargTs, ...} => (mk_applied_form ctxt cargTs (Thm.symmetric f_def), f)) parts |> split_list val all_orig_fdefs = map (fn MutualPart {f_defthm = SOME f_def, ...} => f_def) parts val cargTss = map (fn MutualPart {f = SOME f, cargTs, ...} => cargTs) parts fun mk_mpsimp fqgar sum_psimp = in_context ctxt fqgar (recover_mutual_psimp all_orig_fdefs parts) sum_psimp fun mk_meqvts fqgar sum_psimp = in_context ctxt fqgar (recover_mutual_eqvt eqvt all_orig_fdefs parts) sum_psimp val rew_simpset = put_simpset HOL_basic_ss ctxt addsimps all_f_defs val mpsimps = map2 mk_mpsimp fqgars psimps val minducts = mutual_induct_rules ctxt simple_pinduct all_f_defs m val mtermination = full_simplify rew_simpset termination val mdomintros = Option.map (map (full_simplify rew_simpset)) domintros val meqvts = map2 mk_meqvts fqgars psimps val meqvt_funs = prove_eqvt ctxt fs cargTss meqvts minducts in NominalFunctionResult { fs=fs, G=G, R=R, psimps=mpsimps, simple_pinducts=minducts, cases=cases, termination=mtermination, domintros=mdomintros, eqvts=meqvt_funs } end (* nominal *) fun subst_all s (Q $ Abs(_, _, t)) = let val vs = map Free (Term.add_frees s []) in fold Logic.all vs (subst_bound (s, t)) end fun mk_comp_dummy t s = Const (@{const_name comp}, dummyT) $ t $ s fun all v t = let val T = Term.fastype_of v in Logic.all_const T $ absdummy T (abstract_over (v, t)) end (* nominal *) fun prepare_nominal_function_mutual config defname fixes eqss lthy = let val mutual as Mutual {fsum_var=(n, T), qglrs, ...} = analyze_eqs lthy defname (map fst fixes) (map Envir.beta_eta_contract eqss) - val ((fsum, G, GIntro_thms, G_induct, goalstate, cont), lthy') = + val ((fsum, G, GIntro_thms, G_induct, goalstate, cont), lthy1) = Nominal_Function_Core.prepare_nominal_function config defname [((n, T), NoSyn)] qglrs lthy - val (mutual' as Mutual {n', parts, ST, RST, ...}, lthy'') = define_projections fixes mutual fsum lthy' + val (mutual' as Mutual {n', parts, ST, RST, ...}, lthy2) = define_projections fixes mutual fsum lthy1 (* defining the auxiliary graph *) fun mk_cases (MutualPart {i', fvar as (n, T), ...}) = let val (tys, ty) = strip_type T val fun_var = Free (n ^ "_aux", HOLogic.mk_tupleT tys --> ty) val inj_fun = absdummy dummyT (Sum_Tree.mk_inj RST n' i' (Bound 0)) in - Syntax.check_term lthy'' (mk_comp_dummy inj_fun fun_var) + Syntax.check_term lthy2 (mk_comp_dummy inj_fun fun_var) end val case_sum_exp = map mk_cases parts |> Sum_Tree.mk_sumcases RST val (G_name, G_type) = dest_Free G val G_name_aux = G_name ^ "_aux" val subst = [(G, Free (G_name_aux, G_type))] val GIntros_aux = GIntro_thms |> map Thm.prop_of |> map (Term.subst_free subst) |> map (subst_all case_sum_exp) - val ((G_aux, GIntro_aux_thms, _, G_aux_induct), lthy''') = - Nominal_Function_Core.inductive_def ((Binding.name G_name_aux, G_type), NoSyn) GIntros_aux lthy'' + val ((G_aux, GIntro_aux_thms, _, G_aux_induct), lthy3) = + Nominal_Function_Core.inductive_def ((Binding.name G_name_aux, G_type), NoSyn) GIntros_aux lthy2 - fun mutual_cont ctxt = mk_partial_rules_mutual lthy''' (cont ctxt) mutual' + fun mutual_cont ctxt = mk_partial_rules_mutual lthy3 (cont ctxt) mutual' (* proof of equivalence between graph and auxiliary graph *) val x = Var(("x", 0), ST) val y = Var(("y", 1), RST) val G_aux_prem = HOLogic.mk_Trueprop (G_aux $ x $ y) val G_prem = HOLogic.mk_Trueprop (G $ x $ y) fun mk_inj_goal (MutualPart {i', ...}) = let val injs = Sum_Tree.mk_inj ST n' i' (Bound 0) val projs = y |> Sum_Tree.mk_proj RST n' i' |> Sum_Tree.mk_inj RST n' i' in Const (@{const_name "All"}, dummyT) $ absdummy dummyT (HOLogic.mk_imp (HOLogic.mk_eq(x, injs), HOLogic.mk_eq(projs, y))) end val goal_inj = Logic.mk_implies (G_aux_prem, HOLogic.mk_Trueprop (fold_conj (map mk_inj_goal parts))) |> all x |> all y - |> Syntax.check_term lthy''' + |> Syntax.check_term lthy3 val goal_iff1 = Logic.mk_implies (G_aux_prem, G_prem) |> all x |> all y val goal_iff2 = Logic.mk_implies (G_prem, G_aux_prem) |> all x |> all y val simp_thms = @{thms sum.sel sum.inject sum.case sum.distinct o_apply} - val simpset0 = put_simpset HOL_basic_ss lthy''' addsimps simp_thms - val simpset1 = put_simpset HOL_ss lthy''' addsimps simp_thms - - val inj_thm = Goal.prove lthy''' [] [] goal_inj - (K (HEADGOAL (DETERM o eresolve_tac lthy''' [G_aux_induct] - THEN_ALL_NEW asm_simp_tac simpset1))) - - fun aux_tac thm = - resolve_tac lthy''' [Variable.gen_all lthy''' thm] THEN_ALL_NEW - asm_full_simp_tac (simpset1 addsimps [inj_thm]) + fun simpset0 goal_ctxt = put_simpset HOL_basic_ss goal_ctxt addsimps simp_thms + fun simpset1 goal_ctxt = put_simpset HOL_ss goal_ctxt addsimps simp_thms - val iff1_thm = Goal.prove lthy''' [] [] goal_iff1 - (K (HEADGOAL (DETERM o eresolve_tac lthy''' [G_aux_induct] - THEN' RANGE (map aux_tac GIntro_thms)))) - |> Variable.gen_all lthy''' - val iff2_thm = Goal.prove lthy''' [] [] goal_iff2 - (K (HEADGOAL (DETERM o eresolve_tac lthy''' [G_induct] THEN' RANGE - (map (aux_tac o simplify simpset0) GIntro_aux_thms)))) - |> Variable.gen_all lthy''' + val inj_thm = Goal.prove lthy3 [] [] goal_inj + (fn {context = goal_ctxt, ...} => + HEADGOAL (DETERM o eresolve_tac goal_ctxt [G_aux_induct] + THEN_ALL_NEW asm_simp_tac (simpset1 goal_ctxt))) - val iff_thm = Goal.prove lthy''' [] [] (HOLogic.mk_Trueprop (HOLogic.mk_eq (G, G_aux))) - (K (HEADGOAL (EVERY' ((map (resolve_tac lthy''' o single) @{thms ext ext iffI}) @ - [eresolve_tac lthy''' [iff2_thm], eresolve_tac lthy''' [iff1_thm]])))) + fun aux_tac goal_ctxt thm = + resolve_tac goal_ctxt [Variable.gen_all goal_ctxt thm] THEN_ALL_NEW + asm_full_simp_tac (simpset1 goal_ctxt addsimps [inj_thm]) - val tac = HEADGOAL (simp_tac (put_simpset HOL_basic_ss lthy''' addsimps [iff_thm])) + val iff1_thm = Goal.prove lthy3 [] [] goal_iff1 + (fn {context = goal_ctxt, ...} => + HEADGOAL (DETERM o eresolve_tac goal_ctxt [G_aux_induct] + THEN' RANGE (map (aux_tac goal_ctxt) GIntro_thms))) + |> Variable.gen_all lthy3 + val iff2_thm = Goal.prove lthy3 [] [] goal_iff2 + (fn {context = goal_ctxt, ...} => + HEADGOAL (DETERM o eresolve_tac lthy3 [G_induct] + THEN' RANGE (map (aux_tac goal_ctxt o simplify (simpset0 goal_ctxt)) GIntro_aux_thms))) + |> Variable.gen_all lthy3 + + val iff_thm = Goal.prove lthy3 [] [] (HOLogic.mk_Trueprop (HOLogic.mk_eq (G, G_aux))) + (fn {context = goal_ctxt, ...} => + HEADGOAL (EVERY' ((map (resolve_tac goal_ctxt o single) @{thms ext ext iffI}) @ + [eresolve_tac goal_ctxt [iff2_thm], eresolve_tac goal_ctxt [iff1_thm]]))) + + val tac = HEADGOAL (simp_tac (put_simpset HOL_basic_ss lthy3 addsimps [iff_thm])) val goalstate' = case (SINGLE tac) goalstate of NONE => error "auxiliary equivalence proof failed" | SOME st => st in - ((goalstate', mutual_cont), lthy''') + ((goalstate', mutual_cont), lthy3) end end diff --git a/thys/Nominal2/nominal_thmdecls.ML b/thys/Nominal2/nominal_thmdecls.ML --- a/thys/Nominal2/nominal_thmdecls.ML +++ b/thys/Nominal2/nominal_thmdecls.ML @@ -1,328 +1,328 @@ (* Title: nominal_thmdecls.ML Author: Christian Urban Author: Tjark Weber Infrastructure for the lemma collections "eqvts", "eqvts_raw". Provides the attributes [eqvt] and [eqvt_raw], and the theorem lists "eqvts" and "eqvts_raw". The [eqvt] attribute expects a theorem of the form ?p \ (c ?x1 ?x2 ...) = c (?p \ ?x1) (?p \ ?x2) ... (1) or, if c is a relation with arity >= 1, of the form c ?x1 ?x2 ... ==> c (?p \ ?x1) (?p \ ?x2) ... (2) [eqvt] will store this theorem in the form (1) or, if c is a relation with arity >= 1, in the form c (?p \ ?x1) (?p \ ?x2) ... = c ?x1 ?x2 ... (3) in "eqvts". (The orientation of (3) was chosen because Isabelle's simplifier uses equations from left to right.) [eqvt] will also derive and store the theorem ?p \ c == c (4) in "eqvts_raw". (1)-(4) are all logically equivalent. We consider (1) and (2) to be more end-user friendly, i.e., slightly more natural to understand and prove, while (3) and (4) make the rewriting system for equivariance more predictable and less prone to looping in Isabelle. The [eqvt_raw] attribute expects a theorem of the form (4), and merely stores it in "eqvts_raw". [eqvt_raw] is provided because certain equivariance theorems would lead to looping when used for simplification in the form (1): notably, equivariance of permute (infix \), i.e., ?p \ (?q \ ?x) = (?p \ ?q) \ (?p \ ?x). To support binders such as All/Ex/Ball/Bex etc., which are typically applied to abstractions, argument terms ?xi (as well as permuted arguments ?p \ ?xi) in (1)-(3) need not be eta- contracted, i.e., they may be of the form "%z. ?xi z" or "%z. (?p \ ?x) z", respectively. For convenience, argument terms ?xi (as well as permuted arguments ?p \ ?xi) in (1)-(3) may actually be tuples, e.g., "(?xi, ?xj)" or "(?p \ ?xi, ?p \ ?xj)", respectively. In (1)-(4), "c" is either a (global) constant or a locally fixed parameter, e.g., of a locale or type class. *) signature NOMINAL_THMDECLS = sig val eqvt_add: attribute val eqvt_del: attribute val eqvt_raw_add: attribute val eqvt_raw_del: attribute val get_eqvts_thms: Proof.context -> thm list val get_eqvts_raw_thms: Proof.context -> thm list val eqvt_transform: Proof.context -> thm -> thm val is_eqvt: Proof.context -> term -> bool end; structure Nominal_ThmDecls: NOMINAL_THMDECLS = struct structure EqvtData = Generic_Data ( type T = thm Item_Net.T; val empty = Thm.item_net; val extend = I; val merge = Item_Net.merge); (* EqvtRawData is implemented with a Termtab (rather than an Item_Net) so that we can efficiently decide whether a given constant has a corresponding equivariance theorem stored, cf. the function is_eqvt. *) structure EqvtRawData = Generic_Data ( type T = thm Termtab.table; val empty = Termtab.empty; val extend = I; val merge = Termtab.merge (K true)); val eqvts = Item_Net.content o EqvtData.get val eqvts_raw = map snd o Termtab.dest o EqvtRawData.get val _ = Theory.setup (Global_Theory.add_thms_dynamic (@{binding "eqvts"}, eqvts) #> Global_Theory.add_thms_dynamic (@{binding "eqvts_raw"}, eqvts_raw)) val get_eqvts_thms = eqvts o Context.Proof val get_eqvts_raw_thms = eqvts_raw o Context.Proof (** raw equivariance lemmas **) (* Returns true iff an equivariance lemma exists in "eqvts_raw" for a given term. *) val is_eqvt = Termtab.defined o EqvtRawData.get o Context.Proof (* Returns c if thm is of the form (4), raises an error otherwise. *) fun key_of_raw_thm context thm = let fun error_msg () = error ("Theorem must be of the form \"?p \ c \ c\", with c a constant or fixed parameter:\n" ^ Syntax.string_of_term (Context.proof_of context) (Thm.prop_of thm)) in case Thm.prop_of thm of \<^Const_>\Pure.eq _ for \<^Const_>\permute _ for p c\ c'\ => if is_Var p andalso is_fixed (Context.proof_of context) c andalso c aconv c' then c else error_msg () | _ => error_msg () end fun add_raw_thm thm context = let val c = key_of_raw_thm context thm in if Termtab.defined (EqvtRawData.get context) c then warning ("Replacing existing raw equivariance theorem for \"" ^ Syntax.string_of_term (Context.proof_of context) c ^ "\".") else (); EqvtRawData.map (Termtab.update (c, thm)) context end fun del_raw_thm thm context = let val c = key_of_raw_thm context thm in if Termtab.defined (EqvtRawData.get context) c then EqvtRawData.map (Termtab.delete c) context else ( warning ("Cannot delete non-existing raw equivariance theorem for \"" ^ Syntax.string_of_term (Context.proof_of context) c ^ "\"."); context ) end (** adding/deleting lemmas to/from "eqvts" **) fun add_thm thm context = ( if Item_Net.member (EqvtData.get context) thm then warning ("Theorem already declared as equivariant:\n" ^ Syntax.string_of_term (Context.proof_of context) (Thm.prop_of thm)) else (); EqvtData.map (Item_Net.update thm) context ) fun del_thm thm context = ( if Item_Net.member (EqvtData.get context) thm then EqvtData.map (Item_Net.remove thm) context else ( warning ("Cannot delete non-existing equivariance theorem:\n" ^ Syntax.string_of_term (Context.proof_of context) (Thm.prop_of thm)); context ) ) (** transformation of equivariance lemmas **) (* Transforms a theorem of the form (1) into the form (4). *) local fun tac ctxt thm = let val ss_thms = @{thms "permute_minus_cancel" "permute_prod.simps" "split_paired_all"} in REPEAT o FIRST' [CHANGED o simp_tac (put_simpset HOL_basic_ss ctxt addsimps ss_thms), resolve_tac ctxt [thm RS @{thm trans}], resolve_tac ctxt @{thms trans[OF "permute_fun_def"]} THEN' resolve_tac ctxt @{thms ext}] end in fun thm_4_of_1 ctxt thm = let val (p, c) = thm |> Thm.prop_of |> HOLogic.dest_Trueprop |> HOLogic.dest_eq |> fst |> dest_perm ||> fst o (fixed_nonfixed_args ctxt) val goal = HOLogic.mk_Trueprop (HOLogic.mk_eq (mk_perm p c, c)) val ([goal', p'], ctxt') = Variable.import_terms false [goal, p] ctxt in - Goal.prove ctxt [] [] goal' (fn {context, ...} => tac context thm 1) + Goal.prove ctxt [] [] goal' (fn {context = goal_ctxt, ...} => tac goal_ctxt thm 1) |> singleton (Proof_Context.export ctxt' ctxt) |> (fn th => th RS @{thm "eq_reflection"}) |> zero_var_indexes end handle TERM _ => raise THM ("thm_4_of_1", 0, [thm]) end (* local *) (* Transforms a theorem of the form (2) into the form (1). *) local fun tac ctxt thm thm' = let val ss_thms = @{thms "permute_minus_cancel"(2)} in EVERY' [resolve_tac ctxt @{thms iffI}, dresolve_tac ctxt @{thms permute_boolE}, resolve_tac ctxt [thm], assume_tac ctxt, resolve_tac ctxt @{thms permute_boolI}, dresolve_tac ctxt [thm'], full_simp_tac (put_simpset HOL_basic_ss ctxt addsimps ss_thms)] end in fun thm_1_of_2 ctxt thm = let val (prem, concl) = thm |> Thm.prop_of |> Logic.dest_implies |> apply2 HOLogic.dest_Trueprop (* since argument terms "?p \ ?x1" may actually be eta-expanded or tuples, we need the following function to find ?p *) fun find_perm \<^Const_>\permute _ for \p as Var _\ _\ = p | find_perm \<^Const_>\Pair _ _ for x _\ = find_perm x | find_perm (Abs (_, _, body)) = find_perm body | find_perm _ = raise THM ("thm_3_of_2", 0, [thm]) val p = concl |> dest_comb |> snd |> find_perm val goal = HOLogic.mk_Trueprop (HOLogic.mk_eq (mk_perm p prem, concl)) val ([goal', p'], ctxt') = Variable.import_terms false [goal, p] ctxt val thm' = infer_instantiate ctxt' [(#1 (dest_Var p), Thm.cterm_of ctxt' (mk_minus p'))] thm in - Goal.prove ctxt' [] [] goal' (fn {context = ctxt'', ...} => tac ctxt'' thm thm' 1) + Goal.prove ctxt' [] [] goal' (fn {context = goal_ctxt, ...} => tac goal_ctxt thm thm' 1) |> singleton (Proof_Context.export ctxt' ctxt) end handle TERM _ => raise THM ("thm_1_of_2", 0, [thm]) end (* local *) (* Transforms a theorem of the form (1) into the form (3). *) fun thm_3_of_1 _ thm = (thm RS (@{thm "permute_bool_def"} RS @{thm "sym"} RS @{thm "trans"}) RS @{thm "sym"}) |> zero_var_indexes local val msg = cat_lines ["Equivariance theorem must be of the form", " ?p \ (c ?x1 ?x2 ...) = c (?p \ ?x1) (?p \ ?x2) ...", "or, if c is a relation with arity >= 1, of the form", " c ?x1 ?x2 ... ==> c (?p \ ?x1) (?p \ ?x2) ..."] in (* Transforms a theorem of the form (1) or (2) into the form (4). *) fun eqvt_transform ctxt thm = (case Thm.prop_of thm of \<^Const_>\Trueprop for _\ => thm_4_of_1 ctxt thm | \<^Const_>\Pure.imp for _ _\ => thm_4_of_1 ctxt (thm_1_of_2 ctxt thm) | _ => error msg) handle THM _ => error msg (* Transforms a theorem of the form (1) into theorems of the form (1) (or, if c is a relation with arity >= 1, of the form (3)) and (4); transforms a theorem of the form (2) into theorems of the form (3) and (4). *) fun eqvt_and_raw_transform ctxt thm = (case Thm.prop_of thm of \<^Const_>\Trueprop for \<^Const_>\HOL.eq _ for _ c_args\\ => let val th' = if fastype_of c_args = @{typ "bool"} andalso (not o null) (snd (fixed_nonfixed_args ctxt c_args)) then thm_3_of_1 ctxt thm else thm in (th', thm_4_of_1 ctxt thm) end | \<^Const_>\Pure.imp for _ _\ => let val th1 = thm_1_of_2 ctxt thm in (thm_3_of_1 ctxt th1, thm_4_of_1 ctxt th1) end | _ => error msg) handle THM _ => error msg end (* local *) (** attributes **) val eqvt_raw_add = Thm.declaration_attribute add_raw_thm val eqvt_raw_del = Thm.declaration_attribute del_raw_thm fun eqvt_add_or_del eqvt_fn raw_fn = Thm.declaration_attribute (fn thm => fn context => let val (eqvt, raw) = eqvt_and_raw_transform (Context.proof_of context) thm in context |> eqvt_fn eqvt |> raw_fn raw end) val eqvt_add = eqvt_add_or_del add_thm add_raw_thm val eqvt_del = eqvt_add_or_del del_thm del_raw_thm val _ = Theory.setup (Attrib.setup @{binding "eqvt"} (Attrib.add_del eqvt_add eqvt_del) "Declaration of equivariance lemmas - they will automatically be brought into the form ?p \ c \ c" #> Attrib.setup @{binding "eqvt_raw"} (Attrib.add_del eqvt_raw_add eqvt_raw_del) "Declaration of raw equivariance lemmas - no transformation is performed") 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 diff --git a/thys/Simpl/generalise_state.ML b/thys/Simpl/generalise_state.ML --- a/thys/Simpl/generalise_state.ML +++ b/thys/Simpl/generalise_state.ML @@ -1,301 +1,301 @@ (* Title: generalise_state.ML Author: Norbert Schirmer, TU Muenchen Copyright (C) 2005-2007 Norbert Schirmer This library is free software; you can redistribute it and/or modify it under the terms of the GNU Lesser General Public License as published by the Free Software Foundation; either version 2.1 of the License, or (at your option) any later version. This library is distributed in the hope that it will be useful, but WITHOUT ANY WARRANTY; without even the implied warranty of MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU Lesser General Public License for more details. You should have received a copy of the GNU Lesser General Public License along with this library; if not, write to the Free Software Foundation, Inc., 59 Temple Place, Suite 330, Boston, MA 02111-1307 USA *) signature SPLIT_STATE = sig val isState: term -> bool val abs_state: term -> term option val abs_var: Proof.context -> term -> (string * typ) val split_state: Proof.context -> string -> typ -> term -> (term * term list) val ex_tac: Proof.context -> term list -> tactic (* the term-list is the list of selectors as returned by split_state. They may be used to construct the instatiation of the existentially quantified state. *) end; functor GeneraliseFun (structure SplitState: SPLIT_STATE) = struct val genConj = @{thm generaliseConj}; val genImp = @{thm generaliseImp}; val genImpl = @{thm generaliseImpl}; val genAll = @{thm generaliseAll}; val gen_all = @{thm generalise_all}; val genEx = @{thm generaliseEx}; val genRefl = @{thm generaliseRefl}; val genRefl' = @{thm generaliseRefl'}; val genTrans = @{thm generaliseTrans}; val genAllShift = @{thm generaliseAllShift}; val gen_allShift = @{thm generalise_allShift}; val meta_spec = @{thm meta_spec}; val protectRefl = @{thm protectRefl}; val protectImp = @{thm protectImp}; fun gen_thm decomp (t,ct) = let val (ts,cts,recomb) = decomp (t,ct) in recomb (map (gen_thm decomp) (ts~~cts)) end; fun dest_prop (Const (@{const_name Pure.prop}, _) $ P) = P | dest_prop t = raise TERM ("dest_prop", [t]); fun prem_of thm = #1 (Logic.dest_implies (dest_prop (Thm.prop_of thm))); fun conc_of thm = #2 (Logic.dest_implies (dest_prop (Thm.prop_of thm))); fun dest_All (Const (@{const_name "All"},_)$t) = t | dest_All t = raise TERM ("dest_All",[t]); fun SIMPLE_OF ctxt rule prems = let val mx = fold (fn thm => fn i => Int.max (Thm.maxidx_of thm,i)) prems 0; in DistinctTreeProver.discharge ctxt prems (Thm.incr_indexes (mx + 1) rule) end; infix 0 OF_RAW fun tha OF_RAW thb = thb COMP (Drule.incr_indexes thb tha); fun SIMPLE_OF_RAW ctxt tha thb = SIMPLE_OF ctxt tha [thb]; datatype qantifier = Meta_all | Hol_all | Hol_ex fun list_exists (vs, x) = fold_rev (fn (x, T) => fn P => HOLogic.exists_const T $ Abs (x, T, P)) vs x; fun spec' cv thm = let (* thm = Pure.prop ((all x. P x) ==> Q), where "all" is meta or HOL *) val (ct1,ct2) = thm |> Thm.cprop_of |> Thm.dest_comb |> #2 |> Thm.dest_comb |> #2 |> Thm.dest_comb; in (case Thm.term_of ct1 of Const (@{const_name "Trueprop"},_) => let val (Var (sP,_)$Var (sV,sVT)) = HOLogic.dest_Trueprop (Thm.concl_of spec); val cvT = Thm.ctyp_of_cterm cv; val vT = Thm.typ_of cvT; in Thm.instantiate (TVars.make [(dest_TVar sVT, cvT)], Vars.make [((sP, vT --> HOLogic.boolT), #2 (Thm.dest_comb ct2)), ((sV, vT), cv)]) spec end | Const (@{const_name Pure.all},_) => let val (Var (sP,_)$Var (sV,sVT)) = Thm.concl_of meta_spec; val cvT = Thm.ctyp_of_cterm cv; val vT = Thm.typ_of cvT; in Thm.instantiate (TVars.make [(dest_TVar sVT, cvT)], Vars.make [((sP, vT --> propT), ct2), ((sV, vT),cv)]) meta_spec end | _ => raise THM ("spec'",0,[thm])) end; fun split_thm qnt ctxt s T t = let val (t',vars) = SplitState.split_state ctxt s T t; val vs = map (SplitState.abs_var ctxt) vars; val prop = (case qnt of Meta_all => Logic.list_all (vs,t') | Hol_all => HOLogic.mk_Trueprop (HOLogic.list_all (vs, t')) | Hol_ex => Logic.mk_implies (HOLogic.mk_Trueprop (list_exists (vs, t')), HOLogic.mk_Trueprop (HOLogic.mk_exists (s,T,t)))) in (case qnt of Hol_ex => Goal.prove ctxt [] [] prop (fn _ => SplitState.ex_tac ctxt vars) | _ => let val rP = conc_of genRefl'; val thm0 = Thm.instantiate (TVars.empty, Vars.make [(dest_Var rP, Thm.cterm_of ctxt prop)]) genRefl'; fun elim_all v thm = let val cv = Thm.cterm_of ctxt v; val spc = Goal.protect 0 (spec' cv thm); in SIMPLE_OF ctxt genTrans [thm,spc] end; val thm = fold elim_all vars thm0; in thm end) end; fun eta_expand ctxt ct = let val mi = Thm.maxidx_of_cterm ct; val T = domain_type (Thm.typ_of_cterm ct); val x = Thm.cterm_of ctxt (Var (("s",mi+1),T)); in Thm.lambda x (Thm.apply ct x) end; fun split_abs ctxt ct = (case Thm.term_of ct of - Abs x => (x, Thm.dest_abs NONE ct) + Abs x => (x, Thm.dest_abs_global ct) | _ => split_abs ctxt (eta_expand ctxt ct)) fun decomp ctxt (Const (@{const_name HOL.conj}, _) $ t $ t', ct) = ([t,t'],snd (Drule.strip_comb ct), fn [thm,thm'] => SIMPLE_OF ctxt genConj [thm,thm']) | decomp ctxt ((allc as Const (@{const_name "All"},aT)) $ f, ct) = let val cf = snd (Thm.dest_comb ct); val (abst as (x,T,_),(cx',cb)) = split_abs ctxt cf; val Free (x',_) = Thm.term_of cx'; val (Const (@{const_name Pure.all},_)$Abs (s,_,_)) = genAll |> Thm.prems_of |> hd |> dest_prop; val genAll' = Drule.rename_bvars [(s,x)] genAll; val (Const (@{const_name Pure.all},_)$Abs (s',_,_)) = genAllShift |> Thm.prems_of |> hd |> dest_prop; val genAllShift' = Drule.rename_bvars [(s',x)] genAllShift; in if SplitState.isState (allc$Abs abst) then ([Thm.term_of cb],[cb], fn [thm] => let val P = HOLogic.dest_Trueprop (dest_prop (prem_of thm)); val thm' = split_thm Hol_all ctxt x' T P; val thm1 = genAllShift' OF_RAW Goal.protect 0 (Thm.forall_intr cx' (Goal.conclude thm')); val thm2 = genAll' OF_RAW Goal.protect 0 (Thm.forall_intr cx' (Goal.conclude thm)); in SIMPLE_OF ctxt genTrans [thm1,thm2] end) else ([Thm.term_of cb],[cb], fn [thm] => genAll' OF_RAW Goal.protect 0 (Thm.forall_intr cx' (Goal.conclude thm))) end | decomp ctxt ((exc as Const (@{const_name "Ex"},_)) $ f, ct) = let val cf = snd (Thm.dest_comb ct); val (abst as (x,T,_),(cx',cb)) = split_abs ctxt cf; val Free (x',_) = Thm.term_of cx'; val (Const (@{const_name Pure.all},_)$Abs (s,_,_)) = genEx |> Thm.prems_of |> hd |> dest_prop; val genEx' = Drule.rename_bvars [(s,x)] genEx; in if SplitState.isState (exc$Abs abst) then ([Thm.term_of cb],[cb], fn [thm] => let val P = HOLogic.dest_Trueprop (dest_prop (prem_of thm)); val thm' = split_thm Hol_ex ctxt x' T P; in SIMPLE_OF_RAW ctxt protectImp (Goal.protect 0 thm') end ) else ([Thm.term_of cb],[cb], fn [thm] => genEx' OF_RAW Goal.protect 0 (Thm.forall_intr cx' (Goal.conclude thm))) end | decomp ctxt (Const (@{const_name HOL.implies},_)$P$Q, ct) = let val [cP,cQ] = (snd (Drule.strip_comb ct)); in ([Q],[cQ],fn [thm] => let val X = genImp |> Thm.concl_of |> dest_prop |> Logic.dest_implies |> #1 |> dest_prop |> HOLogic.dest_Trueprop |> HOLogic.dest_imp |> #1 |> dest_Var; val genImp' = Thm.instantiate (TVars.empty, Vars.make [(X,cP)]) genImp; in SIMPLE_OF ctxt genImp' [thm] end) end | decomp ctxt (Const (@{const_name Pure.imp},_)$P$Q, ct) = let val [cP,cQ] = (snd (Drule.strip_comb ct)); in ([Q],[cQ],fn [thm] => let val X = genImpl |> Thm.concl_of |> dest_prop |> Logic.dest_implies |> #1 |> dest_prop |> Logic.dest_implies |> #1 |> dest_Var; val genImpl' = Thm.instantiate (TVars.empty, Vars.make [(X,cP)]) genImpl; in SIMPLE_OF ctxt genImpl' [thm] end) end | decomp ctxt ((allc as Const (@{const_name Pure.all},_)) $ f, ct) = let val cf = snd (Thm.dest_comb ct); val (abst as (x,T,_),(cx',cb)) = split_abs ctxt cf; val Free (x',_) = Thm.term_of cx'; val (Const (@{const_name Pure.all},_)$Abs (s,_,_)) = gen_all |> Thm.prems_of |> hd |> dest_prop; val gen_all' = Drule.rename_bvars [(s,x)] gen_all; val (Const (@{const_name Pure.all},_)$Abs (s',_,_)) = gen_allShift |> Thm.prems_of |> hd |> dest_prop; val gen_allShift' = Drule.rename_bvars [(s',x)] gen_allShift; in if SplitState.isState (allc$Abs abst) then ([Thm.term_of cb],[cb], fn [thm] => let val P = dest_prop (prem_of thm); val thm' = split_thm Meta_all ctxt x' T P; val thm1 = gen_allShift' OF_RAW Goal.protect 0 (Thm.forall_intr cx' (Goal.conclude thm')); val thm2 = gen_all' OF_RAW Goal.protect 0 (Thm.forall_intr cx' (Goal.conclude thm)); in SIMPLE_OF ctxt genTrans [thm1,thm2] end) else ([Thm.term_of cb],[cb], fn [thm] => gen_all' OF_RAW Goal.protect 0 (Thm.forall_intr cx' (Goal.conclude thm))) end | decomp ctxt (Const (@{const_name "Trueprop"},_)$P, ct) = ([P],snd (Drule.strip_comb ct),fn [thm] => thm) | decomp ctxt (t, ct) = ([],[], fn [] => let val rP = HOLogic.dest_Trueprop (dest_prop (conc_of genRefl)); in Thm.instantiate (TVars.empty, Vars.make [(dest_Var rP, ct)]) genRefl end) fun generalise ctxt ct = gen_thm (decomp ctxt) (Thm.term_of ct,ct); (* -------- (init) #C ==> #C *) fun init ct = Thm.instantiate' [] [SOME ct] protectRefl; fun generalise_over_tac ctxt P i st = let val t = List.nth (Thm.prems_of st, i - 1); (* FIXME !? *) in (case P t of SOME t' => let val ct = Thm.cterm_of ctxt t'; val meta_spec_protect' = infer_instantiate ctxt [(("x", 0), ct)] @{thm meta_spec_protect}; in (init (Thm.adjust_maxidx_cterm 0 (List.nth (Drule.cprems_of st, i - 1))) |> resolve_tac ctxt [meta_spec_protect'] 1 |> Seq.maps (fn st' => Thm.bicompose NONE {flatten = true, match = false, incremented = false} (false, Goal.conclude st', Thm.nprems_of st') i st)) end | NONE => no_tac st) end; fun generalise_over_all_states_tac ctxt i = REPEAT (generalise_over_tac ctxt SplitState.abs_state i); fun generalise_tac ctxt i st = let val ct = List.nth (Drule.cprems_of st, i - 1); val ct' = Thm.dest_equals_rhs (Thm.cprop_of (Thm.eta_conversion ct)); val r = Goal.conclude (generalise ctxt ct'); in (init (Thm.adjust_maxidx_cterm 0 (List.nth (Drule.cprems_of st, i - 1))) |> (resolve_tac ctxt [r] 1 THEN resolve_tac ctxt [Drule.protectI] 1) |> Seq.maps (fn st' => Thm.bicompose NONE {flatten = true, match = false, incremented = false} (false, Goal.conclude st', Thm.nprems_of st') i st)) end fun GENERALISE ctxt i = generalise_over_all_states_tac ctxt i THEN generalise_tac ctxt i end; diff --git a/thys/Simpl/hoare.ML b/thys/Simpl/hoare.ML --- a/thys/Simpl/hoare.ML +++ b/thys/Simpl/hoare.ML @@ -1,3410 +1,3408 @@ (* Title: hoare.ML Author: Norbert Schirmer, TU Muenchen Copyright (C) 2004-2007 Norbert Schirmer This library is free software; you can redistribute it and/or modify it under the terms of the GNU Lesser General Public License as published by the Free Software Foundation; either version 2.1 of the License, or (at your option) any later version. This library is distributed in the hope that it will be useful, but WITHOUT ANY WARRANTY; without even the implied warranty of MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU Lesser General Public License for more details. You should have received a copy of the GNU Lesser General Public License along with this library; if not, write to the Free Software Foundation, Inc., 59 Temple Place, Suite 330, Boston, MA 02111-1307 USA *) signature HOARE = sig datatype hoareMode = Partial | Total val gen_proc_rec: Proof.context -> hoareMode -> int -> thm datatype state_kind = Record | Function datatype par_kind = In | Out val deco: string val proc_deco: string val par_deco: string -> string val chopsfx: string -> string -> string val is_state_var: string -> bool val extern: Proof.context -> string -> string val remdeco: Proof.context -> string -> string val remdeco': string -> string val undeco: Proof.context -> term -> term val varname: string -> string val resuffix: string -> string -> string -> string type proc_info = {params: ((par_kind * string) list), recursive: bool, state_kind: state_kind} type hoare_tac = (bool -> int -> tactic) -> Proof.context -> hoareMode -> int -> tactic type hoare_data = {proc_info: proc_info Symtab.table, active_procs: string list list, default_state_kind: state_kind, generate_guard: (stamp * (Proof.context -> term -> term option)), wp_tacs: (string * hoare_tac) list, hoare_tacs: (string * hoare_tac) list, vcg_simps: thm list} val get_data: Proof.context -> hoare_data val get_params: string -> Proof.context -> (par_kind * string) list option val get_default_state_kind: Proof.context -> state_kind val get_state_kind: string -> Proof.context -> state_kind option val clique_name: string list -> string val install_generate_guard: (Proof.context -> term -> term option) -> Context.generic -> Context.generic val generate_guard: Proof.context -> term -> term option val BasicSimpTac: Proof.context -> state_kind -> bool -> thm list -> (int -> tactic) -> int -> tactic val hoare: (Proof.context -> Proof.method) context_parser val hoare_raw: (Proof.context -> Proof.method) context_parser val vcg: (Proof.context -> Proof.method) context_parser val vcg_step: (Proof.context -> Proof.method) context_parser val hoare_rule: (Proof.context -> Proof.method) context_parser val add_foldcongsimps: thm list -> theory -> theory val get_foldcong_ss : theory -> simpset val add_foldcongs : thm list -> theory -> theory val modeqN : string val modexN : string val implementationN : string val specL : string val vcg_tac : string -> string -> string list -> Proof.context -> int -> tactic val hoare_rule_tac : Proof.context -> thm list -> int -> tactic datatype 'a bodykind = BodyTyp of 'a | BodyTerm of 'a val proc_specs : (bstring * string) list parser val add_params : morphism -> string -> (par_kind * string) list -> Context.generic -> Context.generic val set_default_state_kind : state_kind -> Context.generic -> Context.generic val add_state_kind : morphism -> string -> state_kind -> Context.generic -> Context.generic val add_recursive : morphism -> string -> Context.generic -> Context.generic end; structure Hoare: HOARE = struct (* Misc *) val record_vanish = Attrib.setup_config_bool @{binding hoare_record_vanish} (K true); val use_generalise = Attrib.setup_config_bool @{binding hoare_use_generalise} (K false); val sort_variables = Attrib.setup_config_bool @{binding hoare_sort_variables} (K true); val use_cond_inv_modifies = Attrib.setup_config_bool @{binding hoare_use_cond_inv_modifies} (K true); val hoare_trace = Attrib.setup_config_bool @{binding hoare_trace} (K false); val body_def_sfx = "_body"; val programN = "\"; val hoare_ctxtL = "hoare"; val specL = "_spec"; val procL = "_proc"; val bodyP = "_impl"; val modifysfx = "_modifies"; val modexN = "Hoare.mex"; val modeqN = "Hoare.meq"; val KNF = @{const_name StateFun.K_statefun}; (* Some abstract syntax operations *) val Trueprop = HOLogic.mk_Trueprop; infix 0 ===; val (op ===) = Trueprop o HOLogic.mk_eq; fun is_empty_set (Const (@{const_name Orderings.bot}, _)) = true | is_empty_set _ = false; fun mk_Int Ts A B = let val T = fastype_of1 (Ts, A) in Const (@{const_name Lattices.inf}, T --> T --> T) $ A $ B end; fun mk_Un T (A, B) = Const (@{const_name Lattices.sup}, T --> T --> T) $ A $ B; fun dest_Un (Const (@{const_name Lattices.sup}, _) $ t1 $ t2) = dest_Un t1 @ dest_Un t2 | dest_Un t = [t] fun mk_UN' dT rT t = let val dTs = HOLogic.mk_setT dT; val rTs = HOLogic.mk_setT rT; in Const (@{const_name Complete_Lattices.Sup}, rTs --> rT) $ (Const (@{const_name image}, (dT --> rT) --> dTs --> rTs) $ t $ Const (@{const_name Orderings.top}, dTs)) end; fun mk_UN ((x, T), P) = mk_UN' T (fastype_of P) (absfree (x, T) P); fun dest_UN (Const (@{const_name Complete_Lattices.Sup}, _) $ (Const (@{const_name Set.image}, _) $ Abs (x, T, t) $ Const (@{const_name Orderings.top}, _))) = let val (vars, body) = dest_UN t in ((x, T) :: vars, body) end | dest_UN t = ([], t); fun tap_UN (Const (@{const_name Complete_Lattices.Sup}, _) $ (Const (@{const_name Set.image}, _) $ t $ Const (@{const_name Orderings.top}, _))) = SOME t | tap_UN _ = NONE; (* Fetching the rules *) datatype hoareMode = Partial | Total fun get_rule p t Partial = p | get_rule p t Total = t val Guard = get_rule @{thm HoarePartial.Guard} @{thm HoareTotal.Guard}; val GuardStrip = get_rule @{thm HoarePartial.GuardStrip} @{thm HoareTotal.GuardStrip}; val GuaranteeAsGuard = get_rule @{thm HoarePartial.GuaranteeAsGuard} @{thm HoareTotal.GuaranteeAsGuard}; val Guarantee = get_rule @{thm HoarePartial.Guarantee} @{thm HoareTotal.Guarantee}; val GuaranteeStrip = get_rule @{thm HoarePartial.GuaranteeStrip} @{thm HoareTotal.GuaranteeStrip}; val GuardsNil = get_rule @{thm HoarePartial.GuardsNil} @{thm HoareTotal.GuardsNil}; val GuardsCons = get_rule @{thm HoarePartial.GuardsCons} @{thm HoareTotal.GuardsCons}; val GuardsConsGuaranteeStrip = get_rule @{thm HoarePartial.GuardsConsGuaranteeStrip} @{thm HoareTotal.GuardsConsGuaranteeStrip}; val Skip = get_rule @{thm HoarePartial.Skip} @{thm HoareTotal.Skip}; val Basic = get_rule @{thm HoarePartial.Basic} @{thm HoareTotal.Basic}; val BasicCond = get_rule @{thm HoarePartial.BasicCond} @{thm HoareTotal.BasicCond}; val Spec = get_rule @{thm HoarePartial.Spec} @{thm HoareTotal.Spec}; val SpecIf = get_rule @{thm HoarePartial.SpecIf} @{thm HoareTotal.SpecIf}; val Throw = get_rule @{thm HoarePartial.Throw} @{thm HoareTotal.Throw}; val Raise = get_rule @{thm HoarePartial.raise} @{thm HoareTotal.raise}; val Catch = get_rule @{thm HoarePartial.Catch} @{thm HoareTotal.Catch}; val CondCatch = get_rule @{thm HoarePartial.condCatch} @{thm HoareTotal.condCatch}; val CatchSwap = get_rule @{thm HoarePartial.CatchSwap} @{thm HoareTotal.CatchSwap}; val CondCatchSwap = get_rule @{thm HoarePartial.condCatchSwap} @{thm HoareTotal.condCatchSwap}; val Seq = get_rule @{thm HoarePartial.Seq} @{thm HoareTotal.Seq}; val SeqSwap = get_rule @{thm HoarePartial.SeqSwap} @{thm HoareTotal.SeqSwap}; val BSeq = get_rule @{thm HoarePartial.BSeq} @{thm HoareTotal.BSeq}; val Cond = get_rule @{thm HoarePartial.Cond} @{thm HoareTotal.Cond}; val CondInv'Partial = @{thm HoarePartial.CondInv'}; val CondInv'Total = @{thm HoareTotal.CondInv'}; val CondInv' = get_rule CondInv'Partial CondInv'Total; val SwitchNil = get_rule @{thm HoarePartial.switchNil} @{thm HoareTotal.switchNil}; val SwitchCons = get_rule @{thm HoarePartial.switchCons} @{thm HoareTotal.switchCons}; val CondSwap = get_rule @{thm HoarePartial.CondSwap} @{thm HoareTotal.CondSwap}; val While = get_rule @{thm HoarePartial.While} @{thm HoareTotal.While}; val WhileAnnoG = get_rule @{thm HoarePartial.WhileAnnoG} @{thm HoareTotal.WhileAnnoG}; val WhileAnnoFix = get_rule @{thm HoarePartial.WhileAnnoFix'} @{thm HoareTotal.WhileAnnoFix'}; val WhileAnnoGFix = get_rule @{thm HoarePartial.WhileAnnoGFix} @{thm HoareTotal.WhileAnnoGFix}; val BindR = get_rule @{thm HoarePartial.Bind} @{thm HoareTotal.Bind}; val Block = get_rule @{thm HoarePartial.Block} @{thm HoareTotal.Block}; val BlockSwap = get_rule @{thm HoarePartial.BlockSwap} @{thm HoareTotal.BlockSwap}; val Proc = get_rule @{thm HoarePartial.ProcSpec} @{thm HoareTotal.ProcSpec}; val ProcNoAbr = get_rule @{thm HoarePartial.ProcSpecNoAbrupt} @{thm HoareTotal.ProcSpecNoAbrupt}; val ProcBody = get_rule @{thm HoarePartial.ProcBody} @{thm HoareTotal.ProcBody}; val CallBody = get_rule @{thm HoarePartial.CallBody} @{thm HoareTotal.CallBody}; val FCall = get_rule @{thm HoarePartial.FCall} @{thm HoareTotal.FCall}; val ProcRecSpecs = get_rule @{thm HoarePartial.ProcRecSpecs} @{thm HoareTotal.ProcRecSpecs}; val ProcModifyReturnSameFaults = get_rule @{thm HoarePartial.ProcModifyReturnSameFaults} @{thm HoareTotal.ProcModifyReturnSameFaults}; val ProcModifyReturn = get_rule @{thm HoarePartial.ProcModifyReturn} @{thm HoareTotal.ProcModifyReturn}; val ProcModifyReturnNoAbr = get_rule @{thm HoarePartial.ProcModifyReturnNoAbr} @{thm HoareTotal.ProcModifyReturnNoAbr}; val ProcModifyReturnNoAbrSameFaultsPartial = @{thm HoarePartial.ProcModifyReturnNoAbrSameFaults}; val ProcModifyReturnNoAbrSameFaultsTotal = @{thm HoareTotal.ProcModifyReturnNoAbrSameFaults}; val ProcModifyReturnNoAbrSameFaults = get_rule ProcModifyReturnNoAbrSameFaultsPartial ProcModifyReturnNoAbrSameFaultsTotal; val TrivPost = get_rule @{thm HoarePartial.TrivPost} @{thm HoareTotal.TrivPost}; val TrivPostNoAbr = get_rule @{thm HoarePartial.TrivPostNoAbr} @{thm HoareTotal.TrivPostNoAbr}; val DynProcProcPar = get_rule @{thm HoarePartial.DynProcProcPar} @{thm HoareTotal.DynProcProcPar}; val DynProcProcParNoAbr = get_rule @{thm HoarePartial.DynProcProcParNoAbrupt} @{thm HoareTotal.DynProcProcParNoAbrupt}; val ProcProcParModifyReturn = get_rule @{thm HoarePartial.ProcProcParModifyReturn} @{thm HoareTotal.ProcProcParModifyReturn}; val ProcProcParModifyReturnSameFaultsPartial = @{thm HoarePartial.ProcProcParModifyReturnSameFaults}; val ProcProcParModifyReturnSameFaultsTotal = @{thm HoareTotal.ProcProcParModifyReturnSameFaults}; val ProcProcParModifyReturnSameFaults = get_rule ProcProcParModifyReturnSameFaultsPartial ProcProcParModifyReturnSameFaultsTotal; val ProcProcParModifyReturnNoAbr = get_rule @{thm HoarePartial.ProcProcParModifyReturnNoAbr} @{thm HoareTotal.ProcProcParModifyReturnNoAbr}; val ProcProcParModifyReturnNoAbrSameFaultsPartial = @{thm HoarePartial.ProcProcParModifyReturnNoAbrSameFaults}; val ProcProcParModifyReturnNoAbrSameFaultsTotal = @{thm HoareTotal.ProcProcParModifyReturnNoAbrSameFaults}; val ProcProcParModifyReturnNoAbrSameFaults = get_rule ProcProcParModifyReturnNoAbrSameFaultsPartial ProcProcParModifyReturnNoAbrSameFaultsTotal; val DynCom = get_rule @{thm HoarePartial.DynComConseq} @{thm HoareTotal.DynComConseq}; val AugmentContext = get_rule @{thm HoarePartial.augment_context'} @{thm HoareTotal.augment_context'}; val AugmentEmptyFaults = get_rule @{thm HoarePartial.augment_emptyFaults} @{thm HoareTotal.augment_emptyFaults}; val AsmUN = get_rule @{thm HoarePartial.AsmUN} @{thm HoareTotal.AsmUN}; val SpecAnno = get_rule @{thm HoarePartial.SpecAnno'} @{thm HoareTotal.SpecAnno'}; val SpecAnnoNoAbrupt = get_rule @{thm HoarePartial.SpecAnnoNoAbrupt} @{thm HoareTotal.SpecAnnoNoAbrupt}; val LemAnno = get_rule @{thm HoarePartial.LemAnno} @{thm HoareTotal.LemAnno}; val LemAnnoNoAbrupt = get_rule @{thm HoarePartial.LemAnnoNoAbrupt} @{thm HoareTotal.LemAnnoNoAbrupt}; val singleton_conv_sym = @{thm Set.singleton_conv2} RS sym; val anno_defs = [@{thm Language.whileAnno_def},@{thm Language.whileAnnoG_def},@{thm Language.specAnno_def}, @{thm Language.whileAnnoGFix_def},@{thm Language.whileAnnoFix_def},@{thm Language.lem_def}]; val strip_simps = @{thm Language.strip_simp} :: @{thm Option.option.map(2)} :: @{thms Language.strip_guards_simps}; val normalize_simps = [@{thm Language.while_def}, @{thm Language.bseq_def}, @{thm List.append_Nil}, @{thm List.append_Cons}] @ @{thms List.list.cases} @ @{thms Language.flatten_simps} @ @{thms Language.sequence.simps} @ @{thms Language.normalize_simps} @ @{thms Language.guards.simps} @ [@{thm fst_conv}, @{thm snd_conv}]; val K_rec_convs = []; val K_fun_convs = [@{thm StateFun.K_statefun_apply}, @{thm StateFun.K_statefun_comp}]; val K_convs = K_rec_convs @ K_fun_convs; val K_rec_congs = []; val K_fun_congs = [@{thm StateFun.K_statefun_cong}]; val K_congs = K_rec_congs @ K_fun_congs; (* misc. aux. functions *) (* first_subterm * yields result x of P for first subterm for which P is (SOME x), and all bound variables on the path * to that term *) fun first_subterm_dest P = let fun first abs_vars t = (case P t of SOME x => SOME (abs_vars,x) |_=> (case t of u $ v => (case first abs_vars u of NONE => first abs_vars v | SOME x => SOME x) | Abs (c,T,u) => first (abs_vars @ [(c,T)]) u | _ => NONE)) in first [] end; (* first_subterm * yields first subterm for which P holds, and all bound variables on the path * to that term *) fun first_subterm P = let fun P' t = if P t then SOME t else NONE; in first_subterm_dest P' end; (* max_subterm_dest * yields results of P for all maximal subterms for which P is (SOME x), * and all bound variables on the path to that subterm *) fun max_subterms_dest P = let fun collect abs_vars t = (case P t of SOME x => [(abs_vars,x)] |_=> (case t of u $ v => collect abs_vars u @ collect abs_vars v | Abs (c,T,u) => collect (abs_vars @ [(c,T)]) u | _ => [])) in collect [] end; fun last [] = raise Empty | last [x] = x | last (_::xs) = last xs; fun dest_splits (Const (@{const_name case_prod},_)$Abs (n,T,t)) = (n,T)::dest_splits t | dest_splits (Const (@{const_name case_prod},_)$Abs (n,T,t)$_) = (n,T)::dest_splits t | dest_splits (Abs (n,T,_)) = [(n,T)] | dest_splits _ = []; fun idx eq [] x = ~1 | idx eq (x::rs) y = if eq x y then 0 else let val i = idx eq rs y in if i < 0 then i else i+1 end; fun resuffix sfx1 sfx2 s = suffix sfx2 (unsuffix sfx1 s) handle Fail _ => s; (* state space representation dependent functions *) datatype state_kind = Record | Function fun state_simprocs Record = [Record.simproc] | state_simprocs Function = [Record.simproc, StateFun.lookup_simproc]; fun state_upd_simproc Record = Record.upd_simproc | state_upd_simproc Function = StateFun.update_simproc; fun state_ex_sel_eq_simproc Record = Record.ex_sel_eq_simproc | state_ex_sel_eq_simproc Function = StateFun.ex_lookup_eq_simproc; val state_split_simp_tac = Record.split_simp_tac val state_hierarchy = Record.dest_recTs fun stateT_id T = case (state_hierarchy T) of [] => NONE | Ts => SOME (last Ts); fun globalsT (Type (_, T :: _)) = SOME T | globalsT _ = NONE; fun stateT_ids T = (case stateT_id T of NONE => NONE | SOME sT => (case globalsT T of NONE => SOME [sT] | SOME gT => (case stateT_id gT of NONE => SOME [sT] | SOME gT' => SOME [sT,gT']))); datatype par_kind = In | Out (*** utilities ***) (* utils for variable name decorations *) val deco = "_'"; val proc_deco = "_'proc"; fun par_deco name = deco ^ name ^ deco; fun chopsfx sfx str = (case try (unsuffix sfx) str of SOME s => s | NONE => str) val is_state_var = can (unsuffix deco); (* removes the suffix of the string beginning with deco. * "xys_'a" --> "xys"; * The a is also chopped, since sometimes the bound variables * are renamed, I think SELECT_GOAL in rename_goal is to blame *) fun remdeco' str = let fun chop (p::ps) (x::xs) = chop ps xs | chop [] xs = [] | chop (p::ps) [] = error "remdeco: code should never be reached"; fun remove prf (s as (x::xs)) = if is_prefix (op =) prf s then chop prf s else (x::remove prf xs) | remove prf [] = []; in String.implode (remove (String.explode deco) (String.explode str)) end; fun extern ctxt s = (case try (Proof_Context.extern_const ctxt o Lexicon.unmark_const) s of NONE => s | SOME s' => s'); fun remdeco ctxt s = remdeco' (extern ctxt s); fun undeco ctxt (Const (c, T)) = Const (remdeco ctxt c, T) | undeco ctxt ((f as Const (@{syntax_const "_free"},_)) $ Free (x, T)) = (*f$*)Const (remdeco' x, T) | undeco ctxt (Const _ $ _ $ ((Const (@{syntax_const "_free"},_)) $ Free (x, T))) = (*f$*)Const (remdeco' x, T) | undeco ctxt (Free (c, T)) = Const (remdeco' c, T) | undeco ctxt x = x fun varname x = x ^ deco val dest_string = map (chr o HOLogic.dest_char) o HOLogic.dest_list; fun dest_string' t = (case try dest_string t of SOME s => implode s | NONE => (case t of Free (s,_) => s | Const (s,_) => Long_Name.base_name s | _ => raise TERM ("dest_string'",[t]))) fun is_state_space_var Tids t = let fun is_stateT T = (case stateT_id T of NONE => 0 | SOME id => if member (op =) Tids id then ~1 else 0); in (case t of Const _ $ Abs (_,T,_) => is_stateT T | Free (_,T) => is_stateT T | _ => 0) end; datatype callMode = Static | Parameter fun proc_name Static (Const (p,_)$_) = resuffix deco proc_deco (Long_Name.base_name p) | proc_name Static (Const (@{const_name StateFun.lookup},_)$_$Free (p,_)$_) = suffix proc_deco (remdeco' (Long_Name.base_name p)) | proc_name Static p = dest_string' p | proc_name Parameter (Const (p,_)) = resuffix deco proc_deco (Long_Name.base_name p) | proc_name Parameter (Abs (_,_,Const (p,_)$Bound 0)) = resuffix deco proc_deco (Long_Name.base_name p) | proc_name Parameter (Abs (_,_,Const (@{const_name StateFun.lookup},_)$_$Free (p,_)$_)) = suffix proc_deco (remdeco' (Long_Name.base_name p)) | proc_name _ t = raise TERM ("proc_name",[t]); fun dest_call (Const (@{const_name Language.call},_)$init$pname$return$c) = (init,pname,return,c,Static,true) | dest_call (Const (@{const_name Language.fcall},_)$init$pname$return$_$c) = (init,pname,return,c,Static,true) | dest_call (Const (@{const_name Language.com.Call},_)$pname) = (Bound 0,pname,Bound 0,Bound 0,Static,false) | dest_call (Const (@{const_name Language.dynCall},_)$init$pname$return$c) = (init,pname,return,c,Parameter,true) | dest_call t = raise TERM ("Hoare.dest_call: unexpected term",[t]); fun dest_whileAnno (Const (@{const_name Language.whileAnnoG},_) $gs$b$I$V$c) = (SOME gs,b,I,V,c,false) | dest_whileAnno (Const (@{const_name Language.whileAnno},_) $b$I$V$c) = (NONE,b,I,V,c,false) | dest_whileAnno (Const (@{const_name Language.whileAnnoGFix},_)$gs$b$I$V$c) = (SOME gs,b,I,V,c,true) | dest_whileAnno (Const (@{const_name Language.whileAnnoFix},_) $b$I$V$c) = (NONE,b,I,V,c,true) | dest_whileAnno t = raise TERM ("Hoare.dest_while: unexpected term",[t]); fun dest_Guard (Const (@{const_name Language.com.Guard},_)$f$g$c) = (f,g,c,false) | dest_Guard (Const (@{const_name Language.guaranteeStrip},_)$f$g$c) = (f,g,c,true) | dest_Guard t = raise TERM ("Hoare.dest_guard: unexpected term",[t]); (*** extend theory by procedure definition ***) fun add_declaration name decl thy = thy |> Named_Target.init [] name |> Local_Theory.declaration {syntax = false, pervasive = false} decl |> Local_Theory.exit |> Proof_Context.theory_of; (* data kind 'HOL/hoare' *) type proc_info = {params: ((par_kind * string) list), recursive: bool, state_kind: state_kind} type hoare_tac = (bool -> int -> tactic) -> Proof.context -> hoareMode -> int -> tactic; type hoare_data = {proc_info: proc_info Symtab.table, active_procs: string list list, default_state_kind: state_kind, generate_guard: (stamp * (Proof.context -> term -> term option)), wp_tacs: (string * hoare_tac) list, hoare_tacs: (string * hoare_tac) list, vcg_simps: thm list}; fun make_hoare_data proc_info active_procs default_state_kind generate_guard wp_tacs hoare_tacs vcg_simps = {proc_info = proc_info, active_procs = active_procs, default_state_kind = default_state_kind, generate_guard = generate_guard, wp_tacs = wp_tacs, hoare_tacs = hoare_tacs, vcg_simps = vcg_simps}; structure Hoare_Data = Generic_Data ( type T = hoare_data; val empty = make_hoare_data (Symtab.empty: proc_info Symtab.table) ([]:string list list) (Function) (stamp (),(K (K NONE)): Proof.context -> term -> term option) ([]:(string * hoare_tac) list) ([]:(string * hoare_tac) list) ([]:thm list); val extend = I; (* FIXME exponential blowup due to append !? *) fun merge ({proc_info = proc_info1, active_procs = active_procs1, default_state_kind = _, generate_guard = (stmp1,generate_gaurd1), wp_tacs = wp_tacs1, hoare_tacs = hoare_tacs1, vcg_simps = vcg_simps1}, {proc_info = proc_info2, active_procs = active_procs2, default_state_kind = default_state_kind2, generate_guard = (stmp2, _), wp_tacs = wp_tacs2, hoare_tacs = hoare_tacs2, vcg_simps=vcg_simps2}) : T = if stmp1=stmp2 then make_hoare_data (Symtab.merge (K true) (proc_info1,proc_info2)) (active_procs1 @ active_procs2) (default_state_kind2) (stmp1,generate_gaurd1) (wp_tacs1 @ wp_tacs2) (hoare_tacs1 @ hoare_tacs2) (Thm.merge_thms (vcg_simps1,vcg_simps2)) else error ("Theories have different aux. functions to generate guards") ); val get_data = Hoare_Data.get o Context.Proof; (* access 'params' *) -fun mk_free ctxt name = +fun mk_free context name = let - val ctxt' = Context.proof_of ctxt; - val n' = Variable.intern_fixed ctxt' name |> perhaps Long_Name.dest_hidden; - val T' = Proof_Context.infer_type ctxt' (n', dummyT) handle ERROR _ => dummyT + val ctxt = Context.proof_of context; + val n' = Variable.intern_fixed ctxt name |> perhaps Long_Name.dest_hidden; + val T' = Proof_Context.infer_type ctxt (n', dummyT) handle ERROR _ => dummyT in (Free (n',T')) end; -fun morph_name ctxt phi name = - (case Morphism.term phi (mk_free ctxt name) of +fun morph_name context phi name = + (case Morphism.term phi (mk_free context name) of Free (x,_) => x | _ => name); datatype 'a bodykind = BodyTyp of 'a | BodyTerm of 'a -fun set_default_state_kind sk ctxt = +fun set_default_state_kind sk context = let val {proc_info,active_procs,default_state_kind,generate_guard,wp_tacs,hoare_tacs, vcg_simps,...} - = Hoare_Data.get ctxt; + = Hoare_Data.get context; val data = make_hoare_data proc_info active_procs sk generate_guard wp_tacs hoare_tacs vcg_simps; - in Hoare_Data.put data ctxt end; + in Hoare_Data.put data context end; val get_default_state_kind = #default_state_kind o get_data; -fun add_active_procs phi ps ctxt = +fun add_active_procs phi ps context = let val {proc_info,active_procs,default_state_kind,generate_guard,wp_tacs,hoare_tacs, vcg_simps,...} - = Hoare_Data.get ctxt; + = Hoare_Data.get context; val data = make_hoare_data proc_info - ((map (morph_name ctxt phi) ps)::active_procs) + ((map (morph_name context phi) ps)::active_procs) default_state_kind generate_guard wp_tacs hoare_tacs vcg_simps; - in Hoare_Data.put data ctxt end; - - -fun add_hoare_tacs tacs ctxt = + in Hoare_Data.put data context end; + + +fun add_hoare_tacs tacs context = let val {proc_info,active_procs, default_state_kind, generate_guard,wp_tacs,hoare_tacs, vcg_simps,...} - = Hoare_Data.get ctxt; + = Hoare_Data.get context; val data = make_hoare_data proc_info active_procs default_state_kind generate_guard wp_tacs (hoare_tacs@tacs) vcg_simps; - in Hoare_Data.put data ctxt end; - -fun map_vcg_simps f ctxt = + in Hoare_Data.put data context end; + +fun map_vcg_simps f context = let val {proc_info,active_procs,default_state_kind,generate_guard,wp_tacs,hoare_tacs, vcg_simps,...} - = Hoare_Data.get ctxt; + = Hoare_Data.get context; val data = make_hoare_data proc_info active_procs default_state_kind generate_guard wp_tacs hoare_tacs (f vcg_simps); - in Hoare_Data.put data ctxt end; + in Hoare_Data.put data context end; fun thy_attrib f = Thm.declaration_attribute (fn thm => map_vcg_simps (f thm)); val vcg_simpadd = Thm.add_thm val vcg_simpdel = Thm.del_thm val vcg_simp_add = thy_attrib vcg_simpadd; val vcg_simp_del = thy_attrib vcg_simpdel; (* add 'procedure' *) fun mk_proc_info params recursive state_kind = {params=params,recursive=recursive,state_kind=state_kind}; val empty_proc_info = {params=[],recursive=false,state_kind=Record}; fun map_proc_info_params f {params,recursive,state_kind} = mk_proc_info (f params) recursive state_kind; fun map_proc_info_recursive f {params,recursive,state_kind} = mk_proc_info params (f recursive) state_kind; fun map_proc_info_state_kind f {params,recursive,state_kind} = mk_proc_info params recursive (f state_kind); -fun add_params phi name frmls ctxt = +fun add_params phi name frmls context = let val {proc_info,active_procs,default_state_kind,generate_guard,wp_tacs,hoare_tacs, vcg_simps,...} - = Hoare_Data.get ctxt; - val params = map (apsnd (morph_name ctxt phi)) frmls; + = Hoare_Data.get context; + val params = map (apsnd (morph_name context phi)) frmls; val f = map_proc_info_params (K params); val default = f empty_proc_info; - val proc_info' = Symtab.map_default (morph_name ctxt phi name,default) f proc_info; + val proc_info' = Symtab.map_default (morph_name context phi name,default) f proc_info; val data = make_hoare_data proc_info' active_procs default_state_kind generate_guard wp_tacs hoare_tacs vcg_simps; - in Hoare_Data.put data ctxt end; + in Hoare_Data.put data context end; fun get_params name ctxt = Option.map #params (Symtab.lookup (#proc_info (get_data ctxt)) name); -fun add_recursive phi name ctxt = +fun add_recursive phi name context = let val {proc_info,active_procs,default_state_kind,generate_guard,wp_tacs,hoare_tacs, vcg_simps,...} - = Hoare_Data.get ctxt; + = Hoare_Data.get context; val f = map_proc_info_recursive (K true); val default = f empty_proc_info; - val proc_info'= Symtab.map_default (morph_name ctxt phi name,default) f proc_info; + val proc_info'= Symtab.map_default (morph_name context phi name,default) f proc_info; val data = make_hoare_data proc_info' active_procs default_state_kind generate_guard wp_tacs hoare_tacs vcg_simps; - in Hoare_Data.put data ctxt end; + in Hoare_Data.put data context end; fun get_recursive name ctxt = Option.map #recursive (Symtab.lookup (#proc_info (get_data ctxt)) name); -fun add_state_kind phi name sk ctxt = +fun add_state_kind phi name sk context = let val {proc_info,active_procs,default_state_kind,generate_guard,wp_tacs,hoare_tacs, vcg_simps,...} - = Hoare_Data.get ctxt; + = Hoare_Data.get context; val f = map_proc_info_state_kind (K sk); val default = f empty_proc_info; - val proc_info'= Symtab.map_default (morph_name ctxt phi name,default) f proc_info; + val proc_info'= Symtab.map_default (morph_name context phi name,default) f proc_info; val data = make_hoare_data proc_info' active_procs default_state_kind generate_guard wp_tacs hoare_tacs vcg_simps; - in Hoare_Data.put data ctxt end; + in Hoare_Data.put data context end; fun get_state_kind name ctxt = Option.map #state_kind (Symtab.lookup (#proc_info (get_data ctxt)) name); -fun install_generate_guard f ctxt = +fun install_generate_guard f context = let val {proc_info,active_procs, default_state_kind, generate_guard,wp_tacs,hoare_tacs, vcg_simps,...} = - Hoare_Data.get ctxt; + Hoare_Data.get context; val data = make_hoare_data proc_info active_procs default_state_kind (stamp (), f) wp_tacs hoare_tacs vcg_simps - in Hoare_Data.put data ctxt end; + in Hoare_Data.put data context end; fun generate_guard ctxt = snd (#generate_guard (get_data ctxt)) ctxt; fun check_procedures_definition procs thy = let val ctxt = Proof_Context.init_global thy; fun already_defined name = if is_some (get_params name ctxt) then ["procedure " ^ quote name ^ " already defined"] else [] val err_already_defined = maps (already_defined o #1) procs; fun duplicate_procs names = (case duplicates (op =) names of [] => [] | dups => ["Duplicate procedures " ^ commas_quote dups]); val err_duplicate_procs = duplicate_procs (map #1 procs); fun duplicate_pars name pars = (case duplicates (op =) (map fst pars) of [] => [] | dups => ["Duplicate parameters in procedure " ^ quote name ^ ": " ^ commas_quote dups]); val err_duplicate_pars = maps (fn (name,inpars,outpars,locals,_,_,_) => duplicate_pars name (inpars @ locals) @ duplicate_pars name (outpars @ locals)) procs; (* FIXME: Check that no global variables are used as result parameters *) val errs = err_already_defined @ err_duplicate_procs @ err_duplicate_pars; in if null errs then () else error (cat_lines errs) end; -fun add_parameter_info phi cname (name,(inpars,outpars,state_kind)) ctxt = +fun add_parameter_info phi cname (name,(inpars,outpars,state_kind)) context = let fun par_deco' T = if T = "" then deco else par_deco (cname name); val pars = map (fn (par,T) => (In,suffix (par_deco' T) par)) inpars@ map (fn (par,T) => (Out,suffix (par_deco' T) par)) outpars; - - val ctxt_decl = ctxt - |> add_params phi name pars - |> add_state_kind phi name state_kind - in ctxt_decl + in + context + |> add_params phi name pars + |> add_state_kind phi name state_kind end; fun mk_loc_exp xs = let fun mk_expr s = (s,(("",false),(Expression.Named [],[]))) in (map mk_expr xs,[]) end; val parametersN = "_parameters"; val variablesN = "_variables"; val signatureN = "_signature"; val bodyN = "_body"; val implementationN = "_impl"; val cliqueN = "_clique"; val clique_namesN = "_clique_names"; val NoBodyN = @{const_name Vcg.NoBody}; val statetypeN = "StateType"; val proc_nameT = HOLogic.stringT; fun expression_no_pos (expr, fixes) : Expression.expression = (map (fn (name, inst) => ((name, Position.none), inst)) expr, fixes); fun add_locale name expr elems thy = thy |> Expression.add_locale (Binding.name name) (Binding.name name) [] expr elems |> snd |> Local_Theory.exit; fun add_locale' name expr elems thy = thy |> Expression.add_locale (Binding.name name) (Binding.name name) [] expr elems ||> Local_Theory.exit; fun add_locale_cmd name expr elems thy = thy |> Expression.add_locale_cmd (Binding.name name) (Binding.name name) [] (expression_no_pos expr) elems |> snd |> Local_Theory.exit; fun read_typ thy raw_T env = let val ctxt' = Proof_Context.init_global thy |> fold (Variable.declare_typ o TFree) env; val T = Syntax.read_typ ctxt' raw_T; val env' = Term.add_tfreesT T env; in (T, env') end; fun add_variable_statespaces (cname, (inpars, outpars, locvars)) thy = let val inpars' = if forall (fn (_,T) => T = "") inpars then [] else inpars; val outpars' = if forall (fn (_,T) => T = "") outpars then [] else outpars; fun prep_comp (n, T) env = let val (T', env') = read_typ thy T env handle ERROR msg => cat_error msg ("The error(s) above occurred in component " ^ quote n) in ((n, T'), env') end; val (in_outs,in_out_env) = fold_map prep_comp (distinct (op =) (inpars'@outpars')) []; val (locs,var_env) = fold_map prep_comp locvars in_out_env; val parSP = cname ^ parametersN; val in_outs' = map (apfst (suffix (par_deco cname))) in_outs; val in_out_args = map fst in_out_env; val varSP = cname ^ variablesN; val locs' = map (apfst (suffix (par_deco cname))) locs; val var_args = map fst var_env; in if null inpars' andalso null outpars' andalso null locvars then thy |> add_locale_cmd parSP ([],[]) [] |> Proof_Context.theory_of |> add_locale_cmd varSP ([],[]) [] |> Proof_Context.theory_of else thy |> StateSpace.define_statespace_i (SOME statetypeN) in_out_args parSP [] in_outs' |> StateSpace.define_statespace_i (SOME statetypeN) var_args varSP [((cname, false), ((map TFree in_out_env),parSP,[]))] locs' end; fun intern_locale thy = Locale.intern thy #> perhaps Long_Name.dest_hidden; fun apply_in_context thy lexp f t = let fun name_variant lname = if intern_locale thy lname = lname then lname else name_variant (lname ^ "'"); in thy (* Create a dummy locale in dummy theory just to read the term *) |> add_locale_cmd (name_variant "foo") lexp [] |> (fn ctxt => f ctxt t) end; fun add_abbrev loc mode name spec thy = thy |> Named_Target.init [] loc |> (fn lthy => let val t = Syntax.read_term (Local_Theory.target_of lthy) spec; in Local_Theory.abbrev mode ((Binding.name name, NoSyn), t) lthy end) |> #2 |> Local_Theory.exit |> Proof_Context.theory_of; exception TOPSORT of string fun topsort less [] = [] | topsort less xs = let fun list_all P xs = fold (fn x => fn b => b andalso P x) xs true; fun split_min n (x::xs) = if n=0 then raise TOPSORT "no minimum in list" else if list_all (less x) xs then (x,xs) else split_min (n-1) (xs@[x]); fun tsort [] = [] | tsort xs = let val (x,xs') = split_min (length xs) xs; in x::tsort xs' end; in tsort xs end; fun clique_name clique = (foldr1 (fn (a,b) => a ^ "_" ^ b) (map (unsuffix proc_deco) clique)); fun error_to_warning msg f thy = f thy handle ERROR msg' => (warning (msg' ^ "\n" ^ msg); thy); fun procedures_definition locname procs thy = let val procs' = map (fn (name,a,b,c,d,e,f) => (suffix proc_deco name,a,b,c,d,e,f)) procs; val _ = check_procedures_definition procs' thy; val name_pars = map (fn (name,inpars,outpars,_,_,_,sk) => (name,(inpars,outpars,sk))) procs'; val name_vars = map (fn (name,inpars,outpars,locals,_,_,_) => (name,(inpars,outpars,locals))) procs'; val name_body = map (fn (name,_,_,_,body,_,_) => (name,body)) procs'; val name_pars_specs = map (fn (name,inpars,outpars,_,_,specs,sk) => (name,(inpars,outpars,sk),specs)) procs'; val names = map #1 procs'; val sk = #7 (hd procs'); val thy = thy |> Context.theory_map (set_default_state_kind sk); val (all_callss,cliques,is_recursive,has_body) = let - val ctxt = + val context = Context.Theory thy |> fold (add_parameter_info Morphism.identity (unsuffix proc_deco)) name_pars |> StateSpace.set_silent true fun read_body (_, body) = - Syntax.read_term (Context.proof_of ctxt) body; + Syntax.read_term (Context.proof_of context) body; val bodies = map read_body name_body; fun dcall t = (case try dest_call t of SOME (_,p,_,_,m,_) => SOME (proc_name m p) | _ => NONE); fun in_names x = if member (op =) names x then SOME x else NONE; fun add_edges n = fold (fn x => Graph.add_edge (n, x)); val all_callss = map (map snd o max_subterms_dest dcall) bodies; val callss = map (map_filter in_names) all_callss; val graph = fold (fn n => Graph.new_node (n, ())) names Graph.empty; val graph' = fold2 add_edges names callss graph; fun idx x = find_index (fn y => x=y) names; fun name_ord (a,b) = int_ord (idx a, idx b); val cliques = Graph.strong_conn graph'; val cliques' = map (sort name_ord) cliques; val my_calls = these o AList.lookup (op =) (names ~~ map (distinct (op =)) callss); val my_body = AList.lookup (op =) (names ~~ bodies); fun is_recursive n = exists (fn [_] => false | ns => member (op =) ns n) (Graph.strong_conn graph'); fun has_body n = (case my_body n of SOME (Const (c,_)) => c <> NoBodyN | _ => true) fun clique_less c1 c2 = null (inter (op =) (distinct (op =) (maps my_calls c1)) c2); val cliques'' = topsort clique_less cliques'; in (all_callss,cliques'',is_recursive,has_body) end; (* cliques may only depend on ones to the left, so it is safe to * add the locales from the left to the right. *) fun my_clique n = Library.find_first (fn c => member (op =) c n) cliques; fun lname sfx clique = suffix sfx (clique_name clique); fun cname n = clique_name (the (my_clique n)); fun parameter_info_decl phi = fold (add_parameter_info phi cname) name_pars; fun get_loc sfx clique n = if member (op =) clique n then NONE else SOME (resuffix proc_deco sfx n); fun parent_locales thy sfx clique = let val calls = distinct (op =) (flat (map_filter (AList.lookup (op =) (names ~~ all_callss)) clique)); in map (intern_locale thy) (distinct (op =) (map_filter (get_loc sfx clique) calls)) end; val names_all_callss = names ~~ map (distinct (op =)) all_callss; val get_calls = the o AList.lookup (op =) names_all_callss; fun clique_vars clique = let fun add name (ins,outs,locs) = let val (nins,nouts,nlocs) = the (AList.lookup (op =) name_vars name) in (ins@nins,outs@nouts,locs@nlocs) end; val (is,os,ls) = fold add clique ([],[],[]); in (lname "" clique, (distinct (op =) is, distinct (op =) os, distinct (op =) ls)) end; fun add_signature_locale (cname, name) thy = let val name' = unsuffix proc_deco name; val fixes = [Element.Fixes [(Binding.name name, SOME proc_nameT, NoSyn)]]; (* FIXME: may use HOLogic.typeT as soon as locale type-inference works properly *) val pE = mk_loc_exp [intern_locale thy (suffix parametersN cname)]; val sN = suffix signatureN name'; in thy |> add_locale sN pE fixes |> Proof_Context.theory_of |> (fn thy => add_declaration (intern_locale thy sN) parameter_info_decl thy) end; fun mk_bdy_def read_term name = let val name' = unsuffix proc_deco name; val bdy = read_term (the (AList.lookup (op =) name_body name)); val bdy_defN = suffix body_def_sfx name'; val b = Binding.name bdy_defN; in ((b, NoSyn), ((Thm.def_binding b, []), bdy)) end; fun add_body_locale (name, _) thy = let val name' = unsuffix proc_deco name; val callees = filter_out (fn n => n = name) (get_calls name) val fixes = [Element.Fixes [(Binding.name name, SOME proc_nameT, NoSyn)]]; (* FIXME: may use HOLogic.typeT as soon as locale type-inference works properly *) val pE = mk_loc_exp (map (intern_locale thy) ([lname variablesN (the (my_clique name))]@ the_list locname@ map (resuffix proc_deco signatureN) callees)); fun def lthy = let val read = Syntax.read_term (Context.proof_map (add_active_procs Morphism.identity (the (my_clique name))) (Local_Theory.target_of lthy)) in mk_bdy_def read name end; fun add_decl_and_def lname ctxt = ctxt |> Proof_Context.theory_of |> Named_Target.init [] lname |> Local_Theory.declaration {syntax = false, pervasive = false} parameter_info_decl |> (fn lthy => if has_body name then snd (Local_Theory.define (def lthy) lthy) else lthy) |> Local_Theory.exit |> Proof_Context.theory_of; in thy |> add_locale' (suffix bodyN name') pE fixes |-> add_decl_and_def end; fun mk_def_eq thy read_term name = if has_body name then let (* FIXME: All the read_term stuff is just because type-inference/abbrevs for * new locale elements does not work right now; * We read the term to expand the abbreviations, then we print it again * (without folding the abbreviation) and reread as string *) val name' = unsuffix proc_deco name; val bdy_defN = suffix body_def_sfx name'; val rhs = read_term ("Some " ^ bdy_defN); val nt = read_term name; val Free (gamma,_) = read_term programN; val eq = HOLogic.Trueprop$ HOLogic.mk_eq (Free (gamma,fastype_of nt --> fastype_of rhs)$nt,rhs) val consts = Sign.consts_of thy; val eqs = YXML.string_of_body (Term_XML.Encode.term consts (Consts.dummy_types consts eq)); val assms = Element.Assumes [((Binding.name (suffix bodyP name'), []),[(eqs,[])])] in [assms] end else []; fun add_impl_locales clique thy = let val cliqN = lname cliqueN clique; val cnamesN = lname clique_namesN clique; val multiple_procs = length clique > 1; val add_distinct_procs_namespace = if multiple_procs then StateSpace.namespace_definition cnamesN proc_nameT ([],[]) [] clique else I; val bodyLs = map (suffix bodyN o unsuffix proc_deco) clique; fun pE thy = mk_loc_exp (map (intern_locale thy) (hoare_ctxtL::bodyLs) @ (parent_locales thy implementationN clique) @ (if multiple_procs then [intern_locale thy cnamesN] else [])); fun read_term thy = apply_in_context thy (pE thy) Syntax.read_term; fun elems thy = maps (mk_def_eq thy (read_term thy)) clique; fun add_recursive_info phi name = if is_recursive name then (add_recursive phi name) else I; fun proc_declaration phi = add_active_procs phi clique; - fun recursive_declaration phi ctxt = - ctxt |> fold (add_recursive_info phi) clique; + fun recursive_declaration phi context = + context |> fold (add_recursive_info phi) clique; fun add_impl_locale name thy = let val implN = suffix implementationN (unsuffix proc_deco name); val parentN = intern_locale thy cliqN val parent = mk_loc_exp [parentN]; in thy |> add_locale_cmd implN parent [] |> Proof_Context.theory_of |> (fn thy => Interpretation.global_sublocale parentN (mk_loc_exp [intern_locale thy implN]) [] thy) |> Proof.global_terminal_proof ((Method.Basic (fn ctxt => Method.SIMPLE_METHOD (Locale.intro_locales_tac {strict = true, eager = false} ctxt [])), Position.no_range), NONE) |> Proof_Context.theory_of end; in thy |> add_distinct_procs_namespace |> (fn thy => add_locale_cmd cliqN (pE thy) (elems thy) thy) |> Proof_Context.theory_of |> fold add_impl_locale clique |> (fn thy => add_declaration (intern_locale thy cliqN) proc_declaration thy) |> (fn thy => add_declaration (intern_locale thy cliqN) recursive_declaration thy) end; fun add_spec_locales (name, _, specs) thy = let val name' = unsuffix proc_deco name; val ps = (suffix signatureN name' :: the_list locname); val ps' = hoare_ctxtL :: ps ; val pE = mk_loc_exp (map (intern_locale thy) ps) val pE' = mk_loc_exp (map (intern_locale thy) ps') fun read thy = apply_in_context thy (mk_loc_exp [intern_locale thy (suffix cliqueN (cname name))]) (Syntax.read_prop); fun proc_declaration phi = (*parameter_info_decl phi o already in signature *) add_active_procs phi (the (my_clique name)); fun add_locale'' (thm_name,spec) thy = let val spec' = read thy spec; val elem = Element.Assumes [((Binding.name thm_name, []), [(spec', [])])]; in thy |> add_locale thm_name pE' [elem] |> Proof_Context.theory_of |> (fn thy => add_declaration (intern_locale thy thm_name) proc_declaration thy) |> error_to_warning ("abbreviation: '" ^ thm_name ^ "' not added") (add_abbrev (intern_locale thy (suffix cliqueN (cname name))) Syntax.mode_input thm_name spec) end; in thy |> fold add_locale'' specs end; in thy |> fold (add_variable_statespaces o clique_vars) cliques |> fold (fn c => fold (fn n => add_signature_locale (lname "" c, n)) c) cliques |> fold add_body_locale name_pars |> fold add_impl_locales cliques |> fold add_spec_locales name_pars_specs end; (********************* theory extender interface ********************************) (** package setup **) (* outer syntax *) val var_declP = Parse.name -- (@{keyword "::"} |-- Parse.!!! Parse.embedded); val var_declP' = Parse.name >> (fn n => (n,"")); val localsP = Scan.repeat var_declP; val argP = var_declP; val argP' = var_declP'; val not_eqP = Scan.ahead (Scan.unless @{keyword "="} (Scan.one (K true))) val proc_decl_statespace = (Parse.short_ident --| @{keyword "("}) -- ((Parse.list argP) -- (Scan.optional (@{keyword "|"} |-- Parse.list argP) []) --| @{keyword ")"}) --| not_eqP val proc_decl_record = (Parse.short_ident --| @{keyword "("}) -- ((Parse.list argP') -- (Scan.optional (@{keyword "|"} |-- Parse.list argP') []) --| @{keyword ")"}) --| Scan.option @{keyword "="} val proc_decl = proc_decl_statespace >> pair Function || proc_decl_record >> pair Record; val loc_decl = Scan.optional (@{keyword "where"} |-- localsP --| @{keyword "in"}) [] val proc_body = Parse.embedded (*>> BodyTerm*) fun proc_specs x = (Parse.!!! (Scan.repeat (Parse_Spec.thm_name ":" -- Parse.embedded)) >> map (fn ((thm_name, _), prop) => (Binding.name_of thm_name, prop))) x val par_loc = Scan.option (@{keyword "("} |-- @{keyword "imports"} |-- Parse.name --| @{keyword ")"}); val _ = Outer_Syntax.command @{command_keyword procedures} "define procedures" (par_loc -- (Parse.and_list1 (proc_decl -- loc_decl -- proc_body -- proc_specs)) >> (fn (loc,decls) => let val decls' = map (fn ((((state_kind,(name,(ins,outs))),ls),bdy),specs) => (name,ins,outs,ls,bdy,specs,state_kind)) decls in Toplevel.theory (procedures_definition loc decls') end)); val _ = Outer_Syntax.command @{command_keyword hoarestate} "define state space for hoare logic" (StateSpace.statespace_decl >> (fn ((args,name),(parents,comps)) => Toplevel.theory (StateSpace.define_statespace args name parents (map (apfst (suffix deco)) comps)))); (*************************** Auxiliary Functions for integration of ********************) (*************************** automatic program analysers ********************) fun dest_conjs t = (case HOLogic.dest_conj t of [t1,t2] => dest_conjs t1 @ dest_conjs t2 | ts => ts); fun split_guard (Const (@{const_name Collect},CT)$(Abs (s,T,t))) = let fun mkCollect t = Const (@{const_name Collect},CT)$(Abs (s,T,t)); in map mkCollect (dest_conjs t) end | split_guard t = [t]; fun split_guards gs = let fun norm c f g = map (fn g => c$f$g) (split_guard g); fun norm_guard ((c as Const (@{const_name Pair},_))$f$g) = norm c f g | norm_guard ((c as Const (@{const_name Language.guaranteeStripPair},_))$f$g) = norm c f g | norm_guard t = [t]; in maps norm_guard (HOLogic.dest_list gs) end fun fold_com f t = let (* traverse does not descend into abstractions, like in DynCom, call, etc. *) fun traverse cnt (c as Const (@{const_name Language.com.Skip},_)) = (cnt,f cnt c [] []) | traverse cnt ((c as Const (@{const_name Language.com.Basic},_))$g) = (cnt, f cnt c [g] []) | traverse cnt ((c as Const (@{const_name Language.com.Spec},_))$r) = (cnt, f cnt c [r] []) | traverse cnt ((c as Const (@{const_name Language.com.Seq},_))$c1$c2) = let val (cnt1,v1) = traverse cnt c1; val (cnt2,v2) = traverse cnt1 c2; in (cnt2, f cnt c [] [v1,v2]) end | traverse cnt ((c as Const (@{const_name Language.com.Cond},_))$b$c1$c2) = let val (cnt1,v1) = traverse cnt c1; val (cnt2,v2) = traverse cnt1 c2; in (cnt2, f cnt c [b] [v1,v2]) end | traverse cnt ((c as Const (@{const_name Language.com.While},_))$b$c1) = let val (cnt1,v1) = traverse cnt c1 in (cnt1, f cnt c [b] [v1]) end | traverse cnt ((c as Const (@{const_name Language.com.Call},_))$p) = (cnt, f cnt c [p] []) | traverse cnt ((c as Const (@{const_name Language.com.DynCom},_))$c1) = (cnt, f cnt c [c1] []) | traverse cnt ((c as Const (@{const_name Language.com.Guard},_))$flt$g$c1) = let val (cnt1,v1) = traverse (cnt + length (split_guard g)) c1 in (cnt1, f cnt c [flt,g] [v1]) end | traverse cnt (c as Const (@{const_name Language.com.Throw},_)) = (cnt,f cnt c [] []) | traverse cnt ((c as Const (@{const_name Language.com.Catch},_))$c1$c2) = let val (cnt1,v1) = traverse cnt c1; val (cnt2,v2) = traverse cnt1 c2; in (cnt2, f cnt c [] [v1,v2]) end | traverse cnt ((c as Const (@{const_name Language.guards},_))$gs$c1) = let val (cnt1,v1) = traverse (cnt + length (split_guards gs)) c1; in (cnt1, f cnt c [gs] [v1]) end | traverse cnt ((c as Const (@{const_name Language.block},_))$init$c1$return$c2) = let val (cnt1,v1) = traverse cnt c1 in (cnt1, f cnt c [init,return,c2] [v1]) end | traverse cnt ((c as Const (@{const_name Language.call},_))$init$p$return$c1) = (cnt, f cnt c [init,p,return,c1] []) | traverse cnt ((c as Const (@{const_name Language.whileAnno},_))$b$I$V$c1) = let val (cnt1,v1) = traverse cnt c1 in (cnt1, f cnt c [b,I,V] [v1]) end | traverse cnt ((c as Const (@{const_name Language.whileAnnoG},_))$gs$b$I$V$c1) = let val (cnt1,v1) = traverse (cnt + length (split_guards gs)) c1 in (cnt1, f cnt c [gs,b,I,V] [v1]) end | traverse _ t = raise TERM ("fold_com: unknown command",[t]); in snd (traverse 0 t) end; (*************************** Tactics ****************************************************) (*** Aux. tactics ***) fun cond_rename_bvars cond name thm = let fun rename (tm as (Abs (x, T, t))) = if cond tm then Abs (name x, T, rename t) else Abs (x, T, rename t) | rename (t $ u) = rename t $ rename u | rename a = a; in Thm.renamed_prop (rename (Thm.prop_of thm)) thm end; val rename_bvars = cond_rename_bvars (K true); fun trace_tac ctxt str st = (if Config.get ctxt hoare_trace then tracing str else (); all_tac st); fun error_tac str st = (error str;no_tac st); fun rename_goal ctxt name = EVERY' [K (trace_tac ctxt "rename_goal -- START"), SELECT_GOAL (PRIMITIVE (rename_bvars name)), K (trace_tac ctxt "rename_goal -- STOP")]; (* splits applications of tupled arguments to a schematic Variables, e.g. * ALL a b. ?P (a,b) --> ?Q (a,b) gets * ALL a b. ?P a b --> ?Q a b * only tuples nested to the right are splitted. *) fun split_pair_apps ctxt thm = let val t = Thm.prop_of thm; fun mk_subst subst (Abs (x,T,t)) = mk_subst subst t | mk_subst subst (t as (t1$t2)) = (case strip_comb t of (var as Var (v,vT),args) => (if not (AList.defined (op =) subst var) then let val len = length args; val (argTs,bdyT) = strip_type vT; val (z, _) = Name.variant "z" (fold Term.declare_term_frees args Name.context); val frees = map (apfst (fn i => z^string_of_int i)) (0 upto (len - 1) ~~ argTs); fun splitT (Type (@{type_name Product_Type.prod}, [T1, T2])) = T1::splitT T2 | splitT T = [T]; fun pair_depth (Const (@{const_name Pair},aT)$t1$t2) = 1 + (pair_depth t2) | pair_depth _ = 0; fun mk_sel max free i = let val snds = funpow i HOLogic.mk_snd (Free free) in if i=max then snds else HOLogic.mk_fst snds end; fun split (free,arg) = let val depth = (pair_depth arg); in if depth = 0 then [Free free] else map (mk_sel depth free) (0 upto depth) end; val args' = maps split (frees ~~ args); val argTs' = maps splitT argTs; val inst = fold_rev absfree frees (list_comb (Var (v,argTs' ---> bdyT), args')) in subst@[(var,inst)] end else subst) | _ => mk_subst (mk_subst subst t1) t2) | mk_subst subst t = subst; val subst = map (fn (v,t) => (dest_Var v, Thm.cterm_of ctxt t)) (mk_subst [] t); in full_simplify (put_simpset HOL_basic_ss ctxt addsimps [@{thm fst_conv}, @{thm snd_conv}]) (Drule.instantiate_normalize (TVars.empty, Vars.make subst) thm) end; (* Generates split theorems, for !!,!,? quantifiers and for UN, e.g. * ALL x. P x = ALL a b. P a b *) fun mk_split_thms ctxt (vars as _::_) = let val thy = Proof_Context.theory_of ctxt; val names = map fst vars; val types = map snd vars; val free_vars = map Free vars; val pT = foldr1 HOLogic.mk_prodT types; val x = (singleton (Name.variant_list names) "x", pT); val xp = foldr1 HOLogic.mk_prod free_vars; val tfree_names = fold Term.add_tfree_names free_vars []; val zeta = TFree (singleton (Name.variant_list tfree_names) "z", Sign.defaultS thy); val split_meta_prop = let val P = Free (singleton (Name.variant_list names) "P", pT --> Term.propT) in Logic.mk_equals (Logic.all (Free x) (P $ Free x), fold_rev Logic.all free_vars (P $ xp)) end; val P = Free (singleton (Name.variant_list names) "P", pT --> HOLogic.boolT); val split_object_prop = let fun ALL vs t = Library.foldr (fn ((v,T),t) => HOLogic.mk_all (v,T,t)) (vs,t) in (ALL [x] (P $ Free x)) === (ALL vars (P $ xp)) end; val split_ex_prop = let fun EX vs t = Library.foldr (fn ((v,T),t) => HOLogic.mk_exists (v,T,t)) (vs,t) in (EX [x] (P $ Free x)) === (EX vars (P $ xp)) end; val split_UN_prop = let val P = Free (singleton (Name.variant_list names) "P", pT --> HOLogic.mk_setT zeta); fun UN vs t = Library.foldr mk_UN (vs, t) in (UN [x] (P $ Free x)) === (UN vars (P $ xp)) end; fun prove_simp simps prop = let val ([prop'], _) = Variable.importT_terms [prop] ctxt (* FIXME continue context!? *) in Goal.prove_global thy [] [] prop' - (fn {context = ctxt, ...} => - ALLGOALS (simp_tac (put_simpset HOL_basic_ss ctxt addsimps simps))) + (fn {context = goal_ctxt, ...} => + ALLGOALS (simp_tac (put_simpset HOL_basic_ss goal_ctxt addsimps simps))) end; val split_meta = prove_simp [@{thm split_paired_all}] split_meta_prop; val split_object = prove_simp [@{thm split_paired_All}] split_object_prop; val split_ex = prove_simp [@{thm split_paired_Ex}] split_ex_prop; val split_UN = prove_simp [@{thm Hoare.split_paired_UN}] split_UN_prop; in [split_meta,split_object,split_ex,split_UN] end | mk_split_thms _ _ = raise Match; fun rename_aux_var name rule = let fun is_aux_var (Abs ("Z",TVar(_,_),_)) = true | is_aux_var _ = false; in cond_rename_bvars is_aux_var (K name) rule end; (* adapts single auxiliary variable in a rule to potentialy multiple auxiliary * variables in actual specification, e.g. if vars are a b, * split_app=false: ALL Z. ?P Z gets to ALL a b. ?P (a,b) * split_app=true: ALL Z. ?P Z gets to ALL a b. ?P a b * If only one auxiliary variable is given, the variables are just renamed, * If no auxiliary is given, unit is inserted for Z: * ALL Z. ?P Z gets P () *) fun adapt_aux_var ctxt split_app (vars as (_::_::_)) tvar_rules = let val thy = Proof_Context.theory_of ctxt; val max_idx = fold Integer.max (map (Thm.maxidx_of o snd) tvar_rules) 0; val types = map (fn i => TVar (("z",i),Sign.defaultS thy)) (max_idx + 1 upto (max_idx + length vars)); fun tvar n = (n, Sign.defaultS thy); val pT = Thm.ctyp_of ctxt (foldr1 HOLogic.mk_prodT types); val rules' = map (fn (z,r) => (Drule.instantiate_normalize (TVars.make [(tvar z,pT)], Vars.empty) r)) tvar_rules; val splits = mk_split_thms ctxt (vars ~~ types); val rules'' = map (full_simplify (put_simpset HOL_basic_ss ctxt addsimps splits)) rules'; in if split_app then (map (split_pair_apps ctxt) rules'') else rules'' end | adapt_aux_var _ _ ([name]) tvar_rules = map (rename_aux_var name o snd) tvar_rules | adapt_aux_var ctxt _ ([]) tvar_rules = let val thy = Proof_Context.theory_of ctxt; fun tvar n = (n, Sign.defaultS thy); val uT = Thm.ctyp_of ctxt HOLogic.unitT; val rules' = map (fn (z,r) => (Drule.instantiate_normalize (TVars.make [(tvar z,uT)], Vars.empty) r)) tvar_rules; val splits = [@{thm Hoare.unit_meta},@{thm Hoare.unit_object},@{thm Hoare.unit_ex},@{thm Hoare.unit_UN}]; val rules'' = map (full_simplify (put_simpset HOL_basic_ss ctxt addsimps splits)) rules'; in rules'' end (* Generates a rule for recursion for n procedures out of general recursion rule *) fun gen_call_rec_rule ctxt specs_name n rule = let val thy = Proof_Context.theory_of ctxt; val maxidx = Thm.maxidx_of rule; val vars = Term.add_vars (Thm.prop_of rule) []; fun get_type n = the (AList.lookup (op =) vars (n, 0)); val (Type (_, [Type (_, [assT, Type (_, [pT,_])])])) = get_type specs_name; val zT = TVar (("z",maxidx+1),Sign.defaultS thy) fun mk_var i n T = Var ((n ^ string_of_int i,0),T); val quadT = HOLogic.mk_prodT (assT, HOLogic.mk_prodT (pT, HOLogic.mk_prodT (assT,assT))); val quadT_set = HOLogic.mk_setT quadT; fun mk_spec i = let val quadruple = HOLogic.mk_ptuple (HOLogic.flat_tupleT_paths quadT) quadT [mk_var i "P" (zT --> assT)$Bound 0, mk_var i "p" pT, mk_var i "Q" (zT --> assT)$Bound 0, mk_var i "A" (zT --> assT)$Bound 0]; val single = HOLogic.mk_set quadT [quadruple]; in mk_UN' zT quadT_set (Abs ("Z", zT, single)) end; val Specs = foldr1 (mk_Un quadT_set) (map mk_spec (1 upto n)); val rule' = Thm.instantiate' [] [SOME (Thm.cterm_of ctxt Specs)] rule |> full_simplify (put_simpset (simpset_of @{theory_context Main}) ctxt addsimps [@{thm Hoare.conjE_simp},@{thm Hoare.in_Specs_simp},@{thm Hoare.in_set_Un_simp},@{thm split_all_conj}, @{thm image_Un},@{thm image_Un_single_simp}] ) |> rename_bvars (fn s => if member (op =) ["s","\"] s then s else "Z") in rule' end; fun gen_proc_rec ctxt mode n = gen_call_rec_rule ctxt "Specs" n (ProcRecSpecs mode); (*** verification condition generator ***) (* simplifications on "Collect" sets, like {s. P s} Int {s. Q s} = {s. P s & Q s} *) fun merge_assertion_simp_tac ctxt thms = simp_tac (put_simpset HOL_basic_ss ctxt addsimps ([@{thm Hoare.CollectInt_iff},@{thm HOL.conj_assoc},@{thm Hoare.Compl_Collect},singleton_conv_sym, @{thm Set.Int_empty_right},@{thm Set.Int_empty_left},@{thm Un_empty_right},@{thm Un_empty_left}]@thms)) ; (* The following probably shouldn't live here, but refactoring so that Hoare could depend on recursive_records does not look feasible. The upshot is that there's a duplicate foldcong_ss set here. *) structure FoldCongData = Theory_Data ( type T = simpset; val empty = HOL_basic_ss; val copy = I; val extend = I; val merge = merge_ss; ) val get_foldcong_ss = FoldCongData.get fun add_foldcongs congs thy = FoldCongData.map (fn ss => Proof_Context.init_global thy |> put_simpset ss |> fold Simplifier.add_cong congs |> simpset_of) thy fun add_foldcongsimps simps thy = FoldCongData.map (fn ss => Proof_Context.init_global thy |> put_simpset ss |> (fn ctxt => ctxt addsimps simps) |> simpset_of) thy (* propagates state into "Collect" sets and simplifies selections updates like: * s:{s. P s} = P s *) fun in_assertion_simp_tac ctxt state_kind thms i = let val vcg_simps = #vcg_simps (get_data ctxt); val fold_simps = get_foldcong_ss (Proof_Context.theory_of ctxt) in EVERY [simp_tac (put_simpset HOL_basic_ss ctxt addsimps ([mem_Collect_eq,@{thm Set.Un_iff},@{thm Set.Int_iff}, @{thm Set.empty_subsetI}, @{thm Set.empty_iff}, UNIV_I, @{thm Hoare.Collect_False}]@thms@K_convs@vcg_simps) addsimprocs (state_simprocs state_kind) |> fold Simplifier.add_cong K_congs) i THEN_MAYBE (simp_tac (put_simpset fold_simps ctxt addsimprocs [state_upd_simproc state_kind]) i) ] end; fun assertion_simp_tac ctxt state_kind thms i = merge_assertion_simp_tac ctxt [] i THEN_MAYBE in_assertion_simp_tac ctxt state_kind thms i (* simplify equality test on strings (and datatype-constructors) and propagate result*) fun string_eq_simp_tac ctxt = simp_tac (put_simpset HOL_basic_ss ctxt addsimps @{thms list.inject list.distinct char.inject cong_exp_iff_simps simp_thms}); fun assertion_string_eq_simp_tac ctxt state_kind thms i = assertion_simp_tac ctxt state_kind thms i THEN_MAYBE string_eq_simp_tac ctxt i; fun before_set2pred_simp_tac ctxt = (simp_tac (put_simpset HOL_basic_ss ctxt addsimps [singleton_conv_sym, @{thm Hoare.CollectInt_iff}, @{thm Hoare.Compl_Collect}])); (*****************************************************************************) (** set2pred transforms sets inclusion into predicates implication, **) (** maintaining the original variable names. **) (** Ex. "{x. x=0} <= {x. x <= 1}" -set2pred-> "x=0 --> x <= 1" **) (** Subgoals containing intersections (A Int B) or complement sets (-A) **) (** are first simplified by "before_set2pred_simp_tac", that returns only **) (** subgoals of the form "{x. P x} <= {x. Q x}", which are easily **) (** transformed. **) (** This transformation may solve very easy subgoals due to a ligth **) (** simplification done by full_simp_tac **) (*****************************************************************************) fun set2pred_tac ctxt i thm = ((before_set2pred_simp_tac ctxt i) THEN_MAYBE (EVERY [trace_tac ctxt "set2pred", resolve_tac ctxt [subsetI] i, resolve_tac ctxt [CollectI] i, dresolve_tac ctxt [CollectD] i, full_simp_tac (put_simpset HOL_basic_ss ctxt) i ])) thm (*****************************************************************************) (** BasicSimpTac is called to simplify all verification conditions. It does **) (** a light simplification by applying "mem_Collect_eq" **) (** then it tries to solve subgoals of the form "A <= A" and then if **) (** set2pred is true it **) (** transforms any other into predicates, applying then **) (** the tactic chosen by the user, which may solve the subgoal completely **) (** (MaxSimpTac). **) (*****************************************************************************) fun MaxSimpTac ctxt tac i = TRY (FIRST[resolve_tac ctxt [subset_refl] i, set2pred_tac ctxt i THEN_MAYBE tac i, trace_tac ctxt "final_tac failed" ]); fun BasicSimpTac ctxt state_kind set2pred thms tac i = EVERY [(trace_tac ctxt "BasicSimpTac -- START --"), assertion_simp_tac ctxt state_kind thms i THEN_MAYBE (if set2pred then MaxSimpTac ctxt tac i else TRY (resolve_tac ctxt [subset_refl] i)), (trace_tac ctxt "BasicSimpTac -- STOP --")]; (* EVERY [(trace_tac ctxt "BasicSimpTac -- START --"), simp_tac (HOL_basic_ss addsimps [mem_Collect_eq,@{thm Hoare.CollectInt_iff}, @{thm Set.empty_subsetI}, @{thm Set.empty_iff}, UNIV_I] addsimprocs [state_simproc sk]) i THEN_MAYBE simp_tac (HOL_basic_ss addsimprocs [state_upd_simproc sk]) i THEN_MAYBE (if set2pred then MaxSimpTac ctxt tac i else TRY (rtac subset_refl i)), (trace_tac ctxt "BasicSimpTac -- STOP --")]; *) (* fun simp_state_eq_tac Record state_space = full_simp_tac (HOL_basic_ss addsimprocs (state_simprocs Record)) THEN_MAYBE' full_simp_tac (HOL_basic_ss addsimprocs [state_upd_simproc Record]) THEN_MAYBE' (state_split_simp_tac [] state_space) | simp_state_eq_tac StateFun state_space = *) fun post_conforms_tac ctxt state_kind i = EVERY [REPEAT1 (resolve_tac ctxt [allI,impI] i), ((fn i => TRY (resolve_tac ctxt [conjI] i)) THEN_ALL_NEW (fn i => (REPEAT (resolve_tac ctxt [allI,impI] i)) THEN (full_simp_tac (put_simpset HOL_basic_ss ctxt addsimps [mem_Collect_eq,@{thm Set.singleton_iff},@{thm Set.empty_iff},UNIV_I] addsimprocs (state_simprocs state_kind)) i))) i]; fun dest_hoare_raw (Const(@{const_name HoarePartialDef.hoarep},_)$G$T$F$P$C$Q$A) = (P,C,Q,A,Partial,G,T,F) | dest_hoare_raw (Const(@{const_name HoareTotalDef.hoaret},_)$G$T$F$P$C$Q$A) = (P,C,Q,A,Total,G,T,F) | dest_hoare_raw t = raise TERM ("Hoare.dest_hoare_raw: unexpected term",[t]) fun mk_hoare_abs Ts (P,C,Q,A,mode,G,T,F) = let val hoareT = map (curry fastype_of1 Ts) [G,T,F,P,C,Q,A] ---> HOLogic.boolT; val hoareC = (case mode of Partial => Const (@{const_name HoarePartialDef.hoarep},hoareT) | Total => Const (@{const_name HoareTotalDef.hoaret},hoareT)); in hoareC$G$T$F$P$C$Q$A end; val is_hoare = can dest_hoare_raw fun dest_hoare t = let val triple = (strip_qnt_body @{const_name "All"} o HOLogic.dest_Trueprop o strip_qnt_body @{const_name Pure.all}) t; in dest_hoare_raw triple end; fun get_aux_tvar rule = let fun aux_hoare (Abs ("Z",TVar (z,_),t)) = if is_hoare (strip_qnt_body @{const_name All} t) then SOME z else NONE | aux_hoare _ = NONE; in (case first_subterm_dest (aux_hoare) (Thm.prop_of rule) of SOME (_,z) => (z,rule) | NONE => raise TERM ("get_aux_tvar: No auxiliary variable of hoare-rule found", [Thm.prop_of rule])) end; fun strip_vars t = let val bdy = (HOLogic.dest_Trueprop o Logic.strip_assums_concl) t; in strip_qnt_vars @{const_name Pure.all} t @ strip_qnt_vars @{const_name All} bdy end; local (* ex_simps are necessary in case of multiple logical variables. The state will usually be the first variable. EX s a b. s=s' ... . We have to transport EX s to s=s' to perform the substitution *) val conseq1_ss_base = simpset_of (put_simpset HOL_basic_ss @{context} addsimps ([mem_Collect_eq,@{thm Set.singleton_iff},@{thm Set.Int_iff}, @{thm Set.empty_iff},UNIV_I, @{thm HOL.conj_assoc}, @{thm disj_assoc}] @ @{thms Hoare.all_imp_eq_triv} @K_convs @ @{thms simp_thms} @ @{thms ex_simps} @ @{thms all_simps}) delsimps [@{thm Hoare.all_imp_to_ex}] |> fold Simplifier.add_cong K_congs) val conseq1_ss_record = simpset_of (put_simpset conseq1_ss_base @{context} addsimprocs (state_simprocs Record)); val conseq1_ss_fun = simpset_of (put_simpset conseq1_ss_base @{context} addsimprocs (state_simprocs Function)); fun conseq1_ss Record = conseq1_ss_record | conseq1_ss Function = conseq1_ss_fun; val conseq2_ss_base = simpset_of (put_simpset HOL_basic_ss @{context} addsimps (@{thms Hoare.all_imp_eq_triv} @ @{thms simp_thms} @ @{thms ex_simps} @ @{thms all_simps}) delsimps [@{thm Hoare.all_imp_to_ex}] |> Simplifier.add_cong @{thm imp_cong}); val conseq2_ss_record = simpset_of (put_simpset conseq2_ss_base @{context} addsimprocs [state_upd_simproc Record, state_ex_sel_eq_simproc Record]); val conseq2_ss_fun = simpset_of (put_simpset conseq2_ss_base @{context} addsimprocs [state_upd_simproc Function, state_ex_sel_eq_simproc Function]); fun conseq2_ss Record = conseq2_ss_record | conseq2_ss Function = conseq2_ss_fun; in fun raw_conseq_simp_tac ctxt state_kind thms i = let val ctxt' = Config.put simp_depth_limit 0 ctxt; in simp_tac (put_simpset (conseq1_ss state_kind) ctxt' addsimps thms) i THEN_MAYBE simp_tac (put_simpset (conseq2_ss state_kind) ctxt') i end end val conseq_simp_tac = raw_conseq_simp_tac; (* Generates the hoare-quadruples that can be derived out of the hoare-context T *) fun gen_context_thms ctxt mode params G T F = let val Type (_,[comT]) = range_type (fastype_of G); fun destQuadruple (Const (@{const_name Set.insert},_) $ PpQA $ Const (@{const_name Orderings.bot}, _)) = PpQA | destQuadruple t = raise Match; fun mkCallQuadruple (Const (@{const_name Pair}, _) $ P $ (Const (@{const_name Pair}, _) $ p $ (Const(@{const_name Pair}, _) $ Q $ A))) = let val Call_p = Const (@{const_name Language.com.Call}, fastype_of p --> comT) $ p; in (P, Call_p, Q, A) end; fun mkHoare mode G T F (vars,PpQA) = let val hoare = (case mode of Partial => @{const_name HoarePartialDef.hoarep} | Total => @{const_name HoareTotalDef.hoaret}); (* FIXME: Use future Proof_Context.rename_vars or make closed term and remove by hand *) (* fun free_params ps t = foldr (fn ((x,xT),t) => snd (variant_abs (x,xT,t))) (ps,t); val PpQA' = mkCallQuadruple (strip_qnt_body @{const_name Pure.all} (free_params params (Term.list_all (vars,PpQA)))); *) val params' = (Variable.variant_frees ctxt [PpQA] params); val bnds = map Bound (0 upto (length vars - 1)); fun free_params_vars t = subst_bounds (bnds @ rev (map Free params' ), t) fun free_params t = subst_bounds (rev (map Free params' ), t) val (P',p',Q',A') = mkCallQuadruple (free_params_vars PpQA); val T' = free_params T; val G' = free_params G; val F' = free_params F; val bdy = mk_hoare_abs [] (P',p',Q',A',mode,G',T',F'); in (HOLogic.mk_Trueprop (HOLogic.list_all (vars,bdy)), map fst params') end; fun hoare_context_specs mode G T F = let fun mk t = try (mkHoare mode G T F o apsnd destQuadruple o dest_UN) t; in map_filter mk (dest_Un T) end; fun mk_prove mode (prop,params) = let val vars = map fst (strip_qnt_vars @{const_name All} (HOLogic.dest_Trueprop (Logic.strip_assums_concl prop))); - val [asmUN'] = adapt_aux_var ctxt true vars [get_aux_tvar (AsmUN mode)] in Goal.prove ctxt params [] prop - (fn thms => - EVERY[trace_tac ctxt "extracting specifications from hoare context", - resolve_tac ctxt [asmUN'] 1, - DEPTH_SOLVE_1 (resolve_tac ctxt [subset_refl,refl] 1 ORELSE - ((resolve_tac ctxt [@{thm Hoare.subset_unI1}] 1 APPEND resolve_tac ctxt [@{thm Hoare.subset_unI2}] 1) + (fn {context = ctxt', ...} => + EVERY[trace_tac ctxt' "extracting specifications from hoare context", + resolve_tac ctxt' (adapt_aux_var ctxt' true vars [get_aux_tvar (AsmUN mode)]) 1, + DEPTH_SOLVE_1 (resolve_tac ctxt' [subset_refl,refl] 1 ORELSE + ((resolve_tac ctxt' [@{thm Hoare.subset_unI1}] 1 APPEND resolve_tac ctxt' [@{thm Hoare.subset_unI2}] 1) ORELSE - (resolve_tac ctxt [@{thm Hoare.subset_singleton_insert1}] 1 - APPEND resolve_tac ctxt [@{thm Hoare.subset_singleton_insert2}] 1))) + (resolve_tac ctxt' [@{thm Hoare.subset_singleton_insert1}] 1 + APPEND resolve_tac ctxt' [@{thm Hoare.subset_singleton_insert2}] 1))) ORELSE error_tac ("vcg: cannot extract specifications from context") ]) end; val specs = hoare_context_specs mode G T F; in map (mk_prove mode) specs end; fun is_modifies_clause t = exists_subterm (fn (Const (@{const_name Hoare.meq},_)) => true| _ => false) (#3 (dest_hoare (Logic.strip_assums_concl t))) handle (TERM _) => false; val is_spec_clause = not o is_modifies_clause; (* e.g: Intg => the_Intg lift Intg => lift the_Intg map Ingt => map the_Intg Hp o lift Intg => lift the_Intg o the_Hp *) fun swap_constr_destr f (t as (Const (@{const_name Fun.id},_))) = t | swap_constr_destr f (t as (Const (c,Type ("fun",[T,valT])))) = (Const (f c, Type ("fun",[valT,T])) handle Empty => raise TERM ("Hoare.swap_constr_destr",[t])) | swap_constr_destr f (Const ("StateFun.map_fun",Type ("fun", (* FIXME unknown "StateFun.map_fun" !? *) [Type ("fun",[T,valT]), Type ("fun",[Type ("fun",[xT,T']), Type ("fun",[xT',valT'])])]))$g) = Const ("StateFun.map_fun",Type("fun",[Type ("fun",[valT,T]), Type ("fun",[Type ("fun",[xT,valT']), Type ("fun",[xT',T'])])]))$ swap_constr_destr f g | swap_constr_destr f (Const (@{const_name Fun.comp},Type ("fun", [Type ("fun",[bT',cT]), Type ("fun",[Type ("fun",[aT ,bT]), Type ("fun",[aT',cT'])])]))$h$g) = let val h'=swap_constr_destr f h; val g'=swap_constr_destr f g; in Const (@{const_name Fun.comp},Type ("fun", [Type ("fun",[bT,aT]), Type ("fun",[Type ("fun",[cT,bT']), Type ("fun",[cT',aT'])])]))$g'$h' end | swap_constr_destr f (Const (@{const_name List.map},Type ("fun", [Type ("fun",[aT,bT]), Type ("fun",[asT,bsT])]))$g) = (Const (@{const_name List.map},Type ("fun", [Type ("fun",[bT,aT]), Type ("fun",[bsT,asT])]))$swap_constr_destr f g) | swap_constr_destr f t = raise TERM ("Hoare.swap_constr_destr",[t]); (* FIXME: unused? *) val destr_to_constr = let fun convert c = let val (path,base) = split_last (Long_Name.explode c); in Long_Name.implode (path @ ["val",unprefix "the_" base]) end; in swap_constr_destr convert end; fun gen_call_tac cont_tac mode cmode state_kind state_space ctxt asms spec_sfx pname return has_args _ = let val thy = Proof_Context.theory_of ctxt; val pname' = unsuffix proc_deco pname; val spec = (case AList.lookup (op =) asms pname of SOME s => SOME s | NONE => try (Proof_Context.get_thm ctxt) (suffix spec_sfx pname')); fun auxvars_for p t = (case first_subterm_dest (try dest_call) t of SOME (vars,(_,p',_,_,m,_)) => (if m=Static andalso p=(dest_string' p') then SOME vars else NONE) | NONE => NONE); fun get_auxvars_for p t = (case (map_filter ((auxvars_for p) o snd) (max_subterms_dest tap_UN t)) of (vars::_) => vars | _ => []); - fun spec_tac augment_rule augment_emptyFaults _ spec i = + fun spec_tac ctxt' augment_rule augment_emptyFaults _ spec i = let val spec' = augment_emptyFaults OF [spec] handle THM _ => spec; in - EVERY [resolve_tac ctxt [augment_rule] i, - resolve_tac ctxt [spec'] (i+1), - TRY (resolve_tac ctxt [subset_refl, @{thm Set.empty_subsetI}, @{thm Set.Un_upper1}, @{thm Set.Un_upper2}] i)] + EVERY [resolve_tac ctxt' [augment_rule] i, + resolve_tac ctxt' [spec'] (i+1), + TRY (resolve_tac ctxt' [subset_refl, @{thm Set.empty_subsetI}, @{thm Set.Un_upper1}, @{thm Set.Un_upper2}] i)] end; fun check_spec name P thm = (case try dest_hoare (Thm.concl_of thm) of SOME spc => (case try dest_call (#2 spc) of SOME (_,p,_,_,m,_) => if proc_name m p = name andalso P (Thm.concl_of thm) then SOME (#5 spc,thm) else NONE | _ => NONE) | _ => NONE) fun find_dyn_specs name P thms = map_filter (check_spec name P) thms; fun get_spec name P thms = case find_dyn_specs name P thms of (spec_mode,spec)::_ => SOME (spec_mode,spec) | _ => NONE; - fun solve_spec augment_rule _ augment_emptyFaults mode _ (SOME spec_mode) (SOME spec) i= + fun solve_spec ctxt' augment_rule _ augment_emptyFaults mode _ (SOME spec_mode) (SOME spec) i= if mode=Partial andalso spec_mode=Total - then resolve_tac ctxt [@{thm HoareTotal.hoaret_to_hoarep'}] i THEN spec_tac augment_rule augment_emptyFaults mode spec i - else if mode=spec_mode then spec_tac augment_rule augment_emptyFaults mode spec i + then resolve_tac ctxt' [@{thm HoareTotal.hoaret_to_hoarep'}] i THEN spec_tac ctxt' augment_rule augment_emptyFaults mode spec i + else if mode=spec_mode then spec_tac ctxt' augment_rule augment_emptyFaults mode spec i else error("vcg: cannot use a partial correctness specification of " ^ pname' ^ " for a total correctness proof") - | solve_spec _ asmUN_rule _ mode Static _ _ i =(* try to infer spec out of context *) - EVERY[trace_tac ctxt "inferring specification from hoare context1", - resolve_tac ctxt [asmUN_rule] i, - DEPTH_SOLVE_1 (resolve_tac ctxt [subset_refl,refl] i ORELSE - ((resolve_tac ctxt [@{thm Hoare.subset_unI1}] i APPEND resolve_tac ctxt [@{thm Hoare.subset_unI2}] i) + | solve_spec ctxt' _ asmUN_rule _ mode Static _ _ i =(* try to infer spec out of context *) + EVERY[trace_tac ctxt' "inferring specification from hoare context1", + resolve_tac ctxt' [asmUN_rule] i, + DEPTH_SOLVE_1 (resolve_tac ctxt' [subset_refl,refl] i ORELSE + ((resolve_tac ctxt' [@{thm Hoare.subset_unI1}] i APPEND resolve_tac ctxt' [@{thm Hoare.subset_unI2}] i) ORELSE - (resolve_tac ctxt [@{thm Hoare.subset_singleton_insert1}] i - APPEND resolve_tac ctxt [@{thm Hoare.subset_singleton_insert2}] i))) + (resolve_tac ctxt' [@{thm Hoare.subset_singleton_insert1}] i + APPEND resolve_tac ctxt' [@{thm Hoare.subset_singleton_insert2}] i))) ORELSE error_tac ("vcg: cannot infer specification of " ^ pname' ^ " from context") (* if tactic for DEPTH_SOLVE_1 would create new subgoals, use SELECT_GOAL and DEPTH_SOLVE *) ] - | solve_spec augment_rule asmUN_rule augment_emptyFaults mode Parameter _ _ i = + | solve_spec ctxt' augment_rule asmUN_rule augment_emptyFaults mode Parameter _ _ i = (* try to infer spec out of assumptions *) let - fun tac thms = - (case (find_dyn_specs pname is_spec_clause thms) of + fun tac ({context = ctxt'', prems, ...}: Subgoal.focus) = + (case (find_dyn_specs pname is_spec_clause prems) of (spec_mode,spec)::_ - => solve_spec augment_rule asmUN_rule augment_emptyFaults mode Parameter + => solve_spec ctxt'' augment_rule asmUN_rule augment_emptyFaults mode Parameter (SOME spec_mode) (SOME spec) 1 | _ => all_tac) - in Subgoal.FOCUS (tac o #prems) ctxt i end + in Subgoal.FOCUS tac ctxt' i end val strip_spec_vars = strip_qnt_vars @{const_name All} o HOLogic.dest_Trueprop; - fun apply_call_tac ctxt pname mode cmode spec_mode spec_goal is_abr + fun apply_call_tac ctxt' pname mode cmode spec_mode spec_goal is_abr spec (subgoal,i) = let val spec_vars = map fst (case spec of SOME sp => (strip_spec_vars (Thm.concl_of sp)) | NONE => (case try (dest_hoare) subgoal of SOME (_,_,_,_,_,_,Theta,_) => get_auxvars_for pname Theta | _ => [])); fun get_call_rule Static mode is_abr = if is_abr then Proc mode else ProcNoAbr mode | get_call_rule Parameter mode is_abr = if is_abr then DynProcProcPar mode else DynProcProcParNoAbr mode; val [call_rule,augment_ctxt_rule,asmUN_rule, augment_emptyFaults] = - adapt_aux_var ctxt true spec_vars + adapt_aux_var ctxt' true spec_vars (map get_aux_tvar [get_call_rule cmode mode is_abr, AugmentContext mode, AsmUN mode, AugmentEmptyFaults mode]); - in EVERY [resolve_tac ctxt [call_rule] i, - trace_tac ctxt "call_tac -- basic_tac -- solving spec", - solve_spec augment_ctxt_rule asmUN_rule augment_emptyFaults + in EVERY [resolve_tac ctxt' [call_rule] i, + trace_tac ctxt' "call_tac -- basic_tac -- solving spec", + solve_spec ctxt' augment_ctxt_rule asmUN_rule augment_emptyFaults mode cmode spec_mode spec spec_goal] end; - fun basic_tac spec i = + fun basic_tac ctxt' spec i = let val msg ="Theorem " ^pname'^spec_sfx ^ " is no proper specification for procedure " ^pname'^ "; trying to infer specification from hoare context"; fun spec' s mode abr = if is_modifies_clause (Thm.concl_of s) then if abr then (TrivPost mode) OF [s] else (TrivPostNoAbr mode) OF [s] else s; val (is_abr,spec_mode,spec,spec_has_args) = (* is there a proper specification fact? *) case spec of NONE => (true,NONE,NONE,false) | SOME s => case try dest_hoare (Thm.concl_of s) of NONE => (warning msg;(true,NONE,NONE,false)) | SOME (_,c,Q,spec_abr,spec_mode,_,_,_) => case try dest_call c of NONE => (warning msg;(true,NONE,NONE,false)) | SOME (_,p,_,_,m,spec_has_args) => if proc_name m p = pname then if (mode=Total andalso spec_mode=Partial) then (warning msg;(true,NONE,NONE,false)) else if is_empty_set spec_abr then (false,SOME spec_mode, SOME (spec' s spec_mode false),spec_has_args) else (true,SOME spec_mode, SOME (spec' s spec_mode true),spec_has_args) else (warning msg;(true,NONE,NONE,false)); val () = if spec_has_args then error "procedure call in specification must be parameterless!" else (); val spec_goal = i+2; in - EVERY[trace_tac ctxt "call_tac -- basic_tac -- START --", + EVERY[trace_tac ctxt' "call_tac -- basic_tac -- START --", SUBGOAL - (apply_call_tac ctxt pname mode cmode spec_mode spec_goal is_abr spec) i, - resolve_tac ctxt [allI] (i+1), - resolve_tac ctxt [allI] (i+1), - cont_tac (i+1), - trace_tac ctxt "call_tac -- basic_tac -- simplify", - conseq_simp_tac ctxt state_kind [@{thm StateSpace.upd_globals_def}] i, - trace_tac ctxt "call_tac -- basic_tac -- STOP --"] + (apply_call_tac ctxt' pname mode cmode spec_mode spec_goal is_abr spec) i, + resolve_tac ctxt' [allI] (i+1), + resolve_tac ctxt' [allI] (i+1), + cont_tac ctxt' (i+1), + trace_tac ctxt' "call_tac -- basic_tac -- simplify", + conseq_simp_tac ctxt' state_kind [@{thm StateSpace.upd_globals_def}] i, + trace_tac ctxt' "call_tac -- basic_tac -- STOP --"] end; fun get_modifies (Const (@{const_name Collect},_) $ Abs (_,_,m)) = m | get_modifies t = raise TERM ("gen_call_tac.get_modifies",[t]); fun get_updates (Abs (_,_,t)) = get_updates t | get_updates (Const (@{const_name Hoare.mex},_) $ t) = get_updates t | get_updates (Const (@{const_name Hoare.meq},T) $ _ $ upd) = (T,upd) | get_updates t = raise TERM ("gen_call_tac.get_updates",[t]); (* return has the form: %s t. s(|globals:=globals t,...|) * should be translated to %s t. s(|globals := globals s(|m := m (globals t),...|),...|) * for all m in the modifies list. *) fun mk_subst gT meqT = fst (Sign.typ_unify thy (gT,domain_type meqT) (Vartab.empty,0)); fun mk_selR subst gT (upd,uT) = let val vT = range_type (hd (binder_types uT)); in Const (unsuffix Record.updateN upd,gT --> (Envir.norm_type subst vT)) end; (* lookup:: "('v => 'a) => 'n => ('n => 'v) => 'a" * update:: "('v => 'a) => ('a => 'v) => 'n => ('a => 'a) => ('n => 'v) => ('n => 'v)" *) fun mk_selF subst uT d n = let val vT_a::a_vT::nT::aT_aT::n_vT::_ = binder_types uT; val lT = (Envir.norm_type subst (vT_a --> nT --> n_vT --> (domain_type aT_aT))); val d' = map_types (Envir.norm_type subst) d; in Const (@{const_name StateFun.lookup},lT)$d'$n end; fun mk_rupdR subst gT (upd,uT) = let val [vT,_] = binder_types uT in Const (upd,(Envir.norm_type subst vT) --> gT --> gT) end; fun K_fun kn uT = let val T=range_type (hd (binder_types uT)) in Const (kn,T --> T --> T) end; fun K_rec uT t = let val T=range_type (hd (binder_types uT)) in Abs ("_", T, incr_boundvars 1 t) end; fun mk_supdF subst uT d c n = let val uT' = Envir.norm_type subst uT; val c' = map_types (Envir.norm_type subst) c; val d' = map_types (Envir.norm_type subst) d; in Const (@{const_name StateFun.update},uT')$d'$c'$n end; fun modify_updatesR subst gT glob ((Const (upd,uT))$_$(Const _$Z)) = mk_rupdR subst gT (upd,uT)$ (K_rec uT (mk_selR subst gT (upd,uT)$(glob$Bound 0)))$(glob$Bound 1) | modify_updatesR subst gT glob ((Const (upd,uT))$_$s) = mk_rupdR subst gT (upd,uT)$ (K_rec uT (mk_selR subst gT (upd,uT)$(glob$Bound 0)))$ modify_updatesR subst gT glob s | modify_updatesR subst gT glob ((_$Z)) = (glob$Bound 1) (* may_not_modify *) | modify_updatesR _ _ _ t = raise TERM ("gen_call_tac.modify_updatesR",[t]); fun modify_updatesF subst _ glob (Const (@{const_name StateFun.update},uT)$d$c$n$_$(Const globs$Z)) = mk_supdF subst uT d c n$ (K_fun KNF uT$(mk_selF subst uT d n$(glob$Bound 0)))$(glob$Bound 1) | modify_updatesF subst gT glob (Const (@{const_name StateFun.update},uT)$d$c$n$_$s) = mk_supdF subst uT d c n$ (K_fun KNF uT$(mk_selF subst uT d n$(glob$Bound 0)))$modify_updatesF subst gT glob s | modify_updatesF subst _ glob ((globs$Z)) = (glob$Bound 1) (* may_not_modify *) | modify_updatesF _ _ _ t = raise TERM ("gen_call_tac.modify_updatesF",[t]); fun modify_updates Record = modify_updatesR | modify_updates _ = modify_updatesF fun globalsT (Const (gupd,T)) = domain_type (hd (binder_types T)) | globalsT t = raise TERM ("gen_call_tac.globalsT",[t]); fun mk_upd meqT mods (gupd$(Abs (dmy,dmyT,(glob$Bound 1)))$Bound 1) = let val gT = (globalsT gupd); val subst = mk_subst gT meqT; in (gupd$(Abs (dmy,dmyT,incr_boundvars 1 (modify_updates state_kind subst gT glob mods)))$Bound 1) end | mk_upd meqT mods (upd$v$s) = upd$v$mk_upd meqT mods s | mk_upd _ _ t = raise TERM ("gen_call_tac.mk_upd",[t]); fun modify_return (meqT,mods) (Abs (s,T,Abs (t,U,upd))) = (Abs (s,T,Abs (t,U,mk_upd meqT mods upd))) | modify_return _ t = raise TERM ("get_call_tac.modify_return",[t]); - fun modify_tac spec modifies_thm i = + fun modify_tac ctxt' spec modifies_thm i = try (fn () => let val (_,call,modif_spec_nrm,modif_spec_abr,modif_spec_mode,G,Theta,_) = dest_hoare (Thm.concl_of modifies_thm); val is_abr = not (is_empty_set modif_spec_abr); val emptyTheta = is_empty_set Theta; (*val emptyFaults = is_empty_set F;*) val spec_has_args = #6 (dest_call call); val () = if spec_has_args then error "procedure call in modifies-specification must be parameterless!" else (); val (mxprem,ModRet) = (case cmode of Static => (8,if is_abr then if emptyTheta then (ProcModifyReturn mode) else (ProcModifyReturnSameFaults mode) else if emptyTheta then (ProcModifyReturnNoAbr mode) else (ProcModifyReturnNoAbrSameFaults mode)) | Parameter => (9,if is_abr then if emptyTheta then (ProcProcParModifyReturn mode) else (ProcProcParModifyReturnSameFaults mode) else if emptyTheta then (ProcProcParModifyReturnNoAbr mode) else (ProcProcParModifyReturnNoAbrSameFaults mode))); val to_prove_prem = (case cmode of Static => 0 | Parameter => 1); val spec_goal = if is_abr then i + mxprem - 5 else i + mxprem - 6 val mods_nrm = modif_spec_nrm |> get_modifies |> get_updates; val return' = modify_return mods_nrm return; (* val return' = if is_abr then let val mods_abr = modif_spec_abr |> get_modifies |> get_updates; in modify_return mods_abr return end else return;*) - val cret = Thm.cterm_of ctxt return'; + val cret = Thm.cterm_of ctxt' return'; val (_,_,return'_var,_,_,_) = nth (Thm.prems_of ModRet) to_prove_prem |> dest_hoare |> #2 |> dest_call; - val ModRet' = infer_instantiate ctxt [(#1 (dest_Var return'_var), cret)] ModRet; + val ModRet' = infer_instantiate ctxt' [(#1 (dest_Var return'_var), cret)] ModRet; val modifies_thm_partial = if modif_spec_mode = Total then @{thm HoareTotal.hoaret_to_hoarep'} OF [modifies_thm] else modifies_thm; fun solve_modifies_tac i = - (clarsimp_tac ((ctxt + (clarsimp_tac ((ctxt' |> put_claset (claset_of @{theory_context Set}) |> put_simpset (simpset_of @{theory_context Set})) addsimps ([@{thm Hoare.mex_def},@{thm Hoare.meq_def},@{thm StateSpace.upd_globals_def}]@K_convs) addsimprocs (state_upd_simproc Record::(state_simprocs state_kind)) |> fold Simplifier.add_cong K_congs) i) THEN_MAYBE - EVERY [trace_tac ctxt "modify_tac: splitting record", - state_split_simp_tac ctxt [] state_space i]; + EVERY [trace_tac ctxt' "modify_tac: splitting record", + state_split_simp_tac ctxt' [] state_space i]; val cnt = i + mxprem; in - EVERY[trace_tac ctxt "call_tac -- modifies_tac --", - resolve_tac ctxt [ModRet'] i, - solve_spec (AugmentContext Partial) (AsmUN Partial) + EVERY[trace_tac ctxt' "call_tac -- modifies_tac --", + resolve_tac ctxt' [ModRet'] i, + solve_spec ctxt' (AugmentContext Partial) (AsmUN Partial) (AugmentEmptyFaults Partial) Partial Static (SOME Partial) (SOME modifies_thm_partial) spec_goal, if is_abr then - EVERY [trace_tac ctxt "call_tac -- Solving abrupt modifies --", + EVERY [trace_tac ctxt' "call_tac -- Solving abrupt modifies --", solve_modifies_tac (cnt - 6)] else all_tac, - trace_tac ctxt "call_tac -- Solving Modifies --", + trace_tac ctxt' "call_tac -- Solving Modifies --", solve_modifies_tac (cnt - 7), - basic_tac spec (cnt - 8), + basic_tac ctxt' spec (cnt - 8), if cmode = Parameter - then EVERY [resolve_tac ctxt [subset_refl] (cnt - 8), - simp_tac (put_simpset HOL_basic_ss ctxt + then EVERY [resolve_tac ctxt' [subset_refl] (cnt - 8), + simp_tac (put_simpset HOL_basic_ss ctxt' addsimps (@{thm Hoare.CollectInt_iff} :: @{thms simp_thms})) 1] else all_tac] end) () |> (fn SOME res => res | NONE => raise TERM ("get_call_tac.modify_tac: no proper modifies spec", [])); - fun specs_of_assms_tac thms = - (case get_spec pname is_spec_clause thms of - SOME (_,spec) => (case get_spec pname is_modifies_clause thms of - SOME (_,modifies_thm) => modify_tac (SOME spec) modifies_thm 1 - | NONE => basic_tac (SOME spec) 1) + fun specs_of_assms_tac ({context = ctxt', prems, ...}: Subgoal.focus) = + (case get_spec pname is_spec_clause prems of + SOME (_,spec) => (case get_spec pname is_modifies_clause prems of + SOME (_,modifies_thm) => modify_tac ctxt' (SOME spec) modifies_thm 1 + | NONE => basic_tac ctxt' (SOME spec) 1) | NONE => (warning ("no proper specification for procedure " ^pname^ " in assumptions"); all_tac)); val test_modify_in_ctxt_tac = let val mname = (suffix modifysfx pname'); val mspec = (case try (Proof_Context.get_thm ctxt) mname of SOME s => SOME s | NONE => (case AList.lookup (op =) asms pname of SOME s => if is_modifies_clause (Thm.concl_of s) then SOME s else NONE | NONE => NONE)); in (case mspec of - NONE => basic_tac spec + NONE => basic_tac ctxt spec | SOME modifies_thm => (case check_spec pname is_modifies_clause modifies_thm of - SOME _ => modify_tac spec modifies_thm + SOME _ => modify_tac ctxt spec modifies_thm | NONE => (warning ("ignoring theorem " ^ (suffix modifysfx pname') ^ "; no proper modifies specification for procedure "^pname'); - basic_tac spec))) + basic_tac ctxt spec))) end; fun inline_bdy_tac has_args i = (case try (Proof_Context.get_thm ctxt) (suffix bodyP pname') of NONE => no_tac | SOME impl => (case try (Proof_Context.get_thm ctxt) (suffix (body_def_sfx^"_def") pname') of NONE => no_tac | SOME bdy => (tracing ("No specification found for procedure \"" ^ pname' ^ "\". Inlining procedure!"); if has_args then EVERY [trace_tac ctxt "inline_bdy_tac args", resolve_tac ctxt [CallBody mode] i, resolve_tac ctxt [impl] (i+3), resolve_tac ctxt [allI] (i+2), resolve_tac ctxt [allI] (i+2), in_assertion_simp_tac ctxt state_kind [] (i+2), - cont_tac (i+2), + cont_tac ctxt (i+2), resolve_tac ctxt [allI] (i+1),in_assertion_simp_tac ctxt state_kind [bdy] (i+1), - cont_tac (i+1), + cont_tac ctxt (i+1), in_assertion_simp_tac ctxt state_kind [@{thm StateSpace.upd_globals_def}] i] else EVERY [trace_tac ctxt "inline_bdy_tac no args", resolve_tac ctxt [ProcBody mode] i, resolve_tac ctxt [impl] (i+2), simp_tac (put_simpset HOL_basic_ss ctxt addsimps [bdy]) (i+1), - cont_tac (i+1)]))); + cont_tac ctxt (i+1)]))); in (case cmode of Static => if get_recursive pname ctxt = SOME false andalso is_none spec then inline_bdy_tac has_args else test_modify_in_ctxt_tac | Parameter => (case spec of - NONE => (tracing "no spec found!"; Subgoal.FOCUS (specs_of_assms_tac o #prems) ctxt) + NONE => (tracing "no spec found!"; Subgoal.FOCUS specs_of_assms_tac ctxt) | SOME spec => (tracing "found spec!"; case check_spec pname is_spec_clause spec of SOME _ => test_modify_in_ctxt_tac | NONE => (warning ("ignoring theorem " ^ (suffix spec_sfx pname') ^ "; no proper specification for procedure " ^pname'); - Subgoal.FOCUS (specs_of_assms_tac o #prems) ctxt)))) + Subgoal.FOCUS specs_of_assms_tac ctxt)))) end; fun call_tac cont_tac mode state_kind state_space ctxt asms spec_sfx t = let val (_,c,_,_,_,_,_,F) = dest_hoare (Logic.strip_assums_concl t); fun gen_tac (_,pname,return,c,cmode,has_args) = gen_call_tac cont_tac mode cmode state_kind state_space ctxt asms spec_sfx (proc_name cmode pname) return has_args F; in gen_tac (dest_call c) end handle TERM _ => K no_tac; fun solve_in_Faults_tac ctxt i = resolve_tac ctxt [UNIV_I, @{thm in_insert_hd}] i ORELSE SELECT_GOAL (SOLVE (simp_tac (put_simpset (simpset_of @{theory_context Set}) ctxt) 1)) i; local fun triv_simp ctxt = merge_assertion_simp_tac ctxt [mem_Collect_eq] THEN' simp_tac (put_simpset HOL_basic_ss ctxt addsimps @{thms simp_thms} |> fold Simplifier.add_cong [@{thm conj_cong}, @{thm imp_cong}]); (* a guarded while produces stupid things, since the guards are put at the end of the body and in the invariant (rule WhileAnnoG): - guard: g /\ g - guarantee: g --> g *) in fun guard_tac ctxt strip cont_tac mode (t,i) = let val (_,c,_,_,_,_,_,F) = dest_hoare (Logic.strip_assums_concl t); val (_,_,_,doStrip) = dest_Guard c; val guarantees = if strip orelse doStrip then [GuardStrip mode, GuaranteeStrip mode] else [Guarantee mode] fun basic_tac i = EVERY [resolve_tac ctxt [Guard mode, GuaranteeAsGuard mode] i, - trace_tac ctxt "Guard", cont_tac (i+1), + trace_tac ctxt "Guard", cont_tac ctxt (i+1), triv_simp ctxt i] fun guarantee_tac i = EVERY [resolve_tac ctxt guarantees i, solve_in_Faults_tac ctxt (i+2), - cont_tac (i+1), + cont_tac ctxt (i+1), triv_simp ctxt i] in if is_empty_set F then EVERY [trace_tac ctxt "Guard: basic_tac", basic_tac i] else EVERY [trace_tac ctxt "Guard: trying guarantee_tac", guarantee_tac i ORELSE basic_tac i] end handle TERM _ => no_tac end; fun wf_tac ctxt = simp_tac (put_simpset HOL_basic_ss ctxt addsimps [@{thm Wellfounded.wf_measure},@{thm Wellfounded.wf_lex_prod},@{thm Wfrec.wf_same_fst}, @{thm Hoare.wf_measure_lex_prod},@{thm Wellfounded.wf_inv_image}]); fun in_rel_simp ctxt = simp_tac (put_simpset HOL_basic_ss ctxt addsimps [@{thm Hoare.in_measure_iff},@{thm Hoare.in_lex_iff},@{thm Hoare.in_mlex_iff},@{thm Hoare.in_inv_image_iff}, @{thm split_conv}]); fun while_annotate_tac ctxt inv i st = let val annotateWhile = Thm.lift_rule (Thm.cprem_of st i) @{thm HoarePartial.reannotateWhileNoGuard}; val lifted_inv = fold_rev Term.abs (Logic.strip_params (Logic.get_goal (Thm.prop_of st) i)) inv; val invVar = (#1 o strip_comb o #3 o dest_whileAnno o #2 o dest_hoare) (List.last (Thm.prems_of annotateWhile)); val annotate = infer_instantiate ctxt [(#1 (dest_Var invVar), Thm.cterm_of ctxt lifted_inv)] annotateWhile; in ((trace_tac ctxt ("annotating While with: " ^ Syntax.string_of_term ctxt lifted_inv )) THEN compose_tac ctxt (false,annotate,1) i) st end; fun cond_annotate_tac ctxt inv mode (_,i) st = let val annotateCond = Thm.lift_rule (Thm.cprem_of st i) (CondInv' mode); val lifted_inv = fold_rev Term.abs (Logic.strip_params (Logic.get_goal (Thm.prop_of st) i)) inv; val invVar = List.last (Thm.prems_of annotateCond) |> dest_hoare |> #3 |> strip_comb |> #1; val annotate = infer_instantiate ctxt [(#1 (dest_Var invVar), Thm.cterm_of ctxt lifted_inv)] annotateCond; in ((trace_tac ctxt ("annotating Cond with: "^ Syntax.string_of_term ctxt lifted_inv)) THEN compose_tac ctxt (false,annotate,5) i) st end; fun basic_while_tac ctxt state_kind cont_tac tac mode i = let fun common_tac i = EVERY[if mode=Total then wf_tac ctxt (i+3) else all_tac, BasicSimpTac ctxt state_kind true [] tac (i+2), if mode=Total then in_rel_simp ctxt (i+1) THEN (resolve_tac ctxt [allI] (i+1)) else all_tac, - cont_tac (i+1) + cont_tac ctxt (i+1) ] fun basic_tac i = EVERY [resolve_tac ctxt [While mode] i, common_tac i] in EVERY [trace_tac ctxt "basic_while_tac: basic_tac", basic_tac i] end; fun while_tac ctxt state_kind inv cont_tac tac mode t i= let val basic_tac = basic_while_tac ctxt state_kind cont_tac tac mode; in (case inv of NONE => basic_tac i | SOME I => EVERY [while_annotate_tac ctxt I i, basic_tac i]) end handle TERM _ => no_tac fun dest_split (Abs (x,T,t)) = let val (vs,recomb,bdy) = dest_split t; in ((x,T)::vs,fn t' => Abs (x,T,recomb t'),bdy) end | dest_split (c as Const (@{const_name case_prod},_)$Abs(x,T,t)) = let val (vs,recomb,bdy) = dest_split t; in ((x,T)::vs,fn t' => c$Abs (x,T,recomb t'),bdy) end | dest_split t = ([],I,t); fun whileAnnoG_tac ctxt strip_guards mode t i st = let val (_,c,_,_,_,_,_,F) = dest_hoare (Logic.strip_assums_concl t); val (SOME grds,_,I,_,_,fix) = dest_whileAnno c; val Rule = if fix then WhileAnnoGFix mode else WhileAnnoG mode; fun extract_faults (Const (@{const_name Set.insert}, _) $ t $ _) = [t] | extract_faults _ = []; fun leave_grd fs (Const (@{const_name Pair}, _) $ f $ g) = if member (op =) fs f andalso strip_guards then NONE else SOME g | leave_grd fs (Const (@{const_name Language.guaranteeStripPair}, _) $ f $ g) = if member (op =) fs f then NONE else SOME g | leave_grd fs _ = NONE; val (I_vs,I_recomb,I') = dest_split I; val grds' = map_filter (leave_grd (extract_faults F)) (HOLogic.dest_list grds); val pars = (Logic.strip_params (Logic.get_goal (Thm.prop_of st) i)); val J = fold_rev Term.abs pars (I_recomb (fold_rev (mk_Int (map snd (pars@I_vs))) grds' I')); val WhileG = Thm.lift_rule (Thm.cprem_of st i) Rule; val invVar = (fst o strip_comb o #3 o dest_whileAnno o (fn xs => nth xs 1) o snd o strip_comb o #2 o dest_hoare) (List.last (Thm.prems_of WhileG)); val WhileGInst = infer_instantiate ctxt [(#1 (dest_Var invVar), Thm.cterm_of ctxt J)] WhileG; in ((trace_tac ctxt ("WhileAnnoG, adding guards to invariant: " ^ Syntax.string_of_term ctxt J)) THEN compose_tac ctxt (false,WhileGInst,1) i) st end handle TERM _ => no_tac st | Bind => no_tac st (* renames bound state variable according to name given in goal, * before rule specAnno is applied, and solves sidecondition *) fun gen_Anno_tac dest rules tac cont_tac ctxt state_kind (t,i) st = let val vars = (dest o #2 o dest_hoare) (Logic.strip_assums_concl t); val rules' = (case (List.filter (not o null) (map dest_splits vars)) of [] => rules |(xs::_) => adapt_aux_var ctxt false (map fst xs) (map get_aux_tvar rules)); in EVERY [resolve_tac ctxt rules' i, tac, simp_tac (put_simpset HOL_basic_ss ctxt addsimps ([@{thm split_conv}, refl, @{thm Hoare.triv_All_eq}]@anno_defs) addsimprocs [@{simproc case_prod_beta}]) (i+2), simp_tac (put_simpset HOL_basic_ss ctxt addsimps [@{thm split_conv}]) (i+1), simp_tac (put_simpset HOL_basic_ss ctxt addsimps [@{thm split_conv}]) i, REPEAT (resolve_tac ctxt [allI] (i+1)), - cont_tac (i+1), + cont_tac ctxt (i+1), conseq_simp_tac ctxt state_kind [] i] st end handle TERM _ => no_tac st; fun specAnno_tac ctxt state_kind cont_tac mode = let fun dest_specAnno (Const (@{const_name Language.specAnno},_)$P$c$Q$A) = [P,c,Q,A] | dest_specAnno t = raise TERM ("dest_specAnno",[t]); val rules = [SpecAnnoNoAbrupt mode,SpecAnno mode]; in gen_Anno_tac dest_specAnno rules all_tac cont_tac ctxt state_kind end; fun whileAnnoFix_tac ctxt state_kind cont_tac mode (t,i) = let fun dest (Const (@{const_name Language.whileAnnoFix},_)$b$I$V$c) = [I,V,c] | dest t = raise TERM ("dest_whileAnnoFix",[t]); val rules = [WhileAnnoFix mode]; fun wf_tac' i = EVERY [REPEAT (resolve_tac ctxt [allI] i), simp_tac (put_simpset HOL_basic_ss ctxt addsimps [@{thm split_conv}]) i, wf_tac ctxt i]; val tac = if mode=Total then EVERY [wf_tac' (i+3),in_rel_simp ctxt (i+1)] else all_tac in gen_Anno_tac dest rules tac cont_tac ctxt state_kind (t,i) end; fun lemAnno_tac ctxt state_kind mode (t,i) st = let fun dest_name (Const (x,_)) = x | dest_name (Free (x,_)) = x | dest_name t = raise TERM ("dest_name",[t]); - fun dest_lemAnno ctxt (Const (@{const_name Language.lem},_)$n$c) = + fun dest_lemAnno (Const (@{const_name Language.lem},_)$n$c) = let val x = Long_Name.base_name (dest_name n); in (case try (Proof_Context.get_thm ctxt) x of NONE => error ("No lemma: '" ^ x ^ "' found.") | SOME spec => (strip_qnt_vars @{const_name All} (HOLogic.dest_Trueprop (Thm.concl_of spec)),spec)) end - | dest_lemAnno _ t = raise TERM ("dest_lemAnno",[t]); - - val (vars,spec) = dest_lemAnno ctxt (#2 (dest_hoare t)); + | dest_lemAnno t = raise TERM ("dest_lemAnno",[t]); + + val (vars,spec) = dest_lemAnno (#2 (dest_hoare t)); val rules = [LemAnnoNoAbrupt mode,LemAnno mode]; val rules' = (case vars of [] => rules | xs => adapt_aux_var ctxt true (map fst xs) (map get_aux_tvar rules)); in EVERY [resolve_tac ctxt rules' i, resolve_tac ctxt [spec] (i+1), conseq_simp_tac ctxt state_kind [] i] st end handle TERM _ => no_tac st; fun prems_tac ctxt i = TRY (resolve_tac ctxt (Assumption.all_prems_of ctxt) i); fun mk_proc_assoc thms = let fun name (_,p,_,_,cmode,_) = proc_name cmode p; fun proc_name thm = thm |> Thm.concl_of |> dest_hoare |> #2 |> dest_call |> name; in map (fn thm => (proc_name thm,thm)) thms end; fun mk_hoare_tac cont ctxt mode i (name,tac) = EVERY [trace_tac ctxt ("trying: " ^ name),tac cont ctxt mode i]; (* the main hoare tactic *) fun HoareTac annotate_inv exspecs strip_guards mode state_kind state_space spec_sfx ctxt tac st = let val (P,c,Q,A,_,G,T,F) = dest_hoare (Logic.strip_assums_concl (Logic.get_goal (Thm.prop_of st) 1)); val wp_tacs = #wp_tacs (get_data ctxt); val hoare_tacs = #hoare_tacs (get_data ctxt); val params = (strip_vars (Logic.get_goal (Thm.prop_of st) 1)); val Inv = (if annotate_inv then (* Take the postcondition of the triple as invariant for all *) (* while loops (makes sense for the modifies clause) *) SOME Q else NONE); val exspecthms = map (Proof_Context.get_thm ctxt) exspecs; val asms = try (fn () => mk_proc_assoc (gen_context_thms ctxt mode params G T F @ exspecthms)) () |> the_default []; fun while_annoG_tac (t,i) = whileAnnoG_tac ctxt (annotate_inv orelse strip_guards) mode t i; fun WlpTac tac i = (* WlpTac does not end with subset_refl *) FIRST - ([EVERY [resolve_tac ctxt [Seq mode] i,trace_tac ctxt "Seq",HoareRuleTac tac false (i+1)], - EVERY [resolve_tac ctxt [Catch mode] i,trace_tac ctxt "Catch",HoareRuleTac tac false (i+1)], - EVERY [resolve_tac ctxt [CondCatch mode] i,trace_tac ctxt "CondCatch",HoareRuleTac tac false (i+1)], - EVERY [resolve_tac ctxt [BSeq mode] i,trace_tac ctxt "BSeq",HoareRuleTac tac false (i+1)], + ([EVERY [resolve_tac ctxt [Seq mode] i,trace_tac ctxt "Seq",HoareRuleTac tac false ctxt (i+1)], + EVERY [resolve_tac ctxt [Catch mode] i,trace_tac ctxt "Catch",HoareRuleTac tac false ctxt (i+1)], + EVERY [resolve_tac ctxt [CondCatch mode] i,trace_tac ctxt "CondCatch",HoareRuleTac tac false ctxt (i+1)], + EVERY [resolve_tac ctxt [BSeq mode] i,trace_tac ctxt "BSeq",HoareRuleTac tac false ctxt (i+1)], EVERY [resolve_tac ctxt [FCall mode] i,trace_tac ctxt "FCall"], EVERY [resolve_tac ctxt [GuardsNil mode] i,trace_tac ctxt "GuardsNil"], EVERY [resolve_tac ctxt [GuardsConsGuaranteeStrip mode] i, trace_tac ctxt "GuardsConsGuaranteeStrip"], EVERY [resolve_tac ctxt [GuardsCons mode] i,trace_tac ctxt "GuardsCons"], EVERY [SUBGOAL while_annoG_tac i] ] @ - map (mk_hoare_tac (HoareRuleTac tac) ctxt mode i) wp_tacs) - - and HoareRuleTac tac pre_cond i st = + map (mk_hoare_tac (fn p => HoareRuleTac tac p ctxt) ctxt mode i) wp_tacs) + + and HoareRuleTac tac pre_cond ctxt i st = let fun call (t,i) = call_tac (HoareRuleTac tac false) mode state_kind state_space ctxt asms spec_sfx t i fun cond_tac i = if annotate_inv andalso Config.get ctxt use_cond_inv_modifies then EVERY[SUBGOAL (cond_annotate_tac ctxt (the Inv) mode) i, - HoareRuleTac tac false (i+4), - HoareRuleTac tac false (i+3), + HoareRuleTac tac false ctxt (i+4), + HoareRuleTac tac false ctxt (i+3), BasicSimpTac ctxt state_kind true [] tac (i+2), BasicSimpTac ctxt state_kind true [] tac (i+1) ] else EVERY[resolve_tac ctxt [Cond mode] i,trace_tac ctxt "Cond", - HoareRuleTac tac false (i+2), - HoareRuleTac tac false (i+1)]; + HoareRuleTac tac false ctxt (i+2), + HoareRuleTac tac false ctxt (i+1)]; fun switch_tac i = EVERY[resolve_tac ctxt [SwitchNil mode] i, trace_tac ctxt "SwitchNil"] ORELSE EVERY[resolve_tac ctxt [SwitchCons mode] i,trace_tac ctxt "SwitchCons", - HoareRuleTac tac false (i+2), - HoareRuleTac tac false (i+1)]; + HoareRuleTac tac false ctxt (i+2), + HoareRuleTac tac false ctxt (i+1)]; fun while_tac' (t,i) = while_tac ctxt state_kind Inv (HoareRuleTac tac true) tac mode t i; in st |> - ( (WlpTac tac i THEN HoareRuleTac tac pre_cond i) + ( (WlpTac tac i THEN HoareRuleTac tac pre_cond ctxt i) ORELSE (FIRST([EVERY[resolve_tac ctxt [Skip mode] i,trace_tac ctxt "Skip"], EVERY[resolve_tac ctxt [BasicCond mode] i, trace_tac ctxt "BasicCond", assertion_simp_tac ctxt state_kind [] i], (resolve_tac ctxt [Basic mode] i THEN trace_tac ctxt "Basic") THEN_MAYBE (assertion_simp_tac ctxt state_kind [] i), (* we don't really need simplificaton here. The question is if it is better to simplify the assertion after each Basic step, so that intermediate assertions stay "small", or if we just accumulate the raw assertions and leave the simplification to the final BasicSimpTac *) EVERY[resolve_tac ctxt [Throw mode] i, trace_tac ctxt "Throw"], (resolve_tac ctxt [Raise mode] i THEN trace_tac ctxt "Raise") THEN_MAYBE (assertion_string_eq_simp_tac ctxt state_kind [] i), cond_tac i, switch_tac i, EVERY[resolve_tac ctxt [Block mode] i, trace_tac ctxt "Block", resolve_tac ctxt [allI] (i+2), resolve_tac ctxt [allI] (i+2), - HoareRuleTac tac false (i+2), + HoareRuleTac tac false ctxt (i+2), resolve_tac ctxt [allI] (i+1), in_assertion_simp_tac ctxt state_kind [] (i+1), - HoareRuleTac tac false (i+1)], + HoareRuleTac tac false ctxt (i+1)], SUBGOAL while_tac' i, SUBGOAL (guard_tac ctxt (annotate_inv orelse strip_guards) (HoareRuleTac tac false) mode) i, EVERY[SUBGOAL (specAnno_tac ctxt state_kind (HoareRuleTac tac true) mode) i], EVERY[SUBGOAL (whileAnnoFix_tac ctxt state_kind (HoareRuleTac tac true) mode) i], EVERY[resolve_tac ctxt [SpecIf mode] i, trace_tac ctxt "SpecIf", assertion_simp_tac ctxt state_kind [] i], (resolve_tac ctxt [Spec mode] i THEN trace_tac ctxt "Spec") THEN_MAYBE (assertion_simp_tac ctxt state_kind [@{thm split_conv}] i), EVERY[resolve_tac ctxt [BindR mode] i, trace_tac ctxt "Bind", resolve_tac ctxt [allI] (i+1), - HoareRuleTac tac false (i+1)], + HoareRuleTac tac false ctxt (i+1)], EVERY[resolve_tac ctxt [DynCom mode] i, trace_tac ctxt "DynCom"], EVERY[trace_tac ctxt "calling call_tac",SUBGOAL call i], EVERY[trace_tac ctxt "LemmaAnno",SUBGOAL (lemAnno_tac ctxt state_kind mode) i]] @ - map (mk_hoare_tac (HoareRuleTac tac) ctxt mode i) hoare_tacs) + map (mk_hoare_tac (fn p => HoareRuleTac tac p ctxt) ctxt mode i) hoare_tacs) THEN (if pre_cond then EVERY [trace_tac ctxt "pre_cond", TRY (BasicSimpTac ctxt state_kind true [] tac i), (* FIXME: Do we need TRY *) trace_tac ctxt "after BasicSimpTac"] else (resolve_tac ctxt [subset_refl] i)))) end; - in ((K (EVERY [REPEAT (resolve_tac ctxt [allI] 1), HoareRuleTac tac true 1])) + in ((K (EVERY [REPEAT (resolve_tac ctxt [allI] 1), HoareRuleTac tac true ctxt 1])) THEN_ALL_NEW (prems_tac ctxt)) 1 st (*Procedure specifications may have an locale assumption as premise. These are accumulated by the vcg and are be solved afterward by prems_tac *) end; fun prefer_tac i = (Tactic.defer_tac i THEN PRIMITIVE (Thm.permute_prems 0 ~1)); fun HoareStepTac strip_guards mode state_kind state_space spec_sfx ctxt tac st = let val asms = try (fn () => let val (_,_,_,_,_,G,T,F) = dest_hoare (Logic.strip_assums_concl (Logic.get_goal (Thm.prop_of st) 1)); val params = (strip_vars (Logic.get_goal (Thm.prop_of st) 1)); in mk_proc_assoc (gen_context_thms ctxt mode params G T F) end) () |> the_default []; - fun result_tac i = TRY (EVERY [resolve_tac ctxt [Basic mode] i, - resolve_tac ctxt [subset_refl] i]); + fun result_tac ctxt' i = TRY (EVERY [resolve_tac ctxt' [Basic mode] i, + resolve_tac ctxt' [subset_refl] i]); fun call (t,i) = call_tac result_tac mode state_kind state_space ctxt asms spec_sfx t i fun final_simp_tac i = EVERY [full_simp_tac (put_simpset HOL_basic_ss ctxt addsimps [mem_Collect_eq]) i, REPEAT (eresolve_tac ctxt [conjE] i), TRY (hyp_subst_tac_thin true ctxt i), BasicSimpTac ctxt state_kind true [] tac i] fun while_annoG_tac (t,i) = whileAnnoG_tac ctxt strip_guards mode t i; in st |> (REPEAT (resolve_tac ctxt [allI] 1) THEN FIRST [resolve_tac ctxt [subset_refl] 1, EVERY[resolve_tac ctxt [Skip mode] 1,TRY (BasicSimpTac ctxt state_kind false [] tac 1)], EVERY[resolve_tac ctxt [BasicCond mode] 1,trace_tac ctxt "BasicCond", TRY (BasicSimpTac ctxt state_kind false [] tac 1)], EVERY[resolve_tac ctxt [Basic mode] 1,TRY (BasicSimpTac ctxt state_kind false [] tac 1)], EVERY[resolve_tac ctxt [Throw mode] 1,TRY (BasicSimpTac ctxt state_kind false [] tac 1)], EVERY[resolve_tac ctxt [Raise mode] 1,TRY (assertion_string_eq_simp_tac ctxt state_kind [] 1)], resolve_tac ctxt [SeqSwap mode] 1 THEN_MAYBE HoareStepTac strip_guards mode state_kind state_space spec_sfx ctxt tac, EVERY[resolve_tac ctxt [BSeq mode] 1, prefer_tac 2 THEN_MAYBE HoareStepTac strip_guards mode state_kind state_space spec_sfx ctxt tac], resolve_tac ctxt [CondSwap mode] 1, resolve_tac ctxt [SwitchNil mode] 1, resolve_tac ctxt [SwitchCons mode] 1, EVERY [SUBGOAL while_annoG_tac 1], EVERY[resolve_tac ctxt [While mode] 1, if mode=Total then wf_tac ctxt 4 else all_tac, BasicSimpTac ctxt state_kind false [] tac 3, if mode=Total then in_rel_simp ctxt 2 THEN (resolve_tac ctxt [allI] 2) else all_tac, BasicSimpTac ctxt state_kind false [] tac 1], resolve_tac ctxt [CatchSwap mode] 1, resolve_tac ctxt [CondCatchSwap mode] 1, EVERY[resolve_tac ctxt [BlockSwap mode] 1, resolve_tac ctxt [allI] 1, resolve_tac ctxt [allI] 1, resolve_tac ctxt [allI] 2, BasicSimpTac ctxt state_kind false [] tac 2], resolve_tac ctxt [GuardsNil mode] 1, resolve_tac ctxt [GuardsConsGuaranteeStrip mode] 1, resolve_tac ctxt [GuardsCons mode] 1, - SUBGOAL (guard_tac ctxt strip_guards (fn i => all_tac) mode) 1, - EVERY[SUBGOAL (specAnno_tac ctxt state_kind (K all_tac) mode) 1], - EVERY[SUBGOAL (whileAnnoFix_tac ctxt state_kind (K all_tac) mode) 1], + SUBGOAL (guard_tac ctxt strip_guards (K (K all_tac)) mode) 1, + EVERY[SUBGOAL (specAnno_tac ctxt state_kind (K (K all_tac)) mode) 1], + EVERY[SUBGOAL (whileAnnoFix_tac ctxt state_kind (K (K all_tac)) mode) 1], EVERY[resolve_tac ctxt [SpecIf mode] 1,trace_tac ctxt "SpecIf", TRY (BasicSimpTac ctxt state_kind false [] tac 1)], EVERY[resolve_tac ctxt [Spec mode] 1, TRY (BasicSimpTac ctxt state_kind false [@{thm split_conv}] tac 1)], EVERY[resolve_tac ctxt [BindR mode] 1, resolve_tac ctxt [allI] 2, prefer_tac 2], EVERY[resolve_tac ctxt [FCall mode] 1], EVERY[resolve_tac ctxt [DynCom mode] 1], EVERY[SUBGOAL call 1, BasicSimpTac ctxt state_kind false [] tac 1], EVERY[SUBGOAL (lemAnno_tac ctxt state_kind mode) 1, BasicSimpTac ctxt state_kind false [] tac 1], final_simp_tac 1 ]) end; (*****************************************************************************) (** Generalise verification condition **) (*****************************************************************************) structure RecordSplitState : SPLIT_STATE = struct val globals = @{const_name StateSpace.state.globals}; fun isState (Const _$Abs (s,T,t)) = (case (state_hierarchy T) of ((n,_)::_) => n = "StateSpace.state.state" andalso is_none (try dest_hoare_raw (strip_qnt_body @{const_name All} t)) | _ => false) | isState _ = false; fun isFreeState (Free (_,T)) = (case (state_hierarchy T) of ((n,_)::_) => n = "StateSpace.state.state" | _ => false) | isFreeState _ = false; val abs_state = Option.map snd o first_subterm isFreeState; fun sel_eq (Const (x,_)$_) y = (x=y) | sel_eq t y = raise TERM ("RecordSplitState.sel_eq",[t]); val sel_idx = idx sel_eq; fun bound xs (t as (Const (x,_)$_)) = let val i = sel_idx xs x in if i < 0 then (length xs, xs@[t]) else (i,xs) end | bound xs t = raise TERM ("RecordSplitState.bound",[t]); fun abs_var _ (Const (x,T)$_) = (remdeco' (Long_Name.base_name x),range_type T) | abs_var _ t = raise TERM ("RecordSplitState.abs_var",[t]); fun fld_eq (x, _) y = (x = y) fun fld_idx xs x = idx fld_eq xs x; fun sort_vars ctxt T vars = let val thy = Proof_Context.theory_of ctxt; val (flds,_) = Record.get_recT_fields thy T; val gT = the (AList.lookup (fn (x:string,y) => x=y) flds globals); val (gflds,_) = (Record.get_recT_fields thy gT handle TYPE _ => ([],("",dummyT))); fun compare (Const _$Free _, Const _$(Const _$Free _)) = GREATER | compare (Const (s1,_)$Free _, Const (s2,_)$Free _) = int_ord (fld_idx flds s1, fld_idx flds s2) | compare (Const (s1,_)$(Const _$Free _), Const (s2,_)$(Const _$Free _)) = int_ord (fld_idx gflds s1, fld_idx gflds s2) | compare _ = LESS; in sort (rev_order o compare) vars end; fun fold_state_prop loc glob app abs other inc s res (t as (Const (sel,_)$Free (s',_))) = if s'=s then if is_state_var sel then loc inc res t else raise TERM ("RecordSplitState.fold_state_prop",[t]) else other res t | fold_state_prop loc glob app abs other inc s res (t as ((t1 as (Const (sel,_)))$(t2 as (Const (glb,_)$Free (s',_))))) = if s'=s andalso is_state_var sel andalso (glb=globals) then glob inc res t else let val res1 = fold_state_prop loc glob app abs other inc s res t1 val res2 = fold_state_prop loc glob app abs other inc s res1 t2 in app res1 res2 end | fold_state_prop loc glob app abs other inc s res (t as (Free (s',_))) = if s'=s then raise TERM ("RecordSplitState.fold_state_prop",[t]) else other res t | fold_state_prop loc glob app abs other inc s res (t1$t2) = let val res1 = fold_state_prop loc glob app abs other inc s res t1 val res2 = fold_state_prop loc glob app abs other inc s res1 t2 in app res1 res2 end | fold_state_prop loc glob app abs other inc s res (Abs (x,T,t)) = let val res1 = fold_state_prop loc glob app abs other (inc+1) s res t in abs x T res1 end | fold_state_prop loc glob app abs other inc s res t = other res t fun collect_vars s t = let fun loc _ vars t = snd (bound vars t); fun glob _ vars t = snd (bound vars t); fun app _ vars2 = vars2; fun abs _ _ vars = vars; fun other vars _ = vars; in fold_state_prop loc glob app abs other 0 s [] t end; fun abstract_vars vars s t = let fun loc inc _ t = let val i = fst (bound vars t) in Bound (i+inc) end; fun glob inc _ t = let val i = fst (bound vars t) in Bound (i+inc) end; fun app t1 t2 = t1$t2; fun abs x T t = Abs (x,T,t); fun other _ t = t; val dummy = Bound 0; in fold_state_prop loc glob app abs other 0 s dummy t end; fun split_state ctxt s T t = let val vars = collect_vars s t; val vars' = if Config.get ctxt sort_variables then sort_vars ctxt T vars else vars; in (abstract_vars vars' s t,rev vars') end; fun ex_tac ctxt _ st = Record.split_simp_tac ctxt @{thms simp_thms} (K ~1) 1 st; end; structure FunSplitState : SPLIT_STATE = struct val full_globalsN = @{const_name StateSpace.state.globals}; fun isState (Const _$Abs (s,T,t)) = (case (state_hierarchy T) of ((n,_)::_) => n = "StateSpace.state.state" andalso is_none (try dest_hoare_raw (strip_qnt_body @{const_name All} t)) | _ => false) | isState _ = false; fun isFreeState (Free (_,T)) = (case (state_hierarchy T) of ((n,_)::_) => n = "StateSpace.state.state" | _ => false) | isFreeState _ = false; val abs_state = Option.map snd o first_subterm isFreeState; fun comp_name t = case try (implode o dest_string) t of SOME str => str | NONE => (case t of Free (s,_) => s | Const (s,_) => s | t => raise TERM ("FunSplitState.comp_name",[t])) fun sel_name (Const _$_$name$_) = comp_name name | sel_name t = raise TERM ("FunSplitState.sel_name",[t]); fun sel_raw_name (Const _$_$name$_) = name | sel_raw_name t = raise TERM ("FunSplitState.sel_raw_name",[t]); fun component_type (Const _$_$_$(sel$_)) = range_type (fastype_of sel) | component_type t = raise TERM ("FunSplitState.component_type",[t]); fun component_name (Const _$_$_$((Const (sel,_)$_))) = sel | component_name t = raise TERM ("FunSplitState.component_name",[t]); fun sel_type (Const _$destr$_$_) = range_type (fastype_of destr) | sel_type t = raise TERM ("FunSplitState.sel_type",[t]); fun sel_destr (Const _$destr$_$_) = destr | sel_destr t = raise TERM ("FunSplitState.sel_destr",[t]); fun sel_eq t y = (sel_name t = y) | sel_eq t y = raise TERM ("FunSplitState.sel_eq",[t]); val sel_idx = idx sel_eq; fun bound xs t = let val i = sel_idx xs (sel_name t) in if i < 0 then (length xs, xs@[t]) else (i,xs) end | bound xs t = raise TERM ("FunSplitState.bound",[t]); fun fold_state_prop var app abs other inc s res (t as (Const (@{const_name StateFun.lookup},_)$destr$name$(Const _$Free (s',_)))) = if s'=s then var inc res t else other res t (*raise TERM ("FunSplitState.fold_state_prop",[t])*) | fold_state_prop var app abs other inc s res (t as (Free (s',_))) = if s'=s then raise TERM ("FunSplitState.fold_state_prop",[t]) else other res t | fold_state_prop var app abs other inc s res (t1$t2) = let val res1 = fold_state_prop var app abs other inc s res t1 val res2 = fold_state_prop var app abs other inc s res1 t2 in app res1 res2 end | fold_state_prop var app abs other inc s res (Abs (x,T,t)) = let val res1 = fold_state_prop var app abs other (inc+1) s res t in abs x T res1 end | fold_state_prop var app abs other inc s res t = other res t fun collect_vars s t = let fun var _ vars t = snd (bound vars t); fun app _ vars2 = vars2; fun abs _ _ vars = vars; fun other vars _ = vars; in fold_state_prop var app abs other 0 s [] t end; fun abstract_vars vars s t = let fun var inc _ t = let val i = fst (bound vars t) in Bound (i+inc) end; fun app t1 t2 = t1$t2; fun abs x T t = Abs (x,T,t); fun other _ t = t; val dummy = Bound 0; in fold_state_prop var app abs other 0 s dummy t end; -fun sort_vars ctxt vars = +fun sort_vars _ vars = let val fld_idx = idx (fn s1:string => fn s2 => s1 = s2); fun compare (_$_$n$(Const (s1,_)$_),_$_$m$(Const (s2,_)$_)) = let val n' = remdeco' (comp_name n); val m' = remdeco' (comp_name m); in if s1 = full_globalsN then if s2 = full_globalsN then string_ord (n',m') else LESS else if s2 = full_globalsN then GREATER else string_ord (n',m') end | compare (t1,t2) = raise TERM ("FunSplitState.sort_vars.compare",[t1,t2]); in sort (rev_order o compare) vars end; fun split_state ctxt s _ t = let val vars = collect_vars s t; val vars' = if Config.get ctxt sort_variables then sort_vars ctxt vars else vars; in (abstract_vars vars' s t,rev vars') end; fun abs_var _ t = (remdeco' (sel_name t), sel_type t); (* Proof for: EX x_1 ... x_n. P x_1 ... x_n * ==> EX s. P (lookup destr_1 "x_1" s) ... (lookup destr_n "x_n" s) * Implementation: * 1. Eliminate existential quantifiers in premise * 2. Instantiate s with: (%x. undefined)("x_1" := constr_1 x_1, ..., "x_n" := constr_n x_n) * 3. Simplify *) local val ss = simpset_of (put_simpset (simpset_of @{theory_context Fun}) @{context} addsimps (@{thm StateFun.lookup_def} :: @{thm StateFun.id_id_cancel} :: @{thms list.inject list.distinct char.inject cong_exp_iff_simps simp_thms}) addsimprocs [Record.simproc, StateFun.lazy_conj_simproc] |> fold Simplifier.add_cong @{thms block_conj_cong}); in fun ex_tac ctxt vs st = let val vs' = rev vs; val (Const (_,exT)$_) = HOLogic.dest_Trueprop (Logic.strip_imp_concl (Logic.get_goal (Thm.prop_of st) 1)); val sT = domain_type (domain_type exT); val s0 = Const (@{const_name HOL.undefined},sT); fun streq (s1:string,s2) = s1=s2 ; fun mk_init [] = [] | mk_init (t::ts) = let val xs = mk_init ts; val n = component_name t; val T = component_type t; in if AList.defined streq xs n then xs else (n,(T,Const (n,sT --> component_type t)$s0))::xs end; fun mk_upd (i,t) xs = let val selN = component_name t; val selT = component_type t; val (_,s) = the (AList.lookup streq xs selN); val strT = domain_type selT; val valT = range_type selT; val constr = destr_to_constr (sel_destr t); val name = (sel_raw_name t); val upd = Const (@{const_name Fun.fun_upd}, (strT --> valT)-->strT-->valT--> (strT --> valT)) $ s $ name $ (constr $ Bound i) in AList.update streq (selN,(selT,upd)) xs end; val upds = fold_index mk_upd vs' (mk_init vs'); val upd = fold (fn (n,(T,upd)) => fn s => Const (n ^ Record.updateN, T --> sT --> sT)$upd$s) upds s0; val inst = fold_rev (Term.abs o (fn t => (sel_name t, sel_type t))) vs upd; fun lift_inst_ex_tac i st = let val rule = Thm.lift_rule (Thm.cprem_of st i) (Drule.incr_indexes st exI); val (_$x) = HOLogic.dest_Trueprop (Logic.strip_assums_concl (hd (Thm.prems_of rule))); val inst_rule = infer_instantiate ctxt [(#1 (dest_Var (head_of x)), Thm.cterm_of ctxt inst)] rule; in (compose_tac ctxt (false,inst_rule, Thm.nprems_of exI) i st) end; in EVERY [REPEAT_DETERM_N (length vs) (eresolve_tac ctxt [exE] 1), lift_inst_ex_tac 1, simp_tac (put_simpset ss ctxt) 1 ] st end end (* Test: What happens when there are no lookups., EX s. True *) end; structure GeneraliseRecord = GeneraliseFun (structure SplitState=RecordSplitState); structure GeneraliseStateFun = GeneraliseFun (structure SplitState=FunSplitState); fun generalise Record = GeneraliseRecord.GENERALISE | generalise Function = GeneraliseStateFun.GENERALISE; (*****************************************************************************) (** record_vanish_tac splits up the records of a verification condition, **) (** trying to generate a predicate without records. **) (** A typical verification condition with a procedure call will have the **) (** form "!!s Z. s=Z ==> ..., where s and Z are records **) (*****************************************************************************) (* FIXME: Check out if removing the useless vars is a performance issue. If so, maybe we can remove all useless vars at once (no iterated simplification) or try to avoid introducing them. Bevore splitting the gaol we can simplifiy the goal with state_simproc this may leed to better performance... *) fun record_vanish_tac ctxt state_kind state_space i = if Config.get ctxt record_vanish then let val rem_useless_vars_simps = [Drule.triv_forall_equality,@{thm Hoare.triv_All_eq},@{thm Hoare.triv_Ex_eq}]; val rem_useless_vars_simpset = empty_simpset ctxt addsimps rem_useless_vars_simps; fun no_spec (t as (Const (@{const_name All},_)$_)) = is_none (try dest_hoare_raw (strip_qnt_body @{const_name All} t)) | no_spec _ = true; fun state_space_no_spec t = if state_space t <> 0 andalso no_spec t then ~1 else 0; in EVERY [trace_tac ctxt "record_vanish_tac -- START --", REPEAT (eresolve_tac ctxt [conjE] i), trace_tac ctxt "record_vanish_tac -- hyp_subst_tac ctxt --", TRY (hyp_subst_tac_thin true ctxt i), full_simp_tac rem_useless_vars_simpset i, (* hyp_subst_tac may have made some state variables unnecessary. We do not want to split them to avoid naming conflicts and increase performance *) trace_tac ctxt "record_vanish_tac -- Splitting records --", if Config.get ctxt use_generalise orelse state_kind = Function then generalise state_kind ctxt i else state_split_simp_tac ctxt rem_useless_vars_simps state_space_no_spec i (*THEN_MAYBE EVERY [trace_tac ctxt "record_vanish_tac -- removing useless vars --", full_simp_tac rem_useless_vars_simpset i, trace_tac ctxt "record_vanish_tac -- STOP --"]*) ] end else all_tac; (* solve_modifies_tac tries to solve modifies-clauses automatically; * The following strategy is followed: * After clar-simplifying the modifies clause we remain with a goal of the form * * EX a b. s(|A := x|) = s(|A:=a,B:=b|) * * (or because of conditional statements conjunctions of these kind of goals) * We split up the state-records and simplify (record_vanish_tac) and get to a goal of the form: * * EX a b. (|A=x,B=B|) = (|A=a,B=b|). * * If the modifies clause was correct we just have to introduce the existential quantifies * and apply reflexivity. * If not we just simplify the goal. *) local val state_fun_update_ss = simpset_of (put_simpset HOL_basic_ss @{context} addsimps ([@{thm StateFun.update_def}, @{thm id_def}, @{thm fun_upd_apply}, @{thm if_True}, @{thm if_False}] @ @{thms list.inject list.distinct char.inject cong_exp_iff_simps simp_thms} @ K_fun_convs) addsimprocs [DistinctTreeProver.distinct_simproc ["distinct_fields", "distinct_fields_globals"]] |> Simplifier.add_cong @{thm imp_cong} (* K_fun_congs FIXME: Stefan fragen*) |> Splitter.add_split @{thm if_split}); in fun solve_modifies_tac ctxt state_kind state_space i st = let val thy = Proof_Context.theory_of ctxt; fun is_split_state (trm as (Const (@{const_name Pure.all},_)$Abs(x,T,t))) = if state_space trm <> 0 then try (fn () => let fun seed (_ $ v $ st) = seed st | seed (_ $ t) = (1,t) (* split only state pair *) | seed t = (~1,t) (* split globals completely *) val all_vars = strip_qnt_vars @{const_name Pure.all} t; val concl = HOLogic.dest_Trueprop (Logic.strip_assums_concl t); val ex_vars = strip_qnt_vars @{const_name Ex} concl; val state = Bound (length all_vars + length ex_vars); val (Const (@{const_name HOL.eq},_)$x_upd$x_upd') = strip_qnt_body @{const_name Ex} concl; val (split,sd) = seed x_upd; in if sd = state then split else 0 end) () |> the_default 0 else 0 | is_split_state t = 0; val simp_ctxt = put_simpset HOL_ss ctxt addsimps (@{thm Ex_True} :: @{thm Ex_False} :: Record.get_extinjects thy); fun try_solve Record i = (*(SOLVE*) (((fn k => (TRY (REPEAT_ALL_NEW (resolve_tac ctxt [conjI, impI, allI]) k))) THEN_ALL_NEW (fn k => EVERY [state_split_simp_tac ctxt [] is_split_state k, simp_tac simp_ctxt k THEN_MAYBE rename_goal ctxt remdeco' k ])) i) (*)*) | try_solve _ i = ((fn k => (TRY (REPEAT_ALL_NEW (resolve_tac ctxt [conjI, impI, allI]) k))) THEN_ALL_NEW (fn k => REPEAT (resolve_tac ctxt [exI] k) THEN resolve_tac ctxt [ext] k THEN simp_tac (put_simpset state_fun_update_ss ctxt) k THEN_MAYBE (REPEAT_ALL_NEW (resolve_tac ctxt [conjI,impI,refl]) k))) i in ((trace_tac ctxt "solve_modifies_tac" THEN clarsimp_tac ((ctxt |> put_claset (claset_of @{theory_context HOL}) |> put_simpset (simpset_of @{theory_context Set})) addsimps ([@{thm Hoare.mex_def},@{thm Hoare.meq_def}]@K_convs) addsimprocs (state_upd_simproc Record::(state_simprocs Record)) |> fold Simplifier.add_cong K_congs) i) THEN_MAYBE try_solve state_kind i) st end; end fun proc_lookup_simp_tac ctxt i st = try (fn () => let val name = (Logic.concl_of_goal (Thm.prop_of st) i) |> dest_hoare |> #2 |> strip_comb |> #2 |> last |> strip_comb |> #2 |> last; (* the$(Gamma$name) or the$(strip$Gamma$name) *) val pname = (unsuffix proc_deco (dest_string' name)); val thms = map_filter I (map (try (Proof_Context.get_thm ctxt)) [suffix bodyP pname, suffix (body_def_sfx^"_def") pname, suffix procL pname^"."^ (suffix (body_def_sfx^"_def") pname)]); in simp_tac (put_simpset HOL_basic_ss ctxt addsimps thms @ strip_simps @ @{thms option.sel}) i st end) () |> the_default (Seq.single st); fun proc_lookup_in_dom_simp_tac ctxt i st = try (fn () => let val _$name$_ = (HOLogic.dest_Trueprop (Logic.concl_of_goal (Thm.prop_of st) i)); (* name : Gamma *) val pname = (unsuffix proc_deco (dest_string' name)); val thms = map_filter I (map (try (Proof_Context.get_thm ctxt)) [suffix bodyP pname]); in simp_tac (put_simpset HOL_basic_ss ctxt addsimps (@{thm Hoare.lookup_Some_in_dom} :: @{thm dom_strip} :: thms)) i st end) () |> the_default (Seq.single st); fun HoareRuleTac ctxt insts fixes st = let val annotate_simp_tac = simp_tac (put_simpset HOL_basic_ss ctxt addsimps (anno_defs@normalize_simps) addsimprocs [@{simproc case_prod_beta}]); fun is_com_eq (Const (@{const_name Trueprop},_)$(Const (@{const_name HOL.eq},T)$_$_)) = (case (binder_types T) of (Type (@{type_name Language.com},_)::_) => true | _ => false) | is_com_eq _ = false; fun annotate_tac i st = if is_com_eq (Logic.concl_of_goal (Thm.prop_of st) i) then annotate_simp_tac i st else all_tac st; in ((fn i => REPEAT (resolve_tac ctxt [allI] i)) THEN' Rule_Insts.res_inst_tac ctxt insts fixes st) THEN_ALL_NEW annotate_tac end; fun HoareCallRuleTac state_kind state_space ctxt thms i st = let fun dest_All (Const (@{const_name All},_)$t) = SOME t | dest_All _ = NONE; fun auxvars t = (case (map_filter ((first_subterm is_hoare) o snd) (max_subterms_dest dest_All t)) of ((vars,_)::_) => vars | _ => []); fun auxtype rule = (case (auxvars (Thm.prop_of rule)) of [] => NONE | vs => (case (last vs) of (_,TVar (z,_)) => SOME (z,rule) | _ => NONE)); val thms' = let val auxvs = map fst (auxvars (Logic.concl_of_goal (Thm.prop_of st) i)); val tvar_thms = map_filter auxtype thms in if length thms = length tvar_thms then adapt_aux_var ctxt true auxvs tvar_thms else thms end; val is_sidecondition = not o can dest_hoare; fun solve_sidecondition_tac (t,i) = if is_sidecondition t then FIRST [CHANGED_PROP (wf_tac ctxt i), (*init_conforms_tac state_kind state_space i,*) post_conforms_tac ctxt state_kind i THEN_MAYBE (if is_modifies_clause t then solve_modifies_tac ctxt state_kind state_space i else all_tac), proc_lookup_in_dom_simp_tac ctxt i ] else in_rel_simp ctxt i THEN simp_tac (put_simpset HOL_basic_ss ctxt addsimps [@{thm Un_empty_left},@{thm Un_empty_right}]) i THEN proc_lookup_simp_tac ctxt i fun basic_tac i = (((resolve_tac ctxt thms') THEN_ALL_NEW (fn k => (SUBGOAL solve_sidecondition_tac k))) i) in (basic_tac ORELSE' (fn k => (REPEAT (resolve_tac ctxt [allI] k)) THEN EVERY [resolve_tac ctxt thms' k])) i st end; (* vcg_polish_tac tries to solve modifies-clauses automatically; for other specifications the * records are only splitted and simplified. *) fun vcg_polish_tac solve_modifies ctxt state_kind state_space i = if solve_modifies then solve_modifies_tac ctxt state_kind state_space i else record_vanish_tac ctxt state_kind state_space i THEN_MAYBE EVERY [rename_goal ctxt remdeco' i(*, simp_tac (HOL_basic_ss addsimps @{thms simp_thms})) i*)]; fun is_funtype (Type ("fun", _)) = true | is_funtype _ = false; fun state_kind_of ctxt T = let val thy = Proof_Context.theory_of ctxt; val (s,sT) = nth (fst (Record.get_recT_fields thy T)) 1; in if Long_Name.base_name s = "locals" andalso is_funtype sT then Function else Record end handle Subscript => Record; fun find_state_space_in_triple ctxt t = try (fn () => (case first_subterm is_hoare t of NONE => NONE | SOME (abs_vars,triple) => let val (_,com,_,_,mode,_,_,_) = dest_hoare_raw triple; val T = fastype_of1 (map snd abs_vars,com) val Type(_,state_spaceT::_) = T; val SOME Tids = stateT_ids state_spaceT; in SOME (Tids,mode, state_kind_of ctxt state_spaceT) end)) () |> Option.join; fun get_state_space_in_subset_eq ctxt t = (* get state type from the following kind of terms: P <= Q, s: P *) try (fn () => let val (subset_eq,_) = (strip_comb o HOLogic.dest_Trueprop o strip_qnt_body @{const_name Pure.all}) t; val Ts = map snd (strip_vars t); val T = fastype_of1 (Ts,subset_eq); val Type (_, [_,Type (_, [Type (_, [state_spaceT]), _])]) = T; (* also works for "in": x : P *) val SOME Tids = stateT_ids state_spaceT; in (Tids,Partial, state_kind_of ctxt state_spaceT) end) (); fun get_state_space ctxt i st = (case try (Logic.concl_of_goal (Thm.prop_of st)) i of SOME t => (case find_state_space_in_triple ctxt t of SOME sp => SOME sp | NONE => get_state_space_in_subset_eq ctxt t) | NONE => NONE); fun mk_hoare_tac hoare_tac finish_tac annotate_inv exnames strip_guards spec_sfx ctxt i st = case get_state_space ctxt i st of SOME (Tids,mode,kind) => SELECT_GOAL (hoare_tac annotate_inv exnames strip_guards mode kind (is_state_space_var Tids) spec_sfx ctxt (finish_tac kind (is_state_space_var Tids))) i st | NONE => no_tac st fun vcg_tac spec_sfx strip_guards exnames ctxt i st = mk_hoare_tac HoareTac (vcg_polish_tac (spec_sfx="_modifies") ctxt) (spec_sfx="_modifies") exnames (strip_guards="true") spec_sfx ctxt i st; fun hoare_tac spec_sfx strip_guards _ ctxt i st = let fun tac state_kind state_space i = if spec_sfx="_modifies" then solve_modifies_tac ctxt state_kind state_space i else all_tac; in mk_hoare_tac HoareTac tac (spec_sfx="_modifies") [] (strip_guards="true") spec_sfx ctxt i st end; fun hoare_raw_tac spec_sfx strip_guards exnames ctxt i st = mk_hoare_tac HoareTac (K (K (K all_tac))) (spec_sfx="_modifies") [] (strip_guards="true") spec_sfx ctxt i st; fun hoare_step_tac spec_sfx strip_guards exnames ctxt i st = mk_hoare_tac (K (K HoareStepTac)) (vcg_polish_tac (spec_sfx="_modifies") ctxt) false [] (strip_guards="true") spec_sfx ctxt i st; fun hoare_rule_tac ctxt thms i st = SUBGOAL (fn _ => (case get_state_space ctxt i st of SOME (Tids,_,kind) => HoareCallRuleTac kind (is_state_space_var Tids) ctxt thms i | NONE => error "could not find proper state space type (structure or record) in goal")) i st; (*** Methods ***) val hoare_rule = Rule_Insts.method HoareRuleTac hoare_rule_tac; val argP = Args.name --| @{keyword "="} -- Args.name val argsP = Scan.repeat argP val default_args = [("spec","spec"),("strip_guards","false")] val vcg_simp_modifiers = [Args.add -- Args.colon >> K (Method.modifier vcg_simp_add \<^here>), Args.del -- Args.colon >> K (Method.modifier vcg_simp_del \<^here>)]; fun assocs2 key = map snd o filter (curry (op =) key o fst); fun gen_simp_method tac = Scan.lift (argsP >> (fn args => args @ default_args)) --| Method.sections vcg_simp_modifiers >> (fn args => fn ctxt => Method.SIMPLE_METHOD' (tac ("_" ^ the (AList.lookup (op =) args "spec")) (the (AList.lookup (op =) args "strip_guards")) (assocs2 "exspec" args) ctxt)); val hoare = gen_simp_method hoare_tac; val hoare_raw = gen_simp_method hoare_raw_tac; val vcg = gen_simp_method vcg_tac; val vcg_step = gen_simp_method hoare_step_tac; val trace_hoare_users = Unsynchronized.ref false fun print_subgoal_tac ctxt s i = SUBGOAL (fn (prem, _) => trace_tac ctxt (s ^ (Syntax.string_of_term ctxt prem))) i fun mk_hoare_thm thm _ ctxt _ i = EVERY [resolve_tac ctxt [thm] i, if !trace_hoare_users then print_subgoal_tac ctxt "Tracing: " i else all_tac] val vcg_hoare_add = Thm.declaration_attribute (fn thm => add_hoare_tacs [(Thm.derivation_name thm, mk_hoare_thm thm)]) exception UNDEF val vcg_hoare_del = Thm.declaration_attribute (fn _ => fn _ => raise UNDEF) (* setup theory *) val _ = Theory.setup (Attrib.setup @{binding vcg_simp} (Attrib.add_del vcg_simp_add vcg_simp_del) "declaration of Simplifier rule for vcg" #> Attrib.setup @{binding vcg_hoare} (Attrib.add_del vcg_hoare_add vcg_hoare_del) "declaration of wp rule for vcg") (*#> add_wp_tacs initial_wp_tacs*) end; diff --git a/thys/Timed_Automata/Closure.thy b/thys/Timed_Automata/Closure.thy --- a/thys/Timed_Automata/Closure.thy +++ b/thys/Timed_Automata/Closure.thy @@ -1,683 +1,683 @@ theory Closure imports Regions begin section \Correct Approximation of Zones with \\\-regions\ locale AlphaClosure = fixes X k \ and V :: "('c, t) cval set" defines "\ \ {region X I r | I r. valid_region X k I r}" defines "V \ {v . \ x \ X. v x \ 0}" assumes finite: "finite X" begin lemmas set_of_regions_spec = set_of_regions[OF _ _ _ finite, of _ k, folded \_def] lemmas region_cover_spec = region_cover[of X _ k, folded \_def] lemmas region_unique_spec = region_unique[of \ X k, folded \_def, simplified] lemmas regions_closed'_spec = regions_closed'[of \ X k, folded \_def, simplified] lemma valid_regions_distinct_spec: "R \ \ \ R' \ \ \ v \ R \ v \ R' \ R = R'" unfolding \_def using valid_regions_distinct by auto (drule valid_regions_distinct, assumption+, simp)+ definition cla ("Closure\<^sub>\ _" [71] 71) where "cla Z = \ {R \ \. R \ Z \ {}}" subsubsection \The nice and easy properties proved by Bouyer\ lemma closure_constraint_id: "\(x, m)\collect_clock_pairs g. m \ real (k x) \ x \ X \ m \ \ \ Closure\<^sub>\ \g\ = \g\ \ V" proof goal_cases case 1 show ?case proof auto fix v assume v: "v \ Closure\<^sub>\ \g\" then obtain R where R: "v \ R" "R \ \" "R \ \g\ \ {}" unfolding cla_def by auto with ccompatible[OF 1, folded \_def] show "v \ \g\" unfolding ccompatible_def by auto from R show "v \ V" unfolding V_def \_def by auto next fix v assume v: "v \ \g\" "v \ V" with region_cover[of X v k, folded \_def] obtain R where "R \ \" "v \ R" unfolding V_def by auto then show "v \ Closure\<^sub>\ \g\" unfolding cla_def using v by auto qed qed lemma closure_id': "Z \ {} \ Z \ R \ R \ \ \ Closure\<^sub>\ Z = R" proof goal_cases case 1 note A = this then have "R \ Closure\<^sub>\ Z" unfolding cla_def by auto moreover { fix R' assume R': "Z \ R' \ {}" "R' \ \" "R \ R'" with A obtain v where "v \ R" "v \ R'" by auto with \_regions_distinct[OF _ A(3) this(1) R'(2-)] \_def have False by auto } ultimately show ?thesis unfolding cla_def by auto qed lemma closure_id: "Closure\<^sub>\ Z \ {} \ Z \ R \ R \ \ \ Closure\<^sub>\ Z = R" proof goal_cases case 1 then have "Z \ {}" unfolding cla_def by auto with 1 closure_id' show ?case by blast qed lemma closure_update_mono: "Z \ V \ set r \ X \ zone_set (Closure\<^sub>\ Z) r \ Closure\<^sub>\(zone_set Z r)" proof - assume A: "Z \ V" "set r \ X" let ?U = "{R \ \. Z \ R \ {}}" from A(1) region_cover_spec have "\ v \ Z. \ R. R \ \ \ v \ R" unfolding V_def by auto then have "Z = \ {Z \ R | R. R \ ?U}" proof (auto, goal_cases) case (1 v) then obtain R where "R \ \" "v \ R" by auto moreover with 1 have "Z \ R \ {}" "v \ Z \ R" by auto ultimately show ?case by auto qed then obtain U where U: "Z = \ {Z \ R | R. R \ U}" "\ R \ U. R \ \" by blast { fix R assume R: "R \ U" { fix v' assume v': "v' \ zone_set (Closure\<^sub>\ (Z \ R)) r - Closure\<^sub>\(zone_set (Z \ R) r)" then obtain v where *: "v \ Closure\<^sub>\ (Z \ R)" "v' = [r \ 0]v" unfolding zone_set_def by auto with closure_id[of "Z \ R" R] R U(2) have **: "Closure\<^sub>\ (Z \ R) = R" "Closure\<^sub>\ (Z \ R) \ \" by fastforce+ with region_set'_id[OF _ *(1) finite _ _ A(2), of k 0, folded \_def, OF this(2)] have ***: "zone_set R r \ \" "[r\0]v \ zone_set R r" unfolding zone_set_def region_set'_def by auto from * have "Z \ R \ {}" unfolding cla_def by auto then have "zone_set (Z \ R) r \ {}" unfolding zone_set_def by auto from closure_id'[OF this _ ***(1)] have "Closure\<^sub>\ zone_set (Z \ R) r = zone_set R r" unfolding zone_set_def by auto with v' **(1) have False by auto } then have "zone_set (Closure\<^sub>\ (Z \ R)) r \ Closure\<^sub>\(zone_set (Z \ R) r)" by auto } note Z_i = this from U(1) have "Closure\<^sub>\ Z = \ {Closure\<^sub>\ (Z \ R) | R. R \ U}" unfolding cla_def by auto then have "zone_set (Closure\<^sub>\ Z) r = \ {zone_set (Closure\<^sub>\ (Z \ R)) r | R. R \ U}" unfolding zone_set_def by auto also have "\ \ \ {Closure\<^sub>\(zone_set (Z \ R) r) | R. R \ U}" using Z_i by auto also have "\ = Closure\<^sub>\ \ {(zone_set (Z \ R) r) | R. R \ U}" unfolding cla_def by auto also have "\ = Closure\<^sub>\ zone_set (\ {Z \ R| R. R \ U}) r" proof goal_cases case 1 have "zone_set (\ {Z \ R| R. R \ U}) r = \ {(zone_set (Z \ R) r) | R. R \ U}" unfolding zone_set_def by auto then show ?case by auto qed finally show "zone_set (Closure\<^sub>\ Z) r \ Closure\<^sub>\(zone_set Z r)" using U by simp qed lemma SuccI3: "R \ \ \ v \ R \ t \ 0 \ (v \ t) \ R' \ R' \ \ \ R' \ Succ \ R" apply (intro SuccI2[of \ X k, folded \_def, simplified]) apply assumption+ apply (intro region_unique[of \ X k, folded \_def, simplified, symmetric]) by assumption+ lemma closure_delay_mono: "Z \ V \ (Closure\<^sub>\ Z)\<^sup>\ \ Closure\<^sub>\ (Z\<^sup>\)" proof fix v assume v: "v \ (Closure\<^sub>\ Z)\<^sup>\" and Z: "Z \ V" then obtain u u' t R where A: "u \ Closure\<^sub>\ Z" "v = (u \ t)" "u \ R" "u' \ R" "R \ \" "u' \ Z" "t \ 0" unfolding cla_def zone_delay_def by blast from A(3,5) have "\ x \ X. u x \ 0" unfolding \_def by fastforce with region_cover_spec[of v] A(2,7) obtain R' where R': "R' \ \" "v \ R'" unfolding cval_add_def by auto with set_of_regions_spec[OF A(5,4), OF SuccI3, of u] A obtain t where t: "t \ 0" "[u' \ t]\<^sub>\ = R'" by auto with A have "(u' \ t) \ Z\<^sup>\" unfolding zone_delay_def by auto moreover from regions_closed'_spec[OF A(5,4)] t have "(u' \ t) \ R'" by auto ultimately have "R' \ (Z\<^sup>\) \ {}" by auto with R' show "v \ Closure\<^sub>\ (Z\<^sup>\)" unfolding cla_def by auto qed lemma region_V: "R \ \ \ R \ V" using V_def \_def region.cases by auto lemma closure_V: "Closure\<^sub>\ Z \ V" unfolding cla_def using region_V by auto lemma closure_V_int: "Closure\<^sub>\ Z = Closure\<^sub>\ (Z \ V)" unfolding cla_def using region_V by auto lemma closure_constraint_mono: "Closure\<^sub>\ g = g \ g \ (Closure\<^sub>\ Z) \ Closure\<^sub>\ (g \ Z)" unfolding cla_def by auto lemma closure_constraint_mono': assumes "Closure\<^sub>\ g = g \ V" shows "g \ (Closure\<^sub>\ Z) \ Closure\<^sub>\ (g \ Z)" proof - from assms closure_V_int have "Closure\<^sub>\ (g \ V) = g \ V" by auto from closure_constraint_mono[OF this, of Z] have "g \ (V \ Closure\<^sub>\ Z) \ Closure\<^sub>\ (g \ Z \ V)" by (metis Int_assoc Int_commute) with closure_V[of Z] closure_V_int[of "g \ Z"] show ?thesis by auto qed lemma cla_empty_iff: "Z \ V \ Z = {} \ Closure\<^sub>\ Z = {}" unfolding cla_def V_def using region_cover_spec by fast lemma closure_involutive_aux: "U \ \ \ Closure\<^sub>\ \ U = \ U" unfolding cla_def using valid_regions_distinct_spec by blast lemma closure_involutive_aux': "\ U. U \ \ \ Closure\<^sub>\ Z = \ U" unfolding cla_def by (rule exI[where x = "{R \ \. R \ Z \ {}}"]) auto lemma closure_involutive: "Closure\<^sub>\ Closure\<^sub>\ Z = Closure\<^sub>\ Z" using closure_involutive_aux closure_involutive_aux' by metis lemma closure_involutive': "Z \ Closure\<^sub>\ W \ Closure\<^sub>\ Z \ Closure\<^sub>\ W" unfolding cla_def using valid_regions_distinct_spec by fast lemma closure_subs: "Z \ V \ Z \ Closure\<^sub>\ Z" unfolding cla_def V_def using region_cover_spec by fast lemma cla_mono': "Z' \ V \ Z \ Z' \ Closure\<^sub>\ Z \ Closure\<^sub>\ Z'" by (meson closure_involutive' closure_subs subset_trans) lemma cla_mono: "Z \ Z' \ Closure\<^sub>\ Z \ Closure\<^sub>\ Z'" using closure_V_int cla_mono'[of "Z' \ V" "Z \ V"] by auto section \A New Zone Semantics Abstracting with \Closure\<^sub>\\\ subsection \Single step\ inductive step_z_alpha :: "('a, 'c, t, 's) ta \ 's \ ('c, t) zone \ 's \ ('c, t) zone \ bool" ("_ \ \_, _\ \\<^sub>\ \_, _\" [61,61,61] 61) where step_alpha: "A \ \l, Z\ \ \l', Z'\ \ A \ \l, Z\ \\<^sub>\ \l', Closure\<^sub>\ Z'\" inductive_cases[elim!]: "A \ \l, u\ \\<^sub>\ \l',u'\" declare step_z_alpha.intros[intro] lemma up_V: "Z \ V \ Z\<^sup>\ \ V" unfolding V_def zone_delay_def cval_add_def by auto lemma reset_V: "Z \ V \ (zone_set Z r) \ V" unfolding V_def unfolding zone_set_def by (induction r, auto) lemma step_z_V: "A \ \l, Z\ \ \l',Z'\ \ Z \ V \ Z' \ V" apply (induction rule: step_z.induct) apply (rule le_infI1) apply (rule up_V) apply blast apply (rule le_infI1) apply (rule reset_V) by blast text \Single-step soundness and completeness follows trivially from \cla_empty_iff\.\ lemma step_z_alpha_sound: "A \ \l, Z\ \\<^sub>\ \l',Z'\ \ Z \ V \ Z' \ {} \ \ Z''. A \ \l, Z\ \ \l',Z''\ \ Z'' \ {}" apply (induction rule: step_z_alpha.induct) apply (frule step_z_V) apply assumption apply (rotate_tac 3) apply (drule cla_empty_iff) by auto lemma step_z_alpha_complete: "A \ \l, Z\ \ \l',Z'\ \ Z \ V \ Z' \ {} \ \ Z''. A \ \l, Z\ \\<^sub>\ \l',Z''\ \ Z'' \ {}" apply (frule step_z_V) apply assumption apply (rotate_tac 3) apply (drule cla_empty_iff) by auto lemma zone_set_mono: "A \ B \ zone_set A r \ zone_set B r" unfolding zone_set_def by auto lemma zone_delay_mono: "A \ B \ A\<^sup>\ \ B\<^sup>\" unfolding zone_delay_def by auto subsection \Multi step\ inductive steps_z_alpha :: "('a, 'c, t, 's) ta \ 's \ ('c, t) zone \ 's \ ('c, t) zone \ bool" ("_ \ \_, _\ \\<^sub>\* \_, _\" [61,61,61] 61) where refl: "A \ \l, Z\ \\<^sub>\* \l, Z\" | step: "A \ \l, Z\ \\<^sub>\* \l', Z'\ \ A \ \l', Z'\ \\<^sub>\ \l'', Z''\ \ A \ \l, Z\ \\<^sub>\* \l'', Z''\" declare steps_z_alpha.intros[intro] lemma subset_int_mono: "A \ B \ A \ C \ B \ C" by blast text \P. Bouyer's calculation for @{term "Post(Closure\<^sub>\ Z, e) \ Closure\<^sub>\(Post (Z, e))"}\ text \This is now obsolete as we argue solely with monotonicty of \steps_z\ w.r.t \Closure\<^sub>\\\ lemma calc: "valid_abstraction A X k \ Z \ V \ A \ \l, Closure\<^sub>\ Z\ \ \l', Z'\ \ \Z''. A \ \l, Z\ \\<^sub>\ \l', Z''\ \ Z' \ Z''" proof (cases rule: step_z.cases, assumption, goal_cases) case 1 note A = this from A(1) have "\(x, m)\clkp_set A. m \ real (k x) \ x \ X \ m \ \" by (fastforce elim: valid_abstraction.cases) then have "\(x, m)\collect_clock_pairs (inv_of A l). m \ real (k x) \ x \ X \ m \ \" unfolding clkp_set_def collect_clki_def inv_of_def by auto from closure_constraint_id[OF this] have *: "Closure\<^sub>\ \inv_of A l\ = \inv_of A l\ \ V" . from closure_constraint_mono'[OF *, of Z] have "(Closure\<^sub>\ Z) \ {u. u \ inv_of A l} \ Closure\<^sub>\ (Z \ {u. u \ inv_of A l})" unfolding ccval_def by (subst Int_commute) (subst (asm) (2) Int_commute, assumption) moreover have "\\<^sup>\ \ Closure\<^sub>\ ((Z \ {u. u \ inv_of A l})\<^sup>\)" using A(2) by (blast intro!: closure_delay_mono) ultimately have "Z' \ Closure\<^sub>\ ((Z \ {u. u \ inv_of A l})\<^sup>\ \ {u. u \ inv_of A l})" using closure_constraint_mono'[OF *, of "(Z \ {u. u \ inv_of A l})\<^sup>\"] unfolding ccval_def apply (subst A(5)) apply (subst (asm) (5 7) Int_commute) apply (rule subset_trans) defer apply assumption apply (subst subset_int_mono) defer apply rule apply (rule subset_trans) defer apply assumption apply (rule zone_delay_mono) apply assumption done with A(4,3) show ?thesis by auto next case (2 g a r) note A = this from A(1) have *: "\(x, m)\clkp_set A. m \ real (k x) \ x \ X \ m \ \" "collect_clkvt (trans_of A) \ X" "finite X" by (auto elim: valid_abstraction.cases) from *(1) A(5) have "\(x, m)\collect_clock_pairs (inv_of A l'). m \ real (k x) \ x \ X \ m \ \" unfolding clkp_set_def collect_clki_def inv_of_def by fastforce from closure_constraint_id[OF this] have **: "Closure\<^sub>\ \inv_of A l'\ = \inv_of A l'\ \ V" . from *(1) A(5) have "\(x, m)\collect_clock_pairs g. m \ real (k x) \ x \ X \ m \ \" unfolding clkp_set_def collect_clkt_def by fastforce from closure_constraint_id[OF this] have ***: "Closure\<^sub>\ \g\ = \g\ \ V" . from *(2) A(5) have ****: "set r \ X" unfolding collect_clkvt_def by fastforce from closure_constraint_mono'[OF ***, of Z] have "(Closure\<^sub>\ Z) \ {u. u \ g} \ Closure\<^sub>\ (Z \ {u. u \ g})" unfolding ccval_def by (subst Int_commute) (subst (asm) (2) Int_commute, assumption) moreover have "zone_set \ r \ Closure\<^sub>\ (zone_set (Z \ {u. u \ g}) r)" using **** A(2) by (intro closure_update_mono, auto) ultimately have "Z' \ Closure\<^sub>\ (zone_set (Z \ {u. u \ g}) r \ {u. u \ inv_of A l'})" using closure_constraint_mono'[OF **, of "zone_set (Z \ {u. u \ g}) r"] unfolding ccval_def apply (subst A(4)) apply (subst (asm) (5 7) Int_commute) apply (rule subset_trans) defer apply assumption apply (subst subset_int_mono) defer apply rule apply (rule subset_trans) defer apply assumption apply (rule zone_set_mono) apply assumption done with A(5) show ?thesis by auto qed text \ Turning P. Bouyers argument for multiple steps into an inductive proof is not direct. With this initial argument we can get to a point where the induction hypothesis is applicable. This breaks the "information hiding" induced by the different variants of steps. \ lemma steps_z_alpha_closure_involutive'_aux: "A \ \l, Z\ \ \l',Z'\ \ Closure\<^sub>\ Z \ Closure\<^sub>\ W \ valid_abstraction A X k \ Z \ V \ \ W'. A \ \l, W\ \ \l',W'\ \ Closure\<^sub>\ Z' \ Closure\<^sub>\ W'" proof (induction rule: step_z.induct) case A: (step_t_z A l Z) let ?Z' = "(Z \ {u. u \ inv_of A l})\<^sup>\ \ {u. u \ inv_of A l}" let ?W' = "(W \ {u. u \ inv_of A l})\<^sup>\ \ {u. u \ inv_of A l}" from \_def have \_def': "\ = {region X I r |I r. valid_region X k I r}" by simp have step_z: "A \ \l, W\ \ \l,?W'\" by auto moreover have "Closure\<^sub>\ ?Z' \ Closure\<^sub>\ ?W'" proof fix v assume v: "v \ Closure\<^sub>\ ?Z'" then obtain R' v' where 1: "R' \ \" "v \ R'" "v' \ R'" "v' \ ?Z'" unfolding cla_def by auto then obtain u d where "u \ Z" and v': "v' = u \ d" "u \ inv_of A l" "u \ d \ inv_of A l" "0 \ d" unfolding zone_delay_def by blast with closure_subs[OF A(3)] A(1) obtain u' R where u': "u' \ W" "u \ R" "u' \ R" "R \ \" unfolding cla_def by blast then have "\x\X. 0 \ u x" unfolding \_def by fastforce from region_cover'[OF \_def' this] have R: "[u]\<^sub>\ \ \" "u \ [u]\<^sub>\" by auto from SuccI2[OF \_def' this(2,1) v'(4), of "[v']\<^sub>\"] v'(1) have v'1: "[v']\<^sub>\ \ Succ \ ([u]\<^sub>\)" "[v']\<^sub>\ \ \" by auto from regions_closed'_spec[OF R(1,2) v'(4)] v'(1) have v'2: "v' \ [v']\<^sub>\" by simp from A(2) have *: "\(x, m)\clkp_set A. m \ real (k x) \ x \ X \ m \ \" "collect_clkvt (trans_of A) \ X" "finite X" by (auto elim: valid_abstraction.cases) from *(1) u'(2) have "\(x, m)\collect_clock_pairs (inv_of A l). m \ real (k x) \ x \ X \ m \ \" unfolding clkp_set_def collect_clki_def inv_of_def by fastforce from ccompatible[OF this, folded \_def'] v'1(2) v'2 v'(1,2,3) R have 3: "[v']\<^sub>\ \ \inv_of A l\" "([u]\<^sub>\) \ \inv_of A l\" unfolding ccompatible_def ccval_def by auto with A v'1 R(1) \_def' have "A,\ \ \l, ([u]\<^sub>\)\ \ \l,([v']\<^sub>\)\" by auto with valid_regions_distinct_spec[OF v'1(2) 1(1) v'2 1(3)] region_unique_spec[OF u'(2,4)] have step_r: "A,\ \ \l, R\ \ \l, R'\" and 2: "[v']\<^sub>\ = R'" "[u]\<^sub>\ = R" by auto from set_of_regions_spec[OF u'(4,3)] v'1(1) 2 obtain t where t: "t \ 0" "[u' \ t]\<^sub>\ = R'" by auto with regions_closed'_spec[OF u'(4,3) this(1)] step_t_r(1) have *: "u' \ t \ R'" by auto with t(1) 3 2 u'(1,3) have "A \ \l, u'\ \ \l, u' \ t\" "u' \ t \ ?W'" unfolding zone_delay_def ccval_def by auto with * 1(1) have "R' \ Closure\<^sub>\ ?W'" unfolding cla_def by auto with 1(2) show "v \ Closure\<^sub>\ ?W'" .. qed ultimately show ?case by auto next case A: (step_a_z A l g a r l' Z) let ?Z' = "zone_set (Z \ {u. u \ g}) r \ {u. u \ inv_of A l'}" let ?W' = "zone_set (W \ {u. u \ g}) r \ {u. u \ inv_of A l'}" from \_def have \_def': "\ = {region X I r |I r. valid_region X k I r}" by simp from A(1) have step_z: "A \ \l, W\ \ \l',?W'\" by auto moreover have "Closure\<^sub>\ ?Z' \ Closure\<^sub>\ ?W'" proof fix v assume v: "v \ Closure\<^sub>\ ?Z'" then obtain R' v' where 1: "R' \ \" "v \ R'" "v' \ R'" "v' \ ?Z'" unfolding cla_def by auto then obtain u where "u \ Z" and v': "v' = [r\0]u" "u \ g" "v' \ inv_of A l'" unfolding zone_set_def by blast let ?R'= "region_set' (([u]\<^sub>\) \ {u. u \ g}) r 0 \ {u. u \ inv_of A l'}" from \u \ Z\ closure_subs[OF A(4)] A(2) obtain u' R where u': "u' \ W" "u \ R" "u' \ R" "R \ \" unfolding cla_def by blast then have "\x\X. 0 \ u x" unfolding \_def by fastforce from region_cover'[OF \_def' this] have R: "[u]\<^sub>\ \ \" "u \ [u]\<^sub>\" by auto from step_r_complete_aux[OF \_def' A(3) this(2,1) A(1) v'(2)] v' have *: "[u]\<^sub>\ = ([u]\<^sub>\) \ {u. u \ g}" "?R' = region_set' ([u]\<^sub>\) r 0" "?R' \ \" by auto from \_def' A(3) have "collect_clkvt (trans_of A) \ X" "finite X" by (auto elim: valid_abstraction.cases) with A(1) have r: "set r \ X" unfolding collect_clkvt_def by fastforce from * v'(1) R(2) have "v' \ ?R'" unfolding region_set'_def by auto moreover have "A,\ \ \l,([u]\<^sub>\)\ \ \l',?R'\" using R(1) \_def' A(1,3) v'(2) by auto thm valid_regions_distinct_spec with valid_regions_distinct_spec[OF *(3) 1(1) \v' \ ?R'\ 1(3)] region_unique_spec[OF u'(2,4)] have 2: "?R' = R'" "[u]\<^sub>\ = R" by auto with * u' have *: "[r\0]u' \ ?R'" "u' \ g" "[r\0]u' \ inv_of A l'" unfolding region_set'_def by auto with A(1) have "A \ \l, u'\ \ \l',[r\0]u'\" apply (intro step.intros(1)) apply rule by auto moreover from * u'(1) have "[r\0]u' \ ?W'" unfolding zone_set_def by auto ultimately have "R' \ Closure\<^sub>\ ?W'" using *(1) 1(1) 2(1) unfolding cla_def by auto with 1(2) show "v \ Closure\<^sub>\ ?W'" .. qed ultimately show ?case by meson qed lemma steps_z_alpha_closure_involutive'_aux': "A \ \l, Z\ \ \l',Z'\ \ Closure\<^sub>\ Z \ Closure\<^sub>\ W \ valid_abstraction A X k \ Z \ V \ W \ Z \ \ W'. A \ \l, W\ \ \l',W'\ \ Closure\<^sub>\ Z' \ Closure\<^sub>\ W' \ W' \ Z'" proof (induction rule: step_z.induct) case A: (step_t_z A l Z) let ?Z' = "(Z \ {u. u \ inv_of A l})\<^sup>\ \ {u. u \ inv_of A l}" let ?W' = "(W \ {u. u \ inv_of A l})\<^sup>\ \ {u. u \ inv_of A l}" from \_def have \_def': "\ = {region X I r |I r. valid_region X k I r}" by simp have step_z: "A \ \l, W\ \ \l,?W'\" by auto moreover have "Closure\<^sub>\ ?Z' \ Closure\<^sub>\ ?W'" proof fix v assume v: "v \ Closure\<^sub>\ ?Z'" then obtain R' v' where 1: "R' \ \" "v \ R'" "v' \ R'" "v' \ ?Z'" unfolding cla_def by auto then obtain u d where "u \ Z" and v': "v' = u \ d" "u \ inv_of A l" "u \ d \ inv_of A l" "0 \ d" unfolding zone_delay_def by blast with closure_subs[OF A(3)] A(1) obtain u' R where u': "u' \ W" "u \ R" "u' \ R" "R \ \" unfolding cla_def by blast then have "\x\X. 0 \ u x" unfolding \_def by fastforce from region_cover'[OF \_def' this] have R: "[u]\<^sub>\ \ \" "u \ [u]\<^sub>\" by auto from SuccI2[OF \_def' this(2,1) v'(4), of "[v']\<^sub>\"] v'(1) have v'1: "[v']\<^sub>\ \ Succ \ ([u]\<^sub>\)" "[v']\<^sub>\ \ \" by auto from regions_closed'_spec[OF R(1,2) v'(4)] v'(1) have v'2: "v' \ [v']\<^sub>\" by simp from A(2) have *: "\(x, m)\clkp_set A. m \ real (k x) \ x \ X \ m \ \" "collect_clkvt (trans_of A) \ X" "finite X" by (auto elim: valid_abstraction.cases) from *(1) u'(2) have "\(x, m)\collect_clock_pairs (inv_of A l). m \ real (k x) \ x \ X \ m \ \" unfolding clkp_set_def collect_clki_def inv_of_def by fastforce from ccompatible[OF this, folded \_def'] v'1(2) v'2 v'(1,2,3) R have 3: "[v']\<^sub>\ \ \inv_of A l\" "([u]\<^sub>\) \ \inv_of A l\" unfolding ccompatible_def ccval_def by auto with A v'1 R(1) \_def' have "A,\ \ \l, ([u]\<^sub>\)\ \ \l,([v']\<^sub>\)\" by auto with valid_regions_distinct_spec[OF v'1(2) 1(1) v'2 1(3)] region_unique_spec[OF u'(2,4)] have step_r: "A,\ \ \l, R\ \ \l, R'\" and 2: "[v']\<^sub>\ = R'" "[u]\<^sub>\ = R" by auto from set_of_regions_spec[OF u'(4,3)] v'1(1) 2 obtain t where t: "t \ 0" "[u' \ t]\<^sub>\ = R'" by auto with regions_closed'_spec[OF u'(4,3) this(1)] step_t_r(1) have *: "u' \ t \ R'" by auto with t(1) 3 2 u'(1,3) have "A \ \l, u'\ \ \l, u' \ t\" "u' \ t \ ?W'" unfolding zone_delay_def ccval_def by auto with * 1(1) have "R' \ Closure\<^sub>\ ?W'" unfolding cla_def by auto with 1(2) show "v \ Closure\<^sub>\ ?W'" .. qed moreover have "?W' \ ?Z'" using \W \ Z\ unfolding zone_delay_def by auto ultimately show ?case by auto next case A: (step_a_z A l g a r l' Z) let ?Z' = "zone_set (Z \ {u. u \ g}) r \ {u. u \ inv_of A l'}" let ?W' = "zone_set (W \ {u. u \ g}) r \ {u. u \ inv_of A l'}" from \_def have \_def': "\ = {region X I r |I r. valid_region X k I r}" by simp from A(1) have step_z: "A \ \l, W\ \ \l',?W'\" by auto moreover have "Closure\<^sub>\ ?Z' \ Closure\<^sub>\ ?W'" proof fix v assume v: "v \ Closure\<^sub>\ ?Z'" then obtain R' v' where 1: "R' \ \" "v \ R'" "v' \ R'" "v' \ ?Z'" unfolding cla_def by auto then obtain u where "u \ Z" and v': "v' = [r\0]u" "u \ g" "v' \ inv_of A l'" unfolding zone_set_def by blast let ?R'= "region_set' (([u]\<^sub>\) \ {u. u \ g}) r 0 \ {u. u \ inv_of A l'}" from \u \ Z\ closure_subs[OF A(4)] A(2) obtain u' R where u': "u' \ W" "u \ R" "u' \ R" "R \ \" unfolding cla_def by blast then have "\x\X. 0 \ u x" unfolding \_def by fastforce from region_cover'[OF \_def' this] have R: "[u]\<^sub>\ \ \" "u \ [u]\<^sub>\" by auto from step_r_complete_aux[OF \_def' A(3) this(2,1) A(1) v'(2)] v' have *: "[u]\<^sub>\ = ([u]\<^sub>\) \ {u. u \ g}" "?R' = region_set' ([u]\<^sub>\) r 0" "?R' \ \" by auto from \_def' A(3) have "collect_clkvt (trans_of A) \ X" "finite X" by (auto elim: valid_abstraction.cases) with A(1) have r: "set r \ X" unfolding collect_clkvt_def by fastforce from * v'(1) R(2) have "v' \ ?R'" unfolding region_set'_def by auto moreover have "A,\ \ \l,([u]\<^sub>\)\ \ \l',?R'\" using R(1) \_def' A(1,3) v'(2) by auto thm valid_regions_distinct_spec with valid_regions_distinct_spec[OF *(3) 1(1) \v' \ ?R'\ 1(3)] region_unique_spec[OF u'(2,4)] have 2: "?R' = R'" "[u]\<^sub>\ = R" by auto with * u' have *: "[r\0]u' \ ?R'" "u' \ g" "[r\0]u' \ inv_of A l'" unfolding region_set'_def by auto with A(1) have "A \ \l, u'\ \ \l',[r\0]u'\" apply (intro step.intros(1)) apply rule by auto moreover from * u'(1) have "[r\0]u' \ ?W'" unfolding zone_set_def by auto ultimately have "R' \ Closure\<^sub>\ ?W'" using *(1) 1(1) 2(1) unfolding cla_def by auto with 1(2) show "v \ Closure\<^sub>\ ?W'" .. qed moreover have "?W' \ ?Z'" using \W \ Z\ unfolding zone_set_def by auto ultimately show ?case by meson qed lemma steps_z_alt: "A \ \l, Z\ \* \l',Z'\ \ A \ \l', Z'\ \ \l'',Z''\ \ A \ \l, Z\ \* \l'',Z''\" by (induction rule: steps_z.induct) auto lemma steps_z_alpha_V: "A \ \l, Z\ \\<^sub>\* \l',Z'\ \ Z \ V \ Z' \ V" apply (induction rule: steps_z_alpha.induct) using closure_V by auto lemma steps_z_alpha_closure_involutive': "A \ \l, Z\ \\<^sub>\* \l',Z'\ \ A \ \l', Z'\ \ \l'',Z''\ \ valid_abstraction A X k \ Z \ V \ \ Z'''. A \ \l, Z\ \* \l'',Z'''\ \ Closure\<^sub>\ Z'' \ Closure\<^sub>\ Z''' \ Z''' \ Z''" proof (induction A l Z l' Z' arbitrary: Z'' l'' rule: steps_z_alpha.induct, goal_cases) case refl from this(1) show ?case by blast next case A: (2 A l Z l' Z' l'' Z'' Z''a l''a) from A(3) obtain \ where Z'': "Z'' = Closure\<^sub>\ \" "A \ \l', Z'\ \ \l'',\\" by auto from A(2)[OF Z''(2) A(5,6)] obtain Z''' where Z''': "A \ \l, Z\ \* \l'',Z'''\" "Closure\<^sub>\ \ \ Closure\<^sub>\ Z'''" "Z''' \ \" by auto from closure_subs have *: "Z''' \ Closure\<^sub>\ \" by (metis A(1,6) Z'''(3) Z''(2) step_z_V steps_z_alpha_V subset_trans) from A(4) Z'' have "A \ \l'', Closure\<^sub>\ \\ \ \l''a,Z''a\" by auto from steps_z_alpha_closure_involutive'_aux'[OF this(1) _ A(5) closure_V *] Z'''(2,3) obtain W' where ***: "A \ \l'', Z'''\ \ \l''a,W'\" "Closure\<^sub>\ Z''a \ Closure\<^sub>\ W'" "W' \ Z''a" by (auto simp: closure_involutive) with Z'''(1) have "A \ \l, Z\ \* \l''a,W'\" by (blast intro: steps_z_alt) with ***(2,3) show ?case by blast qed text \Old proof using Bouyer's calculation\ lemma steps_z_alpha_closure_involutive'': "A \ \l, Z\ \\<^sub>\* \l',Z'\ \ A \ \l', Z'\ \ \l'',Z''\ \ valid_abstraction A X k \ Z \ V \ \ Z'''. A \ \l, Z\ \* \l'',Z'''\ \ Closure\<^sub>\ Z'' \ Closure\<^sub>\ Z'''" proof (induction A l Z l' Z' arbitrary: Z'' l'' rule: steps_z_alpha.induct, goal_cases) case refl from this(1) show ?case by blast next case A: (2 A l Z l' Z' l'' Z'' Z''a l''a) from A(3) obtain \ where Z'': "Z'' = Closure\<^sub>\ \" "A \ \l', Z'\ \ \l'',\\" by auto from A(2)[OF Z''(2) A(5,6)] obtain Z''' where Z''': "A \ \l, Z\ \* \l'',Z'''\" "Closure\<^sub>\ \ \ Closure\<^sub>\ Z'''" by auto from steps_z_alpha_V[OF A(1,6)] step_z_V[OF Z''(2)] have *: "\ \ V" by blast from A Z'' have "A \ \l'', Closure\<^sub>\ \\ \ \l''a,Z''a\" by auto from calc[OF A(5) * this] obtain \' where **: "A \ \l'', \\ \ \l''a,\'\" "Z''a \ Closure\<^sub>\ \'" by auto from steps_z_alpha_closure_involutive'_aux[OF this(1) Z'''(2) A(5) *] obtain W' where ***: "A \ \l'', Z'''\ \ \l''a,W'\" "Closure\<^sub>\ \' \ Closure\<^sub>\ W'" by auto with Z'''(1) have "A \ \l, Z\ \* \l''a,W'\" by (blast intro: steps_z_alt) with closure_involutive'[OF **(2)] ***(2) show ?case by blast qed lemma steps_z_alpha_closure_involutive: "A \ \l, Z\ \\<^sub>\* \l',Z'\ \ valid_abstraction A X k \ Z \ V \ \ Z''. A \ \l, Z\ \* \l',Z''\ \ Closure\<^sub>\ Z' \ Closure\<^sub>\ Z'' \ Z'' \ Z'" proof (induction A l Z l' Z' rule: steps_z_alpha.induct) case refl show ?case by blast next case 2: (step A l Z l' Z' l'' Z'') then obtain Z''a where *: "A \ \l', Z'\ \ \l'',Z''a\" "Z'' = Closure\<^sub>\ Z''a" by auto from steps_z_alpha_closure_involutive'[OF 2(1) this(1) 2(4,5)] obtain Z''' where Z''': "A \ \l, Z\ \* \l'',Z'''\" "Closure\<^sub>\ Z''a \ Closure\<^sub>\ Z'''" "Z''' \ Z''a" by blast have "Z''' \ Z''" by (metis *(1,2) 2(1,5) Z'''(3) closure_subs step_z_V steps_z_alpha_V subset_trans) with * closure_involutive Z''' show ?case by auto qed lemma steps_z_V: "A \ \l, Z\ \* \l',Z'\ \ Z \ V \ Z' \ V" apply (induction rule: steps_z.induct) apply blast using step_z_V by metis lemma steps_z_alpha_sound: "A \ \l, Z\ \\<^sub>\* \l',Z'\ \ valid_abstraction A X k \ Z \ V \ Z' \ {} \ \ Z''. A \ \l, Z\ \* \l',Z''\ \ Z'' \ {} \ Z'' \ Z'" proof goal_cases case 1 from steps_z_alpha_closure_involutive[OF 1(1-3)] obtain Z'' where Z'': "A \ \l, Z\ \* \l',Z''\" "Closure\<^sub>\ Z' \ Closure\<^sub>\ Z''" "Z'' \ Z'" by blast with 1(4) cla_empty_iff[OF steps_z_alpha_V[OF 1(1)], OF 1(3)] cla_empty_iff[OF steps_z_V, OF this(1) 1(3)] have "Z'' \ {}" by auto with Z'' show ?case by auto qed lemma step_z_mono: "A \ \l, Z\ \ \l',Z'\ \ Z \ W \ \ W'. A \ \l, W\ \ \l',W'\ \ Z' \ W'" proof (cases rule: step_z.cases, assumption, goal_cases) case A: 1 let ?W' = "(W \ {u. u \ inv_of A l})\<^sup>\ \ {u. u \ inv_of A l}" from A have "A \ \l, W\ \ \l',?W'\" by auto moreover have "Z' \ ?W'" apply (subst A(4)) apply (rule subset_int_mono) apply (rule zone_delay_mono) apply (rule subset_int_mono) apply (rule A(2)) done ultimately show ?thesis by auto next case A: (2 g a r) let ?W' = "zone_set (W \ {u. u \ g}) r \ {u. u \ inv_of A l'}" from A have "A \ \l, W\ \ \l',?W'\" by auto moreover have "Z' \ ?W'" apply (subst A(3)) apply (rule subset_int_mono) apply (rule zone_set_mono) apply (rule subset_int_mono) apply (rule A(2)) done ultimately show ?thesis by auto qed lemma step_z_alpha_mono: "A \ \l, Z\ \\<^sub>\ \l',Z'\ \ Z \ W \ W \ V \ \ W'. A \ \l, W\ \\<^sub>\ \l',W'\ \ Z' \ W'" proof goal_cases case 1 then obtain Z'' where *: "A \ \l, Z\ \ \l',Z''\" "Z' = Closure\<^sub>\ Z''" by auto from step_z_mono[OF this(1) 1(2)] obtain W' where "A \ \l, W\ \ \l',W'\" "Z'' \ W'" by auto moreover with *(2) have "Z' \ Closure\<^sub>\ W'" unfolding cla_def by auto ultimately show ?case by blast qed lemma steps_z_alpha_mono: "A \ \l, Z\ \\<^sub>\* \l',Z'\ \ Z \ W \ W \ V \ \ W'. A \ \l, W\ \\<^sub>\* \l',W'\ \ Z' \ W'" proof (induction rule: steps_z_alpha.induct, goal_cases) case refl then show ?case by auto next case (2 A l Z l' Z' l'' Z'') then obtain W' where "A \ \l, W\ \\<^sub>\* \l',W'\" "Z' \ W'" by auto with step_z_alpha_mono[OF 2(3) this(2) steps_z_alpha_V[OF this(1) 2(5)]] show ?case by blast qed lemma steps_z_alpha_alt: "A \ \l, Z\ \\<^sub>\ \l', Z'\ \ A \ \l', Z'\ \\<^sub>\* \l'', Z''\ \ A \ \l, Z\ \\<^sub>\* \l'', Z''\" by (rotate_tac, induction rule: steps_z_alpha.induct) blast+ lemma steps_z_alpha_complete: "A \ \l, Z\ \* \l',Z'\ \ valid_abstraction A X k \ Z \ V \ Z' \ {} \ \ Z''. A \ \l, Z\ \\<^sub>\* \l',Z''\ \ Z' \ Z''" proof (induction rule: steps_z.induct, goal_cases) case refl with cla_empty_iff show ?case by blast next case (2 A l Z l' Z' l'' Z'') with step_z_V[OF this(1,5)] obtain Z''' where "A \ \l', Z'\ \\<^sub>\* \l'',Z'''\" "Z'' \ Z'''" by blast with steps_z_alpha_mono[OF this(1) closure_subs[OF step_z_V[OF 2(1,5)]] closure_V] - obtain W' where "A \ \l', Closure\<^sub>\ Z'\ \\<^sub>\* \l'',W'\" " Z'' \ W'" by auto - moreover with 2(1) have "A \ \l, Z\ \\<^sub>\* \l'',W'\" by (auto intro: steps_z_alpha_alt) - ultimately show ?case by auto + obtain W' where W': "A \ \l', Closure\<^sub>\ Z'\ \\<^sub>\* \l'',W'\" " Z'' \ W'" by auto + with 2(1) have "A \ \l, Z\ \\<^sub>\* \l'',W'\" by (auto intro: steps_z_alpha_alt) + with W' show ?case by auto qed lemma steps_z_alpha_complete': "A \ \l, Z\ \* \l',Z'\ \ valid_abstraction A X k \ Z \ V \ Z' \ {} \ \ Z''. A \ \l, Z\ \\<^sub>\* \l',Z''\ \ Z'' \ {}" using steps_z_alpha_complete by fast end end diff --git a/thys/Types_To_Sets_Extension/ETTS/ETTS_Tools/More_Logic.ML b/thys/Types_To_Sets_Extension/ETTS/ETTS_Tools/More_Logic.ML --- a/thys/Types_To_Sets_Extension/ETTS/ETTS_Tools/More_Logic.ML +++ b/thys/Types_To_Sets_Extension/ETTS/ETTS_Tools/More_Logic.ML @@ -1,44 +1,44 @@ (* Title: ETTS/ETTS_Tools/More_Logic.ML Author: Mihails Milehins Copyright 2021 (C) Mihails Milehins An extension of the structure Logic from the standard library of Isabelle/Pure. *) signature LOGIC = sig include LOGIC val forall_elim_all : term -> term * (string * typ) list val get_forall_ftv_permute : term -> term * ((string * typ) list * int list) end structure Logic: LOGIC = struct open Logic; (*forall elimination*) fun forall_elim_all t = let - fun forall_elim_all_impl t ftv_specs = - let val (ftv_spec, t) = Logic.dest_all t - in forall_elim_all_impl t (ftv_spec::ftv_specs) end - handle TERM ("dest_all", _) => (t, ftv_specs) + fun forall_elim_all_impl t ftv_specs = + (case \<^try>\Logic.dest_all_global t\ of + SOME (ftv_spec, t) => forall_elim_all_impl t (ftv_spec::ftv_specs) + | NONE => (t, ftv_specs)) in forall_elim_all_impl t [] ||> rev end; (*indices of the universally quantified variables with respect to the order of their appearance in the term in the sense of Term.add_frees*) fun get_forall_ftv_permute t = let val (t', forall_ftv_specs) = forall_elim_all t val ftv_specs = Term.add_frees t' [] |> rev val call_ftv_specs = ftv_specs |> subtract op= (ftv_specs |> subtract op= forall_ftv_specs) val index_of_ftv = (call_ftv_specs ~~ (0 upto (length call_ftv_specs - 1))) |> AList.lookup op= #> the val forall_ftv_permute = map index_of_ftv forall_ftv_specs in (t', (forall_ftv_specs, forall_ftv_permute)) end; end; \ No newline at end of file