diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/IMP/Abs_Int3.thy --- a/src/HOL/IMP/Abs_Int3.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/IMP/Abs_Int3.thy Fri Sep 20 19:51:08 2024 +0200 @@ -7,10 +7,10 @@ begin class widen = -fixes widen :: "'a \ 'a \ 'a" (infix "\" 65) +fixes widen :: "'a \ 'a \ 'a" (infix \\\ 65) class narrow = -fixes narrow :: "'a \ 'a \ 'a" (infix "\" 65) +fixes narrow :: "'a \ 'a \ 'a" (infix \\\ 65) class wn = widen + narrow + order + assumes widen1: "x \ x \ y" @@ -398,7 +398,7 @@ done -definition n_s :: "'av st \ vname set \ nat" ("n\<^sub>s") where +definition n_s :: "'av st \ vname set \ nat" (\n\<^sub>s\) where "n\<^sub>s S X = (\x\X. n(fun S x))" lemma n_s_narrow_rep: @@ -422,7 +422,7 @@ apply(auto simp add: less_eq_st_rep_iff eq_st_def fun_eq_iff n_s_narrow_rep) done -definition n_o :: "'av st option \ vname set \ nat" ("n\<^sub>o") where +definition n_o :: "'av st option \ vname set \ nat" (\n\<^sub>o\) where "n\<^sub>o opt X = (case opt of None \ 0 | Some S \ n\<^sub>s S X + 1)" lemma n_o_narrow: @@ -433,7 +433,7 @@ done -definition n_c :: "'av st option acom \ nat" ("n\<^sub>c") where +definition n_c :: "'av st option acom \ nat" (\n\<^sub>c\) where "n\<^sub>c C = sum_list (map (\a. n\<^sub>o a (vars C)) (annos C))" lemma less_annos_iff: "(C1 < C2) = (C1 \ C2 \