merged
authornipkow
Mon, 26 Aug 2024 18:26:06 +0200
changeset 80775 de95c82eb31a
parent 80773 83d21da2bc59 (diff)
parent 80774 a2486a4b42da (current diff)
child 80776 3a9e570c916d
merged
--- a/NEWS	Mon Aug 26 18:26:00 2024 +0200
+++ b/NEWS	Mon Aug 26 18:26:06 2024 +0200
@@ -9,6 +9,13 @@
 
 ** General **
 
+* Inner syntax translations now support formal dependencies via commands
+'syntax_types' or 'syntax_consts'. This is essentially an abstract
+specification of the effect of 'translations' (or translation functions
+in ML). Most Isabelle theories have been adapted accordingly, such that
+hyperlinks work properly e.g. for "['a, 'b] \<Rightarrow> 'c" or "\<lbrakk>A; B\<rbrakk> \<Longrightarrow> C" in
+Pure, and "\<forall>x\<in>A. B x" or "\<exists>x\<in>A. B x" in HOL.
+
 * The Simplifier now supports special "congprocs", using keyword
 'congproc' or 'weak_congproc' in the 'simproc_setup' command (or ML
 antiquotation of the same name). See also
--- a/src/Benchmarks/Quickcheck_Benchmark/Needham_Schroeder_Base.thy	Mon Aug 26 18:26:00 2024 +0200
+++ b/src/Benchmarks/Quickcheck_Benchmark/Needham_Schroeder_Base.thy	Mon Aug 26 18:26:06 2024 +0200
@@ -28,6 +28,8 @@
 
 syntax
   "_MTuple"      :: "['a, args] => 'a * 'b"       ("(2\<lbrace>_,/ _\<rbrace>)")
+syntax_consts
+  "_MTuple"     == MPair
 translations
   "\<lbrace>x, y, z\<rbrace>"   == "\<lbrace>x, \<lbrace>y, z\<rbrace>\<rbrace>"
   "\<lbrace>x, y\<rbrace>"      == "CONST MPair x y"
--- a/src/CCL/Set.thy	Mon Aug 26 18:26:00 2024 +0200
+++ b/src/CCL/Set.thy	Mon Aug 26 18:26:06 2024 +0200
@@ -19,6 +19,8 @@
 
 syntax
   "_Coll" :: "[idt, o] \<Rightarrow> 'a set"  ("(1{_./ _})")
+syntax_consts
+  "_Coll" == Collect
 translations
   "{x. P}" == "CONST Collect(\<lambda>x. P)"
 
@@ -50,6 +52,9 @@
 syntax
   "_Ball" :: "[idt, 'a set, o] \<Rightarrow> o"  ("(ALL _:_./ _)" [0, 0, 0] 10)
   "_Bex" :: "[idt, 'a set, o] \<Rightarrow> o"  ("(EX _:_./ _)" [0, 0, 0] 10)
+syntax_consts
+  "_Ball" == Ball and
+  "_Bex" == Bex
 translations
   "ALL x:A. P"  == "CONST Ball(A, \<lambda>x. P)"
   "EX x:A. P"   == "CONST Bex(A, \<lambda>x. P)"
@@ -124,6 +129,9 @@
 syntax
   "_INTER" :: "[idt, 'a set, 'b set] \<Rightarrow> 'b set"  ("(INT _:_./ _)" [0, 0, 0] 10)
   "_UNION" :: "[idt, 'a set, 'b set] \<Rightarrow> 'b set"  ("(UN _:_./ _)" [0, 0, 0] 10)
+syntax_consts
+  "_INTER" == INTER and
+  "_UNION" == UNION
 translations
   "INT x:A. B" == "CONST INTER(A, \<lambda>x. B)"
   "UN x:A. B" == "CONST UNION(A, \<lambda>x. B)"
--- a/src/CCL/Term.thy	Mon Aug 26 18:26:00 2024 +0200
+++ b/src/CCL/Term.thy	Mon Aug 26 18:26:06 2024 +0200
@@ -49,6 +49,7 @@
   where let_def: "let1(t, f) == case(t,f(true),f(false), \<lambda>x y. f(<x,y>), \<lambda>u. f(lam x. u(x)))"
 
 syntax "_let1" :: "[idt,i,i]\<Rightarrow>i"  ("(3let _ be _/ in _)" [0,0,60] 60)
+syntax_consts "_let1" == let1
 translations "let x be a in e" == "CONST let1(a, \<lambda>x. e)"
 
 definition letrec :: "[[i,i\<Rightarrow>i]\<Rightarrow>i,(i\<Rightarrow>i)\<Rightarrow>i]\<Rightarrow>i"
@@ -67,6 +68,10 @@
   "_letrec" :: "[idt,idt,i,i]\<Rightarrow>i"  ("(3letrec _ _ be _/ in _)" [0,0,0,60] 60)
   "_letrec2" :: "[idt,idt,idt,i,i]\<Rightarrow>i"  ("(3letrec _ _ _ be _/ in _)" [0,0,0,0,60] 60)
   "_letrec3" :: "[idt,idt,idt,idt,i,i]\<Rightarrow>i"  ("(3letrec _ _ _ _ be _/ in _)" [0,0,0,0,0,60] 60)
+syntax_consts
+  "_letrec" == letrec and
+  "_letrec2" == letrec2 and
+  "_letrec3" == letrec3
 parse_translation \<open>
   let
     fun abs_tr t u = Syntax_Trans.abs_tr [t, u];
--- a/src/CCL/Type.thy	Mon Aug 26 18:26:00 2024 +0200
+++ b/src/CCL/Type.thy	Mon Aug 26 18:26:06 2024 +0200
@@ -14,6 +14,8 @@
 
 syntax
   "_Subtype" :: "[idt, 'a set, o] \<Rightarrow> 'a set"  ("(1{_: _ ./ _})")
+syntax_consts
+  "_Subtype" == Subtype
 translations
   "{x: A. B}" == "CONST Subtype(A, \<lambda>x. B)"
 
@@ -37,6 +39,9 @@
   "_Sigma" :: "[idt, i set, i set] \<Rightarrow> i set"  ("(3SUM _:_./ _)" [0,0,60] 60)
   "_arrow" :: "[i set, i set] \<Rightarrow> i set"  ("(_ ->/ _)"  [54, 53] 53)
   "_star"  :: "[i set, i set] \<Rightarrow> i set"  ("(_ */ _)" [56, 55] 55)
+syntax_consts
+  "_Pi" "_arrow" \<rightleftharpoons> Pi and
+  "_Sigma" "_star" \<rightleftharpoons> Sigma
 translations
   "PROD x:A. B" \<rightharpoonup> "CONST Pi(A, \<lambda>x. B)"
   "A -> B" \<rightharpoonup> "CONST Pi(A, \<lambda>_. B)"
--- a/src/CTT/CTT.thy	Mon Aug 26 18:26:00 2024 +0200
+++ b/src/CTT/CTT.thy	Mon Aug 26 18:26:06 2024 +0200
@@ -58,6 +58,9 @@
 syntax
   "_PROD"   :: "[idt,t,t]\<Rightarrow>t"       ("(3\<Prod>_:_./ _)" 10)
   "_SUM"    :: "[idt,t,t]\<Rightarrow>t"       ("(3\<Sum>_:_./ _)" 10)
+syntax_consts
+  "_PROD" \<rightleftharpoons> Prod and
+  "_SUM" \<rightleftharpoons> Sum
 translations
   "\<Prod>x:A. B" \<rightleftharpoons> "CONST Prod(A, \<lambda>x. B)"
   "\<Sum>x:A. B" \<rightleftharpoons> "CONST Sum(A, \<lambda>x. B)"
--- a/src/Cube/Cube.thy	Mon Aug 26 18:26:00 2024 +0200
+++ b/src/Cube/Cube.thy	Mon Aug 26 18:26:06 2024 +0200
@@ -39,6 +39,13 @@
   "_Lam" :: "[idt, term, term] \<Rightarrow> term"  ("(3\<^bold>\<lambda>_:_./ _)" [0, 0, 0] 10)
   "_Pi" :: "[idt, term, term] \<Rightarrow> term"  ("(3\<Prod>_:_./ _)" [0, 0] 10)
   "_arrow" :: "[term, term] \<Rightarrow> term"  (infixr "\<rightarrow>" 10)
+syntax_consts
+  "_Trueprop" \<rightleftharpoons> Trueprop and
+  "_MT_context" \<rightleftharpoons> MT_context and
+  "_Context" \<rightleftharpoons> Context and
+  "_Has_type" \<rightleftharpoons> Has_type and
+  "_Lam" \<rightleftharpoons> Abs and
+  "_Pi" "_arrow" \<rightleftharpoons> Prod
 translations
   "_Trueprop(G, t)" \<rightleftharpoons> "CONST Trueprop(G, t)"
   ("prop") "x:X" \<rightleftharpoons> ("prop") "\<turnstile> x:X"
--- a/src/Doc/Isar_Ref/Inner_Syntax.thy	Mon Aug 26 18:26:00 2024 +0200
+++ b/src/Doc/Isar_Ref/Inner_Syntax.thy	Mon Aug 26 18:26:06 2024 +0200
@@ -1078,6 +1078,8 @@
     @{command_def "nonterminal"} & : & \<open>theory \<rightarrow> theory\<close> \\
     @{command_def "syntax"} & : & \<open>local_theory \<rightarrow> local_theory\<close> \\
     @{command_def "no_syntax"} & : & \<open>local_theory \<rightarrow> local_theory\<close> \\
+    @{command_def "syntax_types"} & : & \<open>theory \<rightarrow> theory\<close> \\
+    @{command_def "syntax_consts"} & : & \<open>theory \<rightarrow> theory\<close> \\
     @{command_def "translations"} & : & \<open>theory \<rightarrow> theory\<close> \\
     @{command_def "no_translations"} & : & \<open>theory \<rightarrow> theory\<close> \\
     @{attribute_def syntax_ast_trace} & : & \<open>attribute\<close> & default \<open>false\<close> \\
@@ -1089,8 +1091,9 @@
   raw syntax declarations provide full access to the priority grammar of the
   inner syntax, without any sanity checks. This includes additional syntactic
   categories (via @{command nonterminal}) and free-form grammar productions
-  (via @{command syntax}). Additional syntax translations (or macros, via
-  @{command translations}) are required to turn resulting parse trees into
+  (via @{command syntax} with formal dependencies via @{command syntax_types}
+  and @{command syntax_consts}). Additional syntax translations (or macros,
+  via @{command translations}) are required to turn resulting parse trees into
   proper representations of formal entities again.
 
   \<^rail>\<open>
@@ -1098,6 +1101,8 @@
     ;
     (@@{command syntax} | @@{command no_syntax}) @{syntax mode}? (constdecl +)
     ;
+    (@@{command syntax_types} | @@{command syntax_consts}) (syntaxdeps + @'and')
+    ;
     (@@{command translations} | @@{command no_translations})
       (transpat ('==' | '=>' | '<=' | '\<rightleftharpoons>' | '\<rightharpoonup>' | '\<leftharpoondown>') transpat +)
     ;
@@ -1106,6 +1111,8 @@
     ;
     mode: ('(' ( @{syntax name} | @'output' | @{syntax name} @'output' ) ')')
     ;
+    syntaxdeps: (@{syntax name}+) ('==' | '\<rightleftharpoons>') (@{syntax name}+)
+    ;
     transpat: ('(' @{syntax name} ')')? @{syntax string}
   \<close>
 
@@ -1157,6 +1164,14 @@
   translations) resulting from \<open>decls\<close>, which are interpreted in the same
   manner as for @{command "syntax"} above.
 
+  \<^descr> @{command "syntax_types"}~\<open>syntax \<rightleftharpoons> types\<close> and @{command
+  "syntax_consts"}~\<open>syntax \<rightleftharpoons> consts\<close> declare dependencies of syntax constants
+  wrt.\ formal entities of the logic: multiple names may be given on both
+  sides. This tells the inner-syntax engine how to markup concrete syntax, to
+  support hyperlinks in the browser or editor. It is essentially an abstract
+  specification of the effect of translation rules (see below) or translation
+  functions (see \secref{sec:tr-funs}).
+
   \<^descr> @{command "translations"}~\<open>rules\<close> specifies syntactic translation rules
   (i.e.\ macros) as first-order rewrite rules on ASTs (\secref{sec:ast}). The
   theory context maintains two independent lists translation rules: parse
--- a/src/Doc/Tutorial/Protocol/Message.thy	Mon Aug 26 18:26:00 2024 +0200
+++ b/src/Doc/Tutorial/Protocol/Message.thy	Mon Aug 26 18:26:06 2024 +0200
@@ -81,6 +81,8 @@
 text\<open>Concrete syntax: messages appear as \<open>\<lbrace>A,B,NA\<rbrace>\<close>, etc...\<close>
 syntax
   "_MTuple"      :: "['a, args] \<Rightarrow> 'a * 'b"       ("(2\<lbrace>_,/ _\<rbrace>)")
+syntax_consts
+  "_MTuple"     == MPair
 translations
   "\<lbrace>x, y, z\<rbrace>"   == "\<lbrace>x, \<lbrace>y, z\<rbrace>\<rbrace>"
   "\<lbrace>x, y\<rbrace>"      == "CONST MPair x y"
--- a/src/FOL/IFOL.thy	Mon Aug 26 18:26:00 2024 +0200
+++ b/src/FOL/IFOL.thy	Mon Aug 26 18:26:06 2024 +0200
@@ -102,6 +102,7 @@
   where \<open>x \<noteq> y \<equiv> \<not> (x = y)\<close>
 
 syntax "_Uniq" :: "pttrn \<Rightarrow> o \<Rightarrow> o"  ("(2\<exists>\<^sub>\<le>\<^sub>1 _./ _)" [0, 10] 10)
+syntax_consts "_Uniq" \<rightleftharpoons> Uniq
 translations "\<exists>\<^sub>\<le>\<^sub>1x. P" \<rightleftharpoons> "CONST Uniq (\<lambda>x. P)"
 
 print_translation \<open>
@@ -746,7 +747,7 @@
   ""            :: \<open>letbind => letbinds\<close>              (\<open>_\<close>)
   "_binds"      :: \<open>[letbind, letbinds] => letbinds\<close>  (\<open>_;/ _\<close>)
   "_Let"        :: \<open>[letbinds, 'a] => 'a\<close>             (\<open>(let (_)/ in (_))\<close> 10)
-
+syntax_consts "_Let" \<rightleftharpoons> Let
 translations
   "_Let(_binds(b, bs), e)"  == "_Let(b, _Let(bs, e))"
   "let x = a in e"          == "CONST Let(a, \<lambda>x. e)"
--- a/src/HOL/Algebra/FiniteProduct.thy	Mon Aug 26 18:26:00 2024 +0200
+++ b/src/HOL/Algebra/FiniteProduct.thy	Mon Aug 26 18:26:06 2024 +0200
@@ -308,6 +308,8 @@
 syntax
   "_finprod" :: "index \<Rightarrow> idt \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b"
       ("(3\<Otimes>__\<in>_. _)" [1000, 0, 51, 10] 10)
+syntax_consts
+  "_finprod" \<rightleftharpoons> finprod
 translations
   "\<Otimes>\<^bsub>G\<^esub>i\<in>A. b" \<rightleftharpoons> "CONST finprod G (%i. b) A"
   \<comment> \<open>Beware of argument permutation!\<close>
--- a/src/HOL/Algebra/Lattice.thy	Mon Aug 26 18:26:00 2024 +0200
+++ b/src/HOL/Algebra/Lattice.thy	Mon Aug 26 18:26:06 2024 +0200
@@ -36,6 +36,10 @@
   "_sup1"     :: "('a, 'b) gorder_scheme \<Rightarrow> pttrns \<Rightarrow> 'a \<Rightarrow> 'a" ("(3SSUP\<index> _./ _)" [0, 10] 10)
   "_sup"      :: "('a, 'b) gorder_scheme \<Rightarrow> pttrn \<Rightarrow> 'c set \<Rightarrow> 'a \<Rightarrow> 'a"  ("(3SSUP\<index> _:_./ _)" [0, 0, 10] 10)
 
+syntax_consts
+  "_inf1" "_inf" == infi and
+  "_sup1" "_sup" == supr
+
 translations
   "IINF\<^bsub>L\<^esub> x. B"     == "CONST infi L CONST UNIV (%x. B)"
   "IINF\<^bsub>L\<^esub> x:A. B"   == "CONST infi L A (%x. B)"
--- a/src/HOL/Algebra/Ring.thy	Mon Aug 26 18:26:00 2024 +0200
+++ b/src/HOL/Algebra/Ring.thy	Mon Aug 26 18:26:06 2024 +0200
@@ -46,6 +46,8 @@
 syntax
   "_finsum" :: "index \<Rightarrow> idt \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b"
       ("(3\<Oplus>__\<in>_. _)" [1000, 0, 51, 10] 10)
+syntax_consts
+  "_finsum" \<rightleftharpoons> finsum
 translations
   "\<Oplus>\<^bsub>G\<^esub>i\<in>A. b" \<rightleftharpoons> "CONST finsum G (\<lambda>i. b) A"
   \<comment> \<open>Beware of argument permutation!\<close>
--- a/src/HOL/Analysis/Abstract_Metric_Spaces.thy	Mon Aug 26 18:26:00 2024 +0200
+++ b/src/HOL/Analysis/Abstract_Metric_Spaces.thy	Mon Aug 26 18:26:06 2024 +0200
@@ -1588,7 +1588,7 @@
   by (force simp: Cauchy_def Met_TC.MCauchy_def)
 
 lemma mcomplete_iff_complete [iff]:
-  "Met_TC.mcomplete (Pure.type ::'a::metric_space itself) \<longleftrightarrow> complete (UNIV::'a set)"
+  "Met_TC.mcomplete TYPE('a::metric_space) \<longleftrightarrow> complete (UNIV::'a set)"
   by (auto simp: Met_TC.mcomplete_def complete_def)
 
 context Submetric
--- a/src/HOL/Analysis/Bochner_Integration.thy	Mon Aug 26 18:26:00 2024 +0200
+++ b/src/HOL/Analysis/Bochner_Integration.thy	Mon Aug 26 18:26:06 2024 +0200
@@ -892,12 +892,18 @@
 syntax
   "_lebesgue_integral" :: "pttrn \<Rightarrow> real \<Rightarrow> 'a measure \<Rightarrow> real" ("\<integral>((2 _./ _)/ \<partial>_)" [60,61] 110)
 
+syntax_consts
+  "_lebesgue_integral" == lebesgue_integral
+
 translations
   "\<integral> x. f \<partial>M" == "CONST lebesgue_integral M (\<lambda>x. f)"
 
 syntax
   "_ascii_lebesgue_integral" :: "pttrn \<Rightarrow> 'a measure \<Rightarrow> real \<Rightarrow> real" ("(3LINT (1_)/|(_)./ _)" [0,110,60] 60)
 
+syntax_consts
+  "_ascii_lebesgue_integral" == lebesgue_integral
+
 translations
   "LINT x|M. f" == "CONST lebesgue_integral M (\<lambda>x. f)"
 
--- a/src/HOL/Analysis/Euclidean_Space.thy	Mon Aug 26 18:26:00 2024 +0200
+++ b/src/HOL/Analysis/Euclidean_Space.thy	Mon Aug 26 18:26:06 2024 +0200
@@ -34,6 +34,7 @@
     "(\<forall>u\<in>Basis. inner x u = 0) \<longleftrightarrow> (x = 0)"
 
 syntax "_type_dimension" :: "type \<Rightarrow> nat"  ("(1DIM/(1'(_')))")
+syntax_consts "_type_dimension" \<rightleftharpoons> card
 translations "DIM('a)" \<rightharpoonup> "CONST card (CONST Basis :: 'a set)"
 typed_print_translation \<open>
   [(\<^const_syntax>\<open>card\<close>,
--- a/src/HOL/Analysis/Finite_Cartesian_Product.thy	Mon Aug 26 18:26:00 2024 +0200
+++ b/src/HOL/Analysis/Finite_Cartesian_Product.thy	Mon Aug 26 18:26:06 2024 +0200
@@ -40,6 +40,7 @@
     \<^item> \<open>'a^'b::_\<close> becomes \<open>('a, 'b) vec\<close> without extra sort-constraint
 \<close>
 syntax "_vec_type" :: "type \<Rightarrow> type \<Rightarrow> type" (infixl "^" 15)
+syntax_types "_vec_type" \<rightleftharpoons> vec
 parse_translation \<open>
   let
     fun vec t u = Syntax.const \<^type_syntax>\<open>vec\<close> $ t $ u;
--- a/src/HOL/Analysis/Finite_Product_Measure.thy	Mon Aug 26 18:26:00 2024 +0200
+++ b/src/HOL/Analysis/Finite_Product_Measure.thy	Mon Aug 26 18:26:06 2024 +0200
@@ -170,6 +170,8 @@
 
 syntax
   "_PiM" :: "pttrn \<Rightarrow> 'i set \<Rightarrow> 'a measure \<Rightarrow> ('i => 'a) measure"  ("(3\<Pi>\<^sub>M _\<in>_./ _)"  10)
+syntax_consts
+  "_PiM" == PiM
 translations
   "\<Pi>\<^sub>M x\<in>I. M" == "CONST PiM I (%x. M)"
 
--- a/src/HOL/Analysis/Infinite_Set_Sum.thy	Mon Aug 26 18:26:00 2024 +0200
+++ b/src/HOL/Analysis/Infinite_Set_Sum.thy	Mon Aug 26 18:26:06 2024 +0200
@@ -100,6 +100,8 @@
 syntax
   "_infsetsum" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b::{banach, second_countable_topology}"
   ("(2\<Sum>\<^sub>a_\<in>_./ _)" [0, 51, 10] 10)
+syntax_consts
+  "_infsetsum" \<rightleftharpoons> infsetsum
 translations \<comment> \<open>Beware of argument permutation!\<close>
   "\<Sum>\<^sub>ai\<in>A. b" \<rightleftharpoons> "CONST infsetsum (\<lambda>i. b) A"
 
@@ -109,6 +111,8 @@
 syntax
   "_uinfsetsum" :: "pttrn \<Rightarrow> 'b \<Rightarrow> 'b::{banach, second_countable_topology}"
   ("(2\<Sum>\<^sub>a_./ _)" [0, 10] 10)
+syntax_consts
+  "_uinfsetsum" \<rightleftharpoons> infsetsum
 translations \<comment> \<open>Beware of argument permutation!\<close>
   "\<Sum>\<^sub>ai. b" \<rightleftharpoons> "CONST infsetsum (\<lambda>i. b) (CONST UNIV)"
 
@@ -118,6 +122,8 @@
 syntax
   "_qinfsetsum" :: "pttrn \<Rightarrow> bool \<Rightarrow> 'a \<Rightarrow> 'a::{banach, second_countable_topology}"
   ("(2\<Sum>\<^sub>a_ | (_)./ _)" [0, 0, 10] 10)
+syntax_consts
+  "_qinfsetsum" \<rightleftharpoons> infsetsum
 translations
   "\<Sum>\<^sub>ax|P. t" => "CONST infsetsum (\<lambda>x. t) {x. P}"
 
--- a/src/HOL/Analysis/Infinite_Sum.thy	Mon Aug 26 18:26:00 2024 +0200
+++ b/src/HOL/Analysis/Infinite_Sum.thy	Mon Aug 26 18:26:06 2024 +0200
@@ -50,6 +50,8 @@
   "_infsum" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b::topological_comm_monoid_add"  ("(3INFSUM (_/:_)./ _)" [0, 51, 10] 10)
 syntax
   "_infsum" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b::topological_comm_monoid_add"  ("(2\<Sum>\<^sub>\<infinity>(_/\<in>_)./ _)" [0, 51, 10] 10)
+syntax_consts
+  "_infsum" \<rightleftharpoons> infsum
 translations \<comment> \<open>Beware of argument permutation!\<close>
   "\<Sum>\<^sub>\<infinity>i\<in>A. b" \<rightleftharpoons> "CONST infsum (\<lambda>i. b) A"
 
@@ -57,6 +59,8 @@
   "_univinfsum" :: "pttrn \<Rightarrow> 'a \<Rightarrow> 'a"  ("(3INFSUM _./ _)" [0, 10] 10)
 syntax
   "_univinfsum" :: "pttrn \<Rightarrow> 'a \<Rightarrow> 'a"  ("(2\<Sum>\<^sub>\<infinity>_./ _)" [0, 10] 10)
+syntax_consts
+  "_univinfsum" \<rightleftharpoons> infsum
 translations
   "\<Sum>\<^sub>\<infinity>x. t" \<rightleftharpoons> "CONST infsum (\<lambda>x. t) (CONST UNIV)"
 
@@ -64,6 +68,8 @@
   "_qinfsum" :: "pttrn \<Rightarrow> bool \<Rightarrow> 'a \<Rightarrow> 'a"  ("(3INFSUM _ |/ _./ _)" [0, 0, 10] 10)
 syntax
   "_qinfsum" :: "pttrn \<Rightarrow> bool \<Rightarrow> 'a \<Rightarrow> 'a"  ("(2\<Sum>\<^sub>\<infinity>_ | (_)./ _)" [0, 0, 10] 10)
+syntax_consts
+  "_qinfsum" \<rightleftharpoons> infsum
 translations
   "\<Sum>\<^sub>\<infinity>x|P. t" => "CONST infsum (\<lambda>x. t) {x. P}"
 
--- a/src/HOL/Analysis/Interval_Integral.thy	Mon Aug 26 18:26:00 2024 +0200
+++ b/src/HOL/Analysis/Interval_Integral.thy	Mon Aug 26 18:26:06 2024 +0200
@@ -148,6 +148,9 @@
   "_ascii_interval_lebesgue_integral" :: "pttrn \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real measure \<Rightarrow> real \<Rightarrow> real"
   ("(5LINT _=_.._|_. _)" [0,60,60,61,100] 60)
 
+syntax_consts
+  "_ascii_interval_lebesgue_integral" == interval_lebesgue_integral
+
 translations
   "LINT x=a..b|M. f" == "CONST interval_lebesgue_integral M a b (\<lambda>x. f)"
 
@@ -159,6 +162,9 @@
   "_ascii_interval_lebesgue_borel_integral" :: "pttrn \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real"
   ("(4LBINT _=_.._. _)" [0,60,60,61] 60)
 
+syntax_consts
+  "_ascii_interval_lebesgue_borel_integral" == interval_lebesgue_integral
+
 translations
   "LBINT x=a..b. f" == "CONST interval_lebesgue_integral CONST lborel a b (\<lambda>x. f)"
 
@@ -1045,11 +1051,17 @@
 syntax "_complex_lebesgue_borel_integral" :: "pttrn \<Rightarrow> real \<Rightarrow> complex"
   ("(2CLBINT _. _)" [0,60] 60)
 
+syntax_consts
+  "_complex_lebesgue_borel_integral" == complex_lebesgue_integral
+
 translations "CLBINT x. f" == "CONST complex_lebesgue_integral CONST lborel (\<lambda>x. f)"
 
 syntax "_complex_set_lebesgue_borel_integral" :: "pttrn \<Rightarrow> real set \<Rightarrow> real \<Rightarrow> complex"
   ("(3CLBINT _:_. _)" [0,60,61] 60)
 
+syntax_consts
+  "_complex_set_lebesgue_borel_integral" == complex_set_lebesgue_integral
+
 translations
   "CLBINT x:A. f" == "CONST complex_set_lebesgue_integral CONST lborel A (\<lambda>x. f)"
 
@@ -1065,6 +1077,9 @@
   "_ascii_complex_interval_lebesgue_borel_integral" :: "pttrn \<Rightarrow> ereal \<Rightarrow> ereal \<Rightarrow> real \<Rightarrow> complex"
   ("(4CLBINT _=_.._. _)" [0,60,60,61] 60)
 
+syntax_consts
+  "_ascii_complex_interval_lebesgue_borel_integral" == complex_interval_lebesgue_integral
+
 translations
   "CLBINT x=a..b. f" == "CONST complex_interval_lebesgue_integral CONST lborel a b (\<lambda>x. f)"
 
--- a/src/HOL/Analysis/Measure_Space.thy	Mon Aug 26 18:26:00 2024 +0200
+++ b/src/HOL/Analysis/Measure_Space.thy	Mon Aug 26 18:26:06 2024 +0200
@@ -1006,6 +1006,9 @@
 syntax
   "_almost_everywhere" :: "pttrn \<Rightarrow> 'a \<Rightarrow> bool \<Rightarrow> bool" ("AE _ in _. _" [0,0,10] 10)
 
+syntax_consts
+  "_almost_everywhere" \<rightleftharpoons> almost_everywhere
+
 translations
   "AE x in M. P" \<rightleftharpoons> "CONST almost_everywhere M (\<lambda>x. P)"
 
@@ -1016,6 +1019,9 @@
   "_set_almost_everywhere" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'a \<Rightarrow> bool \<Rightarrow> bool"
   ("AE _\<in>_ in _./ _" [0,0,0,10] 10)
 
+syntax_consts
+  "_set_almost_everywhere" \<rightleftharpoons> set_almost_everywhere
+
 translations
   "AE x\<in>A in M. P" \<rightleftharpoons> "CONST set_almost_everywhere A M (\<lambda>x. P)"
 
--- a/src/HOL/Analysis/Nonnegative_Lebesgue_Integration.thy	Mon Aug 26 18:26:00 2024 +0200
+++ b/src/HOL/Analysis/Nonnegative_Lebesgue_Integration.thy	Mon Aug 26 18:26:06 2024 +0200
@@ -586,6 +586,9 @@
 syntax
   "_simple_integral" :: "pttrn \<Rightarrow> ennreal \<Rightarrow> 'a measure \<Rightarrow> ennreal" ("\<integral>\<^sup>S _. _ \<partial>_" [60,61] 110)
 
+syntax_consts
+  "_simple_integral" == simple_integral
+
 translations
   "\<integral>\<^sup>S x. f \<partial>M" == "CONST simple_integral M (%x. f)"
 
@@ -819,6 +822,9 @@
 syntax
   "_nn_integral" :: "pttrn \<Rightarrow> ennreal \<Rightarrow> 'a measure \<Rightarrow> ennreal" ("\<integral>\<^sup>+((2 _./ _)/ \<partial>_)" [60,61] 110)
 
+syntax_consts
+  "_nn_integral" == nn_integral
+
 translations
   "\<integral>\<^sup>+x. f \<partial>M" == "CONST nn_integral M (\<lambda>x. f)"
 
--- a/src/HOL/Analysis/Set_Integral.thy	Mon Aug 26 18:26:00 2024 +0200
+++ b/src/HOL/Analysis/Set_Integral.thy	Mon Aug 26 18:26:06 2024 +0200
@@ -26,6 +26,9 @@
   "_ascii_set_lebesgue_integral" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'a measure \<Rightarrow> real \<Rightarrow> real"
   ("(4LINT (_):(_)/|(_)./ _)" [0,0,0,10] 10)
 
+syntax_consts
+  "_ascii_set_lebesgue_integral" == set_lebesgue_integral
+
 translations
   "LINT x:A|M. f" == "CONST set_lebesgue_integral M A (\<lambda>x. f)"
 
@@ -576,6 +579,9 @@
   "_complex_lebesgue_integral" :: "pttrn \<Rightarrow> complex \<Rightarrow> 'a measure \<Rightarrow> complex"
  ("\<integral>\<^sup>C _. _ \<partial>_" [0,0] 110)
 
+syntax_consts
+  "_complex_lebesgue_integral" == complex_lebesgue_integral
+
 translations
   "\<integral>\<^sup>Cx. f \<partial>M" == "CONST complex_lebesgue_integral M (\<lambda>x. f)"
 
@@ -583,6 +589,9 @@
   "_ascii_complex_lebesgue_integral" :: "pttrn \<Rightarrow> 'a measure \<Rightarrow> real \<Rightarrow> real"
   ("(3CLINT _|_. _)" [0,0,10] 10)
 
+syntax_consts
+  "_ascii_complex_lebesgue_integral" == complex_lebesgue_integral
+
 translations
   "CLINT x|M. f" == "CONST complex_lebesgue_integral M (\<lambda>x. f)"
 
@@ -617,6 +626,9 @@
   "_ascii_complex_set_lebesgue_integral" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'a measure \<Rightarrow> real \<Rightarrow> real"
   ("(4CLINT _:_|_. _)" [0,0,0,10] 10)
 
+syntax_consts
+  "_ascii_complex_set_lebesgue_integral" == complex_set_lebesgue_integral
+
 translations
   "CLINT x:A|M. f" == "CONST complex_set_lebesgue_integral M A (\<lambda>x. f)"
 
@@ -640,6 +652,10 @@
 "_set_lebesgue_integral" :: "pttrn => 'a set => 'a measure => ereal => ereal"
 ("(\<integral>((_)\<in>(_)./ _)/\<partial>_)" [0,0,0,110] 10)
 
+syntax_consts
+"_set_nn_integral" == set_nn_integral and
+"_set_lebesgue_integral" == set_lebesgue_integral
+
 translations
 "\<integral>\<^sup>+x \<in> A. f \<partial>M" == "CONST set_nn_integral M A (\<lambda>x. f)"
 "\<integral>x \<in> A. f \<partial>M" == "CONST set_lebesgue_integral M A (\<lambda>x. f)"
--- a/src/HOL/Analysis/Sparse_In.thy	Mon Aug 26 18:26:00 2024 +0200
+++ b/src/HOL/Analysis/Sparse_In.thy	Mon Aug 26 18:26:06 2024 +0200
@@ -166,11 +166,15 @@
 
 syntax
   "_eventually_cosparse" :: "pttrn => 'a set => bool => bool"  ("(3\<forall>\<^sub>\<approx>_\<in>_./ _)" [0, 0, 10] 10)
+syntax_consts
+  "_eventually_cosparse" == eventually
 translations
   "\<forall>\<^sub>\<approx>x\<in>A. P" == "CONST eventually (\<lambda>x. P) (CONST cosparse A)"
 
 syntax
   "_qeventually_cosparse" :: "pttrn \<Rightarrow> bool \<Rightarrow> 'a \<Rightarrow> 'a"  ("(3\<forall>\<^sub>\<approx>_ | (_)./ _)" [0, 0, 10] 10)
+syntax_consts
+  "_qeventually_cosparse" == eventually
 translations
   "\<forall>\<^sub>\<approx>x|P. t" => "CONST eventually (\<lambda>x. t) (CONST cosparse {x. P})"
 
--- a/src/HOL/Auth/Message.thy	Mon Aug 26 18:26:00 2024 +0200
+++ b/src/HOL/Auth/Message.thy	Mon Aug 26 18:26:06 2024 +0200
@@ -51,6 +51,8 @@
 text\<open>Concrete syntax: messages appear as \<open>\<lbrace>A,B,NA\<rbrace>\<close>, etc...\<close>
 syntax
   "_MTuple" :: "['a, args] \<Rightarrow> 'a * 'b"  ("(2\<lbrace>_,/ _\<rbrace>)")
+syntax_consts
+  "_MTuple" \<rightleftharpoons> MPair
 translations
   "\<lbrace>x, y, z\<rbrace>" \<rightleftharpoons> "\<lbrace>x, \<lbrace>y, z\<rbrace>\<rbrace>"
   "\<lbrace>x, y\<rbrace>" \<rightleftharpoons> "CONST MPair x y"
--- a/src/HOL/BNF_Cardinal_Arithmetic.thy	Mon Aug 26 18:26:00 2024 +0200
+++ b/src/HOL/BNF_Cardinal_Arithmetic.thy	Mon Aug 26 18:26:06 2024 +0200
@@ -275,6 +275,9 @@
   "pttrn => ('a * 'a) set => 'b * 'b set => (('a * 'b) * ('a * 'b)) set"
   ("(3CSUM _:_. _)" [0, 51, 10] 10)
 
+syntax_consts
+  "_Csum" == Csum
+
 translations
   "CSUM i:r. rs" == "CONST Csum r (%i. rs)"
 
--- a/src/HOL/Bali/AxSem.thy	Mon Aug 26 18:26:00 2024 +0200
+++ b/src/HOL/Bali/AxSem.thy	Mon Aug 26 18:26:06 2024 +0200
@@ -205,6 +205,8 @@
 
 syntax
   "_peek_res" :: "pttrn \<Rightarrow> 'a assn \<Rightarrow> 'a assn"            ("\<lambda>_:. _" [0,3] 3)
+syntax_consts
+  "_peek_res" == peek_res
 translations
   "\<lambda>w:. P"   == "CONST peek_res (\<lambda>w. P)"
 
@@ -266,6 +268,8 @@
 
 syntax
   "_peek_st"   :: "pttrn \<Rightarrow> 'a assn \<Rightarrow> 'a assn"            ("\<lambda>_.. _" [0,3] 3)
+syntax_consts
+  "_peek_st" == peek_st
 translations
   "\<lambda>s.. P"   == "CONST peek_st (\<lambda>s. P)"
 
--- a/src/HOL/Bali/Basis.thy	Mon Aug 26 18:26:00 2024 +0200
+++ b/src/HOL/Bali/Basis.thy	Mon Aug 26 18:26:06 2024 +0200
@@ -195,6 +195,10 @@
   "_Oall" :: "[pttrn, 'a option, bool] \<Rightarrow> bool"   ("(3\<forall>_\<in>_:/ _)"  [0,0,10] 10)
   "_Oex"  :: "[pttrn, 'a option, bool] \<Rightarrow> bool"   ("(3\<exists>_\<in>_:/ _)"  [0,0,10] 10)
 
+syntax_consts
+  "_Oall" \<rightleftharpoons> Ball and
+  "_Oex" \<rightleftharpoons> Bex
+
 translations
   "\<forall>x\<in>A: P" \<rightleftharpoons> "\<forall>x\<in>CONST set_option A. P"
   "\<exists>x\<in>A: P" \<rightleftharpoons> "\<exists>x\<in>CONST set_option A. P"
@@ -257,6 +261,8 @@
 text \<open>list patterns -- extends pre-defined type "pttrn" used in abstractions\<close>
 syntax
   "_lpttrn" :: "[pttrn, pttrn] \<Rightarrow> pttrn"    ("_#/_" [901,900] 900)
+syntax_consts
+  "_lpttrn" \<rightleftharpoons> lsplit
 translations
   "\<lambda>y # x # xs. b" \<rightleftharpoons> "CONST lsplit (\<lambda>y x # xs. b)"
   "\<lambda>x # xs. b" \<rightleftharpoons> "CONST lsplit (\<lambda>x xs. b)"
--- a/src/HOL/Bali/State.thy	Mon Aug 26 18:26:00 2024 +0200
+++ b/src/HOL/Bali/State.thy	Mon Aug 26 18:26:06 2024 +0200
@@ -135,6 +135,10 @@
   Heap  :: "loc   \<Rightarrow> oref"
   Stat  :: "qtname \<Rightarrow> oref"
 
+syntax_consts
+  Heap == Inl and
+  Stat == Inr
+
 translations
   "Heap" => "CONST Inl"
   "Stat" => "CONST Inr"
--- a/src/HOL/Bit_Operations.thy	Mon Aug 26 18:26:00 2024 +0200
+++ b/src/HOL/Bit_Operations.thy	Mon Aug 26 18:26:06 2024 +0200
@@ -148,7 +148,7 @@
 
 lemma bit_imp_possible_bit:
   \<open>possible_bit TYPE('a) n\<close> if \<open>bit a n\<close>
-  by (rule ccontr) (use that in \<open>auto simp add: bit_iff_odd possible_bit_def\<close>)
+  by (rule ccontr) (use that in \<open>auto simp: bit_iff_odd possible_bit_def\<close>)
 
 lemma impossible_bit:
   \<open>\<not> bit a n\<close> if \<open>\<not> possible_bit TYPE('a) n\<close>
@@ -161,7 +161,7 @@
 
 lemma possible_bit_min [simp]:
   \<open>possible_bit TYPE('a) (min i j) \<longleftrightarrow> possible_bit TYPE('a) i \<or> possible_bit TYPE('a) j\<close>
-  by (auto simp add: min_def elim: possible_bit_less_imp)
+  by (auto simp: min_def elim: possible_bit_less_imp)
 
 lemma bit_eqI:
   \<open>a = b\<close> if \<open>\<And>n. possible_bit TYPE('a) n \<Longrightarrow> bit a n \<longleftrightarrow> bit b n\<close>
@@ -213,7 +213,7 @@
     also have \<open>\<dots> = b\<close>
       by (fact mod_mult_div_eq)
     finally show ?case
-      by (auto simp add: mod2_eq_if)
+      by (auto simp: mod2_eq_if)
   qed
 qed
 
@@ -262,7 +262,7 @@
   \<open>bit (2 * a) (Suc n) \<longleftrightarrow> possible_bit TYPE('a) (Suc n) \<and> bit a n\<close>
   using even_double_div_exp_iff [of n a]
   by (cases \<open>possible_bit TYPE('a) (Suc n)\<close>)
-    (auto simp add: bit_iff_odd possible_bit_def)
+    (auto simp: bit_iff_odd possible_bit_def)
 
 lemma bit_double_iff [bit_simps]:
   \<open>bit (2 * a) n \<longleftrightarrow> possible_bit TYPE('a) n \<and> n \<noteq> 0 \<and> bit a (n - 1)\<close>
@@ -382,7 +382,7 @@
     with rec [of n True] show ?case
       by simp
   qed
-qed (auto simp add: div_mult2_eq bit_nat_def)
+qed (auto simp: div_mult2_eq bit_nat_def)
 
 end
 
@@ -422,12 +422,12 @@
     case (even m)
     then show ?case
       by (cases n)
-        (auto simp add: bit_double_iff Bit_Operations.bit_double_iff possible_bit_def bit_0 dest: mult_not_zero)
+        (auto simp: bit_double_iff Bit_Operations.bit_double_iff possible_bit_def bit_0 dest: mult_not_zero)
   next
     case (odd m)
     then show ?case
       by (cases n)
-        (auto simp add: bit_double_iff even_bit_succ_iff possible_bit_def
+        (auto simp: bit_double_iff even_bit_succ_iff possible_bit_def
           Bit_Operations.bit_Suc Bit_Operations.bit_0 dest: mult_not_zero)
   qed
   with True show ?thesis
@@ -521,7 +521,7 @@
     with rec [of k True] show ?case
       by (simp add: ac_simps)
   qed
-qed (auto simp add: zdiv_zmult2_eq bit_int_def)
+qed (auto simp: zdiv_zmult2_eq bit_int_def)
 
 end
 
@@ -637,25 +637,25 @@
 qed
 
 sublocale "and": semilattice \<open>(AND)\<close>
-  by standard (auto simp add: bit_eq_iff bit_and_iff)
+  by standard (auto simp: bit_eq_iff bit_and_iff)
 
 sublocale or: semilattice_neutr \<open>(OR)\<close> 0
-  by standard (auto simp add: bit_eq_iff bit_or_iff)
+  by standard (auto simp: bit_eq_iff bit_or_iff)
 
 sublocale xor: comm_monoid \<open>(XOR)\<close> 0
-  by standard (auto simp add: bit_eq_iff bit_xor_iff)
+  by standard (auto simp: bit_eq_iff bit_xor_iff)
 
 lemma even_and_iff:
   \<open>even (a AND b) \<longleftrightarrow> even a \<or> even b\<close>
-  using bit_and_iff [of a b 0] by (auto simp add: bit_0)
+  using bit_and_iff [of a b 0] by (auto simp: bit_0)
 
 lemma even_or_iff:
   \<open>even (a OR b) \<longleftrightarrow> even a \<and> even b\<close>
-  using bit_or_iff [of a b 0] by (auto simp add: bit_0)
+  using bit_or_iff [of a b 0] by (auto simp: bit_0)
 
 lemma even_xor_iff:
   \<open>even (a XOR b) \<longleftrightarrow> (even a \<longleftrightarrow> even b)\<close>
-  using bit_xor_iff [of a b 0] by (auto simp add: bit_0)
+  using bit_xor_iff [of a b 0] by (auto simp: bit_0)
 
 lemma zero_and_eq [simp]:
   \<open>0 AND a = 0\<close>
@@ -667,7 +667,7 @@
 
 lemma one_and_eq:
   \<open>1 AND a = a mod 2\<close>
-  by (simp add: bit_eq_iff bit_and_iff) (auto simp add: bit_1_iff bit_0)
+  by (simp add: bit_eq_iff bit_and_iff) (auto simp: bit_1_iff bit_0)
 
 lemma and_one_eq:
   \<open>a AND 1 = a mod 2\<close>
@@ -676,7 +676,7 @@
 lemma one_or_eq:
   \<open>1 OR a = a + of_bool (even a)\<close>
   by (simp add: bit_eq_iff bit_or_iff add.commute [of _ 1] even_bit_succ_iff)
-    (auto simp add: bit_1_iff bit_0)
+    (auto simp: bit_1_iff bit_0)
 
 lemma or_one_eq:
   \<open>a OR 1 = a + of_bool (even a)\<close>
@@ -685,7 +685,7 @@
 lemma one_xor_eq:
   \<open>1 XOR a = a + of_bool (even a) - of_bool (odd a)\<close>
   by (simp add: bit_eq_iff bit_xor_iff add.commute [of _ 1] even_bit_succ_iff)
-    (auto simp add: bit_1_iff odd_bit_iff_bit_pred bit_0 elim: oddE)
+    (auto simp: bit_1_iff odd_bit_iff_bit_pred bit_0 elim: oddE)
 
 lemma xor_one_eq:
   \<open>a XOR 1 = a + of_bool (even a) - of_bool (odd a)\<close>
@@ -762,7 +762,7 @@
 
 lemma even_mask_iff:
   \<open>even (mask n) \<longleftrightarrow> n = 0\<close>
-  using bit_mask_iff [of n 0] by (auto simp add: bit_0)
+  using bit_mask_iff [of n 0] by (auto simp: bit_0)
 
 lemma mask_Suc_0 [simp]:
   \<open>mask (Suc 0) = 1\<close>
@@ -770,7 +770,7 @@
 
 lemma mask_Suc_exp:
   \<open>mask (Suc n) = 2 ^ n OR mask n\<close>
-  by (auto simp add: bit_eq_iff bit_simps)
+  by (auto simp: bit_eq_iff bit_simps)
 
 lemma mask_numeral:
   \<open>mask (numeral n) = 1 + 2 * mask (pred_numeral n)\<close>
@@ -793,20 +793,19 @@
 proof (induction n arbitrary: m)
   case 0
   then show ?case
-    by (auto simp add: bit_0 push_bit_eq_mult)
+    by (auto simp: bit_0 push_bit_eq_mult)
 next
   case (Suc n)
   show ?case
   proof (cases m)
     case 0
     then show ?thesis
-      by (auto simp add: bit_imp_possible_bit)
+      by (auto simp: bit_imp_possible_bit)
   next
-    case (Suc m)
-    with Suc.prems Suc.IH [of m] show ?thesis
+    case (Suc m')
+    with Suc.prems Suc.IH [of m'] show ?thesis
       apply (simp add: push_bit_double)
-      apply (simp add: bit_simps mult.commute [of _ 2])
-      apply (auto simp add: possible_bit_less_imp)
+      apply (auto simp: possible_bit_less_imp bit_simps mult.commute [of _ 2])
       done
   qed
 qed
@@ -849,7 +848,7 @@
 
 lemma disjunctive_xor_eq_or:
   \<open>a XOR b = a OR b\<close> if \<open>a AND b = 0\<close>
-  using that by (auto simp add: bit_eq_iff bit_simps)
+  using that by (auto simp: bit_eq_iff bit_simps)
 
 lemma disjunctive_add_eq_or:
   \<open>a + b = a OR b\<close> if \<open>a AND b = 0\<close>
@@ -864,11 +863,11 @@
   proof (induction n arbitrary: a b)
     case 0
     from "0"(2)[of 0] show ?case
-      by (auto simp add: even_or_iff bit_0)
+      by (auto simp: even_or_iff bit_0)
   next
     case (Suc n)
     from Suc.prems(2) [of 0] have even: \<open>even a \<or> even b\<close>
-      by (auto simp add: bit_0)
+      by (auto simp: bit_0)
     have bit: \<open>\<not> bit (a div 2) n \<or> \<not> bit (b div 2) n\<close> for n
       using Suc.prems(2) [of \<open>Suc n\<close>] by (simp add: bit_Suc)
     from Suc.prems have \<open>possible_bit TYPE('a) n\<close>
@@ -879,7 +878,7 @@
     have \<open>a + b = (a div 2 * 2 + a mod 2) + (b div 2 * 2 + b mod 2)\<close>
       using div_mult_mod_eq [of a 2] div_mult_mod_eq [of b 2] by simp
     also have \<open>\<dots> = of_bool (odd a \<or> odd b) + 2 * (a div 2 + b div 2)\<close>
-      using even by (auto simp add: algebra_simps mod2_eq_if)
+      using even by (auto simp: algebra_simps mod2_eq_if)
     finally have \<open>bit ((a + b) div 2) n \<longleftrightarrow> bit (a div 2 + b div 2) n\<close>
       using \<open>possible_bit TYPE('a) (Suc n)\<close> by simp (simp_all flip: bit_Suc add: bit_double_iff possible_bit_def)
     also have \<open>\<dots> \<longleftrightarrow> bit (a div 2 OR b div 2) n\<close>
@@ -929,9 +928,9 @@
   also have \<open>\<dots> \<longleftrightarrow> m \<le> n \<and> bit a n \<or> n < m \<and> bit a n\<close>
     by auto
   also have \<open>m \<le> n \<and> bit a n \<longleftrightarrow> bit (push_bit m (drop_bit m a)) n\<close>
-    by (auto simp add: bit_simps bit_imp_possible_bit)
+    by (auto simp: bit_simps bit_imp_possible_bit)
   finally show ?thesis
-    by (auto simp add: bit_simps)
+    by (auto simp: bit_simps)
 qed
 
 lemma take_bit_Suc:
@@ -1004,19 +1003,19 @@
 
 lemma push_bit_take_bit:
   \<open>push_bit m (take_bit n a) = take_bit (m + n) (push_bit m a)\<close>
-  by (rule bit_eqI) (auto simp add: bit_simps)
+  by (rule bit_eqI) (auto simp: bit_simps)
 
 lemma take_bit_push_bit:
   \<open>take_bit m (push_bit n a) = push_bit n (take_bit (m - n) a)\<close>
-  by (rule bit_eqI) (auto simp add: bit_simps)
+  by (rule bit_eqI) (auto simp: bit_simps)
 
 lemma take_bit_drop_bit:
   \<open>take_bit m (drop_bit n a) = drop_bit n (take_bit (m + n) a)\<close>
-  by (rule bit_eqI) (auto simp add: bit_simps)
+  by (rule bit_eqI) (auto simp: bit_simps)
 
 lemma drop_bit_take_bit:
   \<open>drop_bit m (take_bit n a) = take_bit (n - m) (drop_bit m a)\<close>
-  by (rule bit_eqI) (auto simp add: bit_simps)
+  by (rule bit_eqI) (auto simp: bit_simps)
 
 lemma even_push_bit_iff [simp]:
   \<open>even (push_bit n a) \<longleftrightarrow> n \<noteq> 0 \<or> even a\<close>
@@ -1091,49 +1090,49 @@
     then have \<open> \<not> bit a (n + (m - n))\<close>
       by (simp add: bit_simps)
     then show \<open>bit (take_bit n a) m \<longleftrightarrow> bit a m\<close>
-      by (cases \<open>m < n\<close>) (auto simp add: bit_simps)
+      by (cases \<open>m < n\<close>) (auto simp: bit_simps)
   qed
 qed
 
 lemma drop_bit_exp_eq:
   \<open>drop_bit m (2 ^ n) = of_bool (m \<le> n \<and> possible_bit TYPE('a) n) * 2 ^ (n - m)\<close>
-  by (auto simp add: bit_eq_iff bit_simps)
+  by (auto simp: bit_eq_iff bit_simps)
 
 lemma take_bit_and [simp]:
   \<open>take_bit n (a AND b) = take_bit n a AND take_bit n b\<close>
-  by (auto simp add: bit_eq_iff bit_simps)
+  by (auto simp: bit_eq_iff bit_simps)
 
 lemma take_bit_or [simp]:
   \<open>take_bit n (a OR b) = take_bit n a OR take_bit n b\<close>
-  by (auto simp add: bit_eq_iff bit_simps)
+  by (auto simp: bit_eq_iff bit_simps)
 
 lemma take_bit_xor [simp]:
   \<open>take_bit n (a XOR b) = take_bit n a XOR take_bit n b\<close>
-  by (auto simp add: bit_eq_iff bit_simps)
+  by (auto simp: bit_eq_iff bit_simps)
 
 lemma push_bit_and [simp]:
   \<open>push_bit n (a AND b) = push_bit n a AND push_bit n b\<close>
-  by (auto simp add: bit_eq_iff bit_simps)
+  by (auto simp: bit_eq_iff bit_simps)
 
 lemma push_bit_or [simp]:
   \<open>push_bit n (a OR b) = push_bit n a OR push_bit n b\<close>
-  by (auto simp add: bit_eq_iff bit_simps)
+  by (auto simp: bit_eq_iff bit_simps)
 
 lemma push_bit_xor [simp]:
   \<open>push_bit n (a XOR b) = push_bit n a XOR push_bit n b\<close>
-  by (auto simp add: bit_eq_iff bit_simps)
+  by (auto simp: bit_eq_iff bit_simps)
 
 lemma drop_bit_and [simp]:
   \<open>drop_bit n (a AND b) = drop_bit n a AND drop_bit n b\<close>
-  by (auto simp add: bit_eq_iff bit_simps)
+  by (auto simp: bit_eq_iff bit_simps)
 
 lemma drop_bit_or [simp]:
   \<open>drop_bit n (a OR b) = drop_bit n a OR drop_bit n b\<close>
-  by (auto simp add: bit_eq_iff bit_simps)
+  by (auto simp: bit_eq_iff bit_simps)
 
 lemma drop_bit_xor [simp]:
   \<open>drop_bit n (a XOR b) = drop_bit n a XOR drop_bit n b\<close>
-  by (auto simp add: bit_eq_iff bit_simps)
+  by (auto simp: bit_eq_iff bit_simps)
 
 lemma take_bit_of_mask [simp]:
   \<open>take_bit m (mask n) = mask (min m n)\<close>
@@ -1141,11 +1140,11 @@
 
 lemma take_bit_eq_mask:
   \<open>take_bit n a = a AND mask n\<close>
-  by (auto simp add: bit_eq_iff bit_simps)
+  by (auto simp: bit_eq_iff bit_simps)
 
 lemma or_eq_0_iff:
   \<open>a OR b = 0 \<longleftrightarrow> a = 0 \<and> b = 0\<close>
-  by (auto simp add: bit_eq_iff bit_or_iff)
+  by (auto simp: bit_eq_iff bit_or_iff)
 
 lemma bit_iff_and_drop_bit_eq_1:
   \<open>bit a n \<longleftrightarrow> drop_bit n a AND 1 = 1\<close>
@@ -1157,23 +1156,23 @@
 
 lemma bit_set_bit_iff [bit_simps]:
   \<open>bit (set_bit m a) n \<longleftrightarrow> bit a n \<or> (m = n \<and> possible_bit TYPE('a) n)\<close>
-  by (auto simp add: set_bit_eq_or bit_or_iff bit_exp_iff)
+  by (auto simp: set_bit_eq_or bit_or_iff bit_exp_iff)
 
 lemma even_set_bit_iff:
   \<open>even (set_bit m a) \<longleftrightarrow> even a \<and> m \<noteq> 0\<close>
-  using bit_set_bit_iff [of m a 0] by (auto simp add: bit_0)
+  using bit_set_bit_iff [of m a 0] by (auto simp: bit_0)
 
 lemma bit_unset_bit_iff [bit_simps]:
   \<open>bit (unset_bit m a) n \<longleftrightarrow> bit a n \<and> m \<noteq> n\<close>
-  by (auto simp add: unset_bit_eq_or_xor bit_simps dest: bit_imp_possible_bit)
+  by (auto simp: unset_bit_eq_or_xor bit_simps dest: bit_imp_possible_bit)
 
 lemma even_unset_bit_iff:
   \<open>even (unset_bit m a) \<longleftrightarrow> even a \<or> m = 0\<close>
-  using bit_unset_bit_iff [of m a 0] by (auto simp add: bit_0)
+  using bit_unset_bit_iff [of m a 0] by (auto simp: bit_0)
 
 lemma bit_flip_bit_iff [bit_simps]:
   \<open>bit (flip_bit m a) n \<longleftrightarrow> (m = n \<longleftrightarrow> \<not> bit a n) \<and> possible_bit TYPE('a) n\<close>
-  by (auto simp add: bit_eq_iff bit_simps flip_bit_eq_xor bit_imp_possible_bit)
+  by (auto simp: bit_eq_iff bit_simps flip_bit_eq_xor bit_imp_possible_bit)
 
 lemma even_flip_bit_iff:
   \<open>even (flip_bit m a) \<longleftrightarrow> \<not> (even a \<longleftrightarrow> m = 0)\<close>
@@ -1182,7 +1181,7 @@
 lemma and_exp_eq_0_iff_not_bit:
   \<open>a AND 2 ^ n = 0 \<longleftrightarrow> \<not> bit a n\<close> (is \<open>?P \<longleftrightarrow> ?Q\<close>)
   using bit_imp_possible_bit[of a n]
-  by (auto simp add: bit_eq_iff bit_simps)
+  by (auto simp: bit_eq_iff bit_simps)
 
 lemma bit_sum_mult_2_cases:
   assumes a: \<open>\<forall>j. \<not> bit a (Suc j)\<close>
@@ -1191,52 +1190,52 @@
   from a have \<open>n = 0\<close> if \<open>bit a n\<close> for n using that
     by (cases n) simp_all
   then have \<open>a = 0 \<or> a = 1\<close>
-    by (auto simp add: bit_eq_iff bit_1_iff)
+    by (auto simp: bit_eq_iff bit_1_iff)
   then show ?thesis
-    by (cases n) (auto simp add: bit_0 bit_double_iff even_bit_succ_iff)
+    by (cases n) (auto simp: bit_0 bit_double_iff even_bit_succ_iff)
 qed
 
 lemma set_bit_0 [simp]:
   \<open>set_bit 0 a = 1 + 2 * (a div 2)\<close>
-  by (auto simp add: bit_eq_iff bit_simps even_bit_succ_iff simp flip: bit_Suc)
+  by (auto simp: bit_eq_iff bit_simps even_bit_succ_iff simp flip: bit_Suc)
 
 lemma set_bit_Suc:
   \<open>set_bit (Suc n) a = a mod 2 + 2 * set_bit n (a div 2)\<close>
-  by (auto simp add: bit_eq_iff bit_sum_mult_2_cases bit_simps bit_0 simp flip: bit_Suc
+  by (auto simp: bit_eq_iff bit_sum_mult_2_cases bit_simps bit_0 simp flip: bit_Suc
     elim: possible_bit_less_imp)
 
 lemma unset_bit_0 [simp]:
   \<open>unset_bit 0 a = 2 * (a div 2)\<close>
-  by (auto simp add: bit_eq_iff bit_simps simp flip: bit_Suc)
+  by (auto simp: bit_eq_iff bit_simps simp flip: bit_Suc)
 
 lemma unset_bit_Suc:
   \<open>unset_bit (Suc n) a = a mod 2 + 2 * unset_bit n (a div 2)\<close>
-  by (auto simp add: bit_eq_iff bit_sum_mult_2_cases bit_simps bit_0 simp flip: bit_Suc)
+  by (auto simp: bit_eq_iff bit_sum_mult_2_cases bit_simps bit_0 simp flip: bit_Suc)
 
 lemma flip_bit_0 [simp]:
   \<open>flip_bit 0 a = of_bool (even a) + 2 * (a div 2)\<close>
-  by (auto simp add: bit_eq_iff bit_simps even_bit_succ_iff bit_0 simp flip: bit_Suc)
+  by (auto simp: bit_eq_iff bit_simps even_bit_succ_iff bit_0 simp flip: bit_Suc)
 
 lemma flip_bit_Suc:
   \<open>flip_bit (Suc n) a = a mod 2 + 2 * flip_bit n (a div 2)\<close>
-  by (auto simp add: bit_eq_iff bit_sum_mult_2_cases bit_simps bit_0 simp flip: bit_Suc
+  by (auto simp: bit_eq_iff bit_sum_mult_2_cases bit_simps bit_0 simp flip: bit_Suc
     elim: possible_bit_less_imp)
 
 lemma flip_bit_eq_if:
   \<open>flip_bit n a = (if bit a n then unset_bit else set_bit) n a\<close>
-  by (rule bit_eqI) (auto simp add: bit_set_bit_iff bit_unset_bit_iff bit_flip_bit_iff)
+  by (rule bit_eqI) (auto simp: bit_set_bit_iff bit_unset_bit_iff bit_flip_bit_iff)
 
 lemma take_bit_set_bit_eq:
   \<open>take_bit n (set_bit m a) = (if n \<le> m then take_bit n a else set_bit m (take_bit n a))\<close>
-  by (rule bit_eqI) (auto simp add: bit_take_bit_iff bit_set_bit_iff)
+  by (rule bit_eqI) (auto simp: bit_take_bit_iff bit_set_bit_iff)
 
 lemma take_bit_unset_bit_eq:
   \<open>take_bit n (unset_bit m a) = (if n \<le> m then take_bit n a else unset_bit m (take_bit n a))\<close>
-  by (rule bit_eqI) (auto simp add: bit_take_bit_iff bit_unset_bit_iff)
+  by (rule bit_eqI) (auto simp: bit_take_bit_iff bit_unset_bit_iff)
 
 lemma take_bit_flip_bit_eq:
   \<open>take_bit n (flip_bit m a) = (if n \<le> m then take_bit n a else flip_bit m (take_bit n a))\<close>
-  by (rule bit_eqI) (auto simp add: bit_take_bit_iff bit_flip_bit_iff)
+  by (rule bit_eqI) (auto simp: bit_take_bit_iff bit_flip_bit_iff)
 
 lemma bit_1_0 [simp]:
   \<open>bit 1 0\<close>
@@ -1272,17 +1271,17 @@
     with bit_rec [of _ n] Cons.prems Cons.IH [of m]
     show ?thesis
       by (simp add: bit_simps)
-        (auto simp add: possible_bit_less_imp bit_simps simp flip: bit_Suc)
+        (auto simp: possible_bit_less_imp bit_simps simp flip: bit_Suc)
   qed
 qed
 
 lemma horner_sum_bit_eq_take_bit:
   \<open>horner_sum of_bool 2 (map (bit a) [0..<n]) = take_bit n a\<close>
-  by (rule bit_eqI) (auto simp add: bit_simps)
+  by (rule bit_eqI) (auto simp: bit_simps)
 
 lemma take_bit_horner_sum_bit_eq:
   \<open>take_bit n (horner_sum of_bool 2 bs) = horner_sum of_bool 2 (take n bs)\<close>
-  by (auto simp add: bit_eq_iff bit_take_bit_iff bit_horner_sum_bit_iff)
+  by (auto simp: bit_eq_iff bit_take_bit_iff bit_horner_sum_bit_iff)
 
 lemma take_bit_sum:
   \<open>take_bit n a = (\<Sum>k = 0..<n. push_bit k (of_bool (bit a k)))\<close>
@@ -1292,9 +1291,9 @@
   \<open>set_bit n a = a + of_bool (\<not> bit a n) * 2 ^ n\<close>
 proof -
   have \<open>a AND of_bool (\<not> bit a n) * 2 ^ n = 0\<close>
-    by (auto simp add: bit_eq_iff bit_simps)
+    by (auto simp: bit_eq_iff bit_simps)
   then show ?thesis
-    by (auto simp add: bit_eq_iff bit_simps disjunctive_add_eq_or)
+    by (auto simp: bit_eq_iff bit_simps disjunctive_add_eq_or)
 qed
 
 end
@@ -1360,7 +1359,7 @@
 
 lemma bit_not_exp_iff [bit_simps]:
   \<open>bit (NOT (2 ^ m)) n \<longleftrightarrow> possible_bit TYPE('a) n \<and> n \<noteq> m\<close>
-  by (auto simp add: bit_not_iff bit_exp_iff)
+  by (auto simp: bit_not_iff bit_exp_iff)
 
 lemma bit_minus_iff [bit_simps]:
   \<open>bit (- a) n \<longleftrightarrow> possible_bit TYPE('a) n \<and> \<not> bit (a - 1) n\<close>
@@ -1372,7 +1371,7 @@
 
 lemma bit_minus_exp_iff [bit_simps]:
   \<open>bit (- (2 ^ m)) n \<longleftrightarrow> possible_bit TYPE('a) n \<and> n \<ge> m\<close>
-  by (auto simp add: bit_simps simp flip: mask_eq_exp_minus_1)
+  by (auto simp: bit_simps simp flip: mask_eq_exp_minus_1)
 
 lemma bit_minus_2_iff [simp]:
   \<open>bit (- 2) n \<longleftrightarrow> possible_bit TYPE('a) n \<and> n > 0\<close>
@@ -1394,13 +1393,13 @@
   by standard (rule bit_eqI, simp add: bit_and_iff)
 
 sublocale bit: abstract_boolean_algebra \<open>(AND)\<close> \<open>(OR)\<close> NOT 0 \<open>- 1\<close>
-  by standard (auto simp add: bit_and_iff bit_or_iff bit_not_iff intro: bit_eqI)
+  by standard (auto simp: bit_and_iff bit_or_iff bit_not_iff intro: bit_eqI)
 
 sublocale bit: abstract_boolean_algebra_sym_diff \<open>(AND)\<close> \<open>(OR)\<close> NOT 0 \<open>- 1\<close> \<open>(XOR)\<close>
-  apply standard
-  apply (rule bit_eqI)
-  apply (auto simp add: bit_simps)
-  done
+proof
+  show \<open>\<And>x y. x XOR y = x AND NOT y OR NOT x AND y\<close>
+    by (intro bit_eqI) (auto simp: bit_simps)
+qed
 
 lemma and_eq_not_not_or:
   \<open>a AND b = NOT (NOT a OR NOT b)\<close>
@@ -1424,7 +1423,7 @@
 
 lemma disjunctive_and_not_eq_xor:
   \<open>a AND NOT b = a XOR b\<close> if \<open>NOT a AND b = 0\<close>
-  using that by (auto simp add: bit_eq_iff bit_simps)
+  using that by (auto simp: bit_eq_iff bit_simps)
 
 lemma disjunctive_diff_eq_and_not:
   \<open>a - b = a AND NOT b\<close> if \<open>NOT a AND b = 0\<close>
@@ -1447,11 +1446,11 @@
 
 lemma take_bit_not_take_bit:
   \<open>take_bit n (NOT (take_bit n a)) = take_bit n (NOT a)\<close>
-  by (auto simp add: bit_eq_iff bit_take_bit_iff bit_not_iff)
+  by (auto simp: bit_eq_iff bit_take_bit_iff bit_not_iff)
 
 lemma take_bit_not_iff:
   \<open>take_bit n (NOT a) = take_bit n (NOT b) \<longleftrightarrow> take_bit n a = take_bit n b\<close>
-  by (auto simp add: bit_eq_iff bit_simps)
+  by (auto simp: bit_eq_iff bit_simps)
 
 lemma take_bit_not_eq_mask_diff:
   \<open>take_bit n (NOT a) = mask n - take_bit n a\<close>
@@ -1459,7 +1458,7 @@
   have \<open>NOT (mask n) AND take_bit n a = 0\<close>
     by (simp add: bit_eq_iff bit_simps)
   moreover have \<open>take_bit n (NOT a) = mask n AND NOT (take_bit n a)\<close>
-    by (auto simp add: bit_eq_iff bit_simps)
+    by (auto simp: bit_eq_iff bit_simps)
   ultimately show ?thesis
     by (simp add: disjunctive_diff_eq_and_not)
 qed
@@ -1486,14 +1485,11 @@
 
 lemma unset_bit_eq_and_not:
   \<open>unset_bit n a = a AND NOT (push_bit n 1)\<close>
-  by (rule bit_eqI) (auto simp add: bit_simps)
+  by (rule bit_eqI) (auto simp: bit_simps)
 
 lemma push_bit_Suc_minus_numeral [simp]:
   \<open>push_bit (Suc n) (- numeral k) = push_bit n (- numeral (Num.Bit0 k))\<close>
-  apply (simp only: numeral_Bit0)
-  apply simp
-  apply (simp only: numeral_mult mult_2_right numeral_add)
-  done
+  using local.push_bit_Suc_numeral push_bit_minus by presburger
 
 lemma push_bit_minus_numeral [simp]:
   \<open>push_bit (numeral l) (- numeral k) = push_bit (pred_numeral l) (- numeral (Num.Bit0 k))\<close>
@@ -1509,11 +1505,11 @@
 
 lemma push_bit_mask_eq:
   \<open>push_bit m (mask n) = mask (n + m) AND NOT (mask m)\<close>
-  by (rule bit_eqI) (auto simp add: bit_simps not_less possible_bit_less_imp)
+  by (rule bit_eqI) (auto simp: bit_simps not_less possible_bit_less_imp)
 
 lemma slice_eq_mask:
   \<open>push_bit n (take_bit m (drop_bit n a)) = a AND mask (m + n) AND NOT (mask n)\<close>
-  by (rule bit_eqI) (auto simp add: bit_simps)
+  by (rule bit_eqI) (auto simp: bit_simps)
 
 lemma push_bit_numeral_minus_1 [simp]:
   \<open>push_bit (numeral n) (- 1) = - (2 ^ numeral n)\<close>
@@ -1523,9 +1519,9 @@
   \<open>unset_bit n a = a - of_bool (bit a n) * 2 ^ n\<close>
 proof -
   have \<open>NOT a AND of_bool (bit a n) * 2 ^ n = 0\<close>
-    by (auto simp add: bit_eq_iff bit_simps)
+    by (auto simp: bit_eq_iff bit_simps)
   moreover have \<open>unset_bit n a = a AND NOT (of_bool (bit a n) * 2 ^ n)\<close>
-    by (auto simp add: bit_eq_iff bit_simps)
+    by (auto simp: bit_eq_iff bit_simps)
   ultimately show ?thesis
     by (simp add: disjunctive_diff_eq_and_not)
 qed
@@ -1616,7 +1612,7 @@
 
 lemma drop_bit_mask_eq:
   \<open>drop_bit m (mask n) = mask (n - m)\<close>
-  by (rule bit_eqI) (auto simp add: bit_simps possible_bit_def)
+  by (rule bit_eqI) (auto simp: bit_simps possible_bit_def)
 
 lemma bit_push_bit_iff':
   \<open>bit (push_bit m a) n \<longleftrightarrow> m \<le> n \<and> bit a (n - m)\<close>
@@ -1645,7 +1641,7 @@
 proof (rule ccontr)
   assume \<open>\<not> 0 < take_bit m a\<close>
   then have \<open>take_bit m a = 0\<close>
-    by (auto simp add: not_less intro: order_antisym)
+    by (auto simp: not_less intro: order_antisym)
   then have \<open>bit (take_bit m a) n = bit 0 n\<close>
     by simp
   with that show False
@@ -1659,7 +1655,7 @@
 lemma drop_bit_push_bit:
   \<open>drop_bit m (push_bit n a) = drop_bit (m - n) (push_bit (n - m) a)\<close>
   by (cases \<open>m \<le> n\<close>)
-    (auto simp add: mult.left_commute [of _ \<open>2 ^ n\<close>] mult.commute [of _ \<open>2 ^ n\<close>] mult.assoc
+    (auto simp: mult.left_commute [of _ \<open>2 ^ n\<close>] mult.commute [of _ \<open>2 ^ n\<close>] mult.assoc
     mult.commute [of a] drop_bit_eq_div push_bit_eq_mult not_le power_add Orderings.not_le dest!: le_Suc_ex less_imp_Suc_add)
 
 end
@@ -1684,7 +1680,7 @@
   have less_eq: \<open>\<bar>k div 2\<bar> \<le> \<bar>k\<bar>\<close> for k :: int
     by (cases k) (simp_all add: divide_int_def nat_add_distrib)
   then have less: \<open>\<bar>k div 2\<bar> < \<bar>k\<bar>\<close> if \<open>k \<notin> {0, - 1}\<close> for k :: int
-    using that by (auto simp add: less_le [of k])
+    using that by (auto simp: less_le [of k])
   show \<open>wf (measure (\<lambda>(k, l). nat (\<bar>k\<bar> + \<bar>l\<bar>)))\<close>
     by simp
   show \<open>((k div 2, l div 2), k, l) \<in> measure (\<lambda>(k, l). nat (\<bar>k\<bar> + \<bar>l\<bar>))\<close>
@@ -1721,11 +1717,11 @@
 proof (cases \<open>k \<in> {0, - 1} \<and> l \<in> {0, - 1}\<close>)
   case True
   then show ?thesis
-    by (auto simp add: F.simps [of 0] F.simps [of \<open>- 1\<close>])
+    by (auto simp: F.simps [of 0] F.simps [of \<open>- 1\<close>])
 next
   case False
   then show ?thesis
-    by (auto simp add: ac_simps F.simps [of k l])
+    by (auto simp: ac_simps F.simps [of k l])
 qed
 
 lemma bit_iff:
@@ -1800,7 +1796,7 @@
   fix k l :: int and m n :: nat
   show \<open>unset_bit n k = (k OR push_bit n 1) XOR push_bit n 1\<close>
     by (rule bit_eqI)
-      (auto simp add: unset_bit_int_def
+      (auto simp: unset_bit_int_def
         and_int.bit_iff or_int.bit_iff xor_int.bit_iff bit_not_int_iff push_bit_int_def bit_simps)
 qed (fact and_int.rec or_int.rec xor_int.rec mask_int_def set_bit_int_def flip_bit_int_def
   push_bit_int_def drop_bit_int_def take_bit_int_def not_int_def)+
@@ -1837,13 +1833,13 @@
     case (even k)
     then show ?case
       using bit_double_iff [of \<open>of_int k\<close> n] Bit_Operations.bit_double_iff [of k n]
-      by (cases n) (auto simp add: ac_simps possible_bit_def dest: mult_not_zero)
+      by (cases n) (auto simp: ac_simps possible_bit_def dest: mult_not_zero)
   next
     case (odd k)
     then show ?case
       using bit_double_iff [of \<open>of_int k\<close> n]
       by (cases n)
-        (auto simp add: ac_simps bit_double_iff even_bit_succ_iff Bit_Operations.bit_0 Bit_Operations.bit_Suc
+        (auto simp: ac_simps bit_double_iff even_bit_succ_iff Bit_Operations.bit_0 Bit_Operations.bit_Suc
           possible_bit_def dest: mult_not_zero)
   qed
   with True show ?thesis
@@ -1953,7 +1949,7 @@
     by simp
   with and_int.rec [of \<open>1 + k * 2\<close> l]
   show ?case
-    by (auto simp add: zero_le_mult_iff not_le)
+    by (auto simp: zero_le_mult_iff not_le)
 qed
 
 lemma and_negative_int_iff [simp]:
@@ -2018,7 +2014,7 @@
 
 lemma xor_negative_int_iff [simp]:
   \<open>k XOR l < 0 \<longleftrightarrow> (k < 0) \<noteq> (l < 0)\<close> for k l :: int
-  by (subst Not_eq_iff [symmetric]) (auto simp add: not_less)
+  by (subst Not_eq_iff [symmetric]) (auto simp: not_less)
 
 lemma OR_upper: \<^marker>\<open>contributor \<open>Stefan Berghofer\<close>\<close>
   \<open>x OR y < 2 ^ n\<close> if \<open>0 \<le> x\<close> \<open>x < 2 ^ n\<close> \<open>y < 2 ^ n\<close> for x y :: int
@@ -2034,12 +2030,12 @@
   case (even x)
   from even.IH [of \<open>n - 1\<close> \<open>y div 2\<close>] even.prems even.hyps
   show ?case
-    by (cases n) (auto simp add: or_int.rec [of \<open>_ * 2\<close>] elim: oddE)
+    by (cases n) (auto simp: or_int.rec [of \<open>_ * 2\<close>] elim: oddE)
 next
   case (odd x)
   from odd.IH [of \<open>n - 1\<close> \<open>y div 2\<close>] odd.prems odd.hyps
   show ?case
-    by (cases n) (auto simp add: or_int.rec [of \<open>1 + _ * 2\<close>], linarith)
+    by (cases n) (auto simp: or_int.rec [of \<open>1 + _ * 2\<close>], linarith)
 qed
 
 lemma XOR_upper: \<^marker>\<open>contributor \<open>Stefan Berghofer\<close>\<close>
@@ -2056,12 +2052,12 @@
   case (even x)
   from even.IH [of \<open>n - 1\<close> \<open>y div 2\<close>] even.prems even.hyps
   show ?case
-    by (cases n) (auto simp add: xor_int.rec [of \<open>_ * 2\<close>] elim: oddE)
+    by (cases n) (auto simp: xor_int.rec [of \<open>_ * 2\<close>] elim: oddE)
 next
   case (odd x)
   from odd.IH [of \<open>n - 1\<close> \<open>y div 2\<close>] odd.prems odd.hyps
   show ?case
-    by (cases n) (auto simp add: xor_int.rec [of \<open>1 + _ * 2\<close>])
+    by (cases n) (auto simp: xor_int.rec [of \<open>1 + _ * 2\<close>])
 qed
 
 lemma AND_lower [simp]: \<^marker>\<open>contributor \<open>Stefan Berghofer\<close>\<close>
@@ -2120,12 +2116,12 @@
   case (even x)
   from even.IH [of \<open>y div 2\<close>]
   show ?case
-    by (auto simp add: and_int.rec [of _ y] or_int.rec [of _ y] elim: oddE)
+    by (auto simp: and_int.rec [of _ y] or_int.rec [of _ y] elim: oddE)
 next
   case (odd x)
   from odd.IH [of \<open>y div 2\<close>]
   show ?case
-    by (auto simp add: and_int.rec [of _ y] or_int.rec [of _ y] elim: oddE)
+    by (auto simp: and_int.rec [of _ y] or_int.rec [of _ y] elim: oddE)
 qed
 
 lemma push_bit_minus_one:
@@ -2182,7 +2178,7 @@
 
 lemma drop_bit_nonnegative_int_iff [simp]:
   \<open>drop_bit n k \<ge> 0 \<longleftrightarrow> k \<ge> 0\<close> for k :: int
-  by (induction n) (auto simp add: drop_bit_Suc drop_bit_half)
+  by (induction n) (auto simp: drop_bit_Suc drop_bit_half)
 
 lemma drop_bit_negative_int_iff [simp]:
   \<open>drop_bit n k < 0 \<longleftrightarrow> k < 0\<close> for k :: int
@@ -2223,17 +2219,17 @@
 lemma and_int_unfold:
   \<open>k AND l = (if k = 0 \<or> l = 0 then 0 else if k = - 1 then l else if l = - 1 then k
     else (k mod 2) * (l mod 2) + 2 * ((k div 2) AND (l div 2)))\<close> for k l :: int
-  by (auto simp add: and_int.rec [of k l] zmult_eq_1_iff elim: oddE)
+  by (auto simp: and_int.rec [of k l] zmult_eq_1_iff elim: oddE)
 
 lemma or_int_unfold:
   \<open>k OR l = (if k = - 1 \<or> l = - 1 then - 1 else if k = 0 then l else if l = 0 then k
     else max (k mod 2) (l mod 2) + 2 * ((k div 2) OR (l div 2)))\<close> for k l :: int
-  by (auto simp add: or_int.rec [of k l] elim: oddE)
+  by (auto simp: or_int.rec [of k l] elim: oddE)
 
 lemma xor_int_unfold:
   \<open>k XOR l = (if k = - 1 then NOT l else if l = - 1 then NOT k else if k = 0 then l else if l = 0 then k
     else \<bar>k mod 2 - l mod 2\<bar> + 2 * ((k div 2) XOR (l div 2)))\<close> for k l :: int
-  by (auto simp add: xor_int.rec [of k l] not_int_def elim!: oddE)
+  by (auto simp: xor_int.rec [of k l] not_int_def elim!: oddE)
 
 lemma bit_minus_int_iff:
   \<open>bit (- k) n \<longleftrightarrow> bit (NOT (k - 1)) n\<close> for k :: int
@@ -2325,7 +2321,7 @@
 
 lemma take_bit_int_less_self_iff:
   \<open>take_bit n k < k \<longleftrightarrow> 2 ^ n \<le> k\<close> for k :: int
-  by (auto simp add: less_le take_bit_int_less_eq_self_iff take_bit_int_eq_self_iff
+  by (auto simp: less_le take_bit_int_less_eq_self_iff take_bit_int_eq_self_iff
     intro: order_trans [of 0 \<open>2 ^ n\<close> k])
 
 lemma take_bit_int_greater_self_iff:
@@ -2334,7 +2330,7 @@
 
 lemma take_bit_int_greater_eq_self_iff:
   \<open>k \<le> take_bit n k \<longleftrightarrow> k < 2 ^ n\<close> for k :: int
-  by (auto simp add: le_less take_bit_int_greater_self_iff take_bit_int_eq_self_iff
+  by (auto simp: le_less take_bit_int_greater_self_iff take_bit_int_eq_self_iff
     dest: sym not_sym intro: less_trans [of k 0 \<open>2 ^ n\<close>])
 
 lemma take_bit_tightened_less_eq_int:
@@ -2370,7 +2366,7 @@
       fix m
       assume \<open>nat k \<le> m\<close>
       then show \<open>bit k m \<longleftrightarrow> bit k (nat k)\<close>
-        by (auto simp add: * bit_iff_odd power_add zdiv_zmult2_eq dest!: le_Suc_ex)
+        by (auto simp: * bit_iff_odd power_add zdiv_zmult2_eq dest!: le_Suc_ex)
     qed
   next
     case False
@@ -2388,7 +2384,7 @@
       fix m
       assume \<open>nat (- k) \<le> m\<close>
       then show \<open>bit k m \<longleftrightarrow> bit k (nat (- k))\<close>
-        by (auto simp add: * bit_iff_odd power_add zdiv_zmult2_eq minus_1_div_exp_eq_int dest!: le_Suc_ex)
+        by (auto simp: * bit_iff_odd power_add zdiv_zmult2_eq minus_1_div_exp_eq_int dest!: le_Suc_ex)
     qed
   qed
   show thesis
@@ -2408,18 +2404,12 @@
     moreover have \<open>finite N\<close> \<open>r \<in> N\<close>
       using ** N_def \<open>r < q\<close> by auto
     moreover define n where \<open>n = Suc (Max N)\<close>
-    ultimately have \<open>\<And>m. n \<le> m \<Longrightarrow> bit k m \<longleftrightarrow> bit k n\<close>
-      apply auto
-         apply (metis (full_types, lifting) "*" Max_ge_iff Suc_n_not_le_n \<open>finite N\<close> all_not_in_conv mem_Collect_eq not_le)
-        apply (metis "*" Max_ge Suc_n_not_le_n \<open>finite N\<close> linorder_not_less mem_Collect_eq)
-        apply (metis "*" Max_ge Suc_n_not_le_n \<open>finite N\<close> linorder_not_less mem_Collect_eq)
-      apply (metis (full_types, lifting) "*" Max_ge_iff Suc_n_not_le_n \<open>finite N\<close> all_not_in_conv mem_Collect_eq not_le)
-      done
+    ultimately have \<dagger>: \<open>\<And>m. n \<le> m \<Longrightarrow> bit k m \<longleftrightarrow> bit k n\<close>
+      by (smt (verit) "*" Max_ge Suc_n_not_le_n linorder_not_less mem_Collect_eq not_less_eq_eq)
     have \<open>bit k (Max N) \<noteq> bit k n\<close>
       by (metis (mono_tags, lifting) "*" Max_in N_def \<open>\<And>m. n \<le> m \<Longrightarrow> bit k m = bit k n\<close> \<open>finite N\<close> \<open>r \<in> N\<close> empty_iff le_cases mem_Collect_eq)
-    show thesis apply (rule that [of n])
-      using \<open>\<And>m. n \<le> m \<Longrightarrow> bit k m = bit k n\<close> apply blast
-      using \<open>bit k (Max N) \<noteq> bit k n\<close> n_def by auto
+    with \<dagger> n_def that [of n] show thesis
+      by fastforce
   qed
 qed
 
@@ -2591,17 +2581,17 @@
 lemma and_nat_unfold [code]:
   \<open>m AND n = (if m = 0 \<or> n = 0 then 0 else (m mod 2) * (n mod 2) + 2 * ((m div 2) AND (n div 2)))\<close>
     for m n :: nat
-  by (auto simp add: and_rec [of m n] elim: oddE)
+  by (auto simp: and_rec [of m n] elim: oddE)
 
 lemma or_nat_unfold [code]:
   \<open>m OR n = (if m = 0 then n else if n = 0 then m
     else max (m mod 2) (n mod 2) + 2 * ((m div 2) OR (n div 2)))\<close> for m n :: nat
-  by (auto simp add: or_rec [of m n] elim: oddE)
+  by (auto simp: or_rec [of m n] elim: oddE)
 
 lemma xor_nat_unfold [code]:
   \<open>m XOR n = (if m = 0 then n else if n = 0 then m
     else (m mod 2 + n mod 2) mod 2 + 2 * ((m div 2) XOR (n div 2)))\<close> for m n :: nat
-  by (auto simp add: xor_rec [of m n] elim!: oddE)
+  by (auto simp: xor_rec [of m n] elim!: oddE)
 
 lemma [code]:
   \<open>unset_bit 0 m = 2 * (m div 2)\<close>
@@ -2680,10 +2670,11 @@
 
 lemma drop_bit_nat_eq:
   \<open>drop_bit n (nat k) = nat (drop_bit n k)\<close>
-  apply (cases \<open>k \<ge> 0\<close>)
-   apply (simp_all add: drop_bit_eq_div nat_div_distrib nat_power_eq not_le)
-  apply (simp add: divide_int_def)
-  done
+proof (cases \<open>k \<ge> 0\<close>)
+  case True
+  then show ?thesis
+    by (metis drop_bit_of_nat int_nat_eq nat_int) 
+qed (simp add: nat_eq_iff2)
 
 lemma take_bit_nat_eq:
   \<open>take_bit n (nat k) = nat (take_bit n k)\<close> if \<open>k \<ge> 0\<close>
@@ -2932,7 +2923,7 @@
   \<open>numeral (Num.Bit1 m) AND NOT (1 :: int) = numeral (Num.Bit0 m)\<close>
   \<open>numeral (Num.Bit1 m) AND NOT (numeral (Num.Bit0 n)) = 1 + (2 :: int) * (numeral m AND NOT (numeral n))\<close>
   \<open>numeral (Num.Bit1 m) AND NOT (numeral (Num.Bit1 n)) = (2 :: int) * (numeral m AND NOT (numeral n))\<close>
-  by (simp_all add: bit_eq_iff) (auto simp add: bit_0 bit_simps bit_Suc bit_numeral_rec BitM_inc_eq sub_inc_One_eq split: nat.split)
+  by (simp_all add: bit_eq_iff) (auto simp: bit_0 bit_simps bit_Suc bit_numeral_rec BitM_inc_eq sub_inc_One_eq split: nat.split)
 
 fun and_not_num :: \<open>num \<Rightarrow> num \<Rightarrow> num option\<close> \<^marker>\<open>contributor \<open>Andreas Lochbihler\<close>\<close>
 where
@@ -2989,7 +2980,7 @@
   \<open>numeral (Num.Bit1 m) OR NOT (numeral (Num.Bit0 n)) = 1 + (2 :: int) * (numeral m OR NOT (numeral n))\<close>
   \<open>numeral (Num.Bit1 m) OR NOT (numeral (Num.Bit1 n)) = 1 + (2 :: int) * (numeral m OR NOT (numeral n))\<close>
   by (simp_all add: bit_eq_iff)
-    (auto simp add: bit_0 bit_simps bit_Suc bit_numeral_rec sub_inc_One_eq split: nat.split)
+    (auto simp: bit_0 bit_simps bit_Suc bit_numeral_rec sub_inc_One_eq split: nat.split)
 
 fun or_not_num_neg :: \<open>num \<Rightarrow> num \<Rightarrow> num\<close> \<^marker>\<open>contributor \<open>Andreas Lochbihler\<close>\<close>
 where
@@ -3055,7 +3046,7 @@
     (case take_bit_num (pred_numeral r) m of None \<Rightarrow> None | Some q \<Rightarrow> Some (Num.Bit0 q))\<close>
   \<open>take_bit_num (numeral r) (Num.Bit1 m) =
     Some (case take_bit_num (pred_numeral r) m of None \<Rightarrow> Num.One | Some q \<Rightarrow> Num.Bit1 q)\<close>
-  by (auto simp add: take_bit_num_def ac_simps mult_2 num_of_nat_double
+  by (auto simp: take_bit_num_def ac_simps mult_2 num_of_nat_double
     take_bit_Suc_bit0 take_bit_Suc_bit1 take_bit_numeral_bit0 take_bit_numeral_bit1)
 
 lemma take_bit_num_code [code]:
@@ -3085,7 +3076,7 @@
   \<open>take_bit m (numeral n) = numeral q\<close> if \<open>take_bit_num m n = Some q\<close>
 proof -
   from that have \<open>take_bit m (numeral n :: nat) = numeral q\<close>
-    by (auto simp add: take_bit_num_def Num.numeral_num_of_nat_unfold split: if_splits)
+    by (auto simp: take_bit_num_def Num.numeral_num_of_nat_unfold split: if_splits)
   then have \<open>of_nat (take_bit m (numeral n)) = of_nat (numeral q)\<close>
     by simp
   then show ?thesis
@@ -3193,7 +3184,7 @@
   \<open>Int.Neg num.One AND Int.Pos m = Int.Pos m\<close>
   \<open>Int.Neg (num.Bit0 n) AND Int.Pos m = Num.sub (or_not_num_neg (Num.BitM n) m) num.One\<close>
   \<open>Int.Neg (num.Bit1 n) AND Int.Pos m = Num.sub (or_not_num_neg (num.Bit0 n) m) num.One\<close>
-  apply (auto simp add: and_num_eq_None_iff [where ?'a = int] and_num_eq_Some_iff [where ?'a = int]
+  apply (auto simp: and_num_eq_None_iff [where ?'a = int] and_num_eq_Some_iff [where ?'a = int]
     split: option.split)
      apply (simp_all only: sub_one_eq_not_neg numeral_or_not_num_eq minus_minus and_not_numerals
        bit.de_Morgan_disj bit.double_compl and_not_num_eq_None_iff and_not_num_eq_Some_iff ac_simps)
@@ -3236,7 +3227,7 @@
   \<open>Int.Neg num.One OR Int.Pos m = Int.Neg num.One\<close>
   \<open>Int.Neg (num.Bit0 n) OR Int.Pos m = (case and_not_num (Num.BitM n) m of None \<Rightarrow> -1 | Some n' \<Rightarrow> Int.Neg (Num.inc n'))\<close>
   \<open>Int.Neg (num.Bit1 n) OR Int.Pos m = (case and_not_num (num.Bit0 n) m of None \<Rightarrow> -1 | Some n' \<Rightarrow> Int.Neg (Num.inc n'))\<close>
-  apply (auto simp add: numeral_or_num_eq split: option.splits)
+  apply (auto simp: numeral_or_num_eq split: option.splits)
          apply (simp_all only: and_not_num_eq_None_iff and_not_num_eq_Some_iff and_not_numerals
            numeral_or_not_num_eq or_eq_not_not_and bit.double_compl ac_simps flip: numeral_eq_iff [where ?'a = int])
          apply simp_all
@@ -3372,11 +3363,11 @@
 
 lemma concat_bit_assoc:
   \<open>concat_bit n k (concat_bit m l r) = concat_bit (m + n) (concat_bit n k l) r\<close>
-  by (rule bit_eqI) (auto simp add: bit_concat_bit_iff ac_simps)
+  by (rule bit_eqI) (auto simp: bit_concat_bit_iff ac_simps)
 
 lemma concat_bit_assoc_sym:
   \<open>concat_bit m (concat_bit n k l) r = concat_bit (min m n) k (concat_bit (m - n) l r)\<close>
-  by (rule bit_eqI) (auto simp add: bit_concat_bit_iff ac_simps min_def)
+  by (rule bit_eqI) (auto simp: bit_concat_bit_iff ac_simps min_def)
 
 lemma concat_bit_eq_iff:
   \<open>concat_bit n k l = concat_bit n r s
@@ -3394,14 +3385,14 @@
     fix m
     from * [of m]
     show \<open>bit (take_bit n k) m \<longleftrightarrow> bit (take_bit n r) m\<close>
-      by (auto simp add: bit_take_bit_iff bit_concat_bit_iff)
+      by (auto simp: bit_take_bit_iff bit_concat_bit_iff)
   qed
   moreover have \<open>push_bit n l = push_bit n s\<close>
   proof (rule bit_eqI)
     fix m
     from * [of m]
     show \<open>bit (push_bit n l) m \<longleftrightarrow> bit (push_bit n s) m\<close>
-      by (auto simp add: bit_push_bit_iff bit_concat_bit_iff)
+      by (auto simp: bit_push_bit_iff bit_concat_bit_iff)
   qed
   then have \<open>l = s\<close>
     by (simp add: push_bit_eq_mult)
@@ -3412,7 +3403,7 @@
 lemma take_bit_concat_bit_eq:
   \<open>take_bit m (concat_bit n k l) = concat_bit (min m n) k (take_bit (m - n) l)\<close>
   by (rule bit_eqI)
-    (auto simp add: bit_take_bit_iff bit_concat_bit_iff min_def)
+    (auto simp: bit_take_bit_iff bit_concat_bit_iff min_def)
 
 lemma concat_bit_take_bit_eq:
   \<open>concat_bit n (take_bit n b) = concat_bit n b\<close>
@@ -3437,7 +3428,7 @@
 
 lemma even_signed_take_bit_iff:
   \<open>even (signed_take_bit m a) \<longleftrightarrow> even a\<close>
-  by (auto simp add: bit_0 signed_take_bit_def even_or_iff even_mask_iff bit_double_iff)
+  by (auto simp: bit_0 signed_take_bit_def even_or_iff even_mask_iff bit_double_iff)
 
 lemma bit_signed_take_bit_iff [bit_simps]:
   \<open>bit (signed_take_bit m a) n \<longleftrightarrow> possible_bit TYPE('a) n \<and> bit a (min m n)\<close>
@@ -3479,21 +3470,21 @@
     by (simp add: fun_eq_iff bit_signed_take_bit_iff bit_take_bit_iff not_le less_Suc_eq_le min_def)
       (use bit_imp_possible_bit in fastforce)
   then show ?thesis
-    by (auto simp add: fun_eq_iff intro: bit_eqI)
+    by (auto simp: fun_eq_iff intro: bit_eqI)
 qed
 
 lemma signed_take_bit_signed_take_bit [simp]:
   \<open>signed_take_bit m (signed_take_bit n a) = signed_take_bit (min m n) a\<close>
-  by (auto simp add: bit_eq_iff bit_simps ac_simps)
+  by (auto simp: bit_eq_iff bit_simps ac_simps)
 
 lemma signed_take_bit_take_bit:
   \<open>signed_take_bit m (take_bit n a) = (if n \<le> m then take_bit n else signed_take_bit m) a\<close>
-  by (rule bit_eqI) (auto simp add: bit_signed_take_bit_iff min_def bit_take_bit_iff)
+  by (rule bit_eqI) (auto simp: bit_signed_take_bit_iff min_def bit_take_bit_iff)
 
 lemma take_bit_signed_take_bit:
   \<open>take_bit m (signed_take_bit n a) = take_bit m a\<close> if \<open>m \<le> Suc n\<close>
   using that by (rule le_SucE; intro bit_eqI)
-   (auto simp add: bit_take_bit_iff bit_signed_take_bit_iff min_def less_Suc_eq)
+   (auto simp: bit_take_bit_iff bit_signed_take_bit_iff min_def less_Suc_eq)
 
 lemma signed_take_bit_eq_take_bit_add:
   \<open>signed_take_bit n k = take_bit (Suc n) k + NOT (mask (Suc n)) * of_bool (bit k n)\<close>
@@ -3504,7 +3495,7 @@
 next
   case True
   have \<open>signed_take_bit n k = take_bit (Suc n) k OR NOT (mask (Suc n))\<close>
-    by (rule bit_eqI) (auto simp add: bit_signed_take_bit_iff min_def bit_take_bit_iff bit_or_iff bit_not_iff bit_mask_iff less_Suc_eq True)
+    by (rule bit_eqI) (auto simp: bit_signed_take_bit_iff min_def bit_take_bit_iff bit_or_iff bit_not_iff bit_mask_iff less_Suc_eq True)
   also have \<open>\<dots> = take_bit (Suc n) k + NOT (mask (Suc n))\<close>
     by (simp add: disjunctive_add_eq_or bit_eq_iff bit_simps)
   finally show ?thesis
@@ -3623,7 +3614,7 @@
 lemma signed_take_bit_int_eq_self_iff:
   \<open>signed_take_bit n k = k \<longleftrightarrow> - (2 ^ n) \<le> k \<and> k < 2 ^ n\<close>
   for k :: int
-  by (auto simp add: signed_take_bit_eq_take_bit_shift take_bit_int_eq_self_iff algebra_simps)
+  by (auto simp: signed_take_bit_eq_take_bit_shift take_bit_int_eq_self_iff algebra_simps)
 
 lemma signed_take_bit_int_eq_self:
   \<open>signed_take_bit n k = k\<close> if \<open>- (2 ^ n) \<le> k\<close> \<open>k < 2 ^ n\<close>
@@ -3782,9 +3773,8 @@
 
 lemma exp_div_exp_eq:
   \<open>2 ^ m div 2 ^ n = of_bool (2 ^ m \<noteq> 0 \<and> m \<ge> n) * 2 ^ (m - n)\<close>
-  apply (rule bit_eqI)
-  using bit_exp_iff div_exp_eq apply (auto simp add: bit_iff_odd possible_bit_def)
-  done
+  using bit_exp_iff div_exp_eq
+  by (intro bit_eqI) (auto simp: bit_iff_odd possible_bit_def)
 
 lemma bits_1_div_2:
   \<open>1 div 2 = 0\<close>
@@ -3855,11 +3845,11 @@
   with that show ?thesis proof (induction n arbitrary: a b)
     case 0
     from "0.prems"(1) [of 0] show ?case
-      by (auto simp add: bit_0)
+      by (auto simp: bit_0)
   next
     case (Suc n)
     from Suc.prems(1) [of 0] have even: \<open>even a \<or> even b\<close>
-      by (auto simp add: bit_0)
+      by (auto simp: bit_0)
     have bit: \<open>\<not> bit (a div 2) n \<or> \<not> bit (b div 2) n\<close> for n
       using Suc.prems(1) [of \<open>Suc n\<close>] by (simp add: bit_Suc)
     from Suc.prems(2) have \<open>possible_bit TYPE('a) (Suc n)\<close> \<open>possible_bit TYPE('a) n\<close>
@@ -3867,7 +3857,7 @@
     have \<open>a + b = (a div 2 * 2 + a mod 2) + (b div 2 * 2 + b mod 2)\<close>
       using div_mult_mod_eq [of a 2] div_mult_mod_eq [of b 2] by simp
     also have \<open>\<dots> = of_bool (odd a \<or> odd b) + 2 * (a div 2 + b div 2)\<close>
-      using even by (auto simp add: algebra_simps mod2_eq_if)
+      using even by (auto simp: algebra_simps mod2_eq_if)
     finally have \<open>bit ((a + b) div 2) n \<longleftrightarrow> bit (a div 2 + b div 2) n\<close>
       using \<open>possible_bit TYPE('a) (Suc n)\<close> by simp (simp_all flip: bit_Suc add: bit_double_iff possible_bit_def)
     also have \<open>\<dots> \<longleftrightarrow> bit (a div 2) n \<or> bit (b div 2) n\<close>
@@ -3884,7 +3874,7 @@
 
 lemma even_mask_div_iff:
   \<open>even ((2 ^ m - 1) div 2 ^ n) \<longleftrightarrow> 2 ^ n = 0 \<or> m \<le> n\<close>
-  using bit_mask_iff [of m n] by (auto simp add: mask_eq_exp_minus_1 bit_iff_odd possible_bit_def)
+  using bit_mask_iff [of m n] by (auto simp: mask_eq_exp_minus_1 bit_iff_odd possible_bit_def)
 
 lemma mod_exp_eq:
   \<open>a mod 2 ^ m mod 2 ^ n = a mod 2 ^ min m n\<close>
@@ -3920,7 +3910,7 @@
 
 lemma even_mod_exp_div_exp_iff:
   \<open>even (a mod 2 ^ m div 2 ^ n) \<longleftrightarrow> m \<le> n \<or> even (a div 2 ^ n)\<close>
-  by (auto simp add: even_drop_bit_iff_not_bit bit_simps simp flip: drop_bit_eq_div take_bit_eq_mod)
+  by (auto simp: even_drop_bit_iff_not_bit bit_simps simp flip: drop_bit_eq_div take_bit_eq_mod)
 
 end
 
@@ -3931,7 +3921,7 @@
   \<open>a - b = a AND NOT b\<close> if \<open>\<And>n. bit b n \<Longrightarrow> bit a n\<close>
 proof -
   have \<open>NOT a + b = NOT a OR b\<close>
-    by (rule disjunctive_add) (auto simp add: bit_not_iff dest: that)
+    by (rule disjunctive_add) (auto simp: bit_not_iff dest: that)
   then have \<open>NOT (NOT a + b) = NOT (NOT a OR b)\<close>
     by simp
   then show ?thesis
--- a/src/HOL/Codegenerator_Test/Code_Lazy_Test.thy	Mon Aug 26 18:26:00 2024 +0200
+++ b/src/HOL/Codegenerator_Test/Code_Lazy_Test.thy	Mon Aug 26 18:26:06 2024 +0200
@@ -24,6 +24,7 @@
 
 no_notation lazy_llist ("_")
 syntax "_llist" :: "args => 'a list"    ("\<^bold>[(_)\<^bold>]")
+syntax_consts "_llist" == lazy_llist
 translations
   "\<^bold>[x, xs\<^bold>]" == "x\<^bold>#\<^bold>[xs\<^bold>]"
   "\<^bold>[x\<^bold>]" == "x\<^bold>#\<^bold>[\<^bold>]"
--- a/src/HOL/Complete_Lattices.thy	Mon Aug 26 18:26:00 2024 +0200
+++ b/src/HOL/Complete_Lattices.thy	Mon Aug 26 18:26:06 2024 +0200
@@ -32,6 +32,10 @@
   "_SUP1"     :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b"           ("(3\<Squnion>_./ _)" [0, 10] 10)
   "_SUP"      :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b"  ("(3\<Squnion>_\<in>_./ _)" [0, 0, 10] 10)
 
+syntax_consts
+  "_INF1" "_INF" \<rightleftharpoons> Inf and
+  "_SUP1" "_SUP" \<rightleftharpoons> Sup
+
 translations
   "\<Sqinter>x y. f"   \<rightleftharpoons> "\<Sqinter>x. \<Sqinter>y. f"
   "\<Sqinter>x. f"     \<rightleftharpoons> "\<Sqinter>(CONST range (\<lambda>x. f))"
@@ -850,6 +854,9 @@
   "_INTER1"     :: "pttrns \<Rightarrow> 'b set \<Rightarrow> 'b set"           ("(3\<Inter>(\<open>unbreakable\<close>\<^bsub>_\<^esub>)/ _)" [0, 10] 10)
   "_INTER"      :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b set \<Rightarrow> 'b set"  ("(3\<Inter>(\<open>unbreakable\<close>\<^bsub>_\<in>_\<^esub>)/ _)" [0, 0, 10] 10)
 
+syntax_consts
+  "_INTER1" "_INTER" \<rightleftharpoons> Inter
+
 translations
   "\<Inter>x y. f"  \<rightleftharpoons> "\<Inter>x. \<Inter>y. f"
   "\<Inter>x. f"    \<rightleftharpoons> "\<Inter>(CONST range (\<lambda>x. f))"
@@ -1013,6 +1020,9 @@
   "_UNION1"     :: "pttrns => 'b set => 'b set"           ("(3\<Union>(\<open>unbreakable\<close>\<^bsub>_\<^esub>)/ _)" [0, 10] 10)
   "_UNION"      :: "pttrn => 'a set => 'b set => 'b set"  ("(3\<Union>(\<open>unbreakable\<close>\<^bsub>_\<in>_\<^esub>)/ _)" [0, 0, 10] 10)
 
+syntax_consts
+  "_UNION1" "_UNION" \<rightleftharpoons> Union
+
 translations
   "\<Union>x y. f"   \<rightleftharpoons> "\<Union>x. \<Union>y. f"
   "\<Union>x. f"     \<rightleftharpoons> "\<Union>(CONST range (\<lambda>x. f))"
--- a/src/HOL/Filter.thy	Mon Aug 26 18:26:00 2024 +0200
+++ b/src/HOL/Filter.thy	Mon Aug 26 18:26:06 2024 +0200
@@ -41,6 +41,8 @@
 
 syntax
   "_eventually" :: "pttrn => 'a filter => bool => bool"  ("(3\<forall>\<^sub>F _ in _./ _)" [0, 0, 10] 10)
+syntax_consts
+  "_eventually" == eventually
 translations
   "\<forall>\<^sub>Fx in F. P" == "CONST eventually (\<lambda>x. P) F"
 
@@ -159,6 +161,8 @@
 
 syntax
   "_frequently" :: "pttrn \<Rightarrow> 'a filter \<Rightarrow> bool \<Rightarrow> bool"  ("(3\<exists>\<^sub>F _ in _./ _)" [0, 0, 10] 10)
+syntax_consts
+  "_frequently" == frequently
 translations
   "\<exists>\<^sub>Fx in F. P" == "CONST frequently (\<lambda>x. P) F"
 
@@ -1305,6 +1309,9 @@
 syntax
   "_LIM" :: "pttrns \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'a \<Rightarrow> bool" ("(3LIM (_)/ (_)./ (_) :> (_))" [1000, 10, 0, 10] 10)
 
+syntax_consts
+  "_LIM" == filterlim
+
 translations
   "LIM x F1. f :> F2" == "CONST filterlim (\<lambda>x. f) F2 F1"
 
--- a/src/HOL/Fun.thy	Mon Aug 26 18:26:00 2024 +0200
+++ b/src/HOL/Fun.thy	Mon Aug 26 18:26:06 2024 +0200
@@ -846,6 +846,9 @@
   "_updbinds":: "updbind \<Rightarrow> updbinds \<Rightarrow> updbinds" ("_,/ _")
   "_Update"  :: "'a \<Rightarrow> updbinds \<Rightarrow> 'a"            ("_/'((_)')" [1000, 0] 900)
 
+syntax_consts
+  "_updbind" "_updbinds" "_Update" \<rightleftharpoons> fun_upd
+
 translations
   "_Update f (_updbinds b bs)" \<rightleftharpoons> "_Update (_Update f b) bs"
   "f(x:=y)" \<rightleftharpoons> "CONST fun_upd f x y"
--- a/src/HOL/GCD.thy	Mon Aug 26 18:26:00 2024 +0200
+++ b/src/HOL/GCD.thy	Mon Aug 26 18:26:06 2024 +0200
@@ -153,6 +153,10 @@
   "_LCM1"     :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b"           ("(3LCM _./ _)" [0, 10] 10)
   "_LCM"      :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b"  ("(3LCM _\<in>_./ _)" [0, 0, 10] 10)
 
+syntax_consts
+  "_GCD1" "_GCD" \<rightleftharpoons> Gcd and
+  "_LCM1" "_LCM" \<rightleftharpoons> Lcm
+
 translations
   "GCD x y. f"   \<rightleftharpoons> "GCD x. GCD y. f"
   "GCD x. f"     \<rightleftharpoons> "CONST Gcd (CONST range (\<lambda>x. f))"
@@ -2847,18 +2851,23 @@
 definition (in semiring_1) semiring_char :: "'a itself \<Rightarrow> nat" 
   where "semiring_char _ = Gcd {n. of_nat n = (0 :: 'a)}"
 
+syntax "_type_char" :: "type => nat" ("(1CHAR/(1'(_')))")
+syntax_consts "_type_char" \<rightleftharpoons> semiring_char
+translations "CHAR('t)" \<rightharpoonup> "CONST semiring_char (CONST Pure.type :: 't itself)"
+print_translation \<open>
+  let
+    fun char_type_tr' ctxt [Const (\<^const_syntax>\<open>Pure.type\<close>, Type (_, [T]))] =
+      Syntax.const \<^syntax_const>\<open>_type_char\<close> $ Syntax_Phases.term_of_typ ctxt T
+  in [(\<^const_syntax>\<open>semiring_char\<close>, char_type_tr')] end
+\<close>
+
 context semiring_1
 begin
 
-context
-  fixes CHAR :: nat
-  defines "CHAR \<equiv> semiring_char (Pure.type :: 'a itself)"
-begin
-
-lemma of_nat_CHAR [simp]: "of_nat CHAR = (0 :: 'a)"
+lemma of_nat_CHAR [simp]: "of_nat CHAR('a) = (0 :: 'a)"
 proof -
-  have "CHAR \<in> {n. of_nat n = (0::'a)}"
-    unfolding CHAR_def semiring_char_def
+  have "CHAR('a) \<in> {n. of_nat n = (0::'a)}"
+    unfolding semiring_char_def
   proof (rule Gcd_in, clarify)
     fix a b :: nat
     assume *: "of_nat a = (0 :: 'a)" "of_nat b = (0 :: 'a)"
@@ -2882,14 +2891,14 @@
     by simp
 qed
 
-lemma of_nat_eq_0_iff_char_dvd: "of_nat n = (0 :: 'a) \<longleftrightarrow> CHAR dvd n"
+lemma of_nat_eq_0_iff_char_dvd: "of_nat n = (0 :: 'a) \<longleftrightarrow> CHAR('a) dvd n"
 proof
   assume "of_nat n = (0 :: 'a)"
-  thus "CHAR dvd n"
-    unfolding CHAR_def semiring_char_def by (intro Gcd_dvd) auto
+  thus "CHAR('a) dvd n"
+    unfolding semiring_char_def by (intro Gcd_dvd) auto
 next
-  assume "CHAR dvd n"
-  then obtain m where "n = CHAR * m"
+  assume "CHAR('a) dvd n"
+  then obtain m where "n = CHAR('a) * m"
     by auto
   thus "of_nat n = (0 :: 'a)"
     by simp
@@ -2898,51 +2907,36 @@
 lemma CHAR_eqI:
   assumes "of_nat c = (0 :: 'a)"
   assumes "\<And>x. of_nat x = (0 :: 'a) \<Longrightarrow> c dvd x"
-  shows   "CHAR = c"
+  shows   "CHAR('a) = c"
   using assms by (intro dvd_antisym) (auto simp: of_nat_eq_0_iff_char_dvd)
 
-lemma CHAR_eq0_iff: "CHAR = 0 \<longleftrightarrow> (\<forall>n>0. of_nat n \<noteq> (0::'a))"
+lemma CHAR_eq0_iff: "CHAR('a) = 0 \<longleftrightarrow> (\<forall>n>0. of_nat n \<noteq> (0::'a))"
   by (auto simp: of_nat_eq_0_iff_char_dvd)
 
-lemma CHAR_pos_iff: "CHAR > 0 \<longleftrightarrow> (\<exists>n>0. of_nat n = (0::'a))"
+lemma CHAR_pos_iff: "CHAR('a) > 0 \<longleftrightarrow> (\<exists>n>0. of_nat n = (0::'a))"
   using CHAR_eq0_iff neq0_conv by blast
 
 lemma CHAR_eq_posI:
   assumes "c > 0" "of_nat c = (0 :: 'a)" "\<And>x. x > 0 \<Longrightarrow> x < c \<Longrightarrow> of_nat x \<noteq> (0 :: 'a)"
-  shows   "CHAR = c"
+  shows   "CHAR('a) = c"
 proof (rule antisym)
-  from assms have "CHAR > 0"
+  from assms have "CHAR('a) > 0"
     by (auto simp: CHAR_pos_iff)
-  from assms(3)[OF this] show "CHAR \<ge> c"
+  from assms(3)[OF this] show "CHAR('a) \<ge> c"
     by force
 next
-  have "CHAR dvd c"
+  have "CHAR('a) dvd c"
     using assms by (auto simp: of_nat_eq_0_iff_char_dvd)
-  thus "CHAR \<le> c"
+  thus "CHAR('a) \<le> c"
     using \<open>c > 0\<close> by (intro dvd_imp_le) auto
 qed
 
 end
-end
-
-lemma (in semiring_char_0) CHAR_eq_0 [simp]: "semiring_char (Pure.type :: 'a itself) = 0"
+
+lemma (in semiring_char_0) CHAR_eq_0 [simp]: "CHAR('a) = 0"
   by (simp add: CHAR_eq0_iff)
 
 
-(* nicer notation *)
-
-syntax "_type_char" :: "type => nat" ("(1CHAR/(1'(_')))")
-
-translations "CHAR('t)" => "CONST semiring_char (CONST Pure.type :: 't itself)"
-
-print_translation \<open>
-  let
-    fun char_type_tr' ctxt [Const (\<^const_syntax>\<open>Pure.type\<close>, Type (_, [T]))] =
-      Syntax.const \<^syntax_const>\<open>_type_char\<close> $ Syntax_Phases.term_of_typ ctxt T
-  in [(\<^const_syntax>\<open>semiring_char\<close>, char_type_tr')] end
-\<close>
-
-
 lemma CHAR_not_1 [simp]: "CHAR('a :: {semiring_1, zero_neq_one}) \<noteq> Suc 0"
   by (metis One_nat_def of_nat_1 of_nat_CHAR zero_neq_one)
 
--- a/src/HOL/Groups_Big.thy	Mon Aug 26 18:26:00 2024 +0200
+++ b/src/HOL/Groups_Big.thy	Mon Aug 26 18:26:06 2024 +0200
@@ -654,32 +654,41 @@
   "_sum" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b::comm_monoid_add"  ("(3SUM (_/:_)./ _)" [0, 51, 10] 10)
 syntax
   "_sum" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b::comm_monoid_add"  ("(2\<Sum>(_/\<in>_)./ _)" [0, 51, 10] 10)
+
+syntax_consts
+  "_sum" \<rightleftharpoons> sum
+
 translations \<comment> \<open>Beware of argument permutation!\<close>
   "\<Sum>i\<in>A. b" \<rightleftharpoons> "CONST sum (\<lambda>i. b) A"
 
+
 text \<open>Instead of \<^term>\<open>\<Sum>x\<in>{x. P}. e\<close> we introduce the shorter \<open>\<Sum>x|P. e\<close>.\<close>
 
 syntax (ASCII)
   "_qsum" :: "pttrn \<Rightarrow> bool \<Rightarrow> 'a \<Rightarrow> 'a"  ("(3SUM _ |/ _./ _)" [0, 0, 10] 10)
 syntax
   "_qsum" :: "pttrn \<Rightarrow> bool \<Rightarrow> 'a \<Rightarrow> 'a"  ("(2\<Sum>_ | (_)./ _)" [0, 0, 10] 10)
+
+syntax_consts
+  "_qsum" == sum
+
 translations
   "\<Sum>x|P. t" => "CONST sum (\<lambda>x. t) {x. P}"
 
 print_translation \<open>
-let
-  fun sum_tr' [Abs (x, Tx, t), Const (\<^const_syntax>\<open>Collect\<close>, _) $ Abs (y, Ty, P)] =
-        if x <> y then raise Match
-        else
-          let
-            val x' = Syntax_Trans.mark_bound_body (x, Tx);
-            val t' = subst_bound (x', t);
-            val P' = subst_bound (x', P);
-          in
-            Syntax.const \<^syntax_const>\<open>_qsum\<close> $ Syntax_Trans.mark_bound_abs (x, Tx) $ P' $ t'
-          end
-    | sum_tr' _ = raise Match;
-in [(\<^const_syntax>\<open>sum\<close>, K sum_tr')] end
+  let
+    fun sum_tr' [Abs (x, Tx, t), Const (\<^const_syntax>\<open>Collect\<close>, _) $ Abs (y, Ty, P)] =
+          if x <> y then raise Match
+          else
+            let
+              val x' = Syntax_Trans.mark_bound_body (x, Tx);
+              val t' = subst_bound (x', t);
+              val P' = subst_bound (x', P);
+            in
+              Syntax.const \<^syntax_const>\<open>_qsum\<close> $ Syntax_Trans.mark_bound_abs (x, Tx) $ P' $ t'
+            end
+      | sum_tr' _ = raise Match;
+  in [(\<^const_syntax>\<open>sum\<close>, K sum_tr')] end
 \<close>
 
 
@@ -1415,18 +1424,43 @@
   "_prod" :: "pttrn => 'a set => 'b => 'b::comm_monoid_mult"  ("(4PROD (_/:_)./ _)" [0, 51, 10] 10)
 syntax
   "_prod" :: "pttrn => 'a set => 'b => 'b::comm_monoid_mult"  ("(2\<Prod>(_/\<in>_)./ _)" [0, 51, 10] 10)
+
+syntax_consts
+  "_prod" == prod
+
 translations \<comment> \<open>Beware of argument permutation!\<close>
   "\<Prod>i\<in>A. b" == "CONST prod (\<lambda>i. b) A"
 
+
 text \<open>Instead of \<^term>\<open>\<Prod>x\<in>{x. P}. e\<close> we introduce the shorter \<open>\<Prod>x|P. e\<close>.\<close>
 
 syntax (ASCII)
   "_qprod" :: "pttrn \<Rightarrow> bool \<Rightarrow> 'a \<Rightarrow> 'a"  ("(4PROD _ |/ _./ _)" [0, 0, 10] 10)
 syntax
   "_qprod" :: "pttrn \<Rightarrow> bool \<Rightarrow> 'a \<Rightarrow> 'a"  ("(2\<Prod>_ | (_)./ _)" [0, 0, 10] 10)
+
+syntax_consts
+  "_qprod" == prod
+
 translations
   "\<Prod>x|P. t" => "CONST prod (\<lambda>x. t) {x. P}"
 
+print_translation \<open>
+  let
+    fun prod_tr' [Abs (x, Tx, t), Const (\<^const_syntax>\<open>Collect\<close>, _) $ Abs (y, Ty, P)] =
+          if x <> y then raise Match
+          else
+            let
+              val x' = Syntax_Trans.mark_bound_body (x, Tx);
+              val t' = subst_bound (x', t);
+              val P' = subst_bound (x', P);
+            in
+              Syntax.const \<^syntax_const>\<open>_qprod\<close> $ Syntax_Trans.mark_bound_abs (x, Tx) $ P' $ t'
+            end
+      | prod_tr' _ = raise Match;
+  in [(\<^const_syntax>\<open>prod\<close>, K prod_tr')] end
+\<close>
+
 context comm_monoid_mult
 begin
 
--- a/src/HOL/Groups_List.thy	Mon Aug 26 18:26:00 2024 +0200
+++ b/src/HOL/Groups_List.thy	Mon Aug 26 18:26:06 2024 +0200
@@ -101,6 +101,8 @@
   "_sum_list" :: "pttrn => 'a list => 'b => 'b"    ("(3SUM _<-_. _)" [0, 51, 10] 10)
 syntax
   "_sum_list" :: "pttrn => 'a list => 'b => 'b"    ("(3\<Sum>_\<leftarrow>_. _)" [0, 51, 10] 10)
+syntax_consts
+  "_sum_list" == sum_list
 translations \<comment> \<open>Beware of argument permutation!\<close>
   "\<Sum>x\<leftarrow>xs. b" == "CONST sum_list (CONST map (\<lambda>x. b) xs)"
 
@@ -598,6 +600,8 @@
   "_prod_list" :: "pttrn => 'a list => 'b => 'b"    ("(3PROD _<-_. _)" [0, 51, 10] 10)
 syntax
   "_prod_list" :: "pttrn => 'a list => 'b => 'b"    ("(3\<Prod>_\<leftarrow>_. _)" [0, 51, 10] 10)
+syntax_consts
+  "_prod_list" \<rightleftharpoons> prod_list
 translations \<comment> \<open>Beware of argument permutation!\<close>
   "\<Prod>x\<leftarrow>xs. b" \<rightleftharpoons> "CONST prod_list (CONST map (\<lambda>x. b) xs)"
 
--- a/src/HOL/HOL.thy	Mon Aug 26 18:26:00 2024 +0200
+++ b/src/HOL/HOL.thy	Mon Aug 26 18:26:06 2024 +0200
@@ -129,6 +129,9 @@
 
 syntax (ASCII) "_Uniq" :: "pttrn \<Rightarrow> bool \<Rightarrow> bool"  ("(4?< _./ _)" [0, 10] 10)
 syntax "_Uniq" :: "pttrn \<Rightarrow> bool \<Rightarrow> bool"  ("(2\<exists>\<^sub>\<le>\<^sub>1 _./ _)" [0, 10] 10)
+
+syntax_consts "_Uniq" \<rightleftharpoons> Uniq
+
 translations "\<exists>\<^sub>\<le>\<^sub>1x. P" \<rightleftharpoons> "CONST Uniq (\<lambda>x. P)"
 
 print_translation \<open>
@@ -141,6 +144,9 @@
 syntax (input)
   "_Ex1" :: "pttrn \<Rightarrow> bool \<Rightarrow> bool"  ("(3?! _./ _)" [0, 10] 10)
 syntax "_Ex1" :: "pttrn \<Rightarrow> bool \<Rightarrow> bool"  ("(3\<exists>!_./ _)" [0, 10] 10)
+
+syntax_consts "_Ex1" \<rightleftharpoons> Ex1
+
 translations "\<exists>!x. P" \<rightleftharpoons> "CONST Ex1 (\<lambda>x. P)"
 
 print_translation \<open>
@@ -151,6 +157,9 @@
 syntax
   "_Not_Ex" :: "idts \<Rightarrow> bool \<Rightarrow> bool"  ("(3\<nexists>_./ _)" [0, 10] 10)
   "_Not_Ex1" :: "pttrn \<Rightarrow> bool \<Rightarrow> bool"  ("(3\<nexists>!_./ _)" [0, 10] 10)
+syntax_consts
+  "_Not_Ex" \<rightleftharpoons> Ex and
+  "_Not_Ex1" \<rightleftharpoons> Ex1
 translations
   "\<nexists>x. P" \<rightleftharpoons> "\<not> (\<exists>x. P)"
   "\<nexists>!x. P" \<rightleftharpoons> "\<not> (\<exists>!x. P)"
@@ -171,6 +180,7 @@
   where "A \<longleftrightarrow> B \<equiv> A = B"
 
 syntax "_The" :: "[pttrn, bool] \<Rightarrow> 'a"  ("(3THE _./ _)" [0, 10] 10)
+syntax_consts "_The" \<rightleftharpoons> The
 translations "THE x. P" \<rightleftharpoons> "CONST The (\<lambda>x. P)"
 print_translation \<open>
   [(\<^const_syntax>\<open>The\<close>, fn _ => fn [Abs abs] =>
@@ -178,13 +188,6 @@
       in Syntax.const \<^syntax_const>\<open>_The\<close> $ x $ t end)]
 \<close>  \<comment> \<open>To avoid eta-contraction of body\<close>
 
-nonterminal letbinds and letbind
-syntax
-  "_bind"       :: "[pttrn, 'a] \<Rightarrow> letbind"              ("(2_ =/ _)" 10)
-  ""            :: "letbind \<Rightarrow> letbinds"                 ("_")
-  "_binds"      :: "[letbind, letbinds] \<Rightarrow> letbinds"     ("_;/ _")
-  "_Let"        :: "[letbinds, 'a] \<Rightarrow> 'a"                ("(let (_)/ in (_))" [0, 10] 10)
-
 nonterminal case_syn and cases_syn
 syntax
   "_case_syntax" :: "['a, cases_syn] \<Rightarrow> 'b"  ("(case _ of/ _)" 10)
@@ -227,6 +230,14 @@
 definition Let :: "'a \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'b"
   where "Let s f \<equiv> f s"
 
+nonterminal letbinds and letbind
+syntax
+  "_bind"       :: "[pttrn, 'a] \<Rightarrow> letbind"              ("(2_ =/ _)" 10)
+  ""            :: "letbind \<Rightarrow> letbinds"                 ("_")
+  "_binds"      :: "[letbind, letbinds] \<Rightarrow> letbinds"     ("_;/ _")
+  "_Let"        :: "[letbinds, 'a] \<Rightarrow> 'a"                ("(let (_)/ in (_))" [0, 10] 10)
+syntax_consts
+  "_bind" "_binds" "_Let" \<rightleftharpoons> Let
 translations
   "_Let (_binds b bs) e"  \<rightleftharpoons> "_Let b (_Let bs e)"
   "let x = a in e"        \<rightleftharpoons> "CONST Let a (\<lambda>x. e)"
--- a/src/HOL/HOLCF/Cfun.thy	Mon Aug 26 18:26:00 2024 +0200
+++ b/src/HOL/HOLCF/Cfun.thy	Mon Aug 26 18:26:06 2024 +0200
@@ -52,6 +52,9 @@
 syntax
   "_Lambda" :: "[cargs, logic] \<Rightarrow> logic" ("(3\<Lambda> _./ _)" [1000, 10] 10)
 
+syntax_consts
+  "_Lambda" \<rightleftharpoons> Abs_cfun
+
 parse_ast_translation \<open>
 (* rewrite (LAM x y z. t) => (_cabs x (_cabs y (_cabs z t))) *)
 (* cf. Syntax.lambda_ast_tr from src/Pure/Syntax/syn_trans.ML *)
--- a/src/HOL/HOLCF/ConvexPD.thy	Mon Aug 26 18:26:00 2024 +0200
+++ b/src/HOL/HOLCF/ConvexPD.thy	Mon Aug 26 18:26:06 2024 +0200
@@ -181,6 +181,9 @@
 syntax
   "_convex_pd" :: "args \<Rightarrow> logic" ("{_}\<natural>")
 
+syntax_consts
+  "_convex_pd" == convex_add
+
 translations
   "{x,xs}\<natural>" == "{x}\<natural> \<union>\<natural> {xs}\<natural>"
   "{x}\<natural>" == "CONST convex_unit\<cdot>x"
@@ -348,6 +351,9 @@
   "_convex_bind" :: "[logic, logic, logic] \<Rightarrow> logic"
     ("(3\<Union>\<natural>_\<in>_./ _)" [0, 0, 10] 10)
 
+syntax_consts
+  "_convex_bind" == convex_bind
+
 translations
   "\<Union>\<natural>x\<in>xs. e" == "CONST convex_bind\<cdot>xs\<cdot>(\<Lambda> x. e)"
 
--- a/src/HOL/HOLCF/IOA/Sequence.thy	Mon Aug 26 18:26:00 2024 +0200
+++ b/src/HOL/HOLCF/IOA/Sequence.thy	Mon Aug 26 18:26:06 2024 +0200
@@ -72,6 +72,8 @@
 syntax
   "_totlist" :: "args \<Rightarrow> 'a Seq"  ("[(_)!]")
   "_partlist" :: "args \<Rightarrow> 'a Seq"  ("[(_)?]")
+syntax_consts
+  "_totlist" "_partlist" \<rightleftharpoons> Consq
 translations
   "[x, xs!]" \<rightleftharpoons> "x \<leadsto> [xs!]"
   "[x!]" \<rightleftharpoons> "x\<leadsto>nil"
--- a/src/HOL/HOLCF/LowerPD.thy	Mon Aug 26 18:26:00 2024 +0200
+++ b/src/HOL/HOLCF/LowerPD.thy	Mon Aug 26 18:26:06 2024 +0200
@@ -136,6 +136,9 @@
 syntax
   "_lower_pd" :: "args \<Rightarrow> logic" ("{_}\<flat>")
 
+syntax_consts
+  "_lower_pd" == lower_add
+
 translations
   "{x,xs}\<flat>" == "{x}\<flat> \<union>\<flat> {xs}\<flat>"
   "{x}\<flat>" == "CONST lower_unit\<cdot>x"
@@ -342,6 +345,9 @@
   "_lower_bind" :: "[logic, logic, logic] \<Rightarrow> logic"
     ("(3\<Union>\<flat>_\<in>_./ _)" [0, 0, 10] 10)
 
+syntax_consts
+  "_lower_bind" == lower_bind
+
 translations
   "\<Union>\<flat>x\<in>xs. e" == "CONST lower_bind\<cdot>xs\<cdot>(\<Lambda> x. e)"
 
--- a/src/HOL/HOLCF/Pcpo.thy	Mon Aug 26 18:26:00 2024 +0200
+++ b/src/HOL/HOLCF/Pcpo.thy	Mon Aug 26 18:26:06 2024 +0200
@@ -153,6 +153,7 @@
 text \<open>Old "UU" syntax:\<close>
 
 syntax UU :: logic
+syntax_consts UU \<rightleftharpoons> bottom
 translations "UU" \<rightharpoonup> "CONST bottom"
 
 text \<open>Simproc to rewrite \<^term>\<open>\<bottom> = x\<close> to \<^term>\<open>x = \<bottom>\<close>.\<close>
--- a/src/HOL/HOLCF/Porder.thy	Mon Aug 26 18:26:00 2024 +0200
+++ b/src/HOL/HOLCF/Porder.thy	Mon Aug 26 18:26:06 2024 +0200
@@ -118,6 +118,9 @@
 syntax
   "_BLub" :: "[pttrn, 'a set, 'b] \<Rightarrow> 'b" ("(3\<Squnion>_\<in>_./ _)" [0,0, 10] 10)
 
+syntax_consts
+  "_BLub" \<rightleftharpoons> lub
+
 translations
   "LUB x:A. t" \<rightleftharpoons> "CONST lub ((\<lambda>x. t) ` A)"
 
--- a/src/HOL/HOLCF/Representable.thy	Mon Aug 26 18:26:00 2024 +0200
+++ b/src/HOL/HOLCF/Representable.thy	Mon Aug 26 18:26:06 2024 +0200
@@ -32,6 +32,7 @@
   assumes cast_liftdefl: "cast\<cdot>(liftdefl TYPE('a)) = liftemb oo liftprj"
 
 syntax "_LIFTDEFL" :: "type \<Rightarrow> logic"  ("(1LIFTDEFL/(1'(_')))")
+syntax_consts "_LIFTDEFL" \<rightleftharpoons> liftdefl
 translations "LIFTDEFL('t)" \<rightleftharpoons> "CONST liftdefl TYPE('t)"
 
 definition liftdefl_of :: "udom defl \<rightarrow> udom u defl"
@@ -51,6 +52,7 @@
   assumes liftdefl_eq: "liftdefl TYPE('a) = liftdefl_of\<cdot>(defl TYPE('a))"
 
 syntax "_DEFL" :: "type \<Rightarrow> logic"  ("(1DEFL/(1'(_')))")
+syntax_consts "_DEFL" \<rightleftharpoons> defl
 translations "DEFL('t)" \<rightleftharpoons> "CONST defl TYPE('t)"
 
 instance "domain" \<subseteq> predomain
--- a/src/HOL/HOLCF/Sprod.thy	Mon Aug 26 18:26:00 2024 +0200
+++ b/src/HOL/HOLCF/Sprod.thy	Mon Aug 26 18:26:06 2024 +0200
@@ -41,6 +41,7 @@
   where "ssplit = (\<Lambda> f p. seq\<cdot>p\<cdot>(f\<cdot>(sfst\<cdot>p)\<cdot>(ssnd\<cdot>p)))"
 
 syntax "_stuple" :: "[logic, args] \<Rightarrow> logic"  ("(1'(:_,/ _:'))")
+syntax_consts "_stuple" \<rightleftharpoons> spair
 translations
   "(:x, y, z:)" \<rightleftharpoons> "(:x, (:y, z:):)"
   "(:x, y:)" \<rightleftharpoons> "CONST spair\<cdot>x\<cdot>y"
--- a/src/HOL/HOLCF/UpperPD.thy	Mon Aug 26 18:26:00 2024 +0200
+++ b/src/HOL/HOLCF/UpperPD.thy	Mon Aug 26 18:26:06 2024 +0200
@@ -134,6 +134,9 @@
 syntax
   "_upper_pd" :: "args \<Rightarrow> logic" ("{_}\<sharp>")
 
+syntax_consts
+  "_upper_pd" == upper_add
+
 translations
   "{x,xs}\<sharp>" == "{x}\<sharp> \<union>\<sharp> {xs}\<sharp>"
   "{x}\<sharp>" == "CONST upper_unit\<cdot>x"
@@ -335,6 +338,9 @@
   "_upper_bind" :: "[logic, logic, logic] \<Rightarrow> logic"
     ("(3\<Union>\<sharp>_\<in>_./ _)" [0, 0, 10] 10)
 
+syntax_consts
+  "_upper_bind" == upper_bind
+
 translations
   "\<Union>\<sharp>x\<in>xs. e" == "CONST upper_bind\<cdot>xs\<cdot>(\<Lambda> x. e)"
 
--- a/src/HOL/HOLCF/ex/Letrec.thy	Mon Aug 26 18:26:00 2024 +0200
+++ b/src/HOL/HOLCF/ex/Letrec.thy	Mon Aug 26 18:26:06 2024 +0200
@@ -22,6 +22,9 @@
   "_recbinds" :: "[recbindt, recbinds] \<Rightarrow> recbinds"  ("_;/ _")
   "_Letrec"   :: "[recbinds, logic] \<Rightarrow> logic"        ("(Letrec (_)/ in (_))" 10)
 
+syntax_consts
+  "_recbind" "_recbinds" "_recbindt" "_Letrec" == CLetrec
+
 translations
   (recbindt) "x = a, (y,ys) = (b,bs)" == (recbindt) "(x,y,ys) = (a,b,bs)"
   (recbindt) "x = a, y = b"          == (recbindt) "(x,y) = (a,b)"
--- a/src/HOL/Hilbert_Choice.thy	Mon Aug 26 18:26:00 2024 +0200
+++ b/src/HOL/Hilbert_Choice.thy	Mon Aug 26 18:26:06 2024 +0200
@@ -22,6 +22,9 @@
   "_Eps" :: "pttrn \<Rightarrow> bool \<Rightarrow> 'a"  ("(3@ _./ _)" [0, 10] 10)
 syntax
   "_Eps" :: "pttrn \<Rightarrow> bool \<Rightarrow> 'a"  ("(3SOME _./ _)" [0, 10] 10)
+
+syntax_consts "_Eps" \<rightleftharpoons> Eps
+
 translations
   "SOME x. P" \<rightleftharpoons> "CONST Eps (\<lambda>x. P)"
 
--- a/src/HOL/Lattices_Big.thy	Mon Aug 26 18:26:00 2024 +0200
+++ b/src/HOL/Lattices_Big.thy	Mon Aug 26 18:26:06 2024 +0200
@@ -470,6 +470,10 @@
   "_MAX1"     :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b"           ("(3MAX _./ _)" [0, 10] 10)
   "_MAX"      :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b"  ("(3MAX _\<in>_./ _)" [0, 0, 10] 10)
 
+syntax_consts
+  "_MIN1" "_MIN" \<rightleftharpoons> Min and
+  "_MAX1" "_MAX" \<rightleftharpoons> Max
+
 translations
   "MIN x y. f"   \<rightleftharpoons> "MIN x. MIN y. f"
   "MIN x. f"     \<rightleftharpoons> "CONST Min (CONST range (\<lambda>x. f))"
@@ -478,6 +482,7 @@
   "MAX x. f"     \<rightleftharpoons> "CONST Max (CONST range (\<lambda>x. f))"
   "MAX x\<in>A. f"   \<rightleftharpoons> "CONST Max ((\<lambda>x. f) ` A)"
 
+
 text \<open>An aside: \<^const>\<open>Min\<close>/\<^const>\<open>Max\<close> on linear orders as special case of \<^const>\<open>Inf_fin\<close>/\<^const>\<open>Sup_fin\<close>\<close>
 
 lemma Inf_fin_Min:
@@ -919,6 +924,8 @@
 syntax
   "_arg_min" :: "('b \<Rightarrow> 'a) \<Rightarrow> pttrn \<Rightarrow> bool \<Rightarrow> 'b"
     ("(3ARG'_MIN _ _./ _)" [1000, 0, 10] 10)
+syntax_consts
+  "_arg_min" \<rightleftharpoons> arg_min
 translations
   "ARG_MIN f x. P" \<rightleftharpoons> "CONST arg_min f (\<lambda>x. P)"
 
@@ -935,41 +942,33 @@
     \<And>y. P y \<Longrightarrow> \<not> f y < f x;
     \<And>x. \<lbrakk> P x; \<forall>y. P y \<longrightarrow> \<not> f y < f x \<rbrakk> \<Longrightarrow> Q x \<rbrakk>
   \<Longrightarrow> Q (arg_min f P)"
-apply (simp add: arg_min_def is_arg_min_def)
-apply (rule someI2_ex)
- apply blast
-apply blast
-done
+  unfolding arg_min_def is_arg_min_def
+  by (blast intro!: someI2_ex)
 
 lemma arg_min_equality:
   "\<lbrakk> P k; \<And>x. P x \<Longrightarrow> f k \<le> f x \<rbrakk> \<Longrightarrow> f (arg_min f P) = f k"
   for f :: "_ \<Rightarrow> 'a::order"
-apply (rule arg_minI)
-  apply assumption
- apply (simp add: less_le_not_le)
-by (metis le_less)
+  by (rule arg_minI; force simp: not_less less_le_not_le)
 
 lemma wf_linord_ex_has_least:
   "\<lbrakk> wf r; \<forall>x y. (x, y) \<in> r\<^sup>+ \<longleftrightarrow> (y, x) \<notin> r\<^sup>*; P k \<rbrakk>
    \<Longrightarrow> \<exists>x. P x \<and> (\<forall>y. P y \<longrightarrow> (m x, m y) \<in> r\<^sup>*)"
-apply (drule wf_trancl [THEN wf_eq_minimal [THEN iffD1]])
-apply (drule_tac x = "m ` Collect P" in spec)
-by force
+  by (force dest!:  wf_trancl [THEN wf_eq_minimal [THEN iffD1, THEN spec], where x = "m ` Collect P"])
 
 lemma ex_has_least_nat: "P k \<Longrightarrow> \<exists>x. P x \<and> (\<forall>y. P y \<longrightarrow> m x \<le> m y)"
   for m :: "'a \<Rightarrow> nat"
-apply (simp only: pred_nat_trancl_eq_le [symmetric])
-apply (rule wf_pred_nat [THEN wf_linord_ex_has_least])
- apply (simp add: less_eq linorder_not_le pred_nat_trancl_eq_le)
-by assumption
+  unfolding pred_nat_trancl_eq_le [symmetric]
+  apply (rule wf_pred_nat [THEN wf_linord_ex_has_least])
+   apply (simp add: less_eq linorder_not_le pred_nat_trancl_eq_le)
+  by assumption
 
 lemma arg_min_nat_lemma:
   "P k \<Longrightarrow> P(arg_min m P) \<and> (\<forall>y. P y \<longrightarrow> m (arg_min m P) \<le> m y)"
   for m :: "'a \<Rightarrow> nat"
-apply (simp add: arg_min_def is_arg_min_linorder)
-apply (rule someI_ex)
-apply (erule ex_has_least_nat)
-done
+  unfolding arg_min_def is_arg_min_linorder
+  apply (rule someI_ex)
+  apply (erule ex_has_least_nat)
+  done
 
 lemmas arg_min_natI = arg_min_nat_lemma [THEN conjunct1]
 
@@ -979,35 +978,35 @@
 
 lemma arg_min_nat_le: "P x \<Longrightarrow> m (arg_min m P) \<le> m x"
   for m :: "'a \<Rightarrow> nat"
-by (rule arg_min_nat_lemma [THEN conjunct2, THEN spec, THEN mp])
+  by (rule arg_min_nat_lemma [THEN conjunct2, THEN spec, THEN mp])
 
 lemma ex_min_if_finite:
   "\<lbrakk> finite S; S \<noteq> {} \<rbrakk> \<Longrightarrow> \<exists>m\<in>S. \<not>(\<exists>x\<in>S. x < (m::'a::order))"
-by(induction rule: finite.induct) (auto intro: order.strict_trans)
+  by(induction rule: finite.induct) (auto intro: order.strict_trans)
 
 lemma ex_is_arg_min_if_finite: fixes f :: "'a \<Rightarrow> 'b :: order"
-shows "\<lbrakk> finite S; S \<noteq> {} \<rbrakk> \<Longrightarrow> \<exists>x. is_arg_min f (\<lambda>x. x \<in> S) x"
-unfolding is_arg_min_def
-using ex_min_if_finite[of "f ` S"]
-by auto
+  shows "\<lbrakk> finite S; S \<noteq> {} \<rbrakk> \<Longrightarrow> \<exists>x. is_arg_min f (\<lambda>x. x \<in> S) x"
+  unfolding is_arg_min_def
+  using ex_min_if_finite[of "f ` S"]
+  by auto
 
 lemma arg_min_SOME_Min:
   "finite S \<Longrightarrow> arg_min_on f S = (SOME y. y \<in> S \<and> f y = Min(f ` S))"
-unfolding arg_min_on_def arg_min_def is_arg_min_linorder
-apply(rule arg_cong[where f = Eps])
-apply (auto simp: fun_eq_iff intro: Min_eqI[symmetric])
-done
+  unfolding arg_min_on_def arg_min_def is_arg_min_linorder
+  apply(rule arg_cong[where f = Eps])
+  apply (auto simp: fun_eq_iff intro: Min_eqI[symmetric])
+  done
 
 lemma arg_min_if_finite: fixes f :: "'a \<Rightarrow> 'b :: order"
-assumes "finite S" "S \<noteq> {}"
-shows  "arg_min_on f S \<in> S" and "\<not>(\<exists>x\<in>S. f x < f (arg_min_on f S))"
-using ex_is_arg_min_if_finite[OF assms, of f]
-unfolding arg_min_on_def arg_min_def is_arg_min_def
-by(auto dest!: someI_ex)
+  assumes "finite S" "S \<noteq> {}"
+  shows  "arg_min_on f S \<in> S" and "\<not>(\<exists>x\<in>S. f x < f (arg_min_on f S))"
+  using ex_is_arg_min_if_finite[OF assms, of f]
+  unfolding arg_min_on_def arg_min_def is_arg_min_def
+  by(auto dest!: someI_ex)
 
 lemma arg_min_least: fixes f :: "'a \<Rightarrow> 'b :: linorder"
-shows "\<lbrakk> finite S;  S \<noteq> {};  y \<in> S \<rbrakk> \<Longrightarrow> f(arg_min_on f S) \<le> f y"
-by(simp add: arg_min_SOME_Min inv_into_def2[symmetric] f_inv_into_f)
+  shows "\<lbrakk> finite S;  S \<noteq> {};  y \<in> S \<rbrakk> \<Longrightarrow> f(arg_min_on f S) \<le> f y"
+  by(simp add: arg_min_SOME_Min inv_into_def2[symmetric] f_inv_into_f)
 
 lemma arg_min_inj_eq: fixes f :: "'a \<Rightarrow> 'b :: order"
 shows "\<lbrakk> inj_on f {x. P x}; P a; \<forall>y. P y \<longrightarrow> f a \<le> f y \<rbrakk> \<Longrightarrow> arg_min f P = a"
@@ -1036,6 +1035,8 @@
 syntax
   "_arg_max" :: "('b \<Rightarrow> 'a) \<Rightarrow> pttrn \<Rightarrow> bool \<Rightarrow> 'a"
     ("(3ARG'_MAX _ _./ _)" [1000, 0, 10] 10)
+syntax_consts
+  "_arg_max" \<rightleftharpoons> arg_max
 translations
   "ARG_MAX f x. P" \<rightleftharpoons> "CONST arg_max f (\<lambda>x. P)"
 
@@ -1048,11 +1049,8 @@
     (\<And>y. P y \<Longrightarrow> \<not> f y > f x) \<Longrightarrow>
     (\<And>x. P x \<Longrightarrow> \<forall>y. P y \<longrightarrow> \<not> f y > f x \<Longrightarrow> Q x) \<Longrightarrow>
     Q (arg_max f P)"
-apply (simp add: arg_max_def is_arg_max_def)
-apply (rule someI2_ex)
- apply blast
-apply blast
-done
+  unfolding arg_max_def is_arg_max_def
+  by (blast intro!: someI2_ex elim: )
 
 lemma arg_max_equality:
   "\<lbrakk> P k; \<And>x. P x \<Longrightarrow> f x \<le> f k \<rbrakk> \<Longrightarrow> f (arg_max f P) = f k"
@@ -1086,15 +1084,13 @@
   "\<lbrakk> P k;  \<forall>y. P y \<longrightarrow> f y < b \<rbrakk>
   \<Longrightarrow> P (arg_max f P) \<and> (\<forall>y. P y \<longrightarrow> f y \<le> f (arg_max f P))"
   for f :: "'a \<Rightarrow> nat"
-apply (simp add: arg_max_def is_arg_max_linorder)
-apply (rule someI_ex)
-apply (erule (1) ex_has_greatest_nat)
-done
+  unfolding arg_max_def is_arg_max_linorder
+  by (rule someI_ex) (metis ex_has_greatest_nat)
 
 lemmas arg_max_natI = arg_max_nat_lemma [THEN conjunct1]
 
 lemma arg_max_nat_le: "P x \<Longrightarrow> \<forall>y. P y \<longrightarrow> f y < b \<Longrightarrow> f x \<le> f (arg_max f P)"
   for f :: "'a \<Rightarrow> nat"
-by (blast dest: arg_max_nat_lemma [THEN conjunct2, THEN spec, of P])
+  using arg_max_nat_lemma by metis
 
 end
--- a/src/HOL/Library/Cardinality.thy	Mon Aug 26 18:26:00 2024 +0200
+++ b/src/HOL/Library/Cardinality.thy	Mon Aug 26 18:26:06 2024 +0200
@@ -32,6 +32,8 @@
 
 syntax "_type_card" :: "type => nat" ("(1CARD/(1'(_')))")
 
+syntax_consts "_type_card" == card
+
 translations "CARD('t)" => "CONST card (CONST UNIV :: 't set)"
 
 print_translation \<open>
--- a/src/HOL/Library/FSet.thy	Mon Aug 26 18:26:00 2024 +0200
+++ b/src/HOL/Library/FSet.thy	Mon Aug 26 18:26:06 2024 +0200
@@ -167,6 +167,9 @@
 syntax
   "_insert_fset"     :: "args => 'a fset"  ("{|(_)|}")
 
+syntax_consts
+  "_insert_fset" == finsert
+
 translations
   "{|x, xs|}" == "CONST finsert x {|xs|}"
   "{|x|}"     == "CONST finsert x {||}"
@@ -200,6 +203,10 @@
   "_fBall"       :: "pttrn \<Rightarrow> 'a fset \<Rightarrow> bool \<Rightarrow> bool"      ("(3\<forall>(_/|\<in>|_)./ _)" [0, 0, 10] 10)
   "_fBex"        :: "pttrn \<Rightarrow> 'a fset \<Rightarrow> bool \<Rightarrow> bool"      ("(3\<exists>(_/|\<in>|_)./ _)" [0, 0, 10] 10)
 
+syntax_consts
+  "_fBall" \<rightleftharpoons> FSet.Ball and
+  "_fBex" \<rightleftharpoons> FSet.Bex
+
 translations
   "\<forall>x|\<in>|A. P" \<rightleftharpoons> "CONST FSet.Ball A (\<lambda>x. P)"
   "\<exists>x|\<in>|A. P" \<rightleftharpoons> "CONST FSet.Bex A (\<lambda>x. P)"
--- a/src/HOL/Library/FuncSet.thy	Mon Aug 26 18:26:00 2024 +0200
+++ b/src/HOL/Library/FuncSet.thy	Mon Aug 26 18:26:06 2024 +0200
@@ -25,6 +25,9 @@
 syntax
   "_Pi" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b set \<Rightarrow> ('a \<Rightarrow> 'b) set"  ("(3\<Pi> _\<in>_./ _)"   10)
   "_lam" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> ('a \<Rightarrow> 'b)"  ("(3\<lambda>_\<in>_./ _)" [0,0,3] 3)
+syntax_consts
+  "_Pi" \<rightleftharpoons> Pi and
+  "_lam" \<rightleftharpoons> restrict
 translations
   "\<Pi> x\<in>A. B" \<rightleftharpoons> "CONST Pi A (\<lambda>x. B)"
   "\<lambda>x\<in>A. f" \<rightleftharpoons> "CONST restrict (\<lambda>x. f) A"
@@ -350,6 +353,8 @@
 
 syntax
   "_PiE" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b set \<Rightarrow> ('a \<Rightarrow> 'b) set"  ("(3\<Pi>\<^sub>E _\<in>_./ _)" 10)
+syntax_consts
+  "_PiE" \<rightleftharpoons> Pi\<^sub>E
 translations
   "\<Pi>\<^sub>E x\<in>A. B" \<rightleftharpoons> "CONST Pi\<^sub>E A (\<lambda>x. B)"
 
--- a/src/HOL/Library/Groups_Big_Fun.thy	Mon Aug 26 18:26:00 2024 +0200
+++ b/src/HOL/Library/Groups_Big_Fun.thy	Mon Aug 26 18:26:06 2024 +0200
@@ -202,6 +202,8 @@
   "_Sum_any" :: "pttrn \<Rightarrow> 'a \<Rightarrow> 'a::comm_monoid_add"    ("(3SUM _. _)" [0, 10] 10)
 syntax
   "_Sum_any" :: "pttrn \<Rightarrow> 'a \<Rightarrow> 'a::comm_monoid_add"    ("(3\<Sum>_. _)" [0, 10] 10)
+syntax_consts
+  "_Sum_any" \<rightleftharpoons> Sum_any
 translations
   "\<Sum>a. b" \<rightleftharpoons> "CONST Sum_any (\<lambda>a. b)"
 
@@ -255,6 +257,8 @@
   "_Prod_any" :: "pttrn \<Rightarrow> 'a \<Rightarrow> 'a::comm_monoid_mult"  ("(3PROD _. _)" [0, 10] 10)
 syntax
   "_Prod_any" :: "pttrn \<Rightarrow> 'a \<Rightarrow> 'a::comm_monoid_mult"  ("(3\<Prod>_. _)" [0, 10] 10)
+syntax_consts
+  "_Prod_any" == Prod_any
 translations
   "\<Prod>a. b" == "CONST Prod_any (\<lambda>a. b)"
 
--- a/src/HOL/Library/Monad_Syntax.thy	Mon Aug 26 18:26:00 2024 +0200
+++ b/src/HOL/Library/Monad_Syntax.thy	Mon Aug 26 18:26:06 2024 +0200
@@ -46,6 +46,11 @@
   "_do_bind" :: "[pttrn, 'a] \<Rightarrow> do_bind" ("(2_ <-/ _)" 13)
   "_thenM" :: "['a, 'b] \<Rightarrow> 'c" (infixl ">>" 54)
 
+syntax_consts
+  "_do_block" "_do_cons" \<rightleftharpoons> bind_do and
+  "_do_bind" "_thenM" \<rightleftharpoons> bind and
+  "_do_let" \<rightleftharpoons> Let
+
 translations
   "_do_block (_do_cons (_do_then t) (_do_final e))"
     \<rightleftharpoons> "CONST bind_do t (\<lambda>_. e)"
--- a/src/HOL/Library/Multiset.thy	Mon Aug 26 18:26:00 2024 +0200
+++ b/src/HOL/Library/Multiset.thy	Mon Aug 26 18:26:06 2024 +0200
@@ -93,6 +93,8 @@
 
 syntax
   "_multiset" :: "args \<Rightarrow> 'a multiset"    ("{#(_)#}")
+syntax_consts
+  "_multiset" == add_mset
 translations
   "{#x, xs#}" == "CONST add_mset x {#xs#}"
   "{#x#}" == "CONST add_mset x {#}"
@@ -169,6 +171,10 @@
   "_MBall"       :: "pttrn \<Rightarrow> 'a set \<Rightarrow> bool \<Rightarrow> bool"      ("(3\<forall>_:#_./ _)" [0, 0, 10] 10)
   "_MBex"        :: "pttrn \<Rightarrow> 'a set \<Rightarrow> bool \<Rightarrow> bool"      ("(3\<exists>_:#_./ _)" [0, 0, 10] 10)
 
+syntax_consts
+  "_MBall" \<rightleftharpoons> Multiset.Ball and
+  "_MBex" \<rightleftharpoons> Multiset.Bex
+
 translations
   "\<forall>x\<in>#A. P" \<rightleftharpoons> "CONST Multiset.Ball A (\<lambda>x. P)"
   "\<exists>x\<in>#A. P" \<rightleftharpoons> "CONST Multiset.Bex A (\<lambda>x. P)"
@@ -1248,6 +1254,8 @@
   "_MCollect" :: "pttrn \<Rightarrow> 'a multiset \<Rightarrow> bool \<Rightarrow> 'a multiset"    ("(1{#_ :# _./ _#})")
 syntax
   "_MCollect" :: "pttrn \<Rightarrow> 'a multiset \<Rightarrow> bool \<Rightarrow> 'a multiset"    ("(1{#_ \<in># _./ _#})")
+syntax_consts
+  "_MCollect" == filter_mset
 translations
   "{#x \<in># M. P#}" == "CONST filter_mset (\<lambda>x. P) M"
 
@@ -1821,6 +1829,8 @@
   "_comprehension_mset" :: "'a \<Rightarrow> 'b \<Rightarrow> 'b multiset \<Rightarrow> 'a multiset"  ("({#_/. _ :# _#})")
 syntax
   "_comprehension_mset" :: "'a \<Rightarrow> 'b \<Rightarrow> 'b multiset \<Rightarrow> 'a multiset"  ("({#_/. _ \<in># _#})")
+syntax_consts
+  "_comprehension_mset" \<rightleftharpoons> image_mset
 translations
   "{#e. x \<in># M#}" \<rightleftharpoons> "CONST image_mset (\<lambda>x. e) M"
 
@@ -1828,6 +1838,8 @@
   "_comprehension_mset'" :: "'a \<Rightarrow> 'b \<Rightarrow> 'b multiset \<Rightarrow> bool \<Rightarrow> 'a multiset"  ("({#_/ | _ :# _./ _#})")
 syntax
   "_comprehension_mset'" :: "'a \<Rightarrow> 'b \<Rightarrow> 'b multiset \<Rightarrow> bool \<Rightarrow> 'a multiset"  ("({#_/ | _ \<in># _./ _#})")
+syntax_consts
+  "_comprehension_mset'" \<rightleftharpoons> image_mset
 translations
   "{#e | x\<in>#M. P#}" \<rightharpoonup> "{#e. x \<in># {# x\<in>#M. P#}#}"
 
@@ -2675,6 +2687,8 @@
   "_sum_mset_image" :: "pttrn \<Rightarrow> 'b set \<Rightarrow> 'a \<Rightarrow> 'a::comm_monoid_add"  ("(3SUM _:#_. _)" [0, 51, 10] 10)
 syntax
   "_sum_mset_image" :: "pttrn \<Rightarrow> 'b set \<Rightarrow> 'a \<Rightarrow> 'a::comm_monoid_add"  ("(3\<Sum>_\<in>#_. _)" [0, 51, 10] 10)
+syntax_consts
+  "_sum_mset_image" \<rightleftharpoons> sum_mset
 translations
   "\<Sum>i \<in># A. b" \<rightleftharpoons> "CONST sum_mset (CONST image_mset (\<lambda>i. b) A)"
 
@@ -2854,6 +2868,8 @@
   "_prod_mset_image" :: "pttrn \<Rightarrow> 'b set \<Rightarrow> 'a \<Rightarrow> 'a::comm_monoid_mult"  ("(3PROD _:#_. _)" [0, 51, 10] 10)
 syntax
   "_prod_mset_image" :: "pttrn \<Rightarrow> 'b set \<Rightarrow> 'a \<Rightarrow> 'a::comm_monoid_mult"  ("(3\<Prod>_\<in>#_. _)" [0, 51, 10] 10)
+syntax_consts
+  "_prod_mset_image" \<rightleftharpoons> prod_mset
 translations
   "\<Prod>i \<in># A. b" \<rightleftharpoons> "CONST prod_mset (CONST image_mset (\<lambda>i. b) A)"
 
--- a/src/HOL/Library/Numeral_Type.thy	Mon Aug 26 18:26:00 2024 +0200
+++ b/src/HOL/Library/Numeral_Type.thy	Mon Aug 26 18:26:06 2024 +0200
@@ -540,6 +540,10 @@
   "_NumeralType0" :: type ("0")
   "_NumeralType1" :: type ("1")
 
+syntax_types
+  "_NumeralType0" == num0 and
+  "_NumeralType1" == num1
+
 translations
   (type) "1" == (type) "num1"
   (type) "0" == (type) "num0"
--- a/src/HOL/Library/Phantom_Type.thy	Mon Aug 26 18:26:00 2024 +0200
+++ b/src/HOL/Library/Phantom_Type.thy	Mon Aug 26 18:26:06 2024 +0200
@@ -18,6 +18,7 @@
 by(simp_all add: o_def id_def)
 
 syntax "_Phantom" :: "type \<Rightarrow> logic" ("(1Phantom/(1'(_')))")
+syntax_consts "_Phantom" == phantom
 translations
   "Phantom('t)" => "CONST phantom :: _ \<Rightarrow> ('t, _) phantom"
 
--- a/src/HOL/Library/Type_Length.thy	Mon Aug 26 18:26:00 2024 +0200
+++ b/src/HOL/Library/Type_Length.thy	Mon Aug 26 18:26:06 2024 +0200
@@ -19,6 +19,9 @@
 
 syntax "_type_length" :: "type \<Rightarrow> nat" (\<open>(1LENGTH/(1'(_')))\<close>)
 
+syntax_consts
+  "_type_length" \<rightleftharpoons> len_of
+
 translations "LENGTH('a)" \<rightharpoonup>
   "CONST len_of (CONST Pure.type :: 'a itself)"
 
--- a/src/HOL/List.thy	Mon Aug 26 18:26:00 2024 +0200
+++ b/src/HOL/List.thy	Mon Aug 26 18:26:06 2024 +0200
@@ -46,6 +46,9 @@
   \<comment> \<open>list Enumeration\<close>
   "_list" :: "args => 'a list"    ("[(_)]")
 
+syntax_consts
+  "_list" == Cons
+
 translations
   "[x, xs]" == "x#[xs]"
   "[x]" == "x#[]"
@@ -83,6 +86,8 @@
   "_filter" :: "[pttrn, 'a list, bool] => 'a list"  ("(1[_<-_./ _])")
 syntax
   "_filter" :: "[pttrn, 'a list, bool] => 'a list"  ("(1[_\<leftarrow>_ ./ _])")
+syntax_consts
+  "_filter" \<rightleftharpoons> filter
 translations
   "[x<-xs . P]" \<rightharpoonup> "CONST filter (\<lambda>x. P) xs"
 
@@ -132,6 +137,9 @@
   "_lupdbinds" :: "[lupdbind, lupdbinds] => lupdbinds"    ("_,/ _")
   "_LUpdate" :: "['a, lupdbinds] => 'a"    ("_/[(_)]" [1000,0] 900)
 
+syntax_consts
+  "_lupdbind" "_lupdbinds" "_LUpdate" == list_update
+
 translations
   "_LUpdate xs (_lupdbinds b bs)" == "_LUpdate (_LUpdate xs b) bs"
   "xs[i:=x]" == "CONST list_update xs i x"
@@ -3359,12 +3367,7 @@
 lemma nth_take_lemma:
   "k \<le> length xs \<Longrightarrow> k \<le> length ys \<Longrightarrow>
      (\<And>i. i < k \<longrightarrow> xs!i = ys!i) \<Longrightarrow> take k xs = take k ys"
-proof (induct k arbitrary: xs ys)
-  case (Suc k)
-  then show ?case
-    apply (simp add: less_Suc_eq_0_disj)
-    by (simp add: Suc.prems(3) take_Suc_conv_app_nth)
-qed simp
+  by (induct k arbitrary: xs ys) (simp_all add: take_Suc_conv_app_nth)
 
 lemma nth_equalityI:
   "\<lbrakk>length xs = length ys; \<And>i. i < length xs \<Longrightarrow> xs!i = ys!i\<rbrakk> \<Longrightarrow> xs = ys"
@@ -3712,7 +3715,7 @@
 
 lemma nth_eq_iff_index_eq:
   "\<lbrakk> distinct xs; i < length xs; j < length xs \<rbrakk> \<Longrightarrow> (xs!i = xs!j) = (i = j)"
-by(auto simp: distinct_conv_nth)
+  by(auto simp: distinct_conv_nth)
 
 lemma distinct_Ex1:
   "distinct xs \<Longrightarrow> x \<in> set xs \<Longrightarrow> (\<exists>!i. i < length xs \<and> xs ! i = x)"
@@ -3733,9 +3736,7 @@
 
 lemma distinct_swap[simp]: "\<lbrakk> i < size xs; j < size xs\<rbrakk> \<Longrightarrow>
   distinct(xs[i := xs!j, j := xs!i]) = distinct xs"
-  apply (simp add: distinct_conv_nth nth_list_update)
-  apply (safe; metis)
-  done
+  by (smt (verit, del_insts) distinct_conv_nth length_list_update nth_list_update)
 
 lemma set_swap[simp]:
   "\<lbrakk> i < size xs; j < size xs \<rbrakk> \<Longrightarrow> set(xs[i := xs!j, j := xs!i]) = set xs"
@@ -4320,7 +4321,7 @@
 next
   case (Cons x xs) thus ?case
     apply(auto simp: nth_Cons' split: if_splits)
-    using diff_Suc_1[unfolded One_nat_def] less_Suc_eq_0_disj by fastforce
+    using diff_Suc_1 less_Suc_eq_0_disj by fastforce
 qed
 
 lemmas find_Some_iff2 = find_Some_iff[THEN eq_iff_swap]
@@ -4514,9 +4515,16 @@
   distinct (removeAll [] xs) \<and>
   (\<forall>ys. ys \<in> set xs \<longrightarrow> distinct ys) \<and>
   (\<forall>ys zs. ys \<in> set xs \<and> zs \<in> set xs \<and> ys \<noteq> zs \<longrightarrow> set ys \<inter> set zs = {})"
-apply (induct xs)
- apply(simp_all, safe, auto)
-by (metis Int_iff UN_I empty_iff equals0I set_empty)
+proof (induct xs)
+  case Nil
+  then show ?case by auto
+next
+  case (Cons a xs)
+  have "\<lbrakk>set a \<inter> \<Union> (set ` set xs) = {}; a \<in> set xs\<rbrakk> \<Longrightarrow> a=[]"
+    by (metis Int_iff UN_I empty_iff equals0I set_empty)
+  then show ?case
+    by (auto simp: Cons)
+qed
 
 
 subsubsection \<open>\<^const>\<open>replicate\<close>\<close>
@@ -4535,24 +4543,24 @@
 qed
 
 lemma Ex_list_of_length: "\<exists>xs. length xs = n"
-by (rule exI[of _ "replicate n undefined"]) simp
+  by (rule exI[of _ "replicate n undefined"]) simp
 
 lemma map_replicate [simp]: "map f (replicate n x) = replicate n (f x)"
-by (induct n) auto
+  by (induct n) auto
 
 lemma map_replicate_const:
   "map (\<lambda> x. k) lst = replicate (length lst) k"
-by (induct lst) auto
+  by (induct lst) auto
 
 lemma replicate_app_Cons_same:
-"(replicate n x) @ (x # xs) = x # replicate n x @ xs"
-by (induct n) auto
+  "(replicate n x) @ (x # xs) = x # replicate n x @ xs"
+  by (induct n) auto
 
 lemma rev_replicate [simp]: "rev (replicate n x) = replicate n x"
-by (induct n) (auto simp: replicate_app_Cons_same)
+  by (metis length_rev map_replicate map_replicate_const rev_map)
 
 lemma replicate_add: "replicate (n + m) x = replicate n x @ replicate m x"
-by (induct n) auto
+  by (induct n) auto
 
 text\<open>Courtesy of Matthias Daum:\<close>
 lemma append_replicate_commute:
@@ -4899,10 +4907,7 @@
 
 lemma nth_rotate:
   \<open>rotate m xs ! n = xs ! ((m + n) mod length xs)\<close> if \<open>n < length xs\<close>
-  using that apply (auto simp add: rotate_drop_take nth_append not_less less_diff_conv ac_simps dest!: le_Suc_ex)
-   apply (metis add.commute mod_add_right_eq mod_less)
-  apply (metis (no_types, lifting) Nat.diff_diff_right add.commute add_diff_cancel_right' diff_le_self dual_order.strict_trans2 length_greater_0_conv less_nat_zero_code list.size(3) mod_add_right_eq mod_add_self2 mod_le_divisor mod_less)
-  done
+  by (smt (verit) add.commute hd_rotate_conv_nth length_rotate not_less0 list.size(3) mod_less rotate_rotate that)
 
 lemma nth_rotate1:
   \<open>rotate1 xs ! n = xs ! (Suc n mod length xs)\<close> if \<open>n < length xs\<close>
@@ -4944,10 +4949,8 @@
   by (auto simp add: nths_def)
 
 lemma nths_all: "\<forall>i < length xs. i \<in> I \<Longrightarrow> nths xs I = xs"
-apply (simp add: nths_def)
-apply (subst filter_True)
- apply (auto simp: in_set_zip subset_iff)
-done
+  unfolding nths_def
+  by (metis add_0 diff_zero filter_True in_set_zip length_upt nth_upt zip_eq_conv)
 
 lemma length_nths:
   "length (nths xs I) = card{i. i < length xs \<and> i \<in> I}"
@@ -5343,8 +5346,7 @@
 
     show ?thesis
       unfolding transpose.simps \<open>i = Suc j\<close> nth_Cons_Suc "3.hyps"[OF j_less]
-      apply (auto simp: transpose_aux_filter_tail filter_map comp_def length_transpose * ** *** XS_def[symmetric])
-      by (simp add: nth_tl)
+      by (auto simp: nth_tl transpose_aux_filter_tail filter_map comp_def length_transpose * ** *** XS_def[symmetric])
   qed
 qed simp_all
 
@@ -6644,14 +6646,9 @@
   have "(xs,ys) \<in> ?L (Suc n)" if r: "(xs,ys) \<in> ?R (Suc n)" for xs ys
   proof -
     from r obtain xys where r': "?len (Suc n) xs" "?len (Suc n) ys" "?P xs ys xys" by auto
-    then show ?thesis
-    proof (cases xys)
-      case Nil
-      thus ?thesis using r' by (auto simp: image_Collect lex_prod_def)
-    next
-      case Cons
-      thus ?thesis using r' Suc by (fastforce simp: image_Collect lex_prod_def)
-    qed
+    then show ?thesis 
+      using r' Suc
+      by (cases xys; fastforce simp: image_Collect lex_prod_def)
   qed
   moreover have "(xs,ys) \<in> ?L (Suc n) \<Longrightarrow> (xs,ys) \<in> ?R (Suc n)" for xs ys
     using Suc by (auto simp add: image_Collect lex_prod_def)(blast, meson Cons_eq_appendI)
@@ -6867,25 +6864,25 @@
 
 lemma lexord_same_pref_iff:
   "(xs @ ys, xs @ zs) \<in> lexord r \<longleftrightarrow> (\<exists>x \<in> set xs. (x,x) \<in> r) \<or> (ys, zs) \<in> lexord r"
-by(induction xs) auto
+  by(induction xs) auto
 
 lemma lexord_same_pref_if_irrefl[simp]:
   "irrefl r \<Longrightarrow> (xs @ ys, xs @ zs) \<in> lexord r \<longleftrightarrow> (ys, zs) \<in> lexord r"
-by (simp add: irrefl_def lexord_same_pref_iff)
+  by (simp add: irrefl_def lexord_same_pref_iff)
 
 lemma lexord_append_rightI: "\<exists> b z. y = b # z \<Longrightarrow> (x, x @ y) \<in> lexord r"
-by (metis append_Nil2 lexord_Nil_left lexord_same_pref_iff)
+  by (metis append_Nil2 lexord_Nil_left lexord_same_pref_iff)
 
 lemma lexord_append_left_rightI:
   "(a,b) \<in> r \<Longrightarrow> (u @ a # x, u @ b # y) \<in> lexord r"
-by (simp add: lexord_same_pref_iff)
+  by (simp add: lexord_same_pref_iff)
 
 lemma lexord_append_leftI: "(u,v) \<in> lexord r \<Longrightarrow> (x @ u, x @ v) \<in> lexord r"
-by (simp add: lexord_same_pref_iff)
+  by (simp add: lexord_same_pref_iff)
 
 lemma lexord_append_leftD:
   "\<lbrakk>(x @ u, x @ v) \<in> lexord r; (\<forall>a. (a,a) \<notin> r) \<rbrakk> \<Longrightarrow> (u,v) \<in> lexord r"
-by (simp add: lexord_same_pref_iff)
+  by (simp add: lexord_same_pref_iff)
 
 lemma lexord_take_index_conv:
    "((x,y) \<in> lexord r) =
@@ -6897,9 +6894,13 @@
   moreover
   have "(\<exists>u a b. (a, b) \<in> r \<and> (\<exists>v. x = u @ a # v) \<and> (\<exists>w. y = u @ b # w)) =
         (\<exists>i<length x. i < length y \<and> take i x = take i y \<and> (x ! i, y ! i) \<in> r)"
-    apply safe
-    using less_iff_Suc_add apply auto[1]
-    by (metis id_take_nth_drop)
+        (is "?L=?R")
+  proof
+    show "?L\<Longrightarrow>?R"
+      by (metis append_eq_conv_conj drop_all leI list.simps(3) nth_append_length)
+    show "?R\<Longrightarrow>?L"
+      by (metis id_take_nth_drop)
+  qed
   ultimately show ?thesis
     by (auto simp: lexord_def Let_def)
 qed
@@ -7118,19 +7119,19 @@
 by(subst lexordp_eq.simps, fastforce)+
 
 lemma lexordp_append_rightI: "ys \<noteq> Nil \<Longrightarrow> lexordp xs (xs @ ys)"
-by(induct xs)(auto simp add: neq_Nil_conv)
+  by(induct xs)(auto simp add: neq_Nil_conv)
 
 lemma lexordp_append_left_rightI: "x < y \<Longrightarrow> lexordp (us @ x # xs) (us @ y # ys)"
-by(induct us) auto
+  by(induct us) auto
 
 lemma lexordp_eq_refl: "lexordp_eq xs xs"
-by(induct xs) simp_all
+  by(induct xs) simp_all
 
 lemma lexordp_append_leftI: "lexordp us vs \<Longrightarrow> lexordp (xs @ us) (xs @ vs)"
-by(induct xs) auto
+  by(induct xs) auto
 
 lemma lexordp_append_leftD: "\<lbrakk> lexordp (xs @ us) (xs @ vs); \<forall>a. \<not> a < a \<rbrakk> \<Longrightarrow> lexordp us vs"
-by(induct xs) auto
+  by(induct xs) auto
 
 lemma lexordp_irreflexive:
   assumes irrefl: "\<forall>x. \<not> x < x"
@@ -7251,21 +7252,21 @@
 definition "measures fs = inv_image (lex less_than) (%a. map (%f. f a) fs)"
 
 lemma wf_measures[simp]: "wf (measures fs)"
-unfolding measures_def
-by blast
+  unfolding measures_def
+  by blast
 
 lemma in_measures[simp]:
   "(x, y) \<in> measures [] = False"
   "(x, y) \<in> measures (f # fs)
          = (f x < f y \<or> (f x = f y \<and> (x, y) \<in> measures fs))"
-unfolding measures_def
-by auto
+  unfolding measures_def
+  by auto
 
 lemma measures_less: "f x < f y \<Longrightarrow> (x, y) \<in> measures (f#fs)"
-by simp
+  by simp
 
 lemma measures_lesseq: "f x \<le> f y \<Longrightarrow> (x, y) \<in> measures fs \<Longrightarrow> (x, y) \<in> measures (f#fs)"
-by auto
+  by auto
 
 
 subsubsection \<open>Lifting Relations to Lists: one element\<close>
@@ -7286,27 +7287,27 @@
 unfolding listrel1_def by auto
 
 lemma not_Nil_listrel1 [iff]: "([], xs) \<notin> listrel1 r"
-unfolding listrel1_def by blast
+  unfolding listrel1_def by blast
 
 lemma not_listrel1_Nil [iff]: "(xs, []) \<notin> listrel1 r"
-unfolding listrel1_def by blast
+  unfolding listrel1_def by blast
 
 lemma Cons_listrel1_Cons [iff]:
   "(x # xs, y # ys) \<in> listrel1 r \<longleftrightarrow>
    (x,y) \<in> r \<and> xs = ys \<or> x = y \<and> (xs, ys) \<in> listrel1 r"
-by (simp add: listrel1_def Cons_eq_append_conv) (blast)
+  by (simp add: listrel1_def Cons_eq_append_conv) (blast)
 
 lemma listrel1I1: "(x,y) \<in> r \<Longrightarrow> (x # xs, y # xs) \<in> listrel1 r"
-by fast
+  by fast
 
 lemma listrel1I2: "(xs, ys) \<in> listrel1 r \<Longrightarrow> (x # xs, x # ys) \<in> listrel1 r"
-by fast
+  by fast
 
 lemma append_listrel1I:
   "(xs, ys) \<in> listrel1 r \<and> us = vs \<or> xs = ys \<and> (us, vs) \<in> listrel1 r
     \<Longrightarrow> (xs @ us, ys @ vs) \<in> listrel1 r"
-unfolding listrel1_def
-by auto (blast intro: append_eq_appendI)+
+  unfolding listrel1_def
+  by auto (blast intro: append_eq_appendI)+
 
 lemma Cons_listrel1E1[elim!]:
   assumes "(x # xs, ys) \<in> listrel1 r"
@@ -7333,19 +7334,19 @@
 qed
 
 lemma listrel1_eq_len: "(xs,ys) \<in> listrel1 r \<Longrightarrow> length xs = length ys"
-unfolding listrel1_def by auto
+  unfolding listrel1_def by auto
 
 lemma listrel1_mono:
   "r \<subseteq> s \<Longrightarrow> listrel1 r \<subseteq> listrel1 s"
-unfolding listrel1_def by blast
+  unfolding listrel1_def by blast
 
 
 lemma listrel1_converse: "listrel1 (r\<inverse>) = (listrel1 r)\<inverse>"
-unfolding listrel1_def by blast
+  unfolding listrel1_def by blast
 
 lemma in_listrel1_converse:
   "(x,y) \<in> listrel1 (r\<inverse>) \<longleftrightarrow> (x,y) \<in> (listrel1 r)\<inverse>"
-unfolding listrel1_def by blast
+  unfolding listrel1_def by blast
 
 lemma listrel1_iff_update:
   "(xs,ys) \<in> (listrel1 r)
@@ -7618,19 +7619,19 @@
 stack overflow in some target languages.\<close>
 
 fun map_tailrec_rev ::  "('a \<Rightarrow> 'b) \<Rightarrow> 'a list \<Rightarrow> 'b list \<Rightarrow> 'b list" where
-"map_tailrec_rev f [] bs = bs" |
-"map_tailrec_rev f (a#as) bs = map_tailrec_rev f as (f a # bs)"
+  "map_tailrec_rev f [] bs = bs" |
+  "map_tailrec_rev f (a#as) bs = map_tailrec_rev f as (f a # bs)"
 
 lemma map_tailrec_rev:
   "map_tailrec_rev f as bs = rev(map f as) @ bs"
-by(induction as arbitrary: bs) simp_all
+  by(induction as arbitrary: bs) simp_all
 
 definition map_tailrec :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a list \<Rightarrow> 'b list" where
-"map_tailrec f as = rev (map_tailrec_rev f as [])"
+  "map_tailrec f as = rev (map_tailrec_rev f as [])"
 
 text\<open>Code equation:\<close>
 lemma map_eq_map_tailrec: "map = map_tailrec"
-by(simp add: fun_eq_iff map_tailrec_def map_tailrec_rev)
+  by(simp add: fun_eq_iff map_tailrec_def map_tailrec_rev)
 
 
 subsubsection \<open>Counterparts for set-related operations\<close>
@@ -8375,7 +8376,7 @@
 
 lemma list_all_transfer [transfer_rule]:
   "((A ===> (=)) ===> list_all2 A ===> (=)) list_all list_all"
-  unfolding list_all_iff [abs_def] by transfer_prover
+  using list.pred_transfer by blast
 
 lemma list_ex_transfer [transfer_rule]:
   "((A ===> (=)) ===> list_all2 A ===> (=)) list_ex list_ex"
--- a/src/HOL/MacLaurin.thy	Mon Aug 26 18:26:00 2024 +0200
+++ b/src/HOL/MacLaurin.thy	Mon Aug 26 18:26:06 2024 +0200
@@ -341,9 +341,7 @@
     show "\<forall>m t. m < n \<and> 0 \<le> t \<and> t \<le> x \<longrightarrow>
                 ((\<lambda>u. sin (u + 1/2 * real m * pi)) has_real_derivative
                  sin (t + 1/2 * real (Suc m) * pi)) (at t)"
-      apply (simp add: sin_expansion_lemma del: of_nat_Suc)
-      apply (force intro!: derivative_eq_intros)
-      done
+      using DERIV_shift sin_expansion_lemma by fastforce
   qed (use assms in auto)
   then show ?thesis
     apply (rule ex_forward, simp)
@@ -362,9 +360,7 @@
     show "\<forall>m t. m < n \<and> 0 \<le> t \<and> t \<le> x \<longrightarrow>
                 ((\<lambda>u. sin (u + 1/2 * real m * pi)) has_real_derivative
                  sin (t + 1/2 * real (Suc m) * pi)) (at t)"
-      apply (simp add: sin_expansion_lemma del: of_nat_Suc)
-      apply (force intro!: derivative_eq_intros)
-      done
+      using DERIV_shift sin_expansion_lemma by fastforce
   qed (use assms in auto)
   then show ?thesis
     apply (rule ex_forward, simp)
@@ -477,8 +473,7 @@
     apply (subst t2)
     apply (rule sin_bound_lemma)
      apply (rule sum.cong[OF refl])
-    unfolding sin_coeff_def
-     apply (subst diff_m_0, simp)
+    apply (simp add: diff_m_0 sin_coeff_def)
     using est
     apply (auto intro: mult_right_mono [where b=1, simplified] mult_right_mono
         simp: ac_simps divide_inverse power_abs [symmetric] abs_mult)
--- a/src/HOL/Map.thy	Mon Aug 26 18:26:00 2024 +0200
+++ b/src/HOL/Map.thy	Mon Aug 26 18:26:06 2024 +0200
@@ -67,6 +67,9 @@
 syntax (ASCII)
   "_maplet"  :: "['a, 'a] \<Rightarrow> maplet"             ("_ /|->/ _")
 
+syntax_consts
+  "_maplet" "_Maplets" "_Map" \<rightleftharpoons> fun_upd
+
 translations
   "_Update f (_maplet x y)" \<rightleftharpoons> "f(x := CONST Some y)"
   "_Maplets m ms" \<rightharpoonup> "_updbinds m ms"
@@ -76,6 +79,7 @@
   "_Map (_maplet x y)"  \<leftharpoondown> "_Update (\<lambda>u. CONST None) (_maplet x y)"
   "_Map (_updbinds m (_maplet x y))"  \<leftharpoondown> "_Update (_Map m) (_maplet x y)"
 
+
 text \<open>Updating with lists:\<close>
 
 primrec map_of :: "('a \<times> 'b) list \<Rightarrow> 'a \<rightharpoonup> 'b" where
@@ -98,6 +102,9 @@
 syntax (ASCII)
   "_maplets" :: "['a, 'a] \<Rightarrow> maplet"             ("_ /[|->]/ _")
 
+syntax_consts
+  "_maplets" \<rightleftharpoons> map_upds
+
 translations
   "_Update m (_maplets xs ys)" \<rightleftharpoons> "CONST map_upds m xs ys"
 
--- a/src/HOL/Matrix_LP/Matrix.thy	Mon Aug 26 18:26:00 2024 +0200
+++ b/src/HOL/Matrix_LP/Matrix.thy	Mon Aug 26 18:26:06 2024 +0200
@@ -1,5 +1,5 @@
 (*  Title:      HOL/Matrix_LP/Matrix.thy
-    Author:     Steven Obua
+    Author:     Steven Obua; updated to Isar by LCP
 *)
 
 theory Matrix
@@ -69,43 +69,26 @@
 by ((rule ext)+, simp)
 
 lemma transpose_infmatrix: "transpose_infmatrix (\<lambda>j i. P j i) = (\<lambda>j i. P i j)"
-  apply (rule ext)+
-  by simp
+  by force
 
 lemma transpose_infmatrix_closed[simp]: "Rep_matrix (Abs_matrix (transpose_infmatrix (Rep_matrix x))) = transpose_infmatrix (Rep_matrix x)"
-apply (rule Abs_matrix_inverse)
-apply (simp add: matrix_def nonzero_positions_def image_def)
 proof -
   let ?A = "{pos. Rep_matrix x (snd pos) (fst pos) \<noteq> 0}"
+  let ?B = "{pos. Rep_matrix x (fst pos) (snd pos) \<noteq> 0}"
   let ?swap = "\<lambda>pos. (snd pos, fst pos)"
-  let ?B = "{pos. Rep_matrix x (fst pos) (snd pos) \<noteq> 0}"
-  have swap_image: "?swap`?A = ?B"
-    apply (simp add: image_def)
-    apply (rule set_eqI)
-    apply (simp)
-    proof
-      fix y
-      assume hyp: "\<exists>a b. Rep_matrix x b a \<noteq> 0 \<and> y = (b, a)"
-      thus "Rep_matrix x (fst y) (snd y) \<noteq> 0"
-        proof -
-          from hyp obtain a b where "(Rep_matrix x b a \<noteq> 0 & y = (b,a))" by blast
-          then show "Rep_matrix x (fst y) (snd y) \<noteq> 0" by (simp)
-        qed
-    next
-      fix y
-      assume hyp: "Rep_matrix x (fst y) (snd y) \<noteq> 0"
-      show "\<exists> a b. (Rep_matrix x b a \<noteq> 0 & y = (b,a))"
-        by (rule exI[of _ "snd y"], rule exI[of _ "fst y"]) (simp add: hyp)
-    qed
-  then have "finite (?swap`?A)"
-    proof -
-      have "finite (nonzero_positions (Rep_matrix x))" by (simp add: finite_nonzero_positions)
-      then have "finite ?B" by (simp add: nonzero_positions_def)
-      with swap_image show "finite (?swap`?A)" by (simp)
-    qed
-  moreover
-  have "inj_on ?swap ?A" by (simp add: inj_on_def)
-  ultimately show "finite ?A"by (rule finite_imageD[of ?swap ?A])
+  have "finite ?A"
+  proof -
+    have swap_image: "?swap`?A = ?B"
+      by (force simp add: image_def)
+    then have "finite (?swap`?A)"
+      by (metis (full_types) finite_nonzero_positions nonzero_positions_def)
+    moreover
+    have "inj_on ?swap ?A" by (simp add: inj_on_def)
+    ultimately show "finite ?A"
+      using finite_imageD by blast
+  qed
+  then show ?thesis
+    by (simp add: Abs_matrix_inverse matrix_def nonzero_positions_def)
 qed
 
 lemma infmatrixforward: "(x::'a infmatrix) = y \<Longrightarrow> \<forall> a b. x a b = y a b"
@@ -134,52 +117,38 @@
   by (metis nrows nrows_transpose transpose_matrix)
 
 lemma ncols_le: "(ncols A \<le> n) \<longleftrightarrow> (\<forall>j i. n \<le> i \<longrightarrow> (Rep_matrix A j i) = 0)" (is "_ = ?st")
-apply (auto)
-apply (simp add: ncols)
-proof (simp add: ncols_def, auto)
-  let ?P = "nonzero_positions (Rep_matrix A)"
-  let ?p = "snd`?P"
-  have a:"finite ?p" by (simp add: finite_nonzero_positions)
-  let ?m = "Max ?p"
-  assume "~(Suc (?m) \<le> n)"
-  then have b:"n \<le> ?m" by (simp)
-  fix a b
-  assume "(a,b) \<in> ?P"
-  then have "?p \<noteq> {}" by (auto)
-  with a have "?m \<in>  ?p" by (simp)
-  moreover have "\<forall>x. (x \<in> ?p \<longrightarrow> (\<exists>y. (Rep_matrix A y x) \<noteq> 0))" by (simp add: nonzero_positions_def image_def)
-  ultimately have "\<exists>y. (Rep_matrix A y ?m) \<noteq> 0" by (simp)
-  moreover assume ?st
-  ultimately show "False" using b by (simp)
+proof -
+  have "Rep_matrix A j i = 0"
+    if "ncols A \<le> n" "n \<le> i" for j i
+    by (meson that le_trans ncols)
+  moreover have "ncols A \<le> n"
+    if "\<forall>j i. n \<le> i \<longrightarrow> Rep_matrix A j i = 0"
+    unfolding ncols_def
+  proof (clarsimp split: if_split_asm)
+    assume \<section>: "nonzero_positions (Rep_matrix A) \<noteq> {}"
+    let ?P = "nonzero_positions (Rep_matrix A)"
+    let ?p = "snd`?P"
+    have a:"finite ?p" by (simp add: finite_nonzero_positions)
+    let ?m = "Max ?p"
+    show "Suc (Max (snd ` nonzero_positions (Rep_matrix A))) \<le> n"
+      using \<section> that  obtains_MAX [OF finite_nonzero_positions]
+      by (metis (mono_tags, lifting) mem_Collect_eq nonzero_positions_def not_less_eq_eq)
+  qed
+  ultimately show ?thesis
+    by auto
 qed
 
-lemma less_ncols: "(n < ncols A) = (\<exists>j i. n \<le> i & (Rep_matrix A j i) \<noteq> 0)"
-proof -
-  have a: "!! (a::nat) b. (a < b) = (~(b \<le> a))" by arith
-  show ?thesis by (simp add: a ncols_le)
-qed
+lemma less_ncols: "(n < ncols A) = (\<exists>j i. n \<le> i \<and> (Rep_matrix A j i) \<noteq> 0)"
+  by (meson linorder_not_le ncols_le)
 
 lemma le_ncols: "(n \<le> ncols A) = (\<forall> m. (\<forall> j i. m \<le> i \<longrightarrow> (Rep_matrix A j i) = 0) \<longrightarrow> n \<le> m)"
-apply (auto)
-apply (subgoal_tac "ncols A \<le> m")
-apply (simp)
-apply (simp add: ncols_le)
-apply (drule_tac x="ncols A" in spec)
-by (simp add: ncols)
+  by (meson le_trans ncols ncols_le)
 
 lemma nrows_le: "(nrows A \<le> n) = (\<forall>j i. n \<le> j \<longrightarrow> (Rep_matrix A j i) = 0)" (is ?s)
-proof -
-  have "(nrows A \<le> n) = (ncols (transpose_matrix A) \<le> n)" by (simp)
-  also have "\<dots> = (\<forall>j i. n \<le> i \<longrightarrow> (Rep_matrix (transpose_matrix A) j i = 0))" by (rule ncols_le)
-  also have "\<dots> = (\<forall>j i. n \<le> i \<longrightarrow> (Rep_matrix A i j) = 0)" by (simp)
-  finally show "(nrows A \<le> n) = (\<forall>j i. n \<le> j \<longrightarrow> (Rep_matrix A j i) = 0)" by (auto)
-qed
+  by (metis ncols_le ncols_transpose transpose_matrix)
 
-lemma less_nrows: "(m < nrows A) = (\<exists>j i. m \<le> j & (Rep_matrix A j i) \<noteq> 0)"
-proof -
-  have a: "!! (a::nat) b. (a < b) = (~(b \<le> a))" by arith
-  show ?thesis by (simp add: a nrows_le)
-qed
+lemma less_nrows: "(m < nrows A) = (\<exists>j i. m \<le> j \<and> (Rep_matrix A j i) \<noteq> 0)"
+  by (meson linorder_not_le nrows_le)
 
 lemma le_nrows: "(n \<le> nrows A) = (\<forall> m. (\<forall> j i. m \<le> j \<longrightarrow> (Rep_matrix A j i) = 0) \<longrightarrow> n \<le> m)"
   by (meson order.trans nrows nrows_le)
@@ -191,33 +160,31 @@
   by (meson leI ncols)
 
 lemma finite_natarray1: "finite {x. x < (n::nat)}"
-  by (induct n) auto
+  by simp
 
-lemma finite_natarray2: "finite {(x, y). x < (m::nat) & y < (n::nat)}"
+lemma finite_natarray2: "finite {(x, y). x < (m::nat) \<and> y < (n::nat)}"
   by simp
 
 lemma RepAbs_matrix:
-  assumes aem: "\<exists>m. \<forall>j i. m \<le> j \<longrightarrow> x j i = 0" (is ?em) and aen:"\<exists>n. \<forall>j i. (n \<le> i \<longrightarrow> x j i = 0)" (is ?en)
+  assumes "\<exists>m. \<forall>j i. m \<le> j \<longrightarrow> x j i = 0"  
+    and "\<exists>n. \<forall>j i. (n \<le> i \<longrightarrow> x j i = 0)" 
   shows "(Rep_matrix (Abs_matrix x)) = x"
-apply (rule Abs_matrix_inverse)
-apply (simp add: matrix_def nonzero_positions_def)
 proof -
-  from aem obtain m where a: "\<forall>j i. m \<le> j \<longrightarrow> x j i = 0" by (blast)
-  from aen obtain n where b: "\<forall>j i. n \<le> i \<longrightarrow> x j i = 0" by (blast)
-  let ?u = "{(i, j). x i j \<noteq> 0}"
-  let ?v = "{(i, j). i < m & j < n}"
-  have c: "!! (m::nat) a. ~(m \<le> a) \<Longrightarrow> a < m" by (arith)
-  from a b have "(?u \<inter> (-?v)) = {}"
-    apply (simp)
-    apply (rule set_eqI)
-    apply (simp)
-    apply auto
-    by (rule c, auto)+
-  then have d: "?u \<subseteq> ?v" by blast
-  moreover have "finite ?v" by (simp add: finite_natarray2)
-  moreover have "{pos. x (fst pos) (snd pos) \<noteq> 0} = ?u" by auto
-  ultimately show "finite {pos. x (fst pos) (snd pos) \<noteq> 0}"
-    by (metis (lifting) finite_subset)
+  have "finite {pos. x (fst pos) (snd pos) \<noteq> 0}"
+  proof -
+    from assms obtain m n where a: "\<forall>j i. m \<le> j \<longrightarrow> x j i = 0" 
+      and b: "\<forall>j i. n \<le> i \<longrightarrow> x j i = 0" by (blast)
+    let ?u = "{(i, j). x i j \<noteq> 0}"
+    let ?v = "{(i, j). i < m \<and> j < n}"
+    have c: "\<And>(m::nat) a. ~(m \<le> a) \<Longrightarrow> a < m" by (arith)
+    with a b have d: "?u \<subseteq> ?v" by blast
+    moreover have "finite ?v" by (simp add: finite_natarray2)
+    moreover have "{pos. x (fst pos) (snd pos) \<noteq> 0} = ?u" by auto
+    ultimately show "finite {pos. x (fst pos) (snd pos) \<noteq> 0}"
+      by (metis (lifting) finite_subset)
+  qed
+  then show ?thesis
+    by (simp add: Abs_matrix_inverse matrix_def nonzero_positions_def)
 qed
 
 definition apply_infmatrix :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a infmatrix \<Rightarrow> 'b infmatrix" where
@@ -277,10 +244,9 @@
 
 lemma combine_infmatrix_closed [simp]:
   "f 0 0 = 0 \<Longrightarrow> Rep_matrix (Abs_matrix (combine_infmatrix f (Rep_matrix A) (Rep_matrix B))) = combine_infmatrix f (Rep_matrix A) (Rep_matrix B)"
-apply (rule Abs_matrix_inverse)
-apply (simp add: matrix_def)
-apply (rule finite_subset[of _ "(nonzero_positions (Rep_matrix A)) \<union> (nonzero_positions (Rep_matrix B))"])
-by (simp_all)
+  apply (rule Abs_matrix_inverse)
+  apply (simp add: matrix_def)
+  by (meson finite_Un finite_nonzero_positions_Rep finite_subset nonzero_positions_combine_infmatrix)
 
 text \<open>We need the next two lemmas only later, but it is analog to the above one, so we prove them now:\<close>
 lemma nonzero_positions_apply_infmatrix[simp]: "f 0 = 0 \<Longrightarrow> nonzero_positions (apply_infmatrix f A) \<subseteq> nonzero_positions A"
@@ -290,31 +256,25 @@
   "f 0 = 0 \<Longrightarrow> Rep_matrix (Abs_matrix (apply_infmatrix f (Rep_matrix A))) = apply_infmatrix f (Rep_matrix A)"
 apply (rule Abs_matrix_inverse)
 apply (simp add: matrix_def)
-apply (rule finite_subset[of _ "nonzero_positions (Rep_matrix A)"])
-by (simp_all)
+  by (meson finite_nonzero_positions_Rep finite_subset nonzero_positions_apply_infmatrix)
 
 lemma combine_infmatrix_assoc[simp]: "f 0 0 = 0 \<Longrightarrow> associative f \<Longrightarrow> associative (combine_infmatrix f)"
   by (simp add: associative_def combine_infmatrix_def)
 
-lemma comb: "f = g \<Longrightarrow> x = y \<Longrightarrow> f x = g y"
-  by (auto)
-
 lemma combine_matrix_assoc: "f 0 0 = 0 \<Longrightarrow> associative f \<Longrightarrow> associative (combine_matrix f)"
-apply (simp(no_asm) add: associative_def combine_matrix_def, auto)
-apply (rule comb [of Abs_matrix Abs_matrix])
-by (auto, insert combine_infmatrix_assoc[of f], simp add: associative_def)
+  by (smt (verit) associative_def combine_infmatrix_assoc combine_infmatrix_closed combine_matrix_def)
 
 lemma Rep_apply_matrix[simp]: "f 0 = 0 \<Longrightarrow> Rep_matrix (apply_matrix f A) j i = f (Rep_matrix A j i)"
-by (simp add: apply_matrix_def)
+  by (simp add: apply_matrix_def)
 
 lemma Rep_combine_matrix[simp]: "f 0 0 = 0 \<Longrightarrow> Rep_matrix (combine_matrix f A B) j i = f (Rep_matrix A j i) (Rep_matrix B j i)"
   by(simp add: combine_matrix_def)
 
 lemma combine_nrows_max: "f 0 0 = 0  \<Longrightarrow> nrows (combine_matrix f A B) \<le> max (nrows A) (nrows B)"
-by (simp add: nrows_le)
+  by (simp add: nrows_le)
 
 lemma combine_ncols_max: "f 0 0 = 0 \<Longrightarrow> ncols (combine_matrix f A B) \<le> max (ncols A) (ncols B)"
-by (simp add: ncols_le)
+  by (simp add: ncols_le)
 
 lemma combine_nrows: "f 0 0 = 0 \<Longrightarrow> nrows A \<le> q \<Longrightarrow> nrows B \<le> q \<Longrightarrow> nrows(combine_matrix f A B) \<le> q"
   by (simp add: nrows_le)
@@ -329,7 +289,7 @@
   "zero_l_neutral f == \<forall>a. f 0 a = a"
 
 definition zero_closed :: "(('a::zero) \<Rightarrow> ('b::zero) \<Rightarrow> ('c::zero)) \<Rightarrow> bool" where
-  "zero_closed f == (\<forall>x. f x 0 = 0) & (\<forall>y. f 0 y = 0)"
+  "zero_closed f == (\<forall>x. f x 0 = 0) \<and> (\<forall>y. f 0 y = 0)"
 
 primrec foldseq :: "('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> (nat \<Rightarrow> 'a) \<Rightarrow> nat \<Rightarrow> 'a"
 where
@@ -341,74 +301,61 @@
   "foldseq_transposed f s 0 = s 0"
 | "foldseq_transposed f s (Suc n) = f (foldseq_transposed f s n) (s (Suc n))"
 
-lemma foldseq_assoc : "associative f \<Longrightarrow> foldseq f = foldseq_transposed f"
+lemma foldseq_assoc: 
+  assumes a:"associative f"
+  shows "associative f \<Longrightarrow> foldseq f = foldseq_transposed f"
 proof -
-  assume a:"associative f"
-  then have sublemma: "\<And>n. \<forall>N s. N \<le> n \<longrightarrow> foldseq f s N = foldseq_transposed f s N"
-  proof -
-    fix n
-    show "\<forall>N s. N \<le> n \<longrightarrow> foldseq f s N = foldseq_transposed f s N"
-    proof (induct n)
-      show "\<forall>N s. N \<le> 0 \<longrightarrow> foldseq f s N = foldseq_transposed f s N" by simp
+  have "N \<le> n \<Longrightarrow> foldseq f s N = foldseq_transposed f s N" for N s n
+  proof (induct n arbitrary: N s)
+    case 0
+    then show ?case
+      by auto
+  next
+    case (Suc n)
+    show ?case
+    proof cases
+      assume "N \<le> n"
+      then show ?thesis
+        by (simp add: Suc.hyps)
     next
-      fix n
-      assume b: "\<forall>N s. N \<le> n \<longrightarrow> foldseq f s N = foldseq_transposed f s N"
-      have c:"\<And>N s. N \<le> n \<Longrightarrow> foldseq f s N = foldseq_transposed f s N" by (simp add: b)
-      show "\<forall>N t. N \<le> Suc n \<longrightarrow> foldseq f t N = foldseq_transposed f t N"
-      proof (auto)
-        fix N t
-        assume Nsuc: "N \<le> Suc n"
-        show "foldseq f t N = foldseq_transposed f t N"
-        proof cases
-          assume "N \<le> n"
-          then show "foldseq f t N = foldseq_transposed f t N" by (simp add: b)
-        next
-          assume "~(N \<le> n)"
-          with Nsuc have Nsuceq: "N = Suc n" by simp
-          have neqz: "n \<noteq> 0 \<Longrightarrow> \<exists>m. n = Suc m & Suc m \<le> n" by arith
-          have assocf: "!! x y z. f x (f y z) = f (f x y) z" by (insert a, simp add: associative_def)
-          show "foldseq f t N = foldseq_transposed f t N"
-            apply (simp add: Nsuceq)
-            apply (subst c)
-            apply (simp)
-            apply (case_tac "n = 0")
-            apply (simp)
-            apply (drule neqz)
-            apply (erule exE)
-            apply (simp)
-            apply (subst assocf)
-            proof -
-              fix m
-              assume "n = Suc m & Suc m \<le> n"
-              then have mless: "Suc m \<le> n" by arith
-              then have step1: "foldseq_transposed f (\<lambda>k. t (Suc k)) m = foldseq f (\<lambda>k. t (Suc k)) m" (is "?T1 = ?T2")
-                apply (subst c)
-                by simp+
-              have step2: "f (t 0) ?T2 = foldseq f t (Suc m)" (is "_ = ?T3") by simp
-              have step3: "?T3 = foldseq_transposed f t (Suc m)" (is "_ = ?T4")
-                apply (subst c)
-                by (simp add: mless)+
-              have step4: "?T4 = f (foldseq_transposed f t m) (t (Suc m))" (is "_=?T5") by simp
-              from step1 step2 step3 step4 show sowhat: "f (f (t 0) ?T1) (t (Suc (Suc m))) = f ?T5 (t (Suc (Suc m)))" by simp
-            qed
-          qed
-        qed
+      assume "~(N \<le> n)"
+      then have Nsuceq: "N = Suc n"
+        using Suc.prems by linarith
+      have neqz: "n \<noteq> 0 \<Longrightarrow> \<exists>m. n = Suc m \<and> Suc m \<le> n" 
+        by arith
+      have assocf: "!! x y z. f x (f y z) = f (f x y) z"
+        by (metis a associative_def) 
+      have "f (f (s 0) (foldseq_transposed f (\<lambda>k. s (Suc k)) m)) (s (Suc (Suc m))) =
+                f (f (foldseq_transposed f s m) (s (Suc m))) (s (Suc (Suc m)))"
+        if "n = Suc m" for m
+      proof -
+        have \<section>: "foldseq_transposed f (\<lambda>k. s (Suc k)) m = foldseq f (\<lambda>k. s (Suc k)) m" (is "?T1 = ?T2")
+          by (simp add: Suc.hyps that)
+        have "f (s 0) ?T2 = foldseq f s (Suc m)" by simp
+        also have "\<dots> = foldseq_transposed f s (Suc m)" 
+          using Suc.hyps that by blast
+        also have "\<dots> = f (foldseq_transposed f s m) (s (Suc m))"  
+          by simp
+        finally show ?thesis
+          by (simp add: \<section>)
       qed
+      then show "foldseq f s N = foldseq_transposed f s N"
+        unfolding Nsuceq using assocf Suc.hyps neqz by force
     qed
-    show "foldseq f = foldseq_transposed f" by ((rule ext)+, insert sublemma, auto)
   qed
+  then show ?thesis
+    by blast
+qed
 
-lemma foldseq_distr: "\<lbrakk>associative f; commutative f\<rbrakk> \<Longrightarrow> foldseq f (\<lambda>k. f (u k) (v k)) n = f (foldseq f u n) (foldseq f v n)"
+lemma foldseq_distr: 
+  assumes assoc: "associative f" and comm: "commutative f"
+  shows "foldseq f (\<lambda>k. f (u k) (v k)) n = f (foldseq f u n) (foldseq f v n)"
 proof -
-  assume assoc: "associative f"
-  assume comm: "commutative f"
   from assoc have a:"!! x y z. f (f x y) z = f x (f y z)" by (simp add: associative_def)
   from comm have b: "!! x y. f x y = f y x" by (simp add: commutative_def)
   from assoc comm have c: "!! x y z. f x (f y z) = f y (f x z)" by (simp add: commutative_def associative_def)
-  have "\<And>n. (\<forall>u v. foldseq f (\<lambda>k. f (u k) (v k)) n = f (foldseq f u n) (foldseq f v n))"
-    apply (induct_tac n)
-    apply (simp+, auto)
-    by (simp add: a b c)
+  have "(\<forall>u v. foldseq f (\<lambda>k. f (u k) (v k)) n = f (foldseq f u n) (foldseq f v n))" for n
+    by (induct n) (simp_all add: assoc b c foldseq_assoc)
   then show "foldseq f (\<lambda>k. f (u k) (v k)) n = f (foldseq f u n) (foldseq f v n)" by simp
 qed
 
@@ -431,76 +378,60 @@
 *)
 
 lemma foldseq_zero:
-assumes fz: "f 0 0 = 0" and sz: "\<forall>i. i \<le> n \<longrightarrow> s i = 0"
-shows "foldseq f s n = 0"
+  assumes fz: "f 0 0 = 0" and sz: "\<forall>i. i \<le> n \<longrightarrow> s i = 0"
+  shows "foldseq f s n = 0"
 proof -
-  have "\<And>n. \<forall>s. (\<forall>i. i \<le> n \<longrightarrow> s i = 0) \<longrightarrow> foldseq f s n = 0"
-    apply (induct_tac n)
-    apply (simp)
-    by (simp add: fz)
-  then show "foldseq f s n = 0" by (simp add: sz)
+  have "\<forall>s. (\<forall>i. i \<le> n \<longrightarrow> s i = 0) \<longrightarrow> foldseq f s n = 0" for n
+    by (induct n) (simp_all add: fz)
+  then show ?thesis 
+    by (simp add: sz)
 qed
 
 lemma foldseq_significant_positions:
   assumes p: "\<forall>i. i \<le> N \<longrightarrow> S i = T i"
   shows "foldseq f S N = foldseq f T N"
-proof -
-  have "\<And>m. \<forall>s t. (\<forall>i. i<=m \<longrightarrow> s i = t i) \<longrightarrow> foldseq f s m = foldseq f t m"
-    apply (induct_tac m)
-    apply (simp)
-    apply (simp)
-    apply (auto)
-    proof -
-      fix n
-      fix s::"nat\<Rightarrow>'a"
-      fix t::"nat\<Rightarrow>'a"
-      assume a: "\<forall>s t. (\<forall>i\<le>n. s i = t i) \<longrightarrow> foldseq f s n = foldseq f t n"
-      assume b: "\<forall>i\<le>Suc n. s i = t i"
-      have c:"!! a b. a = b \<Longrightarrow> f (t 0) a = f (t 0) b" by blast
-      have d:"!! s t. (\<forall>i\<le>n. s i = t i) \<Longrightarrow> foldseq f s n = foldseq f t n" by (simp add: a)
-      show "f (t 0) (foldseq f (\<lambda>k. s (Suc k)) n) = f (t 0) (foldseq f (\<lambda>k. t (Suc k)) n)" by (rule c, simp add: d b)
-    qed
-  with p show ?thesis by simp
+  using assms
+proof (induction N arbitrary: S T)
+  case 0
+  then show ?case by simp
+next
+  case (Suc N)
+  then show ?case
+    unfolding foldseq.simps by (metis not_less_eq_eq le0)
 qed
 
 lemma foldseq_tail:
   assumes "M \<le> N"
   shows "foldseq f S N = foldseq f (\<lambda>k. (if k < M then (S k) else (foldseq f (\<lambda>k. S(k+M)) (N-M)))) M"
-proof -
-  have suc: "\<And>a b. \<lbrakk>a \<le> Suc b; a \<noteq> Suc b\<rbrakk> \<Longrightarrow> a \<le> b" by arith
-  have a: "\<And>a b c . a = b \<Longrightarrow> f c a = f c b" by blast
-  have "\<And>n. \<forall>m s. m \<le> n \<longrightarrow> foldseq f s n = foldseq f (\<lambda>k. (if k < m then (s k) else (foldseq f (\<lambda>k. s(k+m)) (n-m)))) m"
-    apply (induct_tac n)
-    apply (simp)
-    apply (simp)
-    apply (auto)
-    apply (case_tac "m = Suc na")
-    apply (simp)
-    apply (rule a)
-    apply (rule foldseq_significant_positions)
-    apply (auto)
-    apply (drule suc, simp+)
-    proof -
-      fix na m s
-      assume suba:"\<forall>m\<le>na. \<forall>s. foldseq f s na = foldseq f (\<lambda>k. if k < m then s k else foldseq f (\<lambda>k. s (k + m)) (na - m))m"
-      assume subb:"m \<le> na"
-      from suba have subc:"!! m s. m \<le> na \<Longrightarrow>foldseq f s na = foldseq f (\<lambda>k. if k < m then s k else foldseq f (\<lambda>k. s (k + m)) (na - m))m" by simp
-      have subd: "foldseq f (\<lambda>k. if k < m then s (Suc k) else foldseq f (\<lambda>k. s (Suc (k + m))) (na - m)) m =
-        foldseq f (\<lambda>k. s(Suc k)) na"
-        by (rule subc[of m "\<lambda>k. s(Suc k)", THEN sym], simp add: subb)
-      from subb have sube: "m \<noteq> 0 \<Longrightarrow> \<exists>mm. m = Suc mm & mm \<le> na" by arith
-      show "f (s 0) (foldseq f (\<lambda>k. if k < m then s (Suc k) else foldseq f (\<lambda>k. s (Suc (k + m))) (na - m)) m) =
-        foldseq f (\<lambda>k. if k < m then s k else foldseq f (\<lambda>k. s (k + m)) (Suc na - m)) m"
-        apply (simp add: subd)
-        apply (cases "m = 0")
-         apply simp
-        apply (drule sube)
-        apply auto
-        apply (rule a)
-        apply (simp add: subc cong del: if_weak_cong)
-        done
+  using assms
+proof (induction N arbitrary: M S)
+  case 0
+  then show ?case by auto
+next
+  case (Suc N)
+  show ?case
+  proof (cases "M = Suc N")
+    case True
+    then show ?thesis
+      by (auto intro!: arg_cong [of concl: "f (S 0)"] foldseq_significant_positions)
+  next
+    case False
+    then have "M\<le>N"
+      using Suc.prems by force
+    show ?thesis
+    proof (cases "M = 0")
+      case True
+      then show ?thesis
+        by auto
+    next
+      case False
+      then obtain M' where M': "M = Suc M'" "M' \<le> N"
+        by (metis Suc_leD \<open>M \<le> N\<close> nat.nchotomy)
+      then show ?thesis
+        apply (simp add:  Suc.IH [OF \<open>M'\<le>N\<close>])
+        using add_Suc_right diff_Suc_Suc by presburger
     qed
-  then show ?thesis using assms by simp
+  qed
 qed
 
 lemma foldseq_zerotail:
@@ -513,99 +444,73 @@
   assumes "\<forall>x. f x 0 = x"
   and "\<forall>i. n < i \<longrightarrow> s i = 0"
   and nm: "n \<le> m"
-  shows "foldseq f s n = foldseq f s m"
+shows "foldseq f s n = foldseq f s m"
 proof -
-  have "f 0 0 = 0" by (simp add: assms)
-  have b: "\<And>m n. n \<le> m \<Longrightarrow> m \<noteq> n \<Longrightarrow> \<exists>k. m-n = Suc k" by arith
-  have c: "0 \<le> m" by simp
-  have d: "\<And>k. k \<noteq> 0 \<Longrightarrow> \<exists>l. k = Suc l" by arith
-  show ?thesis
-    apply (subst foldseq_tail[OF nm])
-    apply (rule foldseq_significant_positions)
-    apply (auto)
-    apply (case_tac "m=n")
-    apply (simp+)
-    apply (drule b[OF nm])
-    apply (auto)
-    apply (case_tac "k=0")
-    apply (simp add: assms)
-    apply (drule d)
-    apply (auto)
-    apply (simp add: assms foldseq_zero)
-    done
+  have "s i = (if i < n then s i else foldseq f (\<lambda>k. s (k + n)) (m - n))"
+    if "i\<le>n" for i
+  proof (cases "m=n")
+    case True
+    then show ?thesis
+      using that by auto
+  next
+    case False
+    then obtain k where "m-n = Suc k"
+      by (metis Suc_diff_Suc le_neq_implies_less nm)
+    then show ?thesis
+      apply simp
+      by (simp add: assms(1,2) foldseq_zero nat_less_le that)
+  qed
+  then show ?thesis
+    unfolding foldseq_tail[OF nm]
+    by (auto intro: foldseq_significant_positions)
 qed
 
 lemma foldseq_zerostart:
-  "\<forall>x. f 0 (f 0 x) = f 0 x \<Longrightarrow> \<forall>i. i \<le> n \<longrightarrow> s i = 0 \<Longrightarrow> foldseq f s (Suc n) = f 0 (s (Suc n))"
-proof -
-  assume f00x: "\<forall>x. f 0 (f 0 x) = f 0 x"
-  have "\<forall>s. (\<forall>i. i<=n \<longrightarrow> s i = 0) \<longrightarrow> foldseq f s (Suc n) = f 0 (s (Suc n))"
-    apply (induct n)
-    apply (simp)
-    apply (rule allI, rule impI)
-    proof -
-      fix n
-      fix s
-      have a:"foldseq f s (Suc (Suc n)) = f (s 0) (foldseq f (\<lambda>k. s(Suc k)) (Suc n))" by simp
-      assume b: "\<forall>s. ((\<forall>i\<le>n. s i = 0) \<longrightarrow> foldseq f s (Suc n) = f 0 (s (Suc n)))"
-      from b have c:"!! s. (\<forall>i\<le>n. s i = 0) \<Longrightarrow> foldseq f s (Suc n) = f 0 (s (Suc n))" by simp
-      assume d: "\<forall>i. i \<le> Suc n \<longrightarrow> s i = 0"
-      show "foldseq f s (Suc (Suc n)) = f 0 (s (Suc (Suc n)))"
-        apply (subst a)
-        apply (subst c)
-        by (simp add: d f00x)+
-    qed
-  then show "\<forall>i. i \<le> n \<longrightarrow> s i = 0 \<Longrightarrow> foldseq f s (Suc n) = f 0 (s (Suc n))" by simp
+  assumes f00x: "\<forall>x. f 0 (f 0 x) = f 0 x" and 0: "\<forall>i. i \<le> n \<longrightarrow> s i = 0"
+  shows "foldseq f s (Suc n) = f 0 (s (Suc n))"
+  using 0
+proof (induction n arbitrary: s)
+  case 0
+  then show ?case by auto
+next
+  case (Suc n s)
+  then show ?case 
+    apply (simp add: le_Suc_eq)
+    by (smt (verit, ccfv_threshold) Suc.prems Suc_le_mono f00x foldseq_significant_positions le0)
 qed
 
 lemma foldseq_zerostart2:
-  "\<forall>x. f 0 x = x \<Longrightarrow> \<forall>i. i < n \<longrightarrow> s i = 0 \<Longrightarrow> foldseq f s n = s n"
+  assumes x: "\<forall>x. f 0 x = x" and 0: "\<forall>i. i < n \<longrightarrow> s i = 0"
+  shows  "foldseq f s n = s n"
 proof -
-  assume a: "\<forall>i. i<n \<longrightarrow> s i = 0"
-  assume x: "\<forall>x. f 0 x = x"
-  from x have f00x: "\<forall>x. f 0 (f 0 x) = f 0 x" by blast
-  have b: "\<And>i l. i < Suc l = (i \<le> l)" by arith
-  have d: "\<And>k. k \<noteq> 0 \<Longrightarrow> \<exists>l. k = Suc l" by arith
   show "foldseq f s n = s n"
-  apply (case_tac "n=0")
-  apply (simp)
-  apply (insert a)
-  apply (drule d)
-  apply (auto)
-  apply (simp add: b)
-  apply (insert f00x)
-  apply (drule foldseq_zerostart)
-  by (simp add: x)+
+  proof (cases n)
+    case 0
+    then show ?thesis
+      by auto
+  next
+    case (Suc n')
+    then show ?thesis
+      by (metis "0" foldseq_zerostart le_imp_less_Suc x)
+  qed
 qed
 
 lemma foldseq_almostzero:
   assumes f0x: "\<forall>x. f 0 x = x" and fx0: "\<forall>x. f x 0 = x" and s0: "\<forall>i. i \<noteq> j \<longrightarrow> s i = 0"
   shows "foldseq f s n = (if (j \<le> n) then (s j) else 0)"
-proof -
-  from s0 have a: "\<forall>i. i < j \<longrightarrow> s i = 0" by simp
-  from s0 have b: "\<forall>i. j < i \<longrightarrow> s i = 0" by simp
-  show ?thesis
-    apply auto
-    apply (subst foldseq_zerotail2[of f, OF fx0, of j, OF b, of n, THEN sym])
-    apply simp
-    apply (subst foldseq_zerostart2)
-    apply (simp add: f0x a)+
-    apply (subst foldseq_zero)
-    by (simp add: s0 f0x)+
-qed
+  by (smt (verit, ccfv_SIG) f0x foldseq_zerostart2 foldseq_zerotail2 fx0 le_refl nat_less_le s0)
 
 lemma foldseq_distr_unary:
-  assumes "!! a b. g (f a b) = f (g a) (g b)"
+  assumes "\<And>a b. g (f a b) = f (g a) (g b)"
   shows "g(foldseq f s n) = foldseq f (\<lambda>x. g(s x)) n"
-proof -
-  have "\<forall>s. g(foldseq f s n) = foldseq f (\<lambda>x. g(s x)) n"
-    apply (induct_tac n)
-    apply (simp)
-    apply (simp)
-    apply (auto)
-    apply (drule_tac x="\<lambda>k. s (Suc k)" in spec)
-    by (simp add: assms)
-  then show ?thesis by simp
+proof (induction n arbitrary: s)
+  case 0
+  then show ?case
+    by auto
+next
+  case (Suc n)
+  then show ?case
+    using assms by fastforce
 qed
 
 definition mult_matrix_n :: "nat \<Rightarrow> (('a::zero) \<Rightarrow> ('b::zero) \<Rightarrow> ('c::zero)) \<Rightarrow> ('c \<Rightarrow> 'c \<Rightarrow> 'c) \<Rightarrow> 'a matrix \<Rightarrow> 'b matrix \<Rightarrow> 'c matrix" where
@@ -615,14 +520,15 @@
   "mult_matrix fmul fadd A B == mult_matrix_n (max (ncols A) (nrows B)) fmul fadd A B"
 
 lemma mult_matrix_n:
-  assumes "ncols A \<le>  n" (is ?An) "nrows B \<le> n" (is ?Bn) "fadd 0 0 = 0" "fmul 0 0 = 0"
-  shows c:"mult_matrix fmul fadd A B = mult_matrix_n n fmul fadd A B"
+  assumes "ncols A \<le>  n" "nrows B \<le> n" "fadd 0 0 = 0" "fmul 0 0 = 0"
+  shows "mult_matrix fmul fadd A B = mult_matrix_n n fmul fadd A B"
 proof -
-  show ?thesis using assms
-    apply (simp add: mult_matrix_def mult_matrix_n_def)
-    apply (rule comb[of "Abs_matrix" "Abs_matrix"], simp, (rule ext)+)
-    apply (rule foldseq_zerotail, simp_all add: nrows_le ncols_le assms)
-    done
+  have "foldseq fadd (\<lambda>k. fmul (Rep_matrix A j k) (Rep_matrix B k i))
+                (max (ncols A) (nrows B)) =
+        foldseq fadd (\<lambda>k. fmul (Rep_matrix A j k) (Rep_matrix B k i)) n" for i j
+    using assms by (simp add: foldseq_zerotail nrows_le ncols_le)
+  then show ?thesis
+    by (simp add: mult_matrix_def mult_matrix_n_def)
 qed
 
 lemma mult_matrix_nm:
@@ -643,7 +549,7 @@
   "l_distributive fmul fadd == \<forall>a u v. fmul (fadd u v) a = fadd (fmul u a) (fmul v a)"
 
 definition distributive :: "('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> ('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> bool" where
-  "distributive fmul fadd == l_distributive fmul fadd & r_distributive fmul fadd"
+  "distributive fmul fadd == l_distributive fmul fadd \<and> r_distributive fmul fadd"
 
 lemma max1: "!! a x y. (a::nat) \<le> x \<Longrightarrow> a \<le> max x y" by (arith)
 lemma max2: "!! b x y. (b::nat) \<le> y \<Longrightarrow> b \<le> max x y" by (arith)
@@ -675,7 +581,7 @@
         apply (simp add: max1 max2 combine_nrows combine_ncols)+
         apply (simp add: mult_matrix_n_def r_distributive_def foldseq_distr[of fadd])
         apply (simp add: combine_matrix_def combine_infmatrix_def)
-        apply (rule comb[of "Abs_matrix" "Abs_matrix"], simp, (rule ext)+)
+        apply (intro ext arg_cong[of concl: "Abs_matrix"])
         apply (simplesubst RepAbs_matrix)
         apply (simp, auto)
         apply (rule exI[of _ "nrows a"], simp add: nrows_le foldseq_zero)
@@ -715,7 +621,7 @@
         apply (simp add: max1 max2 combine_nrows combine_ncols)+
         apply (simp add: mult_matrix_n_def l_distributive_def foldseq_distr[of fadd])
         apply (simp add: combine_matrix_def combine_infmatrix_def)
-        apply (rule comb[of "Abs_matrix" "Abs_matrix"], simp, (rule ext)+)
+        apply (intro ext arg_cong[of concl: "Abs_matrix"])
         apply (simplesubst RepAbs_matrix)
         apply (simp, auto)
         apply (rule exI[of _ "nrows v"], simp add: nrows_le foldseq_zero)
@@ -738,21 +644,13 @@
 end
 
 lemma Rep_zero_matrix_def[simp]: "Rep_matrix 0 j i = 0"
-  apply (simp add: zero_matrix_def)
-  apply (subst RepAbs_matrix)
-  by (auto)
+  by (simp add: RepAbs_matrix zero_matrix_def)
 
 lemma zero_matrix_def_nrows[simp]: "nrows 0 = 0"
-proof -
-  have a:"!! (x::nat). x \<le> 0 \<Longrightarrow> x = 0" by (arith)
-  show "nrows 0 = 0" by (rule a, subst nrows_le, simp)
-qed
+  using nrows_le by force
 
 lemma zero_matrix_def_ncols[simp]: "ncols 0 = 0"
-proof -
-  have a:"!! (x::nat). x \<le> 0 \<Longrightarrow> x = 0" by (arith)
-  show "ncols 0 = 0" by (rule a, subst ncols_le, simp)
-qed
+  using ncols_le by fastforce
 
 lemma combine_matrix_zero_l_neutral: "zero_l_neutral f \<Longrightarrow> zero_l_neutral (combine_matrix f)"
   by (simp add: zero_l_neutral_def combine_matrix_def combine_infmatrix_def)
@@ -762,42 +660,34 @@
 
 lemma mult_matrix_zero_closed: "\<lbrakk>fadd 0 0 = 0; zero_closed fmul\<rbrakk> \<Longrightarrow> zero_closed (mult_matrix fmul fadd)"
   apply (simp add: zero_closed_def mult_matrix_def mult_matrix_n_def)
-  apply (auto)
-  by (subst foldseq_zero, (simp add: zero_matrix_def)+)+
+  by (simp add: foldseq_zero zero_matrix_def)
 
 lemma mult_matrix_n_zero_right[simp]: "\<lbrakk>fadd 0 0 = 0; \<forall>a. fmul a 0 = 0\<rbrakk> \<Longrightarrow> mult_matrix_n n fmul fadd A 0 = 0"
-  apply (simp add: mult_matrix_n_def)
-  apply (subst foldseq_zero)
-  by (simp_all add: zero_matrix_def)
+  by (simp add: RepAbs_matrix foldseq_zero matrix_eqI mult_matrix_n_def)
 
 lemma mult_matrix_n_zero_left[simp]: "\<lbrakk>fadd 0 0 = 0; \<forall>a. fmul 0 a = 0\<rbrakk> \<Longrightarrow> mult_matrix_n n fmul fadd 0 A = 0"
-  apply (simp add: mult_matrix_n_def)
-  apply (subst foldseq_zero)
-  by (simp_all add: zero_matrix_def)
+  by (simp add: RepAbs_matrix foldseq_zero matrix_eqI mult_matrix_n_def)
 
 lemma mult_matrix_zero_left[simp]: "\<lbrakk>fadd 0 0 = 0; \<forall>a. fmul 0 a = 0\<rbrakk> \<Longrightarrow> mult_matrix fmul fadd 0 A = 0"
-by (simp add: mult_matrix_def)
+  by (simp add: mult_matrix_def)
 
 lemma mult_matrix_zero_right[simp]: "\<lbrakk>fadd 0 0 = 0; \<forall>a. fmul a 0 = 0\<rbrakk> \<Longrightarrow> mult_matrix fmul fadd A 0 = 0"
-by (simp add: mult_matrix_def)
+  by (simp add: mult_matrix_def)
 
 lemma apply_matrix_zero[simp]: "f 0 = 0 \<Longrightarrow> apply_matrix f 0 = 0"
-  apply (simp add: apply_matrix_def apply_infmatrix_def)
-  by (simp add: zero_matrix_def)
+  by (simp add: matrix_eqI)
 
 lemma combine_matrix_zero: "f 0 0 = 0 \<Longrightarrow> combine_matrix f 0 0 = 0"
-  apply (simp add: combine_matrix_def combine_infmatrix_def)
-  by (simp add: zero_matrix_def)
+  by (simp add: matrix_eqI)
 
 lemma transpose_matrix_zero[simp]: "transpose_matrix 0 = 0"
-  by (simp add: transpose_matrix_def zero_matrix_def RepAbs_matrix transpose_infmatrix)
+  by (simp add: matrix_eqI)
 
 lemma apply_zero_matrix_def[simp]: "apply_matrix (\<lambda>x. 0) A = 0"
-  apply (simp add: apply_matrix_def apply_infmatrix_def)
-  by (simp add: zero_matrix_def)
+  by (simp add: matrix_eqI)
 
 definition singleton_matrix :: "nat \<Rightarrow> nat \<Rightarrow> ('a::zero) \<Rightarrow> 'a matrix" where
-  "singleton_matrix j i a == Abs_matrix(\<lambda>m n. if j = m & i = n then a else 0)"
+  "singleton_matrix j i a == Abs_matrix(\<lambda>m n. if j = m \<and> i = n then a else 0)"
 
 definition move_matrix :: "('a::zero) matrix \<Rightarrow> int \<Rightarrow> int \<Rightarrow> 'a matrix" where
   "move_matrix A y x == Abs_matrix(\<lambda>j i. if (((int j)-y) < 0) | (((int i)-x) < 0) then 0 else Rep_matrix A (nat ((int j)-y)) (nat ((int i)-x)))"
@@ -814,13 +704,9 @@
 definition row_of_matrix :: "('a::zero) matrix \<Rightarrow> nat \<Rightarrow> 'a matrix" where
   "row_of_matrix A m == take_rows (move_matrix A (- int m) 0) 1"
 
-lemma Rep_singleton_matrix[simp]: "Rep_matrix (singleton_matrix j i e) m n = (if j = m & i = n then e else 0)"
-apply (simp add: singleton_matrix_def)
-apply (auto)
-apply (subst RepAbs_matrix)
-apply (rule exI[of _ "Suc m"], simp)
-apply (rule exI[of _ "Suc n"], simp+)
-by (subst RepAbs_matrix, rule exI[of _ "Suc j"], simp, rule exI[of _ "Suc i"], simp+)+
+lemma Rep_singleton_matrix[simp]: "Rep_matrix (singleton_matrix j i e) m n = (if j = m \<and> i = n then e else 0)"
+  unfolding singleton_matrix_def
+  by (smt (verit, del_insts) RepAbs_matrix Suc_n_not_le_n)
 
 lemma apply_singleton_matrix[simp]: "f 0 = 0 \<Longrightarrow> apply_matrix f (singleton_matrix j i x) = (singleton_matrix j i (f x))"
   by (simp add: matrix_eqI)
@@ -829,47 +715,20 @@
   by (simp add: singleton_matrix_def zero_matrix_def)
 
 lemma nrows_singleton[simp]: "nrows(singleton_matrix j i e) = (if e = 0 then 0 else Suc j)"
-proof-
-  have th: "\<not> (\<forall>m. m \<le> j)" "\<exists>n. \<not> n \<le> i" by arith+
-from th show ?thesis 
-apply (auto)
-apply (rule le_antisym)
-apply (subst nrows_le)
-apply (simp add: singleton_matrix_def, auto)
-apply (subst RepAbs_matrix)
-apply auto
-apply (simp add: Suc_le_eq)
-apply (rule not_le_imp_less)
-apply (subst nrows_le)
-by simp
+proof -
+  have "e \<noteq> 0 \<Longrightarrow> Suc j \<le> nrows (singleton_matrix j i e)"
+    by (metis Rep_singleton_matrix not_less_eq_eq nrows)
+  then show ?thesis
+    by (simp add: le_antisym nrows_le)
 qed
 
 lemma ncols_singleton[simp]: "ncols(singleton_matrix j i e) = (if e = 0 then 0 else Suc i)"
-proof-
-have th: "\<not> (\<forall>m. m \<le> j)" "\<exists>n. \<not> n \<le> i" by arith+
-from th show ?thesis 
-apply (auto)
-apply (rule le_antisym)
-apply (subst ncols_le)
-apply (simp add: singleton_matrix_def, auto)
-apply (subst RepAbs_matrix)
-apply auto
-apply (simp add: Suc_le_eq)
-apply (rule not_le_imp_less)
-apply (subst ncols_le)
-by simp
-qed
+  by (simp add: Suc_leI le_antisym ncols_le ncols_notzero)
 
 lemma combine_singleton: "f 0 0 = 0 \<Longrightarrow> combine_matrix f (singleton_matrix j i a) (singleton_matrix j i b) = singleton_matrix j i (f a b)"
-apply (simp add: singleton_matrix_def combine_matrix_def combine_infmatrix_def)
-apply (subst RepAbs_matrix)
-apply (rule exI[of _ "Suc j"], simp)
-apply (rule exI[of _ "Suc i"], simp)
-apply (rule comb[of "Abs_matrix" "Abs_matrix"], simp, (rule ext)+)
-apply (subst RepAbs_matrix)
-apply (rule exI[of _ "Suc j"], simp)
-apply (rule exI[of _ "Suc i"], simp)
-by simp
+  apply (simp add: singleton_matrix_def combine_matrix_def combine_infmatrix_def)
+  apply (intro ext arg_cong[of concl: "Abs_matrix"])
+  by (metis Rep_singleton_matrix singleton_matrix_def)
 
 lemma transpose_singleton[simp]: "transpose_matrix (singleton_matrix j i a) = singleton_matrix i j a"
   by (simp add: matrix_eqI)
@@ -877,14 +736,13 @@
 lemma Rep_move_matrix[simp]:
   "Rep_matrix (move_matrix A y x) j i =
   (if (((int j)-y) < 0) | (((int i)-x) < 0) then 0 else Rep_matrix A (nat((int j)-y)) (nat((int i)-x)))"
-apply (simp add: move_matrix_def)
-apply (auto)
+  apply (simp add: move_matrix_def)
 by (subst RepAbs_matrix,
   rule exI[of _ "(nrows A)+(nat \<bar>y\<bar>)"], auto, rule nrows, arith,
   rule exI[of _ "(ncols A)+(nat \<bar>x\<bar>)"], auto, rule ncols, arith)+
 
 lemma move_matrix_0_0[simp]: "move_matrix A 0 0 = A"
-by (simp add: move_matrix_def)
+  by (simp add: move_matrix_def)
 
 lemma move_matrix_ortho: "move_matrix A j i = move_matrix (move_matrix A j 0) 0 i"
   by (simp add: matrix_eqI)
@@ -895,28 +753,17 @@
 
 lemma move_matrix_singleton[simp]: "move_matrix (singleton_matrix u v x) j i =
   (if (j + int u < 0) | (i + int v < 0) then 0 else (singleton_matrix (nat (j + int u)) (nat (i + int v)) x))"
-  apply (intro matrix_eqI)
-  apply (split if_split)
-  apply (auto simp: split: if_split_asm)
-  done
+  by (auto intro!: matrix_eqI split: if_split_asm)
 
 lemma Rep_take_columns[simp]:
-  "Rep_matrix (take_columns A c) j i =
-  (if i < c then (Rep_matrix A j i) else 0)"
-apply (simp add: take_columns_def)
-apply (simplesubst RepAbs_matrix)
-apply (rule exI[of _ "nrows A"], auto, simp add: nrows_le)
-apply (rule exI[of _ "ncols A"], auto, simp add: ncols_le)
-done
+  "Rep_matrix (take_columns A c) j i = (if i < c then (Rep_matrix A j i) else 0)"
+  unfolding take_columns_def
+  by (smt (verit, best) RepAbs_matrix leD nrows)
 
 lemma Rep_take_rows[simp]:
-  "Rep_matrix (take_rows A r) j i =
-  (if j < r then (Rep_matrix A j i) else 0)"
-apply (simp add: take_rows_def)
-apply (simplesubst RepAbs_matrix)
-apply (rule exI[of _ "nrows A"], auto, simp add: nrows_le)
-apply (rule exI[of _ "ncols A"], auto, simp add: ncols_le)
-done
+  "Rep_matrix (take_rows A r) j i = (if j < r then (Rep_matrix A j i) else 0)"
+  unfolding take_rows_def
+  by (smt (verit, best) RepAbs_matrix leD ncols)
 
 lemma Rep_column_of_matrix[simp]:
   "Rep_matrix (column_of_matrix A c) j i = (if i = 0 then (Rep_matrix A j c) else 0)"
@@ -933,22 +780,14 @@
   by (simp add: matrix_eqI nrows)
 
 lemma mult_matrix_singleton_right[simp]:
-  assumes
-  "\<forall>x. fmul x 0 = 0"
-  "\<forall>x. fmul 0 x = 0"
-  "\<forall>x. fadd 0 x = x"
-  "\<forall>x. fadd x 0 = x"
+  assumes "\<forall>x. fmul x 0 = 0" "\<forall>x. fmul 0 x = 0" "\<forall>x. fadd 0 x = x" "\<forall>x. fadd x 0 = x"
   shows "(mult_matrix fmul fadd A (singleton_matrix j i e)) = apply_matrix (\<lambda>x. fmul x e) (move_matrix (column_of_matrix A j) 0 (int i))"
-  apply (simp add: mult_matrix_def)
-  apply (subst mult_matrix_nm[of _ _ _ "max (ncols A) (Suc j)"])
-  apply (auto)
-  apply (simp add: assms)+
-  apply (simp add: mult_matrix_n_def apply_matrix_def apply_infmatrix_def)
-  apply (rule comb[of "Abs_matrix" "Abs_matrix"], auto, (rule ext)+)
-  apply (subst foldseq_almostzero[of _ j])
-  apply (simp add: assms)+
-  apply (auto)
-  done
+  using assms
+  unfolding mult_matrix_def
+  apply (subst mult_matrix_nm[of _ _ _ "max (ncols A) (Suc j)"];
+         simp add: mult_matrix_n_def apply_matrix_def apply_infmatrix_def)
+  apply (intro ext arg_cong[of concl: "Abs_matrix"])
+  by (simp add: max_def assms foldseq_almostzero[of _ j])
 
 lemma mult_matrix_ext:
   assumes
@@ -961,12 +800,10 @@
   "\<forall>a. fadd 0 a = a"
   and contraprems: "mult_matrix fmul fadd A = mult_matrix fmul fadd B"
   shows "A = B"
-proof(rule contrapos_np[of "False"], simp)
-  assume a: "A \<noteq> B"
-  have b: "\<And>f g. (\<forall>x y. f x y = g x y) \<Longrightarrow> f = g" by ((rule ext)+, auto)
-  have "\<exists>j i. (Rep_matrix A j i) \<noteq> (Rep_matrix B j i)"
-    using Rep_matrix_inject a by blast
-  then obtain J I where c:"(Rep_matrix A J I) \<noteq> (Rep_matrix B J I)" by blast
+proof(rule ccontr)
+  assume "A \<noteq> B"
+  then obtain J I where ne: "(Rep_matrix A J I) \<noteq> (Rep_matrix B J I)"
+    by (meson matrix_eqI)
   from eprem obtain e where eprops:"(\<forall>a b. a \<noteq> b \<longrightarrow> fmul a e \<noteq> fmul b e)" by blast
   let ?S = "singleton_matrix I 0 e"
   let ?comp = "mult_matrix fmul fadd"
@@ -975,8 +812,8 @@
   have "Rep_matrix (apply_matrix (\<lambda>x. fmul x e) (column_of_matrix A I)) \<noteq>
         Rep_matrix (apply_matrix (\<lambda>x. fmul x e) (column_of_matrix B I))"
     using fprems
-    by (metis Rep_apply_matrix Rep_column_of_matrix eprops c)
-  then have "~(?comp A ?S = ?comp B ?S)"
+    by (metis Rep_apply_matrix Rep_column_of_matrix eprops ne)
+  then have "?comp A ?S \<noteq> ?comp B ?S"
     by (simp add: fprems eprops Rep_matrix_inject)
   with contraprems show "False" by simp
 qed
@@ -988,95 +825,90 @@
   "foldmatrix_transposed f g A m n == foldseq g (\<lambda>j. foldseq_transposed f (A j) n) m"
 
 lemma foldmatrix_transpose:
-  assumes
-  "\<forall>a b c d. g(f a b) (f c d) = f (g a c) (g b d)"
-  shows
-  "foldmatrix f g A m n = foldmatrix_transposed g f (transpose_infmatrix A) n m"
+  assumes "\<forall>a b c d. g(f a b) (f c d) = f (g a c) (g b d)"
+  shows "foldmatrix f g A m n = foldmatrix_transposed g f (transpose_infmatrix A) n m"
 proof -
   have forall:"\<And>P x. (\<forall>x. P x) \<Longrightarrow> P x" by auto
   have tworows:"\<forall>A. foldmatrix f g A 1 n = foldmatrix_transposed g f (transpose_infmatrix A) n 1"
-    apply (induct n)
-    apply (simp add: foldmatrix_def foldmatrix_transposed_def assms)+
-    apply (auto)
-    by (drule_tac x="(\<lambda>j i. A j (Suc i))" in forall, simp)
-  show "foldmatrix f g A m n = foldmatrix_transposed g f (transpose_infmatrix A) n m"
-    apply (simp add: foldmatrix_def foldmatrix_transposed_def)
-    apply (induct m, simp)
-    apply (simp)
-    apply (insert tworows)
-    apply (drule_tac x="\<lambda>j i. (if j = 0 then (foldseq_transposed g (\<lambda>u. A u i) m) else (A (Suc m) i))" in spec)
+  proof (induct n)
+    case 0
+    then show ?case 
+      by (simp add: foldmatrix_def foldmatrix_transposed_def)
+  next
+    case (Suc n)
+    then show ?case
+      apply (clarsimp simp: foldmatrix_def foldmatrix_transposed_def assms)
+      apply (rule arg_cong [of concl: "f _"])
+      by meson 
+  qed
+  have "foldseq_transposed g (\<lambda>j. foldseq f (A j) n) m =
+    foldseq f (\<lambda>j. foldseq_transposed g (transpose_infmatrix A j) m) n"
+  proof (induct m)
+    case 0
+    then show ?case by auto
+  next
+    case (Suc m)
+    then show ?case
+      using tworows
+      apply (drule_tac x="\<lambda>j i. (if j = 0 then (foldseq_transposed g (\<lambda>u. A u i) m) else (A (Suc m) i))" in spec)
+      by (simp add: Suc foldmatrix_def foldmatrix_transposed_def)
+  qed
+  then show "foldmatrix f g A m n = foldmatrix_transposed g f (transpose_infmatrix A) n m"
     by (simp add: foldmatrix_def foldmatrix_transposed_def)
 qed
 
 lemma foldseq_foldseq:
-assumes
-  "associative f"
-  "associative g"
-  "\<forall>a b c d. g(f a b) (f c d) = f (g a c) (g b d)"
+assumes "associative f" "associative g" "\<forall>a b c d. g(f a b) (f c d) = f (g a c) (g b d)"
 shows
   "foldseq g (\<lambda>j. foldseq f (A j) n) m = foldseq f (\<lambda>j. foldseq g ((transpose_infmatrix A) j) m) n"
-  apply (insert foldmatrix_transpose[of g f A m n])
+  using foldmatrix_transpose[of g f A m n]
   by (simp add: foldmatrix_def foldmatrix_transposed_def foldseq_assoc[THEN sym] assms)
 
 lemma mult_n_nrows:
-assumes
-"\<forall>a. fmul 0 a = 0"
-"\<forall>a. fmul a 0 = 0"
-"fadd 0 0 = 0"
-shows "nrows (mult_matrix_n n fmul fadd A B) \<le> nrows A"
-apply (subst nrows_le)
-apply (simp add: mult_matrix_n_def)
-apply (subst RepAbs_matrix)
-apply (rule_tac x="nrows A" in exI)
-apply (simp add: nrows assms foldseq_zero)
-apply (rule_tac x="ncols B" in exI)
-apply (simp add: ncols assms foldseq_zero)
-apply (simp add: nrows assms foldseq_zero)
-done
+  assumes "\<forall>a. fmul 0 a = 0" "\<forall>a. fmul a 0 = 0" "fadd 0 0 = 0"
+  shows "nrows (mult_matrix_n n fmul fadd A B) \<le> nrows A"
+  unfolding nrows_le mult_matrix_n_def
+  apply (subst RepAbs_matrix)
+    apply (rule_tac x="nrows A" in exI)
+    apply (simp add: nrows assms foldseq_zero)
+   apply (rule_tac x="ncols B" in exI)
+   apply (simp add: ncols assms foldseq_zero)
+  apply (simp add: nrows assms foldseq_zero)
+  done
 
 lemma mult_n_ncols:
-assumes
-"\<forall>a. fmul 0 a = 0"
-"\<forall>a. fmul a 0 = 0"
-"fadd 0 0 = 0"
-shows "ncols (mult_matrix_n n fmul fadd A B) \<le> ncols B"
-apply (subst ncols_le)
-apply (simp add: mult_matrix_n_def)
-apply (subst RepAbs_matrix)
-apply (rule_tac x="nrows A" in exI)
-apply (simp add: nrows assms foldseq_zero)
-apply (rule_tac x="ncols B" in exI)
-apply (simp add: ncols assms foldseq_zero)
-apply (simp add: ncols assms foldseq_zero)
-done
+  assumes "\<forall>a. fmul 0 a = 0" "\<forall>a. fmul a 0 = 0" "fadd 0 0 = 0"
+  shows "ncols (mult_matrix_n n fmul fadd A B) \<le> ncols B"
+  unfolding ncols_le mult_matrix_n_def
+  apply (subst RepAbs_matrix)
+    apply (rule_tac x="nrows A" in exI)
+    apply (simp add: nrows assms foldseq_zero)
+   apply (rule_tac x="ncols B" in exI)
+   apply (simp add: ncols assms foldseq_zero)
+  apply (simp add: ncols assms foldseq_zero)
+  done
 
 lemma mult_nrows:
-assumes
-"\<forall>a. fmul 0 a = 0"
-"\<forall>a. fmul a 0 = 0"
-"fadd 0 0 = 0"
-shows "nrows (mult_matrix fmul fadd A B) \<le> nrows A"
-by (simp add: mult_matrix_def mult_n_nrows assms)
+  assumes
+    "\<forall>a. fmul 0 a = 0"
+    "\<forall>a. fmul a 0 = 0"
+    "fadd 0 0 = 0"
+  shows "nrows (mult_matrix fmul fadd A B) \<le> nrows A"
+  by (simp add: mult_matrix_def mult_n_nrows assms)
 
 lemma mult_ncols:
-assumes
-"\<forall>a. fmul 0 a = 0"
-"\<forall>a. fmul a 0 = 0"
-"fadd 0 0 = 0"
-shows "ncols (mult_matrix fmul fadd A B) \<le> ncols B"
-by (simp add: mult_matrix_def mult_n_ncols assms)
+  assumes
+    "\<forall>a. fmul 0 a = 0"
+    "\<forall>a. fmul a 0 = 0"
+    "fadd 0 0 = 0"
+  shows "ncols (mult_matrix fmul fadd A B) \<le> ncols B"
+  by (simp add: mult_matrix_def mult_n_ncols assms)
 
 lemma nrows_move_matrix_le: "nrows (move_matrix A j i) \<le> nat((int (nrows A)) + j)"
-  apply (auto simp: nrows_le)
-  apply (rule nrows)
-  apply (arith)
-  done
+  by (smt (verit) Rep_move_matrix int_nat_eq nrows nrows_le of_nat_le_iff)
 
 lemma ncols_move_matrix_le: "ncols (move_matrix A j i) \<le> nat((int (ncols A)) + i)"
-  apply (auto simp: ncols_le)
-  apply (rule ncols)
-  apply (arith)
-  done
+  by (metis nrows_move_matrix_le nrows_transpose transpose_move_matrix)
 
 lemma mult_matrix_assoc:
   assumes
@@ -1104,78 +936,47 @@
     apply (intro matrix_eqI)
     apply (simp add: mult_matrix_def)
     apply (simplesubst mult_matrix_nm[of _ "max (ncols (mult_matrix_n (max (ncols A) (nrows B)) fmul1 fadd1 A B)) (nrows C)" _ "max (ncols B) (nrows C)"])
-    apply (simp add: max1 max2 mult_n_ncols mult_n_nrows assms)+
+          apply (simp add: max1 max2 mult_n_ncols mult_n_nrows assms)+
     apply (simplesubst mult_matrix_nm[of _ "max (ncols A) (nrows (mult_matrix_n (max (ncols B) (nrows C)) fmul2 fadd2 B C))" _ "max (ncols A) (nrows B)"])
-    apply (simp add: max1 max2 mult_n_ncols mult_n_nrows assms)+
+          apply (simp add: max1 max2 mult_n_ncols mult_n_nrows assms)+
     apply (simplesubst mult_matrix_nm[of _ _ _ "?N"])
-    apply (simp add: max1 max2 mult_n_ncols mult_n_nrows assms)+
+          apply (simp add: max1 max2 mult_n_ncols mult_n_nrows assms)+
     apply (simplesubst mult_matrix_nm[of _ _ _ "?N"])
-    apply (simp add: max1 max2 mult_n_ncols mult_n_nrows assms)+
+          apply (simp add: max1 max2 mult_n_ncols mult_n_nrows assms)+
     apply (simplesubst mult_matrix_nm[of _ _ _ "?N"])
-    apply (simp add: max1 max2 mult_n_ncols mult_n_nrows assms)+
+          apply (simp add: max1 max2 mult_n_ncols mult_n_nrows assms)+
     apply (simplesubst mult_matrix_nm[of _ _ _ "?N"])
-    apply (simp add: max1 max2 mult_n_ncols mult_n_nrows assms)+
+          apply (simp add: max1 max2 mult_n_ncols mult_n_nrows assms)+
     apply (simp add: mult_matrix_n_def)
     apply (rule comb_left)
     apply ((rule ext)+, simp)
     apply (simplesubst RepAbs_matrix)
-    apply (rule exI[of _ "nrows B"])
-    apply (simp add: nrows assms foldseq_zero)
-    apply (rule exI[of _ "ncols C"])
-    apply (simp add: assms ncols foldseq_zero)
+      apply (rule exI[of _ "nrows B"])
+      apply (simp add: nrows assms foldseq_zero)
+     apply (rule exI[of _ "ncols C"])
+     apply (simp add: assms ncols foldseq_zero)
     apply (subst RepAbs_matrix)
-    apply (rule exI[of _ "nrows A"])
-    apply (simp add: nrows assms foldseq_zero)
-    apply (rule exI[of _ "ncols B"])
-    apply (simp add: assms ncols foldseq_zero)
+      apply (rule exI[of _ "nrows A"])
+      apply (simp add: nrows assms foldseq_zero)
+     apply (rule exI[of _ "ncols B"])
+     apply (simp add: assms ncols foldseq_zero)
     apply (simp add: fmul2fadd1fold fmul1fadd2fold assms)
     apply (subst foldseq_foldseq)
-    apply (simp add: assms)+
+       apply (simp add: assms)+
     apply (simp add: transpose_infmatrix)
     done
 qed
 
-lemma
-  assumes
-  "\<forall>a. fmul1 0 a = 0"
-  "\<forall>a. fmul1 a 0 = 0"
-  "\<forall>a. fmul2 0 a = 0"
-  "\<forall>a. fmul2 a 0 = 0"
-  "fadd1 0 0 = 0"
-  "fadd2 0 0 = 0"
-  "\<forall>a b c d. fadd2 (fadd1 a b) (fadd1 c d) = fadd1 (fadd2 a c) (fadd2 b d)"
-  "associative fadd1"
-  "associative fadd2"
-  "\<forall>a b c. fmul2 (fmul1 a b) c = fmul1 a (fmul2 b c)"
-  "\<forall>a b c. fmul2 (fadd1 a b) c = fadd1 (fmul2 a c) (fmul2 b c)"
-  "\<forall>a b c. fmul1 c (fadd2 a b) = fadd2 (fmul1 c a) (fmul1 c b)"
-  shows
-  "(mult_matrix fmul1 fadd1 A) o (mult_matrix fmul2 fadd2 B) = mult_matrix fmul2 fadd2 (mult_matrix fmul1 fadd1 A B)"
-apply (rule ext)+
-apply (simp add: comp_def )
-apply (simp add: mult_matrix_assoc assms)
-done
-
 lemma mult_matrix_assoc_simple:
   assumes
   "\<forall>a. fmul 0 a = 0"
   "\<forall>a. fmul a 0 = 0"
-  "fadd 0 0 = 0"
   "associative fadd"
   "commutative fadd"
   "associative fmul"
   "distributive fmul fadd"
   shows "mult_matrix fmul fadd (mult_matrix fmul fadd A B) C = mult_matrix fmul fadd A (mult_matrix fmul fadd B C)"
-proof -
-  have "!! a b c d. fadd (fadd a b) (fadd c d) = fadd (fadd a c) (fadd b d)"
-    using assms by (simp add: associative_def commutative_def)
-  then show ?thesis
-    apply (subst mult_matrix_assoc)
-    using assms
-    apply simp_all
-    apply (simp_all add: associative_def distributive_def l_distributive_def r_distributive_def)
-    done
-qed
+  by (smt (verit) assms associative_def commutative_def distributive_def l_distributive_def mult_matrix_assoc r_distributive_def)
 
 lemma transpose_apply_matrix: "f 0 = 0 \<Longrightarrow> transpose_matrix (apply_matrix f A) = apply_matrix f (transpose_matrix A)"
   by (simp add: matrix_eqI)
@@ -1184,19 +985,17 @@
   by (simp add: matrix_eqI)
 
 lemma Rep_mult_matrix:
-  assumes
-  "\<forall>a. fmul 0 a = 0"
-  "\<forall>a. fmul a 0 = 0"
-  "fadd 0 0 = 0"
+  assumes "\<forall>a. fmul 0 a = 0" "\<forall>a. fmul a 0 = 0" "fadd 0 0 = 0"
   shows
-  "Rep_matrix(mult_matrix fmul fadd A B) j i =
+    "Rep_matrix(mult_matrix fmul fadd A B) j i =
   foldseq fadd (\<lambda>k. fmul (Rep_matrix A j k) (Rep_matrix B k i)) (max (ncols A) (nrows B))"
-apply (simp add: mult_matrix_def mult_matrix_n_def)
-apply (subst RepAbs_matrix)
-apply (rule exI[of _ "nrows A"], insert assms, simp add: nrows foldseq_zero)
-apply (rule exI[of _ "ncols B"], insert assms, simp add: ncols foldseq_zero)
-apply simp
-done
+  using assms
+  apply (simp add: mult_matrix_def mult_matrix_n_def)
+  apply (subst RepAbs_matrix)
+    apply (rule exI[of _ "nrows A"], simp add: nrows foldseq_zero)
+   apply (rule exI[of _ "ncols B"], simp add: ncols foldseq_zero)
+  apply simp
+  done
 
 lemma transpose_mult_matrix:
   assumes
@@ -1252,7 +1051,7 @@
 lemma le_combine_matrix:
   assumes
   "f 0 0 = 0"
-  "\<forall>a b c d. a \<le> b & c \<le> d \<longrightarrow> f a c \<le> f b d"
+  "\<forall>a b c d. a \<le> b \<and> c \<le> d \<longrightarrow> f a c \<le> f b d"
   "A \<le> B"
   "C \<le> D"
   shows "combine_matrix f A C \<le> combine_matrix f B D"
@@ -1263,8 +1062,7 @@
   "f 0 0 = 0"
   "\<forall>a b c. a \<le> b \<longrightarrow> f c a \<le> f c b"
   "A \<le> B"
-  shows
-  "combine_matrix f C A \<le> combine_matrix f C B"
+  shows "combine_matrix f C A \<le> combine_matrix f C B"
   using assms by (simp add: le_matrix_def)
 
 lemma le_right_combine_matrix:
@@ -1272,8 +1070,7 @@
   "f 0 0 = 0"
   "\<forall>a b c. a \<le> b \<longrightarrow> f a c \<le> f b c"
   "A \<le> B"
-  shows
-  "combine_matrix f A C \<le> combine_matrix f B C"
+  shows "combine_matrix f A C \<le> combine_matrix f B C"
   using assms by (simp add: le_matrix_def)
 
 lemma le_transpose_matrix: "(A \<le> B) = (transpose_matrix A \<le> transpose_matrix B)"
@@ -1281,10 +1078,9 @@
 
 lemma le_foldseq:
   assumes
-  "\<forall>a b c d . a \<le> b & c \<le> d \<longrightarrow> f a c \<le> f b d"
+  "\<forall>a b c d . a \<le> b \<and> c \<le> d \<longrightarrow> f a c \<le> f b d"
   "\<forall>i. i \<le> n \<longrightarrow> s i \<le> t i"
-  shows
-  "foldseq f s n \<le> foldseq f t n"
+  shows "foldseq f s n \<le> foldseq f t n"
 proof -
   have "\<forall>s t. (\<forall>i. i<=n \<longrightarrow> s i \<le> t i) \<longrightarrow> foldseq f s n \<le> foldseq f t n"
     by (induct n) (simp_all add: assms)
@@ -1293,18 +1089,16 @@
 
 lemma le_left_mult:
   assumes
-  "\<forall>a b c d. a \<le> b & c \<le> d \<longrightarrow> fadd a c \<le> fadd b d"
-  "\<forall>c a b.   0 \<le> c & a \<le> b \<longrightarrow> fmul c a \<le> fmul c b"
+  "\<forall>a b c d. a \<le> b \<and> c \<le> d \<longrightarrow> fadd a c \<le> fadd b d"
+  "\<forall>c a b.   0 \<le> c \<and> a \<le> b \<longrightarrow> fmul c a \<le> fmul c b"
   "\<forall>a. fmul 0 a = 0"
   "\<forall>a. fmul a 0 = 0"
   "fadd 0 0 = 0"
   "0 \<le> C"
   "A \<le> B"
-  shows
-  "mult_matrix fmul fadd C A \<le> mult_matrix fmul fadd C B"
+  shows "mult_matrix fmul fadd C A \<le> mult_matrix fmul fadd C B"
   using assms
-  apply (simp add: le_matrix_def Rep_mult_matrix)
-  apply (auto)
+  apply (auto simp: le_matrix_def Rep_mult_matrix)
   apply (simplesubst foldseq_zerotail[of _ _ _ "max (ncols C) (max (nrows A) (nrows B))"], simp_all add: nrows ncols max1 max2)+
   apply (rule le_foldseq)
   apply (auto)
@@ -1312,62 +1106,70 @@
 
 lemma le_right_mult:
   assumes
-  "\<forall>a b c d. a \<le> b & c \<le> d \<longrightarrow> fadd a c \<le> fadd b d"
-  "\<forall>c a b. 0 \<le> c & a \<le> b \<longrightarrow> fmul a c \<le> fmul b c"
+  "\<forall>a b c d. a \<le> b \<and> c \<le> d \<longrightarrow> fadd a c \<le> fadd b d"
+  "\<forall>c a b. 0 \<le> c \<and> a \<le> b \<longrightarrow> fmul a c \<le> fmul b c"
   "\<forall>a. fmul 0 a = 0"
   "\<forall>a. fmul a 0 = 0"
   "fadd 0 0 = 0"
   "0 \<le> C"
   "A \<le> B"
-  shows
-  "mult_matrix fmul fadd A C \<le> mult_matrix fmul fadd B C"
+  shows "mult_matrix fmul fadd A C \<le> mult_matrix fmul fadd B C"
   using assms
-  apply (simp add: le_matrix_def Rep_mult_matrix)
-  apply (auto)
+  apply (auto simp: le_matrix_def Rep_mult_matrix)
   apply (simplesubst foldseq_zerotail[of _ _ _ "max (nrows C) (max (ncols A) (ncols B))"], simp_all add: nrows ncols max1 max2)+
   apply (rule le_foldseq)
   apply (auto)
   done
 
 lemma spec2: "\<forall>j i. P j i \<Longrightarrow> P j i" by blast
-lemma neg_imp: "(\<not> Q \<longrightarrow> \<not> P) \<Longrightarrow> P \<longrightarrow> Q" by blast
 
 lemma singleton_matrix_le[simp]: "(singleton_matrix j i a \<le> singleton_matrix j i b) = (a \<le> (b::_::order))"
   by (auto simp: le_matrix_def)
 
 lemma singleton_le_zero[simp]: "(singleton_matrix j i x \<le> 0) = (x \<le> (0::'a::{order,zero}))"
-  apply (auto)
-  apply (simp add: le_matrix_def)
-  apply (drule_tac j=j and i=i in spec2)
-  apply (simp)
-  apply (simp add: le_matrix_def)
-  done
+  by (metis singleton_matrix_le singleton_matrix_zero)
 
 lemma singleton_ge_zero[simp]: "(0 \<le> singleton_matrix j i x) = ((0::'a::{order,zero}) \<le> x)"
-  apply (auto)
-  apply (simp add: le_matrix_def)
-  apply (drule_tac j=j and i=i in spec2)
-  apply (simp)
-  apply (simp add: le_matrix_def)
-  done
+  by (metis singleton_matrix_le singleton_matrix_zero)
+
+lemma move_matrix_le_zero[simp]: 
+  fixes A:: "'a::{order,zero} matrix"
+  assumes "0 \<le> j" "0 \<le> i"
+  shows "(move_matrix A j i \<le> 0) = (A \<le> 0)"
+proof -
+  have "Rep_matrix A j' i' \<le> 0"
+    if "\<forall>n m. \<not> int n < j \<and> \<not> int m < i \<longrightarrow> Rep_matrix A (nat (int n - j)) (nat (int m - i)) \<le> 0"
+    for j' i'
+    using that[rule_format, of "j' + nat j" "i' + nat i"] by (simp add: assms)
+  then show ?thesis
+    by (auto simp: le_matrix_def)
+qed
 
-lemma move_matrix_le_zero[simp]: "0 \<le> j \<Longrightarrow> 0 \<le> i \<Longrightarrow> (move_matrix A j i \<le> 0) = (A \<le> (0::('a::{order,zero}) matrix))"
-  apply (auto simp: le_matrix_def)
-  apply (drule_tac j="ja+(nat j)" and i="ia+(nat i)" in spec2)
-  apply (auto)
-  done
+lemma move_matrix_zero_le[simp]:
+  fixes A:: "'a::{order,zero} matrix"
+  assumes "0 \<le> j" "0 \<le> i"
+  shows  "(0 \<le> move_matrix A j i) = (0 \<le> A)"
+proof -
+  have "0 \<le> Rep_matrix A j' i'"
+    if "\<forall>n m. \<not> int n < j \<and> \<not> int m < i \<longrightarrow> 0 \<le> Rep_matrix A (nat (int n - j)) (nat (int m - i))"
+    for j' i'
+    using that[rule_format, of "j' + nat j" "i' + nat i"] by (simp add: assms)
+  then show ?thesis
+    by (auto simp: le_matrix_def)
+qed
 
-lemma move_matrix_zero_le[simp]: "0 \<le> j \<Longrightarrow> 0 \<le> i \<Longrightarrow> (0 \<le> move_matrix A j i) = ((0::('a::{order,zero}) matrix) \<le> A)"
-  apply (auto simp: le_matrix_def)
-  apply (drule_tac j="ja+(nat j)" and i="ia+(nat i)" in spec2)
-  apply (auto)
-  done
-
-lemma move_matrix_le_move_matrix_iff[simp]: "0 \<le> j \<Longrightarrow> 0 \<le> i \<Longrightarrow> (move_matrix A j i \<le> move_matrix B j i) = (A \<le> (B::('a::{order,zero}) matrix))"
-  apply (auto simp: le_matrix_def)
-  apply (drule_tac j="ja+(nat j)" and i="ia+(nat i)" in spec2)
-  apply (auto)
-  done  
+lemma move_matrix_le_move_matrix_iff[simp]:
+  fixes A:: "'a::{order,zero} matrix"
+  assumes "0 \<le> j" "0 \<le> i"
+  shows "(move_matrix A j i \<le> move_matrix B j i) = (A \<le> B)"
+proof -
+  have "Rep_matrix A j' i' \<le> Rep_matrix B j' i'"
+    if "\<forall>n m. \<not> int n < j \<and> \<not> int m < i \<longrightarrow> Rep_matrix A (nat (int n - j)) (nat (int m - i)) \<le> Rep_matrix B (nat (int n - j)) (nat (int m - i))"
+    for j' i'
+    using that[rule_format, of "j' + nat j" "i' + nat i"] by (simp add: assms)
+  then show ?thesis
+    by (auto simp: le_matrix_def)
+qed
 
 instantiation matrix :: ("{lattice, zero}") lattice
 begin
@@ -1434,36 +1236,21 @@
 instance matrix :: (monoid_add) monoid_add
 proof
   fix A B C :: "'a matrix"
-  show "A + B + C = A + (B + C)"    
-    apply (simp add: plus_matrix_def)
-    apply (rule combine_matrix_assoc[simplified associative_def, THEN spec, THEN spec, THEN spec])
-    apply (simp_all add: add.assoc)
-    done
+  show "A + B + C = A + (B + C)"
+    by (simp add: add.assoc matrix_eqI plus_matrix_def)    
   show "0 + A = A"
-    apply (simp add: plus_matrix_def)
-    apply (rule combine_matrix_zero_l_neutral[simplified zero_l_neutral_def, THEN spec])
-    apply (simp)
-    done
+    by (simp add: matrix_eqI plus_matrix_def)
   show "A + 0 = A"
-    apply (simp add: plus_matrix_def)
-    apply (rule combine_matrix_zero_r_neutral [simplified zero_r_neutral_def, THEN spec])
-    apply (simp)
-    done
+    by (simp add: matrix_eqI plus_matrix_def)
 qed
 
 instance matrix :: (comm_monoid_add) comm_monoid_add
 proof
   fix A B :: "'a matrix"
   show "A + B = B + A"
-    apply (simp add: plus_matrix_def)
-    apply (rule combine_matrix_commute[simplified commutative_def, THEN spec, THEN spec])
-    apply (simp_all add: add.commute)
-    done
+    by (simp add: add.commute matrix_eqI plus_matrix_def)
   show "0 + A = A"
-    apply (simp add: plus_matrix_def)
-    apply (rule combine_matrix_zero_l_neutral[simplified zero_l_neutral_def, THEN spec])
-    apply (simp)
-    done
+    by (simp add: plus_matrix_def matrix_eqI)
 qed
 
 instance matrix :: (group_add) group_add
@@ -1489,10 +1276,7 @@
   fix A B C :: "'a matrix"
   assume "A \<le> B"
   then show "C + A \<le> C + B"
-    apply (simp add: plus_matrix_def)
-    apply (rule le_left_combine_matrix)
-    apply (simp_all)
-    done
+    by (simp add: le_matrix_def plus_matrix_def)
 qed
   
 instance matrix :: (lattice_ab_group_add) semilattice_inf_ab_group_add ..
@@ -1502,41 +1286,28 @@
 proof
   fix A B C :: "'a matrix"
   show "A * B * C = A * (B * C)"
-    apply (simp add: times_matrix_def)
-    apply (rule mult_matrix_assoc)
-    apply (simp_all add: associative_def algebra_simps)
-    done
+    unfolding times_matrix_def
+    by (smt (verit, best) add.assoc associative_def distrib_left distrib_right group_cancel.add2 mult.assoc mult_matrix_assoc mult_not_zero)
   show "(A + B) * C = A * C + B * C"
-    apply (simp add: times_matrix_def plus_matrix_def)
-    apply (rule l_distributive_matrix[simplified l_distributive_def, THEN spec, THEN spec, THEN spec])
-    apply (simp_all add: associative_def commutative_def algebra_simps)
-    done
+    unfolding times_matrix_def plus_matrix_def
+    using l_distributive_matrix
+    by (metis (full_types) add.assoc add.commute associative_def commutative_def distrib_right l_distributive_def mult_not_zero)
   show "A * (B + C) = A * B + A * C"
-    apply (simp add: times_matrix_def plus_matrix_def)
-    apply (rule r_distributive_matrix[simplified r_distributive_def, THEN spec, THEN spec, THEN spec])
-    apply (simp_all add: associative_def commutative_def algebra_simps)
-    done
-  show "0 * A = 0" by (simp add: times_matrix_def)
-  show "A * 0 = 0" by (simp add: times_matrix_def)
-qed
+    unfolding times_matrix_def plus_matrix_def
+    using r_distributive_matrix
+    by (metis (no_types, lifting) add.assoc add.commute associative_def commutative_def distrib_left mult_zero_left mult_zero_right r_distributive_def) 
+qed (auto simp: times_matrix_def)
 
 instance matrix :: (ring) ring ..
 
 instance matrix :: (ordered_ring) ordered_ring
 proof
   fix A B C :: "'a matrix"
-  assume a: "A \<le> B"
-  assume b: "0 \<le> C"
-  from a b show "C * A \<le> C * B"
-    apply (simp add: times_matrix_def)
-    apply (rule le_left_mult)
-    apply (simp_all add: add_mono mult_left_mono)
-    done
-  from a b show "A * C \<le> B * C"
-    apply (simp add: times_matrix_def)
-    apply (rule le_right_mult)
-    apply (simp_all add: add_mono mult_right_mono)
-    done
+  assume \<section>: "A \<le> B" "0 \<le> C"
+  from \<section> show "C * A \<le> C * B"
+    by (simp add: times_matrix_def add_mono le_left_mult mult_left_mono)
+  from \<section> show "A * C \<le> B * C"
+    by (simp add: times_matrix_def add_mono le_right_mult mult_right_mono)
 qed
 
 instance matrix :: (lattice_ring) lattice_ring
@@ -1558,9 +1329,7 @@
 
 lemma Rep_matrix_mult: "Rep_matrix ((a::('a::semiring_0) matrix) * b) j i = 
   foldseq (+) (\<lambda>k.  (Rep_matrix a j k) * (Rep_matrix b k i)) (max (ncols a) (nrows b))"
-apply (simp add: times_matrix_def)
-apply (simp add: Rep_mult_matrix)
-done
+  by (simp add: times_matrix_def Rep_mult_matrix)
 
 lemma apply_matrix_add: "\<forall>x y. f (x+y) = (f x) + (f y) \<Longrightarrow> f 0 = (0::'a)
   \<Longrightarrow> apply_matrix f ((a::('a::monoid_add) matrix) + b) = (apply_matrix f a) + (apply_matrix f b)"
@@ -1570,20 +1339,18 @@
   by (simp add: matrix_eqI)
 
 lemma nrows_mult: "nrows ((A::('a::semiring_0) matrix) * B) \<le> nrows A"
-by (simp add: times_matrix_def mult_nrows)
+  by (simp add: times_matrix_def mult_nrows)
 
 lemma ncols_mult: "ncols ((A::('a::semiring_0) matrix) * B) \<le> ncols B"
-by (simp add: times_matrix_def mult_ncols)
+  by (simp add: times_matrix_def mult_ncols)
 
 definition
   one_matrix :: "nat \<Rightarrow> ('a::{zero,one}) matrix" where
-  "one_matrix n = Abs_matrix (\<lambda>j i. if j = i & j < n then 1 else 0)"
+  "one_matrix n = Abs_matrix (\<lambda>j i. if j = i \<and> j < n then 1 else 0)"
 
-lemma Rep_one_matrix[simp]: "Rep_matrix (one_matrix n) j i = (if (j = i & j < n) then 1 else 0)"
-apply (simp add: one_matrix_def)
-apply (simplesubst RepAbs_matrix)
-apply (rule exI[of _ n], simp add: if_split)+
-by (simp add: if_split)
+lemma Rep_one_matrix[simp]: "Rep_matrix (one_matrix n) j i = (if (j = i \<and> j < n) then 1 else 0)"
+  unfolding one_matrix_def
+  by (smt (verit, del_insts) RepAbs_matrix not_le)
 
 lemma nrows_one_matrix[simp]: "nrows ((one_matrix n) :: ('a::zero_neq_one)matrix) = n" (is "?r = _")
 proof -
@@ -1599,32 +1366,41 @@
   ultimately show "?r = n" by simp
 qed
 
-lemma one_matrix_mult_right[simp]: "ncols A \<le> n \<Longrightarrow> (A::('a::{semiring_1}) matrix) * (one_matrix n) = A"
+lemma one_matrix_mult_right[simp]:
+  fixes A :: "('a::semiring_1) matrix"
+  shows "ncols A \<le> n \<Longrightarrow> A * (one_matrix n) = A"
   apply (intro matrix_eqI)
   apply (simp add: times_matrix_def Rep_mult_matrix)
   apply (subst foldseq_almostzero, auto simp: ncols)
   done
 
-lemma one_matrix_mult_left[simp]: "nrows A \<le> n \<Longrightarrow> (one_matrix n) * A = (A::('a::semiring_1) matrix)"
+lemma one_matrix_mult_left[simp]: 
+  fixes A :: "('a::semiring_1) matrix"
+  shows "nrows A \<le> n \<Longrightarrow> (one_matrix n) * A = A"
   apply (intro matrix_eqI)
   apply (simp add: times_matrix_def Rep_mult_matrix)
   apply (subst foldseq_almostzero, auto simp: nrows)
   done
 
-lemma transpose_matrix_mult: "transpose_matrix ((A::('a::comm_ring) matrix)*B) = (transpose_matrix B) * (transpose_matrix A)"
-apply (simp add: times_matrix_def)
-apply (subst transpose_mult_matrix)
-apply (simp_all add: mult.commute)
-done
+lemma transpose_matrix_mult:
+  fixes A :: "('a::comm_ring) matrix"
+  shows  "transpose_matrix (A*B) = (transpose_matrix B) * (transpose_matrix A)"
+  by (simp add: times_matrix_def transpose_mult_matrix mult.commute)
+
+lemma transpose_matrix_add:
+  fixes A :: "('a::monoid_add) matrix"
+  shows  "transpose_matrix (A+B) = transpose_matrix A + transpose_matrix B"
+  by (simp add: plus_matrix_def transpose_combine_matrix)
 
-lemma transpose_matrix_add: "transpose_matrix ((A::('a::monoid_add) matrix)+B) = transpose_matrix A + transpose_matrix B"
-by (simp add: plus_matrix_def transpose_combine_matrix)
+lemma transpose_matrix_diff:
+  fixes A :: "('a::group_add) matrix"
+  shows "transpose_matrix (A-B) = transpose_matrix A - transpose_matrix B"
+  by (simp add: diff_matrix_def transpose_combine_matrix)
 
-lemma transpose_matrix_diff: "transpose_matrix ((A::('a::group_add) matrix)-B) = transpose_matrix A - transpose_matrix B"
-by (simp add: diff_matrix_def transpose_combine_matrix)
-
-lemma transpose_matrix_minus: "transpose_matrix (-(A::('a::group_add) matrix)) = - transpose_matrix (A::'a matrix)"
-by (simp add: minus_matrix_def transpose_apply_matrix)
+lemma transpose_matrix_minus: 
+  fixes A :: "('a::group_add) matrix"
+  shows "transpose_matrix (-A) = - transpose_matrix (A::'a matrix)"
+  by (simp add: minus_matrix_def transpose_apply_matrix)
 
 definition right_inverse_matrix :: "('a::{ring_1}) matrix \<Rightarrow> 'a matrix \<Rightarrow> bool" where
   "right_inverse_matrix A X == (A * X = one_matrix (max (nrows A) (ncols X))) \<and> nrows X \<le> ncols A" 
@@ -1636,32 +1412,25 @@
   "inverse_matrix A X == (right_inverse_matrix A X) \<and> (left_inverse_matrix A X)"
 
 lemma right_inverse_matrix_dim: "right_inverse_matrix A X \<Longrightarrow> nrows A = ncols X"
-apply (insert ncols_mult[of A X], insert nrows_mult[of A X])
-by (simp add: right_inverse_matrix_def)
+  using ncols_mult[of A X] nrows_mult[of A X]
+  by (simp add: right_inverse_matrix_def)
 
 lemma left_inverse_matrix_dim: "left_inverse_matrix A Y \<Longrightarrow> ncols A = nrows Y"
-apply (insert ncols_mult[of Y A], insert nrows_mult[of Y A]) 
-by (simp add: left_inverse_matrix_def)
+  using ncols_mult[of Y A] nrows_mult[of Y A]
+  by (simp add: left_inverse_matrix_def)
 
 lemma left_right_inverse_matrix_unique: 
   assumes "left_inverse_matrix A Y" "right_inverse_matrix A X"
   shows "X = Y"
 proof -
-  have "Y = Y * one_matrix (nrows A)" 
-    apply (subst one_matrix_mult_right)
-    using assms
-    apply (simp_all add: left_inverse_matrix_def)
-    done
-  also have "\<dots> = Y * (A * X)" 
-    apply (insert assms)
-    apply (frule right_inverse_matrix_dim)
-    by (simp add: right_inverse_matrix_def)
+  have "Y = Y * one_matrix (nrows A)"
+    by (metis assms(1) left_inverse_matrix_def one_matrix_mult_right) 
+  also have "\<dots> = Y * (A * X)"
+    by (metis assms(2) max.idem right_inverse_matrix_def right_inverse_matrix_dim) 
   also have "\<dots> = (Y * A) * X" by (simp add: mult.assoc)
-  also have "\<dots> = X" 
-    apply (insert assms)
-    apply (frule left_inverse_matrix_dim)
-    apply (simp_all add:  left_inverse_matrix_def right_inverse_matrix_def one_matrix_mult_left)
-    done
+  also have "\<dots> = X"
+    using assms left_inverse_matrix_def right_inverse_matrix_def
+    by (metis left_inverse_matrix_dim max.idem one_matrix_mult_left)
   ultimately show "X = Y" by (simp)
 qed
 
@@ -1672,19 +1441,18 @@
   by (simp add: inverse_matrix_def left_inverse_matrix_def right_inverse_matrix_def)
 
 lemma zero_imp_mult_zero: "(a::'a::semiring_0) = 0 | b = 0 \<Longrightarrow> a * b = 0"
-by auto
+  by auto
 
 lemma Rep_matrix_zero_imp_mult_zero:
   "\<forall>j i k. (Rep_matrix A j k = 0) | (Rep_matrix B k i) = 0  \<Longrightarrow> A * B = (0::('a::lattice_ring) matrix)"
   by (simp add: matrix_eqI Rep_matrix_mult foldseq_zero zero_imp_mult_zero)
 
 lemma add_nrows: "nrows (A::('a::monoid_add) matrix) \<le> u \<Longrightarrow> nrows B \<le> u \<Longrightarrow> nrows (A + B) \<le> u"
-apply (simp add: plus_matrix_def)
-apply (rule combine_nrows)
-apply (simp_all)
-done
+  by (simp add: nrows_le)
 
-lemma move_matrix_row_mult: "move_matrix ((A::('a::semiring_0) matrix) * B) j 0 = (move_matrix A j 0) * B"
+lemma move_matrix_row_mult:
+  fixes A :: "('a::semiring_0) matrix"
+  shows "move_matrix (A * B) j 0 = (move_matrix A j 0) * B"
 proof -
   have "\<And>m. \<not> int m < j \<Longrightarrow> ncols (move_matrix A j 0) \<le> max (ncols A) (nrows B)"
     by (smt (verit, best) max1 nat_int ncols_move_matrix_le)
@@ -1696,7 +1464,9 @@
     done
 qed
 
-lemma move_matrix_col_mult: "move_matrix ((A::('a::semiring_0) matrix) * B) 0 i = A * (move_matrix B 0 i)"
+lemma move_matrix_col_mult:
+  fixes A :: "('a::semiring_0) matrix"
+  shows  "move_matrix (A * B) 0 i = A * (move_matrix B 0 i)"
 proof -
   have "\<And>n. \<not> int n < i \<Longrightarrow> nrows (move_matrix B 0 i) \<le> max (ncols A) (nrows B)"
     by (smt (verit, del_insts) max2 nat_int nrows_move_matrix_le)
--- a/src/HOL/Matrix_LP/SparseMatrix.thy	Mon Aug 26 18:26:00 2024 +0200
+++ b/src/HOL/Matrix_LP/SparseMatrix.thy	Mon Aug 26 18:26:06 2024 +0200
@@ -1,5 +1,5 @@
 (*  Title:      HOL/Matrix_LP/SparseMatrix.thy
-    Author:     Steven Obua
+    Author:     Steven Obua; updated to Isar by LCP
 *)
 
 theory SparseMatrix
@@ -291,7 +291,7 @@
         using sorted_sparse_row_matrix_zero apply fastforce
         apply (subst Rep_matrix_zero_imp_mult_zero)
          apply (metis Rep_move_matrix comp_1 nrows_le nrows_spvec sorted_sparse_row_vector_zero verit_comp_simplify1(3))
-        apply (simp add: )
+        apply simp
         done
     next
       case greater
@@ -434,9 +434,7 @@
 qed
 
 lemma sorted_spmat_add_spmat[rule_format]: "sorted_spmat A \<Longrightarrow> sorted_spmat B \<Longrightarrow> sorted_spmat (add_spmat A B)"
-  apply (induct A B rule: add_spmat.induct)
-  apply (simp_all add: sorted_spvec_add_spvec)
-  done
+  by (induct A B rule: add_spmat.induct) (simp_all add: sorted_spvec_add_spvec)
 
 fun le_spvec :: "('a::lattice_ab_group_add) spvec \<Rightarrow> 'a spvec \<Rightarrow> bool"
 where
@@ -464,8 +462,6 @@
   "disj_matrices A B \<longleftrightarrow>
     (\<forall>j i. (Rep_matrix A j i \<noteq> 0) \<longrightarrow> (Rep_matrix B j i = 0)) & (\<forall>j i. (Rep_matrix B j i \<noteq> 0) \<longrightarrow> (Rep_matrix A j i = 0))"  
 
-declare [[simp_depth_limit = 6]]
-
 lemma disj_matrices_contr1: "disj_matrices A B \<Longrightarrow> Rep_matrix A j i \<noteq> 0 \<Longrightarrow> Rep_matrix B j i = 0"
    by (simp add: disj_matrices_def)
 
@@ -473,73 +469,47 @@
    by (simp add: disj_matrices_def)
 
 
-lemma disj_matrices_add: "disj_matrices A B \<Longrightarrow> disj_matrices C D \<Longrightarrow> disj_matrices A D \<Longrightarrow> disj_matrices B C \<Longrightarrow> 
-  (A + B \<le> C + D) = (A \<le> C & B \<le> (D::('a::lattice_ab_group_add) matrix))"
-  apply (auto)
-  apply (simp (no_asm_use) only: le_matrix_def disj_matrices_def)
-  apply (intro strip)
-  apply (erule conjE)+
-  apply (drule_tac j=j and i=i in spec2)+
-  apply (case_tac "Rep_matrix B j i = 0")
-  apply (case_tac "Rep_matrix D j i = 0")
-  apply (simp_all)
-  apply (simp (no_asm_use) only: le_matrix_def disj_matrices_def)
-  apply (intro strip)
-  apply (erule conjE)+
-  apply (drule_tac j=j and i=i in spec2)+
-  apply (case_tac "Rep_matrix A j i = 0")
-  apply (case_tac "Rep_matrix C j i = 0")
-  apply (simp_all)
-  apply (erule add_mono)
-  apply (assumption)
-  done
+lemma disj_matrices_add:
+  fixes A :: "('a::lattice_ab_group_add) matrix"
+  shows  "disj_matrices A B \<Longrightarrow> disj_matrices C D \<Longrightarrow> disj_matrices A D 
+       \<Longrightarrow> disj_matrices B C \<Longrightarrow> (A + B \<le> C + D) = (A \<le> C \<and> B \<le> D)"
+  apply (intro iffI conjI)
+  unfolding le_matrix_def disj_matrices_def
+  apply (metis Rep_matrix_add group_cancel.rule0 order_refl)
+  apply (metis (no_types, lifting) Rep_matrix_add add_cancel_right_left dual_order.refl)
+  by (meson add_mono le_matrix_def)
 
 lemma disj_matrices_zero1[simp]: "disj_matrices 0 B"
-by (simp add: disj_matrices_def)
+  by (simp add: disj_matrices_def)
 
 lemma disj_matrices_zero2[simp]: "disj_matrices A 0"
-by (simp add: disj_matrices_def)
+  by (simp add: disj_matrices_def)
 
 lemma disj_matrices_commute: "disj_matrices A B = disj_matrices B A"
-by (auto simp: disj_matrices_def)
+  by (auto simp: disj_matrices_def)
 
 lemma disj_matrices_add_le_zero: "disj_matrices A B \<Longrightarrow>
   (A + B \<le> 0) = (A \<le> 0 & (B::('a::lattice_ab_group_add) matrix) \<le> 0)"
-by (rule disj_matrices_add[of A B 0 0, simplified])
- 
+  by (rule disj_matrices_add[of A B 0 0, simplified])
+
 lemma disj_matrices_add_zero_le: "disj_matrices A B \<Longrightarrow>
   (0 \<le> A + B) = (0 \<le> A & 0 \<le> (B::('a::lattice_ab_group_add) matrix))"
-by (rule disj_matrices_add[of 0 0 A B, simplified])
+  by (rule disj_matrices_add[of 0 0 A B, simplified])
 
 lemma disj_matrices_add_x_le: "disj_matrices A B \<Longrightarrow> disj_matrices B C \<Longrightarrow> 
   (A \<le> B + C) = (A \<le> C & 0 \<le> (B::('a::lattice_ab_group_add) matrix))"
-by (auto simp: disj_matrices_add[of 0 A B C, simplified])
+  by (auto simp: disj_matrices_add[of 0 A B C, simplified])
 
 lemma disj_matrices_add_le_x: "disj_matrices A B \<Longrightarrow> disj_matrices B C \<Longrightarrow> 
   (B + A \<le> C) = (A \<le> C &  (B::('a::lattice_ab_group_add) matrix) \<le> 0)"
-by (auto simp: disj_matrices_add[of B A 0 C,simplified] disj_matrices_commute)
+  by (auto simp: disj_matrices_add[of B A 0 C,simplified] disj_matrices_commute)
 
 lemma disj_sparse_row_singleton: "i \<le> j \<Longrightarrow> sorted_spvec((j,y)#v) \<Longrightarrow> disj_matrices (sparse_row_vector v) (singleton_matrix 0 i x)"
   apply (simp add: disj_matrices_def)
-  apply (rule conjI)
-  apply (rule neg_imp)
-  apply (simp)
-  apply (intro strip)
-  apply (rule sorted_sparse_row_vector_zero)
-  apply (simp_all)
-  apply (intro strip)
-  apply (rule sorted_sparse_row_vector_zero)
-  apply (simp_all)
-  done 
+  using sorted_sparse_row_vector_zero by blast
 
 lemma disj_matrices_x_add: "disj_matrices A B \<Longrightarrow> disj_matrices A C \<Longrightarrow> disj_matrices (A::('a::lattice_ab_group_add) matrix) (B+C)"
-  apply (simp add: disj_matrices_def)
-  apply (auto)
-  apply (drule_tac j=j and i=i in spec2)+
-  apply (case_tac "Rep_matrix B j i = 0")
-  apply (case_tac "Rep_matrix C j i = 0")
-  apply (simp_all)
-  done
+  by (smt (verit, ccfv_SIG) Rep_matrix_add add_0 disj_matrices_def)
 
 lemma disj_matrices_add_x: "disj_matrices A B \<Longrightarrow> disj_matrices A C \<Longrightarrow> disj_matrices (B+C) (A::('a::lattice_ab_group_add) matrix)" 
   by (simp add: disj_matrices_x_add disj_matrices_commute)
@@ -547,97 +517,106 @@
 lemma disj_singleton_matrices[simp]: "disj_matrices (singleton_matrix j i x) (singleton_matrix u v y) = (j \<noteq> u | i \<noteq> v | x = 0 | y = 0)" 
   by (auto simp: disj_matrices_def)
 
-lemma disj_move_sparse_vec_mat[simplified disj_matrices_commute]: 
-  "j \<le> a \<Longrightarrow> sorted_spvec((a,c)#as) \<Longrightarrow> disj_matrices (move_matrix (sparse_row_vector b) (int j) i) (sparse_row_matrix as)"
-  apply (auto simp: disj_matrices_def)
-  apply (drule nrows_notzero)
-  apply (drule less_le_trans[OF _ nrows_spvec])
-  apply (subgoal_tac "ja = j")
-  apply (simp add: sorted_sparse_row_matrix_zero)
-  apply (arith)
-  apply (rule nrows)
-  apply (rule order_trans[of _ 1 _])
-  apply (simp)
-  apply (case_tac "nat (int ja - int j) = 0")
-  apply (case_tac "ja = j")
-  apply (simp add: sorted_sparse_row_matrix_zero)
-  apply arith+
-  done
+lemma disj_move_sparse_vec_mat:
+  assumes "j \<le> a" and "sorted_spvec ((a, c) # as)"
+  shows "disj_matrices (sparse_row_matrix as) (move_matrix (sparse_row_vector b) (int j) i)"
+proof -
+  have "Rep_matrix (sparse_row_vector b) (n-j) (nat (int m - i)) = 0"
+    if "\<not> n<j" and nz: "Rep_matrix (sparse_row_matrix as) n m \<noteq> 0"
+    for n m
+  proof -
+    have "n \<noteq> j"
+      using assms sorted_sparse_row_matrix_zero nz by blast
+    with that have "j < n"  by auto
+    then show ?thesis
+      by (metis One_nat_def Suc_diff_Suc nrows nrows_spvec plus_1_eq_Suc trans_le_add1)
+  qed
+  then show ?thesis
+    by (auto simp: disj_matrices_def nat_minus_as_int)
+qed
 
 lemma disj_move_sparse_row_vector_twice:
   "j \<noteq> u \<Longrightarrow> disj_matrices (move_matrix (sparse_row_vector a) j i) (move_matrix (sparse_row_vector b) u v)"
-  apply (auto simp: disj_matrices_def)
-  apply (rule nrows, rule order_trans[of _ 1], simp, drule nrows_notzero, drule less_le_trans[OF _ nrows_spvec], arith)+
-  done
-
-lemma le_spvec_iff_sparse_row_le[rule_format]: "(sorted_spvec a) \<longrightarrow> (sorted_spvec b) \<longrightarrow> (le_spvec a b) = (sparse_row_vector a \<le> sparse_row_vector b)"
-  apply (induct a b rule: le_spvec.induct)
-  apply (simp_all add: sorted_spvec_cons1 disj_matrices_add_le_zero disj_matrices_add_zero_le 
-    disj_sparse_row_singleton[OF order_refl] disj_matrices_commute)
-  apply (rule conjI, intro strip)
-  apply (simp add: sorted_spvec_cons1)
-  apply (subst disj_matrices_add_x_le)
-  apply (simp add: disj_sparse_row_singleton[OF less_imp_le] disj_matrices_x_add disj_matrices_commute)
-  apply (simp add: disj_sparse_row_singleton[OF order_refl] disj_matrices_commute)
-  apply (simp, blast)
-  apply (intro strip, rule conjI, intro strip)
-  apply (simp add: sorted_spvec_cons1)
-  apply (subst disj_matrices_add_le_x)
-  apply (simp_all add: disj_sparse_row_singleton[OF order_refl] disj_sparse_row_singleton[OF less_imp_le] disj_matrices_commute disj_matrices_x_add)
-  apply (blast)
-  apply (intro strip)
-  apply (simp add: sorted_spvec_cons1)
-  apply (case_tac "a=b", simp_all)
-  apply (subst disj_matrices_add)
-  apply (simp_all add: disj_sparse_row_singleton[OF order_refl] disj_matrices_commute)
-  done
-
-lemma le_spvec_empty2_sparse_row[rule_format]: "sorted_spvec b \<longrightarrow> le_spvec b [] = (sparse_row_vector b \<le> 0)"
-  apply (induct b)
-  apply (simp_all add: sorted_spvec_cons1)
-  apply (intro strip)
-  apply (subst disj_matrices_add_le_zero)
-  apply (auto simp: disj_matrices_commute disj_sparse_row_singleton[OF order_refl] sorted_spvec_cons1)
-  done
+  unfolding disj_matrices_def
+  by (smt (verit, ccfv_SIG) One_nat_def Rep_move_matrix of_nat_1 le_nat_iff nrows nrows_spvec of_nat_le_iff)
 
-lemma le_spvec_empty1_sparse_row[rule_format]: "(sorted_spvec b) \<longrightarrow> (le_spvec [] b = (0 \<le> sparse_row_vector b))"
-  apply (induct b)
-  apply (simp_all add: sorted_spvec_cons1)
-  apply (intro strip)
-  apply (subst disj_matrices_add_zero_le)
-  apply (auto simp: disj_matrices_commute disj_sparse_row_singleton[OF order_refl] sorted_spvec_cons1)
-  done
+lemma le_spvec_iff_sparse_row_le:
+  "sorted_spvec a \<Longrightarrow> sorted_spvec b \<Longrightarrow> (le_spvec a b) \<longleftrightarrow> (sparse_row_vector a \<le> sparse_row_vector b)"
+proof (induct a b rule: le_spvec.induct)
+  case 1
+  then show ?case
+    by auto
+next
+  case (2 uu a as)
+  then have "sorted_spvec as"
+    by (metis sorted_spvec_cons1)
+  with 2 show ?case
+    apply (simp add: add.commute)
+    by (metis disj_matrices_add_le_zero disj_sparse_row_singleton le_refl singleton_le_zero)
+next
+  case (3 uv b bs)
+  then have "sorted_spvec bs"
+    by (metis sorted_spvec_cons1)
+  with 3 show ?case
+    apply (simp add: add.commute)
+    by (metis disj_matrices_add_zero_le disj_sparse_row_singleton le_refl singleton_ge_zero)
+next
+  case (4 i a as j b bs)
+  then obtain \<section>: "sorted_spvec as" "sorted_spvec bs"
+    by (metis sorted_spvec_cons1)
+  show ?case
+  proof (cases i j rule: linorder_cases)
+    case less
+    with 4 \<section> show ?thesis
+      apply (simp add: )
+      by (metis disj_matrices_add_le_x disj_matrices_add_x disj_matrices_commute disj_singleton_matrices disj_sparse_row_singleton less_imp_le_nat singleton_le_zero not_le)
+  next
+    case equal
+    with 4 \<section> show ?thesis
+      apply (simp add: )
+      by (metis disj_matrices_add disj_matrices_commute disj_sparse_row_singleton order_refl singleton_matrix_le)
+  next
+    case greater
+    with 4 \<section> show ?thesis
+      apply (simp add: )
+      by (metis disj_matrices_add_x disj_matrices_add_x_le disj_matrices_commute disj_singleton_matrices disj_sparse_row_singleton le_refl order_less_le singleton_ge_zero)
+  qed
+qed
 
-lemma le_spmat_iff_sparse_row_le[rule_format]: "(sorted_spvec A) \<longrightarrow> (sorted_spmat A) \<longrightarrow> (sorted_spvec B) \<longrightarrow> (sorted_spmat B) \<longrightarrow> 
+lemma le_spvec_empty2_sparse_row: 
+  "sorted_spvec b \<Longrightarrow> le_spvec b [] = (sparse_row_vector b \<le> 0)"
+  by (simp add: le_spvec_iff_sparse_row_le)
+
+lemma le_spvec_empty1_sparse_row: 
+  "(sorted_spvec b) \<Longrightarrow> (le_spvec [] b = (0 \<le> sparse_row_vector b))"
+  by (simp add: le_spvec_iff_sparse_row_le)
+
+lemma le_spmat_iff_sparse_row_le:
+  "\<lbrakk>sorted_spvec A; sorted_spmat A; sorted_spvec B; sorted_spmat B\<rbrakk> \<Longrightarrow> 
   le_spmat A B = (sparse_row_matrix A \<le> sparse_row_matrix B)"
-  apply (induct A B rule: le_spmat.induct)
-  apply (simp add: sparse_row_matrix_cons disj_matrices_add_le_zero disj_matrices_add_zero_le disj_move_sparse_vec_mat[OF order_refl] 
-    disj_matrices_commute sorted_spvec_cons1 le_spvec_empty2_sparse_row le_spvec_empty1_sparse_row)+ 
-  apply (rule conjI, intro strip)
-  apply (simp add: sorted_spvec_cons1)
-  apply (subst disj_matrices_add_x_le)
-  apply (rule disj_matrices_add_x)
-  apply (simp add: disj_move_sparse_row_vector_twice)
-  apply (simp add: disj_move_sparse_vec_mat[OF less_imp_le] disj_matrices_commute)
-  apply (simp add: disj_move_sparse_vec_mat[OF order_refl] disj_matrices_commute)
-  apply (simp, blast)
-  apply (intro strip, rule conjI, intro strip)
-  apply (simp add: sorted_spvec_cons1)
-  apply (subst disj_matrices_add_le_x)
-  apply (simp add: disj_move_sparse_vec_mat[OF order_refl])
-  apply (rule disj_matrices_x_add)
-  apply (simp add: disj_move_sparse_row_vector_twice)
-  apply (simp add: disj_move_sparse_vec_mat[OF less_imp_le] disj_matrices_commute)
-  apply (simp, blast)
-  apply (intro strip)
-  apply (case_tac "i=j")
-  apply (simp_all)
-  apply (subst disj_matrices_add)
-  apply (simp_all add: disj_matrices_commute disj_move_sparse_vec_mat[OF order_refl])
-  apply (simp add: sorted_spvec_cons1 le_spvec_iff_sparse_row_le)
-  done
+proof (induct A B rule: le_spmat.induct)
+  case (4 i a as j b bs)
+  then obtain \<section>: "sorted_spvec as" "sorted_spvec bs"
+    by (metis sorted_spvec_cons1)
+  show ?case
+  proof (cases i j rule: linorder_cases)
+    case less
+    with 4 \<section> show ?thesis
+      apply (simp add: sparse_row_matrix_cons le_spvec_empty2_sparse_row)
+      by (metis disj_matrices_add_le_x disj_matrices_add_x disj_matrices_commute disj_move_sparse_row_vector_twice disj_move_sparse_vec_mat int_eq_iff less_not_refl move_matrix_le_zero order_le_less)
+  next
+    case equal
+    with 4 \<section> show ?thesis
+      by (simp add: sparse_row_matrix_cons le_spvec_iff_sparse_row_le disj_matrices_commute disj_move_sparse_vec_mat[OF order_refl] disj_matrices_add)
+  next
+    case greater
+    with 4 \<section> show ?thesis
+      apply (simp add: sparse_row_matrix_cons le_spvec_empty1_sparse_row)
+      by (metis disj_matrices_add_x disj_matrices_add_x_le disj_matrices_commute disj_move_sparse_row_vector_twice disj_move_sparse_vec_mat move_matrix_zero_le nat_int nat_less_le of_nat_0_le_iff order_refl)
+  qed
+qed (auto simp add: sparse_row_matrix_cons disj_matrices_add_le_zero disj_matrices_add_zero_le disj_move_sparse_vec_mat[OF order_refl] 
+    disj_matrices_commute sorted_spvec_cons1 le_spvec_empty2_sparse_row le_spvec_empty1_sparse_row)
 
-declare [[simp_depth_limit = 999]]
 
 primrec abs_spmat :: "('a::lattice_ring) spmat \<Rightarrow> 'a spmat"
 where
@@ -792,7 +771,7 @@
 proof (induct v rule: sorted_spvec.induct)
   case (3 m x n y bs)
   then show ?case
-    apply (simp add: )
+    apply simp
     apply (subst pprt_add)
      apply (metis disj_matrices_commute disj_sparse_row_singleton order.refl fst_conv prod.sel(2) sparse_row_vector_cons)
     by (metis pprt_singleton sorted_spvec_cons1)
@@ -804,7 +783,7 @@
 proof (induct v rule: sorted_spvec.induct)
   case (3 m x n y bs)
   then show ?case
-    apply (simp add: )
+    apply simp
     apply (subst nprt_add)
     apply (metis disj_matrices_commute disj_sparse_row_singleton dual_order.refl fst_conv prod.sel(2) sparse_row_vector_cons)
     using sorted_spvec_cons1 by force
@@ -868,10 +847,10 @@
 qed
 
 lemma sorted_nprt_spvec: "sorted_spvec v \<Longrightarrow> sorted_spvec (nprt_spvec v)"
-by (induct v rule: sorted_spvec.induct) (simp_all add: sorted_spvec.simps split:list.split_asm)
+  by (induct v rule: sorted_spvec.induct) (simp_all add: sorted_spvec.simps split:list.split_asm)
 
 lemma sorted_spvec_pprt_spmat: "sorted_spvec m \<Longrightarrow> sorted_spvec (pprt_spmat m)"
-by (induct m rule: sorted_spvec.induct) (simp_all add: sorted_spvec.simps split:list.split_asm)
+  by (induct m rule: sorted_spvec.induct) (simp_all add: sorted_spvec.simps split:list.split_asm)
 
 lemma sorted_spvec_nprt_spmat: "sorted_spvec m \<Longrightarrow> sorted_spvec (nprt_spmat m)"
 by (induct m rule: sorted_spvec.induct) (simp_all add: sorted_spvec.simps split:list.split_asm)
--- a/src/HOL/Metis_Examples/Message.thy	Mon Aug 26 18:26:00 2024 +0200
+++ b/src/HOL/Metis_Examples/Message.thy	Mon Aug 26 18:26:06 2024 +0200
@@ -50,6 +50,8 @@
 text\<open>Concrete syntax: messages appear as \<open>\<lbrace>A,B,NA\<rbrace>\<close>, etc...\<close>
 syntax
   "_MTuple"      :: "['a, args] => 'a * 'b"       ("(2\<lbrace>_,/ _\<rbrace>)")
+syntax_consts
+  "_MTuple"     == MPair
 translations
   "\<lbrace>x, y, z\<rbrace>"   == "\<lbrace>x, \<lbrace>y, z\<rbrace>\<rbrace>"
   "\<lbrace>x, y\<rbrace>"      == "CONST MPair x y"
--- a/src/HOL/Orderings.thy	Mon Aug 26 18:26:00 2024 +0200
+++ b/src/HOL/Orderings.thy	Mon Aug 26 18:26:06 2024 +0200
@@ -745,6 +745,10 @@
   "_All_neq" :: "[idt, 'a, bool] => bool"    ("(3! _~=_./ _)"  [0, 0, 10] 10)
   "_Ex_neq" :: "[idt, 'a, bool] => bool"    ("(3? _~=_./ _)"  [0, 0, 10] 10)
 
+syntax_consts
+  "_All_less" "_All_less_eq" "_All_greater" "_All_greater_eq" "_All_neq" \<rightleftharpoons> All and
+  "_Ex_less" "_Ex_less_eq" "_Ex_greater" "_Ex_greater_eq" "_Ex_neq" \<rightleftharpoons> Ex
+
 translations
   "\<forall>x<y. P" \<rightharpoonup> "\<forall>x. x < y \<longrightarrow> P"
   "\<exists>x<y. P" \<rightharpoonup> "\<exists>x. x < y \<and> P"
--- a/src/HOL/Probability/Fin_Map.thy	Mon Aug 26 18:26:00 2024 +0200
+++ b/src/HOL/Probability/Fin_Map.thy	Mon Aug 26 18:26:06 2024 +0200
@@ -90,6 +90,8 @@
 
 syntax
   "_Pi'" :: "[pttrn, 'a set, 'b set] => ('a => 'b) set"  ("(3\<Pi>'' _\<in>_./ _)"   10)
+syntax_consts
+  "_Pi'" == Pi'
 translations
   "\<Pi>' x\<in>A. B" == "CONST Pi' A (\<lambda>x. B)"
 
@@ -635,6 +637,8 @@
 
 syntax
   "_PiF" :: "pttrn \<Rightarrow> 'i set \<Rightarrow> 'a measure \<Rightarrow> ('i => 'a) measure"  ("(3\<Pi>\<^sub>F _\<in>_./ _)"  10)
+syntax_consts
+  "_PiF" == PiF
 translations
   "\<Pi>\<^sub>F x\<in>I. M" == "CONST PiF I (%x. M)"
 
--- a/src/HOL/Probability/Probability_Measure.thy	Mon Aug 26 18:26:00 2024 +0200
+++ b/src/HOL/Probability/Probability_Measure.thy	Mon Aug 26 18:26:06 2024 +0200
@@ -250,6 +250,9 @@
 syntax
   "_prob" :: "pttrn \<Rightarrow> logic \<Rightarrow> logic \<Rightarrow> logic" (\<open>('\<P>'((/_ in _./ _)'))\<close>)
 
+syntax_consts
+  "_prob" == measure
+
 translations
   "\<P>(x in M. P)" => "CONST measure M {x \<in> CONST space M. P}"
 
@@ -322,6 +325,9 @@
 syntax
   "_conditional_prob" :: "pttrn \<Rightarrow> logic \<Rightarrow> logic \<Rightarrow> logic \<Rightarrow> logic" (\<open>('\<P>'(_ in _. _ \<bar>/ _'))\<close>)
 
+syntax_consts
+  "_conditional_prob" == cond_prob
+
 translations
   "\<P>(x in M. P \<bar> Q)" => "CONST cond_prob M (\<lambda>x. P) (\<lambda>x. Q)"
 
--- a/src/HOL/Product_Type.thy	Mon Aug 26 18:26:00 2024 +0200
+++ b/src/HOL/Product_Type.thy	Mon Aug 26 18:26:06 2024 +0200
@@ -293,6 +293,10 @@
   ""            :: "pttrn \<Rightarrow> patterns"                  ("_")
   "_patterns"   :: "pttrn \<Rightarrow> patterns \<Rightarrow> patterns"      ("_,/ _")
   "_unit"       :: pttrn                                ("'(')")
+syntax_consts
+  "_tuple" "_tuple_arg" "_tuple_args" \<rightleftharpoons> Pair and
+  "_pattern" "_patterns" \<rightleftharpoons> case_prod and
+  "_unit" \<rightleftharpoons> case_unit
 translations
   "(x, y)" \<rightleftharpoons> "CONST Pair x y"
   "_pattern x y" \<rightleftharpoons> "CONST Pair x y"
@@ -1010,6 +1014,8 @@
 
 syntax
   "_Sigma" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b set \<Rightarrow> ('a \<times> 'b) set"  ("(3SIGMA _:_./ _)" [0, 0, 10] 10)
+syntax_consts
+  "_Sigma" \<rightleftharpoons> Sigma
 translations
   "SIGMA x:A. B" \<rightleftharpoons> "CONST Sigma A (\<lambda>x. B)"
 
--- a/src/HOL/Quotient_Examples/Lift_FSet.thy	Mon Aug 26 18:26:00 2024 +0200
+++ b/src/HOL/Quotient_Examples/Lift_FSet.thy	Mon Aug 26 18:26:06 2024 +0200
@@ -98,6 +98,9 @@
 syntax
   "_insert_fset"     :: "args => 'a fset"  ("{|(_)|}")
 
+syntax_consts
+  "_insert_fset" == fcons
+
 translations
   "{|x, xs|}" == "CONST fcons x {|xs|}"
   "{|x|}"     == "CONST fcons x {||}"
--- a/src/HOL/Quotient_Examples/Quotient_FSet.thy	Mon Aug 26 18:26:00 2024 +0200
+++ b/src/HOL/Quotient_Examples/Quotient_FSet.thy	Mon Aug 26 18:26:06 2024 +0200
@@ -291,6 +291,9 @@
 syntax
   "_insert_fset"     :: "args => 'a fset"  ("{|(_)|}")
 
+syntax_consts
+  "_insert_fset" == insert_fset
+
 translations
   "{|x, xs|}" == "CONST insert_fset x {|xs|}"
   "{|x|}"     == "CONST insert_fset x {||}"
--- a/src/HOL/SET_Protocol/Message_SET.thy	Mon Aug 26 18:26:00 2024 +0200
+++ b/src/HOL/SET_Protocol/Message_SET.thy	Mon Aug 26 18:26:06 2024 +0200
@@ -70,6 +70,8 @@
 (*Concrete syntax: messages appear as \<open>\<lbrace>A,B,NA\<rbrace>\<close>, etc...*)
 syntax
   "_MTuple"      :: "['a, args] \<Rightarrow> 'a * 'b"       ("(2\<lbrace>_,/ _\<rbrace>)")
+syntax_consts
+  "_MTuple"     == MPair
 translations
   "\<lbrace>x, y, z\<rbrace>"   == "\<lbrace>x, \<lbrace>y, z\<rbrace>\<rbrace>"
   "\<lbrace>x, y\<rbrace>"      == "CONST MPair x y"
--- a/src/HOL/SPARK/SPARK_Setup.thy	Mon Aug 26 18:26:00 2024 +0200
+++ b/src/HOL/SPARK/SPARK_Setup.thy	Mon Aug 26 18:26:06 2024 +0200
@@ -57,6 +57,9 @@
 syntax
   "_updsbind" :: "['a, 'a] => updbind"             ("(2_ [:=]/ _)")
 
+syntax_consts
+  "_updsbind" == fun_upds
+
 translations
   "f(xs[:=]y)" == "CONST fun_upds f xs y"
 
--- a/src/HOL/Set.thy	Mon Aug 26 18:26:00 2024 +0200
+++ b/src/HOL/Set.thy	Mon Aug 26 18:26:06 2024 +0200
@@ -40,6 +40,8 @@
 
 syntax
   "_Coll" :: "pttrn \<Rightarrow> bool \<Rightarrow> 'a set"    ("(1{_./ _})")
+syntax_consts
+  "_Coll" \<rightleftharpoons> Collect
 translations
   "{x. P}" \<rightleftharpoons> "CONST Collect (\<lambda>x. P)"
 
@@ -47,6 +49,8 @@
   "_Collect" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> bool \<Rightarrow> 'a set"  ("(1{(_/: _)./ _})")
 syntax
   "_Collect" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> bool \<Rightarrow> 'a set"  ("(1{(_/ \<in> _)./ _})")
+syntax_consts
+  "_Collect" \<rightleftharpoons> Collect
 translations
   "{p:A. P}" \<rightharpoonup> "CONST Collect (\<lambda>p. p \<in> A \<and> P)"
 
@@ -141,6 +145,8 @@
 
 syntax
   "_Finset" :: "args \<Rightarrow> 'a set"    ("{(_)}")
+syntax_consts
+  "_Finset" \<rightleftharpoons> insert
 translations
   "{x, xs}" \<rightleftharpoons> "CONST insert x {xs}"
   "{x}" \<rightleftharpoons> "CONST insert x {}"
@@ -203,6 +209,12 @@
   "_Bex1"       :: "pttrn \<Rightarrow> 'a set \<Rightarrow> bool \<Rightarrow> bool"      ("(3\<exists>!(_/\<in>_)./ _)" [0, 0, 10] 10)
   "_Bleast"     :: "id \<Rightarrow> 'a set \<Rightarrow> bool \<Rightarrow> 'a"           ("(3LEAST(_/\<in>_)./ _)" [0, 0, 10] 10)
 
+syntax_consts
+  "_Ball" \<rightleftharpoons> Ball and
+  "_Bex" \<rightleftharpoons> Bex and
+  "_Bex1" \<rightleftharpoons> Ex1 and
+  "_Bleast" \<rightleftharpoons> Least
+
 translations
   "\<forall>x\<in>A. P" \<rightleftharpoons> "CONST Ball A (\<lambda>x. P)"
   "\<exists>x\<in>A. P" \<rightleftharpoons> "CONST Bex A (\<lambda>x. P)"
@@ -223,12 +235,17 @@
   "_setleEx"    :: "[idt, 'a, bool] \<Rightarrow> bool"   ("(3\<exists>_\<subseteq>_./ _)" [0, 0, 10] 10)
   "_setleEx1"   :: "[idt, 'a, bool] \<Rightarrow> bool"   ("(3\<exists>!_\<subseteq>_./ _)" [0, 0, 10] 10)
 
+syntax_consts
+  "_setlessAll" "_setleAll" \<rightleftharpoons> All and
+  "_setlessEx" "_setleEx" \<rightleftharpoons> Ex and
+  "_setleEx1" \<rightleftharpoons> Ex1
+
 translations
- "\<forall>A\<subset>B. P" \<rightharpoonup> "\<forall>A. A \<subset> B \<longrightarrow> P"
- "\<exists>A\<subset>B. P" \<rightharpoonup> "\<exists>A. A \<subset> B \<and> P"
- "\<forall>A\<subseteq>B. P" \<rightharpoonup> "\<forall>A. A \<subseteq> B \<longrightarrow> P"
- "\<exists>A\<subseteq>B. P" \<rightharpoonup> "\<exists>A. A \<subseteq> B \<and> P"
- "\<exists>!A\<subseteq>B. P" \<rightharpoonup> "\<exists>!A. A \<subseteq> B \<and> P"
+  "\<forall>A\<subset>B. P" \<rightharpoonup> "\<forall>A. A \<subset> B \<longrightarrow> P"
+  "\<exists>A\<subset>B. P" \<rightharpoonup> "\<exists>A. A \<subset> B \<and> P"
+  "\<forall>A\<subseteq>B. P" \<rightharpoonup> "\<forall>A. A \<subseteq> B \<longrightarrow> P"
+  "\<exists>A\<subseteq>B. P" \<rightharpoonup> "\<exists>A. A \<subseteq> B \<and> P"
+  "\<exists>!A\<subseteq>B. P" \<rightharpoonup> "\<exists>!A. A \<subseteq> B \<and> P"
 
 print_translation \<open>
   let
@@ -272,6 +289,8 @@
 
 syntax
   "_Setcompr" :: "'a \<Rightarrow> idts \<Rightarrow> bool \<Rightarrow> 'a set"    ("(1{_ |/_./ _})")
+syntax_consts
+  "_Setcompr" \<rightleftharpoons> Collect
 
 parse_translation \<open>
   let
--- a/src/HOL/Set_Interval.thy	Mon Aug 26 18:26:00 2024 +0200
+++ b/src/HOL/Set_Interval.thy	Mon Aug 26 18:26:06 2024 +0200
@@ -84,6 +84,10 @@
   "_INTER_le"   :: "'a => 'a => 'b set => 'b set"       ("(3\<Inter>_\<le>_./ _)" [0, 0, 10] 10)
   "_INTER_less" :: "'a => 'a => 'b set => 'b set"       ("(3\<Inter>_<_./ _)" [0, 0, 10] 10)
 
+syntax_consts
+  "_UNION_le" "_UNION_less" \<rightleftharpoons> Union and
+  "_INTER_le" "_INTER_less" \<rightleftharpoons> Inter
+
 translations
   "\<Union>i\<le>n. A" \<rightleftharpoons> "\<Union>i\<in>{..n}. A"
   "\<Union>i<n. A" \<rightleftharpoons> "\<Union>i\<in>{..<n}. A"
@@ -2005,6 +2009,9 @@
   "_upt_sum" :: "idt \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b"  ("(3\<Sum>_<_./ _)" [0,0,10] 10)
   "_upto_sum" :: "idt \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b"  ("(3\<Sum>_\<le>_./ _)" [0,0,10] 10)
 
+syntax_consts
+  "_from_to_sum" "_from_upto_sum" "_upt_sum" "_upto_sum" == sum
+
 translations
   "\<Sum>x=a..b. t" == "CONST sum (\<lambda>x. t) {a..b}"
   "\<Sum>x=a..<b. t" == "CONST sum (\<lambda>x. t) {a..<b}"
@@ -2695,6 +2702,9 @@
   "_upt_prod" :: "idt \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b"  ("(3\<Prod>_<_./ _)" [0,0,10] 10)
   "_upto_prod" :: "idt \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b"  ("(3\<Prod>_\<le>_./ _)" [0,0,10] 10)
 
+syntax_consts
+  "_from_to_prod" "_from_upto_prod" "_upt_prod" "_upto_prod" \<rightleftharpoons> prod
+
 translations
   "\<Prod>x=a..b. t" \<rightleftharpoons> "CONST prod (\<lambda>x. t) {a..b}"
   "\<Prod>x=a..<b. t" \<rightleftharpoons> "CONST prod (\<lambda>x. t) {a..<b}"
--- a/src/HOL/String.thy	Mon Aug 26 18:26:00 2024 +0200
+++ b/src/HOL/String.thy	Mon Aug 26 18:26:06 2024 +0200
@@ -218,6 +218,8 @@
 syntax
   "_Char" :: "str_position \<Rightarrow> char"    ("CHR _")
   "_Char_ord" :: "num_const \<Rightarrow> char"   ("CHR _")
+syntax_consts
+  "_Char" "_Char_ord" \<rightleftharpoons> Char
 
 type_synonym string = "char list"
 
@@ -522,6 +524,8 @@
 syntax
   "_Literal" :: "str_position \<Rightarrow> String.literal"   ("STR _")
   "_Ascii" :: "num_const \<Rightarrow> String.literal"        ("STR _")
+syntax_consts
+  "_Literal" "_Ascii" \<rightleftharpoons> String.Literal
 
 ML_file \<open>Tools/literal.ML\<close>
 
--- a/src/HOL/Typerep.thy	Mon Aug 26 18:26:00 2024 +0200
+++ b/src/HOL/Typerep.thy	Mon Aug 26 18:26:06 2024 +0200
@@ -19,6 +19,8 @@
 
 syntax
   "_TYPEREP" :: "type => logic"  ("(1TYPEREP/(1'(_')))")
+syntax_consts
+  "_TYPEREP" \<rightleftharpoons> typerep
 
 parse_translation \<open>
   let
--- a/src/HOL/UNITY/PPROD.thy	Mon Aug 26 18:26:00 2024 +0200
+++ b/src/HOL/UNITY/PPROD.thy	Mon Aug 26 18:26:06 2024 +0200
@@ -16,6 +16,8 @@
 
 syntax
   "_PLam" :: "[pttrn, nat set, 'b set] => (nat => 'b) set"  ("(3plam _:_./ _)" 10)
+syntax_consts
+  "_PLam" == PLam
 translations
   "plam x : A. B" == "CONST PLam A (%x. B)"
 
--- a/src/HOL/UNITY/Union.thy	Mon Aug 26 18:26:00 2024 +0200
+++ b/src/HOL/UNITY/Union.thy	Mon Aug 26 18:26:06 2024 +0200
@@ -41,6 +41,8 @@
 syntax
   "_JOIN1" :: "[pttrns, 'b set] => 'b set"              ("(3\<Squnion>_./ _)" 10)
   "_JOIN"  :: "[pttrn, 'a set, 'b set] => 'b set"       ("(3\<Squnion>_\<in>_./ _)" 10)
+syntax_consts
+  "_JOIN1" "_JOIN" == JOIN
 translations
   "\<Squnion>x \<in> A. B" == "CONST JOIN A (\<lambda>x. B)"
   "\<Squnion>x y. B" == "\<Squnion>x. \<Squnion>y. B"
--- a/src/HOL/ex/Code_Lazy_Demo.thy	Mon Aug 26 18:26:00 2024 +0200
+++ b/src/HOL/ex/Code_Lazy_Demo.thy	Mon Aug 26 18:26:06 2024 +0200
@@ -40,6 +40,7 @@
   | LCons (lhd: 'a) (ltl: "'a llist") (infixr "###" 65)
 
 syntax "_llist" :: "args => 'a list"    ("\<^bold>\<lbrakk>(_)\<^bold>\<rbrakk>")
+syntax_consts "_llist" == LCons
 translations
   "\<^bold>\<lbrakk>x, xs\<^bold>\<rbrakk>" == "x###\<^bold>\<lbrakk>xs\<^bold>\<rbrakk>"
   "\<^bold>\<lbrakk>x\<^bold>\<rbrakk>" == "x###\<^bold>\<lbrakk>\<^bold>\<rbrakk>"
--- a/src/Pure/Examples/Higher_Order_Logic.thy	Mon Aug 26 18:26:00 2024 +0200
+++ b/src/Pure/Examples/Higher_Order_Logic.thy	Mon Aug 26 18:26:06 2024 +0200
@@ -402,6 +402,7 @@
   where someI: "P x \<Longrightarrow> P (Eps P)"
 
 syntax "_Eps" :: "pttrn \<Rightarrow> o \<Rightarrow> 'a"  ("(3SOME _./ _)" [0, 10] 10)
+syntax_consts "_Eps" \<rightleftharpoons> Eps
 translations "SOME x. P" \<rightleftharpoons> "CONST Eps (\<lambda>x. P)"
 
 text \<open>
--- a/src/Pure/Isar/isar_cmd.ML	Mon Aug 26 18:26:00 2024 +0200
+++ b/src/Pure/Isar/isar_cmd.ML	Mon Aug 26 18:26:06 2024 +0200
@@ -1,5 +1,5 @@
 (*  Title:      Pure/Isar/isar_cmd.ML
-    Author:     Markus Wenzel, TU Muenchen
+    Author:     Makarius
 
 Miscellaneous Isar commands.
 *)
@@ -15,6 +15,10 @@
   val print_ast_translation: Input.source -> theory -> theory
   val translations: (xstring * string) Syntax.trrule list -> theory -> theory
   val no_translations: (xstring * string) Syntax.trrule list -> theory -> theory
+  val syntax_consts: ((string * Position.T) list * (xstring * Position.T) list) list ->
+    theory -> theory
+  val syntax_types: ((string * Position.T) list * (xstring * Position.T) list) list ->
+    theory -> theory
   val oracle: bstring * Position.range -> Input.source -> theory -> theory
   val declaration: {syntax: bool, pervasive: bool} -> Input.source -> local_theory -> local_theory
   val qed: Method.text_range option -> Toplevel.transition -> Toplevel.transition
@@ -108,6 +112,44 @@
 fun no_translations args thy = Sign.del_trrules (read_trrules thy args) thy;
 
 
+(* syntax consts/types (after translation) *)
+
+local
+
+fun syntax_deps_cmd f args thy =
+  let
+    val ctxt = Proof_Context.init_global thy;
+
+    val check_lhs = Proof_Context.check_syntax_const ctxt;
+    fun check_rhs (b: xstring, pos: Position.T) =
+      let
+        val (c: string, reports) = f ctxt (b, pos);
+        val _ = Context_Position.reports ctxt reports;
+      in c end;
+
+    fun check (raw_lhs, raw_rhs) =
+      let
+        val lhs = map check_lhs raw_lhs;
+        val rhs = map check_rhs raw_rhs;
+      in map (fn l => (l, rhs)) lhs end;
+  in Sign.syntax_deps (maps check args) thy end;
+
+fun check_const ctxt (s, pos) =
+  Proof_Context.check_const {proper = true, strict = false} ctxt (s, [pos])
+  |>> (Term.dest_Const_name #> Lexicon.mark_const);
+
+fun check_type_name ctxt arg =
+  Proof_Context.check_type_name {proper = true, strict = false} ctxt arg
+  |>> (Term.dest_Type_name #> Lexicon.mark_type);
+
+in
+
+val syntax_consts = syntax_deps_cmd check_const;
+val syntax_types = syntax_deps_cmd check_type_name;
+
+end;
+
+
 (* oracles *)
 
 fun oracle (name, range) source =
--- a/src/Pure/Isar/proof_context.ML	Mon Aug 26 18:26:00 2024 +0200
+++ b/src/Pure/Isar/proof_context.ML	Mon Aug 26 18:26:06 2024 +0200
@@ -155,6 +155,7 @@
   val apply_case: Rule_Cases.T -> Proof.context -> (string * term list) list * Proof.context
   val check_case: Proof.context -> bool ->
     string * Position.T -> binding option list -> Rule_Cases.T
+  val is_syntax_const: Proof.context -> string -> bool
   val check_syntax_const: Proof.context -> string * Position.T -> string
   val syntax: bool -> Syntax.mode -> (string * typ * mixfix) list ->
     Proof.context -> Proof.context
@@ -570,8 +571,7 @@
           let
             val reports = ps
               |> filter (Context_Position.is_reported ctxt)
-              |> map (fn pos =>
-                (pos, Markup.name x (if Name.is_skolem x then Markup.skolem else Markup.free)));
+              |> map (fn pos => (pos, Markup.name x (Variable.markup ctxt x)));
           in (Free (x, infer_type ctxt (x, dummyT)), reports) end
       | (_, SOME d) =>
           let
@@ -1126,8 +1126,10 @@
 
 (* syntax *)
 
+val is_syntax_const = Syntax.is_const o syntax_of;
+
 fun check_syntax_const ctxt (c, pos) =
-  if is_some (Syntax.lookup_const (syntax_of ctxt) c) then c
+  if is_syntax_const ctxt c then c
   else error ("Unknown syntax const: " ^ quote c ^ Position.here pos);
 
 fun syntax add mode args ctxt =
--- a/src/Pure/Pure.thy	Mon Aug 26 18:26:00 2024 +0200
+++ b/src/Pure/Pure.thy	Mon Aug 26 18:26:06 2024 +0200
@@ -15,9 +15,10 @@
   and "text" "txt" :: document_body
   and "text_raw" :: document_raw
   and "default_sort" :: thy_decl
-  and "typedecl" "nonterminal" "judgment" "consts" "syntax" "no_syntax" "translations"
-    "no_translations" "type_notation" "no_type_notation" "notation" "no_notation" "alias"
-    "type_alias" "declare" "hide_class" "hide_type" "hide_const" "hide_fact" :: thy_decl
+  and "typedecl" "nonterminal" "judgment" "consts" "syntax" "no_syntax" "syntax_consts"
+    "syntax_types" "translations" "no_translations" "type_notation" "no_type_notation" "notation"
+    "no_notation" "alias" "type_alias" "declare" "hide_class" "hide_type" "hide_const"
+    "hide_fact" :: thy_decl
   and "type_synonym" "definition" "abbreviation" "lemmas" :: thy_defn
   and "axiomatization" :: thy_stmt
   and "external_file" "bibtex_file" "ROOTS_file" :: thy_load
@@ -401,6 +402,19 @@
     (Parse.syntax_mode -- Scan.repeat1 Parse.const_decl
       >> uncurry (Local_Theory.syntax_cmd false));
 
+val syntax_consts =
+  Parse.and_list1
+    ((Scan.repeat1 Parse.name_position --| (\<^keyword>\<open>\<rightleftharpoons>\<close> || \<^keyword>\<open>==\<close>)) --
+      Parse.!!! (Scan.repeat1 Parse.name_position));
+
+val _ =
+  Outer_Syntax.command \<^command_keyword>\<open>syntax_consts\<close> "declare syntax const dependencies"
+    (syntax_consts >> (Toplevel.theory o Isar_Cmd.syntax_consts));
+
+val _ =
+  Outer_Syntax.command \<^command_keyword>\<open>syntax_types\<close> "declare syntax const dependencies (type names)"
+    (syntax_consts >> (Toplevel.theory o Isar_Cmd.syntax_types));
+
 val trans_pat =
   Scan.optional
     (\<^keyword>\<open>(\<close> |-- Parse.!!! (Parse.inner_syntax Parse.name --| \<^keyword>\<open>)\<close>)) "logic"
--- a/src/Pure/Syntax/mixfix.ML	Mon Aug 26 18:26:00 2024 +0200
+++ b/src/Pure/Syntax/mixfix.ML	Mon Aug 26 18:26:06 2024 +0200
@@ -182,7 +182,7 @@
 
     val mfix = map mfix_of type_decls;
     val _ = map2 check_args type_decls mfix;
-    val consts = map (fn (t, _, _) => (t, "")) type_decls;
+    val consts = map (fn (t, _, _) => (t, [])) type_decls;
   in Syntax_Ext.syn_ext [] (map_filter I mfix) consts ([], [], [], []) ([], []) end;
 
 
@@ -235,7 +235,9 @@
       |> map (Syntax_Ext.stamp_trfun binder_stamp o
           apsnd Syntax_Trans.non_typed_tr' o Syntax_Trans.mk_binder_tr' o swap);
 
-    val consts = binders @ map (fn (c, _, _) => (c, "")) const_decls;
+    val consts =
+      map (fn (c, b) => (c, [b])) binders @
+      map (fn (c, _, _) => (c, [])) const_decls;
   in Syntax_Ext.syn_ext logical_types mfix consts ([], binder_trs, binder_trs', []) ([], []) end;
 
 end;
--- a/src/Pure/Syntax/printer.ML	Mon Aug 26 18:26:00 2024 +0200
+++ b/src/Pure/Syntax/printer.ML	Mon Aug 26 18:26:06 2024 +0200
@@ -33,12 +33,13 @@
   val pretty_term_ast: bool -> Proof.context -> prtabs ->
     (string -> Proof.context -> Ast.ast list -> Ast.ast) ->
     (string -> Ast.ast list -> Pretty.T option) ->
-    (string -> Markup.T list * string) ->
+    ((string -> Markup.T list) * (string -> string)) ->
     Ast.ast -> Pretty.T list
   val pretty_typ_ast: Proof.context -> prtabs ->
     (string -> Proof.context -> Ast.ast list -> Ast.ast) ->
     (string -> Ast.ast list -> Pretty.T option) ->
-    (string -> Markup.T list * string) -> Ast.ast -> Pretty.T list
+    ((string -> Markup.T list) * (string -> string)) ->
+    Ast.ast -> Pretty.T list
 end;
 
 structure Printer: PRINTER =
@@ -202,10 +203,14 @@
   | is_chain [Arg _] = true
   | is_chain _  = false;
 
+val pretty_type_mode = Config.declare_bool ("Syntax.pretty_type_mode", \<^here>) (K false);
+val pretty_curried = Config.declare_bool ("Syntax.pretty_curried", \<^here>) (K false);
 val pretty_priority = Config.declare_int ("Syntax.pretty_priority", \<^here>) (K 0);
 
-fun pretty type_mode curried ctxt tabs trf markup_trans markup_extern ast0 =
+fun pretty ctxt tabs trf markup_trans markup_extern ast0 =
   let
+    val type_mode = Config.get ctxt pretty_type_mode;
+    val curried = Config.get ctxt pretty_curried;
     val show_brackets = Config.get ctxt show_brackets;
 
     (*default applications: prefix / postfix*)
@@ -224,8 +229,10 @@
           in
             if type_mode then (astT (t, p) @ Ts, args')
             else
-              (pretty true curried (Config.put pretty_priority p ctxt)
-                tabs trf markup_trans markup_extern t @ Ts, args')
+              let val ctxt' = ctxt
+                |> Config.put pretty_type_mode true
+                |> Config.put pretty_priority p
+              in (pretty ctxt' tabs trf markup_trans markup_extern t @ Ts, args') end
           end
       | synT m (String (do_mark, s) :: symbs, args) =
           let
@@ -254,7 +261,7 @@
             [Block (Syntax_Ext.block_indent 1, String (false, "(") :: pr @ [String (false, ")")])]
            else pr, args))
 
-    and atomT a = Pretty.marks_str (markup_extern a)
+    and atomT a = Pretty.marks_str (#1 markup_extern a, #2 markup_extern a)
 
     and prefixT (_, a, [], _) = [atomT a]
       | prefixT (c, _, args, p) = astT (appT (c, args), p)
@@ -271,7 +278,7 @@
         fun prnt ([], []) = prefixT tup
           | prnt ([], tb :: tbs) = prnt (Symtab.lookup_list tb a, tbs)
           | prnt ((pr, n, p') :: prnps, tbs) =
-              if nargs = n then parT (#1 (markup_extern a)) (pr, args, p, p')
+              if nargs = n then parT (#1 markup_extern a) (pr, args, p, p')
               else if nargs > n andalso not type_mode then
                 astT (appT (splitT n ([c], args)), p)
               else prnt (prnps, tbs);
@@ -291,14 +298,20 @@
 
 (* pretty_term_ast *)
 
-fun pretty_term_ast curried ctxt prtabs trf markup_trans extern ast =
-  pretty false curried ctxt (mode_tabs prtabs (print_mode_value ())) trf markup_trans extern ast;
+fun pretty_term_ast curried ctxt prtabs trf markup_trans markup_extern ast =
+  let val ctxt' = ctxt
+    |> Config.put pretty_type_mode false
+    |> Config.put pretty_curried curried
+  in pretty ctxt' (mode_tabs prtabs (print_mode_value ())) trf markup_trans markup_extern ast end;
 
 
 (* pretty_typ_ast *)
 
-fun pretty_typ_ast ctxt prtabs trf markup_trans extern ast =
-  pretty true false ctxt (mode_tabs prtabs (print_mode_value ())) trf markup_trans extern ast;
+fun pretty_typ_ast ctxt prtabs trf markup_trans markup_extern ast =
+  let val ctxt' = ctxt
+    |> Config.put pretty_type_mode true
+    |> Config.put pretty_curried false
+  in pretty ctxt' (mode_tabs prtabs (print_mode_value ())) trf markup_trans markup_extern ast end;
 
 end;
 
--- a/src/Pure/Syntax/syntax.ML	Mon Aug 26 18:26:00 2024 +0200
+++ b/src/Pure/Syntax/syntax.ML	Mon Aug 26 18:26:06 2024 +0200
@@ -77,7 +77,8 @@
   val eq_syntax: syntax * syntax -> bool
   datatype approx = Prefix of string | Infix of {assoc: Printer.assoc, delim: string, pri: int}
   val get_approx: syntax -> string -> approx option
-  val lookup_const: syntax -> string -> string option
+  val get_consts: syntax -> string -> string list
+  val is_const: syntax -> string -> bool
   val is_keyword: syntax -> string -> bool
   val tokenize: syntax -> bool -> Symbol_Pos.T list -> Lexicon.token list
   val parse: syntax -> string -> Lexicon.token list -> Parser.parsetree list
@@ -111,6 +112,7 @@
   val update_type_gram: bool -> mode -> (string * typ * mixfix) list -> syntax -> syntax
   val update_const_gram: bool -> string list ->
     mode -> (string * typ * mixfix) list -> syntax -> syntax
+  val update_consts: (string * string list) list -> syntax -> syntax
   val update_trrules: Ast.ast trrule list -> syntax -> syntax
   val remove_trrules: Ast.ast trrule list -> syntax -> syntax
   val const: string -> term
@@ -411,7 +413,7 @@
     input: Syntax_Ext.xprod list,
     lexicon: Scan.lexicon,
     gram: Parser.gram Synchronized.cache,
-    consts: string Symtab.table,
+    consts: unit Graph.T,
     prmodes: string list,
     parse_ast_trtab: ((Proof.context -> Ast.ast list -> Ast.ast) * stamp) Symtab.table,
     parse_ruletab: ruletab,
@@ -440,7 +442,9 @@
       | NONE => Printer.get_binder prtabs c)
       |> Option.map Prefix);
 
-fun lookup_const (Syntax ({consts, ...}, _)) = Symtab.lookup consts;
+fun is_const (Syntax ({consts, ...}, _)) c =
+  Graph.defined consts c andalso not (Lexicon.is_marked c);
+
 fun is_keyword (Syntax ({lexicon, ...}, _)) = Scan.is_literal lexicon o Symbol.explode;
 fun tokenize (Syntax ({lexicon, ...}, _)) = Lexicon.tokenize lexicon;
 fun parse (Syntax ({gram, ...}, _)) = Parser.parse (cache_eval gram);
@@ -466,13 +470,35 @@
   Synchronized.cache (fn () => Parser.make_gram new_xprods [] NONE);
 
 
+(* syntax consts *)
+
+fun err_cyclic_consts css =
+  error (cat_lines (map (fn cs =>
+    "Cycle in syntax consts: " ^ (space_implode " \<leadsto> " (map quote cs))) css));
+
+fun get_consts (Syntax ({consts, ...}, _)) c =
+  if Graph.defined consts c then
+    Graph.all_preds consts [c]
+    |> filter (Graph.Keys.is_empty o Graph.imm_preds consts)
+  else [c];
+
+fun update_consts (c, bs) consts =
+  if c = "" orelse (null bs andalso (Lexicon.is_marked c orelse Graph.defined consts c))
+  then consts
+  else
+    consts
+    |> fold (fn a => Graph.default_node (a, ())) (c :: bs)
+    |> Graph.add_deps_acyclic (c, bs)
+    handle Graph.CYCLES css => err_cyclic_consts css;
+
+
 (* empty_syntax *)
 
 val empty_syntax = Syntax
   ({input = [],
     lexicon = Scan.empty_lexicon,
     gram = Synchronized.cache (fn () => Parser.empty_gram),
-    consts = Symtab.empty,
+    consts = Graph.empty,
     prmodes = [],
     parse_ast_trtab = Symtab.empty,
     parse_ruletab = Symtab.empty,
@@ -485,11 +511,6 @@
 
 (* update_syntax *)
 
-fun update_const (c, b) tab =
-  if c = "" orelse (b = "" andalso (Lexicon.is_marked c orelse Symtab.defined tab c))
-  then tab
-  else Symtab.update (c, b) tab;
-
 fun update_syntax (mode, inout) syn_ext (Syntax (tabs, _)) =
   let
     val {input, lexicon, gram, consts = consts1, prmodes, parse_ast_trtab, parse_ruletab,
@@ -505,7 +526,7 @@
     ({input = new_xprods @ input,
       lexicon = fold Scan.extend_lexicon (Syntax_Ext.delims_of new_xprods) lexicon,
       gram = if null new_xprods then gram else extend_gram new_xprods input gram,
-      consts = fold update_const consts2 consts1,
+      consts = fold update_consts consts2 consts1,
       prmodes = insert (op =) mode prmodes,
       parse_ast_trtab =
         update_trtab "parse ast translation" (if_inout parse_ast_translation) parse_ast_trtab,
@@ -577,7 +598,9 @@
     ({input = input',
       lexicon = Scan.merge_lexicons (lexicon1, lexicon2),
       gram = gram',
-      consts = Symtab.merge (K true) (consts1, consts2),
+      consts =
+        Graph.merge_acyclic (op =) (consts1, consts2)
+          handle Graph.CYCLES css => err_cyclic_consts css,
       prmodes = Library.merge (op =) (prmodes1, prmodes2),
       parse_ast_trtab =
         merge_trtabs "parse ast translation" parse_ast_trtab1 parse_ast_trtab2,
@@ -620,7 +643,7 @@
     val {consts, parse_ast_trtab, parse_ruletab, parse_trtab, print_trtab,
       print_ruletab, print_ast_trtab, ...} = tabs;
   in
-   [pretty_tab "consts:" consts,
+   [pretty_strs_qs "consts:" (sort_strings (Graph.keys consts)),
     pretty_tab "parse_ast_translation:" parse_ast_trtab,
     pretty_ruletab "parse_rules:" parse_ruletab,
     pretty_tab "parse_translation:" parse_trtab,
@@ -692,6 +715,8 @@
 fun update_const_gram add logical_types prmode decls =
   (if add then update_syntax else remove_syntax) prmode (Mixfix.syn_ext_consts logical_types decls);
 
+val update_consts = update_syntax mode_default o Syntax_Ext.syn_ext_consts;
+
 val update_trrules = update_syntax mode_default o Syntax_Ext.syn_ext_rules o check_rules;
 val remove_trrules = remove_syntax mode_default o Syntax_Ext.syn_ext_rules o check_rules;
 
--- a/src/Pure/Syntax/syntax_ext.ML	Mon Aug 26 18:26:00 2024 +0200
+++ b/src/Pure/Syntax/syntax_ext.ML	Mon Aug 26 18:26:06 2024 +0200
@@ -23,7 +23,7 @@
   datatype syn_ext =
     Syn_Ext of {
       xprods: xprod list,
-      consts: (string * string) list,
+      consts: (string * string list) list,
       parse_ast_translation: (string * ((Proof.context -> Ast.ast list -> Ast.ast) * stamp)) list,
       parse_rules: (Ast.ast * Ast.ast) list,
       parse_translation: (string * ((Proof.context -> term list -> term) * stamp)) list,
@@ -34,11 +34,12 @@
   val mixfix_args: Input.source -> int
   val escape: string -> string
   val syn_ext: string list -> mfix list ->
-    (string * string) list -> (string * ((Proof.context -> Ast.ast list -> Ast.ast) * stamp)) list *
+    (string * string list) list -> (string * ((Proof.context -> Ast.ast list -> Ast.ast) * stamp)) list *
     (string * ((Proof.context -> term list -> term) * stamp)) list *
     (string * ((Proof.context -> typ -> term list -> term) * stamp)) list *
     (string * ((Proof.context -> Ast.ast list -> Ast.ast) * stamp)) list ->
     (Ast.ast * Ast.ast) list * (Ast.ast * Ast.ast) list -> syn_ext
+  val syn_ext_consts: (string * string list) list -> syn_ext
   val syn_ext_rules: (Ast.ast * Ast.ast) list * (Ast.ast * Ast.ast) list -> syn_ext
   val syn_ext_trfuns:
     (string * ((Proof.context -> Ast.ast list -> Ast.ast) * stamp)) list *
@@ -345,7 +346,7 @@
             Ast.mk_appl (Ast.Constant indexed_const)
               (xs1 @ [Ast.mk_appl (Ast.Constant "_index") [i]] @ xs2);
           val rhs = Ast.mk_appl (Ast.Constant const) (i :: xs);
-        in (indexed_const, rangeT, SOME (indexed_const, const), SOME (lhs, rhs)) end;
+        in (indexed_const, rangeT, SOME (indexed_const, [const]), SOME (lhs, rhs)) end;
 
     val (symbs1, lhs) = add_args symbs0 typ' pris;
 
@@ -389,7 +390,7 @@
 datatype syn_ext =
   Syn_Ext of {
     xprods: xprod list,
-    consts: (string * string) list,
+    consts: (string * string list) list,
     parse_ast_translation: (string * ((Proof.context -> Ast.ast list -> Ast.ast) * stamp)) list,
     parse_rules: (Ast.ast * Ast.ast) list,
     parse_translation: (string * ((Proof.context -> term list -> term) * stamp)) list,
@@ -409,7 +410,7 @@
     val xprods = map #1 xprod_results;
     val consts' = map_filter #2 xprod_results;
     val parse_rules' = rev (map_filter #3 xprod_results);
-    val mfix_consts = map (fn Mfix x => (#3 x, "")) mfixes @ map (fn XProd x => (#3 x, "")) xprods;
+    val mfix_consts = map (fn Mfix x => (#3 x, [])) mfixes @ map (fn XProd x => (#3 x, [])) xprods;
   in
     Syn_Ext {
       xprods = xprods,
@@ -423,6 +424,7 @@
   end;
 
 
+fun syn_ext_consts consts = syn_ext [] [] consts ([], [], [], []) ([], []);
 fun syn_ext_rules rules = syn_ext [] [] [] ([], [], [], []) rules;
 fun syn_ext_trfuns trfuns = syn_ext [] [] [] trfuns ([], []);
 
--- a/src/Pure/Syntax/syntax_phases.ML	Mon Aug 26 18:26:00 2024 +0200
+++ b/src/Pure/Syntax/syntax_phases.ML	Mon Aug 26 18:26:06 2024 +0200
@@ -53,10 +53,13 @@
   [Name_Space.markup (Consts.space_of (Proof_Context.consts_of ctxt)) c];
 
 fun markup_free ctxt x =
-  Variable.markup ctxt x ::
-  (if Variable.is_body ctxt orelse Variable.is_fixed ctxt x
-   then [Variable.markup_fixed ctxt x]
-   else []);
+  let
+    val m1 =
+      if Variable.is_fixed ctxt x then Variable.markup_fixed ctxt x
+      else if not (Variable.is_body ctxt) orelse Syntax.is_pretty_global ctxt then Markup.fixed x
+      else Markup.intensify;
+    val m2 = Variable.markup ctxt x;
+  in [m1, m2] end;
 
 fun markup_var xi = [Markup.name (Term.string_of_vname xi) Markup.var];
 
@@ -64,10 +67,8 @@
   Markup.bound :: map (fn pos => Position.make_entity_markup def id Markup.boundN (name, pos)) ps;
 
 fun markup_entity ctxt c =
-  (case Syntax.lookup_const (Proof_Context.syntax_of ctxt) c of
-    SOME "" => []
-  | SOME b => markup_entity ctxt b
-  | NONE => c |> Lexicon.unmark
+  Syntax.get_consts (Proof_Context.syntax_of ctxt) c
+  |> maps (Lexicon.unmark
      {case_class = markup_class ctxt,
       case_type = markup_type ctxt,
       case_const = markup_const ctxt,
@@ -468,7 +469,7 @@
     fun decode_appl ps asts = Ast.Appl (map (decode ps) asts)
     and decode ps (Ast.Constant c) = decode_const ps c
       | decode ps (Ast.Variable x) =
-          if is_some (Syntax.lookup_const syn x) orelse Long_Name.is_qualified x
+          if Syntax.is_const syn x orelse Long_Name.is_qualified x
           then decode_const ps x
           else decode_var ps x
       | decode ps (Ast.Appl (asts as (Ast.Constant c :: ast :: Ast.Variable x :: args))) =
@@ -619,7 +620,7 @@
           in (t1' $ t2', seen'') end;
   in #1 (prune (tm, [])) end;
 
-fun mark_atoms is_syntax_const ctxt tm =
+fun mark_atoms ctxt tm =
   let
     val {structs, fixes} = Syntax_Trans.get_idents ctxt;
     val show_structs = Config.get ctxt show_structs;
@@ -630,7 +631,7 @@
       | mark (t $ u) = mark t $ mark u
       | mark (Abs (x, T, t)) = Abs (x, T, mark t)
       | mark (t as Const (c, T)) =
-          if is_syntax_const c then t
+          if Proof_Context.is_syntax_const ctxt c then t
           else Const (Lexicon.mark_const c, T)
       | mark (t as Free (x, T)) =
           let val i = find_index (fn s => s = x) structs + 1 in
@@ -648,7 +649,7 @@
 
 in
 
-fun term_to_ast is_syntax_const ctxt trf tm =
+fun term_to_ast ctxt trf tm =
   let
     val show_types = Config.get ctxt show_types orelse Config.get ctxt show_sorts;
     val show_markup = Config.get ctxt show_markup;
@@ -691,7 +692,7 @@
     |> mark_aprop
     |> show_types ? prune_types
     |> Variable.revert_bounds ctxt
-    |> mark_atoms is_syntax_const ctxt
+    |> mark_atoms ctxt
     |> ast_of
   end;
 
@@ -703,23 +704,32 @@
 
 local
 
-fun free_or_skolem ctxt x =
-  let
-    val m =
-      if Variable.is_fixed ctxt x orelse Syntax.is_pretty_global ctxt
-      then Markup.fixed x else Markup.intensify;
-  in
-    if Name.is_skolem x
-    then ([m, Markup.skolem], Variable.revert_fixed ctxt x)
-    else ([m, Markup.free], x)
-  end;
+fun extern_fixed ctxt x =
+  if Name.is_skolem x then Variable.revert_fixed ctxt x else x;
+
+fun extern ctxt c =
+  (case Syntax.get_consts (Proof_Context.syntax_of ctxt) c of
+    [b] => b
+  | bs =>
+      (case filter Lexicon.is_marked bs of
+        [] => c
+      | [b] => b
+      | _ => error ("Multiple logical entities for " ^ quote c ^ ": " ^ commas_quote bs)))
+  |> Lexicon.unmark
+     {case_class = Proof_Context.extern_class ctxt,
+      case_type = Proof_Context.extern_type ctxt,
+      case_const = Proof_Context.extern_const ctxt,
+      case_fixed = extern_fixed ctxt,
+      case_default = I};
+
+fun free_or_skolem ctxt x = (markup_free ctxt x, extern_fixed ctxt x);
 
 fun var_or_skolem s =
   (case Lexicon.read_variable s of
     SOME (x, i) =>
       (case try Name.dest_skolem x of
-        NONE => (Markup.var, s)
-      | SOME x' => (Markup.skolem, Term.string_of_vname (x', i)))
+        SOME x' => (Markup.skolem, Term.string_of_vname (x', i))
+      | NONE => (Markup.var, s))
   | NONE => (Markup.var, s));
 
 val typing_elem = YXML.output_markup_elem Markup.typing;
@@ -734,17 +744,7 @@
     val syn = Proof_Context.syntax_of ctxt;
     val prtabs = Syntax.prtabs syn;
     val trf = Syntax.print_ast_translation syn;
-
-    fun markup_extern c =
-      (case Syntax.lookup_const syn c of
-        SOME "" => ([], c)
-      | SOME b => markup_extern b
-      | NONE => c |> Lexicon.unmark
-         {case_class = fn x => (markup_class ctxt x, Proof_Context.extern_class ctxt x),
-          case_type = fn x => (markup_type ctxt x, Proof_Context.extern_type ctxt x),
-          case_const = fn x => (markup_const ctxt x, Proof_Context.extern_const ctxt x),
-          case_fixed = fn x => free_or_skolem ctxt x,
-          case_default = fn x => ([], x)});
+    val markup_extern = (markup_entity ctxt, extern ctxt);
 
     fun token_trans "_tfree" x = SOME (Pretty.mark_str (Markup.tfree, x))
       | token_trans "_tvar" x = SOME (Pretty.mark_str (Markup.tvar, x))
@@ -801,12 +801,8 @@
 fun unparse_term ctxt =
   let
     val thy = Proof_Context.theory_of ctxt;
-    val syn = Proof_Context.syntax_of ctxt;
-  in
-    unparse_t (term_to_ast (is_some o Syntax.lookup_const syn))
-      (Printer.pretty_term_ast (not (Pure_Thy.old_appl_syntax thy)))
-      (Markup.language_term false) ctxt
-  end;
+    val curried = not (Pure_Thy.old_appl_syntax thy);
+  in unparse_t term_to_ast (Printer.pretty_term_ast curried) (Markup.language_term false) ctxt end;
 
 end;
 
--- a/src/Pure/pure_thy.ML	Mon Aug 26 18:26:00 2024 +0200
+++ b/src/Pure/pure_thy.ML	Mon Aug 26 18:26:06 2024 +0200
@@ -220,6 +220,8 @@
   #> Sign.parse_ast_translation Syntax_Trans.pure_parse_ast_translation
   #> Sign.parse_translation Syntax_Trans.pure_parse_translation
   #> Sign.print_ast_translation Syntax_Trans.pure_print_ast_translation
+  #> Sign.syntax_deps
+   [("_bracket", ["\<^type>fun"]), ("_bigimpl", ["\<^const>Pure.imp"])]
   #> Sign.add_consts
    [(qualify (Binding.make ("term", \<^here>)), typ "'a \<Rightarrow> prop", NoSyn),
     (qualify (Binding.make ("sort_constraint", \<^here>)), typ "'a itself \<Rightarrow> prop", NoSyn),
--- a/src/Pure/sign.ML	Mon Aug 26 18:26:00 2024 +0200
+++ b/src/Pure/sign.ML	Mon Aug 26 18:26:06 2024 +0200
@@ -78,6 +78,7 @@
   val add_type_abbrev: Proof.context -> binding * string list * typ -> theory -> theory
   val syntax: bool -> Syntax.mode -> (string * typ * mixfix) list -> theory -> theory
   val syntax_cmd: bool -> Syntax.mode -> (string * string * mixfix) list -> theory -> theory
+  val syntax_deps: (string * string list) list -> theory -> theory
   val type_notation: bool -> Syntax.mode -> (typ * mixfix) list -> theory -> theory
   val notation: bool -> Syntax.mode -> (term * mixfix) list -> theory -> theory
   val declare_const: Proof.context -> (binding * typ) * mixfix -> theory -> term * theory
@@ -392,6 +393,8 @@
 val syntax = gen_syntax (K I);
 val syntax_cmd = gen_syntax Syntax.read_typ;
 
+val syntax_deps = map_syn o Syntax.update_consts;
+
 fun type_notation add mode args =
   let
     fun type_syntax (Type (c, args), mx) =
--- a/src/ZF/Bin.thy	Mon Aug 26 18:26:00 2024 +0200
+++ b/src/ZF/Bin.thy	Mon Aug 26 18:26:06 2024 +0200
@@ -117,6 +117,9 @@
   "_Int" :: "num_token \<Rightarrow> i"  (\<open>#_\<close> 1000)
   "_Neg_Int" :: "num_token \<Rightarrow> i"  (\<open>#-_\<close> 1000)
 
+syntax_consts
+  "_Int0" "_Int1" "_Int2" "_Int" "_Neg_Int1" "_Neg_Int2" "_Neg_Int" \<rightleftharpoons> integ_of
+
 ML_file \<open>Tools/numeral_syntax.ML\<close>
 
 
--- a/src/ZF/Induct/Multiset.thy	Mon Aug 26 18:26:00 2024 +0200
+++ b/src/ZF/Induct/Multiset.thy	Mon Aug 26 18:26:06 2024 +0200
@@ -75,6 +75,7 @@
 
 syntax
   "_MColl" :: "[pttrn, i, o] \<Rightarrow> i" (\<open>(1{# _ \<in> _./ _#})\<close>)
+syntax_consts "_MColl" \<rightleftharpoons> MCollect
 translations
   "{#x \<in> M. P#}" == "CONST MCollect(M, \<lambda>x. P)"
 
--- a/src/ZF/List.thy	Mon Aug 26 18:26:00 2024 +0200
+++ b/src/ZF/List.thy	Mon Aug 26 18:26:06 2024 +0200
@@ -17,13 +17,12 @@
 syntax
  "_Nil" :: i  (\<open>[]\<close>)
  "_List" :: "is \<Rightarrow> i"  (\<open>[(_)]\<close>)
-
+syntax_consts "_List" \<rightleftharpoons> Cons
 translations
   "[x, xs]"     == "CONST Cons(x, [xs])"
   "[x]"         == "CONST Cons(x, [])"
   "[]"          == "CONST Nil"
 
-
 consts
   length :: "i\<Rightarrow>i"
   hd     :: "i\<Rightarrow>i"
--- a/src/ZF/OrdQuant.thy	Mon Aug 26 18:26:00 2024 +0200
+++ b/src/ZF/OrdQuant.thy	Mon Aug 26 18:26:06 2024 +0200
@@ -26,6 +26,10 @@
   "_oall"     :: "[idt, i, o] \<Rightarrow> o"        (\<open>(3\<forall>_<_./ _)\<close> 10)
   "_oex"      :: "[idt, i, o] \<Rightarrow> o"        (\<open>(3\<exists>_<_./ _)\<close> 10)
   "_OUNION"   :: "[idt, i, i] \<Rightarrow> i"        (\<open>(3\<Union>_<_./ _)\<close> 10)
+syntax_consts
+  "_oall" \<rightleftharpoons> oall and
+  "_oex" \<rightleftharpoons> oex and
+  "_OUNION" \<rightleftharpoons> OUnion
 translations
   "\<forall>x<a. P" \<rightleftharpoons> "CONST oall(a, \<lambda>x. P)"
   "\<exists>x<a. P" \<rightleftharpoons> "CONST oex(a, \<lambda>x. P)"
@@ -194,6 +198,9 @@
 syntax
   "_rall"     :: "[pttrn, i\<Rightarrow>o, o] \<Rightarrow> o"        (\<open>(3\<forall>_[_]./ _)\<close> 10)
   "_rex"      :: "[pttrn, i\<Rightarrow>o, o] \<Rightarrow> o"        (\<open>(3\<exists>_[_]./ _)\<close> 10)
+syntax_consts
+  "_rall" \<rightleftharpoons> rall and
+  "_rex" \<rightleftharpoons> rex
 translations
   "\<forall>x[M]. P" \<rightleftharpoons> "CONST rall(M, \<lambda>x. P)"
   "\<exists>x[M]. P" \<rightleftharpoons> "CONST rex(M, \<lambda>x. P)"
--- a/src/ZF/QPair.thy	Mon Aug 26 18:26:00 2024 +0200
+++ b/src/ZF/QPair.thy	Mon Aug 26 18:26:06 2024 +0200
@@ -46,6 +46,8 @@
 
 syntax
   "_QSUM"   :: "[idt, i, i] \<Rightarrow> i"               (\<open>(3QSUM _ \<in> _./ _)\<close> 10)
+syntax_consts
+  "_QSUM" \<rightleftharpoons> QSigma
 translations
   "QSUM x \<in> A. B" => "CONST QSigma(A, \<lambda>x. B)"
 
--- a/src/ZF/UNITY/Union.thy	Mon Aug 26 18:26:00 2024 +0200
+++ b/src/ZF/UNITY/Union.thy	Mon Aug 26 18:26:06 2024 +0200
@@ -43,7 +43,8 @@
 syntax
   "_JOIN1"  :: "[pttrns, i] \<Rightarrow> i"     (\<open>(3\<Squnion>_./ _)\<close> 10)
   "_JOIN"   :: "[pttrn, i, i] \<Rightarrow> i"   (\<open>(3\<Squnion>_ \<in> _./ _)\<close> 10)
-
+syntax_consts
+  "_JOIN1" "_JOIN" == JOIN
 translations
   "\<Squnion>x \<in> A. B" == "CONST JOIN(A, (\<lambda>x. B))"
   "\<Squnion>x y. B"   == "\<Squnion>x. \<Squnion>y. B"
--- a/src/ZF/ZF_Base.thy	Mon Aug 26 18:26:00 2024 +0200
+++ b/src/ZF/ZF_Base.thy	Mon Aug 26 18:26:06 2024 +0200
@@ -38,6 +38,9 @@
 syntax
   "_Ball" :: "[pttrn, i, o] \<Rightarrow> o"  (\<open>(3\<forall>_\<in>_./ _)\<close> 10)
   "_Bex" :: "[pttrn, i, o] \<Rightarrow> o"  (\<open>(3\<exists>_\<in>_./ _)\<close> 10)
+syntax_consts
+  "_Ball" \<rightleftharpoons> Ball and
+  "_Bex" \<rightleftharpoons> Bex
 translations
   "\<forall>x\<in>A. P" \<rightleftharpoons> "CONST Ball(A, \<lambda>x. P)"
   "\<exists>x\<in>A. P" \<rightleftharpoons> "CONST Bex(A, \<lambda>x. P)"
@@ -52,7 +55,9 @@
   where "Replace(A,P) \<equiv> PrimReplace(A, \<lambda>x y. (\<exists>!z. P(x,z)) \<and> P(x,y))"
 
 syntax
-  "_Replace"  :: "[pttrn, pttrn, i, o] \<Rightarrow> i"  (\<open>(1{_ ./ _ \<in> _, _})\<close>)
+  "_Replace" :: "[pttrn, pttrn, i, o] \<Rightarrow> i"  (\<open>(1{_ ./ _ \<in> _, _})\<close>)
+syntax_consts
+  "_Replace" \<rightleftharpoons> Replace
 translations
   "{y. x\<in>A, Q}" \<rightleftharpoons> "CONST Replace(A, \<lambda>x y. Q)"
 
@@ -64,10 +69,11 @@
 
 syntax
   "_RepFun" :: "[i, pttrn, i] \<Rightarrow> i"  (\<open>(1{_ ./ _ \<in> _})\<close> [51,0,51])
+syntax_consts
+  "_RepFun" \<rightleftharpoons> RepFun
 translations
   "{b. x\<in>A}" \<rightleftharpoons> "CONST RepFun(A, \<lambda>x. b)"
 
-
 (* Separation and Pairing can be derived from the Replacement
    and Powerset Axioms using the following definitions. *)
 definition Collect :: "[i, i \<Rightarrow> o] \<Rightarrow> i"
@@ -75,6 +81,8 @@
 
 syntax
   "_Collect" :: "[pttrn, i, o] \<Rightarrow> i"  (\<open>(1{_ \<in> _ ./ _})\<close>)
+syntax_consts
+  "_Collect" \<rightleftharpoons> Collect
 translations
   "{x\<in>A. P}" \<rightleftharpoons> "CONST Collect(A, \<lambda>x. P)"
 
@@ -87,6 +95,9 @@
 syntax
   "_UNION" :: "[pttrn, i, i] \<Rightarrow> i"  (\<open>(3\<Union>_\<in>_./ _)\<close> 10)
   "_INTER" :: "[pttrn, i, i] \<Rightarrow> i"  (\<open>(3\<Inter>_\<in>_./ _)\<close> 10)
+syntax_consts
+  "_UNION" == Union and
+  "_INTER" == Inter
 translations
   "\<Union>x\<in>A. B" == "CONST Union({B. x\<in>A})"
   "\<Inter>x\<in>A. B" == "CONST Inter({B. x\<in>A})"
@@ -123,6 +134,8 @@
   "" :: "i \<Rightarrow> is"  (\<open>_\<close>)
   "_Enum" :: "[i, is] \<Rightarrow> is"  (\<open>_,/ _\<close>)
   "_Finset" :: "is \<Rightarrow> i"  (\<open>{(_)}\<close>)
+syntax_consts
+  "_Finset" == cons
 translations
   "{x, xs}" == "CONST cons(x, {xs})"
   "{x}" == "CONST cons(x, 0)"
@@ -186,6 +199,9 @@
   ""          :: "pttrn \<Rightarrow> patterns"         (\<open>_\<close>)
   "_patterns" :: "[pttrn, patterns] \<Rightarrow> patterns"  (\<open>_,/_\<close>)
   "_Tuple"    :: "[i, is] \<Rightarrow> i"              (\<open>\<langle>(_,/ _)\<rangle>\<close>)
+syntax_consts
+  "_pattern" "_patterns" == split and
+  "_Tuple" == Pair
 translations
   "\<langle>x, y, z\<rangle>"   == "\<langle>x, \<langle>y, z\<rangle>\<rangle>"
   "\<langle>x, y\<rangle>"      == "CONST Pair(x, y)"
@@ -252,6 +268,10 @@
   "_PROD"     :: "[pttrn, i, i] \<Rightarrow> i"        (\<open>(3\<Prod>_\<in>_./ _)\<close> 10)
   "_SUM"      :: "[pttrn, i, i] \<Rightarrow> i"        (\<open>(3\<Sum>_\<in>_./ _)\<close> 10)
   "_lam"      :: "[pttrn, i, i] \<Rightarrow> i"        (\<open>(3\<lambda>_\<in>_./ _)\<close> 10)
+syntax_consts
+  "_PROD" == Pi and
+  "_SUM" == Sigma and
+  "_lam" == Lambda
 translations
   "\<Prod>x\<in>A. B"   == "CONST Pi(A, \<lambda>x. B)"
   "\<Sum>x\<in>A. B"   == "CONST Sigma(A, \<lambda>x. B)"
--- a/src/ZF/func.thy	Mon Aug 26 18:26:00 2024 +0200
+++ b/src/ZF/func.thy	Mon Aug 26 18:26:06 2024 +0200
@@ -448,19 +448,15 @@
 nonterminal updbinds and updbind
 
 syntax
-
-  (* Let expressions *)
-
   "_updbind"    :: "[i, i] \<Rightarrow> updbind"               (\<open>(2_ :=/ _)\<close>)
   ""            :: "updbind \<Rightarrow> updbinds"             (\<open>_\<close>)
   "_updbinds"   :: "[updbind, updbinds] \<Rightarrow> updbinds" (\<open>_,/ _\<close>)
   "_Update"     :: "[i, updbinds] \<Rightarrow> i"              (\<open>_/'((_)')\<close> [900,0] 900)
-
+syntax_consts "_Update" "_updbind" \<rightleftharpoons> update
 translations
   "_Update (f, _updbinds(b,bs))"  == "_Update (_Update(f,b), bs)"
   "f(x:=y)"                       == "CONST update(f,x,y)"
 
-
 lemma update_apply [simp]: "f(x:=y) ` z = (if z=x then y else f`z)"
 apply (simp add: update_def)
 apply (case_tac "z \<in> domain(f)")