diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Hoare_Parallel/RG_Tran.thy --- a/src/HOL/Hoare_Parallel/RG_Tran.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Hoare_Parallel/RG_Tran.thy Fri Sep 20 19:51:08 2024 +0200 @@ -12,7 +12,7 @@ inductive_set etran :: "('a conf \ 'a conf) set" - and etran' :: "'a conf \ 'a conf \ bool" ("_ -e\ _" [81,81] 80) + and etran' :: "'a conf \ 'a conf \ bool" (\_ -e\ _\ [81,81] 80) where "P -e\ Q \ (P,Q) \ etran" | Env: "(P, s) -e\ (P, t)" @@ -24,8 +24,8 @@ inductive_set ctran :: "('a conf \ 'a conf) set" - and ctran' :: "'a conf \ 'a conf \ bool" ("_ -c\ _" [81,81] 80) - and ctrans :: "'a conf \ 'a conf \ bool" ("_ -c*\ _" [81,81] 80) + and ctran' :: "'a conf \ 'a conf \ bool" (\_ -c\ _\ [81,81] 80) + and ctrans :: "'a conf \ 'a conf \ bool" (\_ -c*\ _\ [81,81] 80) where "P -c\ Q \ (P,Q) \ ctran" | "P -c*\ Q \ (P,Q) \ ctran\<^sup>*" @@ -52,14 +52,14 @@ inductive_set par_etran :: "('a par_conf \ 'a par_conf) set" - and par_etran' :: "['a par_conf,'a par_conf] \ bool" ("_ -pe\ _" [81,81] 80) + and par_etran' :: "['a par_conf,'a par_conf] \ bool" (\_ -pe\ _\ [81,81] 80) where "P -pe\ Q \ (P,Q) \ par_etran" | ParEnv: "(Ps, s) -pe\ (Ps, t)" inductive_set par_ctran :: "('a par_conf \ 'a par_conf) set" - and par_ctran' :: "['a par_conf,'a par_conf] \ bool" ("_ -pc\ _" [81,81] 80) + and par_ctran' :: "['a par_conf,'a par_conf] \ bool" (\_ -pc\ _\ [81,81] 80) where "P -pc\ Q \ (P,Q) \ par_ctran" | ParComp: "\i (r, t)\ \ (Ps, s) -pc\ (Ps[i:=r], t)" @@ -367,7 +367,7 @@ (fst (last c) = None \ snd (last c) \ post)}" definition com_validity :: "'a com \ 'a set \ ('a \ 'a) set \ ('a \ 'a) set \ 'a set \ bool" - ("\ _ sat [_, _, _, _]" [60,0,0,0,0] 45) where + (\\ _ sat [_, _, _, _]\ [60,0,0,0,0] 45) where "\ P sat [pre, rely, guar, post] \ \s. cp (Some P) s \ assum(pre, rely) \ comm(guar, post)" @@ -386,7 +386,7 @@ (All_None (fst (last c)) \ snd( last c) \ post)}" definition par_com_validity :: "'a par_com \ 'a set \ ('a \ 'a) set \ ('a \ 'a) set -\ 'a set \ bool" ("\ _ SAT [_, _, _, _]" [60,0,0,0,0] 45) where +\ 'a set \ bool" (\\ _ SAT [_, _, _, _]\ [60,0,0,0,0] 45) where "\ Ps SAT [pre, rely, guar, post] \ \s. par_cp Ps s \ par_assum(pre, rely) \ par_comm(guar, post)" @@ -409,7 +409,7 @@ (\li \ (clist!l)!j -e\ (clist!l)! Suc j))) \ (c!j -pe\ c!Suc j \ (\i (clist!i)! Suc j)))" -definition conjoin :: "'a par_confs \ ('a confs) list \ bool" ("_ \ _" [65,65] 64) where +definition conjoin :: "'a par_confs \ ('a confs) list \ bool" (\_ \ _\ [65,65] 64) where "c \ clist \ (same_length c clist) \ (same_state c clist) \ (same_program c clist) \ (compat_label c clist)" subsubsection \Some previous lemmas\