merged
authorpaulson
Mon, 24 Apr 2023 15:00:54 +0100
changeset 77932 1c7ce7943396
parent 77912 430e6c477ba4 (diff)
parent 77931 aca0b615ec4f (current diff)
child 77933 0649ff183dcf
merged
--- a/NEWS	Wed Apr 12 12:00:40 2023 +0100
+++ b/NEWS	Mon Apr 24 15:00:54 2023 +0100
@@ -230,6 +230,7 @@
       mult_mono_strong
       multeqp_code_iff_reflclp_multp
       multp_code_iff_multp
+      multp_image_mset_image_msetI
       multp_mono_strong
       multp_repeat_mset_repeat_msetI
       total_mult
@@ -245,8 +246,10 @@
       irreflp_on_multpHO[simp]
       multpDM_mono_strong
       multpDM_plus_plusI[simp]
+      multpHO_double_double[simp]
       multpHO_mono_strong
       multpHO_plus_plus[simp]
+      multpHO_repeat_mset_repeat_mset[simp]
       strict_subset_implies_multpDM
       strict_subset_implies_multpHO
       totalp_multpDM
@@ -280,6 +283,10 @@
 
 *** ML ***
 
+* ML antiquotation \<^if_none>\<open>expr\<close> inlines a function (fn SOME x => x
+| NONE => expr); this is a non-strict version of the regular function
+"the_default". Both are particularly useful with the |> combinator.
+
 * Improved implementations and signatures of functor Table() and
 corresponding functor Set().
 
@@ -292,12 +299,6 @@
   - The new "size" operation works with constant time complexity and
     minimal space overhead: small structures have no size descriptor.
 
-  - Various internal interfaces now use scalable Set() instances instead
-    of plain list, notably Thm.hyps_of and Thm.shyps_of. Minor
-    INCOMPATIBILITY: use e.g. Termset.dest to adapt the result, better
-    use proper Termset operations such as Termset.exists or Termset.fold
-    without the overhead of destruction.
-
 * Operations ML_Heap.sizeof1 and ML_Heap.sizeof determine the object
 size on the heap in bytes. The latter works simultaneously on multiple
 objects, taking implicit sharing of values into account. Examples for
--- a/etc/symbols	Wed Apr 12 12:00:40 2023 +0100
+++ b/etc/symbols	Mon Apr 24 15:00:54 2023 +0100
@@ -506,6 +506,7 @@
 \<^if_macos>            argument: cartouche
 \<^if_windows>          argument: cartouche
 \<^if_unix>             argument: cartouche
+\<^if_none>             argument: cartouche
 \<^cite>                argument: cartouche
 \<^nocite>              argument: cartouche
 \<^citet>               argument: cartouche
--- a/lib/texinputs/isabellesym.sty	Wed Apr 12 12:00:40 2023 +0100
+++ b/lib/texinputs/isabellesym.sty	Mon Apr 24 15:00:54 2023 +0100
@@ -495,6 +495,7 @@
 \newcommand{\isactrlifUNDERSCOREmacos}{\isakeywordcontrol{if{\isacharunderscore}macos}}
 \newcommand{\isactrlifUNDERSCOREwindows}{\isakeywordcontrol{if{\isacharunderscore}windows}}
 \newcommand{\isactrlifUNDERSCOREunix}{\isakeywordcontrol{if{\isacharunderscore}unix}}
+\newcommand{\isactrlifUNDERSCOREnone}{\isakeywordcontrol{if{\isacharunderscore}none}}
 
 \newcommand{\isactrlcite}{\isakeywordcontrol{cite}}
 \newcommand{\isactrlnocite}{\isakeywordcontrol{nocite}}
--- a/src/Doc/Implementation/Logic.thy	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/Doc/Implementation/Logic.thy	Mon Apr 24 15:00:54 2023 +0100
@@ -701,7 +701,7 @@
   \begin{mldecls}
   @{define_ML Theory.add_deps: "Defs.context -> string ->
   Defs.entry -> Defs.entry list -> theory -> theory"} \\
-  @{define_ML Thm_Deps.all_oracles: "thm list -> Oracles.T"} \\
+  @{define_ML Thm_Deps.all_oracles: "thm list -> Proofterm.oracle list"} \\
   \end{mldecls}
 
   \<^descr> \<^ML>\<open>Logic.all\<close>~\<open>a B\<close> produces a Pure quantification \<open>\<And>a. B\<close>, where
@@ -949,7 +949,7 @@
 
 text %mlref \<open>
   \begin{mldecls}
-  @{define_ML Thm.extra_shyps: "thm -> Sortset.T"} \\
+  @{define_ML Thm.extra_shyps: "thm -> sort list"} \\
   @{define_ML Thm.strip_shyps: "thm -> thm"} \\
   \end{mldecls}
 
@@ -971,7 +971,7 @@
 theorem (in empty) false: False
   using bad by blast
 
-ML_val \<open>\<^assert> (Sortset.dest (Thm.extra_shyps @{thm false}) = [\<^sort>\<open>empty\<close>])\<close>
+ML_val \<open>\<^assert> (Thm.extra_shyps @{thm false} = [\<^sort>\<open>empty\<close>])\<close>
 
 text \<open>
   Thanks to the inference kernel managing sort hypothesis according to their
--- a/src/Doc/Implementation/ML.thy	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/Doc/Implementation/ML.thy	Mon Apr 24 15:00:54 2023 +0100
@@ -1495,12 +1495,52 @@
   @{define_ML the_list: "'a option -> 'a list"} \\
   @{define_ML the_default: "'a -> 'a option -> 'a"} \\
   \end{mldecls}
+
+  \begin{matharray}{rcl}
+  @{ML_antiquotation_def "if_none"} & : & \<open>ML_antiquotation\<close> \\
+  \end{matharray}
+
+  \<^rail>\<open>@@{ML_antiquotation if_none} embedded\<close>
 \<close>
 
 text \<open>
   Apart from \<^ML>\<open>Option.map\<close> most other operations defined in structure
   \<^ML_structure>\<open>Option\<close> are alien to Isabelle/ML and never used. The
   operations shown above are defined in \<^file>\<open>~~/src/Pure/General/basics.ML\<close>.
+
+  Note that the function \<^ML>\<open>the_default\<close> is strict in all of its
+  arguments, the default value is evaluated beforehand, even if not required
+  later. In contrast, the antiquotation @{ML_antiquotation "if_none"} is
+  non-strict: the given expression is only evaluated for an application to
+  \<^ML>\<open>NONE\<close>. This allows to work with exceptions like this:
+\<close>
+
+ML \<open>
+  fun div_total x y =
+    \<^try>\<open>x div y\<close> |> the_default 0;
+
+  fun div_error x y =
+    \<^try>\<open>x div y\<close> |> \<^if_none>\<open>error "Division by zero"\<close>;
+\<close>
+
+text \<open>
+  Of course, it is also possible to handle exceptions directly, without an
+  intermediate option construction:
+\<close>
+
+ML \<open>
+  fun div_total x y =
+    x div y handle Div => 0;
+
+  fun div_error x y =
+    x div y handle Div => error "Division by zero";
+\<close>
+
+text \<open>
+  The first form works better in longer chains of functional composition, with
+  combinators like \<^ML>\<open>|>\<close> or \<^ML>\<open>#>\<close> or \<^ML>\<open>o\<close>. The second form is more
+  adequate in elementary expressions: there is no need to pretend that
+  Isabelle/ML is actually a version of Haskell.
 \<close>
 
 
--- a/src/HOL/Analysis/measurable.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/HOL/Analysis/measurable.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -34,82 +34,73 @@
 structure Measurable : MEASURABLE =
 struct
 
-type preprocessor = thm -> Proof.context -> (thm list * Proof.context)
+type preprocessor = thm -> Proof.context -> thm list * Proof.context
 
 datatype level = Concrete | Generic;
 
-fun eq_measurable_thms ((th1, d1), (th2, d2)) = 
+type measurable_thm = thm * (bool * level);
+
+fun eq_measurable_thm ((th1, d1): measurable_thm, (th2, d2): measurable_thm) =
   d1 = d2 andalso Thm.eq_thm_prop (th1, th2) ;
 
-fun merge_dups (xs:(string * preprocessor) list) ys =
-  xs @ (filter (fn (name, _) => is_none (find_first (fn (name', _) => name' = name) xs)) ys) 
+fun merge_preprocessors (xs: (string * preprocessor) list, ys) =
+  xs @ (filter (fn (name, _) => is_none (find_first (fn (name', _) => name' = name) xs)) ys)
 
 structure Data = Generic_Data
 (
-  type T = {
-    measurable_thms : (thm * (bool * level)) Item_Net.T,
-    dest_thms : thm Item_Net.T,
-    cong_thms : thm Item_Net.T,
-    preprocessors : (string * preprocessor) list }
-  val empty: T = {
-    measurable_thms = Item_Net.init eq_measurable_thms (single o Thm.prop_of o fst),
-    dest_thms = Thm.item_net,
-    cong_thms = Thm.item_net,
-    preprocessors = [] };
-  fun merge ({measurable_thms = t1, dest_thms = dt1, cong_thms = ct1, preprocessors = i1 },
-      {measurable_thms = t2, dest_thms = dt2, cong_thms = ct2, preprocessors = i2 }) : T = {
-    measurable_thms = Item_Net.merge (t1, t2),
-    dest_thms = Item_Net.merge (dt1, dt2),
-    cong_thms = Item_Net.merge (ct1, ct2),
-    preprocessors = merge_dups i1 i2 
-    };
+  type T =
+    measurable_thm Item_Net.T *
+    (*dest_thms*) thm Item_Net.T *
+    (*cong_thms*) thm Item_Net.T *
+    (string * preprocessor) list
+  val empty: T =
+    (Item_Net.init eq_measurable_thm (single o Thm.full_prop_of o fst), Thm.item_net, Thm.item_net, [])
+  fun merge
+   ((measurable_thms1, dest_thms1, cong_thms1, preprocessors1),
+    (measurable_thms2, dest_thms2, cong_thms2, preprocessors2)) : T =
+   (Item_Net.merge (measurable_thms1, measurable_thms2),
+    Item_Net.merge (dest_thms1, dest_thms2),
+    Item_Net.merge (cong_thms1, cong_thms2),
+    merge_preprocessors (preprocessors1, preprocessors2))
 );
 
-val debug =
-  Attrib.setup_config_bool \<^binding>\<open>measurable_debug\<close> (K false)
-
-val split =
-  Attrib.setup_config_bool \<^binding>\<open>measurable_split\<close> (K true)
+val map_measurable_thms = Data.map o @{apply 4(1)}
+val map_dest_thms = Data.map o @{apply 4(2)}
+val map_cong_thms = Data.map o @{apply 4(3)}
+val map_preprocessors = Data.map o @{apply 4(4)}
 
-fun map_data f1 f2 f3 f4
-  {measurable_thms = t1,    dest_thms = t2,    cong_thms = t3,    preprocessors = t4 } =
-  {measurable_thms = f1 t1, dest_thms = f2 t2, cong_thms = f3 t3, preprocessors = f4 t4}
-
-fun map_measurable_thms f = map_data f I I I
-fun map_dest_thms f = map_data I f I I
-fun map_cong_thms f = map_data I I f I
-fun map_preprocessors f = map_data I I I f
+val debug = Attrib.setup_config_bool \<^binding>\<open>measurable_debug\<close> (K false)
+val split = Attrib.setup_config_bool \<^binding>\<open>measurable_split\<close> (K true)
 
 fun generic_add_del map : attribute context_parser =
   Scan.lift
     (Args.add >> K Item_Net.update || Args.del >> K Item_Net.remove || Scan.succeed Item_Net.update) >>
-    (fn f => Thm.declaration_attribute (Data.map o map o f))
+    (fn f => Thm.declaration_attribute (map o f o Thm.trim_context))
 
 val dest_thm_attr = generic_add_del map_dest_thms
-
 val cong_thm_attr = generic_add_del map_cong_thms
 
 fun del_thm th net =
   let
-    val thms = net |> Item_Net.content |> filter (fn (th', _) => Thm.eq_thm (th, th'))
+    val thms = net |> Item_Net.content |> filter (fn (th', _) => Thm.eq_thm_prop (th, th'))
   in fold Item_Net.remove thms net end ;
 
 fun measurable_thm_attr (do_add, d) = Thm.declaration_attribute
-  (Data.map o map_measurable_thms o (if do_add then Item_Net.update o rpair d else del_thm))
-
-val get_dest = Item_Net.content o #dest_thms o Data.get;
+  (map_measurable_thms o (if do_add then Item_Net.update o rpair d else del_thm) o Thm.trim_context)
 
-val get_cong = Item_Net.content o #cong_thms o Data.get;
-val add_cong = Data.map o map_cong_thms o Item_Net.update;
-val del_cong = Data.map o map_cong_thms o Item_Net.remove;
+fun get_dest context = map (Thm.transfer'' context) (Item_Net.content (#2 (Data.get context)));
+fun get_cong context = map (Thm.transfer'' context) (Item_Net.content (#3 (Data.get context)));
+
+val add_cong = map_cong_thms o Item_Net.update o Thm.trim_context;
+val del_cong = map_cong_thms o Item_Net.remove o Thm.trim_context;
 fun add_del_cong_thm true = add_cong
   | add_del_cong_thm false = del_cong
 
-fun add_preprocessor name f = Data.map (map_preprocessors (fn xs => xs @ [(name, f)]))
-fun del_preprocessor name = Data.map (map_preprocessors (filter (fn (n, _) => n <> name)))
+fun add_preprocessor name f = map_preprocessors (fn xs => xs @ [(name, f)])
+fun del_preprocessor name = map_preprocessors (filter (fn (n, _) => n <> name))
 val add_local_cong = Context.proof_map o add_cong
 
-val get_preprocessors = Context.Proof #> Data.get #> #preprocessors ;
+val get_preprocessors = Context.Proof #> Data.get #> #4;
 
 fun is_too_generic thm =
   let 
@@ -117,9 +108,10 @@
     val concl' = HOLogic.dest_Trueprop concl handle TERM _ => concl
   in is_Var (head_of concl') end
 
-val get_thms = Data.get #> #measurable_thms #> Item_Net.content ;
+fun get_thms context =
+  map (apfst (Thm.transfer'' context)) (Item_Net.content (#1 (Data.get context)));
 
-val get_all = get_thms #> map fst ;
+val get_all = get_thms #> map fst;
 
 fun debug_tac ctxt msg f =
   if Config.get ctxt debug then print_tac ctxt (msg ()) THEN f else f
--- a/src/HOL/Data_Structures/Tries_Binary.thy	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/HOL/Data_Structures/Tries_Binary.thy	Mon Apr 24 15:00:54 2023 +0100
@@ -91,9 +91,39 @@
 lemma set_trie_delete: "set_trie(delete xs t) = set_trie t - {xs}"
 by(auto simp add: isin_delete set_trie_def)
 
+text \<open>Invariant: tries are fully shrunk:\<close>
+fun invar where
+"invar Lf = True" |
+"invar (Nd b (l,r)) = (invar l \<and> invar r \<and> (l = Lf \<and> r = Lf \<longrightarrow> b))"
+
+lemma insert_Lf: "insert xs t \<noteq> Lf"
+using insert.elims by blast
+
+lemma invar_insert: "invar t \<Longrightarrow> invar(insert xs t)"
+proof(induction xs t rule: insert.induct)
+  case 1 thus ?case by simp
+next
+  case (2 b lr)
+  thus ?case by(cases lr; simp)
+next
+  case (3 k ks)
+  thus ?case by(simp; cases ks; auto)
+next
+  case (4 k ks b lr)
+  then show ?case by(cases lr; auto simp: insert_Lf)
+qed
+
+lemma invar_delete: "invar t \<Longrightarrow> invar(delete xs t)"
+proof(induction t arbitrary: xs)
+  case Lf thus ?case by simp
+next
+  case (Nd b lr)
+  thus ?case by(cases lr)(auto split: list.split)
+qed
+
 interpretation S: Set
 where empty = empty and isin = isin and insert = insert and delete = delete
-and set = set_trie and invar = "\<lambda>t. True"
+and set = set_trie and invar = invar
 proof (standard, goal_cases)
   case 1 show ?case by (rule set_trie_empty)
 next
@@ -102,13 +132,24 @@
   case 3 thus ?case by(auto simp: set_trie_insert)
 next
   case 4 show ?case by(rule set_trie_delete)
-qed (rule TrueI)+
+next
+  case 5 show ?case by(simp)
+next
+  case 6 thus ?case by(rule invar_insert)
+next
+  case 7 thus ?case by(rule invar_delete)
+qed
 
 
 subsection "Patricia Trie"
 
 datatype trieP = LfP | NdP "bool list" bool "trieP * trieP"
 
+text \<open>Fully shrunk:\<close>
+fun invarP where
+"invarP LfP = True" |
+"invarP (NdP ps b (l,r)) = (invarP l \<and> invarP r \<and> (l = LfP \<or> r = LfP \<longrightarrow> b))"
+
 fun isinP :: "trieP \<Rightarrow> bool list \<Rightarrow> bool" where
 "isinP LfP ks = False" |
 "isinP (NdP ps b lr) ks =
@@ -120,12 +161,12 @@
 definition emptyP :: trieP where
 [simp]: "emptyP = LfP"
 
-fun split where
-"split [] ys = ([],[],ys)" |
-"split xs [] = ([],xs,[])" |
-"split (x#xs) (y#ys) =
+fun lcp :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list \<times> 'a list \<times> 'a list" where
+"lcp [] ys = ([],[],ys)" |
+"lcp xs [] = ([],xs,[])" |
+"lcp (x#xs) (y#ys) =
   (if x\<noteq>y then ([],x#xs,y#ys)
-   else let (ps,xs',ys') = split xs ys in (x#ps,xs',ys'))"
+   else let (ps,xs',ys') = lcp xs ys in (x#ps,xs',ys'))"
 
 
 lemma mod2_cong[fundef_cong]:
@@ -137,7 +178,7 @@
 fun insertP :: "bool list \<Rightarrow> trieP \<Rightarrow> trieP" where
 "insertP ks LfP  = NdP ks True (LfP,LfP)" |
 "insertP ks (NdP ps b lr) =
-  (case split ks ps of
+  (case lcp ks ps of
      (qs, k#ks', p#ps') \<Rightarrow>
        let tp = NdP ps' b lr; tk = NdP ks' True (LfP,LfP) in
        NdP qs False (if k then (tp,tk) else (tk,tp)) |
@@ -149,18 +190,26 @@
      (qs,[],[]) \<Rightarrow> NdP ps True lr)"
 
 
-fun nodeP :: "bool list \<Rightarrow> bool \<Rightarrow> trieP * trieP \<Rightarrow> trieP" where
-"nodeP ps b lr = (if \<not> b \<and> lr = (LfP,LfP) then LfP else NdP ps b lr)"
+text \<open>Smart constructor that shrinks:\<close>
+definition nodeP :: "bool list \<Rightarrow> bool \<Rightarrow> trieP * trieP \<Rightarrow> trieP" where
+"nodeP ps b lr =
+ (if b then  NdP ps b lr
+  else case lr of
+   (LfP,LfP) \<Rightarrow> LfP |
+   (LfP, NdP ks b lr) \<Rightarrow> NdP (ps @ True # ks) b lr |
+   (NdP ks b lr, LfP) \<Rightarrow> NdP (ps @ False # ks) b lr |
+   _ \<Rightarrow> NdP ps b lr)"
 
 fun deleteP :: "bool list \<Rightarrow> trieP \<Rightarrow> trieP" where
 "deleteP ks LfP  = LfP" |
 "deleteP ks (NdP ps b lr) =
-  (case split ks ps of
+  (case lcp ks ps of
      (_, _, _#_) \<Rightarrow> NdP ps b lr |
      (_, k#ks', []) \<Rightarrow> nodeP ps b (mod2 (deleteP ks') k lr) |
      (_, [], []) \<Rightarrow> nodeP ps False lr)"
 
 
+
 subsubsection \<open>Functional Correctness\<close>
 
 text \<open>First step: @{typ trieP} implements @{typ trie} via the abstraction function \<open>abs_trieP\<close>:\<close>
@@ -214,9 +263,9 @@
 apply auto
 done
 
-lemma split_if: "split ks ps = (qs, ks', ps') \<Longrightarrow>
+lemma lcp_if: "lcp ks ps = (qs, ks', ps') \<Longrightarrow>
   ks = qs @ ks' \<and> ps = qs @ ps' \<and> (ks' \<noteq> [] \<and> ps' \<noteq> [] \<longrightarrow> hd ks' \<noteq> hd ps')"
-apply(induction ks ps arbitrary: qs ks' ps' rule: split.induct)
+apply(induction ks ps arbitrary: qs ks' ps' rule: lcp.induct)
 apply(auto split: prod.splits if_splits)
 done
 
@@ -224,7 +273,7 @@
   "abs_trieP (insertP ks t) = insert ks (abs_trieP t)"
 apply(induction t arbitrary: ks)
 apply(auto simp: prefix_trie_Lfs insert_prefix_trie_same insert_append prefix_trie_append
-           dest!: split_if split: list.split prod.split if_splits)
+           dest!: lcp_if split: list.split prod.split if_splits)
 done
 
 
@@ -246,14 +295,52 @@
    = (if delete ys t = Lf then Lf else prefix_trie xs (delete ys t))"
 by(induction xs)(auto simp: prefix_trie_Lf)
 
+lemma nodeP_LfP2: "nodeP xs False (LfP, LfP) = LfP"
+by(simp add: nodeP_def)
+
+text \<open>Some non-inductive aux. lemmas:\<close>
+
+lemma abs_trieP_nodeP: "a\<noteq>LfP \<or> b \<noteq> LfP \<Longrightarrow>
+  abs_trieP (nodeP xs f (a, b)) = prefix_trie xs (Nd f (abs_trieP a, abs_trieP b))"
+by(auto simp add: nodeP_def prefix_trie_append split: trieP.split)
+
+lemma nodeP_True: "nodeP ps True lr = NdP ps True lr"
+by(simp add: nodeP_def)
+
 lemma delete_abs_trieP:
   "delete ks (abs_trieP t) = abs_trieP (deleteP ks t)"
 apply(induction t arbitrary: ks)
 apply(auto simp: delete_prefix_trie delete_append_prefix_trie
-        prefix_trie_append prefix_trie_Lf abs_trieP_Lf
-        dest!: split_if split: if_splits list.split prod.split)
+        prefix_trie_append prefix_trie_Lf abs_trieP_Lf nodeP_LfP2 abs_trieP_nodeP nodeP_True
+        dest!: lcp_if split: if_splits list.split prod.split)
 done
 
+text \<open>Invariant preservation:\<close>
+
+lemma insertP_LfP: "insertP xs t \<noteq> LfP"
+by(cases t)(auto split: prod.split list.split)
+
+lemma invarP_insertP: "invarP t \<Longrightarrow> invarP(insertP xs t)"
+proof(induction t arbitrary: xs)
+  case LfP thus ?case by simp
+next
+  case (NdP bs b lr)
+  then show ?case
+    by(cases lr)(auto simp: insertP_LfP split: prod.split list.split)
+qed
+
+(* Inlining this proof leads to nontermination *)
+lemma invarP_nodeP: "\<lbrakk> invarP t1; invarP t2\<rbrakk> \<Longrightarrow> invarP (nodeP xs b (t1, t2))"
+by (auto simp add: nodeP_def split: trieP.split)
+
+lemma invarP_deleteP: "invarP t \<Longrightarrow> invarP(deleteP xs t)"
+proof(induction t arbitrary: xs)
+  case LfP thus ?case by simp
+next
+  case (NdP ks b lr)
+  thus ?case by(cases lr)(auto simp: invarP_nodeP split: prod.split list.split)
+qed
+
 
 text \<open>The overall correctness proof. Simply composes correctness lemmas.\<close>
 
@@ -271,7 +358,7 @@
 
 interpretation SP: Set
 where empty = emptyP and isin = isinP and insert = insertP and delete = deleteP
-and set = set_trieP and invar = "\<lambda>t. True"
+and set = set_trieP and invar = invarP
 proof (standard, goal_cases)
   case 1 show ?case by (simp add: set_trieP_def set_trie_def)
 next
@@ -280,6 +367,12 @@
   case 3 thus ?case by (auto simp: set_trieP_insertP)
 next
   case 4 thus ?case by(auto simp: set_trieP_deleteP)
-qed (rule TrueI)+
+next
+  case 5 thus ?case by(simp)
+next
+  case 6 thus ?case by(rule invarP_insertP)
+next
+  case 7 thus ?case by(rule invarP_deleteP)
+qed
 
 end
--- a/src/HOL/Decision_Procs/ferrante_rackoff.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/HOL/Decision_Procs/ferrante_rackoff.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -115,11 +115,11 @@
    val [pinf_conj, pinf_disj, pinf_eq, pinf_neq, pinf_lt,
         pinf_le, pinf_gt, pinf_ge, pinf_P] = pinf
    val [nmi_conj, nmi_disj, nmi_eq, nmi_neq, nmi_lt,
-        nmi_le, nmi_gt, nmi_ge, nmi_P] = map (Drule.instantiate_normalize (TVars.empty, Vars.make [(U_v,cU)])) nmi
+        nmi_le, nmi_gt, nmi_ge, nmi_P] = map (Drule.instantiate_normalize (TVars.empty, Vars.make1 (U_v,cU))) nmi
    val [npi_conj, npi_disj, npi_eq, npi_neq, npi_lt,
-        npi_le, npi_gt, npi_ge, npi_P] = map (Drule.instantiate_normalize (TVars.empty, Vars.make [(U_v,cU)])) npi
+        npi_le, npi_gt, npi_ge, npi_P] = map (Drule.instantiate_normalize (TVars.empty, Vars.make1 (U_v,cU))) npi
    val [ld_conj, ld_disj, ld_eq, ld_neq, ld_lt,
-        ld_le, ld_gt, ld_ge, ld_P] = map (Drule.instantiate_normalize (TVars.empty, Vars.make [(U_v,cU)])) ld
+        ld_le, ld_gt, ld_ge, ld_P] = map (Drule.instantiate_normalize (TVars.empty, Vars.make1 (U_v,cU))) ld
 
    fun decomp_mpinf fm =
      case Thm.term_of fm of
--- a/src/HOL/Examples/Gauss_Numbers.thy	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/HOL/Examples/Gauss_Numbers.thy	Mon Apr 24 15:00:54 2023 +0100
@@ -4,7 +4,7 @@
 section \<open>Gauss Numbers: integral gauss numbers\<close>
 
 theory Gauss_Numbers
-  imports "HOL-Library.Rounded_Division"
+  imports "HOL-Library.Centered_Division"
 begin
 
 codatatype gauss = Gauss (Re: int) (Im: int)
@@ -308,17 +308,17 @@
 
 primcorec divide_gauss :: \<open>gauss \<Rightarrow> gauss \<Rightarrow> gauss\<close>
   where
-    \<open>Re (x div y) = (Re x * Re y + Im x * Im y) rdiv ((Re y)\<^sup>2 + (Im y)\<^sup>2)\<close>
-  | \<open>Im (x div y) = (Im x * Re y - Re x * Im y) rdiv ((Re y)\<^sup>2 + (Im y)\<^sup>2)\<close>
+    \<open>Re (x div y) = (Re x * Re y + Im x * Im y) cdiv ((Re y)\<^sup>2 + (Im y)\<^sup>2)\<close>
+  | \<open>Im (x div y) = (Im x * Re y - Re x * Im y) cdiv ((Re y)\<^sup>2 + (Im y)\<^sup>2)\<close>
 
 primcorec modulo_gauss :: \<open>gauss \<Rightarrow> gauss \<Rightarrow> gauss\<close>
   where
     \<open>Re (x mod y) = Re x -
-      ((Re x * Re y + Im x * Im y) rdiv ((Re y)\<^sup>2 + (Im y)\<^sup>2) * Re y -
-       (Im x * Re y - Re x * Im y) rdiv ((Re y)\<^sup>2 + (Im y)\<^sup>2) * Im y)\<close>
+      ((Re x * Re y + Im x * Im y) cdiv ((Re y)\<^sup>2 + (Im y)\<^sup>2) * Re y -
+       (Im x * Re y - Re x * Im y) cdiv ((Re y)\<^sup>2 + (Im y)\<^sup>2) * Im y)\<close>
   | \<open>Im (x mod y) = Im x -
-      ((Re x * Re y + Im x * Im y) rdiv ((Re y)\<^sup>2 + (Im y)\<^sup>2) * Im y +
-       (Im x * Re y - Re x * Im y) rdiv ((Re y)\<^sup>2 + (Im y)\<^sup>2) * Re y)\<close>
+      ((Re x * Re y + Im x * Im y) cdiv ((Re y)\<^sup>2 + (Im y)\<^sup>2) * Im y +
+       (Im x * Re y - Re x * Im y) cdiv ((Re y)\<^sup>2 + (Im y)\<^sup>2) * Re y)\<close>
 
 instance proof
   fix x y :: gauss
@@ -334,7 +334,7 @@
       \<open>(Im y)\<^sup>2 + (Re y)\<^sup>2 = Y\<close>
       by (simp_all add: power2_eq_square algebra_simps Y_def)
     from \<open>Y > 0\<close> show ?thesis
-      by (simp add: gauss_eq_iff algebra_simps) (simp add: * nonzero_mult_rdiv_cancel_right)
+      by (simp add: gauss_eq_iff algebra_simps) (simp add: * nonzero_mult_cdiv_cancel_right)
   qed
   show \<open>x div y * y + x mod y = x\<close>
     by (simp add: gauss_eq_iff)
@@ -360,22 +360,22 @@
       by (simp_all add: gauss_neq_0 euclidean_size_gauss_def)
     have \<open>X * Y = R\<^sup>2 + I\<^sup>2\<close>
       by (simp add: R_def I_def X_def Y_def power2_eq_square algebra_simps)
-    let ?lhs = \<open>X - I * (I rdiv Y) - R * (R rdiv Y)
-        - I rdiv Y * (I rmod Y) - R rdiv Y * (R rmod Y)\<close>
-    have \<open>?lhs = X + Y * (R rdiv Y) * (R rdiv Y) + Y * (I rdiv Y) * (I rdiv Y)
-        - 2 * (R rdiv Y * R + I rdiv Y * I)\<close>
-      by (simp flip: minus_rmod_eq_mult_rdiv add: algebra_simps)
+    let ?lhs = \<open>X - I * (I cdiv Y) - R * (R cdiv Y)
+        - I cdiv Y * (I cmod Y) - R cdiv Y * (R cmod Y)\<close>
+    have \<open>?lhs = X + Y * (R cdiv Y) * (R cdiv Y) + Y * (I cdiv Y) * (I cdiv Y)
+        - 2 * (R cdiv Y * R + I cdiv Y * I)\<close>
+      by (simp flip: minus_cmod_eq_mult_cdiv add: algebra_simps)
     also have \<open>\<dots> = (Re (x mod y))\<^sup>2 + (Im (x mod y))\<^sup>2\<close>
       by (simp add: X_def Y_def R_def I_def algebra_simps power2_eq_square)
     finally have lhs: \<open>int (euclidean_size (x mod y)) = ?lhs\<close>
       by (simp add: euclidean_size_gauss_def)
-    have \<open>?lhs * Y = (I rmod Y)\<^sup>2 + (R rmod Y)\<^sup>2\<close>
+    have \<open>?lhs * Y = (I cmod Y)\<^sup>2 + (R cmod Y)\<^sup>2\<close>
       apply (simp add: algebra_simps power2_eq_square \<open>X * Y = R\<^sup>2 + I\<^sup>2\<close>)
-      apply (simp flip: mult.assoc add.assoc minus_rmod_eq_mult_rdiv)
+      apply (simp flip: mult.assoc add.assoc minus_cmod_eq_mult_cdiv)
       apply (simp add: algebra_simps)
       done
     also have \<open>\<dots> \<le> (Y div 2)\<^sup>2 + (Y div 2)\<^sup>2\<close>
-      by (rule add_mono) (use \<open>Y > 0\<close> abs_rmod_less_equal [of Y] in \<open>simp_all add: power2_le_iff_abs_le\<close>)
+      by (rule add_mono) (use \<open>Y > 0\<close> abs_cmod_less_equal [of Y] in \<open>simp_all add: power2_le_iff_abs_le\<close>)
     also have \<open>\<dots> < Y\<^sup>2\<close>
       using \<open>Y > 0\<close> by (cases \<open>Y = 1\<close>) (simp_all add: power2_eq_square mult_le_less_imp_less flip: mult.assoc)
     finally have \<open>?lhs * Y < Y\<^sup>2\<close> .
--- a/src/HOL/Examples/Iff_Oracle.thy	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/HOL/Examples/Iff_Oracle.thy	Mon Apr 24 15:00:54 2023 +0100
@@ -36,8 +36,7 @@
 ML \<open>iff_oracle (\<^theory>, 10)\<close>
 
 ML \<open>
-  \<^assert> (map (#1 o #1) (Oracles.dest (Thm_Deps.all_oracles [iff_oracle (\<^theory>, 10)])) =
-    [\<^oracle_name>\<open>iff_oracle\<close>]);
+  \<^assert> (map (#1 o #1) (Thm_Deps.all_oracles [iff_oracle (\<^theory>, 10)]) = [\<^oracle_name>\<open>iff_oracle\<close>]);
 \<close>
 
 text \<open>These oracle calls had better fail.\<close>
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Library/Centered_Division.thy	Mon Apr 24 15:00:54 2023 +0100
@@ -0,0 +1,178 @@
+(*  Author:  Florian Haftmann, TU Muenchen
+*)
+
+section \<open>Division with modulus centered towards zero.\<close>
+
+theory Centered_Division
+  imports Main
+begin
+
+lemma off_iff_abs_mod_2_eq_one:
+  \<open>odd l \<longleftrightarrow> \<bar>l\<bar> mod 2 = 1\<close> for l :: int
+  by (simp flip: odd_iff_mod_2_eq_one)
+
+text \<open>
+  \noindent The following specification of division on integers centers
+  the modulus around zero.  This is useful e.g.~to define division
+  on Gauss numbers.
+  N.b.: This is not mentioned \cite{leijen01}.
+\<close>
+
+definition centered_divide :: \<open>int \<Rightarrow> int \<Rightarrow> int\<close>  (infixl \<open>cdiv\<close> 70)
+  where \<open>k cdiv l = sgn l * ((k + \<bar>l\<bar> div 2) div \<bar>l\<bar>)\<close>
+
+definition centered_modulo :: \<open>int \<Rightarrow> int \<Rightarrow> int\<close>  (infixl \<open>cmod\<close> 70)
+  where \<open>k cmod l = (k + \<bar>l\<bar> div 2) mod \<bar>l\<bar> - \<bar>l\<bar> div 2\<close>
+
+text \<open>
+  \noindent Example: @{lemma \<open>k cmod 5 \<in> {-2, -1, 0, 1, 2}\<close> by (auto simp add: centered_modulo_def)}
+\<close>
+
+lemma signed_take_bit_eq_cmod:
+  \<open>signed_take_bit n k = k cmod (2 ^ Suc n)\<close>
+  by (simp only: centered_modulo_def power_abs abs_numeral flip: take_bit_eq_mod)
+    (simp add: signed_take_bit_eq_take_bit_shift)
+
+text \<open>
+  \noindent Property @{thm signed_take_bit_eq_cmod [no_vars]} is the key to generalize
+  centered division to arbitrary structures satisfying \<^class>\<open>ring_bit_operations\<close>,
+  but so far it is not clear what practical relevance that would have.
+\<close>
+
+lemma cdiv_mult_cmod_eq:
+  \<open>k cdiv l * l + k cmod l = k\<close>
+proof -
+  have *: \<open>l * (sgn l * j) = \<bar>l\<bar> * j\<close> for j
+    by (simp add: ac_simps abs_sgn)
+  show ?thesis
+    by (simp add: centered_divide_def centered_modulo_def algebra_simps *)
+qed
+
+lemma mult_cdiv_cmod_eq:
+  \<open>l * (k cdiv l) + k cmod l = k\<close>
+  using cdiv_mult_cmod_eq [of k l] by (simp add: ac_simps)
+
+lemma cmod_cdiv_mult_eq:
+  \<open>k cmod l + k cdiv l * l = k\<close>
+  using cdiv_mult_cmod_eq [of k l] by (simp add: ac_simps)
+
+lemma cmod_mult_cdiv_eq:
+  \<open>k cmod l + l * (k cdiv l) = k\<close>
+  using cdiv_mult_cmod_eq [of k l] by (simp add: ac_simps)
+
+lemma minus_cdiv_mult_eq_cmod:
+  \<open>k - k cdiv l * l = k cmod l\<close>
+  by (rule add_implies_diff [symmetric]) (fact cmod_cdiv_mult_eq)
+
+lemma minus_mult_cdiv_eq_cmod:
+  \<open>k - l * (k cdiv l) = k cmod l\<close>
+  by (rule add_implies_diff [symmetric]) (fact cmod_mult_cdiv_eq)
+
+lemma minus_cmod_eq_cdiv_mult:
+  \<open>k - k cmod l = k cdiv l * l\<close>
+  by (rule add_implies_diff [symmetric]) (fact cdiv_mult_cmod_eq)
+
+lemma minus_cmod_eq_mult_cdiv:
+  \<open>k - k cmod l = l * (k cdiv l)\<close>
+  by (rule add_implies_diff [symmetric]) (fact mult_cdiv_cmod_eq)
+
+lemma cdiv_0_eq [simp]:
+  \<open>k cdiv 0 = 0\<close>
+  by (simp add: centered_divide_def)
+
+lemma cmod_0_eq [simp]:
+  \<open>k cmod 0 = k\<close>
+  by (simp add: centered_modulo_def)
+
+lemma cdiv_1_eq [simp]:
+  \<open>k cdiv 1 = k\<close>
+  by (simp add: centered_divide_def)
+
+lemma cmod_1_eq [simp]:
+  \<open>k cmod 1 = 0\<close>
+  by (simp add: centered_modulo_def)
+
+lemma zero_cdiv_eq [simp]:
+  \<open>0 cdiv k = 0\<close>
+  by (auto simp add: centered_divide_def not_less zdiv_eq_0_iff)
+
+lemma zero_cmod_eq [simp]:
+  \<open>0 cmod k = 0\<close>
+  by (auto simp add: centered_modulo_def not_less zmod_trivial_iff)
+
+lemma cdiv_minus_eq:
+  \<open>k cdiv - l = - (k cdiv l)\<close>
+  by (simp add: centered_divide_def)
+
+lemma cmod_minus_eq [simp]:
+  \<open>k cmod - l = k cmod l\<close>
+  by (simp add: centered_modulo_def)
+
+lemma cdiv_abs_eq:
+  \<open>k cdiv \<bar>l\<bar> = sgn l * (k cdiv l)\<close>
+  by (simp add: centered_divide_def)
+
+lemma cmod_abs_eq [simp]:
+  \<open>k cmod \<bar>l\<bar> = k cmod l\<close>
+  by (simp add: centered_modulo_def)
+
+lemma nonzero_mult_cdiv_cancel_right:
+  \<open>k * l cdiv l = k\<close> if \<open>l \<noteq> 0\<close>
+proof -
+  have \<open>sgn l * k * \<bar>l\<bar> cdiv l = k\<close>
+    using that by (simp add: centered_divide_def)
+  with that show ?thesis
+    by (simp add: ac_simps abs_sgn)
+qed
+
+lemma cdiv_self_eq [simp]:
+  \<open>k cdiv k = 1\<close> if \<open>k \<noteq> 0\<close>
+  using that nonzero_mult_cdiv_cancel_right [of k 1] by simp
+
+lemma cmod_self_eq [simp]:
+  \<open>k cmod k = 0\<close>
+proof -
+  have \<open>(sgn k * \<bar>k\<bar> + \<bar>k\<bar> div 2) mod \<bar>k\<bar> = \<bar>k\<bar> div 2\<close>
+    by (auto simp add: zmod_trivial_iff)
+  also have \<open>sgn k * \<bar>k\<bar> = k\<close>
+    by (simp add: abs_sgn)
+  finally show ?thesis
+    by (simp add: centered_modulo_def algebra_simps)
+qed
+
+lemma cmod_less_divisor:
+  \<open>k cmod l < \<bar>l\<bar> - \<bar>l\<bar> div 2\<close> if \<open>l \<noteq> 0\<close>
+  using that pos_mod_bound [of \<open>\<bar>l\<bar>\<close>] by (simp add: centered_modulo_def)
+
+lemma cmod_less_equal_divisor:
+  \<open>k cmod l \<le> \<bar>l\<bar> div 2\<close> if \<open>l \<noteq> 0\<close>
+proof -
+  from that cmod_less_divisor [of l k]
+  have \<open>k cmod l < \<bar>l\<bar> - \<bar>l\<bar> div 2\<close>
+    by simp
+  also have \<open>\<bar>l\<bar> - \<bar>l\<bar> div 2 = \<bar>l\<bar> div 2 + of_bool (odd l)\<close>
+    by auto
+  finally show ?thesis
+    by (cases \<open>even l\<close>) simp_all
+qed
+
+lemma divisor_less_equal_cmod':
+  \<open>\<bar>l\<bar> div 2 - \<bar>l\<bar> \<le> k cmod l\<close> if \<open>l \<noteq> 0\<close>
+proof -
+  have \<open>0 \<le> (k + \<bar>l\<bar> div 2) mod \<bar>l\<bar>\<close>
+    using that pos_mod_sign [of \<open>\<bar>l\<bar>\<close>] by simp
+  then show ?thesis
+    by (simp_all add: centered_modulo_def)
+qed
+
+lemma divisor_less_equal_cmod:
+  \<open>- (\<bar>l\<bar> div 2) \<le> k cmod l\<close> if \<open>l \<noteq> 0\<close>
+  using that divisor_less_equal_cmod' [of l k]
+  by (simp add: centered_modulo_def)
+
+lemma abs_cmod_less_equal:
+  \<open>\<bar>k cmod l\<bar> \<le> \<bar>l\<bar> div 2\<close> if \<open>l \<noteq> 0\<close>
+  using that divisor_less_equal_cmod [of l k]
+  by (simp add: abs_le_iff cmod_less_equal_divisor)
+
+end
--- a/src/HOL/Library/Library.thy	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/HOL/Library/Library.thy	Mon Apr 24 15:00:54 2023 +0100
@@ -6,6 +6,7 @@
   BNF_Axiomatization
   BNF_Corec
   Bourbaki_Witt_Fixpoint
+  Centered_Division
   Char_ord
   Code_Cardinality
   Code_Lazy
@@ -78,7 +79,6 @@
   Ramsey
   Reflection
   Rewrite
-  Rounded_Division
   Saturated
   Set_Algebras
   Set_Idioms
--- a/src/HOL/Library/Multiset.thy	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/HOL/Library/Multiset.thy	Mon Apr 24 15:00:54 2023 +0100
@@ -3257,6 +3257,28 @@
   shows "monotone (multp orda) (multp ordb) (image_mset f)"
   by (rule monotone_on_multp_multp_image_mset[OF assms, simplified])
 
+lemma multp_image_mset_image_msetI:
+  assumes "multp (\<lambda>x y. R (f x) (f y)) M1 M2" and "transp R"
+  shows "multp R (image_mset f M1) (image_mset f M2)"
+proof -
+  from \<open>transp R\<close> have "transp (\<lambda>x y. R (f x) (f y))"
+    by (auto intro: transpI dest: transpD)
+  with \<open>multp (\<lambda>x y. R (f x) (f y)) M1 M2\<close> obtain I J K where
+    "M2 = I + J" and "M1 = I + K" and "J \<noteq> {#}" and "\<forall>k\<in>#K. \<exists>x\<in>#J. R (f k) (f x)"
+    using multp_implies_one_step by blast
+
+  have "multp R (image_mset f I + image_mset f K) (image_mset f I + image_mset f J)"
+  proof (rule one_step_implies_multp)
+    show "image_mset f J \<noteq> {#}"
+      by (simp add: \<open>J \<noteq> {#}\<close>)
+  next
+    show "\<forall>k\<in>#image_mset f K. \<exists>j\<in>#image_mset f J. R k j"
+      by (simp add: \<open>\<forall>k\<in>#K. \<exists>x\<in>#J. R (f k) (f x)\<close>)
+  qed
+  thus ?thesis
+    by (simp add: \<open>M1 = I + K\<close> \<open>M2 = I + J\<close>)
+qed
+
 lemma multp_image_mset_image_msetD:
   assumes
     "multp R (image_mset f A) (image_mset f B)" and
--- a/src/HOL/Library/Multiset_Order.thy	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/HOL/Library/Multiset_Order.thy	Mon Apr 24 15:00:54 2023 +0100
@@ -489,6 +489,47 @@
   by (meson less_eq_multiset_def mset_lt_single_right_iff)
 
 
+subsubsection \<open>Simplifications\<close>
+
+lemma multp\<^sub>H\<^sub>O_repeat_mset_repeat_mset[simp]:
+  assumes "n \<noteq> 0"
+  shows "multp\<^sub>H\<^sub>O R (repeat_mset n A) (repeat_mset n B) \<longleftrightarrow> multp\<^sub>H\<^sub>O R A B"
+proof (rule iffI)
+  assume hyp: "multp\<^sub>H\<^sub>O R (repeat_mset n A) (repeat_mset n B)"
+  hence
+    1: "repeat_mset n A \<noteq> repeat_mset n B" and
+    2: "\<forall>y. n * count B y < n * count A y \<longrightarrow> (\<exists>x. R y x \<and> n * count A x < n * count B x)"
+    by (simp_all add: multp\<^sub>H\<^sub>O_def)
+
+  from 1 \<open>n \<noteq> 0\<close> have "A \<noteq> B"
+    by auto
+
+  moreover from 2 \<open>n \<noteq> 0\<close> have "\<forall>y. count B y < count A y \<longrightarrow> (\<exists>x. R y x \<and> count A x < count B x)"
+    by auto
+
+  ultimately show "multp\<^sub>H\<^sub>O R A B"
+    by (simp add: multp\<^sub>H\<^sub>O_def)
+next
+  assume "multp\<^sub>H\<^sub>O R A B"
+  hence 1: "A \<noteq> B" and 2: "\<forall>y. count B y < count A y \<longrightarrow> (\<exists>x. R y x \<and> count A x < count B x)"
+    by (simp_all add: multp\<^sub>H\<^sub>O_def)
+
+  from 1 have "repeat_mset n A \<noteq> repeat_mset n B"
+    by (simp add: assms repeat_mset_cancel1)
+
+  moreover from 2 have "\<forall>y. n * count B y < n * count A y \<longrightarrow>
+    (\<exists>x. R y x \<and> n * count A x < n * count B x)"
+    by auto
+
+  ultimately show "multp\<^sub>H\<^sub>O R (repeat_mset n A) (repeat_mset n B)"
+    by (simp add: multp\<^sub>H\<^sub>O_def)
+qed
+
+lemma multp\<^sub>H\<^sub>O_double_double[simp]: "multp\<^sub>H\<^sub>O R (A + A) (B + B) \<longleftrightarrow> multp\<^sub>H\<^sub>O R A B"
+  using multp\<^sub>H\<^sub>O_repeat_mset_repeat_mset[of 2]
+  by (simp add: numeral_Bit0)
+
+
 subsection \<open>Simprocs\<close>
 
 lemma mset_le_add_iff1:
--- a/src/HOL/Library/Rounded_Division.thy	Wed Apr 12 12:00:40 2023 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,178 +0,0 @@
-(*  Author:  Florian Haftmann, TU Muenchen
-*)
-
-section \<open>Rounded division: modulus centered towards zero.\<close>
-
-theory Rounded_Division
-  imports Main
-begin
-
-lemma off_iff_abs_mod_2_eq_one:
-  \<open>odd l \<longleftrightarrow> \<bar>l\<bar> mod 2 = 1\<close> for l :: int
-  by (simp flip: odd_iff_mod_2_eq_one)
-
-text \<open>
-  \noindent The following specification of division on integers centers
-  the modulus around zero.  This is useful e.g.~to define division
-  on Gauss numbers.
-  N.b.: This is not mentioned \cite{leijen01}.
-\<close>
-
-definition rounded_divide :: \<open>int \<Rightarrow> int \<Rightarrow> int\<close>  (infixl \<open>rdiv\<close> 70)
-  where \<open>k rdiv l = sgn l * ((k + \<bar>l\<bar> div 2) div \<bar>l\<bar>)\<close>
-
-definition rounded_modulo :: \<open>int \<Rightarrow> int \<Rightarrow> int\<close>  (infixl \<open>rmod\<close> 70)
-  where \<open>k rmod l = (k + \<bar>l\<bar> div 2) mod \<bar>l\<bar> - \<bar>l\<bar> div 2\<close>
-
-text \<open>
-  \noindent Example: @{lemma \<open>k rmod 5 \<in> {-2, -1, 0, 1, 2}\<close> by (auto simp add: rounded_modulo_def)}
-\<close>
-
-lemma signed_take_bit_eq_rmod:
-  \<open>signed_take_bit n k = k rmod (2 ^ Suc n)\<close>
-  by (simp only: rounded_modulo_def power_abs abs_numeral flip: take_bit_eq_mod)
-    (simp add: signed_take_bit_eq_take_bit_shift)
-
-text \<open>
-  \noindent Property @{thm signed_take_bit_eq_rmod [no_vars]} is the key to generalize
-  rounded division to arbitrary structures satisfying \<^class>\<open>ring_bit_operations\<close>,
-  but so far it is not clear what practical relevance that would have.
-\<close>
-
-lemma rdiv_mult_rmod_eq:
-  \<open>k rdiv l * l + k rmod l = k\<close>
-proof -
-  have *: \<open>l * (sgn l * j) = \<bar>l\<bar> * j\<close> for j
-    by (simp add: ac_simps abs_sgn)
-  show ?thesis
-    by (simp add: rounded_divide_def rounded_modulo_def algebra_simps *)
-qed
-
-lemma mult_rdiv_rmod_eq:
-  \<open>l * (k rdiv l) + k rmod l = k\<close>
-  using rdiv_mult_rmod_eq [of k l] by (simp add: ac_simps)
-
-lemma rmod_rdiv_mult_eq:
-  \<open>k rmod l + k rdiv l * l = k\<close>
-  using rdiv_mult_rmod_eq [of k l] by (simp add: ac_simps)
-
-lemma rmod_mult_rdiv_eq:
-  \<open>k rmod l + l * (k rdiv l) = k\<close>
-  using rdiv_mult_rmod_eq [of k l] by (simp add: ac_simps)
-
-lemma minus_rdiv_mult_eq_rmod:
-  \<open>k - k rdiv l * l = k rmod l\<close>
-  by (rule add_implies_diff [symmetric]) (fact rmod_rdiv_mult_eq)
-
-lemma minus_mult_rdiv_eq_rmod:
-  \<open>k - l * (k rdiv l) = k rmod l\<close>
-  by (rule add_implies_diff [symmetric]) (fact rmod_mult_rdiv_eq)
-
-lemma minus_rmod_eq_rdiv_mult:
-  \<open>k - k rmod l = k rdiv l * l\<close>
-  by (rule add_implies_diff [symmetric]) (fact rdiv_mult_rmod_eq)
-
-lemma minus_rmod_eq_mult_rdiv:
-  \<open>k - k rmod l = l * (k rdiv l)\<close>
-  by (rule add_implies_diff [symmetric]) (fact mult_rdiv_rmod_eq)
-
-lemma rdiv_0_eq [simp]:
-  \<open>k rdiv 0 = 0\<close>
-  by (simp add: rounded_divide_def)
-
-lemma rmod_0_eq [simp]:
-  \<open>k rmod 0 = k\<close>
-  by (simp add: rounded_modulo_def)
-
-lemma rdiv_1_eq [simp]:
-  \<open>k rdiv 1 = k\<close>
-  by (simp add: rounded_divide_def)
-
-lemma rmod_1_eq [simp]:
-  \<open>k rmod 1 = 0\<close>
-  by (simp add: rounded_modulo_def)
-
-lemma zero_rdiv_eq [simp]:
-  \<open>0 rdiv k = 0\<close>
-  by (auto simp add: rounded_divide_def not_less zdiv_eq_0_iff)
-
-lemma zero_rmod_eq [simp]:
-  \<open>0 rmod k = 0\<close>
-  by (auto simp add: rounded_modulo_def not_less zmod_trivial_iff)
-
-lemma rdiv_minus_eq:
-  \<open>k rdiv - l = - (k rdiv l)\<close>
-  by (simp add: rounded_divide_def)
-
-lemma rmod_minus_eq [simp]:
-  \<open>k rmod - l = k rmod l\<close>
-  by (simp add: rounded_modulo_def)
-
-lemma rdiv_abs_eq:
-  \<open>k rdiv \<bar>l\<bar> = sgn l * (k rdiv l)\<close>
-  by (simp add: rounded_divide_def)
-
-lemma rmod_abs_eq [simp]:
-  \<open>k rmod \<bar>l\<bar> = k rmod l\<close>
-  by (simp add: rounded_modulo_def)
-
-lemma nonzero_mult_rdiv_cancel_right:
-  \<open>k * l rdiv l = k\<close> if \<open>l \<noteq> 0\<close>
-proof -
-  have \<open>sgn l * k * \<bar>l\<bar> rdiv l = k\<close>
-    using that by (simp add: rounded_divide_def)
-  with that show ?thesis
-    by (simp add: ac_simps abs_sgn)
-qed
-
-lemma rdiv_self_eq [simp]:
-  \<open>k rdiv k = 1\<close> if \<open>k \<noteq> 0\<close>
-  using that nonzero_mult_rdiv_cancel_right [of k 1] by simp
-
-lemma rmod_self_eq [simp]:
-  \<open>k rmod k = 0\<close>
-proof -
-  have \<open>(sgn k * \<bar>k\<bar> + \<bar>k\<bar> div 2) mod \<bar>k\<bar> = \<bar>k\<bar> div 2\<close>
-    by (auto simp add: zmod_trivial_iff)
-  also have \<open>sgn k * \<bar>k\<bar> = k\<close>
-    by (simp add: abs_sgn)
-  finally show ?thesis
-    by (simp add: rounded_modulo_def algebra_simps)
-qed
-
-lemma rmod_less_divisor:
-  \<open>k rmod l < \<bar>l\<bar> - \<bar>l\<bar> div 2\<close> if \<open>l \<noteq> 0\<close>
-  using that pos_mod_bound [of \<open>\<bar>l\<bar>\<close>] by (simp add: rounded_modulo_def)
-
-lemma rmod_less_equal_divisor:
-  \<open>k rmod l \<le> \<bar>l\<bar> div 2\<close> if \<open>l \<noteq> 0\<close>
-proof -
-  from that rmod_less_divisor [of l k]
-  have \<open>k rmod l < \<bar>l\<bar> - \<bar>l\<bar> div 2\<close>
-    by simp
-  also have \<open>\<bar>l\<bar> - \<bar>l\<bar> div 2 = \<bar>l\<bar> div 2 + of_bool (odd l)\<close>
-    by auto
-  finally show ?thesis
-    by (cases \<open>even l\<close>) simp_all
-qed
-
-lemma divisor_less_equal_rmod':
-  \<open>\<bar>l\<bar> div 2 - \<bar>l\<bar> \<le> k rmod l\<close> if \<open>l \<noteq> 0\<close>
-proof -
-  have \<open>0 \<le> (k + \<bar>l\<bar> div 2) mod \<bar>l\<bar>\<close>
-    using that pos_mod_sign [of \<open>\<bar>l\<bar>\<close>] by simp
-  then show ?thesis
-    by (simp_all add: rounded_modulo_def)
-qed
-
-lemma divisor_less_equal_rmod:
-  \<open>- (\<bar>l\<bar> div 2) \<le> k rmod l\<close> if \<open>l \<noteq> 0\<close>
-  using that divisor_less_equal_rmod' [of l k]
-  by (simp add: rounded_modulo_def)
-
-lemma abs_rmod_less_equal:
-  \<open>\<bar>k rmod l\<bar> \<le> \<bar>l\<bar> div 2\<close> if \<open>l \<noteq> 0\<close>
-  using that divisor_less_equal_rmod [of l k]
-  by (simp add: abs_le_iff rmod_less_equal_divisor)
-
-end
--- a/src/HOL/Library/Sum_of_Squares/positivstellensatz.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/HOL/Library/Sum_of_Squares/positivstellensatz.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -169,7 +169,7 @@
      (Thm.implies_intr (Thm.cprop_of tha) thb);
 
 fun prove_hyp tha thb =
-  if Termset.exists (curry op aconv (Thm.concl_of tha)) (Thm.hyps_of thb)  (* FIXME !? *)
+  if exists (curry op aconv (Thm.concl_of tha)) (Thm.hyps_of thb)  (* FIXME !? *)
   then Thm.equal_elim (Thm.symmetric (deduct_antisym_rule tha thb)) tha else thb;
 
 val pth = @{lemma "(((x::real) < y) \<equiv> (y - x > 0))" and "((x \<le> y) \<equiv> (y - x \<ge> 0))" and
--- a/src/HOL/Library/cconv.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/HOL/Library/cconv.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -132,7 +132,7 @@
                end
              val rule = abstract_rule_thm x
              val inst = Thm.match (hd (Drule.cprems_of rule), mk_concl eq)
-             val gen = (Names.empty, Names.make_set [#1 (dest_Free v)])
+             val gen = (Names.empty, Names.make1_set (#1 (dest_Free v)))
            in
              (Drule.instantiate_normalize inst rule OF [Drule.generalize gen eq])
              |> Drule.zero_var_indexes
--- a/src/HOL/Library/old_recdef.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/HOL/Library/old_recdef.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -308,7 +308,7 @@
       Conv.fconv_rule Drule.beta_eta_conversion
          (case_thm
             |> Thm.instantiate (TVars.make type_cinsts, Vars.empty)
-            |> Thm.instantiate (TVars.empty, Vars.make [(Pv, abs_ct), (Dv, free_ct)]))
+            |> Thm.instantiate (TVars.empty, Vars.make2 (Pv, abs_ct) (Dv, free_ct)))
     end;
 
 
@@ -920,9 +920,8 @@
 fun chyps thm = map Dcterm.drop_prop (Thm.chyps_of thm);
 
 fun dest_thm thm =
-  (map HOLogic.dest_Trueprop (Termset.dest (Thm.hyps_of thm)),
-    HOLogic.dest_Trueprop (Thm.prop_of thm))
-  handle TERM _ => raise RULES_ERR "dest_thm" "missing Trueprop";
+  (map HOLogic.dest_Trueprop (Thm.hyps_of thm), HOLogic.dest_Trueprop (Thm.prop_of thm))
+    handle TERM _ => raise RULES_ERR "dest_thm" "missing Trueprop";
 
 
 (* Inference rules *)
@@ -1102,7 +1101,7 @@
 fun DISJ_CASESL disjth thl =
    let val c = cconcl disjth
        fun eq th atm =
-        Termset.exists (fn t => HOLogic.dest_Trueprop t aconv Thm.term_of atm) (Thm.hyps_of th)
+        exists (fn t => HOLogic.dest_Trueprop t aconv Thm.term_of atm) (Thm.hyps_of th)
        val tml = Dcterm.strip_disj c
        fun DL _ [] = raise RULES_ERR "DISJ_CASESL" "no cases"
          | DL th [th1] = PROVE_HYP th th1
@@ -1125,7 +1124,7 @@
 in
 fun SPEC tm thm =
    let val gspec' =
-    Drule.instantiate_normalize (TVars.make [(TV, Thm.ctyp_of_cterm tm)], Vars.empty) gspec
+    Drule.instantiate_normalize (TVars.make1 (TV, Thm.ctyp_of_cterm tm), Vars.empty) gspec
    in thm RS (Thm.forall_elim tm gspec') end
 end;
 
--- a/src/HOL/Library/refute.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/HOL/Library/refute.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -2375,7 +2375,7 @@
                               val result = fold (fn arg_i => fn i =>
                                 interpretation_apply (i, arg_i)) arg_intrs intr
                               (* update 'REC_OPERATORS' *)
-                              val _ = Array.update (arr, elem, (true, result))
+                              val _ = Array.upd arr elem (true, result)
                             in
                               result
                             end
--- a/src/HOL/Matrix_LP/Compute_Oracle/compute.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/HOL/Matrix_LP/Compute_Oracle/compute.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -16,8 +16,8 @@
     val make : machine -> theory -> thm list -> computer
     val make_with_cache : machine -> theory -> term list -> thm list -> computer
     val theory_of : computer -> theory
-    val hyps_of : computer -> Termset.T
-    val shyps_of : computer -> Sortset.T
+    val hyps_of : computer -> term list
+    val shyps_of : computer -> sort list
     (* ! *) val update : computer -> thm list -> unit
     (* ! *) val update_with_cache : computer -> term list -> thm list -> unit
     
@@ -169,12 +169,13 @@
 fun default_naming i = "v_" ^ string_of_int i
 
 datatype computer = Computer of
-  (theory * Encode.encoding * Termset.T * Sortset.T * prog * unit Unsynchronized.ref * naming)
+  (theory * Encode.encoding * term list * unit Sorttab.table * prog * unit Unsynchronized.ref * naming)
     option Unsynchronized.ref
 
 fun theory_of (Computer (Unsynchronized.ref (SOME (thy,_,_,_,_,_,_)))) = thy
 fun hyps_of (Computer (Unsynchronized.ref (SOME (_,_,hyps,_,_,_,_)))) = hyps
-fun shyps_of (Computer (Unsynchronized.ref (SOME (_,_,_,shypset,_,_,_)))) = shypset
+fun shyps_of (Computer (Unsynchronized.ref (SOME (_,_,_,shyptable,_,_,_)))) = Sorttab.keys (shyptable)
+fun shyptab_of (Computer (Unsynchronized.ref (SOME (_,_,_,shyptable,_,_,_)))) = shyptable
 fun stamp_of (Computer (Unsynchronized.ref (SOME (_,_,_,_,_,stamp,_)))) = stamp
 fun prog_of (Computer (Unsynchronized.ref (SOME (_,_,_,_,prog,_,_)))) = prog
 fun encoding_of (Computer (Unsynchronized.ref (SOME (_,encoding,_,_,_,_,_)))) = encoding
@@ -187,7 +188,7 @@
 fun ref_of (Computer r) = r
 
 
-datatype cthm = ComputeThm of Termset.T * Sortset.T * term
+datatype cthm = ComputeThm of term list * sort list * term
 
 fun thm2cthm th = 
     (if not (null (Thm.tpairs_of th)) then raise Make "theorems may not contain tpairs" else ();
@@ -219,11 +220,11 @@
                     (n, vars, AbstractMachine.PConst (c, args@[pb]))
             end
 
-        fun thm2rule (encoding, hypset, shypset) th =
+        fun thm2rule (encoding, hyptable, shyptable) th =
             let
                 val (ComputeThm (hyps, shyps, prop)) = th
-                val hypset = Termset.merge (hyps, hypset)
-                val shypset = Sortset.merge (shyps, shypset)
+                val hyptable = fold (fn h => Termtab.update (h, ())) hyps hyptable
+                val shyptable = fold (fn sh => Sorttab.update (sh, ())) shyps shyptable
                 val (prems, prop) = (Logic.strip_imp_prems prop, Logic.strip_imp_concl prop)
                 val (a, b) = Logic.dest_equals prop
                   handle TERM _ => raise (Make "theorems must be meta-level equations (with optional guards)")
@@ -269,15 +270,15 @@
                 fun rename_guard (AbstractMachine.Guard (a,b)) = 
                     AbstractMachine.Guard (rename 0 vars a, rename 0 vars b)
             in
-                ((encoding, hypset, shypset), (map rename_guard prems, pattern, rename 0 vars right))
+                ((encoding, hyptable, shyptable), (map rename_guard prems, pattern, rename 0 vars right))
             end
 
-        val ((encoding, hypset, shypset), rules) =
-          fold_rev (fn th => fn (encoding_hypset, rules) =>
+        val ((encoding, hyptable, shyptable), rules) =
+          fold_rev (fn th => fn (encoding_hyptable, rules) =>
             let
-              val (encoding_hypset, rule) = thm2rule encoding_hypset th
-            in (encoding_hypset, rule::rules) end)
-          ths ((encoding, Termset.empty, Sortset.empty), [])
+              val (encoding_hyptable, rule) = thm2rule encoding_hyptable th
+            in (encoding_hyptable, rule::rules) end)
+          ths ((encoding, Termtab.empty, Sorttab.empty), [])
 
         fun make_cache_pattern t (encoding, cache_patterns) =
             let
@@ -287,7 +288,7 @@
                 (encoding, p::cache_patterns)
             end
         
-        val (encoding, _) = Termset.fold_rev make_cache_pattern cache_pattern_terms (encoding, [])
+        val (encoding, _) = fold_rev make_cache_pattern cache_pattern_terms (encoding, [])
 
         val prog = 
             case machine of 
@@ -298,20 +299,19 @@
 
         fun has_witness s = not (null (Sign.witness_sorts thy [] [s]))
 
-        val shypset = Sortset.fold (fn s => has_witness s ? Sortset.remove s) shypset shypset
+        val shyptable = fold Sorttab.delete (filter has_witness (Sorttab.keys (shyptable))) shyptable
 
-    in (thy, encoding, hypset, shypset, prog, stamp, default_naming) end
+    in (thy, encoding, Termtab.keys hyptable, shyptable, prog, stamp, default_naming) end
 
 fun make_with_cache machine thy cache_patterns raw_thms =
-  Computer (Unsynchronized.ref
-    (SOME (make_internal machine thy (Unsynchronized.ref ()) Encode.empty (Termset.make cache_patterns) raw_thms)))
+  Computer (Unsynchronized.ref (SOME (make_internal machine thy (Unsynchronized.ref ()) Encode.empty cache_patterns raw_thms)))
 
 fun make machine thy raw_thms = make_with_cache machine thy [] raw_thms
 
 fun update_with_cache computer cache_patterns raw_thms =
     let 
         val c = make_internal (machine_of_prog (prog_of computer)) (theory_of computer) (stamp_of computer) 
-                              (encoding_of computer) (Termset.make cache_patterns) raw_thms
+                              (encoding_of computer) cache_patterns raw_thms
         val _ = (ref_of computer) := SOME c     
     in
         ()
@@ -328,19 +328,30 @@
 (* An oracle for exporting theorems; must only be accessible from inside this structure! *)
 (* ------------------------------------------------------------------------------------- *)
 
+fun merge_hyps hyps1 hyps2 = 
+let
+    fun add hyps tab = fold (fn h => fn tab => Termtab.update (h, ()) tab) hyps tab
+in
+    Termtab.keys (add hyps2 (add hyps1 Termtab.empty))
+end
+
+fun add_shyps shyps tab = fold (fn h => fn tab => Sorttab.update (h, ()) tab) shyps tab
+
+fun merge_shyps shyps1 shyps2 = Sorttab.keys (add_shyps shyps2 (add_shyps shyps1 Sorttab.empty))
+
 val (_, export_oracle) = Context.>>> (Context.map_theory_result
   (Thm.add_oracle (\<^binding>\<open>compute\<close>, fn (thy, hyps, shyps, prop) =>
     let
-        fun remove_term t = Sortset.subtract (Sortset.build (Sortset.insert_term t))
+        val shyptab = add_shyps shyps Sorttab.empty
+        fun delete s shyptab = Sorttab.delete s shyptab handle Sorttab.UNDEF _ => shyptab
+        fun delete_term t shyptab = fold delete (Sorts.insert_term t []) shyptab
         fun has_witness s = not (null (Sign.witness_sorts thy [] [s]))
-        val shyps = Sortset.fold (fn s => has_witness s ? Sortset.remove s) shyps shyps
-        val shyps =
-          if Sortset.is_empty shyps then Sortset.empty
-          else fold remove_term (prop::hyps) shyps
+        val shyptab = fold Sorttab.delete (filter has_witness (Sorttab.keys (shyptab))) shyptab
+        val shyps = if Sorttab.is_empty shyptab then [] else Sorttab.keys (fold delete_term (prop::hyps) shyptab)
         val _ =
-          if not (Sortset.is_empty shyps) then
+          if not (null shyps) then
             raise Compute ("dangling sort hypotheses: " ^
-              commas (map (Syntax.string_of_sort_global thy) (Sortset.dest shyps)))
+              commas (map (Syntax.string_of_sort_global thy) shyps))
           else ()
     in
         Thm.global_cterm_of thy (fold_rev (fn hyp => fn p => Logic.mk_implies (hyp, p)) hyps prop)
@@ -368,7 +379,7 @@
         val t = infer_types naming encoding ty t
         val eq = Logic.mk_equals (t', t)
     in
-        export_thm thy (Termset.dest (hyps_of computer)) (shyps_of computer) eq
+        export_thm thy (hyps_of computer) (Sorttab.keys (shyptab_of computer)) eq
     end
 
 (* --------- Simplify ------------ *)
@@ -376,7 +387,7 @@
 datatype prem = EqPrem of AbstractMachine.term * AbstractMachine.term * Term.typ * int 
               | Prem of AbstractMachine.term
 datatype theorem = Theorem of theory * unit Unsynchronized.ref * (int * typ) Symtab.table * (AbstractMachine.term option) Inttab.table  
-               * prem list * AbstractMachine.term * Termset.T * Sortset.T
+               * prem list * AbstractMachine.term * term list * sort list
 
 
 exception ParamSimplify of computer * theorem
@@ -607,8 +618,8 @@
         let
             val th = update_varsubst varsubst th
             val th = update_prems (splicein prem_no (prems_of_theorem th') prems) th
-            val th = update_hyps (Termset.merge (hyps_of_theorem th, hyps_of_theorem th')) th
-            val th = update_shyps (Sortset.merge (shyps_of_theorem th, shyps_of_theorem th')) th
+            val th = update_hyps (merge_hyps (hyps_of_theorem th) (hyps_of_theorem th')) th
+            val th = update_shyps (merge_shyps (shyps_of_theorem th) (shyps_of_theorem th')) th
         in
             update_theory thy th
         end
@@ -624,10 +635,11 @@
     fun run t = infer (runprog (prog_of computer) (apply_subst true varsubst t))
     fun runprem p = run (prem2term p)
     val prop = Logic.list_implies (map runprem (prems_of_theorem th), run (concl_of_theorem th))
-    val hyps = Termset.merge (hyps_of computer, hyps_of_theorem th)
-    val shyps = Sortset.merge (shyps_of_theorem th, shyps_of computer)
+    val hyps = merge_hyps (hyps_of computer) (hyps_of_theorem th)
+    val shyps = merge_shyps (shyps_of_theorem th) (Sorttab.keys (shyptab_of computer))
 in
-    export_thm (theory_of_theorem th) (Termset.dest hyps) shyps prop
+    export_thm (theory_of_theorem th) hyps shyps prop
 end
 
 end
+
--- a/src/HOL/Matrix_LP/Compute_Oracle/linker.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/HOL/Matrix_LP/Compute_Oracle/linker.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -365,12 +365,12 @@
         (!changed, result)
     end
 
-datatype cthm = ComputeThm of Termset.T * Sortset.T * term
+datatype cthm = ComputeThm of term list * sort list * term
 
 fun thm2cthm th = ComputeThm (Thm.hyps_of th, Thm.shyps_of th, Thm.prop_of th)
 
 val cthm_ord' =
-  prod_ord (prod_ord Termset.ord Sortset.ord) Term_Ord.term_ord
+  prod_ord (prod_ord (list_ord Term_Ord.term_ord) (list_ord Term_Ord.sort_ord)) Term_Ord.term_ord
 
 fun cthm_ord (ComputeThm (h1, sh1, p1), ComputeThm (h2, sh2, p2)) =
   cthm_ord' (((h1,sh1), p1), ((h2, sh2), p2))
--- a/src/HOL/Mutabelle/mutabelle_extra.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/HOL/Mutabelle/mutabelle_extra.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -239,7 +239,7 @@
   let val consts = Term.add_const_names (Thm.prop_of th) [] in
     exists (member (op =) (Long_Name.explode s)) forbidden_thms orelse
     exists (member (op =) forbidden_consts) consts orelse
-    length (Long_Name.explode s) <> 2 orelse
+    Long_Name.count s <> 2 orelse
     String.isPrefix "type_definition" (List.last (Long_Name.explode s)) orelse
     String.isSuffix "_def" s orelse
     String.isSuffix "_raw" s orelse
@@ -331,7 +331,7 @@
 
 fun thms_of all thy =
   filter
-    (fn th => (all orelse Thm.theory_name th = Context.theory_name thy)
+    (fn th => (all orelse Thm.theory_long_name th = Context.theory_long_name thy)
       (* andalso is_executable_thm thy th *))
     (map snd (filter_out is_forbidden_theorem (Global_Theory.all_thms_of thy false)))
 
--- a/src/HOL/Nitpick_Examples/Mono_Nits.thy	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/HOL/Nitpick_Examples/Mono_Nits.thy	Mon Apr 24 15:00:54 2023 +0100
@@ -144,7 +144,7 @@
 val mono_timeout = seconds 1.0
 
 fun is_forbidden_theorem name =
-  length (Long_Name.explode name) <> 2 orelse
+  Long_Name.count name <> 2 orelse
   String.isPrefix "type_definition" (List.last (Long_Name.explode name)) orelse
   String.isPrefix "arity_" (List.last (Long_Name.explode name)) orelse
   String.isSuffix "_def" name orelse
@@ -153,7 +153,7 @@
 fun theorems_of thy =
   filter (fn (name, th) =>
              not (is_forbidden_theorem name) andalso
-             Thm.theory_name th = Context.theory_name thy)
+             Thm.theory_long_name th = Context.theory_long_name thy)
          (Global_Theory.all_thms_of thy true)
 
 fun check_formulas tsp =
@@ -175,7 +175,7 @@
 
 fun check_theory thy =
   let
-    val path = File.tmp_path (Context.theory_name thy ^ ".out" |> Path.explode)
+    val path = File.tmp_path (Context.theory_base_name thy ^ ".out" |> Path.explode)
     val _ = File.write path ""
     fun check_theorem (name, th) =
       let
--- a/src/HOL/Nominal/nominal_datatype.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/HOL/Nominal/nominal_datatype.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -1571,8 +1571,8 @@
               (Logic.mk_implies (HOLogic.mk_Trueprop Q, HOLogic.mk_Trueprop P)))
             (fn {context = ctxt, ...} =>
               dresolve_tac ctxt [Thm.instantiate (TVars.empty,
-                 Vars.make [((("pi", 0), permT),
-                   Thm.global_cterm_of thy11 (Const (\<^const_name>\<open>rev\<close>, permT --> permT) $ pi))]) th] 1 THEN
+                 Vars.make1 ((("pi", 0), permT),
+                   Thm.global_cterm_of thy11 (Const (\<^const_name>\<open>rev\<close>, permT --> permT) $ pi))) th] 1 THEN
                NominalPermeq.perm_simp_tac (put_simpset HOL_ss ctxt) 1)) (ps ~~ ths)
       in (ths, ths') end) dt_atomTs);
 
@@ -1662,9 +1662,9 @@
                     SUBPROOF (fn {context = ctxt', prems = prems', params = [(_, a), (_, b)], ...} =>
                       EVERY [cut_facts_tac [rec_prem] 1,
                        resolve_tac ctxt' [Thm.instantiate (TVars.empty,
-                         Vars.make [((("pi", 0), mk_permT aT),
+                         Vars.make1 ((("pi", 0), mk_permT aT),
                            Thm.cterm_of ctxt'
-                            (perm_of_pair (Thm.term_of a, Thm.term_of b)))]) eqvt_th] 1,
+                            (perm_of_pair (Thm.term_of a, Thm.term_of b)))) eqvt_th] 1,
                        asm_simp_tac (put_simpset HOL_ss ctxt' addsimps
                          (prems' @ perm_swap @ perm_fresh_fresh)) 1]) ctxt 1,
                     resolve_tac ctxt [rec_prem] 1,
@@ -1882,7 +1882,7 @@
                             val l = find_index (equal T) dt_atomTs;
                             val th = nth (nth rec_equiv_thms' l) k;
                             val th' = Thm.instantiate (TVars.empty,
-                              Vars.make [((("pi", 0), U), Thm.global_cterm_of thy11 p)]) th;
+                              Vars.make1 ((("pi", 0), U), Thm.global_cterm_of thy11 p)) th;
                           in resolve_tac ctxt [th'] 1 end;
                         val th' = Goal.prove context'' [] []
                           (HOLogic.mk_Trueprop (S $ mk_pi x $ mk_pi y))
--- a/src/HOL/Nominal/nominal_fresh_fun.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/HOL/Nominal/nominal_fresh_fun.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -153,7 +153,7 @@
      [Var v] =>
       Seq.single
         (Thm.instantiate (TVars.empty,
-          Vars.make [(v, Thm.cterm_of ctxt (fold_rev Term.abs params (Bound 0)))]) st)
+          Vars.make1 (v, Thm.cterm_of ctxt (fold_rev Term.abs params (Bound 0)))) st)
     | _ => error "fresh_fun_simp: Too many variables, please report."
    end
   in
--- a/src/HOL/Real_Asymp/multiseries_expansion.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/HOL/Real_Asymp/multiseries_expansion.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -218,7 +218,7 @@
         case tvs of
           [v] =>
             let
-              val thm' = Thm.instantiate (TVars.make [(v, Thm.ctyp_of_cterm ct)], Vars.empty) thm
+              val thm' = Thm.instantiate (TVars.make1 (v, Thm.ctyp_of_cterm ct), Vars.empty) thm
               val vs = take (length cts) (rev (Term.add_vars (Thm.concl_of thm') []))
             in
               Thm.instantiate (TVars.empty, Vars.make (vs ~~ cts)) thm'
--- a/src/HOL/Statespace/distinct_tree_prover.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/HOL/Statespace/distinct_tree_prover.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -310,8 +310,8 @@
         val [alphaI] = #2 (dest_Type T);
       in
         Thm.instantiate
-          (TVars.make [(alpha, Thm.ctyp_of ctxt alphaI)],
-           Vars.make [((v, treeT alphaI), ct')]) @{thm subtract_Tip}
+          (TVars.make1 (alpha, Thm.ctyp_of ctxt alphaI),
+           Vars.make1 ((v, treeT alphaI), ct')) @{thm subtract_Tip}
       end
   | subtractProver ctxt (Const (\<^const_name>\<open>Node\<close>, nT) $ l $ x $ d $ r) ct dist_thm =
       let
--- a/src/HOL/TPTP/TPTP_Proof_Reconstruction.thy	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/HOL/TPTP/TPTP_Proof_Reconstruction.thy	Mon Apr 24 15:00:54 2023 +0100
@@ -943,7 +943,7 @@
      end
 
     fun instantiate_tac from to =
-      PRIMITIVE (Thm.instantiate (TVars.empty, Vars.make [(from, to)]))
+      PRIMITIVE (Thm.instantiate (TVars.empty, Vars.make1 (from, to)))
 
     val tactic =
       if is_none var_opt then no_tac
--- a/src/HOL/TPTP/mash_export.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/HOL/TPTP/mash_export.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -58,7 +58,7 @@
   | _ => ("", []))
 
 fun has_thm_thy th thy =
-  Context.theory_name thy = Thm.theory_name th
+  Context.theory_base_name thy = Thm.theory_base_name th
 
 fun has_thys thys th = exists (has_thm_thy th) thys
 
@@ -98,7 +98,7 @@
     fun do_fact ((_, stature), th) =
       let
         val name = nickname_of_thm th
-        val feats = features_of ctxt (Thm.theory_name th) stature [Thm.prop_of th]
+        val feats = features_of ctxt (Thm.theory_base_name th) stature [Thm.prop_of th]
         val s = encode_str name ^ ": " ^ encode_strs (sort string_ord feats) ^ "\n"
       in
         File.append path s
@@ -187,14 +187,14 @@
           val _ = tracing ("Fact " ^ string_of_int j ^ ": " ^ name)
           val isar_deps = isar_dependencies_of name_tabs th
           val do_query = not (is_bad_query ctxt step j th isar_deps)
-          val goal_feats = features_of ctxt (Thm.theory_name th) stature [Thm.prop_of th]
+          val goal_feats = features_of ctxt (Thm.theory_base_name th) stature [Thm.prop_of th]
           val access_facts = filter_accessible_from th new_facts @ old_facts
           val (marker, deps) =
             smart_dependencies_of ctxt params_opt access_facts name_tabs th isar_deps
 
           fun extra_features_of (((_, stature), th), weight) =
             [Thm.prop_of th]
-            |> features_of ctxt (Thm.theory_name th) stature
+            |> features_of ctxt (Thm.theory_base_name th) stature
             |> map (rpair (weight * extra_feature_factor))
 
           val query =
@@ -261,7 +261,7 @@
               val suggs =
                 old_facts
                 |> filter_accessible_from th
-                |> mepo_or_mash_suggested_facts ctxt (Thm.theory_name th)
+                |> mepo_or_mash_suggested_facts ctxt (Thm.theory_base_name th)
                   params max_suggs hyp_ts concl_t
                 |> map (nickname_of_thm o snd)
             in
--- a/src/HOL/Tools/Argo/argo_real.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/HOL/Tools/Argo/argo_real.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -208,7 +208,7 @@
 fun add_cmp_conv ctxt n thm =
   let val v = var_of_add_cmp (Thm.prop_of thm) in
     Conv.rewr_conv
-      (Thm.instantiate (TVars.empty, Vars.make [(v, Thm.cterm_of ctxt (mk_number n))]) thm)
+      (Thm.instantiate (TVars.empty, Vars.make1 (v, Thm.cterm_of ctxt (mk_number n))) thm)
   end
 
 fun mul_cmp_conv ctxt n pos_thm neg_thm =
--- a/src/HOL/Tools/Argo/argo_tactic.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/HOL/Tools/Argo/argo_tactic.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -656,7 +656,7 @@
 
 (* normalizing goals *)
 
-fun instantiate v ct = Thm.instantiate (TVars.empty, Vars.make [(v, ct)])
+fun instantiate v ct = Thm.instantiate (TVars.empty, Vars.make1 (v, ct))
 
 fun instantiate_elim_rule thm =
   let
--- a/src/HOL/Tools/BNF/bnf_fp_def_sugar.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/HOL/Tools/BNF/bnf_fp_def_sugar.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -851,7 +851,7 @@
                 |> unfold_thms lthy [dtor_ctor];
             in
               (fp_map_thm' RS ctor_cong RS (ctor_dtor RS sym RS trans))
-              |> Drule.generalize (Names.empty, Names.make_set [y_s])
+              |> Drule.generalize (Names.empty, Names.make1_set y_s)
             end;
 
         val map_thms =
@@ -930,7 +930,7 @@
               |> infer_instantiate' lthy (replicate live NONE @
                 [SOME (Thm.cterm_of lthy (ctorA $ y)), SOME (Thm.cterm_of lthy (ctorB $ z))])
               |> unfold_thms lthy [dtor_ctor]
-              |> Drule.generalize (Names.empty, Names.make_set [y_s, z_s])
+              |> Drule.generalize (Names.empty, Names.make2_set y_s z_s)
             end;
 
         val rel_inject_thms =
--- a/src/HOL/Tools/BNF/bnf_gfp_rec_sugar_tactics.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/HOL/Tools/BNF/bnf_gfp_rec_sugar_tactics.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -160,7 +160,7 @@
       val eq = Abs (Name.uu, T, HOLogic.mk_eq (Free (s, T), Bound 0));
     in
       Thm.instantiate' [] [SOME (Thm.cterm_of ctxt eq)] split
-      |> Drule.generalize (Names.empty, Names.make_set [s])
+      |> Drule.generalize (Names.empty, Names.make1_set s)
     end
   | _ => split);
 
--- a/src/HOL/Tools/Function/induction_schema.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/HOL/Tools/Function/induction_schema.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -381,7 +381,7 @@
 
     val res = Conjunction.intr_balanced (map_index project branches)
       |> fold_rev Thm.implies_intr (map Thm.cprop_of newgoals @ steps)
-      |> Drule.generalize (Names.empty, Names.make_set [Rn])
+      |> Drule.generalize (Names.empty, Names.make1_set Rn)
 
     val nbranches = length branches
     val npres = length pres
--- a/src/HOL/Tools/Function/partial_function.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/HOL/Tools/Function/partial_function.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -209,7 +209,7 @@
          THEN Simplifier.simp_tac (put_simpset curry_uncurry_ss ctxt) 3
          THEN CONVERSION (split_params_conv ctxt
            then_conv (Conv.forall_conv (K split_paired_all_conv) ctxt)) 3)
-      |> Thm.instantiate (TVars.empty, Vars.make [(P_var, P_inst), (x_var, x_inst)])
+      |> Thm.instantiate (TVars.empty, Vars.make2 (P_var, P_inst) (x_var, x_inst))
       |> Simplifier.full_simplify (put_simpset split_conv_ss ctxt)
       |> singleton (Variable.export ctxt' ctxt)
   in
--- a/src/HOL/Tools/Function/relation.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/HOL/Tools/Function/relation.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -19,7 +19,7 @@
   SUBGOAL (fn (goal, _) =>
     (case Term.add_vars goal [] of
       [v as (_, T)] =>
-        PRIMITIVE (Thm.instantiate (TVars.empty, Vars.make [(v, Thm.cterm_of ctxt (inst T))]))
+        PRIMITIVE (Thm.instantiate (TVars.empty, Vars.make1 (v, Thm.cterm_of ctxt (inst T))))
     | _ => no_tac));
 
 fun relation_tac ctxt rel i =
@@ -39,7 +39,7 @@
             |> map_types Type_Infer.paramify_vars
             |> Type.constraint T
             |> Syntax.check_term ctxt;
-        in PRIMITIVE (Thm.instantiate (TVars.empty, Vars.make [(v, Thm.cterm_of ctxt rel')])) end
+        in PRIMITIVE (Thm.instantiate (TVars.empty, Vars.make1 (v, Thm.cterm_of ctxt rel'))) end
     | _ => no_tac));
 
 fun relation_infer_tac ctxt rel i =
--- a/src/HOL/Tools/Meson/meson.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/HOL/Tools/Meson/meson.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -358,7 +358,7 @@
         Variable.variant_fixes [name_of (HOLogic.dest_Trueprop (Thm.concl_of th))] ctxt;
       val spec' = spec
         |> Thm.instantiate
-          (TVars.empty, Vars.make [(spec_var, Thm.cterm_of ctxt' (Free (x, snd spec_var)))]);
+          (TVars.empty, Vars.make1 (spec_var, Thm.cterm_of ctxt' (Free (x, snd spec_var))));
     in (th RS spec', ctxt') end
 end;
 
--- a/src/HOL/Tools/Meson/meson_clausify.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/HOL/Tools/Meson/meson_clausify.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -42,9 +42,9 @@
 fun transform_elim_theorem th =
   (case Thm.concl_of th of    (*conclusion variable*)
     \<^Const_>\<open>Trueprop for \<open>Var (v as (_, \<^Type>\<open>bool\<close>))\<close>\<close> =>
-      Thm.instantiate (TVars.empty, Vars.make [(v, cfalse)]) th
+      Thm.instantiate (TVars.empty, Vars.make1 (v, cfalse)) th
   | Var (v as (_, \<^Type>\<open>prop\<close>)) =>
-      Thm.instantiate (TVars.empty, Vars.make [(v, ctp_false)]) th
+      Thm.instantiate (TVars.empty, Vars.make1 (v, ctp_false)) th
   | _ => th)
 
 
--- a/src/HOL/Tools/Mirabelle/mirabelle.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/HOL/Tools/Mirabelle/mirabelle.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -302,7 +302,7 @@
 
 fun fold_body_thms f =
   let
-    fun app n (PBody {thms, ...}) = thms |> PThms.fold (fn (i, thm_node) =>
+    fun app n (PBody {thms, ...}) = thms |> fold (fn (i, thm_node) =>
       fn (x, seen) =>
         if Inttab.defined seen i then (x, seen)
         else
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_aux.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_aux.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -943,7 +943,7 @@
       Thm.instantiate
         (TVars.make [(dest_TVar uninst_T, Thm.ctyp_of ctxt' (U --> T')),
           (dest_TVar uninst_T', Thm.ctyp_of ctxt' T')],
-         Vars.make [((fst (dest_Var f), (U --> T') --> T'), Thm.cterm_of ctxt' f')]) th
+         Vars.make1 ((fst (dest_Var f), (U --> T') --> T'), Thm.cterm_of ctxt' f')) th
     val [th'] = Variable.export (Variable.declare_thm th' ctxt') ctxt [th']
   in
     th'
--- a/src/HOL/Tools/Qelim/cooper.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/HOL/Tools/Qelim/cooper.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -507,17 +507,17 @@
    then
     (bl,b0,decomp_minf,
      fn B => (map (fn th =>
-       Thm.implies_elim (Thm.instantiate (TVars.empty, Vars.make [(B_v,B), (D_v,cd)]) th) dp)
+       Thm.implies_elim (Thm.instantiate (TVars.empty, Vars.make2 (B_v,B) (D_v,cd)) th) dp)
                      [bseteq,bsetneq,bsetlt, bsetle, bsetgt,bsetge])@
-                   (map (Thm.instantiate (TVars.empty, Vars.make [(B_v,B), (D_v,cd)]))
+                   (map (Thm.instantiate (TVars.empty, Vars.make2 (B_v,B) (D_v,cd)))
                         [bsetdvd,bsetndvd,bsetP,infDdvd, infDndvd,bsetconj,
                          bsetdisj,infDconj, infDdisj]),
                        cpmi)
      else (al,a0,decomp_pinf,fn A =>
           (map (fn th =>
-            Thm.implies_elim (Thm.instantiate (TVars.empty, Vars.make [(A_v,A), (D_v,cd)]) th) dp)
+            Thm.implies_elim (Thm.instantiate (TVars.empty, Vars.make2 (A_v,A) (D_v,cd)) th) dp)
                    [aseteq,asetneq,asetlt, asetle, asetgt,asetge])@
-                   (map (Thm.instantiate (TVars.empty, Vars.make [(A_v,A), (D_v,cd)]))
+                   (map (Thm.instantiate (TVars.empty, Vars.make2 (A_v,A) (D_v,cd)))
                    [asetdvd,asetndvd, asetP, infDdvd, infDndvd,asetconj,
                          asetdisj,infDconj, infDdisj]),cppi)
  val cpth =
--- a/src/HOL/Tools/Quickcheck/find_unused_assms.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/HOL/Tools/Quickcheck/find_unused_assms.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -16,7 +16,7 @@
 
 fun thms_of thy thy_name =
   Global_Theory.all_thms_of thy false
-  |> filter (fn (_, th) => Thm.theory_name th = thy_name)
+  |> filter (fn (_, th) => Thm.theory_base_name th = thy_name)
 
 fun do_while P f s list =
   if P s then
@@ -71,7 +71,7 @@
     val thy = Proof_Context.theory_of ctxt
     val all_thms =
       thms_of thy thy_name
-      |> filter (fn (name, _) => length (Long_Name.explode name) = 2)  (* FIXME !? *)
+      |> filter (fn (name, _) => Long_Name.count name = 2)  (* FIXME !? *)
       |> sort_by #1
     val check = check_unused_assms ctxt
   in rev (Par_List.map check all_thms) end
@@ -95,7 +95,8 @@
 
 fun print_unused_assms ctxt opt_thy_name =
   let
-    val thy_name = the_default (Context.theory_name (Proof_Context.theory_of ctxt)) opt_thy_name
+    val thy_name =
+      the_default (Context.theory_base_name (Proof_Context.theory_of ctxt)) opt_thy_name
     val results = find_unused_assms ctxt thy_name
     val total = length results
     val with_assumptions = length (filter (is_some o snd) results)
--- a/src/HOL/Tools/Quickcheck/random_generators.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/HOL/Tools/Quickcheck/random_generators.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -92,7 +92,7 @@
       (HOLogic.dest_eq o HOLogic.dest_Trueprop) eq;
     val Type (_, [_, iT]) = T;
     val icT = Thm.ctyp_of lthy iT;
-    val inst = Thm.instantiate_cterm (TVars.make [(a_v, icT)], Vars.empty);
+    val inst = Thm.instantiate_cterm (TVars.make1 (a_v, icT), Vars.empty);
     fun subst_v t' = map_aterms (fn t as Free (w, _) => if v = w then t' else t | t => t);
     val t_rhs = lambda t_k proto_t_rhs;
     val eqs0 = [subst_v \<^term>\<open>0::natural\<close> eq,
@@ -105,9 +105,9 @@
     val cT_rhs = inst pt_rhs |> Thm.term_of |> dest_Var;
     val rule = @{thm random_aux_rec}
       |> Drule.instantiate_normalize
-        (TVars.make [(a_v, icT)],
-          Vars.make [(cT_random_aux, Thm.cterm_of lthy' t_random_aux),
-           (cT_rhs, Thm.cterm_of lthy' t_rhs)]);
+        (TVars.make1 (a_v, icT),
+          Vars.make2 (cT_random_aux, Thm.cterm_of lthy' t_random_aux)
+           (cT_rhs, Thm.cterm_of lthy' t_rhs));
     fun tac ctxt =
       ALLGOALS (resolve_tac ctxt [rule])
       THEN ALLGOALS (simp_tac (put_simpset rew_ss ctxt))
--- a/src/HOL/Tools/Quotient/quotient_tacs.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/HOL/Tools/Quotient/quotient_tacs.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -477,7 +477,7 @@
           else make_inst (Thm.term_of (Thm.lhs_of thm3)) (Thm.term_of ctrm)
         val thm4 =
           Drule.instantiate_normalize
-            (TVars.empty, Vars.make [(dest_Var insp, Thm.cterm_of ctxt inst)]) thm3
+            (TVars.empty, Vars.make1 (dest_Var insp, Thm.cterm_of ctxt inst)) thm3
       in
         Conv.rewr_conv thm4 ctrm
       end
--- a/src/HOL/Tools/SMT/smt_normalize.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/HOL/Tools/SMT/smt_normalize.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -35,7 +35,7 @@
 
   fun inst f ct thm =
     let val cv = f (Drule.strip_imp_concl (Thm.cprop_of thm))
-    in Thm.instantiate (TVars.empty, Vars.make [(dest_Var (Thm.term_of cv), ct)]) thm end
+    in Thm.instantiate (TVars.empty, Vars.make1 (dest_Var (Thm.term_of cv), ct)) thm end
 in
 
 fun instantiate_elim thm =
@@ -401,7 +401,7 @@
   let val (q, cts) = fold (add_apps add_int_of_nat [] o Thm.cprop_of) thms (false, [])
   in
     if q then (thms, nat_int_thms)
-    else (thms, map (fn ct => Thm.instantiate (TVars.empty, Vars.make [(var, ct)]) int_thm) cts)
+    else (thms, map (fn ct => Thm.instantiate (TVars.empty, Vars.make1 (var, ct)) int_thm) cts)
   end
 
 val setup_nat_as_int =
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_fact.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_fact.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -484,7 +484,7 @@
           let
             val n = length ths
             val collection = n > 1
-            val dotted_name = length (Long_Name.explode name0) > 2 (* ignore theory name *)
+            val dotted_name = Long_Name.count name0 > 2 (* ignore theory name *)
 
             fun check_thms a =
               (case try (Proof_Context.get_thms ctxt) a of
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_mash.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_mash.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -796,7 +796,7 @@
       (* There must be a better way to detect local facts. *)
       (case Long_Name.dest_local hint of
         SOME suf =>
-        Long_Name.implode [Thm.theory_name th, suf, crude_printed_term 25 (Thm.prop_of th)]
+        Long_Name.implode [Thm.theory_base_name th, suf, crude_printed_term 25 (Thm.prop_of th)]
       | NONE => hint)
     end
   else
@@ -820,9 +820,9 @@
 fun crude_thm_ord ctxt =
   let
     val ancestor_lengths =
-      fold (fn thy => Symtab.update (Context.theory_name thy, length (Context.ancestors_of thy)))
+      fold (fn thy => Symtab.update (Context.theory_base_name thy, length (Context.ancestors_of thy)))
         (Theory.nodes_of (Proof_Context.theory_of ctxt)) Symtab.empty
-    val ancestor_length = Symtab.lookup ancestor_lengths o Context.theory_id_name
+    val ancestor_length = Symtab.lookup ancestor_lengths o Context.theory_id_name {long = false}
 
     fun crude_theory_ord p =
       if Context.eq_thy_id p then EQUAL
@@ -832,9 +832,9 @@
         (case apply2 ancestor_length p of
           (SOME m, SOME n) =>
             (case int_ord (m, n) of
-              EQUAL => string_ord (apply2 Context.theory_id_name p)
+              EQUAL => string_ord (apply2 (Context.theory_id_name {long = false}) p)
             | ord => ord)
-        | _ => string_ord (apply2 Context.theory_id_name p))
+        | _ => string_ord (apply2 (Context.theory_id_name {long = false}) p))
   in
     fn p =>
       (case crude_theory_ord (apply2 Thm.theory_id p) of
@@ -1125,7 +1125,7 @@
             val chunks = app_hd (cons th) chunks
             val chunks_and_parents' =
               if thm_less_eq (th, th') andalso
-                Thm.theory_name th = Thm.theory_name th'
+                Thm.theory_base_name th = Thm.theory_base_name th'
               then (chunks, [nickname_of_thm th])
               else chunks_and_parents_for chunks th'
           in
@@ -1172,11 +1172,11 @@
 
     val chained = filter (fn ((_, (scope, _)), _) => scope = Chained) facts
 
-    fun fact_has_right_theory (_, th) = thy_name = Thm.theory_name th
+    fun fact_has_right_theory (_, th) = thy_name = Thm.theory_base_name th
 
     fun chained_or_extra_features_of factor (((_, stature), th), weight) =
       [Thm.prop_of th]
-      |> features_of ctxt (Thm.theory_name th) stature
+      |> features_of ctxt (Thm.theory_base_name th) stature
       |> map (rpair (weight * factor))
   in
     (case get_state ctxt of
@@ -1283,7 +1283,7 @@
     launch_thread timeout (fn () =>
       let
         val thy = Proof_Context.theory_of ctxt
-        val feats = features_of ctxt (Context.theory_name thy) (Local, General) [t]
+        val feats = features_of ctxt (Context.theory_base_name thy) (Local, General) [t]
       in
         map_state ctxt
           (fn {access_G, xtabs as ((num_facts0, _), _), ffds, freqs, dirty_facts} =>
@@ -1395,7 +1395,7 @@
                   (learns, (num_nontrivial, next_commit, _)) =
                 let
                   val name = nickname_of_thm th
-                  val feats = features_of ctxt (Thm.theory_name th) stature [Thm.prop_of th]
+                  val feats = features_of ctxt (Thm.theory_base_name th) stature [Thm.prop_of th]
                   val deps = these (deps_of status th)
                   val num_nontrivial = num_nontrivial |> not (null deps) ? Integer.add 1
                   val learns = (name, parents, feats, deps) :: learns
@@ -1544,7 +1544,7 @@
     [("", [])]
   else
     let
-      val thy_name = Context.theory_name (Proof_Context.theory_of ctxt)
+      val thy_name = Context.theory_base_name (Proof_Context.theory_of ctxt)
 
       fun maybe_launch_thread exact min_num_facts_to_learn =
         if not (Async_Manager_Legacy.has_running_threads MaShN) andalso
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_mepo.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_mepo.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -220,7 +220,7 @@
     t
 
 fun theory_const_prop_of fudge th =
-  theory_constify fudge (Thm.theory_name th) (Thm.prop_of th)
+  theory_constify fudge (Thm.theory_base_name th) (Thm.prop_of th)
 
 fun pair_consts_fact thy fudge fact =
   (case fact |> snd |> theory_const_prop_of fudge |> pconsts_in_fact thy of
@@ -546,7 +546,7 @@
        []
      else
        relevance_filter ctxt thres0 decay max_facts fudge facts hyp_ts
-         (concl_t |> theory_constify fudge (Context.theory_name thy)))
+         (concl_t |> theory_constify fudge (Context.theory_base_name thy)))
     |> map fact_of_lazy_fact
   end
 
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_util.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_util.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -104,7 +104,7 @@
             else (num_thms + 1, name' :: names)
           | NONE => accum)
       end
-    and app_body map_name (PBody {thms, ...}) = PThms.fold (app_thm map_name) thms
+    and app_body map_name (PBody {thms, ...}) = fold (app_thm map_name) thms
   in
     snd (app_body map_plain_name body (0, []))
   end
--- a/src/HOL/Tools/numeral.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/HOL/Tools/numeral.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -64,7 +64,7 @@
 val uminus_tvar = tvar \<^sort>\<open>uminus\<close>;
 val uminus = cterm_of (Const (\<^const_name>\<open>uminus\<close>, TVar uminus_tvar --> TVar uminus_tvar));
 
-fun instT T v = Thm.instantiate_cterm (TVars.make [(v, T)], Vars.empty);
+fun instT T v = Thm.instantiate_cterm (TVars.make1 (v, T), Vars.empty);
 
 in
 
--- a/src/HOL/Tools/numeral_simprocs.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/HOL/Tools/numeral_simprocs.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -591,8 +591,8 @@
 
 fun prove_nz ctxt T t =
   let
-    val z = Thm.instantiate_cterm (TVars.make [(zero_tvar, T)], Vars.empty) zero
-    val eq = Thm.instantiate_cterm (TVars.make [(type_tvar, T)], Vars.empty) geq
+    val z = Thm.instantiate_cterm (TVars.make1 (zero_tvar, T), Vars.empty) zero
+    val eq = Thm.instantiate_cterm (TVars.make1 (type_tvar, T), Vars.empty) geq
     val th =
       Simplifier.rewrite (ctxt addsimps @{thms simp_thms})
         (Thm.apply \<^cterm>\<open>Trueprop\<close> (Thm.apply \<^cterm>\<open>Not\<close>
--- a/src/HOL/Tools/record.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/HOL/Tools/record.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -1493,7 +1493,7 @@
 
 (* attributes *)
 
-fun case_names_fields x = Rule_Cases.case_names ["fields"] x;
+val case_names_fields = Rule_Cases.case_names ["fields"];
 fun induct_type_global name = [case_names_fields, Induct.induct_type name];
 fun cases_type_global name = [case_names_fields, Induct.cases_type name];
 
--- a/src/HOL/Tools/sat.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/HOL/Tools/sat.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -153,9 +153,9 @@
             val _ =
               cond_tracing ctxt (fn () =>
                 "Resolving clause: " ^ Thm.string_of_thm ctxt c1 ^
-                " (hyps: " ^ ML_Syntax.print_list (Syntax.string_of_term ctxt) (Termset.dest (Thm.hyps_of c1)) ^
+                " (hyps: " ^ ML_Syntax.print_list (Syntax.string_of_term ctxt) (Thm.hyps_of c1) ^
                 ")\nwith clause: " ^ Thm.string_of_thm ctxt c2 ^
-                " (hyps: " ^ ML_Syntax.print_list (Syntax.string_of_term ctxt) (Termset.dest (Thm.hyps_of c2)) ^ ")")
+                " (hyps: " ^ ML_Syntax.print_list (Syntax.string_of_term ctxt) (Thm.hyps_of c2) ^ ")")
 
             (* the two literals used for resolution *)
             val (hyp1_is_neg, hyp1, hyp2, new_hyps) = find_res_hyps (hyps1, hyps2) []
@@ -186,7 +186,7 @@
               cond_tracing ctxt (fn () =>
                 "Resulting clause: " ^ Thm.string_of_thm ctxt c_new ^
                 " (hyps: " ^
-                ML_Syntax.print_list (Syntax.string_of_term ctxt) (Termset.dest (Thm.hyps_of c_new)) ^ ")")
+                ML_Syntax.print_list (Syntax.string_of_term ctxt) (Thm.hyps_of c_new) ^ ")")
 
             val _ = Unsynchronized.inc counter
           in
@@ -216,7 +216,7 @@
       handle TERM _ => NONE;  (* 'chyp' is not a literal *)
 
     fun prove_clause id =
-      (case Array.sub (clauses, id) of
+      (case Array.nth clauses id of
         RAW_CLAUSE clause => clause
       | ORIG_CLAUSE thm =>
           (* convert the original clause *)
@@ -226,7 +226,7 @@
             val hyps = sort (lit_ord o apply2 fst) (map_filter (fn chyp =>
               Option.map (rpair chyp) (index_of_literal chyp)) (Thm.chyps_of raw))
             val clause = (raw, hyps)
-            val _ = Array.update (clauses, id, RAW_CLAUSE clause)
+            val _ = Array.upd clauses id (RAW_CLAUSE clause)
           in
             clause
           end
@@ -236,7 +236,7 @@
             val _ = cond_tracing ctxt (fn () => "Proving clause #" ^ string_of_int id ^ " ...")
             val ids = the (Inttab.lookup clause_table id)
             val clause = resolve_raw_clauses ctxt (map prove_clause ids)
-            val _ = Array.update (clauses, id, RAW_CLAUSE clause)
+            val _ = Array.upd clauses id (RAW_CLAUSE clause)
             val _ =
               cond_tracing ctxt
                 (fn () => "Replay chain successful; clause stored at #" ^ string_of_int id)
@@ -350,22 +350,22 @@
             (* [c_1, ..., c_n] |- c_1 && ... && c_n *)
             val clauses_thm = Conjunction.intr_balanced (map Thm.assume sorted_clauses)
             (* [c_1 && ... && c_n] |- c_1 && ... && c_n *)
-            val cnf_cterm = Thm.cprop_of clauses_thm
-            val cnf_thm = Thm.assume cnf_cterm
+            val clauses_cprop = Thm.cprop_of clauses_thm
             (* [[c_1 && ... && c_n] |- c_1, ..., [c_1 && ... && c_n] |- c_n] *)
-            val cnf_clauses = Conjunction.elim_balanced (length sorted_clauses) cnf_thm
+            val cnf_clauses =
+              Conjunction.elim_balanced (length sorted_clauses) (Thm.assume clauses_cprop)
             (* initialize the clause array with the given clauses *)
             val max_idx = fst (the (Inttab.max clause_table))
             val clause_arr = Array.array (max_idx + 1, NO_CLAUSE)
             val _ =
-              fold (fn thm => fn idx => (Array.update (clause_arr, idx, ORIG_CLAUSE thm); idx+1))
+              fold (fn thm => fn i => (Array.upd clause_arr i (ORIG_CLAUSE thm); i + 1))
                 cnf_clauses 0
             (* replay the proof to derive the empty clause *)
             (* [c_1 && ... && c_n] |- False *)
             val raw_thm = replay_proof ctxt atom_table clause_arr (clause_table, empty_id)
           in
             (* [c_1, ..., c_n] |- False *)
-            Thm.implies_elim (Thm.implies_intr cnf_cterm raw_thm) clauses_thm
+            Thm.implies_elim (Thm.implies_intr clauses_cprop raw_thm) clauses_thm
           end)
     | SAT_Solver.UNSATISFIABLE NONE =>
         if Config.get ctxt quick_and_dirty then
--- a/src/HOL/Tools/sat_solver.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/HOL/Tools/sat_solver.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -667,7 +667,7 @@
       fun init_array (Prop_Logic.And (fm1, fm2), n) =
         init_array (fm2, init_array (fm1, n))
         | init_array (fm, n) =
-        (Array.update (clauses, n, clause_to_lit_list fm); n+1)
+        (Array.upd clauses n (clause_to_lit_list fm); n+1)
       val _ = init_array (cnf, 0)
       (* optimization for the common case where MiniSat "R"s clauses in their *)
       (* original order:                                                      *)
@@ -682,7 +682,7 @@
             original_clause_id_from 0
           (* both 'lits' and the list of literals used in 'clauses' are sorted, so *)
           (* testing for equality should suffice -- barring duplicate literals     *)
-          else if Array.sub (clauses, index) = lits then (
+          else if Array.nth clauses index = lits then (
             (* success *)
             last_ref_clause := index;
             SOME index
--- a/src/HOL/Tools/semiring_normalizer.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/HOL/Tools/semiring_normalizer.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -144,7 +144,7 @@
       in if b = 1 then Numeral.mk_cnumber cT a
         else Thm.apply
              (Thm.apply
-               (Thm.instantiate_cterm (TVars.make [(divide_tvar, cT)], Vars.empty) divide_const)
+               (Thm.instantiate_cterm (TVars.make1 (divide_tvar, cT), Vars.empty) divide_const)
                (Numeral.mk_cnumber cT a))
              (Numeral.mk_cnumber cT b)
       end
--- a/src/HOL/Tools/split_rule.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/HOL/Tools/split_rule.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -42,7 +42,7 @@
 fun split_rule_var' ctxt (t as Var (v, Type ("fun", [T1, T2]))) rl =
       let val T' = HOLogic.flatten_tupleT T1 ---> T2;
           val newt = ap_split T1 T2 (Var (v, T'));
-      in Thm.instantiate (TVars.empty, Vars.make [(dest_Var t, Thm.cterm_of ctxt newt)]) rl end
+      in Thm.instantiate (TVars.empty, Vars.make1 (dest_Var t, Thm.cterm_of ctxt newt)) rl end
   | split_rule_var' _ _ rl = rl;
 
 
--- a/src/HOL/Types_To_Sets/internalize_sort.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/HOL/Types_To_Sets/internalize_sort.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -31,7 +31,7 @@
     val thy = Thm.theory_of_thm thm;
     val tvar = Thm.typ_of ctvar;
 
-    val (ucontext, _) = Logic.unconstrainT Sortset.empty (Thm.prop_of thm);
+    val (ucontext, _) = Logic.unconstrainT [] (Thm.prop_of thm);
 
     fun is_proper_class thy = can (Axclass.get_info thy); (* trick by FH *)
     fun reduce_to_non_proper_sort (TVar (name, sort)) =
--- a/src/HOL/Types_To_Sets/local_typedef.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/HOL/Types_To_Sets/local_typedef.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -35,7 +35,7 @@
 
     val thy = Thm.theory_of_thm thm;
     val prop = Thm.prop_of thm;
-    val hyps = Termset.dest (Thm.hyps_of thm);
+    val hyps = Thm.hyps_of thm;
 
     val _ = null (Thm.tpairs_of thm) orelse err "the theorem contains unsolved flex-flex pairs";
 
--- a/src/HOL/Types_To_Sets/unoverloading.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/HOL/Types_To_Sets/unoverloading.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -80,7 +80,7 @@
   let
     fun err msg = raise THM ("unoverloading: " ^ msg, 0, [thm]);
 
-    val _ = Termset.is_empty (Thm.hyps_of thm) orelse err "the theorem has meta hypotheses";
+    val _ = null (Thm.hyps_of thm) orelse err "the theorem has meta hypotheses";
     val _ = null (Thm.tpairs_of thm) orelse err "the theorem contains unresolved flex-flex pairs";
 
     val prop = Thm.prop_of thm;
--- a/src/HOL/ex/Join_Theory.thy	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/HOL/ex/Join_Theory.thy	Mon Apr 24 15:00:54 2023 +0100
@@ -35,7 +35,7 @@
 setup \<open>
   fn thy =>
     let val forked_thys = Par_List.map (fn i => Named_Target.theory_map (spec i) thy) (1 upto 10)
-    in Theory.join_theory forked_thys end
+    in Context.join_thys forked_thys end
 \<close>
 
 term test1
--- a/src/HOL/ex/Note_on_signed_division_on_words.thy	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/HOL/ex/Note_on_signed_division_on_words.thy	Mon Apr 24 15:00:54 2023 +0100
@@ -1,5 +1,5 @@
 theory Note_on_signed_division_on_words
-  imports "HOL-Library.Word" "HOL-Library.Rounded_Division"
+  imports "HOL-Library.Word" "HOL-Library.Centered_Division"
 begin
 
 unbundle bit_operations_syntax
@@ -24,26 +24,26 @@
 end
 
 lift_definition signed_divide_word :: \<open>'a::len word \<Rightarrow> 'a word \<Rightarrow> 'a word\<close>  (infixl \<open>wdiv\<close> 70)
-  is \<open>\<lambda>a b. signed_take_bit (LENGTH('a) - Suc 0) a rdiv signed_take_bit (LENGTH('a) - Suc 0) b\<close>
+  is \<open>\<lambda>a b. signed_take_bit (LENGTH('a) - Suc 0) a cdiv signed_take_bit (LENGTH('a) - Suc 0) b\<close>
   by (simp flip: signed_take_bit_decr_length_iff)
 
 lift_definition signed_modulo_word :: \<open>'a::len word \<Rightarrow> 'a word \<Rightarrow> 'a word\<close>  (infixl \<open>wmod\<close> 70)
-  is \<open>\<lambda>a b. signed_take_bit (LENGTH('a) - Suc 0) a rmod signed_take_bit (LENGTH('a) - Suc 0) b\<close>
+  is \<open>\<lambda>a b. signed_take_bit (LENGTH('a) - Suc 0) a cmod signed_take_bit (LENGTH('a) - Suc 0) b\<close>
   by (simp flip: signed_take_bit_decr_length_iff)
 
 lemma signed_take_bit_eq_wmod:
   \<open>signed_take_bit n w = w wmod (2 ^ Suc n)\<close>
 proof (transfer fixing: n)
   show \<open>take_bit LENGTH('a) (signed_take_bit n (take_bit LENGTH('a) k)) =
-    take_bit LENGTH('a) (signed_take_bit (LENGTH('a) - Suc 0) k rmod signed_take_bit (LENGTH('a) - Suc 0) (2 ^ Suc n))\<close> for k :: int
+    take_bit LENGTH('a) (signed_take_bit (LENGTH('a) - Suc 0) k cmod signed_take_bit (LENGTH('a) - Suc 0) (2 ^ Suc n))\<close> for k :: int
   proof (cases \<open>LENGTH('a) = Suc (Suc n)\<close>)
     case True
     then show ?thesis
-      by (simp add: signed_take_bit_exp_eq_int signed_take_bit_take_bit rmod_minus_eq flip: power_Suc signed_take_bit_eq_rmod)
+      by (simp add: signed_take_bit_exp_eq_int signed_take_bit_take_bit cmod_minus_eq flip: power_Suc signed_take_bit_eq_cmod)
   next
     case False
     then show ?thesis
-      by (auto simp add: signed_take_bit_exp_eq_int signed_take_bit_take_bit take_bit_signed_take_bit simp flip: power_Suc signed_take_bit_eq_rmod)
+      by (auto simp add: signed_take_bit_exp_eq_int signed_take_bit_take_bit take_bit_signed_take_bit simp flip: power_Suc signed_take_bit_eq_cmod)
   qed
 qed
 
--- a/src/Provers/Arith/fast_lin_arith.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/Provers/Arith/fast_lin_arith.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -437,7 +437,7 @@
             val T = Thm.typ_of_cterm cv
           in
             mth
-            |> Thm.instantiate (TVars.empty, Vars.make [(dest_Var (Thm.term_of cv), number_of T n)])
+            |> Thm.instantiate (TVars.empty, Vars.make1 (dest_Var (Thm.term_of cv), number_of T n))
             |> rewrite_concl
             |> discharge_prem
             handle CTERM _ => mult_by_add n thm
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Pure/General/array.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -0,0 +1,33 @@
+(*  Title:      Pure/General/array.ML
+    Author:     Makarius
+
+Additional operations for structure Array, following Isabelle/ML conventions.
+*)
+
+signature ARRAY =
+sig
+  include ARRAY
+  val nth: 'a array -> int -> 'a
+  val upd: 'a array -> int -> 'a -> unit
+  val forall: ('a -> bool) -> 'a array -> bool
+  val member: ('a * 'a -> bool) -> 'a array -> 'a -> bool
+  val fold: ('a -> 'b -> 'b) -> 'a array -> 'b -> 'b
+  val fold_rev: ('a -> 'b -> 'b) -> 'a array -> 'b -> 'b
+end;
+
+structure Array: ARRAY =
+struct
+
+open Array;
+
+fun nth arr i = sub (arr, i);
+fun upd arr i x = update (arr, i, x);
+
+val forall = all;
+
+fun member eq arr x = exists (fn y => eq (x, y)) arr;
+
+fun fold f arr x = foldl (fn (a, b) => f a b) x arr;
+fun fold_rev f arr x = foldr (fn (a, b) => f a b) x arr;
+
+end;
--- a/src/Pure/General/binding.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/Pure/General/binding.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -43,7 +43,7 @@
   val bad: binding -> string
   val check: binding -> unit
   val name_spec: scope list -> (string * bool) list -> binding ->
-    {restriction: bool option, concealed: bool, spec: (string * bool) list}
+    {restriction: bool option, concealed: bool, spec: (string * bool) list, full_name: string}
 end;
 
 structure Binding: BINDING =
@@ -209,10 +209,12 @@
     val spec2 = if name = "" then [] else [(name, true)];
     val spec = spec1 @ spec2;
     val _ =
-      exists (fn (a, _) => member (op =) bad_specs a orelse exists_string (fn s => s = "\"") a) spec
+      exists (fn (a, _) => member (op =) bad_specs a orelse member_string a "\"") spec
       andalso error (bad binding);
-  in {restriction = restriction, concealed = concealed, spec = if null spec2 then [] else spec} end;
 
+    val spec' = if null spec2 then [] else spec;
+    val full_name = Long_Name.implode (map #1 spec');
+  in {restriction = restriction, concealed = concealed, spec = spec', full_name = full_name} end;
 
 end;
 
--- a/src/Pure/General/bytes.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/Pure/General/bytes.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -158,7 +158,7 @@
   else if size sep <> 1 then [content bytes]
   else
     let
-      val sep_char = String.sub (sep, 0);
+      val sep_char = String.nth sep 0;
 
       fun add_part s part =
         Buffer.add (Substring.string s) (the_default Buffer.empty part);
--- a/src/Pure/General/file.scala	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/Pure/General/file.scala	Mon Apr 24 15:00:54 2023 +0100
@@ -383,10 +383,10 @@
   final class Content private[File](val path: Path, val content: Bytes) {
     override def toString: String = path.toString
 
-    def write(dir: Path): Unit = {
+    def write(dir: Path, ssh: SSH.System = SSH.Local): Unit = {
       val full_path = dir + path
-      Isabelle_System.make_directory(full_path.expand.dir)
-      Bytes.write(full_path, content)
+      ssh.make_directory(ssh.expand_path(full_path).dir)
+      ssh.write_bytes(full_path, content)
     }
   }
 
--- a/src/Pure/General/long_name.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/Pure/General/long_name.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -7,7 +7,13 @@
 signature LONG_NAME =
 sig
   val separator: string
+  val append: string -> string -> string
+  val qualify: string -> string -> string
   val is_qualified: string -> bool
+  val qualifier: string -> string
+  val base_name: string -> string
+  val map_base_name: (string -> string) -> string -> string
+  val count: string -> int
   val hidden: string -> string
   val is_hidden: string -> bool
   val dest_hidden: string -> string option
@@ -15,20 +21,58 @@
   val dest_local: string -> string option
   val implode: string list -> string
   val explode: string -> string list
-  val append: string -> string -> string
-  val qualification: string -> int
-  val qualify: string -> string -> string
-  val qualifier: string -> string
-  val base_name: string -> string
-  val map_base_name: (string -> string) -> string -> string
+  type chunks
+  val count_chunks: chunks -> int
+  val size_chunks: chunks -> int
+  val empty_chunks: chunks
+  val base_chunks: string -> chunks
+  val make_chunks: int -> bool list -> string -> chunks
+  val chunks: string -> chunks
+  val explode_chunks: chunks -> string list
+  val implode_chunks: chunks -> string
+  val compare_chunks: chunks * chunks -> order
+  val eq_chunks: chunks * chunks -> bool
+  structure Chunks: CHANGE_TABLE
 end;
 
 structure Long_Name: LONG_NAME =
 struct
 
+(* long names separated by "." *)
+
 val separator = ".";
+val separator_char = String.nth separator 0;
+
+fun append name1 "" = name1
+  | append "" name2 = name2
+  | append name1 name2 = name1 ^ separator ^ name2;
+
+fun qualify qual name =
+  if qual = "" orelse name = "" then name
+  else qual ^ separator ^ name;
+
+fun last_separator name =
+  let fun last i = if i < 0 orelse String.nth name i = separator_char then i else last (i - 1)
+  in last (size name - 1) end;
 
-val is_qualified = exists_string (fn s => s = separator);
+fun is_qualified name = last_separator name >= 0;
+
+fun qualifier name =
+  let val i = last_separator name
+  in if i < 0 then "" else String.substring (name, 0, i) end;
+
+fun base_name name =
+  let
+    val i = last_separator name;
+    val j = i + 1;
+  in if i < 0 then name else String.substring (name, j, size name - j) end;
+
+fun map_base_name f name =
+  if name = "" then ""
+  else qualify (qualifier name) (f (base_name name));
+
+fun count "" = 0
+  | count name = String.fold (fn c => c = separator_char ? Integer.add 1) name 1;
 
 val hidden_prefix = "??."
 val hidden = prefix hidden_prefix;
@@ -41,26 +85,129 @@
 val implode = space_implode separator;
 val explode = space_explode separator;
 
-fun append name1 "" = name1
-  | append "" name2 = name2
-  | append name1 name2 = name1 ^ separator ^ name2;
+
+(* compact representation of chunks *)
+
+local
+
+fun mask_bad s = s = 0w0;
+fun mask_set s m = Word.orb (m, s);
+fun mask_push s = Word.<< (s, 0w1);
+fun mask_next m = Word.>> (m, 0w1);
+fun mask_keep m = Word.andb (m, 0w1) = 0w0;
+
+fun string_ops string =
+  let
+    val n = size string;
+    val char = String.nth string;
+    fun string_end i = i >= n;
+    fun chunk_end i = string_end i orelse char i = separator_char;
+  in {string_end = string_end, chunk_end = chunk_end} end;
+
+in
+
+abstype chunks = Chunks of {count: int, size: int, mask: word, string: string}
+with
+
+fun count_chunks (Chunks {count, ...}) = count;
+fun size_chunks (Chunks {size, ...}) = size;
 
-fun qualification "" = 0
-  | qualification name = fold_string (fn s => s = separator ? Integer.add 1) name 1;
+val empty_chunks = Chunks {count = 0, size = 0, mask = 0w0, string = ""};
+
+fun make_chunks suppress1 suppress2 string =
+  let
+    val {string_end, chunk_end, ...} = string_ops string;
+
+    fun err msg = raise Fail ("Malformed qualified name " ^ quote string ^ ":\n  " ^ msg);
+
+    fun make suppr1 suppr2 i k l m s =
+      let
+        val is_end = chunk_end i;
 
-fun qualify qual name =
-  if qual = "" orelse name = "" then name
-  else qual ^ separator ^ name;
+        val suppr =
+          if suppr1 > 0 then true
+          else if suppr1 < 0 then false
+          else if null suppr2 then false
+          else hd suppr2;
+        val suppr1' =
+          if is_end andalso suppr1 > 0 then suppr1 - 1
+          else if is_end andalso suppr1 < 0 then suppr1 + 1
+          else suppr1;
+        val suppr2' =
+          if is_end andalso suppr1 = 0 andalso not (null suppr2)
+          then tl suppr2 else suppr2;
+
+        val keep = not suppr;
 
-fun qualifier "" = ""
-  | qualifier name = implode (#1 (split_last (explode name)));
+        val l' = if keep then l + 1 else l;
+        val k' = if is_end andalso keep andalso l' > 0 then k + 1 else k;
+        val m' =
+          if not is_end orelse keep then m
+          else if mask_bad s then
+            err ("cannot suppress chunks beyond word size " ^ string_of_int Word.wordSize)
+          else mask_set s m;
+        val s' = if is_end then mask_push s else s;
+      in
+        if not (string_end i) then make suppr1' suppr2' (i + 1) k' l' m' s'
+        else if not (suppr1' = 0 andalso null suppr2') then
+          err ("cannot suppress chunks beyond " ^ string_of_int k')
+        else if k' = 0 then empty_chunks
+        else Chunks {count = k', size = Int.max (0, l' - 1), mask = m', string = string}
+      end;
+  in make suppress1 suppress2 0 0 0 0w0 0w1 end;
+
+val chunks = make_chunks 0 [];
+
+fun base_chunks name =
+  make_chunks (Int.max (0, count name - 1)) [] name;
+
+fun expand_chunks f (Chunks {count, size, mask, string}) =
+  let
+    val {string_end, chunk_end, ...} = string_ops string;
+
+    fun explode bg en m acc =
+      let
+        val is_end = chunk_end en;
 
-fun base_name "" = ""
-  | base_name name = List.last (explode name);
+        val en' = en + 1;
+        val bg' = if is_end then en' else bg;
+        val m' = if is_end then mask_next m else m;
+        val acc' = if is_end andalso mask_keep m then f (string, bg, en - bg) :: acc else acc;
+      in if string_end en then rev acc' else explode bg' en' m' acc' end;
+  in
+    if count = 0 then []
+    else if count = 1 andalso size = String.size string then [f (string, 0, size)]
+    else explode 0 0 mask []
+  end;
+
+val explode_chunks = expand_chunks String.substring;
+val implode_chunks = implode o explode_chunks;
 
-fun map_base_name _ "" = ""
-  | map_base_name f name =
-      let val names = explode name
-      in implode (nth_map (length names - 1) f names) end;
+val compare_chunks =
+  pointer_eq_ord (fn (chunks1, chunks2) =>
+    let
+      val Chunks args1 = chunks1;
+      val Chunks args2 = chunks2;
+      val ord1 =
+        int_ord o apply2 #size |||
+        int_ord o apply2 #count;
+      val ord2 =
+        dict_ord int_ord o apply2 (expand_chunks #3) |||
+        dict_ord Substring.compare o apply2 (expand_chunks Substring.substring);
+    in
+      (case ord1 (args1, args2) of
+        EQUAL =>
+          if #mask args1 = #mask args2 andalso #string args1 = #string args2 then EQUAL
+          else ord2 (chunks1, chunks2)
+      | ord => ord)
+    end);
+
+val eq_chunks = is_equal o compare_chunks;
 
 end;
+
+end;
+
+structure Chunks = Change_Table(type key = chunks val ord = compare_chunks);
+
+end;
--- a/src/Pure/General/mercurial.scala	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/Pure/General/mercurial.scala	Mon Apr 24 15:00:54 2023 +0100
@@ -307,20 +307,19 @@
       Isabelle_System.with_tmp_dir("sync") { tmp_dir =>
         Hg_Sync.check_directory(target, ssh = context.ssh)
 
-        Rsync.init(context.no_progress, target)
-
         val id_content = id(rev = rev)
         val is_changed = id_content.endsWith("+")
         val log_content = if (is_changed) "" else log(rev = rev, options = "-l1")
         val diff_content = if (is_changed) diff(rev = rev, options = "--git") else ""
         val stat_content = if (is_changed) diff(rev = rev, options = "--stat") else ""
 
-        Rsync.init(context.no_progress, target,
-          contents =
-            File.content(Hg_Sync.PATH_ID, id_content) ::
-            File.content(Hg_Sync.PATH_LOG, log_content) ::
-            File.content(Hg_Sync.PATH_DIFF, diff_content) ::
-            File.content(Hg_Sync.PATH_STAT, stat_content) :: contents)
+        val all_contents =
+          File.content(Hg_Sync.PATH_ID, id_content) ::
+          File.content(Hg_Sync.PATH_LOG, log_content) ::
+          File.content(Hg_Sync.PATH_DIFF, diff_content) ::
+          File.content(Hg_Sync.PATH_STAT, stat_content) :: contents
+
+        all_contents.foreach(_.write(target, ssh = context.ssh))
 
         val (exclude, source) =
           if (rev.isEmpty) {
--- a/src/Pure/General/name_space.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/Pure/General/name_space.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -126,71 +126,82 @@
 
 (* internal names *)
 
-type internals = (string list * string list) Change_Table.T;  (*xname -> visible, hidden*)
+type internal = string list * string list;  (*visible, hidden*)
+type internals = internal Long_Name.Chunks.T;  (*xname -> internal*)
 
 fun map_internals f xname : internals -> internals =
-  Change_Table.map_default (xname, ([], [])) f;
+  Long_Name.Chunks.map_default (xname, ([], [])) f;
 
 val del_name = map_internals o apfst o remove (op =);
 fun del_name_extra name =
   map_internals (apfst (fn [] => [] | x :: xs => x :: remove (op =) name xs));
 val add_name = map_internals o apfst o update (op =);
-fun hide_name name = map_internals (apsnd (update (op =) name)) name;
+fun hide_name name = map_internals (apsnd (update (op =) name)) (Long_Name.chunks name);
 
 
 (* external accesses *)
 
-type accesses = (xstring list * xstring list);  (*input / output fragments*)
-type entries = (accesses * entry) Change_Table.T;  (*name -> accesses, entry*)
+type external =
+ {accesses: Long_Name.chunks list,
+  accesses': Long_Name.chunks list,
+  entry: entry};
+
+type externals = external Change_Table.T;  (*name -> external*)
 
 
 (* datatype T *)
 
-datatype T = Name_Space of {kind: string, internals: internals, entries: entries};
+datatype T = Name_Space of {kind: string, internals: internals, externals: externals};
 
-fun make_name_space (kind, internals, entries) =
-  Name_Space {kind = kind, internals = internals, entries = entries};
+fun make_name_space (kind, internals, externals) =
+  Name_Space {kind = kind, internals = internals, externals = externals};
+
+fun map_name_space f (Name_Space {kind = kind, internals = internals, externals = externals}) =
+  make_name_space (f (kind, internals, externals));
 
-fun map_name_space f (Name_Space {kind = kind, internals = internals, entries = entries}) =
-  make_name_space (f (kind, internals, entries));
+fun change_base_space begin = map_name_space (fn (kind, internals, externals) =>
+  (kind,
+    Long_Name.Chunks.change_base begin internals,
+    Change_Table.change_base begin externals));
 
-fun change_base_space begin = map_name_space (fn (kind, internals, entries) =>
-  (kind, Change_Table.change_base begin internals, Change_Table.change_base begin entries));
-
-val change_ignore_space = map_name_space (fn (kind, internals, entries) =>
-  (kind, Change_Table.change_ignore internals, Change_Table.change_ignore entries));
+val change_ignore_space = map_name_space (fn (kind, internals, externals) =>
+  (kind,
+    Long_Name.Chunks.change_ignore internals,
+    Change_Table.change_ignore externals));
 
 
-fun empty kind = make_name_space (kind, Change_Table.empty, Change_Table.empty);
+fun empty kind = make_name_space (kind, Long_Name.Chunks.empty, Change_Table.empty);
 
 fun kind_of (Name_Space {kind, ...}) = kind;
+fun lookup_internals (Name_Space {internals, ...}) = Long_Name.Chunks.lookup internals;
+fun lookup_externals (Name_Space {externals, ...}) = Change_Table.lookup externals;
 
-fun gen_markup def (Name_Space {kind, entries, ...}) name =
-  (case Change_Table.lookup entries name of
+fun gen_markup def space name =
+  (case lookup_externals space name of
     NONE => Markup.intensify
-  | SOME (_, entry) => entry_markup def kind (name, entry));
+  | SOME {entry, ...} => entry_markup def (kind_of space) (name, entry));
 
 val markup = gen_markup {def = false};
 val markup_def = gen_markup {def = true};
 
-fun undefined (space as Name_Space {kind, entries, ...}) bad =
+fun undefined (space as Name_Space {kind, externals, ...}) bad =
   let
     val (prfx, sfx) =
       (case Long_Name.dest_hidden bad of
         SOME name =>
-          if Change_Table.defined entries name
+          if Change_Table.defined externals name
           then ("Inaccessible", Markup.markup (markup space name) (quote name))
           else ("Undefined", quote name)
       | NONE => ("Undefined", quote bad));
   in prfx ^ " " ^ plain_words kind ^ ": " ^ sfx end;
 
-fun get_names (Name_Space {entries, ...}) =
-  Change_Table.fold (cons o #1) entries [];
+fun get_names (Name_Space {externals, ...}) =
+  Change_Table.fold (cons o #1) externals [];
 
-fun the_entry (space as Name_Space {entries, ...}) name =
-  (case Change_Table.lookup entries name of
+fun the_entry space name =
+  (case lookup_externals space name of
     NONE => error (undefined space name)
-  | SOME (_, entry) => entry);
+  | SOME {entry, ...} => entry);
 
 fun the_entry_theory_name space name =
   Long_Name.base_name (#theory_long_name (the_entry space name));
@@ -203,22 +214,17 @@
 
 (* intern *)
 
-fun intern' (Name_Space {internals, ...}) xname =
-  (case the_default ([], []) (Change_Table.lookup internals xname) of
+fun intern_chunks space xname =
+  (case the_default ([], []) (lookup_internals space xname) of
     ([name], _) => (name, true)
   | (name :: _, _) => (name, false)
-  | ([], []) => (Long_Name.hidden xname, true)
+  | ([], []) => (Long_Name.hidden (Long_Name.implode_chunks xname), true)
   | ([], name' :: _) => (Long_Name.hidden name', true));
 
-val intern = #1 oo intern';
+fun intern space = #1 o intern_chunks space o Long_Name.chunks;
 
-fun get_accesses (Name_Space {entries, ...}) name =
-  (case Change_Table.lookup entries name of
-    NONE => ([], [])
-  | SOME (accesses, _) => accesses);
-
-fun is_valid_access (Name_Space {internals, ...}) name xname =
-  (case Change_Table.lookup internals xname of
+fun is_valid_access space name xname =
+  (case lookup_internals space xname of
     SOME (name' :: _, _) => name = name'
   | _ => false);
 
@@ -236,15 +242,22 @@
     val names_unique = Config.get ctxt names_unique;
 
     fun valid require_unique xname =
-      let val (name', is_unique) = intern' space xname
+      let val (name', is_unique) = intern_chunks space xname
       in name = name' andalso (not require_unique orelse is_unique) end;
 
-    fun ext [] = if valid false name then name else Long_Name.hidden name
-      | ext (nm :: nms) = if valid names_unique nm then nm else ext nms;
+    fun extern [] =
+          if valid false (Long_Name.chunks name) then name
+          else Long_Name.hidden name
+      | extern (a :: bs) =
+          if valid names_unique a then Long_Name.implode_chunks a
+          else extern bs;
   in
     if names_long then name
     else if names_short then Long_Name.base_name name
-    else ext (#2 (get_accesses space name))
+    else
+      (case lookup_externals space name of
+        NONE => extern []
+      | SOME {accesses', ...} => extern accesses')
   end;
 
 fun extern_ord ctxt space = string_ord o apply2 (extern ctxt space);
@@ -270,7 +283,7 @@
           EQUAL =>
             (case bool_ord (apply2 (is_some o Long_Name.dest_local) (name2, name1)) of
               EQUAL =>
-                (case int_ord (apply2 Long_Name.qualification (xname1, xname2)) of
+                (case int_ord (apply2 Long_Name.count (xname1, xname2)) of
                   EQUAL => string_ord (xname1, xname2)
                 | ord => ord)
             | ord => ord)
@@ -292,7 +305,9 @@
           end
         else I;
     in
-      Change_Table.fold (fn (xname', (name :: _, _)) => complete xname' name | _ => I) internals []
+      Long_Name.Chunks.fold
+        (fn (xname', (name :: _, _)) => complete (Long_Name.implode_chunks xname') name | _ => I)
+        internals []
       |> sort_distinct result_ord
       |> map #2
     end);
@@ -301,23 +316,23 @@
 (* merge *)
 
 fun merge
-  (Name_Space {kind = kind1, internals = internals1, entries = entries1},
-    Name_Space {kind = kind2, internals = internals2, entries = entries2}) =
+  (Name_Space {kind = kind1, internals = internals1, externals = externals1},
+    Name_Space {kind = kind2, internals = internals2, externals = externals2}) =
   let
     val kind' =
       if kind1 = kind2 then kind1
       else error ("Attempt to merge different kinds of name spaces " ^
         quote kind1 ^ " vs. " ^ quote kind2);
-    val internals' = (internals1, internals2) |> Change_Table.join
+    val internals' = (internals1, internals2) |> Long_Name.Chunks.join
       (K (fn ((names1, names1'), (names2, names2')) =>
         if pointer_eq (names1, names2) andalso pointer_eq (names1', names2')
-        then raise Change_Table.SAME
+        then raise Long_Name.Chunks.SAME
         else (Library.merge (op =) (names1, names2), Library.merge (op =) (names1', names2'))));
-    val entries' = (entries1, entries2) |> Change_Table.join
-      (fn name => fn ((_, entry1), (_, entry2)) =>
+    val externals' = (externals1, externals2) |> Change_Table.join
+      (fn name => fn ({entry = entry1, ...}, {entry = entry2, ...}) =>
         if #serial entry1 = #serial entry2 then raise Change_Table.SAME
         else err_dup kind' (name, entry1) (name, entry2) Position.none);
-  in make_name_space (kind', internals', entries') end;
+  in make_name_space (kind', internals', externals') end;
 
 
 
@@ -427,15 +442,15 @@
 fun name_spec naming binding =
   Binding.name_spec (get_scopes naming) (get_path naming) (transform_binding naming binding);
 
-fun full_name naming =
-  name_spec naming #> #spec #> map #1 #> Long_Name.implode;
-
-val base_name = full_name global_naming #> Long_Name.base_name;
+val full_name = #full_name oo name_spec;
+val base_name = Long_Name.base_name o full_name global_naming;
 
 
 (* accesses *)
 
-fun mandatory xs = map_filter (fn (x, true) => SOME x | _ => NONE) xs;
+local
+
+fun mandatory xs = fold_rev (fn (x, b) => b ? cons x) xs [];
 
 fun mandatory_prefixes xs = mandatory xs :: mandatory_prefixes1 xs
 and mandatory_prefixes1 [] = []
@@ -444,46 +459,59 @@
 
 fun mandatory_suffixes xs = map rev (mandatory_prefixes (rev xs));
 
+in
+
 fun make_accesses naming binding =
-  (case name_spec naming binding of
-    {restriction = SOME true, ...} => ([], [])
-  | {restriction, spec, ...} =>
-      let
-        val restrict = is_some restriction ? filter (fn [_] => false | _ => true);
-        val sfxs = restrict (mandatory_suffixes spec);
-        val pfxs = restrict (mandatory_prefixes spec);
-      in apply2 (map Long_Name.implode) (sfxs @ pfxs, sfxs) end);
+  let
+    val args as {restriction, spec, ...} = name_spec naming binding;
+    val accesses =
+      if restriction = SOME true then ([], [])
+      else
+        let
+          val restrict = is_some restriction ? filter (fn [_] => false | _ => true);
+          val sfxs = restrict (mandatory_suffixes spec);
+          val pfxs = restrict (mandatory_prefixes spec);
+        in apply2 (map Long_Name.implode o distinct op =) (sfxs @ pfxs, sfxs) end;
+  in (apply2 (map Long_Name.chunks) accesses (* FIXME proper make_chunks mask*), args) end;
+
+end;
 
 
 (* hide *)
 
 fun hide fully name space =
-  space |> map_name_space (fn (kind, internals, entries) =>
+  space |> map_name_space (fn (kind, internals, externals) =>
     let
       val _ = the_entry space name;
-      val (accs, accs') = get_accesses space name;
-      val xnames = filter (is_valid_access space name) accs;
+      val (accesses, accesses') =
+        (case lookup_externals space name of
+          NONE => ([], [])
+        | SOME {accesses, accesses', ...} => (accesses, accesses'));
+      val xnames = filter (is_valid_access space name) accesses;
+      val xnames' =
+        if fully then xnames
+        else inter Long_Name.eq_chunks [Long_Name.base_chunks name] xnames;
       val internals' = internals
         |> hide_name name
-        |> fold (del_name name)
-          (if fully then xnames else inter (op =) [Long_Name.base_name name] xnames)
-        |> fold (del_name_extra name) accs';
-    in (kind, internals', entries) end);
+        |> fold (del_name name) xnames'
+        |> fold (del_name_extra name) accesses';
+    in (kind, internals', externals) end);
 
 
 (* alias *)
 
 fun alias naming binding name space =
-  space |> map_name_space (fn (kind, internals, entries) =>
+  space |> map_name_space (fn (kind, internals, externals) =>
     let
       val _ = the_entry space name;
-      val (more_accs, more_accs') = make_accesses naming binding;
+      val (more_accs, more_accs') = #1 (make_accesses naming binding);
       val internals' = internals |> fold (add_name name) more_accs;
-      val entries' = entries
-        |> Change_Table.map_entry name (apfst (fn (accs, accs') =>
-            (fold_rev (update op =) more_accs accs,
-             fold_rev (update op =) more_accs' accs')))
-    in (kind, internals', entries') end);
+      val externals' = externals
+        |> Change_Table.map_entry name (fn {accesses, accesses', entry} =>
+            {accesses = fold_rev (update Long_Name.eq_chunks) more_accs accesses,
+             accesses' = fold_rev (update Long_Name.eq_chunks) more_accs' accesses',
+             entry = entry});
+    in (kind, internals', externals') end);
 
 
 
@@ -512,16 +540,13 @@
 
 (* declaration *)
 
-fun declared (Name_Space {entries, ...}) = Change_Table.defined entries;
+fun declared (Name_Space {externals, ...}) = Change_Table.defined externals;
 
 fun declare context strict binding space =
   let
     val naming = naming_of context;
     val Naming {group, theory_long_name, ...} = naming;
-    val {concealed, spec, ...} = name_spec naming binding;
-    val accesses = make_accesses naming binding;
-
-    val name = Long_Name.implode (map fst spec);
+    val ((accesses, accesses'), {concealed, full_name = name, ...}) = make_accesses naming binding;
     val _ = name = "" andalso error (Binding.bad binding);
 
     val (proper_pos, pos) = Position.default (Binding.pos_of binding);
@@ -532,16 +557,16 @@
       pos = pos,
       serial = serial ()};
     val space' =
-      space |> map_name_space (fn (kind, internals, entries) =>
+      space |> map_name_space (fn (kind, internals, externals) =>
         let
-          val internals' = internals |> fold (add_name name) (#1 accesses);
-          val entries' =
+          val internals' = internals |> fold (add_name name) accesses;
+          val externals' =
             (if strict then Change_Table.update_new else Change_Table.update)
-              (name, (accesses, entry)) entries
+              (name, {accesses = accesses, accesses' = accesses', entry = entry}) externals
             handle Change_Table.DUP dup =>
-              err_dup kind (dup, #2 (the (Change_Table.lookup entries dup)))
+              err_dup kind (dup, #entry (the (lookup_externals space dup)))
                 (name, entry) (#pos entry);
-        in (kind, internals', entries') end);
+        in (kind, internals', externals') end);
     val _ =
       if proper_pos andalso Context_Position.is_reported_generic context pos then
         Position.report pos (entry_markup {def = true} (kind_of space) (name, entry))
--- a/src/Pure/General/rsync.scala	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/Pure/General/rsync.scala	Mon Apr 24 15:00:54 2023 +0100
@@ -28,9 +28,6 @@
   ) {
     override def toString: String = directory.toString
 
-    def no_progress: Context = new Context(directory, new Progress, archive, stats)
-    def no_archive: Context = new Context(directory, progress, false, stats)
-
     def ssh: SSH.System = directory.ssh
 
     def command: String = {
@@ -70,17 +67,4 @@
         if_proper(args, " " + Bash.strings(args))
     progress.bash(script, echo = true)
   }
-
-  def init(context: Context, target: Path,
-    contents: List[File.Content] = Nil
-  ): Unit =
-    Isabelle_System.with_tmp_dir("sync") { tmp_dir =>
-      val init_dir = Isabelle_System.make_directory(tmp_dir + Path.explode("init"))
-      contents.foreach(_.write(init_dir))
-      exec(context.no_archive,
-        thorough = true,
-        args =
-          List(if (contents.nonEmpty) "--archive" else "--dirs",
-            File.bash_path(init_dir) + "/.", context.target(target))).check
-    }
 }
--- a/src/Pure/General/set.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/Pure/General/set.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -23,6 +23,7 @@
   val get_first: (elem -> 'a option) -> T -> 'a option
   val member: T -> elem -> bool
   val subset: T * T -> bool
+  val eq_set: T * T -> bool
   val ord: T ord
   val insert: elem -> T -> T
   val make: elem list -> T
@@ -30,6 +31,9 @@
   val merges: T list -> T
   val remove: elem -> T -> T
   val subtract: T -> T -> T
+  val restrict: (elem -> bool) -> T -> T
+  val inter: T -> T -> T
+  val union: T -> T -> T
 end;
 
 functor Set(Key: KEY): SET =
@@ -237,7 +241,7 @@
   in get end;
 
 
-(* member *)
+(* member and subset *)
 
 fun member set elem =
   let
@@ -273,18 +277,25 @@
       | mem (Size (_, arg)) = mem arg;
   in mem set end;
 
+fun subset (set1, set2) = forall (member set2) set1;
 
-(* subset and order *)
+
+(* equality and order *)
 
-fun subset (set1, set2) = forall (member set2) set1;
+fun eq_set (set1, set2) =
+  pointer_eq (set1, set2) orelse size set1 = size set2 andalso subset (set1, set2);
 
 val ord =
-  pointer_eq_ord (fn sets =>
-    (case int_ord (apply2 size sets) of
+  pointer_eq_ord (fn (set1, set2) =>
+    (case int_ord (size set1, size set2) of
       EQUAL =>
-        if subset sets then EQUAL
-        else dict_ord Key.ord (apply2 dest sets)
-    | ord => ord));
+        (case get_first (fn a => if member set2 a then NONE else SOME a) set1 of
+          NONE => EQUAL
+        | SOME a =>
+            (case get_first (fn b => if member set1 b then NONE else SOME b) set2 of
+              NONE => EQUAL
+            | SOME b => Key.ord (a, b)))
+    | order => order));
 
 
 (* insert *)
@@ -295,7 +306,8 @@
   if member set elem then set
   else
     let
-      fun elem_ord e = Key.ord (elem, e)
+      fun elem_ord e = Key.ord (elem, e);
+      val elem_less = is_less o elem_ord;
 
       fun ins Empty = Sprout (Empty, elem, Empty)
         | ins (t as Leaf1 _) = ins (unmake t)
@@ -336,10 +348,23 @@
                         Sprout (make2 (left, e1, mid), e2, make2 (right1, e', right2)))))
         | ins (Size (_, arg)) = ins arg;
     in
-      make_size (size set + 1)
-        (case ins set of
-          Stay set' => set'
-        | Sprout br => make2 br)
+      (case set of
+        Empty => Leaf1 elem
+      | Leaf1 e =>
+          if elem_less e
+          then Leaf2 (elem, e)
+          else Leaf2 (e, elem)
+      | Leaf2 (e1, e2) =>
+          if elem_less e1
+          then Leaf3 (elem, e1, e2)
+          else if elem_less e2
+          then Leaf3 (e1, elem, e2)
+          else Leaf3 (e1, e2, elem)
+      | _ =>
+          make_size (size set + 1)
+            (case ins set of
+              Stay set' => set'
+            | Sprout br => make2 br))
     end;
 
 fun make elems = build (fold insert elems);
@@ -458,6 +483,20 @@
 end;
 
 
+(* conventional set operations *)
+
+fun restrict pred set =
+  fold_set (fn x => not (pred x) ? remove x) set set;
+
+fun inter set1 set2 =
+  if pointer_eq (set1, set2) then set1
+  else if is_empty set1 orelse is_empty set2 then empty
+  else if size set1 < size set2 then restrict (member set2) set1
+  else restrict (member set1) set2;
+
+fun union set1 set2 = merge (set2, set1);
+
+
 (* ML pretty-printing *)
 
 val _ =
@@ -472,5 +511,4 @@
 end;
 
 structure Intset = Set(Inttab.Key);
-structure Intset' = Set(Inttab'.Key);
 structure Symset = Set(Symtab.Key);
--- a/src/Pure/General/sha1.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/Pure/General/sha1.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -101,12 +101,12 @@
     (*hash result -- 5 words*)
     val hash_array : Word32.word Array.array =
       Array.fromList [0wx67452301, 0wxEFCDAB89, 0wx98BADCFE, 0wx10325476, 0wxC3D2E1F0];
-    fun hash i = Array.sub (hash_array, i);
-    fun add_hash x i = Array.update (hash_array, i, hash i + x);
+    val hash = Array.nth hash_array;
+    fun add_hash x i = Array.upd hash_array i (hash i + x);
 
     (*current chunk -- 80 words*)
     val chunk_array = Array.array (80, 0w0: Word32.word);
-    fun chunk i = Array.sub (chunk_array, i);
+    val chunk = Array.nth chunk_array;
     fun init_chunk pos =
       Array.modifyi (fn (i, _) =>
         if i < 16 then text (pos + i)
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Pure/General/string.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -0,0 +1,50 @@
+(*  Title:      Pure/General/string.ML
+    Author:     Makarius
+
+Additional operations for structure String, following Isabelle/ML conventions.
+
+NB: Isabelle normally works with Symbol.symbol, which is represented as a
+small string fragment. Raw type char is rarely useful.
+*)
+
+signature STRING =
+sig
+  include STRING
+  val nth: string -> int -> char
+  val exists: (char -> bool) -> string -> bool
+  val forall: (char -> bool) -> string -> bool
+  val member: string -> char -> bool
+  val fold: (char -> 'a -> 'a) -> string -> 'a -> 'a
+  val fold_rev: (char -> 'a -> 'a) -> string -> 'a -> 'a
+end;
+
+structure String: STRING =
+struct
+
+open String;
+
+fun nth str i = sub (str, i);
+
+fun exists pred str =
+  let
+    val n = size str;
+    fun ex i = Int.< (i, n) andalso (pred (nth str i) orelse ex (i + 1));
+  in ex 0 end;
+
+fun forall pred = not o exists (not o pred);
+
+fun member str c = exists (fn c' => c = c') str;
+
+fun fold f str x0 =
+  let
+    val n = size str;
+    fun iter (x, i) = if Int.< (i, n) then iter (f (nth str i) x, i + 1) else x;
+  in iter (x0, 0) end;
+
+fun fold_rev f str x0 =
+  let
+    val n = size str;
+    fun iter (x, i) = if Int.>= (i, 0) then iter (f (nth str i) x, i - 1) else x;
+  in iter (x0, n - 1) end;
+
+end;
--- a/src/Pure/General/symbol.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/Pure/General/symbol.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -110,13 +110,11 @@
 fun is_space s = s = space;
 
 local
-  val small_spaces = Vector.tabulate (65, fn i => replicate_string i space);
+  val spaces0 = Vector.tabulate (65, fn i => replicate_string i space);
 in
   fun spaces n =
-    if n < 64 then Vector.sub (small_spaces, n)
-    else
-      replicate_string (n div 64) (Vector.sub (small_spaces, 64)) ^
-        Vector.sub (small_spaces, n mod 64);
+    if n < 64 then Vector.nth spaces0 n
+    else replicate_string (n div 64) (Vector.nth spaces0 64) ^ Vector.nth spaces0 (n mod 64);
 end;
 
 val comment = "\<comment>";
@@ -126,7 +124,7 @@
 
 fun is_char s = size s = 1;
 
-fun is_utf8 s = size s > 0 andalso forall_string (fn c => ord c >= 128) s;
+fun is_utf8 s = size s > 0 andalso String.forall (fn c => Char.ord c >= 128) s;
 
 fun raw_symbolic s =
   String.isPrefix "\092<" s andalso String.isSuffix ">" s andalso not (String.isPrefix "\092<^" s);
--- a/src/Pure/General/symbol_explode.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/Pure/General/symbol_explode.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -17,7 +17,7 @@
 
 fun explode string =
   let
-    fun char i = String.sub (string, i);
+    val char = String.nth string;
     fun substring i j = String.substring (string, i, j - i);
 
     val n = size string;
--- a/src/Pure/General/table.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/Pure/General/table.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -454,9 +454,26 @@
       | modfy (Size (_, arg)) = modfy arg;
 
     val tab' =
-      (case modfy tab of
-        Stay tab' => tab'
-      | Sprout br => make2 br);
+      (case tab of
+        Empty => Leaf1 (key, insert ())
+      | Leaf1 (k, x) =>
+          (case key_ord k of
+            LESS => Leaf2 (key, insert (), k, x)
+          | EQUAL => Leaf1 (k, update x)
+          | GREATER => Leaf2 (k, x, key, insert ()))
+      | Leaf2 (k1, x1, k2, x2) =>
+          (case key_ord k1 of
+            LESS => Leaf3 (key, insert (), k1, x1, k2, x2)
+          | EQUAL => Leaf2 (k1, update x1, k2, x2)
+          | GREATER =>
+              (case key_ord k2 of
+                LESS => Leaf3 (k1, x1, key, insert (), k2, x2)
+              | EQUAL => Leaf2 (k1, x1, k2, update x2)
+              | GREATER => Leaf3 (k1, x1, k2, x2, key, insert ())))
+      | _ =>
+          (case modfy tab of
+            Stay tab' => tab'
+          | Sprout br => make2 br));
   in
     make_size (size tab + !inc) tab'
   end handle SAME => tab;
@@ -649,7 +666,6 @@
 end;
 
 structure Inttab = Table(type key = int val ord = int_ord);
-structure Inttab' = Table(type key = int val ord = rev_order o int_ord);
 structure Symtab = Table(type key = string val ord = fast_string_ord);
 structure Symreltab = Table(type key = string * string
   val ord = prod_ord fast_string_ord fast_string_ord);
--- a/src/Pure/General/value.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/Pure/General/value.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -42,7 +42,7 @@
 fun parse_int str =
   let
     fun err () = raise Fail ("Bad integer " ^ quote str);
-    fun char i = Char.ord (String.sub (str, i));
+    val char = Char.ord o String.nth str;
     val n = size str;
 
     fun digits i a =
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Pure/General/vector.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -0,0 +1,31 @@
+(*  Title:      Pure/General/vector.ML
+    Author:     Makarius
+
+Additional operations for structure Vector, following Isabelle/ML conventions.
+*)
+
+signature VECTOR =
+sig
+  include VECTOR
+  val nth: 'a vector -> int -> 'a
+  val forall: ('a -> bool) -> 'a vector -> bool
+  val member: ('a * 'a -> bool) -> 'a vector -> 'a -> bool
+  val fold: ('a -> 'b -> 'b) -> 'a vector -> 'b -> 'b
+  val fold_rev: ('a -> 'b -> 'b) -> 'a vector -> 'b -> 'b
+end;
+
+structure Vector: VECTOR =
+struct
+
+open Vector;
+
+fun nth vec i = sub (vec, i);
+
+val forall = all;
+
+fun member eq vec x = exists (fn y => eq (x, y)) vec;
+
+fun fold f vec x = foldl (fn (a, b) => f a b) x vec;
+fun fold_rev f vec x = foldr (fn (a, b) => f a b) x vec;
+
+end;
--- a/src/Pure/Isar/bundle.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/Pure/Isar/bundle.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -70,9 +70,8 @@
 (* target -- bundle under construction *)
 
 fun the_target thy =
-  (case #2 (Data.get (Context.Theory thy)) of
-    SOME thms => thms
-  | NONE => error "Missing bundle target");
+  #2 (Data.get (Context.Theory thy))
+  |> \<^if_none>\<open>error "Missing bundle target"\<close>;
 
 val reset_target = (Context.theory_map o Data.map o apsnd o K) NONE;
 val set_target = Context.theory_map o Data.map o apsnd o K o SOME o Attrib.trim_context_thms;
--- a/src/Pure/Isar/class.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/Pure/Isar/class.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -121,9 +121,8 @@
   | NONE => NONE);
 
 fun the_class_data thy class =
-  (case lookup_class_data thy class of
-    NONE => error ("Undeclared class " ^ quote class)
-  | SOME data => data);
+  lookup_class_data thy class
+  |> \<^if_none>\<open>error ("Undeclared class " ^ quote class)\<close>;
 
 val is_class = is_some oo lookup_class_data;
 
--- a/src/Pure/Isar/class_declaration.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/Pure/Isar/class_declaration.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -44,7 +44,7 @@
       Element.instantiate_normalize_morphism (TFrees.empty, param_map_inst);
     val typ_morph =
       Element.instantiate_normalize_morphism
-        (TFrees.make [(a_tfree, certT (Term.aT [class]))], Frees.empty)
+        (TFrees.make1 (a_tfree, certT (Term.aT [class])), Frees.empty)
     val (([raw_props], _, [(_, raw_inst_morph)], _, export_morph), _) = thy_ctxt
       |> Expression.cert_goal_expression ([(class, (("", false),
            (Expression.Named param_map_const, [])))], []);
--- a/src/Pure/Isar/context_rules.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/Pure/Isar/context_rules.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -77,7 +77,7 @@
 
 fun add_rule (i, b) opt_w th (Rules {next, rules, netpairs, wrappers}) =
   let
-    val w = (case opt_w of SOME w => w | NONE => Tactic.subgoals_of_brl (b, th));
+    val w = opt_w |> \<^if_none>\<open>Tactic.subgoals_of_brl (b, th)\<close>;
     val th' = Thm.trim_context th;
   in
     make_rules (next - 1) ((w, ((i, b), th')) :: rules)
--- a/src/Pure/Isar/element.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/Pure/Isar/element.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -207,7 +207,7 @@
     SOME C =>
       let
         val thesis = Var ((Auto_Bind.thesisN, Thm.maxidx_of th + 1), fastype_of C);
-        val insts = (TVars.empty, Vars.make [(Term.dest_Var C, Thm.cterm_of ctxt thesis)]);
+        val insts = (TVars.empty, Vars.make1 (Term.dest_Var C, Thm.cterm_of ctxt thesis));
         val th' = Thm.instantiate insts th;
       in (th', true) end
   | NONE => (th, false));
@@ -266,7 +266,7 @@
 
 val mark_witness = Logic.protect;
 fun witness_prop (Witness (t, _)) = t;
-fun witness_hyps (Witness (_, th)) = Termset.dest (Thm.hyps_of th);
+fun witness_hyps (Witness (_, th)) = Thm.hyps_of th;
 fun map_witness f (Witness witn) = Witness (f witn);
 
 fun transform_witness phi = map_witness (fn (t, th) => (Morphism.term phi t, Morphism.thm phi th));
--- a/src/Pure/Isar/locale.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/Pure/Isar/locale.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -355,8 +355,9 @@
     unique registration serial points to mixin list*)
   type T = reg Idtab.table * mixins;
   val empty: T = (Idtab.empty, Inttab.empty);
-  fun merge old_thys =
+  fun merge args =
     let
+      val ctxt0 = Syntax.init_pretty_global (#1 (hd args));
       fun recursive_merge ((regs1, mixins1), (regs2, mixins2)) : T =
         (Idtab.merge eq_reg (regs1, regs2), merge_mixins (mixins1, mixins2))
         handle Idtab.DUP id =>
@@ -373,9 +374,9 @@
             val _ =
               warning ("Removed duplicate interpretation after retrieving its mixins" ^
                 Position.here_list [#pos reg1, #pos reg2] ^ ":\n  " ^
-                Pretty.string_of (pretty_reg_inst (Syntax.init_pretty_global (#1 old_thys)) [] id));
+                Pretty.string_of (pretty_reg_inst ctxt0 [] id));
           in recursive_merge ((regs1, mixins1), (regs2', mixins2')) end;
-    in recursive_merge end;
+    in Library.foldl1 recursive_merge (map #2 args) end;
 );
 
 structure Local_Registrations = Proof_Data
--- a/src/Pure/Isar/method.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/Pure/Isar/method.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -321,9 +321,8 @@
 fun set_tactic ml_tactic = map_data (fn (methods, _, facts) => (methods, SOME ml_tactic, facts));
 
 fun the_tactic context =
-  (case #ml_tactic (Data.get context) of
-    SOME tac => tac
-  | NONE => raise Fail "Undefined ML tactic");
+  #ml_tactic (Data.get context)
+  |> \<^if_none>\<open>raise Fail "Undefined ML tactic"\<close>;
 
 val parse_tactic =
   Scan.state :|-- (fn context =>
--- a/src/Pure/Isar/named_target.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/Pure/Isar/named_target.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -97,7 +97,7 @@
        theory_registration = Generic_Target.theory_registration,
        locale_dependency = fn _ => error "Not possible in theory target",
        pretty = fn ctxt => [Pretty.block [Pretty.keyword1 "theory", Pretty.brk 1,
-        Pretty.str (Context.theory_name (Proof_Context.theory_of ctxt))]]}
+        Pretty.str (Context.theory_base_name (Proof_Context.theory_of ctxt))]]}
   | operations (Locale locale) =
       {define = Generic_Target.define (locale_foundation locale),
        notes = Generic_Target.notes (Generic_Target.locale_target_notes locale),
--- a/src/Pure/Isar/obtain.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/Pure/Isar/obtain.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -105,9 +105,8 @@
   let
     val frees = Frees.build (t |> Term.fold_aterms (fn Free v => Frees.add_set v | _ => I));
     val T =
-      (case Frees.get_first (fn ((x, T), _) => if x = z then SOME T else NONE) frees of
-        SOME T => T
-      | NONE => the_default dummyT (Variable.default_type ctxt z));
+      Frees.get_first (fn ((x, T), _) => if x = z then SOME T else NONE) frees
+      |> \<^if_none>\<open>the_default dummyT (Variable.default_type ctxt z)\<close>;
   in Logic.all_const T $ Term.lambda_name (y, Free (z, T)) t end;
 
 fun prepare_clause prep_var parse_prop mk_all ctxt thesis raw_vars raw_props =
--- a/src/Pure/Isar/outer_syntax.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/Pure/Isar/outer_syntax.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -112,7 +112,7 @@
 (* maintain commands *)
 
 fun add_command name cmd thy =
-  if member (op =) Thy_Header.bootstrap_thys (Context.theory_name thy) then thy
+  if member (op =) Thy_Header.bootstrap_thys (Context.theory_base_name thy) then thy
   else
     let
       val _ =
--- a/src/Pure/Isar/proof.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/Pure/Isar/proof.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -355,9 +355,8 @@
     (ctxt, statement, using, goal, before_qed, after_qed));
 
 fun find_goal state =
-  (case try current_goal state of
-    SOME ctxt_goal => ctxt_goal
-  | NONE => find_goal (close_block state handle ERROR _ => error "No proof goal"));
+  try current_goal state
+  |> \<^if_none>\<open>find_goal (close_block state handle ERROR _ => error "No proof goal")\<close>;
 
 
 
@@ -1298,9 +1297,7 @@
 );
 
 fun the_result ctxt =
-  (case Result.get ctxt of
-    NONE => error "No result of forked proof"
-  | SOME th => th);
+  Result.get ctxt |> \<^if_none>\<open>error "No result of forked proof"\<close>;
 
 val set_result = Result.put o SOME;
 val reset_result = Result.put NONE;
--- a/src/Pure/Isar/proof_context.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/Pure/Isar/proof_context.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -10,7 +10,7 @@
 sig
   val theory_of: Proof.context -> theory
   val init_global: theory -> Proof.context
-  val get_global: theory -> string -> Proof.context
+  val get_global: {long: bool} -> theory -> string -> Proof.context
   type mode
   val mode_default: mode
   val mode_pattern: mode
--- a/src/Pure/Isar/proof_display.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/Pure/Isar/proof_display.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -44,9 +44,8 @@
   (case Context.get_generic_context () of
     SOME (Context.Proof ctxt) => ctxt
   | SOME (Context.Theory thy) =>
-      (case try Syntax.init_pretty_global thy of
-        SOME ctxt => ctxt
-      | NONE => Syntax.init_pretty_global (mk_thy ()))
+      try Syntax.init_pretty_global thy
+      |> \<^if_none>\<open>Syntax.init_pretty_global (mk_thy ())\<close>
   | NONE => Syntax.init_pretty_global (mk_thy ()));
 
 fun pp_thm mk_thy th =
--- a/src/Pure/Isar/token.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/Pure/Isar/token.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -731,7 +731,7 @@
     val range = Position.range (pos, pos');
     val tok =
       if 0 <= k andalso k < Vector.length immediate_kinds then
-        Token ((s, range), (Vector.sub (immediate_kinds, k), s), Slot)
+        Token ((s, range), (Vector.nth immediate_kinds k, s), Slot)
       else
         (case explode Keyword.empty_keywords pos s of
           [tok] => tok
--- a/src/Pure/Isar/toplevel.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/Pure/Isar/toplevel.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -162,7 +162,7 @@
     Toplevel =>
       (case previous_theory_of state of
         NONE => "at top level"
-      | SOME thy => "at top level, result theory " ^ quote (Context.theory_name thy))
+      | SOME thy => "at top level, result theory " ^ quote (Context.theory_base_name thy))
   | Theory (Context.Theory _) => "in theory mode"
   | Theory (Context.Proof _) => "in local theory mode"
   | Proof _ => "in proof mode"
@@ -217,9 +217,8 @@
   |> Presentation_State.put (SOME (State (current, (NONE, NONE))));
 
 fun presentation_state ctxt =
-  (case Presentation_State.get ctxt of
-    NONE => State (node_presentation (Theory (Context.Proof ctxt)), (NONE, NONE))
-  | SOME state => state);
+  Presentation_State.get ctxt
+  |> \<^if_none>\<open>State (node_presentation (Theory (Context.Proof ctxt)), (NONE, NONE))\<close>;
 
 
 (* print state *)
--- a/src/Pure/ML/ml_antiquotations.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/Pure/ML/ml_antiquotations.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -65,7 +65,7 @@
   ML_Antiquotation.value_embedded \<^binding>\<open>theory_context\<close>
     (Args.context -- Scan.lift Parse.embedded_position >> (fn (ctxt, (name, pos)) =>
       (Theory.check {long = false} ctxt (name, pos);
-       "Proof_Context.get_global (Proof_Context.theory_of ML_context) " ^
+       "Proof_Context.get_global {long = false} (Proof_Context.theory_of ML_context) " ^
         ML_Syntax.print_string name))) #>
 
   ML_Antiquotation.inline \<^binding>\<open>context\<close>
@@ -340,11 +340,24 @@
 in end;
 
 
-(* special forms *)
+(* special forms for option type *)
 
 val _ = Theory.setup
  (ML_Antiquotation.special_form \<^binding>\<open>try\<close> "() |> Basics.try" #>
-  ML_Antiquotation.special_form \<^binding>\<open>can\<close> "() |> Basics.can");
+  ML_Antiquotation.special_form \<^binding>\<open>can\<close> "() |> Basics.can" #>
+  ML_Context.add_antiquotation \<^binding>\<open>if_none\<close> true
+    (fn _ => fn src => fn ctxt =>
+      let
+        val (s, _) = Token.syntax (Scan.lift Parse.embedded_input) src ctxt;
+        val tokenize = ML_Lex.tokenize_no_range;
+
+        val (decl, ctxt') = ML_Context.expand_antiquotes (ML_Lex.read_source s) ctxt;
+        fun decl' ctxt'' =
+          let
+            val (ml_env, ml_body) = decl ctxt'';
+            val ml_body' = tokenize "(fn SOME x => x | NONE => " @ ml_body @ tokenize ")";
+          in (ml_env, ml_body') end;
+      in (decl', ctxt') end));
 
 
 (* basic combinators *)
--- a/src/Pure/ML/ml_syntax.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/Pure/ML/ml_syntax.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -139,8 +139,14 @@
       else take (Int.max (max_len, 0)) body @ ["..."];
   in Pretty.str (quote (implode (map print_symbol body'))) end;
 
+fun pretty_string' depth = pretty_string (FixedInt.toInt (depth * 100));
+
 val _ =
   ML_system_pp (fn depth => fn _ => fn str =>
-    Pretty.to_polyml (pretty_string (FixedInt.toInt (depth * 100)) str));
+    Pretty.to_polyml (pretty_string' depth str));
+
+val _ =
+  ML_system_pp (fn depth => fn _ => fn chunks =>
+    Pretty.to_polyml (pretty_string' depth (Long_Name.implode_chunks chunks)));
 
 end;
--- a/src/Pure/Proof/extraction.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/Pure/Proof/extraction.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -180,7 +180,11 @@
         (fn Oracle (name, prop, _) => apfst (cons ((name, Position.none), SOME prop))
           | PThm (header, thm_body) => apsnd (cons (Proofterm.make_thm header thm_body))
           | _ => I);
-    val body = PBody {oracles = Oracles.make oracles, thms = PThms.make thms, proof = prf};
+    val body =
+      PBody
+       {oracles = Ord_List.make Proofterm.oracle_ord oracles,
+        thms = Ord_List.make Proofterm.thm_ord thms,
+        proof = prf};
   in Proofterm.thm_body body end;
 
 
@@ -279,7 +283,7 @@
 val add_typeof_eqns = gen_add_typeof_eqns read_condeq;
 
 fun thaw (T as TFree (a, S)) =
-      if exists_string (fn s => s = ":") a then TVar (unpack_ixn a, S) else T
+      if member_string a ":" then TVar (unpack_ixn a, S) else T
   | thaw (Type (a, Ts)) = Type (a, map thaw Ts)
   | thaw T = T;
 
--- a/src/Pure/ROOT.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/Pure/ROOT.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -19,6 +19,9 @@
 ML_file "ML/ml_system.ML";
 
 ML_file "General/basics.ML";
+ML_file "General/string.ML";
+ML_file "General/vector.ML";
+ML_file "General/array.ML";
 ML_file "General/symbol_explode.ML";
 
 ML_file "Concurrent/multithreading.ML";
@@ -72,6 +75,7 @@
 ML_file "General/ord_list.ML";
 ML_file "General/balanced_tree.ML";
 ML_file "General/linear_set.ML";
+ML_file "General/change_table.ML";
 ML_file "General/buffer.ML";
 ML_file "General/pretty.ML";
 ML_file "General/rat.ML";
@@ -96,7 +100,6 @@
 ML_file "PIDE/document_id.ML";
 ML_file "General/socket_io.ML";
 
-ML_file "General/change_table.ML";
 ML_file "General/graph.ML";
 
 ML_file "System/options.ML";
--- a/src/Pure/Syntax/lexicon.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/Pure/Syntax/lexicon.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -130,7 +130,7 @@
     String, Cartouche, Space, Comment Comment.Comment, Comment Comment.Cancel,
     Comment Comment.Latex, Dummy, EOF];
 
-fun token_kind i = Vector.sub (token_kinds, i);
+val token_kind = Vector.nth token_kinds;
 fun token_kind_index k = #1 (the (Vector.findi (fn (_, k') => k = k') token_kinds));
 
 
--- a/src/Pure/Syntax/parser.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/Pure/Syntax/parser.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -132,16 +132,16 @@
       if not new_chain then ([], lambdas)
       else
         let (*lookahead of chain's source*)
-          val ((_, from_tks), _) = Array.sub (prods, the chain);
+          val ((_, from_tks), _) = Array.nth prods (the chain);
 
           (*copy from's lookahead to chain's destinations*)
           fun copy_lookahead to =
             let
-              val ((to_nts, to_tks), ps) = Array.sub (prods, to);
+              val ((to_nts, to_tks), ps) = Array.nth prods to;
 
               val new_tks = Tokens.subtract to_tks from_tks;  (*added lookahead tokens*)
               val to_tks' = Tokens.merge (to_tks, new_tks);
-              val _ = Array.update (prods, to, ((to_nts, to_tks'), ps));
+              val _ = Array.upd prods to ((to_nts, to_tks'), ps);
             in tokens_add (to, new_tks) end;
 
           val tos = chains_all_succs chains' [lhs];
@@ -170,7 +170,7 @@
           else (NONE, nts_insert nt nts);
 
     (*get all known starting tokens for a nonterminal*)
-    fun starts_for_nt nt = snd (fst (Array.sub (prods, nt)));
+    fun starts_for_nt nt = snd (fst (Array.nth prods nt));
 
     (*update prods, lookaheads, and lambdas according to new lambda NTs*)
     val (added_starts', lambdas') =
@@ -180,7 +180,7 @@
           | propagate_lambda (l :: ls) added_starts lambdas =
               let
                 (*get lookahead for lambda NT*)
-                val ((dependent, l_starts), _) = Array.sub (prods, l);
+                val ((dependent, l_starts), _) = Array.nth prods l;
 
                 (*check productions whose lookahead may depend on lambda NT*)
                 fun examine_prods [] add_lambda nt_dependencies added_tks nt_prods =
@@ -221,7 +221,7 @@
                 (*check each NT whose lookahead depends on new lambda NT*)
                 fun process_nts nt (added_lambdas, added_starts) =
                   let
-                    val ((old_nts, old_tks), nt_prods) = Array.sub (prods, nt);
+                    val ((old_nts, old_tks), nt_prods) = Array.nth prods nt;
 
                     (*existing productions whose lookahead may depend on l*)
                     val tk_prods = prods_lookup nt_prods (get_start l_starts);
@@ -235,7 +235,7 @@
                     val new_tks = Tokens.merge (old_tks, added_tks);
 
                     val added_lambdas' = added_lambdas |> add_lambda ? cons nt;
-                    val _ = Array.update (prods, nt, ((new_nts, new_tks), nt_prods'));
+                    val _ = Array.upd prods nt ((new_nts, new_tks), nt_prods');
                       (*N.B. that because the tks component
                         is used to access existing
                         productions we have to add new
@@ -279,9 +279,9 @@
           (*add lhs NT to list of dependent NTs in lookahead*)
           fun add_nts nt initial =
             (if initial then
-              let val ((old_nts, old_tks), ps) = Array.sub (prods, nt) in
+              let val ((old_nts, old_tks), ps) = Array.nth prods nt in
                 if nts_member old_nts lhs then ()
-                else Array.update (prods, nt, ((nts_insert lhs old_nts, old_tks), ps))
+                else Array.upd prods nt ((nts_insert lhs old_nts, old_tks), ps)
               end
              else (); false);
 
@@ -290,7 +290,7 @@
           fun add_tks [] added = added
             | add_tks (nt :: nts) added =
                 let
-                  val ((old_nts, old_tks), nt_prods) = Array.sub (prods, nt);
+                  val ((old_nts, old_tks), nt_prods) = Array.nth prods nt;
 
                   val new_tks = Tokens.subtract old_tks start_tks;
 
@@ -306,9 +306,7 @@
                     |> nt = lhs ? Tokens.fold store start_tks';
                   val _ =
                     if not changed andalso Tokens.is_empty new_tks then ()
-                    else
-                      Array.update
-                        (prods, nt, ((old_nts, Tokens.merge (old_tks, new_tks)), nt_prods'));
+                    else Array.upd prods nt ((old_nts, Tokens.merge (old_tks, new_tks)), nt_prods');
                 in add_tks nts (tokens_add (nt, new_tks) added) end;
           val _ = nts_fold add_nts start_nts true;
         in add_tks (chains_all_succs chains' [lhs]) [] end;
@@ -365,7 +363,7 @@
                 (*copy existing productions for new starting tokens*)
                 fun process_nts nt =
                   let
-                    val ((nts, tks), nt_prods) = Array.sub (prods, nt);
+                    val ((nts, tks), nt_prods) = Array.nth prods nt;
 
                     val tk_prods = prods_lookup nt_prods key;
 
@@ -379,10 +377,10 @@
 
                     val added_tks = Tokens.subtract tks new_tks;
                     val tks' = Tokens.merge (tks, added_tks);
-                    val _ = Array.update (prods, nt, ((nts, tks'), nt_prods''));
+                    val _ = Array.upd prods nt ((nts, tks'), nt_prods'');
                   in tokens_add (nt, added_tks) end;
 
-                val ((dependent, _), _) = Array.sub (prods, changed_nt);
+                val ((dependent, _), _) = Array.nth prods changed_nt;
               in add_starts (starts @ nts_fold process_nts dependent []) end;
       in add_starts added_starts' end;
   in (chains', lambdas') end;
@@ -407,7 +405,7 @@
     fun prod_of_chain from = ([Nonterminal (from, ~1)], "", ~1);
 
     fun pretty_prod (name, tag) =
-      (prods_content (#2 (Vector.sub (prods, tag))) @ map prod_of_chain (chains_preds chains tag))
+      (prods_content (#2 (Vector.nth prods tag)) @ map prod_of_chain (chains_preds chains tag))
       |> map (fn (symbs, const, p) =>
           Pretty.block (Pretty.breaks
             (Pretty.str (name ^ print_pri p ^ " =") :: map pretty_symb symbs @ pretty_const const)));
@@ -476,7 +474,7 @@
         val prods' =
           let
             fun get_prod i =
-              if i < nt_count then Vector.sub (prods, i)
+              if i < nt_count then Vector.nth prods i
               else nt_gram_empty;
           in Array.tabulate (nt_count', get_prod) end;
 
@@ -574,7 +572,7 @@
   filter
     (fn (_, _, _, Nonterminal (B, prec) :: _, _, _) => A = B andalso prec <= max_prec
       | _ => false)
-    (Array.sub (Estate, j));
+    (Array.nth Estate j);
 
 
 fun movedot_term c (A, j, ts, Terminal a :: sa, id, i) =
@@ -601,7 +599,7 @@
     fun token_prods prods =
       fold cons (prods_lookup prods tk) #>
       fold cons (prods_lookup prods Lexicon.dummy);
-    fun nt_prods nt = #2 (Vector.sub (prods, nt));
+    fun nt_prods nt = #2 (Vector.nth prods nt);
   in fold (token_prods o nt_prods) (chains_all_preds chains [nt]) [] end;
 
 
@@ -652,7 +650,7 @@
 
 
 fun produce gram stateset i indata prev_token =
-  (case Array.sub (stateset, i) of
+  (case Array.nth stateset i of
     [] =>
       let
         val toks = if Lexicon.is_eof prev_token then indata else prev_token :: indata;
@@ -677,8 +675,8 @@
       | c :: cs =>
           let
             val (si, sii) = PROCESSS gram stateset i c s;
-            val _ = Array.update (stateset, i, si);
-            val _ = Array.update (stateset, i + 1, sii);
+            val _ = Array.upd stateset i si;
+            val _ = Array.upd stateset (i + 1) sii;
           in produce gram stateset (i + 1) cs c end));
 
 
@@ -693,7 +691,7 @@
     val S0 = [(~1, 0, [], [Nonterminal (start_tag, 0), Terminal Lexicon.eof], "", 0)];
     val s = length indata + 1;
     val Estate = Array.array (s, []);
-    val _ = Array.update (Estate, 0, S0);
+    val _ = Array.upd Estate 0 S0;
   in
     get_trees (produce gram Estate 0 indata Lexicon.eof)
   end;
--- a/src/Pure/System/bash.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/Pure/System/bash.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -43,7 +43,7 @@
           | "\r" => "$'\\r'"
           | _ =>
               if Symbol.is_ascii_letter ch orelse Symbol.is_ascii_digit ch orelse
-                exists_string (fn c => c = ch) "+,-./:_" then ch
+                member_string "+,-./:_" ch then ch
               else if c < 16 then "$'\\x0" ^ Int.fmt StringCvt.HEX c ^ "'"
               else if c < 32 orelse c >= 127 then "$'\\x" ^ Int.fmt StringCvt.HEX c ^ "'"
               else "\\" ^ ch)
--- a/src/Pure/Thy/export_theory.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/Pure/Thy/export_theory.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -264,7 +264,7 @@
       in ((sorts @ typargs, args, prop), proof) end;
 
     fun standard_prop_of thm =
-      standard_prop Name.context (Sortset.dest (Thm.extra_shyps thm)) (Thm.full_prop_of thm);
+      standard_prop Name.context (Thm.extra_shyps thm) (Thm.full_prop_of thm);
 
     val encode_prop =
       let open XML.Encode Term_XML.Encode
@@ -414,7 +414,7 @@
         get_dependencies parents thy |> map_index (fn (i, dep) =>
           let
             val xname = string_of_int (i + 1);
-            val name = Long_Name.implode [Context.theory_name thy, xname];
+            val name = Long_Name.implode [Context.theory_base_name thy, xname];
             val markup = make_entity_markup name xname (#pos (#1 dep)) (#serial (#1 dep));
             val body = encode_locale_dependency dep;
           in XML.Elem (markup, body) end)
--- a/src/Pure/Thy/latex.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/Pure/Thy/latex.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -85,7 +85,7 @@
       | "\n" => "\\isanewline\n"
       | s =>
           s
-          |> exists_string (fn s' => s = s') "\"#$%&',-<>\\^_`{}~" ? enclose "{\\char`\\" "}"
+          |> member_string "\"#$%&',-<>\\^_`{}~" s ? enclose "{\\char`\\" "}"
           |> suffix "{\\kern0pt}");
 
 fun output_ascii_breakable sep =
@@ -203,7 +203,7 @@
   let
     val _ =
       citations |> List.app (fn (s, pos) =>
-        if exists_string (fn c => c = ",") s
+        if member_string s ","
         then error ("Single citation expected, without commas" ^ Position.here pos)
         else ());
     val citations' = space_implode "," (map #1 citations);
--- a/src/Pure/Thy/thy_info.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/Pure/Thy/thy_info.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -65,7 +65,7 @@
     else
       let
         val keywords = Thy_Header.get_keywords thy;
-        val thy_name = Context.theory_name thy;
+        val thy_name = Context.theory_base_name thy;
         val latex = Document_Output.present_thy options keywords thy_name segments;
       in
         if Options.string options "document" = "false" then ()
--- a/src/Pure/Tools/find_theorems.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/Pure/Tools/find_theorems.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -335,7 +335,7 @@
 
 val index_ord = option_ord (K EQUAL);
 val hidden_ord = bool_ord o apply2 Long_Name.is_hidden;
-val qual_ord = int_ord o apply2 Long_Name.qualification;
+val qual_ord = int_ord o apply2 Long_Name.count;
 val txt_ord = int_ord o apply2 size;
 
 fun nicer_name ((a, x), i) ((b, y), j) =
--- a/src/Pure/Tools/named_theorems.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/Pure/Tools/named_theorems.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -59,7 +59,7 @@
 
 fun clear name = map_entry name (K Thm.item_net);
 
-fun add_thm name th = map_entry name (Item_Net.update (Thm.trim_context th));
+fun add_thm name = map_entry name o Item_Net.update o Thm.trim_context;
 fun del_thm name = map_entry name o Item_Net.remove;
 
 val add = Thm.declaration_attribute o add_thm;
--- a/src/Pure/Tools/named_thms.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/Pure/Tools/named_thms.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -1,7 +1,7 @@
 (*  Title:      Pure/Tools/named_thms.ML
     Author:     Makarius
 
-Named collections of theorems in canonical order.
+Named collections of theorems in canonical (reverse) order: OLD VERSION.
 *)
 
 signature NAMED_THMS =
@@ -27,10 +27,10 @@
 
 val member = Item_Net.member o Data.get o Context.Proof;
 
-val content = Item_Net.content o Data.get;
+fun content context = map (Thm.transfer'' context) (Item_Net.content (Data.get context));
 val get = content o Context.Proof;
 
-val add_thm = Data.map o Item_Net.update;
+val add_thm = Data.map o Item_Net.update o Thm.trim_context;
 val del_thm = Data.map o Item_Net.remove;
 
 val add = Thm.declaration_attribute add_thm;
--- a/src/Pure/Tools/rule_insts.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/Pure/Tools/rule_insts.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -280,7 +280,7 @@
     fun cvar xi = Thm.cterm_of ctxt (Var (xi, propT));
     val revcut_rl' =
       Drule.revcut_rl |> Drule.instantiate_normalize
-        (TVars.empty, Vars.make [(var "V", cvar ("V", maxidx + 1)), (var "W", cvar ("W", maxidx + 1))]);
+        (TVars.empty, Vars.make2 (var "V", cvar ("V", maxidx + 1)) (var "W", cvar ("W", maxidx + 1)));
   in
     (case Seq.list_of
       (Thm.bicompose (SOME ctxt) {flatten = true, match = false, incremented = false}
--- a/src/Pure/Tools/thy_deps.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/Pure/Tools/thy_deps.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -27,8 +27,8 @@
         SOME Bs => (fn thy => exists (fn B => rel (B, thy)) Bs)
       | NONE => K true);
     fun node thy =
-      ((Context.theory_name thy, Graph_Display.content_node (Context.theory_name thy) []),
-        map Context.theory_name (filter pred (Theory.parents_of thy)));
+      ((Context.theory_base_name thy, Graph_Display.content_node (Context.theory_base_name thy) []),
+        map Context.theory_base_name (filter pred (Theory.parents_of thy)));
   in map node (filter pred (Theory.nodes_of (Proof_Context.theory_of ctxt))) end;
 
 val thy_deps =
--- a/src/Pure/axclass.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/Pure/axclass.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -74,6 +74,8 @@
       (*constant name ~> type constructor ~> (constant name, equation)*)
     (string * string) Symtab.table (*constant name ~> (constant name, type constructor)*)};
 
+fun rep_data (Data args) = args;
+
 fun make_data (axclasses, params, inst_params) =
   Data {axclasses = axclasses, params = params, inst_params = inst_params};
 
@@ -81,22 +83,23 @@
 (
   type T = data;
   val empty = make_data (Symtab.empty, [], (Symtab.empty, Symtab.empty));
-  fun merge old_thys
-      (Data {axclasses = axclasses1, params = params1, inst_params = inst_params1},
-       Data {axclasses = axclasses2, params = params2, inst_params = inst_params2}) =
+  fun merge args =
     let
-      val old_ctxt = Syntax.init_pretty_global (fst old_thys);
+      val ctxt0 = Syntax.init_pretty_global (#1 (hd args));
 
-      val axclasses' = Symtab.merge (K true) (axclasses1, axclasses2);
-      val params' =
+      fun merge_params (params1, params2) =
         if null params1 then params2
         else
-          fold_rev (fn p => if member (op =) params1 p then I else add_param old_ctxt p)
+          fold_rev (fn p => if member (op =) params1 p then I else add_param ctxt0 p)
             params2 params1;
 
-      val inst_params' =
+      fun merge_inst_params (inst_params1, inst_params2) =
         (Symtab.join (K (Symtab.merge (K true))) (#1 inst_params1, #1 inst_params2),
           Symtab.merge (K true) (#2 inst_params1, #2 inst_params2));
+
+      val axclasses' = Library.foldl1 (Symtab.merge (K true)) (map (#axclasses o rep_data o #2) args);
+      val params' = Library.foldl1 merge_params (map (#params o rep_data o #2) args);
+      val inst_params' = Library.foldl1 merge_inst_params (map (#inst_params o rep_data o #2) args);
     in make_data (axclasses', params', inst_params') end;
 );
 
@@ -116,11 +119,11 @@
   map_data (fn (axclasses, params, inst_params) =>
     (axclasses, params, f inst_params));
 
-val rep_data = Data.get #> (fn Data args => args);
+val rep_theory_data = Data.get #> rep_data;
 
-val axclasses_of = #axclasses o rep_data;
-val params_of = #params o rep_data;
-val inst_params_of = #inst_params o rep_data;
+val axclasses_of = #axclasses o rep_theory_data;
+val params_of = #params o rep_theory_data;
+val inst_params_of = #inst_params o rep_theory_data;
 
 
 (* axclasses with parameters *)
@@ -212,17 +215,14 @@
 (* declaration and definition of instances of overloaded constants *)
 
 fun inst_tyco_of thy (c, T) =
-  (case get_inst_tyco (Sign.consts_of thy) (c, T) of
-    SOME tyco => tyco
-  | NONE => error ("Illegal type for instantiation of class parameter: " ^
-      quote (c ^ " :: " ^ Syntax.string_of_typ_global thy T)));
+  get_inst_tyco (Sign.consts_of thy) (c, T)
+  |> \<^if_none>\<open>
+    error ("Illegal type for instantiation of class parameter: " ^
+      quote (c ^ " :: " ^ Syntax.string_of_typ_global thy T))\<close>;
 
 fun declare_overloaded (c, T) thy =
   let
-    val class =
-      (case class_of_param thy c of
-        SOME class => class
-      | NONE => error ("Not a class parameter: " ^ quote c));
+    val class = class_of_param thy c |> \<^if_none>\<open>error ("Not a class parameter: " ^ quote c)\<close>;
     val tyco = inst_tyco_of thy (c, T);
     val name_inst = instance_name (tyco, class) ^ "_inst";
     val c' = instance_name (tyco, c);
--- a/src/Pure/conjunction.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/Pure/conjunction.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -102,7 +102,7 @@
 fun intr tha thb =
   Thm.implies_elim
     (Thm.implies_elim
-      (Thm.instantiate (TVars.empty, Vars.make [(vA, Thm.cprop_of tha), (vB, Thm.cprop_of thb)])
+      (Thm.instantiate (TVars.empty, Vars.make2 (vA, Thm.cprop_of tha) (vB, Thm.cprop_of thb))
         conjunctionI)
     tha)
   thb;
@@ -111,7 +111,7 @@
   let
     val (A, B) = dest_conjunction (Thm.cprop_of th)
       handle TERM (msg, _) => raise THM (msg, 0, [th]);
-    val inst = Thm.instantiate (TVars.empty, Vars.make [(vA, A), (vB, B)]);
+    val inst = Thm.instantiate (TVars.empty, Vars.make2 (vA, A) (vB, B));
   in
    (Thm.implies_elim (inst conjunctionD1) th,
     Thm.implies_elim (inst conjunctionD2) th)
--- a/src/Pure/context.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/Pure/context.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -20,26 +20,29 @@
   sig
     val theory_of: Proof.context -> theory
     val init_global: theory -> Proof.context
-    val get_global: theory -> string -> Proof.context
+    val get_global: {long: bool} -> theory -> string -> Proof.context
   end
 end;
 
 signature CONTEXT =
 sig
   include BASIC_CONTEXT
+  (*theory data*)
+  type data_kind = int
+  val data_kinds: unit -> (data_kind * Position.T) list
   (*theory context*)
+  type id = int
   type theory_id
   val theory_id: theory -> theory_id
   val timing: bool Unsynchronized.ref
   val parents_of: theory -> theory list
   val ancestors_of: theory -> theory list
   val theory_id_ord: theory_id ord
-  val theory_id_long_name: theory_id -> string
-  val theory_id_name: theory_id -> string
+  val theory_id_name: {long: bool} -> theory_id -> string
   val theory_long_name: theory -> string
-  val theory_name: theory -> string
-  val theory_name': {long: bool} -> theory -> string
-  val theory_identifier: theory -> serial
+  val theory_base_name: theory -> string
+  val theory_name: {long: bool} -> theory -> string
+  val theory_identifier: theory -> id
   val PureN: string
   val pretty_thy: theory -> Pretty.T
   val pretty_abbrev_thy: theory -> Pretty.T
@@ -52,7 +55,7 @@
   val subthy: theory * theory -> bool
   val trace_theories: bool Unsynchronized.ref
   val theories_trace: unit -> {active_positions: Position.T list, active: int, total: int}
-  val join_thys: theory * theory -> theory
+  val join_thys: theory list -> theory
   val begin_thy: string -> theory list -> theory
   val finish_thy: theory -> theory
   val theory_data_sizeof1: theory -> (Position.T * int) list
@@ -96,15 +99,15 @@
   include CONTEXT
   structure Theory_Data:
   sig
-    val declare: Position.T -> Any.T -> (theory * theory -> Any.T * Any.T -> Any.T) -> serial
-    val get: serial -> (Any.T -> 'a) -> theory -> 'a
-    val put: serial -> ('a -> Any.T) -> 'a -> theory -> theory
+    val declare: Position.T -> Any.T -> ((theory * Any.T) list -> Any.T) -> data_kind
+    val get: data_kind -> (Any.T -> 'a) -> theory -> 'a
+    val put: data_kind -> ('a -> Any.T) -> 'a -> theory -> theory
   end
   structure Proof_Data:
   sig
-    val declare: (theory -> Any.T) -> serial
-    val get: serial -> (Any.T -> 'a) -> Proof.context -> 'a
-    val put: serial -> ('a -> Any.T) -> 'a -> Proof.context -> Proof.context
+    val declare: (theory -> Any.T) -> data_kind
+    val get: data_kind -> (Any.T -> 'a) -> Proof.context -> 'a
+    val put: data_kind -> ('a -> Any.T) -> 'a -> Proof.context -> Proof.context
   end
 end;
 
@@ -113,24 +116,39 @@
 
 (*** theory context ***)
 
+(* theory data *)
+
 (*private copy avoids potential conflict of table exceptions*)
 structure Datatab = Table(type key = int val ord = int_ord);
 
+type data_kind = int;
+val data_kind = Counter.make ();
+
 
 (** datatype theory **)
 
-type stage = int * int Synchronized.var;
-fun init_stage () : stage = (0, Synchronized.var "Context.stage" 1);
-fun next_stage ((_, state): stage) = (Synchronized.change_result state (fn n => (n, n + 1)), state);
+(* implicit state *)
+
+type state = {stage: int} Synchronized.var;
+
+fun make_state () : state =
+  Synchronized.var "Context.state" {stage = 0};
+
+fun next_stage (state: state) =
+  Synchronized.change_result state (fn {stage} => (stage + 1, {stage = stage + 1}));
+
+
+(* theory_id *)
+
+type id = int;
+val new_id = Counter.make ();
 
 abstype theory_id =
   Theory_Id of
-   (*identity*)
-   {id: serial,                   (*identifier*)
-    ids: Intset.T} *              (*cumulative identifiers -- symbolic body content*)
-   (*history*)
-   {name: string,                 (*official theory name*)
-    stage: stage option}          (*index and counter for anonymous updates*)
+   {id: id,                       (*identifier*)
+    ids: Intset.T,                (*cumulative identifiers -- symbolic body content*)
+    name: string,                 (*official theory name*)
+    stage: int}                   (*index for anonymous updates*)
 with
 
 fun rep_theory_id (Theory_Id args) = args;
@@ -138,8 +156,13 @@
 
 end;
 
+
+(* theory *)
+
 datatype theory =
   Theory of
+   (*allocation state*)
+   state *
    (*identity*)
    {theory_id: theory_id,
     token: Position.T Unsynchronized.ref} *
@@ -153,26 +176,28 @@
 
 fun rep_theory (Theory args) = args;
 
-val theory_identity = #1 o rep_theory;
+val state_of = #1 o rep_theory;
+val theory_identity = #2 o rep_theory;
 val theory_id = #theory_id o theory_identity;
-val identity_of_id = #1 o rep_theory_id;
-val identity_of = identity_of_id o theory_id;
-val history_of_id = #2 o rep_theory_id;
-val history_of = history_of_id o theory_id;
-val ancestry_of = #2 o rep_theory;
-val data_of = #3 o rep_theory;
+val identity_of = rep_theory_id o theory_id;
+val ancestry_of = #3 o rep_theory;
+val data_of = #4 o rep_theory;
 
-fun make_identity id ids = {id = id, ids = ids};
-fun make_history name = {name = name, stage = SOME (init_stage ())};
 fun make_ancestry parents ancestors = {parents = parents, ancestors = ancestors};
 
-val theory_id_ord = int_ord o apply2 (#id o identity_of_id);
-val theory_id_long_name = #name o history_of_id;
-val theory_id_name = Long_Name.base_name o theory_id_long_name;
-val theory_long_name = #name o history_of;
-val theory_name = Long_Name.base_name o theory_long_name;
-fun theory_name' {long} = if long then theory_long_name else theory_name;
-val theory_identifier = #id o identity_of_id o theory_id;
+fun stage_final stage = stage = 0;
+
+val theory_id_stage = #stage o rep_theory_id;
+val theory_id_final = stage_final o theory_id_stage;
+val theory_id_ord = int_ord o apply2 (#id o rep_theory_id);
+fun theory_id_name {long} thy_id =
+  let val name = #name (rep_theory_id thy_id)
+  in if long then name else Long_Name.base_name name end;
+
+val theory_long_name = #name o identity_of;
+val theory_base_name = Long_Name.base_name o theory_long_name;
+fun theory_name {long} = if long then theory_long_name else theory_base_name;
+val theory_identifier = #id o identity_of;
 
 val parents_of = #parents o ancestry_of;
 val ancestors_of = #ancestors o ancestry_of;
@@ -183,9 +208,10 @@
 val PureN = "Pure";
 
 fun display_name thy_id =
-  (case history_of_id thy_id of
-    {name, stage = NONE} => name
-  | {name, stage = SOME (i, _)} => name ^ ":" ^ string_of_int i);
+  let
+    val name = theory_id_name {long = false} thy_id;
+    val final = theory_id_final thy_id;
+  in if final then name else name ^ ":" ^ string_of_int (theory_id_stage thy_id) end;
 
 fun display_names thy =
   let
@@ -205,31 +231,24 @@
   in Pretty.str_list "{" "}" abbrev end;
 
 fun get_theory long thy name =
-  if theory_name' long thy <> name then
-    (case find_first (fn thy' => theory_name' long thy' = name) (ancestors_of thy) of
+  if theory_name long thy <> name then
+    (case find_first (fn thy' => theory_name long thy' = name) (ancestors_of thy) of
       SOME thy' => thy'
     | NONE => error ("Unknown ancestor theory " ^ quote name))
-  else if is_none (#stage (history_of thy)) then thy
+  else if theory_id_final (theory_id thy) then thy
   else error ("Unfinished theory " ^ quote name);
 
 
-(* build ids *)
+(* identity *)
 
-val merge_ids =
-  apply2 (theory_id #> rep_theory_id #> #1) #>
-  (fn ({id = id1, ids = ids1, ...}, {id = id2, ids = ids2, ...}) =>
-    Intset.merge (ids1, ids2)
-    |> Intset.insert id1
-    |> Intset.insert id2);
+fun merge_ids thys =
+  fold (identity_of #> (fn {id, ids, ...} => fn acc => Intset.merge (acc, ids) |> Intset.insert id))
+    thys Intset.empty;
 
-
-(* equality and inclusion *)
-
-val eq_thy_id = op = o apply2 (#id o identity_of_id);
+val eq_thy_id = op = o apply2 (#id o rep_theory_id);
 val eq_thy = op = o apply2 (#id o identity_of);
 
-val proper_subthy_id =
-  apply2 (rep_theory_id #> #1) #> (fn ({id, ...}, {ids, ...}) => Intset.member ids id);
+val proper_subthy_id = apply2 rep_theory_id #> (fn ({id, ...}, {ids, ...}) => Intset.member ids id);
 val proper_subthy = proper_subthy_id o apply2 theory_id;
 
 fun subthy_id p = eq_thy_id p orelse proper_subthy_id p;
@@ -240,7 +259,7 @@
 
 fun eq_thy_consistent (thy1, thy2) =
   eq_thy (thy1, thy2) orelse
-    (theory_name thy1 = theory_name thy2 andalso
+    (theory_base_name thy1 = theory_base_name thy2 andalso
       raise THEORY ("Duplicate theory name", [thy1, thy2]));
 
 fun extend_ancestors thy thys =
@@ -250,6 +269,11 @@
 
 val merge_ancestors = merge eq_thy_consistent;
 
+val eq_ancestry =
+  apply2 ancestry_of #>
+    (fn ({parents, ancestors}, {parents = parents', ancestors = ancestors'}) =>
+      eq_list eq_thy (parents, parents') andalso eq_list eq_thy (ancestors, ancestors'));
+
 
 
 (** theory data **)
@@ -263,33 +287,53 @@
 type kind =
  {pos: Position.T,
   empty: Any.T,
-  merge: theory * theory -> Any.T * Any.T -> Any.T};
+  merge: (theory * Any.T) list -> Any.T};
 
 val kinds = Synchronized.var "Theory_Data" (Datatab.empty: kind Datatab.table);
 
-fun invoke name f k x =
+fun the_kind k =
   (case Datatab.lookup (Synchronized.value kinds) k of
-    SOME kind =>
-      if ! timing andalso name <> "" then
-        Timing.cond_timeit true ("Theory_Data." ^ name ^ Position.here (#pos kind))
-          (fn () => f kind x)
-      else f kind x
+    SOME kind => kind
   | NONE => raise Fail "Invalid theory data identifier");
 
 in
 
-fun invoke_pos k = invoke "" (K o #pos) k ();
-fun invoke_empty k = invoke "" (K o #empty) k ();
-fun invoke_merge thys = invoke "merge" (fn kind => #merge kind thys);
+fun data_kinds () =
+  Datatab.fold_rev (fn (k, {pos, ...}) => cons (k, pos)) (Synchronized.value kinds) [];
+
+val invoke_pos = #pos o the_kind;
+val invoke_empty = #empty o the_kind;
 
-fun declare_theory_data pos empty merge =
+fun invoke_merge kind args =
+  if ! timing then
+    Timing.cond_timeit true ("Theory_Data.merge" ^ Position.here (#pos kind))
+      (fn () => #merge kind args)
+  else #merge kind args;
+
+fun declare_data pos empty merge =
   let
-    val k = serial ();
+    val k = data_kind ();
     val kind = {pos = pos, empty = empty, merge = merge};
     val _ = Synchronized.change kinds (Datatab.update (k, kind));
   in k end;
 
-fun merge_data thys = Datatab.join (invoke_merge thys);
+fun lookup_data k thy = Datatab.lookup (data_of thy) k;
+
+fun get_data k thy =
+  (case lookup_data k thy of
+    SOME x => x
+  | NONE => invoke_empty k);
+
+fun merge_data [] = Datatab.empty
+  | merge_data [thy] = data_of thy
+  | merge_data thys =
+      let
+        fun merge (k, kind) data =
+          (case map_filter (fn thy => lookup_data k thy |> Option.map (pair thy)) thys of
+            [] => data
+          | [(_, x)] => Datatab.default (k, x) data
+          | args => Datatab.update (k, invoke_merge kind args) data);
+      in Datatab.fold merge (Synchronized.value kinds) (data_of (hd thys)) end;
 
 end;
 
@@ -331,11 +375,11 @@
      total = length trace}
   end;
 
-fun create_thy ids history ancestry data =
+fun create_thy state ids name stage ancestry data =
   let
-    val theory_id = make_theory_id (make_identity (serial ()) ids, history);
-    val token = make_token ();
-  in Theory ({theory_id = theory_id, token = token}, ancestry, data) end;
+    val theory_id = make_theory_id {id = new_id (), ids = ids, name = name, stage = stage};
+    val identity = {theory_id = theory_id, token = make_token ()};
+  in Theory (state, identity, ancestry, data) end;
 
 end;
 
@@ -343,102 +387,80 @@
 (* primitives *)
 
 val pre_pure_thy =
-  create_thy Intset.empty (make_history PureN) (make_ancestry [] []) Datatab.empty;
+  let
+    val state = make_state ();
+    val stage = next_stage state;
+  in create_thy state Intset.empty PureN stage (make_ancestry [] []) Datatab.empty end;
 
 local
 
 fun change_thy finish f thy =
   let
-    val ({id, ids}, {name, stage}) = rep_theory_id (theory_id thy);
-    val Theory (_, ancestry, data) = thy;
+    val {name, stage, ...} = identity_of thy;
+    val Theory (state, _, ancestry, data) = thy;
     val ancestry' =
-      if is_none stage
+      if stage_final stage
       then make_ancestry [thy] (extend_ancestors thy (ancestors_of thy))
       else ancestry;
-    val history' = {name = name, stage = if finish then NONE else Option.map next_stage stage};
-    val ids' = Intset.insert id ids;
+    val ids' = merge_ids [thy];
+    val stage' = if finish then 0 else next_stage state;
     val data' = f data;
-  in create_thy ids' history' ancestry' data' end;
+  in create_thy state ids' name stage' ancestry' data' end;
 
 in
 
 val update_thy = change_thy false;
-val extend_thy = update_thy I;
 val finish_thy = change_thy true I;
 
 end;
 
 
-(* join: anonymous theory nodes *)
-
-local
-
-fun bad_join (thy1, thy2) = raise THEORY ("Cannot join theories", [thy1, thy2]);
+(* join: unfinished theory nodes *)
 
-fun join_history thys =
-  apply2 history_of thys |>
-  (fn ({name, stage}, {name = name', stage = stage'}) =>
-    if name = name' andalso is_some stage andalso is_some stage' then
-      {name = name, stage = Option.map next_stage stage}
-    else bad_join thys);
+fun join_thys [] = raise List.Empty
+  | join_thys thys =
+      let
+        val thy0 = hd thys;
+        val name0 = theory_long_name thy0;
+        val state0 = state_of thy0;
 
-fun join_ancestry thys =
-  apply2 ancestry_of thys |>
-  (fn (ancestry as {parents, ancestors}, {parents = parents', ancestors = ancestors'}) =>
-    if eq_list eq_thy (parents, parents') andalso eq_list eq_thy (ancestors, ancestors')
-    then ancestry else bad_join thys);
-
-in
+        fun ok thy =
+          not (theory_id_final (theory_id thy)) andalso
+          theory_long_name thy = name0 andalso
+          eq_ancestry (thy0, thy);
+        val _ =
+          (case filter_out ok thys of
+            [] => ()
+          | bad => raise THEORY ("Cannot join theories", bad));
 
-fun join_thys thys =
-  let
-    val ids = merge_ids thys;
-    val history = join_history thys;
-    val ancestry = join_ancestry thys;
-    val data = merge_data thys (apply2 data_of thys);
-  in create_thy ids history ancestry data end;
-
-end;
+        val stage = next_stage state0;
+        val ids = merge_ids thys;
+        val data = merge_data thys;
+      in create_thy state0 ids name0 stage (ancestry_of thy0) data end;
 
 
-(* merge: named theory nodes *)
-
-local
+(* merge: finished theory nodes *)
 
-fun merge_thys thys =
-  let
-    val ids = merge_ids thys;
-    val history = make_history "";
-    val ancestry = make_ancestry [] [];
-    val data = merge_data thys (apply2 data_of thys);
-  in create_thy ids history ancestry data end;
-
-fun maximal_thys thys =
-  thys |> filter_out (fn thy => exists (fn thy' => proper_subthy (thy, thy')) thys);
-
-in
+fun make_parents thys =
+  let val thys' = distinct eq_thy thys
+  in thys' |> filter_out (fn thy => exists (fn thy' => proper_subthy (thy, thy')) thys') end;
 
 fun begin_thy name imports =
   if name = "" then error ("Bad theory name: " ^ quote name)
+  else if null imports then error "Missing theory imports"
   else
     let
-      val parents = maximal_thys (distinct eq_thy imports);
+      val parents = make_parents imports;
       val ancestors =
-        Library.foldl merge_ancestors ([], map ancestors_of parents)
+        Library.foldl1 merge_ancestors (map ancestors_of parents)
         |> fold extend_ancestors parents;
+      val ancestry = make_ancestry parents ancestors;
 
-      val thy0 =
-        (case parents of
-          [] => error "Missing theory imports"
-        | [thy] => extend_thy thy
-        | thy :: thys => Library.foldl merge_thys (thy, thys));
-      val ({ids, ...}, _) = rep_theory_id (theory_id thy0);
-
-      val history = make_history name;
-      val ancestry = make_ancestry parents ancestors;
-    in create_thy ids history ancestry (data_of thy0) |> tap finish_thy end;
-
-end;
+      val state = make_state ();
+      val stage = next_stage state;
+      val ids = merge_ids parents;
+      val data = merge_data parents;
+    in create_thy state ids name stage ancestry data |> tap finish_thy end;
 
 
 (* theory data *)
@@ -446,14 +468,11 @@
 structure Theory_Data =
 struct
 
-val declare = declare_theory_data;
+val declare = declare_data;
 
-fun get k dest thy =
-  (case Datatab.lookup (data_of thy) k of
-    SOME x => x
-  | NONE => invoke_empty k) |> dest;
+fun get k dest thy = dest (get_data k thy);
 
-fun put k mk x = update_thy (Datatab.update (k, mk x));
+fun put k make x = update_thy (Datatab.update (k, make x));
 
 fun sizeof1 k thy =
   Datatab.lookup (data_of thy) k |> Option.map ML_Heap.sizeof1;
@@ -509,7 +528,7 @@
 struct
   fun theory_of (Proof.Context (_, thy)) = thy;
   fun init_global thy = Proof.Context (init_data thy, thy);
-  fun get_global thy name = init_global (get_theory {long = false} thy name);
+  fun get_global long thy name = init_global (get_theory long thy name);
 end;
 
 structure Proof_Data =
@@ -517,7 +536,7 @@
 
 fun declare init =
   let
-    val k = serial ();
+    val k = data_kind ();
     val _ = Synchronized.change kinds (Datatab.update (k, init));
   in k end;
 
@@ -526,8 +545,8 @@
     SOME x => x
   | NONE => init_fallback k thy) |> dest;
 
-fun put k mk x (Proof.Context (data, thy)) =
-  Proof.Context (Datatab.update (k, mk x) data, thy);
+fun put k make x (Proof.Context (data, thy)) =
+  Proof.Context (Datatab.update (k, make x) data, thy);
 
 end;
 
@@ -633,7 +652,7 @@
 sig
   type T
   val empty: T
-  val merge: theory * theory -> T * T -> T
+  val merge: (theory * T) list -> T
 end;
 
 signature THEORY_DATA_ARGS =
@@ -662,7 +681,7 @@
     Context.Theory_Data.declare
       pos
       (Data Data.empty)
-      (fn thys => fn (Data x1, Data x2) => Data (Data.merge thys (x1, x2)))
+      (Data o Data.merge o map (fn (thy, Data x) => (thy, x)))
   end;
 
 val get = Context.Theory_Data.get kind (fn Data x => x);
@@ -676,7 +695,7 @@
   (
     type T = Data.T;
     val empty = Data.empty;
-    fun merge _ = Data.merge;
+    fun merge args = Library.foldl (fn (a, (_, b)) => Data.merge (a, b)) (#2 (hd args), tl args)
   );
 
 
--- a/src/Pure/cterm_items.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/Pure/cterm_items.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -19,6 +19,8 @@
 
 end;
 
+structure Ctermset = Set(Ctermtab.Key);
+
 
 structure Cterms:
 sig
@@ -54,3 +56,5 @@
 fun thm_cache f = Cache.create empty lookup update f;
 
 end;
+
+structure Thmset = Set(Thmtab.Key);
--- a/src/Pure/drule.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/Pure/drule.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -599,7 +599,7 @@
     val cT = Thm.ctyp_of_cterm ct;
     val T = Thm.typ_of cT;
   in
-    Thm.instantiate (TVars.make [((("'a", 0), []), cT)], Vars.make [((("x", 0), T), ct)]) termI
+    Thm.instantiate (TVars.make1 ((("'a", 0), []), cT), Vars.make1 ((("x", 0), T), ct)) termI
   end;
 
 fun dest_term th =
--- a/src/Pure/envir.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/Pure/envir.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -17,7 +17,7 @@
   val empty: int -> env
   val init: env
   val merge: env * env -> env
-  val insert_sorts: env -> Sortset.T -> Sortset.T
+  val insert_sorts: env -> sort list -> sort list
   val genvars: string -> env * typ list -> env * term list
   val genvar: string -> env * typ -> env * term
   val lookup1: tenv -> indexname * typ -> term option
@@ -94,7 +94,7 @@
 
 
 (*NB: type unification may invent new sorts*)  (* FIXME tenv!? *)
-val insert_sorts = Vartab.fold (Sortset.insert_typ o #2 o #2) o type_env;
+val insert_sorts = Vartab.fold (Sorts.insert_typ o #2 o #2) o type_env;
 
 (*Generate a list of distinct variables.
   Increments index to make them distinct from ALL present variables. *)
--- a/src/Pure/ex/Guess.thy	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/Pure/ex/Guess.thy	Mon Apr 24 15:00:54 2023 +0100
@@ -174,7 +174,7 @@
       [] [] [(Binding.empty_atts, [(Logic.mk_term goal, []), (goal, [])])]
     |> snd
     |> Proof.refine_singleton (Method.Basic (fn _ => fn _ => CONTEXT_TACTIC
-        (PRIMITIVE (Thm.instantiate (TVars.empty, Vars.make [(guess, Thm.cterm_of ctxt thesis)]))
+        (PRIMITIVE (Thm.instantiate (TVars.empty, Vars.make1 (guess, Thm.cterm_of ctxt thesis)))
           THEN Goal.conjunction_tac 1
           THEN resolve_tac ctxt [Drule.termI] 1)))
   end;
--- a/src/Pure/global_theory.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/Pure/global_theory.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -172,9 +172,9 @@
   let
     val theories =
       Symtab.build (Theory.nodes_of thy |> fold (fn thy' =>
-        Symtab.update (Context.theory_name thy', thy')));
+        Symtab.update (Context.theory_long_name thy', thy')));
     fun transfer th =
-      Thm.transfer (the_default thy (Symtab.lookup theories (Thm.theory_name th))) th;
+      Thm.transfer (the_default thy (Symtab.lookup theories (Thm.theory_long_name th))) th;
   in transfer end;
 
 fun all_thms_of thy verbose =
--- a/src/Pure/goal.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/Pure/goal.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -57,7 +57,7 @@
   -------- (init)
   C \<Longrightarrow> #C
 *)
-fun init C = Thm.instantiate (TVars.empty, Vars.make [((("A", 0), propT), C)]) Drule.protectI;
+fun init C = Thm.instantiate (TVars.empty, Vars.make1 ((("A", 0), propT), C)) Drule.protectI;
 
 (*
   A1 \<Longrightarrow> ... \<Longrightarrow> An \<Longrightarrow> C
--- a/src/Pure/library.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/Pure/library.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -132,6 +132,7 @@
   (*strings*)
   val nth_string: string -> int -> string
   val fold_string: (string -> 'a -> 'a) -> string -> 'a -> 'a
+  val fold_rev_string: (string -> 'a -> 'a) -> string -> 'a -> 'a
   val exists_string: (string -> bool) -> string -> bool
   val forall_string: (string -> bool) -> string -> bool
   val member_string: string -> string -> bool
@@ -209,6 +210,8 @@
   val dict_ord: ('a * 'b -> order) -> 'a list * 'b list -> order
   val length_ord: 'a list * 'b list -> order
   val list_ord: ('a * 'b -> order) -> 'a list * 'b list -> order
+  val vector_ord: 'a ord -> 'a vector ord
+  val array_ord: 'a ord -> 'a array ord
   val sort: 'a ord -> 'a list -> 'a list
   val sort_distinct: 'a ord -> 'a list -> 'a list
   val sort_strings: string list -> string list
@@ -658,7 +661,7 @@
 fun string_of_int i =
   if i < 0 then Int.toString i
   else if i < 10 then chr (zero + i)
-  else if i < small_int then Vector.sub (small_int_table, i)
+  else if i < small_int then Vector.nth small_int_table i
   else Int.toString i;
 
 end;
@@ -693,30 +696,20 @@
 
 (* functions tuned for strings, avoiding explode *)
 
-fun nth_string str i =
-  (case try String.substring (str, i, 1) of
-    SOME s => s
-  | NONE => raise Subscript);
+fun nth_string str = String.str o String.nth str;
+
+fun fold_string f = String.fold (f o String.str);
 
-fun fold_string f str x0 =
-  let
-    val n = size str;
-    fun iter (x, i) =
-      if i < n then iter (f (String.substring (str, i, 1)) x, i + 1) else x;
-  in iter (x0, 0) end;
+fun fold_rev_string f = String.fold_rev (f o String.str);
 
-fun exists_string pred str =
-  let
-    val n = size str;
-    fun ex i = i < n andalso (pred (String.substring (str, i, 1)) orelse ex (i + 1));
-  in ex 0 end;
+fun exists_string pred = String.exists (pred o String.str);
 
-fun forall_string pred = not o exists_string (not o pred);
+fun forall_string pred = String.forall (pred o String.str);
 
 fun member_string str s = exists_string (fn s' => s = s') str;
 
 fun last_string "" = NONE
-  | last_string s = SOME (str (String.sub (s, size s - 1)));
+  | last_string s = SOME (str (String.nth s (size s - 1)));
 
 fun first_field sep str =
   let
@@ -780,7 +773,7 @@
 val trim_split_lines = trim_line #> split_lines #> map trim_line;
 
 fun normalize_lines str =
-  if exists_string (fn s => s = "\r") str then
+  if member_string str "\r" then
     split_lines str |> map trim_line |> cat_lines
   else str;
 
@@ -829,9 +822,9 @@
 
 fun member eq list x =
   let
-    fun memb [] = false
-      | memb (y :: ys) = eq (x, y) orelse memb ys;
-  in memb list end;
+    fun mem [] = false
+      | mem (y :: ys) = eq (x, y) orelse mem ys;
+  in mem list end;
 
 fun insert eq x xs = if member eq xs x then xs else x :: xs;
 fun remove eq x xs = if member eq xs x then filter_out (fn y => eq (x, y)) xs else xs;
@@ -966,6 +959,12 @@
 fun length_ord (xs, ys) = int_ord (length xs, length ys);
 fun list_ord elem_ord = length_ord ||| dict_ord elem_ord;
 
+fun vector_ord ord =
+  pointer_eq_ord (int_ord o apply2 Vector.length ||| Vector.collate ord);
+
+fun array_ord ord =
+  pointer_eq_ord (int_ord o apply2 Array.length ||| Array.collate ord);
+
 
 (* sorting *)
 
--- a/src/Pure/logic.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/Pure/logic.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -66,7 +66,7 @@
     map_atyps: typ -> typ,
     constraints: ((typ * class) * term) list,
     outer_constraints: (typ * class) list};
-  val unconstrainT: Sortset.T -> term -> unconstrain_context * term
+  val unconstrainT: sort list -> term -> unconstrain_context * term
   val protectC: term
   val protect: term -> term
   val unprotect: term -> term
@@ -365,7 +365,7 @@
 fun unconstrainT shyps prop =
   let
     val present = rev ((fold_types o fold_atyps_sorts) (insert (eq_fst op =)) prop []);
-    val extra = Sortset.dest (fold (Sortset.remove o #2) present shyps);
+    val extra = fold (Sorts.remove_sort o #2) present shyps;
 
     val n = length present;
     val (names1, names2) = Name.invent Name.context Name.aT (n + length extra) |> chop n;
--- a/src/Pure/more_thm.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/Pure/more_thm.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -57,7 +57,7 @@
   val undeclared_hyps: Context.generic -> thm -> term list
   val check_hyps: Context.generic -> thm -> thm
   val declare_term_sorts: term -> Proof.context -> Proof.context
-  val extra_shyps': Proof.context -> thm -> Sortset.T
+  val extra_shyps': Proof.context -> thm -> sort list
   val check_shyps: Proof.context -> thm -> thm
   val weaken_sorts': Proof.context -> cterm -> cterm
   val elim_implies: thm -> thm -> thm
@@ -253,8 +253,8 @@
 
 structure Hyps = Proof_Data
 (
-  type T = {checked_hyps: bool, hyps: Termset.T, shyps: Sortset.T};
-  fun init _ : T = {checked_hyps = true, hyps = Termset.empty, shyps = Sortset.empty};
+  type T = {checked_hyps: bool, hyps: Termset.T, shyps: sort Ord_List.T};
+  fun init _ : T = {checked_hyps = true, hyps = Termset.empty, shyps = []};
 );
 
 fun map_hyps f = Hyps.map (fn {checked_hyps, hyps, shyps} =>
@@ -278,13 +278,14 @@
   map_hyps (fn (_, hyps, shyps) => (#checked_hyps (Hyps.get ctxt), hyps, shyps));
 
 fun undeclared_hyps context th =
-  Termset.fold_rev
+  Thm.hyps_of th
+  |> filter_out
     (case context of
-      Context.Theory _ => cons
+      Context.Theory _ => K false
     | Context.Proof ctxt =>
-        let val {checked_hyps, hyps, ...} = Hyps.get ctxt
-        in fn hyp => if not checked_hyps orelse Termset.member hyps hyp then I else cons hyp end)
-    (Thm.hyps_of th) [];
+        (case Hyps.get ctxt of
+          {checked_hyps = false, ...} => K true
+        | {hyps, ...} => Termset.member hyps));
 
 fun check_hyps context th =
   (case undeclared_hyps context th of
@@ -298,19 +299,19 @@
 
 fun declare_term_sorts t =
   map_hyps (fn (checked_hyps, hyps, shyps) =>
-    (checked_hyps, hyps, Sortset.insert_term t shyps));
+    (checked_hyps, hyps, Sorts.insert_term t shyps));
 
 fun extra_shyps' ctxt th =
-  Sortset.subtract (#shyps (Hyps.get ctxt)) (Thm.extra_shyps th);
+  Sorts.subtract (#shyps (Hyps.get ctxt)) (Thm.extra_shyps th);
 
 fun check_shyps ctxt raw_th =
   let
     val th = Thm.strip_shyps raw_th;
     val extra_shyps = extra_shyps' ctxt th;
   in
-    if Sortset.is_empty extra_shyps then th
+    if null extra_shyps then th
     else error (Pretty.string_of (Pretty.block (Pretty.str "Pending sort hypotheses:" ::
-      Pretty.brk 1 :: Pretty.commas (map (Syntax.pretty_sort ctxt) (Sortset.dest extra_shyps)))))
+      Pretty.brk 1 :: Pretty.commas (map (Syntax.pretty_sort ctxt) extra_shyps))))
   end;
 
 val weaken_sorts' = Thm.weaken_sorts o #shyps o Hyps.get;
@@ -446,7 +447,7 @@
     val fixed =
       Frees.build
        (fold Frees.add_frees (Thm.terms_of_tpairs (Thm.tpairs_of th)) #>
-        Termset.fold Frees.add_frees (Thm.hyps_of th));
+        fold Frees.add_frees (Thm.hyps_of th));
     val is_fixed = Frees.defined fixed o Term.dest_Free o Thm.term_of;
     val frees =
       Cterms.build (th |> Thm.fold_atomic_cterms {hyps = false} Term.is_Free
@@ -459,9 +460,7 @@
 fun unvarify_global thy th =
   let
     val prop = Thm.full_prop_of th;
-    val _ =
-      (Logic.unvarify_global prop;
-       Termset.fold (fn t => fn () => ignore (Logic.unvarify_global t)) (Thm.hyps_of th) ())
+    val _ = map Logic.unvarify_global (prop :: Thm.hyps_of th)
       handle TERM (msg, _) => raise THM (msg, 0, [th]);
 
     val cert = Thm.global_cterm_of thy;
@@ -718,10 +717,8 @@
       |> perhaps (try (Thm.transfer' ctxt))
       |> perhaps (try Thm.strip_shyps);
 
-    val hyps =
-      if show_hyps then Termset.dest (Thm.hyps_of th)
-      else undeclared_hyps (Context.Proof ctxt) th;
-    val extra_shyps = Sortset.dest (extra_shyps' ctxt th);
+    val hyps = if show_hyps then Thm.hyps_of th else undeclared_hyps (Context.Proof ctxt) th;
+    val extra_shyps = extra_shyps' ctxt th;
     val tags = Thm.get_tags th;
     val tpairs = Thm.tpairs_of th;
 
--- a/src/Pure/morphism.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/Pure/morphism.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -22,6 +22,7 @@
     typ: (typ -> typ) list,
     term: (term -> term) list,
     fact: (thm list -> thm list) list} -> morphism
+  val is_identity: morphism -> bool
   val pretty: morphism -> Pretty.T
   val binding: morphism -> binding -> binding
   val binding_prefix: morphism -> (string * bool) list
@@ -82,6 +83,10 @@
     term = map (pair a) term,
     fact = map (pair a) fact};
 
+(*syntactic test only!*)
+fun is_identity (Morphism {names, binding, typ, term, fact}) =
+  null names andalso null binding andalso null typ andalso null term andalso null fact;
+
 fun pretty (Morphism {names, ...}) = Pretty.enum ";" "{" "}" (map Pretty.str (rev names));
 
 val _ = ML_system_pp (fn _ => fn _ => Pretty.to_polyml o pretty);
--- a/src/Pure/name.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/Pure/name.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -59,8 +59,8 @@
   in ":" ^ leading ^ string_of_int i end);
 
 fun bound n =
-  if n < 1000 then Vector.sub (small_int, n)
-  else ":" ^ bound (n div 1000) ^ Vector.sub (small_int, n mod 1000);
+  if n < 1000 then Vector.nth small_int n
+  else ":" ^ bound (n div 1000) ^ Vector.nth small_int (n mod 1000);
 
 val is_bound = String.isPrefix ":";
 
--- a/src/Pure/par_tactical.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/Pure/par_tactical.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -59,10 +59,7 @@
     SOME goals =>
       if Future.relevant goals then
         let
-          fun try_goal g =
-            (case SINGLE tac (Goal.init g) of
-              NONE => raise FAILED ()
-            | SOME st' => st');
+          fun try_goal g = SINGLE tac (Goal.init g) |> \<^if_none>\<open>raise FAILED ()\<close>;
           val results = Par_List.map try_goal goals;
         in retrofit_tac {close = false} (rev results) st end handle FAILED () => Seq.empty
       else DETERM tac st
--- a/src/Pure/proofterm.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/Pure/proofterm.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -6,16 +6,6 @@
 
 infix 8 % %% %>;
 
-structure Oracles = Set(
-  type key = (string * Position.T) * term option
-  val ord = prod_ord (prod_ord fast_string_ord Position.ord) (option_ord Term_Ord.fast_term_ord)
-);
-
-structure PThms = Table(
-  type key = serial
-  val ord = rev_order o int_ord
-);
-
 signature PROOFTERM =
 sig
   type thm_header =
@@ -36,12 +26,11 @@
    | Oracle of string * term * typ list option
    | PThm of thm_header * thm_body
   and proof_body = PBody of
-    {oracles: Oracles.T,
-     thms: thm_node PThms.table,
+    {oracles: ((string * Position.T) * term option) Ord_List.T,
+     thms: (serial * thm_node) Ord_List.T,
      proof: proof}
+  type oracle = (string * Position.T) * term option
   type thm = serial * thm_node
-  type thms = thm_node PThms.table
-  val merge_thms: thms * thms -> thms
   exception MIN_PROOF of unit
   val proof_of: proof_body -> proof
   val join_proof: proof_body future -> proof
@@ -55,13 +44,19 @@
   val thm_node_name: thm_node -> string
   val thm_node_prop: thm_node -> term
   val thm_node_body: thm_node -> proof_body future
-  val thm_node_thms: thm_node -> thms
-  val join_thms: thms -> proof_body list
+  val thm_node_thms: thm_node -> thm list
+  val join_thms: thm list -> proof_body list
   val make_thm: thm_header -> thm_body -> thm
   val fold_proof_atoms: bool -> (proof -> 'a -> 'a) -> proof list -> 'a -> 'a
   val fold_body_thms:
     ({serial: serial, name: string, prop: term, body: proof_body} -> 'a -> 'a) ->
     proof_body list -> 'a -> 'a
+  val oracle_ord: oracle ord
+  val thm_ord: thm ord
+  val union_oracles: oracle Ord_List.T -> oracle Ord_List.T -> oracle Ord_List.T
+  val unions_oracles: oracle Ord_List.T list -> oracle Ord_List.T
+  val union_thms: thm Ord_List.T -> thm Ord_List.T -> thm Ord_List.T
+  val unions_thms: thm Ord_List.T list -> thm Ord_List.T
   val no_proof_body: proof -> proof_body
   val no_thm_names: proof -> proof
   val no_thm_proofs: proof -> proof
@@ -143,7 +138,7 @@
   val abstract_rule_proof: string * term -> proof -> proof
   val combination_proof: term -> term -> term -> term -> proof -> proof -> proof
   val strip_shyps_proof: Sorts.algebra -> (typ * sort) list -> (typ * sort) list ->
-    Sortset.T -> proof -> proof
+    sort list -> proof -> proof
   val of_sort_proof: Sorts.algebra ->
     (class * class -> proof) ->
     (string * class list list * class -> proof) ->
@@ -184,19 +179,19 @@
   val export_proof_boxes: proof_body list -> unit
   val fulfill_norm_proof: theory -> (serial * proof_body) list -> proof_body -> proof_body
   val thm_proof: theory -> (class * class -> proof) ->
-    (string * class list list * class -> proof) -> string * Position.T -> Sortset.T ->
+    (string * class list list * class -> proof) -> string * Position.T -> sort list ->
     term list -> term -> (serial * proof_body future) list -> proof_body -> thm * proof
   val unconstrain_thm_proof: theory -> (class * class -> proof) ->
-    (string * class list list * class -> proof) -> Sortset.T -> term ->
+    (string * class list list * class -> proof) -> sort list -> term ->
     (serial * proof_body future) list -> proof_body -> thm * proof
-  val get_identity: Sortset.T -> term list -> term -> proof ->
+  val get_identity: sort list -> term list -> term -> proof ->
     {serial: serial, theory_name: string, name: string} option
-  val get_approximative_name: Sortset.T -> term list -> term -> proof -> string
+  val get_approximative_name: sort list -> term list -> term -> proof -> string
   type thm_id = {serial: serial, theory_name: string}
   val make_thm_id: serial * string -> thm_id
   val thm_header_id: thm_header -> thm_id
   val thm_id: thm -> thm_id
-  val get_id: Sortset.T -> term list -> term -> proof -> thm_id option
+  val get_id: sort list -> term list -> term -> proof -> thm_id option
   val this_id: thm_id option -> thm_id -> bool
   val proof_boxes: {excluded: thm_id -> bool, included: thm_id -> bool} ->
     proof list -> (thm_header * proof) list  (*exception MIN_PROOF*)
@@ -224,8 +219,8 @@
  | Oracle of string * term * typ list option
  | PThm of thm_header * thm_body
 and proof_body = PBody of
-  {oracles: Oracles.T,
-   thms: thm_node PThms.table,
+  {oracles: ((string * Position.T) * term option) Ord_List.T,
+   thms: (serial * thm_node) Ord_List.T,
    proof: proof}
 and thm_body =
   Thm_Body of {open_proof: proof -> proof, body: proof_body future}
@@ -233,10 +228,13 @@
   Thm_Node of {theory_name: string, name: string, prop: term,
     body: proof_body future, export: unit lazy, consolidate: unit lazy};
 
-type thm = serial * thm_node;
+type oracle = (string * Position.T) * term option;
+val oracle_ord: oracle ord =
+  prod_ord (prod_ord fast_string_ord Position.ord) (option_ord Term_Ord.fast_term_ord);
 
-type thms = thm_node PThms.table;
-val merge_thms: thms * thms -> thms = PThms.merge (K true);
+type thm = serial * thm_node;
+val thm_ord: thm ord = fn ((i, _), (j, _)) => int_ord (j, i);
+
 
 exception MIN_PROOF of unit;
 
@@ -262,11 +260,11 @@
 val thm_node_export = #export o rep_thm_node;
 val thm_node_consolidate = #consolidate o rep_thm_node;
 
-fun join_thms (thms: thms) =
-  Future.joins (PThms.fold_rev (cons o thm_node_body o #2) thms []);
+fun join_thms (thms: thm list) =
+  Future.joins (map (thm_node_body o #2) thms);
 
 val consolidate_bodies =
-  maps (fn PBody {thms, ...} => PThms.fold_rev (cons o thm_node_consolidate o #2) thms [])
+  maps (fn PBody {thms, ...} => map (thm_node_consolidate o #2) thms)
   #> Lazy.consolidate #> map Lazy.force #> ignore;
 
 fun make_thm_node theory_name name prop body export =
@@ -306,7 +304,7 @@
 fun fold_body_thms f =
   let
     fun app (PBody {thms, ...}) =
-      tap join_thms thms |> PThms.fold (fn (i, thm_node) => fn (x, seen) =>
+      tap join_thms thms |> fold (fn (i, thm_node) => fn (x, seen) =>
         if Intset.member seen i then (x, seen)
         else
           let
@@ -320,7 +318,13 @@
 
 (* proof body *)
 
-fun no_proof_body proof = PBody {oracles = Oracles.empty, thms = PThms.empty, proof = proof};
+val union_oracles = Ord_List.union oracle_ord;
+val unions_oracles = Ord_List.unions oracle_ord;
+
+val union_thms = Ord_List.union thm_ord;
+val unions_thms = Ord_List.unions thm_ord;
+
+fun no_proof_body proof = PBody {oracles = [], thms = [], proof = proof};
 val no_thm_body = thm_body (no_proof_body MinProof);
 
 fun no_thm_names (Abst (x, T, prf)) = Abst (x, T, no_thm_names prf)
@@ -376,8 +380,7 @@
         (map Position.properties_of pos, (prop, (types, map_proof_of open_proof (Future.join body)))))]
 and proof_body consts (PBody {oracles, thms, proof = prf}) =
   triple (list (pair (pair string (properties o Position.properties_of))
-      (option (term consts)))) (list (thm consts)) (proof consts)
-      (Oracles.dest oracles, PThms.dest thms, prf)
+      (option (term consts)))) (list (thm consts)) (proof consts) (oracles, thms, prf)
 and thm consts (a, thm_node) =
   pair int (pair string (pair string (pair (term consts) (proof_body consts))))
     (a, (thm_node_theory_name thm_node, (thm_node_name thm_node, (thm_node_prop thm_node,
@@ -443,7 +446,7 @@
     val (a, b, c) =
       triple (list (pair (pair string (Position.of_properties o properties))
         (option (term consts)))) (list (thm consts)) (proof consts) x;
-  in PBody {oracles = Oracles.make a, thms = PThms.make b, proof = c} end
+  in PBody {oracles = a, thms = b, proof = c} end
 and thm consts x =
   let
     val (a, (b, (c, (d, e)))) =
@@ -1088,7 +1091,7 @@
 
 fun strip_shyps_proof algebra present witnessed extra prf =
   let
-    val replacements = present @ witnessed @ map (`Logic.dummy_tfree) (Sortset.dest extra);
+    val replacements = present @ witnessed @ map (`Logic.dummy_tfree) extra;
     fun get_replacement S =
       replacements |> get_first (fn (T', S') =>
         if Sorts.sort_le algebra (S', S) then SOME T' else NONE);
@@ -1184,7 +1187,7 @@
 fun const_proof mk name prop =
   let
     val args = prop_args prop;
-    val ({outer_constraints, ...}, prop1) = Logic.unconstrainT Sortset.empty prop;
+    val ({outer_constraints, ...}, prop1) = Logic.unconstrainT [] prop;
     val head = mk (name, prop1, NONE);
   in proof_combP (proof_combt' (head, args), map PClass outer_constraints) end;
 
@@ -1992,8 +1995,10 @@
     val _ = consolidate_bodies (map #2 ps @ [body0]);
     val PBody {oracles = oracles0, thms = thms0, proof = proof0} = body0;
     val oracles =
-      fold (fn (_, PBody {oracles, ...}) => fn acc => Oracles.merge (acc, oracles)) ps oracles0;
-    val thms = fold (fn (_, PBody {thms, ...}) => fn acc => merge_thms (acc, thms)) ps thms0;
+      unions_oracles
+        (fold (fn (_, PBody {oracles, ...}) => not (null oracles) ? cons oracles) ps [oracles0]);
+    val thms =
+      unions_thms (fold (fn (_, PBody {thms, ...}) => not (null thms) ? cons thms) ps [thms0]);
     val proof = rew_proof thy proof0;
   in PBody {oracles = oracles, thms = thms, proof = proof} end;
 
@@ -2132,7 +2137,7 @@
 fun export_standard_enabled () = Options.default_bool "export_standard_proofs";
 
 fun export_proof_boxes_required thy =
-  Context.theory_name thy = Context.PureN orelse
+  Context.theory_long_name thy = Context.PureN orelse
     (export_enabled () andalso not (export_standard_enabled ()));
 
 fun export_proof_boxes bodies =
@@ -2142,9 +2147,9 @@
       else
         boxes
         |> Inttab.update (i, thm_node_export thm_node)
-        |> PThms.fold export_thm (thm_node_thms thm_node);
+        |> fold export_thm (thm_node_thms thm_node);
 
-    fun export_body (PBody {thms, ...}) = PThms.fold export_thm thms;
+    fun export_body (PBody {thms, ...}) = fold export_thm thms;
 
     val exports = Inttab.build (fold export_body bodies) |> Inttab.dest;
   in List.app (Lazy.force o #2) exports end;
--- a/src/Pure/raw_simplifier.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/Pure/raw_simplifier.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -1216,7 +1216,7 @@
         val (lhs, rhs) = Thm.dest_equals (Thm.cprop_of eq);
         val eq' =
           Thm.implies_elim
-            (Thm.instantiate (TVars.empty, Vars.make [(vA, prem), (vB, lhs), (vC, rhs)])
+            (Thm.instantiate (TVars.empty, Vars.make3 (vA, prem) (vB, lhs) (vC, rhs))
               Drule.imp_cong)
             (Thm.implies_intr prem eq);
       in
@@ -1228,10 +1228,10 @@
           in
             Thm.transitive
               (Thm.transitive
-                (Thm.instantiate (TVars.empty, Vars.make [(vA, prem'), (vB, prem), (vC, concl)])
+                (Thm.instantiate (TVars.empty, Vars.make3 (vA, prem') (vB, prem) (vC, concl))
                   Drule.swap_prems_eq)
                 eq')
-              (Thm.instantiate (TVars.empty, Vars.make [(vA, prem), (vB, prem''), (vC, concl)])
+              (Thm.instantiate (TVars.empty, Vars.make3 (vA, prem) (vB, prem'') (vC, concl))
                 Drule.swap_prems_eq)
           end
       end
@@ -1280,8 +1280,8 @@
             let
               val prem' = Thm.rhs_of eqn;
               val tprems = map Thm.term_of prems;
-              val i = 1 + Termset.fold (fn p =>
-                Integer.max (find_index (fn q => q aconv p) tprems)) (Thm.hyps_of eqn) ~1;
+              val i = 1 + fold Integer.max (map (fn p =>
+                find_index (fn q => q aconv p) tprems) (Thm.hyps_of eqn)) ~1;
               val ((rrs', asm'), ctxt') = rules_of_prem prem' ctxt;
             in
               mut_impc prems concl rrss asms (prem' :: prems')
--- a/src/Pure/sign.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/Pure/sign.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -129,24 +129,23 @@
   tsig: Type.tsig,              (*order-sorted signature of types*)
   consts: Consts.T};            (*polymorphic constants*)
 
+fun rep_sign (Sign args) = args;
 fun make_sign (syn, tsig, consts) = Sign {syn = syn, tsig = tsig, consts = consts};
 
 structure Data = Theory_Data'
 (
   type T = sign;
   val empty = make_sign (Syntax.empty_syntax, Type.empty_tsig, Consts.empty);
-  fun merge old_thys (sign1, sign2) =
+  fun merge args =
     let
-      val Sign {syn = syn1, tsig = tsig1, consts = consts1} = sign1;
-      val Sign {syn = syn2, tsig = tsig2, consts = consts2} = sign2;
-
-      val syn = Syntax.merge_syntax (syn1, syn2);
-      val tsig = Type.merge_tsig (Context.Theory (fst old_thys)) (tsig1, tsig2);
-      val consts = Consts.merge (consts1, consts2);
-    in make_sign (syn, tsig, consts) end;
+      val context0 = Context.Theory (#1 (hd args));
+      val syn' = Library.foldl1 Syntax.merge_syntax (map (#syn o rep_sign o #2) args);
+      val tsig' = Library.foldl1 (Type.merge_tsig context0) (map (#tsig o rep_sign o #2) args);
+      val consts' = Library.foldl1 Consts.merge (map (#consts o rep_sign o #2) args);
+    in make_sign (syn', tsig', consts') end;
 );
 
-fun rep_sg thy = Data.get thy |> (fn Sign args => args);
+val rep_sg = rep_sign o Data.get;
 
 fun map_sign f = Data.map (fn Sign {syn, tsig, consts} => make_sign (f (syn, tsig, consts)));
 
@@ -514,7 +513,7 @@
 val mandatory_path = map_naming o Name_Space.mandatory_path;
 val qualified_path = map_naming oo Name_Space.qualified_path;
 
-fun local_path thy = thy |> root_path |> add_path (Context.theory_name thy);
+fun local_path thy = thy |> root_path |> add_path (Context.theory_base_name thy);
 
 fun init_naming thy =
   let
--- a/src/Pure/sorts.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/Pure/sorts.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -14,6 +14,17 @@
 
 signature SORTS =
 sig
+  val make: sort list -> sort Ord_List.T
+  val subset: sort Ord_List.T * sort Ord_List.T -> bool
+  val union: sort Ord_List.T -> sort Ord_List.T -> sort Ord_List.T
+  val unions: sort Ord_List.T list -> sort Ord_List.T
+  val subtract: sort Ord_List.T -> sort Ord_List.T -> sort Ord_List.T
+  val remove_sort: sort -> sort Ord_List.T -> sort Ord_List.T
+  val insert_sort: sort -> sort Ord_List.T -> sort Ord_List.T
+  val insert_typ: typ -> sort Ord_List.T -> sort Ord_List.T
+  val insert_typs: typ list -> sort Ord_List.T -> sort Ord_List.T
+  val insert_term: term -> sort Ord_List.T -> sort Ord_List.T
+  val insert_terms: term list -> sort Ord_List.T -> sort Ord_List.T
   type algebra
   val classes_of: algebra -> serial Graph.T
   val arities_of: algebra -> (class * sort list) list Symtab.table
@@ -59,6 +70,36 @@
 structure Sorts: SORTS =
 struct
 
+
+(** ordered lists of sorts **)
+
+val make = Ord_List.make Term_Ord.sort_ord;
+val subset = Ord_List.subset Term_Ord.sort_ord;
+val union = Ord_List.union Term_Ord.sort_ord;
+val unions = Ord_List.unions Term_Ord.sort_ord;
+val subtract = Ord_List.subtract Term_Ord.sort_ord;
+
+val remove_sort = Ord_List.remove Term_Ord.sort_ord;
+val insert_sort = Ord_List.insert Term_Ord.sort_ord;
+
+fun insert_typ (TFree (_, S)) Ss = insert_sort S Ss
+  | insert_typ (TVar (_, S)) Ss = insert_sort S Ss
+  | insert_typ (Type (_, Ts)) Ss = insert_typs Ts Ss
+and insert_typs [] Ss = Ss
+  | insert_typs (T :: Ts) Ss = insert_typs Ts (insert_typ T Ss);
+
+fun insert_term (Const (_, T)) Ss = insert_typ T Ss
+  | insert_term (Free (_, T)) Ss = insert_typ T Ss
+  | insert_term (Var (_, T)) Ss = insert_typ T Ss
+  | insert_term (Bound _) Ss = Ss
+  | insert_term (Abs (_, T, t)) Ss = insert_term t (insert_typ T Ss)
+  | insert_term (t $ u) Ss = insert_term t (insert_term u Ss);
+
+fun insert_terms [] Ss = Ss
+  | insert_terms (t :: ts) Ss = insert_terms ts (insert_term t Ss);
+
+
+
 (** order-sorted algebra **)
 
 (*
--- a/src/Pure/term_items.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/Pure/term_items.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -27,9 +27,15 @@
   val defined: 'a table -> key -> bool
   val add: key * 'a -> 'a table -> 'a table
   val make: (key * 'a) list -> 'a table
+  val make1: key * 'a -> 'a table
+  val make2: key * 'a -> key * 'a -> 'a table
+  val make3: key * 'a -> key * 'a -> key * 'a -> 'a table
   type set = int table
   val add_set: key -> set -> set
   val make_set: key list -> set
+  val make1_set: key -> set
+  val make2_set: key -> key -> set
+  val make3_set: key -> key -> key -> set
   val list_set: set -> key list
   val list_set_rev: set -> key list
   val subset: set * set -> bool
@@ -65,7 +71,10 @@
 fun defined (Table tab) = Table.defined tab;
 
 fun add entry (Table tab) = Table (Table.default entry tab);
-fun make entries = build (fold add entries);
+fun make es = build (fold add es);
+fun make1 e = build (add e);
+fun make2 e1 e2 = build (add e1 #> add e2);
+fun make3 e1 e2 e3 = build (add e1 #> add e2 #> add e3);
 
 
 (* set with order of addition *)
@@ -76,9 +85,12 @@
   Table (if Table.defined tab x then tab else Table.update_new (x, Table.size tab + 1) tab);
 
 fun make_set xs = build (fold add_set xs);
+fun make1_set e = build (add_set e);
+fun make2_set e1 e2 = build (add_set e1 #> add_set e2);
+fun make3_set e1 e2 e3 = build (add_set e1 #> add_set e2 #> add_set e3);
 
 fun subset (A: set, B: set) = forall (defined B o #1) A;
-fun eq_set (A: set, B: set) = size A = size B andalso subset (A, B);
+fun eq_set (A: set, B: set) = pointer_eq (A, B) orelse size A = size B andalso subset (A, B);
 
 fun list_set_ord ord (Table tab) = tab |> Table.dest |> sort (ord o apply2 #2) |> map #1;
 val list_set = list_set_ord int_ord;
--- a/src/Pure/term_ord.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/Pure/term_ord.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -62,38 +62,54 @@
   | cons_nr (Abs _) = 4
   | cons_nr (_ $ _) = 5;
 
-fun struct_ord (Abs (_, _, t), Abs (_, _, u)) = struct_ord (t, u)
-  | struct_ord (t1 $ t2, u1 $ u2) =
-      (case struct_ord (t1, u1) of EQUAL => struct_ord (t2, u2) | ord => ord)
-  | struct_ord (t, u) = int_ord (cons_nr t, cons_nr u);
+fun struct_ord tu =
+  if pointer_eq tu then EQUAL
+  else
+    (case tu of
+      (Abs (_, _, t), Abs (_, _, u)) => struct_ord (t, u)
+    | (t1 $ t2, u1 $ u2) =>
+        (case struct_ord (t1, u1) of EQUAL => struct_ord (t2, u2) | ord => ord)
+    | (t, u) => int_ord (cons_nr t, cons_nr u));
 
-fun atoms_ord (Abs (_, _, t), Abs (_, _, u)) = atoms_ord (t, u)
-  | atoms_ord (t1 $ t2, u1 $ u2) =
-      (case atoms_ord (t1, u1) of EQUAL => atoms_ord (t2, u2) | ord => ord)
-  | atoms_ord (Const (a, _), Const (b, _)) = fast_string_ord (a, b)
-  | atoms_ord (Free (x, _), Free (y, _)) = fast_string_ord (x, y)
-  | atoms_ord (Var (xi, _), Var (yj, _)) = fast_indexname_ord (xi, yj)
-  | atoms_ord (Bound i, Bound j) = int_ord (i, j)
-  | atoms_ord _ = EQUAL;
+fun atoms_ord tu =
+  if pointer_eq tu then EQUAL
+  else
+    (case tu of
+      (Abs (_, _, t), Abs (_, _, u)) => atoms_ord (t, u)
+    | (t1 $ t2, u1 $ u2) =>
+        (case atoms_ord (t1, u1) of EQUAL => atoms_ord (t2, u2) | ord => ord)
+    | (Const (a, _), Const (b, _)) => fast_string_ord (a, b)
+    | (Free (x, _), Free (y, _)) => fast_string_ord (x, y)
+    | (Var (xi, _), Var (yj, _)) => fast_indexname_ord (xi, yj)
+    | (Bound i, Bound j) => int_ord (i, j)
+    | _ => EQUAL);
 
-fun types_ord (Abs (_, T, t), Abs (_, U, u)) =
-      (case typ_ord (T, U) of EQUAL => types_ord (t, u) | ord => ord)
-  | types_ord (t1 $ t2, u1 $ u2) =
-      (case types_ord (t1, u1) of EQUAL => types_ord (t2, u2) | ord => ord)
-  | types_ord (Const (_, T), Const (_, U)) = typ_ord (T, U)
-  | types_ord (Free (_, T), Free (_, U)) = typ_ord (T, U)
-  | types_ord (Var (_, T), Var (_, U)) = typ_ord (T, U)
-  | types_ord _ = EQUAL;
+fun types_ord tu =
+  if pointer_eq tu then EQUAL
+  else
+    (case tu of
+      (Abs (_, T, t), Abs (_, U, u)) =>
+        (case typ_ord (T, U) of EQUAL => types_ord (t, u) | ord => ord)
+    | (t1 $ t2, u1 $ u2) =>
+        (case types_ord (t1, u1) of EQUAL => types_ord (t2, u2) | ord => ord)
+    | (Const (_, T), Const (_, U)) => typ_ord (T, U)
+    | (Free (_, T), Free (_, U)) => typ_ord (T, U)
+    | (Var (_, T), Var (_, U)) => typ_ord (T, U)
+    | _ => EQUAL);
 
-fun comments_ord (Abs (x, _, t), Abs (y, _, u)) =
-      (case fast_string_ord (x, y) of EQUAL => comments_ord (t, u) | ord => ord)
-  | comments_ord (t1 $ t2, u1 $ u2) =
-      (case comments_ord (t1, u1) of EQUAL => comments_ord (t2, u2) | ord => ord)
-  | comments_ord _ = EQUAL;
+fun comments_ord tu =
+  if pointer_eq tu then EQUAL
+  else
+    (case tu of
+      (Abs (x, _, t), Abs (y, _, u)) =>
+        (case fast_string_ord (x, y) of EQUAL => comments_ord (t, u) | ord => ord)
+    | (t1 $ t2, u1 $ u2) =>
+        (case comments_ord (t1, u1) of EQUAL => comments_ord (t2, u2) | ord => ord)
+    | _ => EQUAL);
 
 in
 
-val fast_term_ord = pointer_eq_ord (struct_ord ||| atoms_ord ||| types_ord);
+val fast_term_ord = struct_ord ||| atoms_ord ||| types_ord;
 
 val syntax_term_ord = fast_term_ord ||| comments_ord;
 
@@ -224,38 +240,6 @@
 
 structure Termset = Set(Termtab.Key);
 
-structure Sortset:
-sig
-  include SET
-  val insert_typ: typ -> T -> T
-  val insert_typs: typ list -> T -> T
-  val insert_term: term -> T -> T
-  val insert_terms: term list -> T -> T
-end =
-struct
-
-structure Set = Set(Sorttab.Key);
-open Set;
-
-fun insert_typ (TFree (_, S)) Ss = insert S Ss
-  | insert_typ (TVar (_, S)) Ss = insert S Ss
-  | insert_typ (Type (_, Ts)) Ss = insert_typs Ts Ss
-and insert_typs [] Ss = Ss
-  | insert_typs (T :: Ts) Ss = insert_typs Ts (insert_typ T Ss);
-
-fun insert_term (Const (_, T)) Ss = insert_typ T Ss
-  | insert_term (Free (_, T)) Ss = insert_typ T Ss
-  | insert_term (Var (_, T)) Ss = insert_typ T Ss
-  | insert_term (Bound _) Ss = Ss
-  | insert_term (Abs (_, T, t)) Ss = insert_term t (insert_typ T Ss)
-  | insert_term (t $ u) Ss = insert_term t (insert_term u Ss);
-
-fun insert_terms [] Ss = Ss
-  | insert_terms (t :: ts) Ss = insert_terms ts (insert_term t Ss);
-
-end;
-
 structure Var_Graph = Graph(Vartab.Key);
-structure Sort_Graph = Graph(Sorttab.Key);
 structure Typ_Graph = Graph(Typtab.Key);
 structure Term_Graph = Graph(Termtab.Key);
--- a/src/Pure/theory.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/Pure/theory.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -22,7 +22,6 @@
   val defs_of: theory -> Defs.T
   val at_begin: (theory -> theory option) -> theory -> theory
   val at_end: (theory -> theory option) -> theory -> theory
-  val join_theory: theory list -> theory
   val begin_theory: string * Position.T -> theory list -> theory
   val end_theory: theory -> theory
   val add_axiom: Proof.context -> binding * term -> theory -> theory
@@ -83,6 +82,8 @@
   defs: Defs.T,
   wrappers: wrapper list * wrapper list};
 
+fun rep_thy (Thy args) = args;
+
 fun make_thy (pos, id, axioms, defs, wrappers) =
   Thy {pos = pos, id = id, axioms = axioms, defs = defs, wrappers = wrappers};
 
@@ -90,19 +91,22 @@
 (
   type T = thy;
   val empty = make_thy (Position.none, 0, Name_Space.empty_table Markup.axiomN, Defs.empty, ([], []));
-  fun merge old_thys (thy1, thy2) =
+  fun merge args =
     let
-      val Thy {pos, id, axioms = axioms1, defs = defs1, wrappers = (bgs1, ens1)} = thy1;
-      val Thy {pos = _, id = _, axioms = axioms2, defs = defs2, wrappers = (bgs2, ens2)} = thy2;
+      val thy0 = #1 (hd args);
+      val {pos, id, ...} = rep_thy (#2 (hd args));
 
-      val axioms' = Name_Space.merge_tables (axioms1, axioms2);
-      val defs' = Defs.merge (Defs.global_context (fst old_thys)) (defs1, defs2);
-      val bgs' = Library.merge (eq_snd op =) (bgs1, bgs2);
-      val ens' = Library.merge (eq_snd op =) (ens1, ens2);
+      val merge_defs = Defs.merge (Defs.global_context thy0);
+      val merge_wrappers = Library.merge (eq_snd op =);
+
+      val axioms' = Library.foldl1 Name_Space.merge_tables (map (#axioms o rep_thy o #2) args);
+      val defs' = Library.foldl1 merge_defs (map (#defs o rep_thy o #2) args);
+      val bgs' = Library.foldl1 merge_wrappers (map (#1 o #wrappers o rep_thy o #2) args);
+      val ens' = Library.foldl1 merge_wrappers (map (#2 o #wrappers o rep_thy o #2) args);
     in make_thy (pos, id, axioms', defs', (bgs', ens')) end;
 );
 
-fun rep_theory thy = Thy.get thy |> (fn Thy args => args);
+val rep_theory = rep_thy o Thy.get;
 
 fun map_thy f = Thy.map (fn (Thy {pos, id, axioms, defs, wrappers}) =>
   make_thy (f (pos, id, axioms, defs, wrappers)));
@@ -143,7 +147,7 @@
             val completion_report =
               Completion.make_report (name, pos)
                 (fn completed =>
-                  map (Context.theory_name' long) (ancestors_of thy)
+                  map (Context.theory_name long) (ancestors_of thy)
                   |> filter (completed o Long_Name.base_name)
                   |> sort_strings
                   |> map (fn a => (a, (Markup.theoryN, a))));
@@ -162,13 +166,6 @@
 val defs_of = #defs o rep_theory;
 
 
-(* join theory *)
-
-fun join_theory [] = raise List.Empty
-  | join_theory [thy] = thy
-  | join_theory thys = foldl1 Context.join_thys thys;
-
-
 (* begin/end theory *)
 
 val begin_wrappers = rev o #1 o #wrappers o rep_theory;
--- a/src/Pure/thm.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/Pure/thm.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -66,11 +66,13 @@
   val terms_of_tpairs: (term * term) list -> term list
   val full_prop_of: thm -> term
   val theory_id: thm -> Context.theory_id
-  val theory_name: thm -> string
+  val theory_name: {long: bool} -> thm -> string
+  val theory_long_name: thm -> string
+  val theory_base_name: thm -> string
   val maxidx_of: thm -> int
   val maxidx_thm: thm -> int -> int
-  val shyps_of: thm -> Sortset.T
-  val hyps_of: thm -> Termset.T
+  val shyps_of: thm -> sort Ord_List.T
+  val hyps_of: thm -> term list
   val prop_of: thm -> term
   val tpairs_of: thm -> (term * term) list
   val concl_of: thm -> term
@@ -98,7 +100,7 @@
   val join_transfer_context: Proof.context * thm -> Proof.context * thm
   val renamed_prop: term -> thm -> thm
   val weaken: cterm -> thm -> thm
-  val weaken_sorts: Sortset.T -> cterm -> cterm
+  val weaken_sorts: sort list -> cterm -> cterm
   val proof_bodies_of: thm list -> proof_body list
   val proof_body_of: thm -> proof_body
   val proof_of: thm -> proof
@@ -107,8 +109,8 @@
   val expose_proofs: theory -> thm list -> unit
   val expose_proof: theory -> thm -> unit
   val future: thm future -> cterm -> thm
-  val thm_deps: thm -> Proofterm.thms
-  val extra_shyps: thm -> Sortset.T
+  val thm_deps: thm -> Proofterm.thm Ord_List.T
+  val extra_shyps: thm -> sort list
   val strip_shyps: thm -> thm
   val derivation_closed: thm -> bool
   val derivation_name: thm -> string
@@ -177,7 +179,7 @@
   val bicompose: Proof.context option -> {flatten: bool, match: bool, incremented: bool} ->
     bool * thm * int -> int -> thm -> thm Seq.seq
   val biresolution: Proof.context option -> bool -> (bool * thm) list -> int -> thm -> thm Seq.seq
-  val thynames_of_arity: theory -> string * class -> string list
+  val theory_names_of_arity: {long: bool} -> theory -> string * class -> string list
   val add_classrel: thm -> theory -> theory
   val add_arity: thm -> theory -> theory
 end;
@@ -189,7 +191,7 @@
 
 (** certified types **)
 
-datatype ctyp = Ctyp of {cert: Context.certificate, T: typ, maxidx: int, sorts: Sortset.T};
+datatype ctyp = Ctyp of {cert: Context.certificate, T: typ, maxidx: int, sorts: sort Ord_List.T};
 
 fun typ_of (Ctyp {T, ...}) = T;
 
@@ -197,7 +199,7 @@
   let
     val T = Sign.certify_typ thy raw_T;
     val maxidx = Term.maxidx_of_typ T;
-    val sorts = Sortset.build (Sortset.insert_typ T);
+    val sorts = Sorts.insert_typ T [];
   in Ctyp {cert = Context.Certificate thy, T = T, maxidx = maxidx, sorts = sorts} end;
 
 val ctyp_of = global_ctyp_of o Proof_Context.theory_of;
@@ -219,7 +221,7 @@
 val dest_ctyp1 = dest_ctypN 1;
 
 fun join_certificate_ctyp (Ctyp {cert, ...}) cert0 = Context.join_certificate (cert0, cert);
-fun insert_sorts_ctyp (Ctyp {sorts, ...}) sorts0 = Sortset.merge (sorts0, sorts);
+fun union_sorts_ctyp (Ctyp {sorts, ...}) sorts0 = Sorts.union sorts0 sorts;
 fun maxidx_ctyp (Ctyp {maxidx, ...}) maxidx0 = Int.max (maxidx0, maxidx);
 
 fun make_ctyp (Ctyp {cert, T, maxidx = _, sorts = _}) cargs =
@@ -232,7 +234,7 @@
         Ctyp {
           cert = fold join_certificate_ctyp cargs cert,
           maxidx = fold maxidx_ctyp cargs ~1,
-          sorts = Sortset.build (fold insert_sorts_ctyp cargs),
+          sorts = fold union_sorts_ctyp cargs [],
           T = if length args = length cargs then Type (a, As) else err ()}
     | _ => err ())
   end;
@@ -243,7 +245,7 @@
 
 (*certified terms with checked typ, maxidx, and sorts*)
 datatype cterm =
-  Cterm of {cert: Context.certificate, t: term, T: typ, maxidx: int, sorts: Sortset.T};
+  Cterm of {cert: Context.certificate, t: term, T: typ, maxidx: int, sorts: sort Ord_List.T};
 
 exception CTERM of string * cterm list;
 
@@ -259,7 +261,7 @@
 fun global_cterm_of thy tm =
   let
     val (t, T, maxidx) = Sign.certify_term thy tm;
-    val sorts = Sortset.build (Sortset.insert_term t);
+    val sorts = Sorts.insert_term t [];
   in Cterm {cert = Context.Certificate thy, t = t, T = T, maxidx = maxidx, sorts = sorts} end;
 
 val cterm_of = global_cterm_of o Proof_Context.theory_of;
@@ -345,7 +347,7 @@
         t = f $ x,
         T = rty,
         maxidx = Int.max (maxidx1, maxidx2),
-        sorts = Sortset.merge (sorts1, sorts2)}
+        sorts = Sorts.union sorts1 sorts2}
       else raise CTERM ("apply: types don't agree", [cf, cx])
   | apply cf cx = raise CTERM ("apply: first arg is not a function", [cf, cx]);
 
@@ -356,7 +358,7 @@
       Cterm {cert = join_certificate0 (ct1, ct2),
         t = t, T = T1 --> T2,
         maxidx = Int.max (maxidx1, maxidx2),
-        sorts = Sortset.merge (sorts1, sorts2)}
+        sorts = Sorts.union sorts1 sorts2}
     end;
 
 fun lambda t u = lambda_name ("", t) u;
@@ -385,22 +387,21 @@
 
 type constraint = {theory: theory, typ: typ, sort: sort};
 
-structure Constraints =
-  Set(
-    type key = constraint;
-    val ord =
-      Context.theory_id_ord o apply2 (Context.theory_id o #theory)
-      ||| Term_Ord.typ_ord o apply2 #typ
-      ||| Term_Ord.sort_ord o apply2 #sort;
-  );
+local
 
-local
+val constraint_ord : constraint ord =
+  Context.theory_id_ord o apply2 (Context.theory_id o #theory)
+  ||| Term_Ord.typ_ord o apply2 #typ
+  ||| Term_Ord.sort_ord o apply2 #sort;
 
 val smash_atyps =
   map_atyps (fn TVar (_, S) => Term.aT S | TFree (_, S) => Term.aT S | T => T);
 
 in
 
+val union_constraints = Ord_List.union constraint_ord;
+val unions_constraints = Ord_List.unions constraint_ord;
+
 fun insert_constraints thy (T, S) =
   let
     val ignored =
@@ -411,7 +412,7 @@
         | _ => false);
   in
     if ignored then I
-    else Constraints.insert {theory = thy, typ = smash_atyps T, sort = S}
+    else Ord_List.insert constraint_ord {theory = thy, typ = smash_atyps T, sort = S}
   end;
 
 fun insert_constraints_env thy env =
@@ -431,23 +432,15 @@
  {cert: Context.certificate,    (*background theory certificate*)
   tags: Properties.T,           (*additional annotations/comments*)
   maxidx: int,                  (*maximum index of any Var or TVar*)
-  constraints: Constraints.T,   (*implicit proof obligations for sort constraints*)
-  shyps: Sortset.T,             (*sort hypotheses*)
-  hyps: Termset.T,              (*hypotheses*)
+  constraints: constraint Ord_List.T,  (*implicit proof obligations for sort constraints*)
+  shyps: sort Ord_List.T,       (*sort hypotheses*)
+  hyps: term Ord_List.T,        (*hypotheses*)
   tpairs: (term * term) list,   (*flex-flex pairs*)
   prop: term}                   (*conclusion*)
 and deriv = Deriv of
- {promises: thm future Inttab'.table,
+ {promises: (serial * thm future) Ord_List.T,
   body: Proofterm.proof_body};
 
-type promises = thm future Inttab'.table;
-val null_promises : promises -> bool = Inttab'.is_empty;
-val empty_promises : promises = Inttab'.empty;
-val merge_promises : promises * promises -> promises = Inttab'.merge (K true);
-val make_promises : (serial * thm future) list -> promises = Inttab'.make;
-val dest_promises : promises -> (serial * thm future) list = Inttab'.dest;
-fun forall_promises f : promises -> bool = Inttab'.forall (f o fst);
-
 type conv = cterm -> thm;
 
 (*errors involving theorems*)
@@ -456,7 +449,7 @@
 fun rep_thm (Thm (_, args)) = args;
 
 fun fold_terms h f (Thm (_, {tpairs, prop, hyps, ...})) =
-  fold (fn (t, u) => f t #> f u) tpairs #> f prop #> #hyps h ? Termset.fold f hyps;
+  fold (fn (t, u) => f t #> f u) tpairs #> f prop #> #hyps h ? fold f hyps;
 
 fun fold_atomic_ctyps h g f (th as Thm (_, {cert, maxidx, shyps, ...})) =
   let fun ctyp T = Ctyp {cert = cert, T = T, maxidx = maxidx, sorts = shyps}
@@ -487,6 +480,11 @@
 fun full_prop_of (Thm (_, {tpairs, prop, ...})) = attach_tpairs tpairs prop;
 
 
+val union_hyps = Ord_List.union Term_Ord.fast_term_ord;
+val unions_hyps = Ord_List.unions Term_Ord.fast_term_ord;
+val insert_hyps = Ord_List.insert Term_Ord.fast_term_ord;
+val remove_hyps = Ord_List.remove Term_Ord.fast_term_ord;
+
 fun join_certificate1 (Cterm {cert = cert1, ...}, Thm (_, {cert = cert2, ...})) =
   Context.join_certificate (cert1, cert2);
 
@@ -498,7 +496,10 @@
 
 val cert_of = #cert o rep_thm;
 val theory_id = Context.certificate_theory_id o cert_of;
-val theory_name = Context.theory_id_name o theory_id;
+
+fun theory_name long = Context.theory_id_name long o theory_id;
+val theory_long_name = theory_name {long = true};
+val theory_base_name = theory_name {long = false};
 
 val maxidx_of = #maxidx o rep_thm;
 fun maxidx_thm th i = Int.max (maxidx_of th, i);
@@ -533,18 +534,17 @@
     (prems_of th);
 
 fun chyps_of (Thm (_, {cert, shyps, hyps, ...})) =
-  map (fn t => Cterm {cert = cert, maxidx = ~1, T = propT, sorts = shyps, t = t})
-    (Termset.dest hyps);
+  map (fn t => Cterm {cert = cert, maxidx = ~1, T = propT, sorts = shyps, t = t}) hyps;
 
 
 (* thm order: ignores theory context! *)
 
 val thm_ord =
   pointer_eq_ord
-  (Term_Ord.fast_term_ord o apply2 prop_of
-    ||| list_ord (prod_ord Term_Ord.fast_term_ord Term_Ord.fast_term_ord) o apply2 tpairs_of
-    ||| Termset.ord o apply2 hyps_of
-    ||| Sortset.ord o apply2 shyps_of);
+    (Term_Ord.fast_term_ord o apply2 prop_of
+      ||| list_ord (prod_ord Term_Ord.fast_term_ord Term_Ord.fast_term_ord) o apply2 tpairs_of
+      ||| list_ord Term_Ord.fast_term_ord o apply2 hyps_of
+      ||| list_ord Term_Ord.sort_ord o apply2 shyps_of);
 
 
 (* implicit theory context *)
@@ -574,17 +574,17 @@
         t = t, T = T, maxidx = maxidx, sorts = sorts});
 
 fun trim_context_thm th =
-  if Constraints.is_empty (constraints_of th) then
-    (case th of
-      Thm (_, {cert = Context.Certificate_Id _, ...}) => th
-    | Thm (der,
-        {cert = Context.Certificate thy, tags, maxidx, constraints = _, shyps, hyps,
-          tpairs, prop}) =>
-        Thm (der,
-         {cert = Context.Certificate_Id (Context.theory_id thy),
-          tags = tags, maxidx = maxidx, constraints = Constraints.empty,
-          shyps = shyps, hyps = hyps, tpairs = tpairs, prop = prop}))
-  else raise THM ("trim_context: pending sort constraints", 0, [th]);
+  (case th of
+    Thm (_, {constraints = _ :: _, ...}) =>
+      raise THM ("trim_context: pending sort constraints", 0, [th])
+  | Thm (_, {cert = Context.Certificate_Id _, ...}) => th
+  | Thm (der,
+      {cert = Context.Certificate thy, tags, maxidx, constraints = [], shyps, hyps,
+        tpairs, prop}) =>
+      Thm (der,
+       {cert = Context.Certificate_Id (Context.theory_id thy),
+        tags = tags, maxidx = maxidx, constraints = [], shyps = shyps, hyps = hyps,
+        tpairs = tpairs, prop = prop}));
 
 fun transfer_ctyp thy' cT =
   let
@@ -658,7 +658,7 @@
     val thy = Context.certificate_theory cert
       handle ERROR msg => raise CONTEXT (msg, [], [ct1, ct2], [], NONE);
     val (Tinsts, tinsts) = match thy (t1, t2) (Vartab.empty, Vartab.empty);
-    val sorts = Sortset.merge (sorts1, sorts2);
+    val sorts = Sorts.union sorts1 sorts2;
     fun mk_cTinst ((a, i), (S, T)) =
       (((a, i), S), Ctyp {T = T, cert = cert, maxidx = maxidx2, sorts = sorts});
     fun mk_ctinst ((x, i), (U, t)) =
@@ -719,17 +719,18 @@
         tags = tags,
         maxidx = maxidx,
         constraints = constraints,
-        shyps = Sortset.merge (sorts, shyps),
-        hyps = Termset.insert A hyps,
+        shyps = Sorts.union sorts shyps,
+        hyps = insert_hyps A hyps,
         tpairs = tpairs,
         prop = prop})
   end;
 
-fun weaken_sorts more_sorts ct =
+fun weaken_sorts raw_sorts ct =
   let
     val Cterm {cert, t, T, maxidx, sorts} = ct;
     val thy = theory_of_cterm ct;
-    val sorts' = Sortset.fold (Sortset.insert o Sign.certify_sort thy) more_sorts sorts;
+    val more_sorts = Sorts.make (map (Sign.certify_sort thy) raw_sorts);
+    val sorts' = Sorts.union sorts more_sorts;
   in Cterm {cert = cert, t = t, T = T, maxidx = maxidx, sorts = sorts'} end;
 
 
@@ -739,21 +740,23 @@
 fun make_deriv promises oracles thms proof =
   Deriv {promises = promises, body = PBody {oracles = oracles, thms = thms, proof = proof}};
 
-val empty_deriv = make_deriv empty_promises Oracles.empty PThms.empty MinProof;
+val empty_deriv = make_deriv [] [] [] MinProof;
 
 
 (* inference rules *)
 
+val promise_ord: (serial * thm future) ord = fn ((i, _), (j, _)) => int_ord (j, i);
+
 fun bad_proofs i =
   error ("Illegal level of detail for proof objects: " ^ string_of_int i);
 
 fun deriv_rule2 f
-    (Deriv {promises = promises1, body = PBody {oracles = oracles1, thms = thms1, proof = prf1}})
-    (Deriv {promises = promises2, body = PBody {oracles = oracles2, thms = thms2, proof = prf2}}) =
+    (Deriv {promises = ps1, body = PBody {oracles = oracles1, thms = thms1, proof = prf1}})
+    (Deriv {promises = ps2, body = PBody {oracles = oracles2, thms = thms2, proof = prf2}}) =
   let
-    val ps = merge_promises (promises1, promises2);
-    val oracles = Oracles.merge (oracles1, oracles2);
-    val thms = Proofterm.merge_thms (thms1, thms2);
+    val ps = Ord_List.union promise_ord ps1 ps2;
+    val oracles = Proofterm.union_oracles oracles1 oracles2;
+    val thms = Proofterm.union_thms thms1 thms2;
     val prf =
       (case ! Proofterm.proofs of
         2 => f prf1 prf2
@@ -766,7 +769,7 @@
 
 fun deriv_rule0 make_prf =
   if ! Proofterm.proofs <= 1 then empty_deriv
-  else deriv_rule1 I (make_deriv empty_promises Oracles.empty PThms.empty (make_prf ()));
+  else deriv_rule1 I (make_deriv [] [] [] (make_prf ()));
 
 fun deriv_rule_unconditional f (Deriv {promises, body = PBody {oracles, thms, proof}}) =
   make_deriv promises oracles thms (f proof);
@@ -774,18 +777,15 @@
 
 (* fulfilled proofs *)
 
-fun merge_promises_of (Thm (Deriv {promises, ...}, _)) ps = merge_promises (ps, promises);
+fun raw_promises_of (Thm (Deriv {promises, ...}, _)) = promises;
 
 fun join_promises [] = ()
   | join_promises promises = join_promises_of (Future.joins (map snd promises))
-and join_promises_of thms =
-  join_promises (dest_promises (fold merge_promises_of thms empty_promises));
+and join_promises_of thms = join_promises (Ord_List.make promise_ord (maps raw_promises_of thms));
 
 fun fulfill_body (th as Thm (Deriv {promises, body}, _)) =
-  let
-    val pending = dest_promises promises;
-    val fulfilled = map #1 pending ~~ map fulfill_body (Future.joins (map #2 pending));
-  in Proofterm.fulfill_norm_proof (theory_of_thm th) fulfilled body end;
+  let val fulfilled_promises = map #1 promises ~~ map fulfill_body (Future.joins (map #2 promises))
+  in Proofterm.fulfill_norm_proof (theory_of_thm th) fulfilled_promises body end;
 
 fun proof_bodies_of thms = (join_promises_of thms; map fulfill_body thms);
 val proof_body_of = singleton proof_bodies_of;
@@ -813,12 +813,12 @@
 
     val _ = Context.eq_certificate (cert, orig_cert) orelse err "bad theory";
     val _ = prop aconv orig_prop orelse err "bad prop";
-    val _ = Constraints.is_empty constraints orelse err "bad sort constraints";
+    val _ = null constraints orelse err "bad sort constraints";
     val _ = null tpairs orelse err "bad flex-flex constraints";
-    val _ = Termset.is_empty hyps orelse err "bad hyps";
-    val _ = Sortset.subset (shyps, orig_shyps) orelse err "bad shyps";
-    val _ = forall_promises (fn j => i <> j) promises orelse err "bad dependencies";
-    val _ = join_promises (dest_promises promises);
+    val _ = null hyps orelse err "bad hyps";
+    val _ = Sorts.subset (shyps, orig_shyps) orelse err "bad shyps";
+    val _ = forall (fn (j, _) => i <> j) promises orelse err "bad dependencies";
+    val _ = join_promises promises;
   in thm end;
 
 fun future future_thm ct =
@@ -832,13 +832,13 @@
     val i = serial ();
     val future = future_thm |> Future.map (future_result i cert sorts prop);
   in
-    Thm (make_deriv (make_promises [(i, future)]) Oracles.empty PThms.empty MinProof,
+    Thm (make_deriv [(i, future)] [] [] MinProof,
      {cert = cert,
       tags = [],
       maxidx = maxidx,
-      constraints = Constraints.empty,
+      constraints = [],
       shyps = sorts,
-      hyps = Termset.empty,
+      hyps = [],
       tpairs = [],
       prop = prop})
   end;
@@ -854,12 +854,11 @@
         val der = deriv_rule0 (fn () => Proofterm.axm_proof name prop);
         val cert = Context.Certificate thy;
         val maxidx = maxidx_of_term prop;
-        val shyps = Sortset.build (Sortset.insert_term prop);
+        val shyps = Sorts.insert_term prop [];
       in
         Thm (der,
           {cert = cert, tags = [], maxidx = maxidx,
-            constraints = Constraints.empty, shyps = shyps,
-            hyps = Termset.empty, tpairs = [], prop = prop})
+            constraints = [], shyps = shyps, hyps = [], tpairs = [], prop = prop})
       end
   | NONE => raise THEORY ("No axiom " ^ quote name, [thy]));
 
@@ -974,7 +973,7 @@
 local
 
 fun union_digest (oracles1, thms1) (oracles2, thms2) =
-  (Oracles.merge (oracles1, oracles2), Proofterm.merge_thms (thms1, thms2));
+  (Proofterm.union_oracles oracles1 oracles2, Proofterm.union_thms thms1 thms2);
 
 fun thm_digest (Thm (Deriv {body = PBody {oracles, thms, ...}, ...}, _)) =
   (oracles, thms);
@@ -982,53 +981,51 @@
 fun constraint_digest ({theory = thy, typ, sort, ...}: constraint) =
   Sorts.of_sort_derivation (Sign.classes_of thy)
    {class_relation = fn _ => fn _ => fn (digest, c1) => fn c2 =>
-      if c1 = c2 then (Oracles.empty, PThms.empty)
-      else union_digest digest (thm_digest (the_classrel thy (c1, c2))),
+      if c1 = c2 then ([], []) else union_digest digest (thm_digest (the_classrel thy (c1, c2))),
     type_constructor = fn (a, _) => fn dom => fn c =>
       let val arity_digest = thm_digest (the_arity thy (a, (map o map) #2 dom, c))
       in (fold o fold) (union_digest o #1) dom arity_digest end,
-    type_variable = fn T => map (pair (Oracles.empty, PThms.empty)) (Type.sort_of_atyp T)}
+    type_variable = fn T => map (pair ([], [])) (Type.sort_of_atyp T)}
    (typ, sort);
 
 in
 
-fun solve_constraints (thm as Thm (der, args)) =
-  if Constraints.is_empty (constraints_of thm) then thm
-  else
-    let
-      val {cert, tags, maxidx, constraints, shyps, hyps, tpairs, prop} = args;
+fun solve_constraints (thm as Thm (_, {constraints = [], ...})) = thm
+  | solve_constraints (thm as Thm (der, args)) =
+      let
+        val {cert, tags, maxidx, constraints, shyps, hyps, tpairs, prop} = args;
 
-      val thy = Context.certificate_theory cert;
-      val bad_thys =
-        Constraints.fold (fn {theory = thy', ...} =>
-          if Context.eq_thy (thy, thy') then I else cons thy') constraints [];
-      val () =
-        if null bad_thys then ()
-        else
-          raise THEORY ("solve_constraints: bad theories for theorem\n" ^
-            Syntax.string_of_term_global thy (prop_of thm), thy :: bad_thys);
+        val thy = Context.certificate_theory cert;
+        val bad_thys =
+          constraints |> map_filter (fn {theory = thy', ...} =>
+            if Context.eq_thy (thy, thy') then NONE else SOME thy');
+        val () =
+          if null bad_thys then ()
+          else
+            raise THEORY ("solve_constraints: bad theories for theorem\n" ^
+              Syntax.string_of_term_global thy (prop_of thm), thy :: bad_thys);
 
-      val Deriv {promises, body = PBody {oracles, thms, proof}} = der;
-      val (oracles', thms') = (oracles, thms)
-        |> Constraints.fold (fold union_digest o constraint_digest) constraints;
-      val body' = PBody {oracles = oracles', thms = thms', proof = proof};
-    in
-      Thm (Deriv {promises = promises, body = body'},
-        {constraints = Constraints.empty, cert = cert, tags = tags, maxidx = maxidx,
-          shyps = shyps, hyps = hyps, tpairs = tpairs, prop = prop})
-    end;
+        val Deriv {promises, body = PBody {oracles, thms, proof}} = der;
+        val (oracles', thms') = (oracles, thms)
+          |> fold (fold union_digest o constraint_digest) constraints;
+        val body' = PBody {oracles = oracles', thms = thms', proof = proof};
+      in
+        Thm (Deriv {promises = promises, body = body'},
+          {constraints = [], cert = cert, tags = tags, maxidx = maxidx,
+            shyps = shyps, hyps = hyps, tpairs = tpairs, prop = prop})
+      end;
 
 end;
 
 (*Dangling sort constraints of a thm*)
 fun extra_shyps (th as Thm (_, {shyps, ...})) =
-  Sortset.subtract (Sortset.build (fold_terms {hyps = true} Sortset.insert_term th)) shyps;
+  Sorts.subtract (fold_terms {hyps = true} Sorts.insert_term th []) shyps;
 
 (*Remove extra sorts that are witnessed by type signature information*)
-fun strip_shyps (thm as Thm (der, {cert, tags, maxidx, constraints, shyps, hyps, tpairs, prop})) =
-  solve_constraints (
-    if Sortset.is_empty shyps then thm
-    else
+fun strip_shyps thm =
+  (case thm of
+    Thm (_, {shyps = [], ...}) => thm
+  | Thm (der, {cert, tags, maxidx, constraints, shyps, hyps, tpairs, prop}) =>
       let
         val thy = theory_of_thm thm;
 
@@ -1040,29 +1037,30 @@
 
         val present =
           (fold_terms {hyps = true} o fold_types o fold_atyps_sorts) (insert (eq_fst op =)) thm [];
-        val extra = fold (Sortset.remove o #2) present shyps;
-        val witnessed = Sign.witness_sorts thy present (Sortset.dest extra);
-        val non_witnessed =
-          fold (Sortset.remove o #2) witnessed extra |> Sortset.dest |> map (`minimize);
+        val extra = fold (Sorts.remove_sort o #2) present shyps;
+        val witnessed = Sign.witness_sorts thy present extra;
+        val non_witnessed = fold (Sorts.remove_sort o #2) witnessed extra |> map (`minimize);
 
         val extra' =
           non_witnessed |> map_filter (fn (S, _) =>
             if non_witnessed |> exists (fn (S', _) => lt (S', S)) then NONE else SOME S)
-          |> Sortset.make;
+          |> Sorts.make;
 
         val constrs' =
           non_witnessed |> maps (fn (S1, S2) =>
-            let val S0 = the (Sortset.get_first (fn S => if le (S, S1) then SOME S else NONE) extra')
+            let val S0 = the (find_first (fn S => le (S, S1)) extra')
             in rel (S0, S1) @ rel (S1, S2) end);
 
         val constraints' = fold (insert_constraints thy) (witnessed @ constrs') constraints;
-        val shyps' = fold (Sortset.insert o #2) present extra';
+        val shyps' = fold (Sorts.insert_sort o #2) present extra';
       in
         Thm (deriv_rule_unconditional
           (Proofterm.strip_shyps_proof algebra present witnessed extra') der,
          {cert = cert, tags = tags, maxidx = maxidx, constraints = constraints',
           shyps = shyps', hyps = hyps, tpairs = tpairs, prop = prop})
-      end);
+      end)
+  |> solve_constraints;
+
 
 
 (*** Closed theorems with official name ***)
@@ -1073,12 +1071,12 @@
 
 (*non-deterministic, depends on unknown promises*)
 fun raw_derivation_name (Thm (Deriv {body, ...}, {shyps, hyps, prop, ...})) =
-  Proofterm.get_approximative_name shyps (Termset.dest hyps) prop (Proofterm.proof_of body);
+  Proofterm.get_approximative_name shyps hyps prop (Proofterm.proof_of body);
 
 fun expand_name (Thm (Deriv {body, ...}, {shyps, hyps, prop, ...})) =
   let
     val self_id =
-      (case Proofterm.get_identity shyps (Termset.dest hyps) prop (Proofterm.proof_of body) of
+      (case Proofterm.get_identity shyps hyps prop (Proofterm.proof_of body) of
         NONE => K false
       | SOME {serial, ...} => fn (header: Proofterm.thm_header) => serial = #serial header);
     fun expand header = if self_id header orelse #name header = "" then SOME "" else NONE;
@@ -1086,21 +1084,19 @@
 
 (*deterministic name of finished proof*)
 fun derivation_name (thm as Thm (_, {shyps, hyps, prop, ...})) =
-  Proofterm.get_approximative_name shyps (Termset.dest hyps) prop (proof_of thm);
+  Proofterm.get_approximative_name shyps hyps prop (proof_of thm);
 
 (*identified PThm node*)
 fun derivation_id (thm as Thm (_, {shyps, hyps, prop, ...})) =
-  Proofterm.get_id shyps (Termset.dest hyps) prop (proof_of thm);
+  Proofterm.get_id shyps hyps prop (proof_of thm);
 
 (*dependencies of PThm node*)
-fun thm_deps (thm as Thm (Deriv {promises, body = PBody {thms, ...}, ...}, _)) =
-  if not (null_promises promises) then raise THM ("thm_deps: bad promises", 0, [thm])
-  else if PThms.size thms = 1 then
-    (case (derivation_id thm, PThms.dest thms) of
-      (SOME {serial = i, ...}, [(j, thm_node)]) =>
-        if i = j then Proofterm.thm_node_thms thm_node else thms
-    | _ => thms)
-  else thms;
+fun thm_deps (thm as Thm (Deriv {promises = [], body = PBody {thms, ...}, ...}, _)) =
+      (case (derivation_id thm, thms) of
+        (SOME {serial = i, ...}, [(j, thm_node)]) =>
+          if i = j then Proofterm.thm_node_thms thm_node else thms
+      | _ => thms)
+  | thm_deps thm = raise THM ("thm_deps: bad promises", 0, [thm]);
 
 fun name_derivation name_pos =
   strip_shyps #> (fn thm as Thm (der, args) =>
@@ -1112,11 +1108,11 @@
 
       val _ = null tpairs orelse raise THM ("name_derivation: bad flex-flex constraints", 0, [thm]);
 
-      val ps = map (apsnd (Future.map fulfill_body)) (dest_promises promises);
+      val ps = map (apsnd (Future.map fulfill_body)) promises;
       val (pthm, proof) =
         Proofterm.thm_proof thy (classrel_proof thy) (arity_proof thy)
-          name_pos shyps (Termset.dest hyps) prop ps body;
-      val der' = make_deriv empty_promises Oracles.empty (PThms.make [pthm]) proof;
+          name_pos shyps hyps prop ps body;
+      val der' = make_deriv [] [] [pthm] proof;
     in Thm (der', args) end);
 
 fun close_derivation pos =
@@ -1147,13 +1143,13 @@
               | 0 => (((name, Position.none), NONE), MinProof)
               | i => bad_proofs i);
           in
-            Thm (make_deriv empty_promises (Oracles.make [oracle]) PThms.empty prf,
+            Thm (make_deriv [] [oracle] [] prf,
              {cert = Context.join_certificate (Context.Certificate thy', cert2),
               tags = [],
               maxidx = maxidx,
-              constraints = Constraints.empty,
+              constraints = [],
               shyps = sorts,
-              hyps = Termset.empty,
+              hyps = [],
               tpairs = [],
               prop = prop})
           end
@@ -1188,9 +1184,9 @@
      {cert = cert,
       tags = [],
       maxidx = ~1,
-      constraints = Constraints.empty,
+      constraints = [],
       shyps = sorts,
-      hyps = Termset.make [prop],
+      hyps = [prop],
       tpairs = [],
       prop = prop})
   end;
@@ -1213,8 +1209,8 @@
       tags = [],
       maxidx = Int.max (maxidx1, maxidx2),
       constraints = constraints,
-      shyps = Sortset.merge (sorts, shyps),
-      hyps = Termset.remove A hyps,
+      shyps = Sorts.union sorts shyps,
+      hyps = remove_hyps A hyps,
       tpairs = tpairs,
       prop = Logic.mk_implies (A, prop)});
 
@@ -1239,9 +1235,9 @@
            {cert = join_certificate2 (thAB, thA),
             tags = [],
             maxidx = Int.max (maxidx1, maxidx2),
-            constraints = Constraints.merge (constraintsA, constraints),
-            shyps = Sortset.merge (shypsA, shyps),
-            hyps = Termset.merge (hypsA, hyps),
+            constraints = union_constraints constraintsA constraints,
+            shyps = Sorts.union shypsA shyps,
+            hyps = union_hyps hypsA hyps,
             tpairs = union_tpairs tpairsA tpairs,
             prop = B})
         else err ()
@@ -1257,7 +1253,7 @@
 *)
 fun occs x ts tpairs =
   let fun occs t = Logic.occs (x, t)
-  in Termset.exists occs ts orelse exists (occs o fst) tpairs orelse exists (occs o snd) tpairs end;
+  in exists occs ts orelse exists (occs o fst) tpairs orelse exists (occs o snd) tpairs end;
 
 fun forall_intr
     (ct as Cterm {maxidx = maxidx1, t = x, T, sorts, ...})
@@ -1272,14 +1268,14 @@
           tags = [],
           maxidx = Int.max (maxidx1, maxidx2),
           constraints = constraints,
-          shyps = Sortset.merge (sorts, shyps),
+          shyps = Sorts.union sorts shyps,
           hyps = hyps,
           tpairs = tpairs,
           prop = Logic.all_const T $ Abs (a, T, abstract_over (x, prop))});
   in
     (case x of
       Free (a, _) => check_result a hyps
-    | Var ((a, _), _) => check_result a Termset.empty
+    | Var ((a, _), _) => check_result a []
     | _ => raise THM ("forall_intr: not a variable", 0, [th]))
   end;
 
@@ -1301,7 +1297,7 @@
           tags = [],
           maxidx = Int.max (maxidx1, maxidx2),
           constraints = constraints,
-          shyps = Sortset.merge (sorts, shyps),
+          shyps = Sorts.union sorts shyps,
           hyps = hyps,
           tpairs = tpairs,
           prop = Term.betapply (A, t)})
@@ -1318,9 +1314,9 @@
    {cert = cert,
     tags = [],
     maxidx = maxidx,
-    constraints = Constraints.empty,
+    constraints = [],
     shyps = sorts,
-    hyps = Termset.empty,
+    hyps = [],
     tpairs = [],
     prop = Logic.mk_equals (t, t)});
 
@@ -1364,9 +1360,9 @@
            {cert = join_certificate2 (th1, th2),
             tags = [],
             maxidx = Int.max (maxidx1, maxidx2),
-            constraints = Constraints.merge (constraints1, constraints2),
-            shyps = Sortset.merge (shyps1, shyps2),
-            hyps = Termset.merge (hyps1, hyps2),
+            constraints = union_constraints constraints1 constraints2,
+            shyps = Sorts.union shyps1 shyps2,
+            hyps = union_hyps hyps1 hyps2,
             tpairs = union_tpairs tpairs1 tpairs2,
             prop = eq $ t1 $ t2})
      | _ =>  err "premises"
@@ -1387,9 +1383,9 @@
      {cert = cert,
       tags = [],
       maxidx = maxidx,
-      constraints = Constraints.empty,
+      constraints = [],
       shyps = sorts,
-      hyps = Termset.empty,
+      hyps = [],
       tpairs = [],
       prop = Logic.mk_equals (t, t')})
   end;
@@ -1399,9 +1395,9 @@
    {cert = cert,
     tags = [],
     maxidx = maxidx,
-    constraints = Constraints.empty,
+    constraints = [],
     shyps = sorts,
-    hyps = Termset.empty,
+    hyps = [],
     tpairs = [],
     prop = Logic.mk_equals (t, Envir.eta_contract t)});
 
@@ -1410,9 +1406,9 @@
    {cert = cert,
     tags = [],
     maxidx = maxidx,
-    constraints = Constraints.empty,
+    constraints = [],
     shyps = sorts,
-    hyps = Termset.empty,
+    hyps = [],
     tpairs = [],
     prop = Logic.mk_equals (t, Envir.eta_long [] t)});
 
@@ -1437,7 +1433,7 @@
           tags = [],
           maxidx = maxidx,
           constraints = constraints,
-          shyps = Sortset.merge (sorts, shyps),
+          shyps = Sorts.union sorts shyps,
           hyps = hyps,
           tpairs = tpairs,
           prop = Logic.mk_equals
@@ -1445,7 +1441,7 @@
   in
     (case x of
       Free (a, _) => check_result a hyps
-    | Var ((a, _), _) => check_result a Termset.empty
+    | Var ((a, _), _) => check_result a []
     | _ => raise THM ("abstract_rule: not a variable", 0, [th]))
   end;
 
@@ -1476,9 +1472,9 @@
            {cert = join_certificate2 (th1, th2),
             tags = [],
             maxidx = Int.max (maxidx1, maxidx2),
-            constraints = Constraints.merge (constraints1, constraints2),
-            shyps = Sortset.merge (shyps1, shyps2),
-            hyps = Termset.merge (hyps1, hyps2),
+            constraints = union_constraints constraints1 constraints2,
+            shyps = Sorts.union shyps1 shyps2,
+            hyps = union_hyps hyps1 hyps2,
             tpairs = union_tpairs tpairs1 tpairs2,
             prop = Logic.mk_equals (f $ t, g $ u)}))
      | _ => raise THM ("combination: premises", 0, [th1, th2]))
@@ -1504,9 +1500,9 @@
            {cert = join_certificate2 (th1, th2),
             tags = [],
             maxidx = Int.max (maxidx1, maxidx2),
-            constraints = Constraints.merge (constraints1, constraints2),
-            shyps = Sortset.merge (shyps1, shyps2),
-            hyps = Termset.merge (hyps1, hyps2),
+            constraints = union_constraints constraints1 constraints2,
+            shyps = Sorts.union shyps1 shyps2,
+            hyps = union_hyps hyps1 hyps2,
             tpairs = union_tpairs tpairs1 tpairs2,
             prop = Logic.mk_equals (A, B)})
         else err "not equal"
@@ -1533,9 +1529,9 @@
            {cert = join_certificate2 (th1, th2),
             tags = [],
             maxidx = Int.max (maxidx1, maxidx2),
-            constraints = Constraints.merge (constraints1, constraints2),
-            shyps = Sortset.merge (shyps1, shyps2),
-            hyps = Termset.merge (hyps1, hyps2),
+            constraints = union_constraints constraints1 constraints2,
+            shyps = Sorts.union shyps1 shyps2,
+            hyps = union_hyps hyps1 hyps2,
             tpairs = union_tpairs tpairs1 tpairs2,
             prop = B})
         else err "not equal"
@@ -1599,7 +1595,7 @@
         | bad_term (Abs (_, T, t)) = bad_type T orelse bad_term t
         | bad_term (t $ u) = bad_term t orelse bad_term u
         | bad_term (Bound _) = false;
-      val _ = Termset.exists bad_term hyps andalso
+      val _ = exists bad_term hyps andalso
         raise THM ("generalize: variable free in assumptions", 0, [th]);
 
       val generalize = Term_Subst.generalize (tfrees, frees) idx;
@@ -1648,7 +1644,7 @@
 val add_instT_cert = add_cert (fn Ctyp {cert, ...} => cert);
 val add_inst_cert = add_cert (fn Cterm {cert, ...} => cert);
 
-fun add_sorts sorts_of (_, c) sorts = Sortset.merge (sorts_of c, sorts);
+fun add_sorts sorts_of (_, c) sorts = Sorts.union (sorts_of c) sorts;
 val add_instT_sorts = add_sorts (fn Ctyp {sorts, ...} => sorts);
 val add_inst_sorts = add_sorts (fn Cterm {sorts, ...} => sorts);
 
@@ -1757,9 +1753,9 @@
      {cert = cert,
       tags = [],
       maxidx = maxidx,
-      constraints = Constraints.empty,
+      constraints = [],
       shyps = sorts,
-      hyps = Termset.empty,
+      hyps = [],
       tpairs = [],
       prop = Logic.mk_implies (A, A)});
 
@@ -1781,9 +1777,9 @@
        {cert = cert,
         tags = [],
         maxidx = maxidx,
-        constraints = Constraints.build (insert_constraints thy (T, [c])),
+        constraints = insert_constraints thy (T, [c]) [],
         shyps = sorts,
-        hyps = Termset.empty,
+        hyps = [],
         tpairs = [],
         prop = prop})
     else raise THM ("of_class: type not of class " ^ Syntax.string_of_sort_global thy [c], 0, [])
@@ -1798,25 +1794,25 @@
       val thy = theory_of_thm thm;
 
       fun err msg = raise THM ("unconstrainT: " ^ msg, 0, [thm]);
-      val _ = Termset.is_empty hyps orelse err "bad hyps";
+      val _ = null hyps orelse err "bad hyps";
       val _ = null tpairs orelse err "bad flex-flex constraints";
       val tfrees = build_rev (Term.add_tfree_names prop);
       val _ = null tfrees orelse err ("illegal free type variables " ^ commas_quote tfrees);
 
-      val ps = map (apsnd (Future.map fulfill_body)) (dest_promises promises);
+      val ps = map (apsnd (Future.map fulfill_body)) promises;
       val (pthm, proof) =
         Proofterm.unconstrain_thm_proof thy (classrel_proof thy) (arity_proof thy)
           shyps prop ps body;
-      val der' = make_deriv empty_promises Oracles.empty (PThms.make [pthm]) proof;
+      val der' = make_deriv [] [] [pthm] proof;
       val prop' = Proofterm.thm_node_prop (#2 pthm);
     in
       Thm (der',
        {cert = cert,
         tags = [],
         maxidx = maxidx_of_term prop',
-        constraints = Constraints.empty,
-        shyps = Sortset.make [[]],  (*potentially redundant*)
-        hyps = Termset.empty,
+        constraints = [],
+        shyps = [[]],  (*potentially redundant*)
+        hyps = [],
         tpairs = [],
         prop = prop'})
     end);
@@ -1824,7 +1820,7 @@
 (*Replace all TFrees not fixed or in the hyps by new TVars*)
 fun varifyT_global' fixed (Thm (der, {cert, maxidx, constraints, shyps, hyps, tpairs, prop, ...})) =
   let
-    val tfrees = Termset.fold TFrees.add_tfrees hyps fixed;
+    val tfrees = fold TFrees.add_tfrees hyps fixed;
     val prop1 = attach_tpairs tpairs prop;
     val (al, prop2) = Type.varify_global tfrees prop1;
     val (ts, prop3) = Logic.strip_prems (length tpairs, [], prop2);
@@ -1865,9 +1861,9 @@
     val thm = strip_shyps raw_thm;
     fun err msg = raise THM ("plain_prop_of: " ^ msg, 0, [thm]);
   in
-    if not (Termset.is_empty (hyps_of thm)) then
+    if not (null (hyps_of thm)) then
       err "theorem may not contain hypotheses"
-    else if not (Sortset.is_empty (extra_shyps thm)) then
+    else if not (null (extra_shyps thm)) then
       err "theorem may not contain sort hypotheses"
     else if not (null (tpairs_of thm)) then
       err "theorem may not contain flex-flex pairs"
@@ -1903,7 +1899,7 @@
         tags = [],
         maxidx = maxidx + inc,
         constraints = constraints,
-        shyps = Sortset.merge (sorts, shyps),
+        shyps = Sorts.union shyps sorts,  (*sic!*)
         hyps = hyps,
         tpairs = map (apply2 lift_abs) tpairs,
         prop = Logic.list_implies (map lift_all As, lift_all B)})
@@ -2182,7 +2178,7 @@
                   (ntps, (map normt (Bs @ As), normt C))
              end
            val constraints' =
-             Constraints.merge (constraints1, constraints2)
+             union_constraints constraints1 constraints2
              |> insert_constraints_env (Context.certificate_theory cert) env;
            fun bicompose_proof prf1 prf2 =
              Proofterm.bicompose_proof flatten (map normt Bs) (map normt As) A oldAs n (nlift+1)
@@ -2195,8 +2191,8 @@
                 {tags = [],
                  maxidx = Envir.maxidx_of env,
                  constraints = constraints',
-                 shyps = Envir.insert_sorts env (Sortset.merge (shyps1, shyps2)),
-                 hyps = Termset.merge (hyps1, hyps2),
+                 shyps = Envir.insert_sorts env (Sorts.union shyps1 shyps2),
+                 hyps = union_hyps hyps1 hyps2,
                  tpairs = ntpairs,
                  prop = Logic.list_implies normp,
                  cert = cert})
@@ -2347,10 +2343,10 @@
 
 (* type arities *)
 
-fun thynames_of_arity thy (a, c) =
+fun theory_names_of_arity {long} thy (a, c) =
   build (get_arities thy |> Aritytab.fold
     (fn ((a', _, c'), (_, name, ser)) => (a = a' andalso c = c') ? cons (name, ser)))
-  |> sort (int_ord o apply2 #2) |> map #1;
+  |> sort (int_ord o apply2 #2) |> map (if long then #1 else Long_Name.base_name o #1);
 
 fun insert_arity_completions thy ((t, Ss, c), (th, thy_name, ser)) (finished, arities) =
   let
@@ -2408,7 +2404,7 @@
     val th' = th |> unconstrainT |> tap (expose_proof thy) |> trim_context;
     val prop = plain_prop_of th;
     val (t, Ss, c) = Logic.dest_arity prop;
-    val ar = ((t, Ss, c), (th', Context.theory_name thy, serial ()));
+    val ar = ((t, Ss, c), (th', Context.theory_long_name thy, serial ()));
   in
     thy
     |> Sign.primitive_arity (t, Ss, [c])
--- a/src/Pure/thm_deps.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/Pure/thm_deps.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -7,7 +7,7 @@
 
 signature THM_DEPS =
 sig
-  val all_oracles: thm list -> Oracles.T
+  val all_oracles: thm list -> Proofterm.oracle list
   val has_skip_proof: thm list -> bool
   val pretty_thm_oracles: Proof.context -> thm list -> Pretty.T
   val thm_deps: theory -> thm list -> (Proofterm.thm_id * Thm_Name.T) list
@@ -23,17 +23,17 @@
 fun all_oracles thms =
   let
     fun collect (PBody {oracles, thms, ...}) =
-      (if Oracles.is_empty oracles then I else apfst (cons oracles)) #>
-      (tap Proofterm.join_thms thms |> PThms.fold (fn (i, thm_node) => fn (res, seen) =>
+      (if null oracles then I else apfst (cons oracles)) #>
+      (tap Proofterm.join_thms thms |> fold (fn (i, thm_node) => fn (res, seen) =>
         if Intset.member seen i then (res, seen)
         else
           let val body = Future.join (Proofterm.thm_node_body thm_node)
           in collect body (res, Intset.insert i seen) end));
     val bodies = map Thm.proof_body_of thms;
-  in fold collect bodies ([], Intset.empty) |> #1 |> Oracles.merges end;
+  in fold collect bodies ([], Intset.empty) |> #1 |> Proofterm.unions_oracles end;
 
 fun has_skip_proof thms =
-  all_oracles thms |> Oracles.exists (fn ((name, _), _) => name = \<^oracle_name>\<open>skip_proof\<close>);
+  all_oracles thms |> exists (fn ((name, _), _) => name = \<^oracle_name>\<open>skip_proof\<close>);
 
 fun pretty_thm_oracles ctxt thms =
   let
@@ -44,7 +44,7 @@
           prt_ora ora @ [Pretty.str ":", Pretty.brk 1, Syntax.pretty_term_global thy prop];
     val oracles =
       (case try all_oracles thms of
-        SOME oracles => Oracles.dest oracles
+        SOME oracles => oracles
       | NONE => error "Malformed proofs")
   in Pretty.big_list "oracles:" (map (Pretty.item o prt_oracle) oracles) end;
 
@@ -63,11 +63,11 @@
               Inttab.update (i, SOME (thm_id, thm_name)) res
           | NONE =>
               Inttab.update (i, NONE) res
-              |> PThms.fold deps (Proofterm.thm_node_thms thm_node))
+              |> fold deps (Proofterm.thm_node_thms thm_node))
         end;
   in
     fn thms =>
-      (Inttab.build (fold (PThms.fold deps o Thm.thm_deps o Thm.transfer thy) thms), [])
+      (Inttab.build (fold (fold deps o Thm.thm_deps o Thm.transfer thy) thms), [])
       |-> Inttab.fold_rev (fn (_, SOME entry) => cons entry | _ => I)
   end;
 
--- a/src/Tools/Code/code_runtime.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/Tools/Code/code_runtime.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -718,7 +718,7 @@
         val code_binding = Path.map_binding Code_Target.code_path binding;
         val preamble =
           "(* Generated from " ^
-            Path.implode (Resources.thy_path (Path.basic (Context.theory_name thy))) ^
+            Path.implode (Resources.thy_path (Path.basic (Context.theory_base_name thy))) ^
           "; DO NOT EDIT! *)";
         val thy' = Code_Target.export code_binding (Bytes.string (preamble ^ "\n\n" ^ code)) thy;
         val _ = Code_Target.code_export_message thy';
--- a/src/Tools/Code/code_symbol.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/Tools/Code/code_symbol.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -99,7 +99,7 @@
 local
   val thyname_of_type = Name_Space.the_entry_theory_name o Sign.type_space;
   val thyname_of_class = Name_Space.the_entry_theory_name o Sign.class_space;
-  fun thyname_of_instance thy inst = case Thm.thynames_of_arity thy inst
+  fun thyname_of_instance thy inst = case Thm.theory_names_of_arity {long = false} thy inst
    of [] => error ("No such instance: " ^ quote (fst inst ^ " :: " ^  snd inst))
     | thyname :: _ => thyname;
   fun thyname_of_const thy c = case Axclass.class_of_param thy c
--- a/src/Tools/Code/code_target.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/Tools/Code/code_target.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -311,9 +311,14 @@
   (case pretty_modules of
     Singleton (_, p) => Bytes.write root (format p)
   | Hierarchy code_modules =>
-      (Isabelle_System.make_directory root;
-        List.app (fn (names, p) =>
-          Bytes.write (Path.appends (root :: map Path.basic names)) (format p)) code_modules));
+     List.app (fn (names, p) =>
+       let
+         val segments = map Path.basic names;
+       in
+         Isabelle_System.make_directory (Path.appends (root :: (fst (split_last segments))));
+         Bytes.write (Path.appends (root :: segments)) (format p)
+       end)
+     code_modules);
 
 in
 
--- a/src/Tools/Code/code_thingol.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/Tools/Code/code_thingol.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -1009,7 +1009,7 @@
   let
     val thy = Proof_Context.theory_of ctxt;
     fun this_theory name =
-      if Context.theory_name thy = name then thy
+      if Context.theory_base_name thy = name then thy
       else Context.get_theory {long = false} thy name;
 
     fun consts_of thy' =
--- a/src/Tools/IsaPlanner/rw_inst.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/Tools/IsaPlanner/rw_inst.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -64,8 +64,7 @@
        bound vars with binders outside the redex *)
 
     val ns =
-      IsaND.variant_names ctxt (Thm.full_prop_of rule :: Termset.dest (Thm.hyps_of rule))
-        (map fst Ts);
+      IsaND.variant_names ctxt (Thm.full_prop_of rule :: Thm.hyps_of rule) (map fst Ts);
 
     val (fromnames, tonames, Ts') =
       fold (fn (((faken, _), (n, ty)), n2) => fn (rnf, rnt, Ts'') =>
--- a/src/ZF/Tools/cartprod.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/ZF/Tools/cartprod.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -107,7 +107,7 @@
       in
         remove_split ctxt
           (Drule.instantiate_normalize (TVars.empty,
-            Vars.make [((v, \<^Type>\<open>i\<close> --> T2), Thm.cterm_of ctxt newt)]) rl)
+            Vars.make1 ((v, \<^Type>\<open>i\<close> --> T2), Thm.cterm_of ctxt newt)) rl)
       end
   | split_rule_var _ (t,T,rl) = rl;
 
--- a/src/ZF/Tools/inductive_package.ML	Wed Apr 12 12:00:40 2023 +0100
+++ b/src/ZF/Tools/inductive_package.ML	Mon Apr 24 15:00:54 2023 +0100
@@ -506,7 +506,7 @@
      val inst =
          case elem_frees of [_] => I
             | _ => Drule.instantiate_normalize (TVars.empty,
-                    Vars.make [((("x", 1), \<^Type>\<open>i\<close>), Thm.global_cterm_of thy4 elem_tuple)]);
+                    Vars.make1 ((("x", 1), \<^Type>\<open>i\<close>), Thm.global_cterm_of thy4 elem_tuple));
 
      (*strip quantifier and the implication*)
      val induct0 = inst (quant_induct RS @{thm spec} RSN (2, @{thm rev_mp}));