diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Algebra/Galois_Connection.thy --- a/src/HOL/Algebra/Galois_Connection.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Algebra/Galois_Connection.thy Fri Sep 20 19:51:08 2024 +0200 @@ -12,19 +12,19 @@ subsection \Definition and basic properties\ record ('a, 'b, 'c, 'd) galcon = - orderA :: "('a, 'c) gorder_scheme" ("\\") - orderB :: "('b, 'd) gorder_scheme" ("\\") - lower :: "'a \ 'b" ("\\<^sup>*\") - upper :: "'b \ 'a" ("\\<^sub>*\") + orderA :: "('a, 'c) gorder_scheme" (\\\\) + orderB :: "('b, 'd) gorder_scheme" (\\\\) + lower :: "'a \ 'b" (\\\<^sup>*\\) + upper :: "'b \ 'a" (\\\<^sub>*\\) type_synonym ('a, 'b) galois = "('a, 'b, unit, unit) galcon" abbreviation "inv_galcon G \ \ orderA = inv_gorder \\<^bsub>G\<^esub>, orderB = inv_gorder \\<^bsub>G\<^esub>, lower = upper G, upper = lower G \" -definition comp_galcon :: "('b, 'c) galois \ ('a, 'b) galois \ ('a, 'c) galois" (infixr "\\<^sub>g" 85) +definition comp_galcon :: "('b, 'c) galois \ ('a, 'b) galois \ ('a, 'c) galois" (infixr \\\<^sub>g\ 85) where "G \\<^sub>g F = \ orderA = orderA F, orderB = orderB G, lower = lower G \ lower F, upper = upper F \ upper G \" -definition id_galcon :: "'a gorder \ ('a, 'a) galois" ("I\<^sub>g") where +definition id_galcon :: "'a gorder \ ('a, 'a) galois" (\I\<^sub>g\) where "I\<^sub>g(A) = \ orderA = A, orderB = A, lower = id, upper = id \"