diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/IMP/Abs_Int0.thy --- a/src/HOL/IMP/Abs_Int0.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/IMP/Abs_Int0.thy Fri Sep 20 19:51:08 2024 +0200 @@ -317,20 +317,20 @@ 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(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]) -fun m_o :: "'av st option \ vname set \ nat" ("m\<^sub>o") where +fun m_o :: "'av st option \ vname set \ nat" (\m\<^sub>o\) where "m_o (Some S) X = m_s S X" | "m_o None X = h * card X + 1" lemma m_o_h: "finite X \ m_o opt X \ (h*card X + 1)" by(cases opt)(auto simp add: m_s_h le_SucI 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:\ @@ -359,14 +359,14 @@ the finitely many variables in the program change. That the others do not change follows because they remain \<^term>\\\.\ -fun top_on_st :: "'av st \ vname set \ bool" ("top'_on\<^sub>s") where +fun top_on_st :: "'av st \ vname set \ bool" (\top'_on\<^sub>s\) where "top_on_st S X = (\x\X. S x = \)" -fun top_on_opt :: "'av st option \ vname set \ bool" ("top'_on\<^sub>o") where +fun top_on_opt :: "'av 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 :: "'av st option acom \ vname set \ bool" ("top'_on\<^sub>c") where +definition top_on_acom :: "'av 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 \ X"