diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/IMP/Abs_Int1.thy --- a/src/HOL/IMP/Abs_Int1.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/IMP/Abs_Int1.thy Fri Sep 20 19:51:08 2024 +0200 @@ -105,19 +105,19 @@ assumes h: "m x \ h" begin -definition m_s :: "'av st \ vname set \ nat" ("m\<^sub>s") where +definition m_s :: "'av st \ vname set \ nat" (\m\<^sub>s\) where "m_s S X = (\ x \ X. m(fun S x))" lemma m_s_h: "finite X \ m_s S X \ h * card X" by(simp add: m_s_def) (metis mult.commute of_nat_id sum_bounded_above[OF h]) -definition m_o :: "'av st option \ vname set \ nat" ("m\<^sub>o") where +definition m_o :: "'av st option \ vname set \ nat" (\m\<^sub>o\) where "m_o opt X = (case opt of None \ h * card X + 1 | Some S \ m_s S X)" lemma m_o_h: "finite X \ m_o opt X \ (h*card X + 1)" by(auto simp add: m_o_def m_s_h le_SucI split: option.split dest:m_s_h) -definition m_c :: "'av st option acom \ nat" ("m\<^sub>c") where +definition m_c :: "'av st option acom \ nat" (\m\<^sub>c\) where "m_c C = sum_list (map (\a. m_o a (vars C)) (annos C))" text\Upper complexity bound:\ @@ -134,14 +134,14 @@ end -fun top_on_st :: "'a::order_top st \ vname set \ bool" ("top'_on\<^sub>s") where +fun top_on_st :: "'a::order_top st \ vname set \ bool" (\top'_on\<^sub>s\) where "top_on_st S X = (\x\X. fun S x = \)" -fun top_on_opt :: "'a::order_top st option \ vname set \ bool" ("top'_on\<^sub>o") where +fun top_on_opt :: "'a::order_top st option \ vname set \ bool" (\top'_on\<^sub>o\) where "top_on_opt (Some S) X = top_on_st S X" | "top_on_opt None X = True" -definition top_on_acom :: "'a::order_top st option acom \ vname set \ bool" ("top'_on\<^sub>c") where +definition top_on_acom :: "'a::order_top st option acom \ vname set \ bool" (\top'_on\<^sub>c\) where "top_on_acom C X = (\a \ set(annos C). top_on_opt a X)" lemma top_on_top: "top_on_opt (\::_ st option) X"