src/Pure/term_subst.ML
changeset 74220 c49134ee16c1
parent 74200 17090e27aae9
child 74222 bf595bfbdc98
--- a/src/Pure/term_subst.ML	Fri Sep 03 14:34:14 2021 +0200
+++ b/src/Pure/term_subst.ML	Fri Sep 03 18:57:33 2021 +0200
@@ -4,8 +4,31 @@
 Efficient type/term substitution.
 *)
 
+signature INST_TABLE =
+sig
+  include TABLE
+  val add: key * 'a -> 'a table -> 'a table
+  val table: (key * 'a) list -> 'a table
+end;
+
+functor Inst_Table(Key: KEY): INST_TABLE =
+struct
+
+structure Tab = Table(Key);
+
+fun add entry = Tab.insert (K true) entry;
+fun table entries = fold add entries Tab.empty;
+
+open Tab;
+
+end;
+
 signature TERM_SUBST =
 sig
+  structure TFrees: INST_TABLE
+  structure TVars: INST_TABLE
+  structure Frees: INST_TABLE
+  structure Vars: INST_TABLE
   val map_atypsT_same: typ Same.operation -> typ Same.operation
   val map_types_same: typ Same.operation -> term Same.operation
   val map_aterms_same: term Same.operation -> term Same.operation
@@ -13,24 +36,18 @@
   val generalize_same: Symtab.set * Symtab.set -> int -> term Same.operation
   val generalizeT: Symtab.set -> int -> typ -> typ
   val generalize: Symtab.set * Symtab.set -> int -> term -> term
-  val instantiateT_maxidx: ((indexname * sort) * (typ * int)) list -> typ -> int -> typ * int
-  val instantiate_maxidx:
-    ((indexname * sort) * (typ * int)) list * ((indexname * typ) * (term * int)) list ->
+  val instantiateT_maxidx: (typ * int) TVars.table -> typ -> int -> typ * int
+  val instantiate_maxidx: (typ * int) TVars.table * (term * int) Vars.table ->
     term -> int -> term * int
-  val instantiateT_frees_same: ((string * sort) * typ) list -> typ Same.operation
-  val instantiate_frees_same: ((string * sort) * typ) list * ((string * typ) * term) list ->
-    term Same.operation
-  val instantiateT_frees: ((string * sort) * typ) list -> typ -> typ
-  val instantiate_frees: ((string * sort) * typ) list * ((string * typ) * term) list ->
-    term -> term
-  val instantiateT_same: ((indexname * sort) * typ) list -> typ Same.operation
-  val instantiate_same: ((indexname * sort) * typ) list * ((indexname * typ) * term) list ->
-    term Same.operation
-  val instantiateT: ((indexname * sort) * typ) list -> typ -> typ
-  val instantiate: ((indexname * sort) * typ) list * ((indexname * typ) * term) list ->
-    term -> term
-  val zero_var_indexes_inst: Name.context -> term list ->
-    ((indexname * sort) * typ) list * ((indexname * typ) * term) list
+  val instantiateT_frees_same: typ TFrees.table -> typ Same.operation
+  val instantiate_frees_same: typ TFrees.table * term Frees.table -> term Same.operation
+  val instantiateT_frees: typ TFrees.table -> typ -> typ
+  val instantiate_frees: typ TFrees.table * term Frees.table -> term -> term
+  val instantiateT_same: typ TVars.table -> typ Same.operation
+  val instantiate_same: typ TVars.table * term Vars.table -> term Same.operation
+  val instantiateT: typ TVars.table -> typ -> typ
+  val instantiate: typ TVars.table * term Vars.table -> term -> term
+  val zero_var_indexes_inst: Name.context -> term list -> typ TVars.table * term Vars.table
   val zero_var_indexes: term -> term
   val zero_var_indexes_list: term list -> term list
 end;
@@ -38,6 +55,29 @@
 structure Term_Subst: TERM_SUBST =
 struct
 
+(* instantiation as table *)
+
+structure TFrees = Inst_Table(
+  type key = string * sort
+  val ord = prod_ord fast_string_ord Term_Ord.sort_ord
+);
+
+structure TVars = Inst_Table(
+  type key = indexname * sort
+  val ord = prod_ord Term_Ord.fast_indexname_ord Term_Ord.sort_ord
+);
+
+structure Frees = Inst_Table(
+  type key = string * typ
+  val ord = prod_ord fast_string_ord Term_Ord.typ_ord
+);
+
+structure Vars = Inst_Table(
+  type key = indexname * typ
+  val ord = prod_ord Term_Ord.fast_indexname_ord Term_Ord.typ_ord
+);
+
+
 (* generic mapping *)
 
 fun map_atypsT_same f =
@@ -103,35 +143,37 @@
 
 (* instantiation of free variables (types before terms) *)
 
-fun instantiateT_frees_same [] _ = raise Same.SAME
-  | instantiateT_frees_same instT ty =
-      let
-        fun subst (Type (a, Ts)) = Type (a, Same.map subst Ts)
-          | subst (TFree v) =
-              (case AList.lookup (op =) instT v of
-                SOME T => T
-              | NONE => raise Same.SAME)
-          | subst _ = raise Same.SAME;
-      in subst ty end;
+fun instantiateT_frees_same instT ty =
+  if TFrees.is_empty instT then raise Same.SAME
+  else
+    let
+      fun subst (Type (a, Ts)) = Type (a, Same.map subst Ts)
+        | subst (TFree v) =
+            (case TFrees.lookup instT v of
+              SOME T => T
+            | NONE => raise Same.SAME)
+        | subst _ = raise Same.SAME;
+    in subst ty end;
 
-fun instantiate_frees_same ([], []) _ = raise Same.SAME
-  | instantiate_frees_same (instT, inst) tm =
-      let
-        val substT = instantiateT_frees_same instT;
-        fun subst (Const (c, T)) = Const (c, substT T)
-          | subst (Free (x, T)) =
-              let val (T', same) = (substT T, false) handle Same.SAME => (T, true) in
-                (case AList.lookup (op =) inst (x, T') of
-                   SOME t => t
-                 | NONE => if same then raise Same.SAME else Free (x, T'))
-              end
-          | subst (Var (xi, T)) = Var (xi, substT T)
-          | subst (Bound _) = raise Same.SAME
-          | subst (Abs (x, T, t)) =
-              (Abs (x, substT T, Same.commit subst t)
-                handle Same.SAME => Abs (x, T, subst t))
-          | subst (t $ u) = (subst t $ Same.commit subst u handle Same.SAME => t $ subst u);
-      in subst tm end;
+fun instantiate_frees_same (instT, inst) tm =
+  if TFrees.is_empty instT andalso Frees.is_empty inst then raise Same.SAME
+  else
+    let
+      val substT = instantiateT_frees_same instT;
+      fun subst (Const (c, T)) = Const (c, substT T)
+        | subst (Free (x, T)) =
+            let val (T', same) = (substT T, false) handle Same.SAME => (T, true) in
+              (case Frees.lookup inst (x, T') of
+                 SOME t => t
+               | NONE => if same then raise Same.SAME else Free (x, T'))
+            end
+        | subst (Var (xi, T)) = Var (xi, substT T)
+        | subst (Bound _) = raise Same.SAME
+        | subst (Abs (x, T, t)) =
+            (Abs (x, substT T, Same.commit subst t)
+              handle Same.SAME => Abs (x, T, subst t))
+        | subst (t $ u) = (subst t $ Same.commit subst u handle Same.SAME => t $ subst u);
+    in subst tm end;
 
 fun instantiateT_frees instT ty = Same.commit (instantiateT_frees_same instT) ty;
 fun instantiate_frees inst tm = Same.commit (instantiate_frees_same inst) tm;
@@ -141,9 +183,8 @@
 
 local
 
-fun no_index (x, y) = (x, (y, ~1));
-fun no_indexes1 inst = map no_index inst;
-fun no_indexes2 (inst1, inst2) = (map no_index inst1, map no_index inst2);
+fun no_indexesT instT = TVars.map (fn _ => rpair ~1) instT;
+fun no_indexes inst = Vars.map (fn _ => rpair ~1) inst;
 
 fun instT_same maxidx instT ty =
   let
@@ -151,7 +192,7 @@
 
     fun subst_typ (Type (a, Ts)) = Type (a, subst_typs Ts)
       | subst_typ (TVar ((a, i), S)) =
-          (case AList.lookup Term.eq_tvar instT ((a, i), S) of
+          (case TVars.lookup instT ((a, i), S) of
             SOME (T, j) => (maxify j; T)
           | NONE => (maxify i; raise Same.SAME))
       | subst_typ _ = raise Same.SAME
@@ -170,7 +211,7 @@
       | subst (Free (x, T)) = Free (x, substT T)
       | subst (Var ((x, i), T)) =
           let val (T', same) = (substT T, false) handle Same.SAME => (T, true) in
-            (case AList.lookup Term.eq_var inst ((x, i), T') of
+            (case Vars.lookup inst ((x, i), T') of
                SOME (t, j) => (maxify j; t)
              | NONE => (maxify i; if same then raise Same.SAME else Var ((x, i), T')))
           end
@@ -191,11 +232,13 @@
   let val maxidx = Unsynchronized.ref i
   in (Same.commit (inst_same maxidx insts) tm, ! maxidx) end;
 
-fun instantiateT_same [] _ = raise Same.SAME
-  | instantiateT_same instT ty = instT_same (Unsynchronized.ref ~1) (no_indexes1 instT) ty;
+fun instantiateT_same instT ty =
+  if TVars.is_empty instT then raise Same.SAME
+  else instT_same (Unsynchronized.ref ~1) (no_indexesT instT) ty;
 
-fun instantiate_same ([], []) _ = raise Same.SAME
-  | instantiate_same insts tm = inst_same (Unsynchronized.ref ~1) (no_indexes2 insts) tm;
+fun instantiate_same (instT, inst) tm =
+  if TVars.is_empty instT andalso Vars.is_empty inst then raise Same.SAME
+  else inst_same (Unsynchronized.ref ~1) (no_indexesT instT, no_indexes inst) tm;
 
 fun instantiateT instT ty = Same.commit (instantiateT_same instT) ty;
 fun instantiate inst tm = Same.commit (instantiate_same inst) tm;
@@ -205,26 +248,20 @@
 
 (* zero var indexes *)
 
-structure TVars = Table(type key = indexname * sort val ord = Term_Ord.tvar_ord);
-structure Vars = Table(type key = indexname * typ val ord = Term_Ord.var_ord);
-
-fun zero_var_inst mk (v as ((x, i), X)) (inst, used) =
-  let
-    val (x', used') = Name.variant (if Name.is_bound x then "u" else x) used;
-  in if x = x' andalso i = 0 then (inst, used') else ((v, mk ((x', 0), X)) :: inst, used') end;
+fun zero_var_inst ins mk (v as ((x, i), X)) (inst, used) =
+  let val (x', used') = Name.variant (if Name.is_bound x then "u" else x) used
+  in if x = x' andalso i = 0 then (inst, used') else (ins (v, mk ((x', 0), X)) inst, used') end;
 
 fun zero_var_indexes_inst used ts =
   let
     val (instT, _) =
-      TVars.fold (zero_var_inst TVar o #1)
+      (TVars.empty, used) |> TVars.fold (zero_var_inst TVars.add TVar o #1)
         ((fold o fold_types o fold_atyps) (fn TVar v =>
-          TVars.insert (K true) (v, ()) | _ => I) ts TVars.empty)
-        ([], used);
+          TVars.add (v, ()) | _ => I) ts TVars.empty);
     val (inst, _) =
-      Vars.fold (zero_var_inst Var o #1)
+      (Vars.empty, used) |> Vars.fold (zero_var_inst Vars.add Var o #1)
         ((fold o fold_aterms) (fn Var (xi, T) =>
-          Vars.insert (K true) ((xi, instantiateT instT T), ()) | _ => I) ts Vars.empty)
-        ([], used);
+          Vars.add ((xi, instantiateT instT T), ()) | _ => I) ts Vars.empty);
   in (instT, inst) end;
 
 fun zero_var_indexes_list ts = map (instantiate (zero_var_indexes_inst Name.context ts)) ts;