diff --git a/thys/ML_Unification/Binders/ML_Binders.thy b/thys/ML_Unification/Binders/ML_Binders.thy new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/Binders/ML_Binders.thy @@ -0,0 +1,14 @@ +\<^marker>\creator "Kevin Kappelmann"\ +section \ML-Binders\ +theory ML_Binders + imports + ML_General_Utils + ML_Normalisations +begin + +paragraph \Summary\ +text \Binders for ML.\ + +ML_file\binders.ML\ + +end diff --git a/thys/ML_Unification/Binders/binders.ML b/thys/ML_Unification/Binders/binders.ML new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/Binders/binders.ML @@ -0,0 +1,59 @@ +(* Title: ML_Unification/binders.ML + Author: Kevin Kappelmann + +Binders with De Bruijn indexing. + +When we traverse abstractions, e.g. in unification algorithms, we store a binder as +((name, typ), fresh Free variable). The fresh free variable can be used as a replacement for +the corresponding bound variable, e.g. when one has to create a reflexivity theorem for the +binder. +*) +signature BINDERS = +sig + type 'a binder = (string * typ) * 'a + type 'a binders = ('a binder) list + (*add_binders bs1 bs2 adds binders bs1 to binders bs2*) + val add_binders : 'a binders -> 'a binders -> 'a binders + val fix_binders : (string * typ) list -> Proof.context -> term binders * Proof.context + val fix_binder : (string * typ) -> Proof.context -> term binder * Proof.context + (*returns types of binders*) + val binder_types : 'a binders -> typ list + (*returns free variable corresponding to the given binder*) + val nth_binder_data : 'a binders -> int -> 'a + (*replaces all binders by their corresponding free variable in the given term*) + val replace_binders : term binders -> term -> term + (*replaces all free variables that correspond to a binder by their binder*) + val replace_frees : term binders -> term -> term + (*normalises types and associated free variables*) + val norm_binders : Term_Normalisation.term_normaliser -> term binders -> term binders +end + +structure Binders : BINDERS = +struct + +type 'a binder = (string * typ) * 'a +type 'a binders = ('a binder) list + +fun add_binders binders1 binders2 = binders1 @ binders2 + +fun fix_binders nTs ctxt = + Variable.variant_fixes (map fst nTs) ctxt + |>> map2 (fn (n, T) => fn n' => ((n, T), Free (n', T))) nTs + +val fix_binder = yield_singleton fix_binders + +fun binder_types binders = map (snd o fst) binders + +fun nth_binder_data binders = snd o nth binders + +fun replace_binders binders t = let val bvars = map snd binders + in subst_bounds (bvars, t) end + +fun replace_frees binders = + let val subst = General_Util.fold_rev_index (fn (i, (_, free)) => cons (free, Bound i)) binders [] + in subst_free subst end + +fun norm_binders norm = + map (fn ((n, _), t) => let val t' = norm t in ((n, snd (dest_Free t')), t') end) + +end diff --git a/thys/ML_Unification/Examples/E_Unification_Examples.thy b/thys/ML_Unification/Examples/E_Unification_Examples.thy new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/Examples/E_Unification_Examples.thy @@ -0,0 +1,155 @@ +\<^marker>\creator "Kevin Kappelmann"\ +section \E-Unification Examples\ +theory E_Unification_Examples + imports + Main + ML_Unification_HOL_Setup + Unify_Fact_Tactic +begin + +paragraph \Summary\ +text \Sample applications of e-unifiers, methods, etc. introduced in this session.\ + +experiment +begin + + +subsection \Using The Simplifier For Unification.\ + +inductive_set even :: "nat set" where +zero: "0 \ even" | +step: "n \ even \ Suc (Suc n) \ even" + +text \Premises of the form @{term "SIMPS_TO_UNIF lhs rhs"} are solved by +@{ML_structure Simplifier_Unification}. It first normalises @{term lhs} and then unifies the +normalisation with @{term rhs}. See also @{theory ML_Unification.ML_Unification_HOL_Setup}.\ + +lemma [uhint where prio = Prio.LOW]: "n \ 0 \ PROP SIMPS_TO_UNIF (n - 1) m \ n \ Suc m" + unfolding SIMPS_TO_UNIF_eq by linarith + +text \By default, below unification methods use +@{ML Standard_Mixed_Unification.first_higherp_first_comb_higher_unify}, which is a combination of +various practical unification algorithms.\ + +schematic_goal "(\x. x + 4 = n) \ Suc ?x = n" + by uassm + +lemma "6 \ even" + apply (urule step) + apply (urule step) + apply (urule step) + apply (urule zero) + done + +lemma "(220 + (80 - 2 * 2)) \ even" + apply (urule step) + apply (urule step)+ + apply (urule zero) + done + +lemma + assumes "[a,b,c] = [c,b,a]" + shows "[a] @ [b,c] = [c,b,a]" + using assms by uassm + +lemma "x \ ({z, y, x} \ S) \ {x}" + by (ufact TrueI) + +lemma "(x + (y :: nat))^2 \ x^2 + 2*x*y + y^2 + 4 * y + x - y" + supply power2_sum[simp] + by (ufact TrueI) + +lemma + assumes "\s. P (Suc (Suc 0)) (s(x := (1 :: nat), x := 1 + 1 * 4 - 3))" + shows "P 2 (s(x := 2))" + by (ufact assms[of s]) + +subsection \Providing Canonical Solutions With Unification Hints\ + +lemma [uhint]: "xs \ [] \ length xs \ 0" by simp + +schematic_goal "length ?xs = 0" + by (ufact refl) + +lemma [uhint]: "(n :: nat) \ m \ n - m \ 0" by simp + +schematic_goal "n - ?m = (0 :: nat)" + by (ufact refl) + +text \The following fails because, by default, @{ML Standard_Unification_Hints.try_hints} +uses the higher-order pattern unifier to unify hints against a given disagreement pair, and +@{term 0} cannot be higher-order pattern unified with @{term "length []"}. The unification of the +hint requires the use of yet another hint, namely @{term "length xs = 0"} (cf. above).\ +schematic_goal "n - ?m = length []" + \ \by (ufact refl)\ + oops + +text \There are two ways to fix this: +\<^enum> We allow the recursive uses of unification hints when searching for suitable unification hints. +\<^enum> We use a different unification hint that the recursive use of hints explicit.\ + +text \Solution 1: recursive usages of hints. Warning: such recursive applications easily loop.\ +schematic_goal "n - ?m = length []" + using [[uhint where concl_unifier = Standard_Mixed_Unification.first_higherp_first_comb_higher_unify]] + by (ufact refl) + +text \Solution 2: make the recursion explicit in the hint.\ + +lemma [uhint]: "k \ 0 \ (n :: nat) \ m \ n - m \ k" by simp + +schematic_goal "n - ?m = length []" + by (ufact refl) + + +subsection \Strenghten Unification With Unification Hints\ + +lemma + assumes [uhint]: "n = m" + shows "n - m = (0 :: nat)" + by (ufact refl) + +lemma + assumes "x = y" + shows "y = x" + supply eq_commute[uhint] + by (ufact assms) + + +paragraph \Unfolding definitions.\ + +definition "mysuc n = Suc n" + +lemma + assumes "\m. Suc n > mysuc m" + shows "mysuc n > Suc 3" + supply mysuc_def[uhint] + by (ufact assms) + + +paragraph \Discharging meta impliciations with object-level implications\ + +lemma [uhint]: + "Trueprop A \ A' \ Trueprop B \ B' \ Trueprop (A \ B) \ (PROP A' \ PROP B')" + using atomize_imp[symmetric] by simp + +lemma + assumes "A \ (B \ C) \ D" + shows "A \ (B \ C) \ D" + using assms by ufact + + +subsection \Better Control Over Meta Variable Instantiations\ + +text \Consider the following type-inference problem.\ +schematic_goal + assumes app_typeI: "\f x. (\x. ArgT x \ DomT x (f x)) \ ArgT x \ DomT x (f x)" + and f_type: "\x. ArgT x \ DomT x (f x)" + and x_type: "ArgT x" + shows "?T (f x)" + apply (urule app_typeI) \\compare with the following application, creating an (unintuitive) higher-order instantiation\ + (* apply (rule app_typeI) *) + oops + +end + +end diff --git a/thys/ML_Unification/Examples/Unification_Hints_Reification_Examples.thy b/thys/ML_Unification/Examples/Unification_Hints_Reification_Examples.thy new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/Examples/Unification_Hints_Reification_Examples.thy @@ -0,0 +1,212 @@ +\<^marker>\creator "Kevin Kappelmann"\ +\<^marker>\contributor "Paul Bachmann"\ +section \Examples: Reification Via Unification Hints\ +theory Unification_Hints_Reification_Examples + imports + HOL.Rat + ML_Unification_HOL_Setup + Unify_Fact_Tactic +begin + +paragraph \Summary\ +text \Reification via unification hints. For an introduction to unification hints refer +to \<^cite>\"unif-hints"\. We support a generalisation of unification hints as described +in @{theory ML_Unification.ML_Unification_Hints}.\ + +subsection \Setup\ + +text \One-time setup to obtain a unifier with unification hints for the purpose of reification. +We could also simply use the standard unification hints @{attribute uhint}, +but having separate instances is a cleaner approach.\ + +ML\ + @{functor_instance struct_name = Reification_Unification_Hints + and functor_name = Term_Index_Unification_Hints + and id = \"reify"\ + and more_args = \ + structure TI = Discrimination_Tree + val init_args = { + normalisers = SOME Higher_Order_Pattern_Unification.norms_unify, + retrieval = SOME (Term_Index_Unification_Hints_Args.mk_sym_retrieval + TI.norm_term TI.unifiables), + prems_unifier = NONE, + concl_unifier = NONE, + hint_preprocessor = SOME (Standard_Unification_Hints.get_hint_preprocessor + (Context.the_generic_context ())) + }\} + val reify_unify = Unification_Combinator.add_fallback_unifier + (fn unif_theory => + Higher_Order_Pattern_Unification.e_unify Unification_Util.unify_types unif_theory unif_theory) + (Reification_Unification_Hints.try_hints + |> Unification_Combinator.norm_unifier (#norm_term Higher_Order_Pattern_Unification.norms_unify)) +\ +local_setup \Reification_Unification_Hints.setup_attribute NONE\ + +text \Premises of hints should again be unified by the reification unifier.\ +declare [[reify_uhint where prems_unifier = reify_unify]] + +subsection \Formulas with Quantifiers and Environment\ + +text \The following example is taken from HOL-Library.Reflection\_Examples. It is recommended +to compare the approach presented here with the reflection tactic presented in said theory.\ + +datatype form = + TrueF +| FalseF +| Less nat nat +| And form form +| Or form form +| Neg form +| ExQ form + +primrec interp :: "form \ ('a::ord) list \ bool" +where + "interp TrueF vs \ True" +| "interp FalseF vs \ False" +| "interp (Less i j) vs \ vs ! i < vs ! j" +| "interp (And f1 f2) vs \ interp f1 vs \ interp f2 vs" +| "interp (Or f1 f2) vs \ interp f1 vs \ interp f2 vs" +| "interp (Neg f) vs \ \ interp f vs" +| "interp (ExQ f) vs \ (\v. interp f (v # vs))" + +paragraph \Reification with unification and recursive hint unification for conclusion\ + +text \The following illustrates how to use the equations @{thm interp.simps} directly +as unification hints for reification.\ + +experiment +begin + +text \Hints for list lookup.\ +declare List.nth_Cons_Suc[reify_uhint where prio = Prio.LOW] + and List.nth_Cons_0[reify_uhint] + +text \Hints to reify formulas of type @{type bool} into formulas of type @{type form}.\ +declare interp.simps[reify_uhint] + +text \We have to allow the hint unifier to recursively look for hints during unification of +the hint's conclusion.\ +declare [[reify_uhint where concl_unifier = reify_unify]] + +(*uncomment the following to see the hint unification process*) +(* declare [[ML_map_context \Logger.set_log_levels Unification_Base.logger Logger.DEBUG\]] *) + +schematic_goal + "interp ?f (?vs :: ('a :: ord) list) = (\(x :: 'a). x < y \ \(\(z :: 'a). v < z \ \False))" + by (ufact refl where unifier = reify_unify) + +text \While this all works nicely if set up correctly, it can be rather difficult to +understand and debug the recursive unification process for a hint's conclusion. +In the next paragraph, we present an alternative that is closer to the examples presented +in the original unification hints paper \<^cite>\"unif-hints"\.\ + +end + +paragraph \Reification with matching without recursion for conclusion\ + +text \We disallow the hint unifier to recursively look for hints while unifying the conclusion; +instead, we only allow the hint unifier to match the hint's conclusion against the disagreement terms.\ +declare [[reify_uhint where concl_unifier = Higher_Order_Pattern_Unification.match]] + +text \However, this also means that we now have to write our hints such that the hint's +conclusion can successfully be matched against the disagreement terms. In particular, +the disagreement terms may still contain meta variables that we want to instantiate with the help +of the unification hints. Essentially, a hint then describes a canonical instantiation for +these meta variables.\ + +experiment +begin + +lemma [reify_uhint where prio = Prio.LOW]: + "n \ Suc n' \ vs \ v # vs' \ vs' ! n' \ x \ vs ! n \ x" + by simp + +lemma [reify_uhint]: "n \ 0 \ vs \ x # vs' \ vs ! n \ x" + by simp + +lemma [reify_uhint]: + "\e \ ExQ f; \v. interp f (v # vs) \ P v\ \ interp e vs \ \v. P v" + "\e \ Less i j; x \ vs ! i; y \ vs ! j\ \ interp e vs \ x < y" + "\e \ And f1 f2; interp f1 vs \ r1; interp f2 vs \ r2\ \ interp e vs \ r1 \ r2" + "\e \ Or f1 f2; interp f1 vs \ r1; interp f2 vs \ r2\ \ interp e vs \ r1 \ r2" + "e \ Neg f \ interp f vs \ r \ interp e vs \ \r" + "e \ TrueF \ interp e vs \ True" + "e \ FalseF \ interp e vs \ False" + by simp_all + +schematic_goal + "interp ?f (?vs :: ('a :: ord) list) = (\(x :: 'a). x < y \ \(\(z :: 'a). v < z \ \False))" + by (urule refl where unifier = reify_unify) +end + +text \The next examples are modification from \<^cite>\"unif-hints"\.\ + +subsection \Simple Arithmetic\ + +datatype add_expr = Var int | Add add_expr add_expr + +fun eval_add_expr :: "add_expr \ int" where + "eval_add_expr (Var i) = i" +| "eval_add_expr (Add ex1 ex2) = eval_add_expr ex1 + eval_add_expr ex2" + +lemma eval_add_expr_Var [reify_uhint where prio = Prio.LOW]: + "e \ Var i \ eval_add_expr e \ i" by simp + +lemma eval_add_expr_add [reify_uhint]: + "e \ Add e1 e2 \ eval_add_expr e1 \ m \ eval_add_expr e2 \ n \ eval_add_expr e \ m + n" + by simp + +ML_command\ + val t1 = Proof_Context.read_term_pattern @{context} "eval_add_expr ?e" + val t2 = Proof_Context.read_term_pattern @{context} "1 + (2 + 7) :: int" + val _ = Unification_Util.log_unif_results @{context} (t1, t2) (reify_unify []) +\ + +schematic_goal "eval_add_expr ?e = (1 + (2 + 7) :: int)" + by (urule refl where unifier = reify_unify) + + +subsection \Arithmetic with Environment\ + +datatype mul_expr = + Unit +| Var nat +| Mul mul_expr mul_expr +| Inv mul_expr + +fun eval_mul_expr :: "mul_expr \ rat list \ rat" where + "eval_mul_expr (Unit, \) = 1" +| "eval_mul_expr (Var i, \) = \ ! i" +| "eval_mul_expr (Mul e1 e2, \) = eval_mul_expr (e1, \) * eval_mul_expr (e2, \)" +| "eval_mul_expr (Inv e, \) = inverse (eval_mul_expr (e, \))" + +text \Split @{term e} into an expression and an environment.\ +lemma [reify_uhint where prio = Prio.VERY_LOW]: + "e \ (e1, \) \ eval_mul_expr (e1, \) \ n \ eval_mul_expr e \ n" + by simp + +text \Hints for environment lookup.\ +lemma [reify_uhint where prio = Prio.LOW]: + "e \ Var (Suc p) \ \ \ s # \ \ n \ eval_mul_expr (Var p, \) \ eval_mul_expr (e, \) \ n" + by simp + +lemma [reify_uhint]: "e \ Var 0 \ \ \ n # \ \ eval_mul_expr (e, \) \ n" + by simp + +lemma [reify_uhint]: + "e1 \ Inv e2 \ n \ eval_mul_expr (e2, \) \ eval_mul_expr (e1, \) \ inverse n" + "e \ Mul e1 e2 \ m \ eval_mul_expr (e1, \) \ + n \ eval_mul_expr (e2, \) \ eval_mul_expr (e, \) \ m * n" + "e \ Unit \ eval_mul_expr (e, \) \ 1" + by simp_all + +ML_command\ + val t1 = Proof_Context.read_term_pattern @{context} "eval_mul_expr ?e" + val t2 = Proof_Context.read_term_pattern @{context} "1 * inverse 3 * 5 :: rat" + val _ = Unification_Util.log_unif_results' 1 @{context} (t2, t1) (reify_unify []) +\ + +schematic_goal "eval_mul_expr ?e = (1 * inverse 3 * 5 :: rat)" + by (ufact refl where unifier = reify_unify) + +end diff --git a/thys/ML_Unification/Logger/Data_Structures/binding_tree.ML b/thys/ML_Unification/Logger/Data_Structures/binding_tree.ML new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/Logger/Data_Structures/binding_tree.ML @@ -0,0 +1,63 @@ +(* Title: Logger/Data_Structures/binding_tree.ML + Author: Kevin Kappelmann + +A hierarchical tree indexed on bindings. +*) +signature BINDING_TREE = +sig + type 'a binding_tree + val empty : 'a binding_tree + val is_empty : 'a binding_tree -> bool + + val lookup: 'a binding_tree -> binding -> 'a option + + exception INSERT + (*raises INSERT if the entry is already set*) + val insert : (binding * 'a) -> 'a binding_tree -> 'a binding_tree + val insert_safe : (binding * 'a) -> 'a binding_tree -> 'a binding_tree + + exception DELETE + (*deletes entry; raises DELETE if the corresponding entry is not set*) + val delete : binding -> 'a binding_tree -> 'a binding_tree + val delete_safe : binding -> 'a binding_tree -> 'a binding_tree + (*like delete but also cuts off all subtrees*) + val cut : binding -> 'a binding_tree -> 'a binding_tree + (*does not cut subtrees if entry is not set*) + val cut_safe : binding -> 'a binding_tree -> 'a binding_tree + + (*maps entry*) + val map : binding -> ('a option -> 'a option) -> 'a binding_tree -> 'a binding_tree + (*maps entry and all entries in subtrees*) + val map_below : binding -> ('a option -> 'a option) -> 'a binding_tree -> 'a binding_tree + + val join : ('a option * 'a option -> 'a option) -> 'a binding_tree * 'a binding_tree -> + 'a binding_tree + val merge : 'a binding_tree * 'a binding_tree -> 'a binding_tree +end + +structure Binding_Tree : BINDING_TREE = +struct + +structure HOT = HOption_Tree(TableMap(Symtab)) + +fun full_path_of binding = (Binding.path_of binding |> map fst) @ [Binding.name_of binding] + +type 'a binding_tree = 'a HOT.hoption_tree +val empty = HOT.empty +val is_empty = HOT.is_empty +fun lookup bt = HOT.lookup bt o full_path_of +exception INSERT = HOT.INSERT +fun insert p = HOT.insert (apfst full_path_of p) +fun insert_safe p = HOT.insert_safe (apfst full_path_of p) +exception DELETE = HOT.DELETE +fun delete binding = HOT.delete (full_path_of binding) +fun delete_safe binding = HOT.delete_safe (full_path_of binding) +fun cut binding = HOT.cut (full_path_of binding) +fun cut_safe binding = HOT.cut_safe (full_path_of binding) +fun map binding = HOT.map (full_path_of binding) +fun map_below binding = HOT.map_below (full_path_of binding) +val join = HOT.join + +fun merge opttp = join (fn (data1, data2) => if is_none data1 then data2 else data1) opttp + +end diff --git a/thys/ML_Unification/Logger/Data_Structures/hoption_tree.ML b/thys/ML_Unification/Logger/Data_Structures/hoption_tree.ML new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/Logger/Data_Structures/hoption_tree.ML @@ -0,0 +1,116 @@ +(* Title: Logger/Data_Structures/hoption_tree.ML + Author: Kevin Kappelmann + +Hierarchical trees with optional values. +*) +signature HOPTION_TREE = +sig + type letter + type word = letter list + type 'a hoption_tree + + val empty : 'a hoption_tree + val is_empty : 'a hoption_tree -> bool + + val lookup: 'a hoption_tree -> word -> 'a option + + exception INSERT + (*raises INSERT if the entry is already set*) + val insert : (word * 'a) -> 'a hoption_tree -> 'a hoption_tree + val insert_safe : (word * 'a) -> 'a hoption_tree -> 'a hoption_tree + + exception DELETE + (*deletes entry; raises DELETE if the corresponding entry is not set*) + val delete : word -> 'a hoption_tree -> 'a hoption_tree + val delete_safe : word -> 'a hoption_tree -> 'a hoption_tree + (*like delete but also cuts off all subtrees*) + val cut : word -> 'a hoption_tree -> 'a hoption_tree + (*does not cut subtrees if entry is not set*) + val cut_safe : word -> 'a hoption_tree -> 'a hoption_tree + + (*maps entry*) + val map : word -> ('a option -> 'a option) -> 'a hoption_tree -> 'a hoption_tree + (*maps entry and all entries in subtrees*) + val map_below : word -> ('a option -> 'a option) -> 'a hoption_tree -> 'a hoption_tree + + val join : ('a option * 'a option -> 'a option) -> 'a hoption_tree * 'a hoption_tree -> + 'a hoption_tree +end + +functor HOption_Tree(M : MAP) : HOPTION_TREE = +struct + +type letter = M.key +type word = letter list + +datatype 'a hoption_tree = HOption_Tree of 'a option * ('a hoption_tree M.map) + +fun get_data (HOption_Tree (data, _)) = data +fun get_map (HOption_Tree (_, m)) = m + +val empty = HOption_Tree (NONE, M.empty) +fun has_empty_data hoptt = is_none (get_data hoptt) +fun has_empty_map hoptt = M.is_empty (get_map hoptt) +fun is_empty hoptt = has_empty_data hoptt andalso has_empty_map hoptt + +fun hoptt_map f (HOption_Tree p) = HOption_Tree (f p) +fun map_data f = hoptt_map (apfst f) +fun map_map f = hoptt_map (apsnd f) + + +exception INSERT +fun insert_data data = + map_data (fn old_data => if is_some old_data then raise INSERT else SOME data) + +exception DELETE +fun delete_data hoptt = + map_data (fn old_data => if is_none old_data then raise DELETE else NONE) hoptt + +fun lookup hoptt [] = get_data hoptt + | lookup hoptt (l :: ls) = + M.lookup (get_map hoptt) l + |> Option.mapPartial (fn hoptt => lookup hoptt ls) + +fun insert ([], data) = insert_data data + | insert (l :: ls, data) = + map_map (M.map_default (l, empty) (insert (ls, data))) + +fun insert_safe args hoptt = insert args hoptt handle INSERT => hoptt + +local +exception EMPTY +fun check_empty hoptt = if is_empty hoptt then raise EMPTY else hoptt +fun del [] cut hoptt = + delete_data hoptt |> cut ? (map_map (K M.empty)) |> check_empty + | del (l :: ls) cut hoptt = + let val hoptt = map_map (M.map_entry l (del ls cut)) hoptt + handle M.SAME => raise DELETE + | EMPTY => map_map (M.delete l) hoptt + in check_empty hoptt end +in + +fun delete word hoptt = del word false hoptt handle EMPTY => empty +fun cut word hoptt = del word true hoptt handle EMPTY => empty + +end + +fun delete_safe p hoptt = delete p hoptt handle DELETE => hoptt +fun cut_safe p hoptt = cut p hoptt handle DELETE => hoptt + +local +fun map_aux [] map_below f hoptt = + map_data f hoptt |> map_below ? (map_map (M.map (map_aux [] true f |> K))) + | map_aux (l :: ls) map_below f hoptt = + map_map (M.map_entry l (map_aux ls map_below f)) hoptt + handle M.SAME => hoptt +in + +fun map word = map_aux word false +fun map_below word = map_aux word true + +end + +fun join f (HOption_Tree (data1, m1), HOption_Tree (data2, m2)) = + HOption_Tree (f (data1, data2), M.join (join f |> K) (m1, m2)) + +end \ No newline at end of file diff --git a/thys/ML_Unification/Logger/Data_Structures/map.ML b/thys/ML_Unification/Logger/Data_Structures/map.ML new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/Logger/Data_Structures/map.ML @@ -0,0 +1,47 @@ +(* Title: Logger/Data_Structures/map.ML + Author: Kevin Kappelmann + +Key-Value Maps +*) +signature MAP = +sig + type key + type 'a map + + val empty : 'a map + val is_empty : 'a map -> bool + + val lookup: 'a map -> key -> 'a option + + val map : (key -> 'a -> 'b) -> 'a map -> 'b map + + exception SAME + (*raises SAME if entry is not found*) + val map_entry : key -> ('a -> 'a) -> 'a map -> 'a map + val map_default : key * 'a -> ('a -> 'a) -> 'a map -> 'a map + + val delete : key -> 'a map -> 'a map + val join : (key -> 'a * 'a -> 'a) -> ('a map * 'a map) -> 'a map +end + +functor TableMap(T : TABLE) : MAP = +struct + +type key = T.key +type 'a map = 'a T.table + +val empty = T.empty +val is_empty = T.is_empty + +val lookup = T.lookup + +val map = T.map + +exception SAME = T.SAME +val map_entry = T.map_entry +val map_default = T.map_default + +val delete = T.delete +val join = T.join + +end \ No newline at end of file diff --git a/thys/ML_Unification/Logger/ML_Logger.thy b/thys/ML_Unification/Logger/ML_Logger.thy new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/Logger/ML_Logger.thy @@ -0,0 +1,19 @@ +\<^marker>\creator "Kevin Kappelmann"\ +section \ML Logger\ +theory ML_Logger + imports + ML_Attributes +begin + +paragraph \Summary\ +text \Generic logging, at some places inspired by Apache's Log4J 2 +\<^url>\https://logging.apache.org/log4j/2.x/manual/customloglevels.html\.\ + +ML_file\Data_Structures/map.ML\ +ML_file\Data_Structures/hoption_tree.ML\ +ML_file\Data_Structures/binding_tree.ML\ + +ML_file\logger.ML\ +ML_file\logging_antiquotation.ML\ + +end diff --git a/thys/ML_Unification/Logger/ML_Logger_Examples.thy b/thys/ML_Unification/Logger/ML_Logger_Examples.thy new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/Logger/ML_Logger_Examples.thy @@ -0,0 +1,111 @@ +\<^marker>\creator "Kevin Kappelmann"\ +subsection \Examples\ +theory ML_Logger_Examples + imports + ML_Logger + Setup_Result_Commands +begin + +text \First some simple, barebone logging: print some information.\ +ML_command\ + \ \the following two are equivalent\ + val _ = Logger.log Logger.root_logger Logger.INFO @{context} (K "hello root logger") + val _ = @{log Logger.INFO Logger.root_logger} @{context} (K "hello root logger") +\ + +ML_command\ + val logger = Logger.root_logger + val _ = @{log} @{context} (K "hello root logger") + \ \\@{log}\ is equivalent to \Logger.log logger Logger.INFO\\ +\ + +text \To guarantee the existence of a "logger" in an ML structure, one should use the +\HAS_LOGGER\ signature.\ +ML\ + structure My_Struct : sig + include HAS_LOGGER + val get_n : Proof.context -> int + end = struct + val logger = Logger.setup_new_logger Logger.root_logger "My_Struct" + fun get_n ctxt = (@{log} ctxt (K "retrieving n..."); 42) + end +\ + +ML_command\val n = My_Struct.get_n @{context}\ + +text\We can set up a hierarchy of loggers\ +ML\ + val logger = Logger.root_logger + val parent1 = Logger.setup_new_logger Logger.root_logger "Parent1" + val child1 = Logger.setup_new_logger parent1 "Child1" + val child2 = Logger.setup_new_logger parent1 "Child2" + + val parent2 = Logger.setup_new_logger Logger.root_logger "Parent2" +\ + +ML_command\ + (@{log Logger.INFO Logger.root_logger} @{context} (K "Hello root logger"); + @{log Logger.INFO parent1} @{context} (K "Hello parent1"); + @{log Logger.INFO child1} @{context} (K "Hello child1"); + @{log Logger.INFO child2} @{context} (K "Hello child2"); + @{log Logger.INFO parent2} @{context} (K "Hello parent2")) +\ + +text \We can use different log levels to show/surpress messages. The log levels are based on +Apache's Log4J 2 \<^url>\https://logging.apache.org/log4j/2.x/manual/customloglevels.html\.\ +ML_command\@{log Logger.DEBUG parent1} @{context} (K "Hello parent1")\ \\prints nothings\ +declare [[ML_map_context \Logger.set_log_level parent1 Logger.DEBUG\]] +ML_command\@{log Logger.DEBUG parent1} @{context} (K "Hello parent1")\ \\prints message\ +ML_command\Logger.ALL\ \\ctrl+click on the value to see all log levels\ + +text \We can set options for all loggers below a given logger. Below, we set the log level for all +loggers below (and including) \<^ML>\parent1\ to error, thus disabling warning messages.\ +ML_command\ + (@{log Logger.WARN parent1} @{context} (K "Warning from parent1"); + @{log Logger.WARN child1} @{context} (K "Warning from child1")) +\ +declare [[ML_map_context \Logger.set_log_levels parent1 Logger.ERR\]] +ML_command\ + (@{log Logger.WARN parent1} @{context} (K "Warning from parent1"); + @{log Logger.WARN child1} @{context} (K "Warning from child1")) +\ +declare [[ML_map_context \Logger.set_log_levels parent1 Logger.INFO\]] + +text \We can set message filters.\ + +declare [[ML_map_context \Logger.set_msg_filters Logger.root_logger (match_string "Third")\]] +ML_command\ + (@{log Logger.INFO parent1} @{context} (K "First message"); + @{log Logger.INFO child1} @{context} (K "Second message"); + @{log Logger.INFO child2} @{context} (K "Third message"); + @{log Logger.INFO parent2} @{context} (K "Fourth message")) +\ +declare [[ML_map_context \Logger.set_msg_filters Logger.root_logger (K true)\]] + +text \One can also use different output channels (e.g. files) and hide/show some additional +logging information. Ctrl+click on below values and explore.\ +ML_command\Logger.set_output; Logger.set_show_logger; Logging_Antiquotation.show_log_pos\ + +text \To set up (local) loggers outside ML environments, +@{theory ML_Unification.Setup_Result_Commands} contains two commands, +@{command setup_result} and @{command local_setup_result}.\ +experiment +begin +local_setup_result local_logger = \Logger.new_logger Logger.root_logger "Local"\ + +ML_command\@{log Logger.INFO local_logger} @{context} (K "Hello local world")\ +end + +text \\local_logger\ is no longer available. The follow thus does not work:\ +\ \ML_command\@{log Logger.INFO local_logger} @{context} (K "Hello local world")\\ + +text \Let us create another logger in the global context.\ +setup_result some_logger = \Logger.new_logger Logger.root_logger "Some_Logger"\ +ML_command\@{log Logger.INFO some_logger} @{context} (K "Hello world")\ + +text \Let us delete it again.\ +declare [[ML_map_context \Logger.delete_logger some_logger\]] +text \The logger can no longer be found in the logger hierarchy\ +ML_command\@{log Logger.INFO some_logger} @{context} (K "Hello world")\ + +end diff --git a/thys/ML_Unification/Logger/README.md b/thys/ML_Unification/Logger/README.md new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/Logger/README.md @@ -0,0 +1,12 @@ +A hierarchical logger for Isabelle/ML. +See `ML_Logger_Examples.thy` for a quick overview. + +Features: +1. Per logger configuration, including, among other things, + 1. output function (e.g. print to console or file) + 2. log level: suppress and enable log messages based on severity + 3. message filter: suppress and enable log messages based on a filter function +2. Hierarchical configuration: set options based on logger name spaces; + for example, disable logging for all loggers registered below `Root.Unification.*` +3. Logging antiquotation to optionally print positional information of logging message +4. Commands and attributes to setup and configure loggers with ML code. diff --git a/thys/ML_Unification/Logger/Setup_Result_Commands.thy b/thys/ML_Unification/Logger/Setup_Result_Commands.thy new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/Logger/Setup_Result_Commands.thy @@ -0,0 +1,32 @@ +\<^marker>\creator "Kevin Kappelmann"\ +subsection \Setup Result Commands\ +theory Setup_Result_Commands + imports Pure + keywords "setup_result" :: thy_decl + and "local_setup_result" :: thy_decl +begin + +paragraph \Summary\ +text \Setup and local setup with result commands\ + +ML\ + let + fun setup_result finish (name, (source, pos)) = + ML_Context.expression pos + (ML_Lex.read "val" @ name @ ML_Lex.read "= Context.>>> (" @ source @ ML_Lex.read ")") + |> finish + val parse = Parse.embedded_ml + -- ((\<^keyword>\=\ || \<^keyword>\\\) + |-- Parse.position Parse.embedded_ml) + in + Outer_Syntax.command \<^command_keyword>\setup_result\ + "ML setup with result for global theory" + (parse >> (Toplevel.theory o setup_result Context.theory_map)); + Outer_Syntax.local_theory \<^command_keyword>\local_setup_result\ + "ML setup with result for local theory" + (parse >> (setup_result + (Local_Theory.declaration {pos = \<^here>, syntax = false, pervasive = false} o K))) + end +\ + +end diff --git a/thys/ML_Unification/Logger/logger.ML b/thys/ML_Unification/Logger/logger.ML new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/Logger/logger.ML @@ -0,0 +1,225 @@ +(* Title: Logger/logger.ML + Author: Kevin Kappelmann, Paul Bachmann + +Hierarchical logger, indexed on bindings. +The log levels are based on Apache's Log4J 2 +https://logging.apache.org/log4j/2.x/manual/customloglevels.html +*) +signature LOGGER = +sig + type log_level = int + val OFF : log_level + val FATAL : log_level + (*error log level*) + val ERR : log_level + val WARN : log_level + val INFO : log_level + val DEBUG : log_level + val TRACE : log_level + val ALL : log_level + + type logger_name = bstring + type logger_binding + + val root_logger : logger_binding + (*takes parent logger and creates a child logger with the given name*) + val create_child_logger : logger_binding -> logger_name -> logger_binding + val pretty_binding : logger_binding -> Pretty.T + + type logger_output = log_level -> string -> unit + val default_output : logger_output + + type msg_filter = string -> bool + + type logger_config + val config : logger_output -> log_level -> msg_filter -> bool -> logger_config + val default_config : logger_config + val set_config_output : logger_output -> logger_config -> logger_config + val set_config_log_level : log_level -> logger_config -> logger_config + (*set message filter: only log messages with positive results*) + val set_config_msg_filter : msg_filter -> logger_config -> logger_config + (*set whether to print the logger's name when logging*) + val set_config_show_logger : bool -> logger_config -> logger_config + + val lookup_logger : logger_binding -> Context.generic -> logger_config option + + val insert_logger : (logger_binding * logger_config) -> + Context.generic -> Context.generic + val insert_logger_safe : (logger_binding * logger_config) -> + Context.generic -> Context.generic + + val delete_logger : logger_binding -> Context.generic -> Context.generic + val cut_loggers : logger_binding -> Context.generic -> Context.generic + + val set_logger : logger_binding -> (logger_config -> logger_config) -> + Context.generic -> Context.generic + val set_loggers : logger_binding -> (logger_config -> logger_config) -> + Context.generic -> Context.generic + val set_output : logger_binding -> logger_output -> Context.generic -> Context.generic + val set_outputs : logger_binding -> logger_output -> Context.generic -> Context.generic + val set_log_level : logger_binding -> log_level -> Context.generic -> Context.generic + val set_log_levels : logger_binding -> log_level -> Context.generic -> Context.generic + val set_msg_filter : logger_binding -> msg_filter -> Context.generic -> Context.generic + val set_msg_filters : logger_binding -> msg_filter -> Context.generic -> Context.generic + val set_show_logger : logger_binding -> bool -> Context.generic -> Context.generic + val set_show_loggers : logger_binding -> bool -> Context.generic -> Context.generic + + (*creates and inserts child logger with default configuration into context*) + val new_logger : logger_binding -> logger_name -> Context.generic -> + (logger_binding * Context.generic) + (*registers new logger to background theory*) + val setup_new_logger : logger_binding -> logger_name -> logger_binding + + (*prints message created by passed function to the logger's output if the + logger's log_level is greater or equal than the passed level and the message + filter answers positively; + uses lazy computation of the message to avoid computations in case the log + level blocks the logging.*) + val log : logger_binding -> log_level -> Proof.context -> (unit -> string) -> unit + + (* logging functions for different log levels *) + val fatal : logger_binding -> Proof.context -> (unit -> string) -> unit + val err : logger_binding -> Proof.context -> (unit -> string) -> unit + val warn : logger_binding -> Proof.context -> (unit -> string) -> unit + val info : logger_binding -> Proof.context -> (unit -> string) -> unit + val debug : logger_binding -> Proof.context -> (unit -> string) -> unit + val trace : logger_binding -> Proof.context -> (unit -> string) -> unit + +end + +structure Logger : LOGGER = +struct + +type log_level = int +(*values for different log levels*) +val OFF = 0 +val FATAL = 100 +val ERR = 200 +val WARN = 300 +val INFO = 400 +val DEBUG = 500 +val TRACE = 600 +val ALL = 1000 + +type logger_name = bstring +datatype logger_binding = Logger_Binding of binding + +fun binding_of (Logger_Binding binding) = binding + +val root_logger_name = "Root" +val root_logger = Binding.name root_logger_name |> Logger_Binding +fun create_child_logger (Logger_Binding parent) child_name = + let val child = Binding.qualify_name true parent child_name + in + if Symbol_Pos.is_identifier (Binding.name_of child) + then Logger_Binding child + else error ("Bad identifier for child logger " ^ ML_Syntax.print_string child_name) + end +val pretty_binding = Binding.pretty o binding_of + +type logger_output = log_level -> string -> unit + +fun default_output log_level = + if log_level <= WARN then warning + else if log_level < DEBUG then writeln + else tracing + +type msg_filter = string -> bool + +type logger_config = { + log_level : log_level, + msg_filter : msg_filter, + output : logger_output, + show_logger : bool + } +fun config output log_level msg_filter show_logger = + {log_level = log_level, output = output, msg_filter = msg_filter, show_logger = show_logger} +val default_config = config default_output INFO (K true) true +fun set_config_log_level log_level {output, show_logger, msg_filter,...} = + {log_level = log_level, output = output, msg_filter = msg_filter, show_logger = show_logger} +fun set_config_output output {log_level, show_logger, msg_filter,...} = + {log_level = log_level, output = output, msg_filter = msg_filter, show_logger = show_logger} +fun set_config_msg_filter msg_filter {output, log_level, show_logger,...} = + {log_level = log_level, output = output, msg_filter = msg_filter, show_logger = show_logger} +fun set_config_show_logger show_logger {output, log_level, msg_filter,...} = + {log_level = log_level, output = output, msg_filter = msg_filter, show_logger = show_logger} + +fun log_config binding {log_level, output, msg_filter, show_logger} level message_f = + if level > log_level then () + else + let val msg = message_f () + |> show_logger ? (fn msg => cat_lines [ + "Logger: " ^ Pretty.string_of (pretty_binding binding), + msg]) + in if msg_filter msg then output level msg else () end + +structure BT = Binding_Tree + +val init_tree = BT.insert (binding_of root_logger, default_config) BT.empty + +structure Logger_Data = Generic_Data( + type T = logger_config BT.binding_tree + val empty = init_tree + val merge = BT.merge +) + +fun lookup_logger (Logger_Binding binding) = + Logger_Data.get #> (fn bt => BT.lookup bt binding) +fun insert_logger bcp context = + (Logger_Data.map (apfst binding_of bcp |> BT.insert) context) + handle (exn as BT.INSERT) => (warning (Pretty.block [ + Pretty.str "Logger ", + pretty_binding (fst bcp), + Pretty.str " already added." + ] |> Pretty.string_of); + Exn.reraise exn) + +fun insert_logger_safe bcp = Logger_Data.map (apfst binding_of bcp |> BT.insert_safe) +fun delete_logger (Logger_Binding binding) = Logger_Data.map (BT.delete_safe binding) +fun cut_loggers (Logger_Binding binding) = Logger_Data.map (BT.cut_safe binding) + +fun set_logger (Logger_Binding binding) f = + Logger_Data.map (BT.map binding (Option.map f)) +fun set_loggers (Logger_Binding binding) f = + Logger_Data.map (BT.map_below binding (Option.map f)) + +fun set_output binding = set_logger binding o set_config_output +fun set_outputs binding = set_loggers binding o set_config_output +fun set_log_level binding = set_logger binding o set_config_log_level +fun set_log_levels binding = set_loggers binding o set_config_log_level +fun set_msg_filter binding = set_logger binding o set_config_msg_filter +fun set_msg_filters binding = set_loggers binding o set_config_msg_filter +fun set_show_logger binding = set_logger binding o set_config_show_logger +fun set_show_loggers binding = set_loggers binding o set_config_show_logger + +fun new_logger parent child_name context = + let val child = create_child_logger parent child_name + in (child, insert_logger (child, default_config) context) end +fun setup_new_logger parent child_name = Context.>>> (new_logger parent child_name) + +fun log binding log_level ctxt message_f = + case lookup_logger binding (Context.Proof ctxt) of + SOME config => log_config binding config log_level message_f + | NONE => + let fun warn_msg _ = + "Logger " ^ Pretty.string_of (pretty_binding binding) ^ " not found." + in + if binding = root_logger + then default_output WARN (warn_msg ()) + else log root_logger WARN ctxt warn_msg + end + +fun fatal binding = log binding FATAL +fun err binding = log binding ERR +fun warn binding = log binding WARN +fun info binding = log binding INFO +fun debug binding = log binding DEBUG +fun trace binding = log binding TRACE + +end + +(*structures that use a logger should implement this signature*) +signature HAS_LOGGER = +sig + val logger : Logger.logger_binding +end diff --git a/thys/ML_Unification/Logger/logging_antiquotation.ML b/thys/ML_Unification/Logger/logging_antiquotation.ML new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/Logger/logging_antiquotation.ML @@ -0,0 +1,54 @@ +(* Title: Logger/logging_antiquotation.ML + Author: Kevin Kappelmann +*) +signature LOGGING_ANTIQUOTATION = +sig + val show_log_pos: bool Config.T +end + +structure Logging_Antiquotation : LOGGING_ANTIQUOTATION = +struct + +structure Util = ML_Code_Util + +val show_log_pos = Attrib.setup_config_bool @{binding "show_log_pos"} (K false) + +val log = + let + fun body ts ((log_level, logger_binding), show_log_pos_opt) = + let + val (_, pos) = Token.name_of_src ts + val ctxt_internal = Util.internal_name "ctxt" + val show_log_pos = the_default + (Util.spaces ["Config.get", ctxt_internal, " Logging_Antiquotation.show_log_pos"]) + show_log_pos_opt + val additional_info = Util.spaces [ + "if", ML_Syntax.atomic show_log_pos, + "then", + ML_Syntax.print_string "\n", "^", + ML_Syntax.atomic + ("Position.here " ^ ML_Syntax.atomic (ML_Syntax.print_position pos)), + "else", ML_Syntax.print_string "" + ] |> ML_Syntax.atomic + val message_f_internal = Util.internal_name "message_f" + val code = Util.spaces [ + "fn", ctxt_internal, "=> fn", message_f_internal, "=> Logger.log", + ML_Syntax.atomic logger_binding, + ML_Syntax.atomic log_level, + ctxt_internal, + ML_Syntax.atomic (Util.spaces ["fn _ =>", message_f_internal, "() ^", additional_info]) + ] |> ML_Syntax.atomic + in pair (K ("", code)) end + in + ML_Antiquotation.declaration @{binding "log"} + (Scan.optional Parse.embedded "Logger.INFO" + -- Scan.optional Parse.embedded "logger" (*works in particular if structure implements HAS_LOGGER*) + -- Scan.option Parse.embedded + |> Scan.lift) + body + end + +(*setup the antiquotation*) +val _ = Theory.setup log + +end \ No newline at end of file diff --git a/thys/ML_Unification/ML_Unification_Base.thy b/thys/ML_Unification/ML_Unification_Base.thy new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/ML_Unification_Base.thy @@ -0,0 +1,16 @@ +\<^marker>\creator "Kevin Kappelmann"\ +section \ML Unification Basics\ +theory ML_Unification_Base + imports + ML_Logger + ML_Binders + ML_Normalisations +begin + +paragraph \Summary\ +text \Basic definitions and utilities for unification algorithms.\ + +ML_file\unification_base.ML\ +ML_file\unification_util.ML\ + +end diff --git a/thys/ML_Unification/ML_Unification_HOL_Setup.thy b/thys/ML_Unification/ML_Unification_HOL_Setup.thy new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/ML_Unification_HOL_Setup.thy @@ -0,0 +1,32 @@ +\<^marker>\creator "Kevin Kappelmann"\ +section \Setup for HOL\ +theory ML_Unification_HOL_Setup + imports + HOL.HOL + ML_Unification_Hints +begin + +lemma eq_eq_True: "P \ (P \ Trueprop True)" by standard+ simp_all +declare [[uhint where hint_preprocessor = \Unification_Hints_Base.obj_logic_hint_preprocessor + @{thm atomize_eq[symmetric]} (Conv.rewr_conv @{thm eq_eq_True})\]] + +lemma eq_TrueI: "PROP P \ PROP P \ Trueprop True" by (standard) simp +declare [[ucombine add = \Standard_Unification_Combine.eunif_data + (Simplifier_Unification.SIMPS_TO_unify @{thm eq_TrueI} + |> Unification_Combinator.norm_closed_unifier + (#norm_term Standard_Mixed_Unification.norms_first_higherp_first_comb_higher_unify) + |> Unification_Combinator.unifier_from_closed_unifier + |> K) + (Standard_Unification_Combine.default_metadata \<^binding>\SIMPS_TO_unif\)\]] + +declare [[ucombine add = \Standard_Unification_Combine.eunif_data + (Simplifier_Unification.SIMPS_TO_UNIF_unify @{thm eq_TrueI} + Standard_Mixed_Unification.norms_first_higherp_first_comb_higher_unify + (Standard_Mixed_Unification.first_higherp_first_comb_higher_unify + |> Unification_Combinator.norm_unifier Envir_Normalisation.beta_norm_term_unif) + |> Unification_Combinator.norm_unifier + (#norm_term Standard_Mixed_Unification.norms_first_higherp_first_comb_higher_unify) + |> K) + (Standard_Unification_Combine.default_metadata \<^binding>\SIMPS_TO_UNIF_unif\)\]] + +end diff --git a/thys/ML_Unification/ML_Utils/Attributes/ML_Attribute_Utils.thy b/thys/ML_Unification/ML_Utils/Attributes/ML_Attribute_Utils.thy new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/ML_Utils/Attributes/ML_Attribute_Utils.thy @@ -0,0 +1,13 @@ +\<^marker>\creator "Kevin Kappelmann"\ +section \ML Attribute Utils\ +theory ML_Attribute_Utils + imports + Pure +begin + +paragraph \Summary\ +text \Utilities for attributes.\ + +ML_file\attribute_util.ML\ + +end \ No newline at end of file diff --git a/thys/ML_Unification/ML_Utils/Attributes/attribute_util.ML b/thys/ML_Unification/ML_Utils/Attributes/attribute_util.ML new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/ML_Utils/Attributes/attribute_util.ML @@ -0,0 +1,22 @@ +(* Title: ML_Utils/attribute_util.ML + Author: Kevin Kappelmann + +Utilities for attributes. +*) +signature ML_ATTRIBUTE_UTIL = +sig + val apply_attribute : attribute -> Context.generic * thm -> Context.generic * thm + + val attribute_map_context : attribute -> Context.generic -> Context.generic + val attribute_map_ctxt : attribute -> Proof.context -> Proof.context +end + +structure ML_Attribute_Util : ML_ATTRIBUTE_UTIL = +struct + +fun apply_attribute attr (context, thm) = swap (Thm.apply_attribute attr thm context) + +fun attribute_map_context attr = Thm.attribute_declaration attr Drule.dummy_thm +val attribute_map_ctxt = Context.proof_map o attribute_map_context + +end \ No newline at end of file diff --git a/thys/ML_Unification/ML_Utils/Conversions/ML_Conversion_Utils.thy b/thys/ML_Unification/ML_Utils/Conversions/ML_Conversion_Utils.thy new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/ML_Utils/Conversions/ML_Conversion_Utils.thy @@ -0,0 +1,16 @@ +\<^marker>\creator "Kevin Kappelmann"\ +section \ML Conversion Utils\ +theory ML_Conversion_Utils + imports + Pure +begin + +paragraph \Summary\ +text \Utilities for conversions.\ + + +lemma meta_eq_symmetric: "(A \ B) \ (B \ A)" + by (rule equal_intr_rule) simp_all +ML_file\conversion_util.ML\ + +end \ No newline at end of file diff --git a/thys/ML_Unification/ML_Utils/Conversions/conversion_util.ML b/thys/ML_Unification/ML_Utils/Conversions/conversion_util.ML new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/ML_Utils/Conversions/conversion_util.ML @@ -0,0 +1,53 @@ +(* Title: ML_Utils/conversion_util.ML + Author: Kevin Kappelmann + +Conversion utilities. +*) +signature CONVERSION_UTIL = +sig + val move_prem_to_front_conv : int -> conv + val move_prems_to_front_conv : int list -> conv + val thm_conv : conv -> thm -> thm + val binop_conv : conv -> conv -> conv + val lhs_conv : conv -> conv + val rhs_conv : conv -> conv + val repeat_forall_conv : (cterm list -> Proof.context -> conv) -> Proof.context -> conv + val imp_conv : conv -> conv -> conv + val symmetric_conv : conv +end + +structure Conversion_Util : CONVERSION_UTIL = +struct + +fun move_prem_to_front_conv 1 = Conv.all_conv + | move_prem_to_front_conv i = Conv.implies_concl_conv (move_prem_to_front_conv (i - 1)) + then_conv Conv.rewr_conv Drule.swap_prems_eq + +val move_prems_to_front_conv = sort (rev_order o int_ord) + #> map_index (op +) + #> map move_prem_to_front_conv + #> Conv.every_conv + +val thm_conv = Conv.fconv_rule + +fun binop_conv conv_lhs conv_rhs = + Conv.combination_conv (Conv.arg_conv conv_lhs) conv_rhs + +val lhs_conv = Conv.arg1_conv +val rhs_conv = Conv.arg_conv + +fun repeat_forall_conv cv = + let + fun conv fs ctxt = + Conv.forall_conv (fn (f, ctxt) => conv (f :: fs) ctxt) ctxt + else_conv (cv fs ctxt) + in conv [] end + +fun imp_conv cvprem cvconcl ct = (case try Thm.dest_implies ct of + NONE => cvconcl ct + | SOME (A, C) => Drule.imp_cong_rule (cvprem A) (imp_conv cvprem cvconcl C)) + +fun symmetric_conv ceq = let val (cl, cr) = Thm.dest_equals ceq + in Thm.instantiate' [SOME (Thm.ctyp_of_cterm cl)] [SOME cl, SOME cr] @{thm meta_eq_symmetric} end + +end diff --git a/thys/ML_Unification/ML_Utils/Functor_Instances/ML_Functor_Instances.thy b/thys/ML_Unification/ML_Utils/Functor_Instances/ML_Functor_Instances.thy new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/ML_Utils/Functor_Instances/ML_Functor_Instances.thy @@ -0,0 +1,37 @@ +\<^marker>\creator "Kevin Kappelmann"\ +section \ML Functor Instances\ +theory ML_Functor_Instances + imports + ML_Parsing_Utils +begin + +paragraph \Summary\ +text \Utilities for ML functors that create context data.\ + +ML_file\functor_instance.ML\ +ML_file\functor_instance_antiquot.ML\ + +paragraph \Example\ + +ML_command\ + \\some arbitrary functor\ + functor My_Functor(A : sig + structure FIA : FUNCTOR_INSTANCE_ARGS + val n : int + end) = + struct + fun get_n () = (Pretty.writeln (Pretty.block + [Pretty.str "retrieving n from ", Pretty.str A.FIA.full_name]); + A.n) + end + + \\create an instance (structure) called \Test_Functor_Instance\\ + @{functor_instance struct_name = Test_Functor_Instance + and functor_name = My_Functor + and id = \"test"\ + and more_args = \val n = 42\} + + val _ = Test_Functor_Instance.get_n () +\ + +end diff --git a/thys/ML_Unification/ML_Utils/Functor_Instances/functor_instance.ML b/thys/ML_Unification/ML_Utils/Functor_Instances/functor_instance.ML new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/ML_Utils/Functor_Instances/functor_instance.ML @@ -0,0 +1,47 @@ +(* Title: ML_Utils/functor_instance.ML + Author: Kevin Kappelmann + +Functor instance arguments and utilities. +*) +signature FUNCTOR_INSTANCE_ARGS = +sig + (*full name of the functor instance*) + val full_name : string + (*id for the functor instance*) + val id : bstring + (*position where the functor instance is created*) + val pos : Position.T +end + +signature FUNCTOR_INSTANCE_UTIL = +sig + structure FIA : FUNCTOR_INSTANCE_ARGS + + (*create code that calls the given function of the functor instance*) + val struct_op : string -> string + val code_struct_op : string -> ML_Code_Util.code + (*accessor of the functor instance, e.g. "My_Instance."*) + val accessor : string + + val id_prefix : bstring + + val add_id_prefix : string -> bstring + val mk_binding_id_prefix : bstring -> binding +end + +functor Functor_Instance_Util(FIA : FUNCTOR_INSTANCE_ARGS) : FUNCTOR_INSTANCE_UTIL = +struct + +val _ = @{assert} (FIA.full_name <> "") + +structure FIA = FIA +val struct_op = ML_Syntax_Util.mk_struct_access FIA.full_name +val code_struct_op = ML_Code_Util.read o struct_op +val accessor = struct_op "" + +val id_prefix = (FIA.id <> "" ? suffix "_") FIA.id + +val add_id_prefix = prefix id_prefix +fun mk_binding_id_prefix binding = Binding.make (add_id_prefix binding, FIA.pos) + +end \ No newline at end of file diff --git a/thys/ML_Unification/ML_Utils/Functor_Instances/functor_instance_antiquot.ML b/thys/ML_Unification/ML_Utils/Functor_Instances/functor_instance_antiquot.ML new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/ML_Utils/Functor_Instances/functor_instance_antiquot.ML @@ -0,0 +1,58 @@ +(* Title: ML_Utils/functor_instance_antiquot.ML + Author: Kevin Kappelmann + +Antiquotation to create functor instances with FUNCTOR_INSTANCE_ARGS (FIA) arguments. +*) +signature FUNCTOR_INSTANCE_ANTIQUOT = +sig +end + +structure Functor_Instance_Antiquot : FUNCTOR_INSTANCE_ANTIQUOT = +struct + +@{parse_entries (struct) PA + [accessor, struct_name, struct_sig, functor_name, FIA_struct_name, id, more_args]} + +structure PU = Parse_Util +structure MU = ML_Syntax_Util + +val parse_arg_entries = + let + val required_keys = map PA.key [PA.struct_name, PA.functor_name, PA.id] + val parse_nonempty_name = PU.nonempty_name o K o suffix " must not be empty" + val parse_value = PA.parse_entry + Parse.embedded + (parse_nonempty_name "struct_name") + Parse.embedded + (parse_nonempty_name "functor_name") + (parse_nonempty_name "FIA_struct_name") + Parse.embedded + Parse.embedded + val parse_entry = Parse_Key_Value.parse_entry PA.parse_key PU.eq parse_value + val default_entries = PA.empty_entries () + |> fold PA.set_entry [PA.FIA_struct_name "FIA", PA.more_args "", PA.accessor (quote "")] + in + PA.parse_entries_required Parse.and_list1 required_keys parse_entry default_entries + end + +fun mk_functor_instance (entries, pos) = + let val functor_arg = MU.mk_struct [ + MU.mk_structure (PA.get_FIA_struct_name entries) (SOME "FUNCTOR_INSTANCE_ARGS") (MU.mk_struct [ + MU.mk_val_struct "full_name" + (MU.spaces [PA.get_accessor entries, "^", PA.get_struct_name entries |> quote]), + MU.mk_val_struct "id" (PA.get_id entries), + MU.mk_val_struct "pos" (ML_Syntax.print_position pos) + ]), + PA.get_more_args entries + ] + in + MU.mk_structure (PA.get_struct_name entries) (PA.get_struct_sig_safe entries) + (MU.mk_app [PA.get_functor_name entries, MU.atomic functor_arg]) + end + +val functor_instance = ML_Antiquotation.inline @{binding "functor_instance"} + ((Parse.!!! parse_arg_entries |> PU.position |> Scan.lift) >> mk_functor_instance) + +val _ = Theory.setup functor_instance + +end diff --git a/thys/ML_Unification/ML_Utils/General/ML_General_Utils.thy b/thys/ML_Unification/ML_Utils/General/ML_General_Utils.thy new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/ML_Utils/General/ML_General_Utils.thy @@ -0,0 +1,12 @@ +\<^marker>\creator "Kevin Kappelmann"\ +section \General ML Utils\ +theory ML_General_Utils + imports Pure +begin + +paragraph \Summary\ +text \General ML utilities.\ + +ML_file\general_util.ML\ + +end diff --git a/thys/ML_Unification/ML_Utils/General/general_util.ML b/thys/ML_Unification/ML_Utils/General/general_util.ML new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/ML_Utils/General/general_util.ML @@ -0,0 +1,67 @@ +(* Title: ML_Utils/general_util.ML + Author: Kevin Kappelmann + +General ML utilities. +*) +signature GENERAL_UTIL = +sig + val flip : ('a -> 'b -> 'c) -> 'b -> 'a -> 'c + + val find_first_index : ('a -> bool) -> 'a list -> (int * 'a) option + + (*returns false if function throws an exception*) + val try_bool : ('a -> bool) -> 'a -> bool + + (*lists*) + val fold_rev_index: (int * 'a -> 'b -> 'b) -> 'a list -> 'b -> 'b + + (* sequences *) + (*raises exception if sequence is empty and returns the sequence otherwise*) + val seq_try : exn -> 'a Seq.seq -> 'a Seq.seq + val seq_is_empty : 'a Seq.seq -> (bool * 'a Seq.seq) + val seq_merge : 'a ord -> ('a Seq.seq * 'a Seq.seq) -> 'a Seq.seq + +end + +structure General_Util : GENERAL_UTIL = +struct + +fun flip f x y = f y x + +fun find_first_index p = get_index (Option.filter p) + +fun try_bool f = try f #> (Option.getOpt o rpair false) + +(*lists*) +fun insert _ x [] = [x] + | insert ord x (y :: ys) = case ord (x, y) of + LESS => x :: y :: ys + | EQUAL => x :: y :: ys + | GREATER => y :: insert ord x ys + +fun fold_rev_index f = + let fun fold_aux _ [] y = y + | fold_aux i (x :: xs) y = f (i, x) (fold_aux (i + 1) xs y) + in fold_aux 0 end + +(* sequences *) +fun seq_try exn sq = case Seq.pull sq of + NONE => raise exn + | SOME v => Seq.make (fn () => SOME v) + +fun seq_is_empty sq = case Seq.pull sq of + NONE => (true, Seq.empty) + | SOME v => (false, Seq.make (fn () => SOME v)) + +fun seq_merge ord (xq, yq) = + Seq.make (fn () => (case Seq.pull xq of + NONE => Seq.pull yq + | SOME (x, xq') => case Seq.pull yq of + NONE => Seq.pull xq + | SOME (y, yq') => case ord (x, y) of + LESS => SOME (x, seq_merge ord (xq', Seq.cons y yq')) + | EQUAL => SOME (x, seq_merge ord (xq', Seq.cons y yq')) + | GREATER => SOME (y, seq_merge ord (Seq.cons x xq', yq'))) + ) + +end \ No newline at end of file diff --git a/thys/ML_Unification/ML_Utils/Generic_Data/ML_Generic_Data_Utils.thy b/thys/ML_Unification/ML_Utils/Generic_Data/ML_Generic_Data_Utils.thy new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/ML_Utils/Generic_Data/ML_Generic_Data_Utils.thy @@ -0,0 +1,12 @@ +\<^marker>\creator "Kevin Kappelmann"\ +section \ML Generic Data Utils\ +theory ML_Generic_Data_Utils + imports Pure +begin + +paragraph \Summary\ +text \Utilities for @{ML_functor Generic_Data}.\ + +ML_file\pair_generic_data_args.ML\ + +end \ No newline at end of file diff --git a/thys/ML_Unification/ML_Utils/Generic_Data/pair_generic_data_args.ML b/thys/ML_Unification/ML_Utils/Generic_Data/pair_generic_data_args.ML new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/ML_Utils/Generic_Data/pair_generic_data_args.ML @@ -0,0 +1,16 @@ +(* Title: ML_Utils/pair_generic_data_args.ML + Author: Kevin Kappelmann + +Generic data arguments from pairs of generic data arguments. +*) +functor Pair_Generic_Data_Args(P : sig + structure Data1 : GENERIC_DATA_ARGS + structure Data2 : GENERIC_DATA_ARGS + end) : GENERIC_DATA_ARGS = +struct + +type T = P.Data1.T * P.Data2.T +val empty = (P.Data1.empty, P.Data2.empty) +fun merge ((d1, d2), (d1', d2')) = (P.Data1.merge (d1, d1'), P.Data2.merge (d2, d2')) + +end diff --git a/thys/ML_Unification/ML_Utils/ML_Attributes/ML_Attributes.thy b/thys/ML_Unification/ML_Utils/ML_Attributes/ML_Attributes.thy new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/ML_Utils/ML_Attributes/ML_Attributes.thy @@ -0,0 +1,12 @@ +\<^marker>\creator "Kevin Kappelmann"\ +section \ML Attributes\ +theory ML_Attributes + imports ML_Code_Utils +begin + +paragraph \Summary\ +text \ML code as attributes.\ + +ML_file\ml_attribute.ML\ + +end diff --git a/thys/ML_Unification/ML_Utils/ML_Attributes/ml_attribute.ML b/thys/ML_Unification/ML_Utils/ML_Attributes/ml_attribute.ML new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/ML_Utils/ML_Attributes/ml_attribute.ML @@ -0,0 +1,90 @@ +(* Title: ML_Utils/ml_attribute.ML + Author: Kevin Kappelmann + +Attributes from ML code. +*) +signature ML_ATTRIBUTE = +sig + val get_thm : Context.generic -> thm + val put_thm : thm -> Context.generic -> Context.generic + + val run_attribute : ML_Code_Util.code * Position.T -> attribute + val run_mixed_attribute : ML_Code_Util.code * Position.T -> attribute + val run_declaration_attribute : ML_Code_Util.code * Position.T -> attribute + val run_map_context : ML_Code_Util.code * Position.T -> attribute + val run_rule_attribute : ML_Code_Util.code * Position.T -> attribute + val run_Krule_attribute : ML_Code_Util.code * Position.T -> attribute +end + +structure ML_Attribute : ML_ATTRIBUTE = +struct + +structure U = ML_Code_Util + +structure Thm_Data = Generic_Data( + type T = thm + val empty = Drule.dummy_thm + val merge = fst +) + +val get_thm = Thm_Data.get +val put_thm = Thm_Data.put + +fun run_attribute (code, pos) (context, thm) = + let + val context = put_thm thm context + val (context_internal, thm_internal) = apply2 (U.read o U.internal_name) ("context", "thm") + val code = + U.read "Context.>> (fn" @ context_internal @ U.read "=>" @ + U.read "let val" @ thm_internal @ U.read "= ML_Attribute.get_thm" @ context_internal @ + U.read "val" @ U.tuple [thm_internal, context_internal] @ U.read "=" @ + U.read "Thm.apply_attribute" @ U.atomic code @ thm_internal @ context_internal @ + U.read "in ML_Attribute.put_thm" @ thm_internal @ context_internal @ U.read "end)" + val context = ML_Context.expression pos code context + val thm = get_thm context + in (SOME context, SOME thm) end + +fun run_mixed_attribute (code, pos) = + let val code = U.read "Thm.mixed_attribute" @ U.atomic code + in run_attribute (code, pos) end + +fun run_declaration_attribute (code, pos) = + let val code = U.read "Thm.declaration_attribute" @ U.atomic code + in run_attribute (code, pos) end + +fun run_map_context (code, pos) (context, _) = + let val code = U.read "Thm.declaration_attribute" @ (U.atomic (U.read "K" @ U.atomic code)) + in run_attribute (code, pos) (context, Drule.dummy_thm) end + +fun run_rule_attribute (code, pos) = + let val code = U.read "Thm.rule_attribute []" @ U.atomic code + in run_attribute (code, pos) end + +fun run_Krule_attribute (code, pos) = + let val code = + U.read "Thm.rule_attribute []" @ (U.atomic (U.read "K" @ U.atomic code)) + in run_attribute (code, pos) end + +val parse_code_pos = Scan.lift (Parse.position U.parse_code) + +val _ = Theory.setup + (Attrib.setup \<^binding>\ML_attr\ + (parse_code_pos >> run_attribute) + "attribute as ML code" #> + Attrib.setup \<^binding>\ML_mattr\ + (parse_code_pos >> run_mixed_attribute) + "mixed attribute as ML code" #> + Attrib.setup \<^binding>\ML_dattr\ + (parse_code_pos >> run_declaration_attribute) + "declaration attribute as ML code" #> + Attrib.setup \<^binding>\ML_map_context\ + (parse_code_pos >> run_map_context) + "context only attribute as ML code" #> + Attrib.setup \<^binding>\ML_rattr\ + (parse_code_pos >> run_rule_attribute) + "rule attribute as ML code" #> + Attrib.setup \<^binding>\ML_Krattr\ + (parse_code_pos >> run_Krule_attribute) + "theorem only rule attribute as ML code") + +end \ No newline at end of file diff --git a/thys/ML_Unification/ML_Utils/ML_Code/ML_Code_Utils.thy b/thys/ML_Unification/ML_Utils/ML_Code/ML_Code_Utils.thy new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/ML_Utils/ML_Code/ML_Code_Utils.thy @@ -0,0 +1,14 @@ +\<^marker>\creator "Kevin Kappelmann"\ +section \ML Code Utils\ +theory ML_Code_Utils + imports Pure +begin + +paragraph \Summary\ +text \Utilities to generate and manipulate (parsed) ML code.\ + +ML_file\ml_code_util.ML\ +ML_file\ml_syntax_util.ML\ + +end + diff --git a/thys/ML_Unification/ML_Utils/ML_Code/ml_code_util.ML b/thys/ML_Unification/ML_Utils/ML_Code/ml_code_util.ML new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/ML_Utils/ML_Code/ml_code_util.ML @@ -0,0 +1,55 @@ +(* Title: ML_Utils/ml_code_util.ML + Author: Kevin Kappelmann + +Utilities to create and manipulate parsed ML code. +*) +signature ML_CODE_UTIL = +sig + (*returns a name variant that ought to avoid name shadowing*) + val internal_name : string -> string + val spaces : string list -> string + + (*type of parsed ML code*) + type code = ML_Lex.token Antiquote.antiquote list + + val parse_code : code parser + + val flat : code list -> code + val read : string -> code + val flat_read : string list -> code + val reads : string list -> code + + val commas : code list -> code + + val enclose : code -> code -> code -> code + val atomic : code -> code + val tuple : code list -> code + val list : code list -> code + val record : code -> (code * code) list -> code + +end + +structure ML_Code_Util : ML_CODE_UTIL = +struct + +val internal_name = prefix "INTERNAL___" +val spaces = space_implode " " + +type code = ML_Lex.token Antiquote.antiquote list + +val parse_code = Parse.embedded_ml + +val flat = flat +val read = ML_Lex.read +val flat_read = String.concat #> read +val reads = map read #> flat +val commas = separate (read ",") #> flat + +fun enclose start close code = flat [start, code, close] +val atomic = enclose (read "(") (read ")") +val tuple = commas #> enclose (read "(") (read ")") +val list = commas #> enclose (read "[") (read "]") +fun record delim = map (fn (v, t) => v @ delim @ t) + #> commas #> enclose (read "{") (read "}") + +end diff --git a/thys/ML_Unification/ML_Utils/ML_Code/ml_syntax_util.ML b/thys/ML_Unification/ML_Utils/ML_Code/ml_syntax_util.ML new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/ML_Utils/ML_Code/ml_syntax_util.ML @@ -0,0 +1,131 @@ +(* Title: ML_Utils/ml_syntax_util.ML + Author: Kevin Kappelmann + +Utilities to create ML syntax as strings. +*) +signature ML_SYNTAX_UTIL = +sig + val atomic : string -> string + val internal_name : string -> string + val spaces : string list -> string + val lines : string list -> string + + val mk_app : string list -> string + val mk_app_atomic : string list -> string + + val mk_struct_access : string -> string -> string + + val mk_type_annot : string -> string -> string + + val mk_val_prefix : string -> string + val mk_val_sig : string -> string -> string + val mk_val_struct : string -> string -> string + + val mk_fun_prefix : string -> string + val mk_fun_case : string -> string -> string -> string + val mk_fun : string -> string -> string -> string + val mk_fun_cases : string -> (int * 'a -> string * string) -> 'a list -> string + val add_fun_case : string -> string -> string -> string -> string + + val mk_fn : string -> string -> string + val mk_fn_atomic : string -> string -> string + + val mk_type_prefix : string -> string + val mk_type_abstract : string -> string + val mk_type : string -> string -> string + + val mk_datatype_prefix : string -> string + val mk_datatype : string -> string -> string + + val mk_fun_type : string list -> string + val mk_fun_type_atomic : string list -> string + val mk_type_app : string -> string -> string + val mk_poly_type : string -> string + val mk_poly_type_index : string -> int -> string + + val mk_type_args : string list -> string + + val mk_constructor : string * string -> string + val mk_constructors : (string * string) list -> string + + val mk_list : string list -> string + val mk_tuple : string list -> string + val mk_record : string -> (string * string) list -> string + + val mk_record_sel : string -> string + + val mk_signature_name : string -> string + + val mk_sig : string list -> string + val mk_signature : string -> string list -> string + + val mk_struct : string list -> string + val mk_structure : string -> string option -> string -> string +end + +structure ML_Syntax_Util : ML_SYNTAX_UTIL = +struct + +val atomic = ML_Syntax.atomic +val internal_name = ML_Code_Util.internal_name +val spaces = ML_Code_Util.spaces +val lines = cat_lines + +val mk_app = spaces +val mk_app_atomic = mk_app #> atomic + +fun mk_struct_access struct_name operation = implode [struct_name, ".", operation] + +fun mk_type_annot v t = spaces [v, ":", t] |> atomic + +val mk_val_prefix = prefix "val " +fun mk_val_sig s1 s2 = spaces [mk_val_prefix s1, ":", s2] +fun mk_val_struct s1 s2 = spaces [mk_val_prefix s1, "=", s2] + +val mk_fun_prefix = prefix "fun " +fun mk_fun_case fun_name args def = spaces [fun_name, args, "=", def] +val mk_fun = mk_fun_prefix ooo mk_fun_case +fun mk_fun_cases fun_name mk_case = + map_index (uncurry (mk_fun_case fun_name) o mk_case) #> space_implode "\n| " #> mk_fun_prefix +fun add_fun_case fun_name args def = suffix ("\n| " ^ mk_fun_case fun_name args def) + +fun mk_fn arg body = spaces ["fn", arg, "=>", body] +val mk_fn_atomic = atomic oo mk_fn + +val mk_type_prefix = prefix "type " +val mk_type_abstract = mk_type_prefix +fun mk_type s1 s2 = spaces [mk_type_prefix s1, "=", s2] + +val mk_datatype_prefix = prefix "datatype " +fun mk_datatype s1 s2 = spaces [mk_datatype_prefix s1, "=", s2] + +val mk_fun_type = space_implode " -> " +val mk_fun_type_atomic = space_implode " -> " #> atomic +fun mk_type_app t = suffix (" " ^ t) o enclose "(" ")" +val mk_poly_type = prefix "'" +fun mk_poly_type_index t = prefix (mk_poly_type t) o string_of_int + +val mk_type_args = space_implode ", " + +fun mk_constructor (n, args) = spaces [n, "of", args] +val mk_constructors = map mk_constructor #> space_implode "\n| " + +val mk_list = ML_Syntax.print_list I +val mk_tuple = space_implode ", " #> enclose "(" ")" +fun mk_record delim = map (fn (v, t) => spaces [v, delim, t]) + #> space_implode ", " #> enclose "{" "}" + +val mk_record_sel = prefix "#" + +val mk_signature_name = String.map Char.toUpper + +fun mk_sig body = lines ("sig" :: body @ ["end"]) + +fun mk_signature sig_name rhs = spaces ["signature", sig_name, "=", mk_sig rhs] + +fun mk_struct body = lines ("struct" :: body @ ["end"]) + +fun mk_structure struct_name opt_sig_name rhs = spaces + ["structure", struct_name, the_default "" (Option.map (prefix ": ") opt_sig_name), "=", rhs] + +end \ No newline at end of file diff --git a/thys/ML_Unification/ML_Utils/ML_Utils.thy b/thys/ML_Unification/ML_Utils/ML_Utils.thy new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/ML_Utils/ML_Utils.thy @@ -0,0 +1,21 @@ +\<^marker>\creator "Kevin Kappelmann"\ +section \ML Utils\ +theory ML_Utils + imports + ML_Attribute_Utils + ML_Conversion_Utils + ML_Functor_Instances + ML_General_Utils + ML_Generic_Data_Utils + ML_Method_Utils + ML_Attributes + ML_Code_Utils + ML_Parsing_Utils + ML_Priorities + ML_Tactic_Utils + ML_Term_Utils + ML_Theorem_Utils +begin + +end + diff --git a/thys/ML_Unification/ML_Utils/Methods/ML_Method_Utils.thy b/thys/ML_Unification/ML_Utils/Methods/ML_Method_Utils.thy new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/ML_Utils/Methods/ML_Method_Utils.thy @@ -0,0 +1,12 @@ +\<^marker>\creator "Kevin Kappelmann"\ +section \ML Method Utils\ +theory ML_Method_Utils + imports Pure +begin + +paragraph \Summary\ +text \Utilities for methods.\ + +ML_file\method_util.ML\ + +end \ No newline at end of file diff --git a/thys/ML_Unification/ML_Utils/Methods/method_util.ML b/thys/ML_Unification/ML_Utils/Methods/method_util.ML new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/ML_Utils/Methods/method_util.ML @@ -0,0 +1,18 @@ +(* Title: ML_Utils/method_util.ML + Author: Kevin Kappelmann + +Method utilities. +*) +signature METHOD_UTIL = +sig + val METHOD_CONTEXT : (thm list -> 'a -> tactic) -> 'a -> Method.method + val METHOD_CONTEXT' : (thm list -> 'a -> int -> tactic) -> 'a -> Method.method +end + +structure Method_Util : METHOD_UTIL = +struct + +fun METHOD_CONTEXT tac x = METHOD (fn thms => tac thms x) +fun METHOD_CONTEXT' tac x = METHOD (fn thms => HEADGOAL (tac thms x)) + +end diff --git a/thys/ML_Unification/ML_Utils/Parsing/ML_Parsing_Utils.thy b/thys/ML_Unification/ML_Utils/Parsing/ML_Parsing_Utils.thy new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/ML_Utils/Parsing/ML_Parsing_Utils.thy @@ -0,0 +1,47 @@ +\<^marker>\creator "Kevin Kappelmann"\ +section \ML Parsing Utils\ +theory ML_Parsing_Utils + imports + ML_Attributes + ML_Attribute_Utils +begin + +paragraph \Summary\ +text \Parsing utilities for ML. We provide an antiquotation that takes a list of +keys and creates a corresponding record with getters and mappers and a parser for corresponding +key-value pairs.\ + +ML_file\parse_util.ML\ + +ML_file\parse_key_value.ML\ +ML_file\parse_key_value_antiquot.ML\ + +(*use the following to test the code generation*) +(* ML_command\ + Parse_Key_Value_Antiquot.mk_all "Test" NONE ["ABC", "DEFG" ] + |> split_lines |> map Pretty.str |> Pretty.fbreaks |> Pretty.block |> Pretty.writeln +\ *) + +paragraph \Example\ + +ML_command\ + \ \Create record type and utility functions.\ + @{parse_entries (struct) Test [ABC, DEFG]} + + val parser = + let + \ \Create the key-value parser.\ + val parse_entry = Parse_Key_Value.parse_entry + Test.parse_key \\parser for keys\ + (Scan.succeed []) \\delimiter parser\ + (Test.parse_entry \\value parser\ + Parse.string \\parser for ABC\ + Parse.int) \\parser for DEFG\ + val required_keys = [Test.key Test.ABC] \\required keys\ + val default_entries = Test.empty_entries () \\default values for entries\ + in Test.parse_entries_required Parse.and_list1 required_keys parse_entry default_entries end + \ \This parses, for example, \ABC = hello and DEFG = 3\ or \DEFG = 3 and ABC = hello\, + but not \DEFG = 3\ since the key "ABC" is missing.\ +\ + +end diff --git a/thys/ML_Unification/ML_Utils/Parsing/parse_key_value.ML b/thys/ML_Unification/ML_Utils/Parsing/parse_key_value.ML new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/ML_Utils/Parsing/parse_key_value.ML @@ -0,0 +1,81 @@ +(* Title: ML_Utils/parse_key_value.ML + Author: Kevin Kappelmann + +Parsing of key-value pairs. +*) +signature PARSE_KEY_VALUE = +sig + type ('k, 'v) entry = ('k * 'v) + type ('k, 'v) entries = ('k, 'v) entry list + + (*takes key parser, delimiter parser, value parser*) + val parse_entry : 'k parser -> 'd parser -> ('k -> 'v parser) -> (('k, 'v) entry) parser + val parse_entry' : 'k context_parser -> 'd context_parser -> ('k -> 'v context_parser) -> + (('k, 'v) entry) context_parser + + val pretty_keys : ('k -> string) -> 'k list -> Pretty.T + val parse_key : string list -> (string -> 'k) -> 'k parser + + val has_required_keys : ('k * ('k, 'v) entry -> bool) -> 'k list -> ('k, 'v) entries -> bool + + val parse_entries_required : ('k * 'k -> bool) -> ('k -> string) -> 'k list -> + (('k, 'v) entries) parser -> (('k, 'v) entries) parser + val parse_entries_required' : ('k * 'k -> bool) -> ('k -> string) -> 'k list -> + (('k, 'v) entries) context_parser -> (('k, 'v) entries) context_parser +end + +structure Parse_Key_Value : PARSE_KEY_VALUE = +struct + +structure Util = Parse_Util + +type ('k, 'v) entry = ('k * 'v) +type ('k, 'v) entries = ('k, 'v) entry list + +fun gen_parse_entry cut parse_key parse_delim parse_value = + parse_key + :-- (fn n => parse_delim + |-- cut (parse_value n)) +fun parse_entry parse_key = gen_parse_entry Parse.!!! parse_key +fun parse_entry' parse_key = gen_parse_entry Util.!!!+ parse_key + +fun pretty_keys key_to_string = map (Pretty.str o key_to_string) #> Pretty.commas #> Pretty.block + +fun parse_key names key_from_string = + let val pretty_keys = pretty_keys I names |> Pretty.string_of + in + Parse.group (K pretty_keys) + (Util.nonempty_name (K "Key must not be empty.") >> key_from_string) + end + +fun eq_key_entry eq_key (rn, (xn, _)) = eq_key (rn, xn) + +fun has_required_keys eq_key_entry rks ks = subset eq_key_entry (rks, ks) + +fun gen_parse_entries_required filter_parser eq_key key_to_string ks parse_entries = + let + val eq_key_entry = eq_key_entry eq_key + fun failure_required es = Pretty.fbreaks [ + Pretty.block [ + Pretty.str "Missing keys: ", + pretty_keys key_to_string (subtract (eq_key_entry o swap) es ks) + ], + Pretty.block [ + Pretty.str "Required keys: ", + pretty_keys key_to_string ks + ] + ] |> Pretty.block |> Pretty.string_of + fun failure_distinct es _ = Pretty.block [ + Pretty.str "All keys must be distinct. Duplicates: ", + duplicates eq_key (map fst es) |> pretty_keys key_to_string + ] |> Pretty.string_of + in + filter_parser + (has_required_keys eq_key_entry ks) (Util.fail failure_required) parse_entries + |> Util.distinct_list (eq_key o apply2 fst) failure_distinct + end + +fun parse_entries_required eq_key = gen_parse_entries_required Util.filter_cut eq_key +fun parse_entries_required' eq_key = gen_parse_entries_required Util.filter_cut' eq_key + +end \ No newline at end of file diff --git a/thys/ML_Unification/ML_Utils/Parsing/parse_key_value_antiquot.ML b/thys/ML_Unification/ML_Utils/Parsing/parse_key_value_antiquot.ML new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/ML_Utils/Parsing/parse_key_value_antiquot.ML @@ -0,0 +1,520 @@ +(* Title: ML_Utils/parse_key_value_antiquot.ML + Author: Kevin Kappelmann + +Antiquotation that creates boilerplate code to parse key-value pairs from a given list of keys. +*) +signature PARSE_KEY_VALUE_ANTIQUOT = +sig + val mk_signature : string -> string list -> string + val mk_structure : string -> string option -> string list -> string + val mk_all : string -> string option -> string list -> string +end + +structure Parse_Key_Value_Antiquot : PARSE_KEY_VALUE_ANTIQUOT = +struct + +structure U = ML_Syntax_Util + +val internal_param = U.internal_name o U.internal_name + +val mk_poly_type_a = U.mk_poly_type_index "a" +fun mk_type_args mk_string = U.mk_type_args o map_index (mk_string o fst) + +fun mk_entry_constructors mk_poly_type_a = + map_index (fn (i, k) => (k, mk_poly_type_a i)) #> U.mk_constructors + +val entry_type_name = "entry" +val mk_entry_type = U.mk_type_app entry_type_name oo mk_type_args + +fun mk_entry_type_sig ks = + U.mk_datatype (mk_entry_type mk_poly_type_a ks) (mk_entry_constructors mk_poly_type_a ks) + +val mk_entry_type_struct = mk_entry_type_sig + +val key_type_name = "key" +val mk_key_type = key_type_name + +val mk_key_type_sig = U.mk_type_abstract mk_key_type + +fun mk_key_type_struct ks = U.mk_type mk_key_type (mk_entry_type (K "unit") ks) + +val key_name = "key" + +fun mk_key_sig ks = U.mk_val_sig key_name + (U.spaces ["(unit ->", mk_entry_type (K "unit") ks, ") ->", mk_key_type]) + +val mk_key_struct = let val kparam = internal_param "k" + in U.mk_fun key_name kparam (U.mk_app [kparam, "()"]) end + +val keys_name = internal_param "keys" +fun mk_keys_struct ks = U.mk_val_struct + keys_name (U.mk_app ["map", key_name, U.mk_list ks]) + +val key_to_string_name = "key_to_string" + +val mk_key_to_string_sig = U.mk_val_sig + key_to_string_name (U.mk_fun_type [mk_key_type, "string"]) + +fun mk_key_to_string_struct_entry k = (U.mk_app_atomic [k, "_"], quote k) + +val mk_key_to_string_struct = U.mk_fun_cases + key_to_string_name (mk_key_to_string_struct_entry o snd) + +val keys_strings_name = internal_param "keys_strings" +val mk_keys_strings_struct = U.mk_val_struct + keys_strings_name (U.mk_app ["map", key_to_string_name, keys_name]) + +val key_from_string_name = internal_param "key_from_string" + +fun mk_key_from_string_struct_entry k = (quote k, U.mk_app [key_name, k]) + +val mk_key_from_string_struct = + let val kparam = internal_param "k" + in + U.mk_fun_cases key_from_string_name (mk_key_from_string_struct_entry o snd) + #> U.add_fun_case key_from_string_name kparam (U.mk_app [ + U.mk_struct_access "Scan" "fail_with", + U.mk_fn_atomic kparam + (U.mk_app ["fn _ => \"key \"", "^", kparam, "^", "\" not registered\""]), + kparam]) + end + +fun mk_istate_type itype jtype atype = + U.mk_type_app (U.mk_struct_access "Parse_Util" "istate") (U.mk_type_args [itype, jtype, atype]) + +val parse_entry_name = "parse_entry" + +fun mk_parse_entry_sig ks = + let + val mk_parser_type = mk_istate_type "'i" "'j" + val parser_types = map_index (mk_parser_type o mk_poly_type_a o fst) ks + |> U.mk_fun_type + in + U.mk_val_sig parse_entry_name (U.mk_fun_type + [parser_types, mk_key_type, mk_parser_type (mk_entry_type mk_poly_type_a ks)]) + end + +fun mk_parse_entry_struct_entry ks (i, k) = + let + fun farg j = if i = j then "f" ^ string_of_int i else "_" + val fargs = U.spaces (map_index (farg o fst) ks) + in (U.spaces [fargs, U.mk_app_atomic [k, "_"]], U.mk_app [farg i, ">>", k]) end + +fun mk_parse_entry_struct ks = + U.mk_fun_cases parse_entry_name (mk_parse_entry_struct_entry ks) ks + +val mk_parser_type = U.mk_type_app (U.mk_struct_access "Token" "parser") + +val parse_key_name = "parse_key" + +val mk_parse_key_sig = U.mk_val_sig parse_key_name (mk_parser_type mk_key_type) + +val Parse_Key_Value_op = U.mk_struct_access "Parse_Key_Value" + +val mk_parse_key_struct = U.mk_val_struct parse_key_name + (U.mk_app [Parse_Key_Value_op "parse_key", keys_strings_name, key_from_string_name]) + +val entries_type_name = "entries" +val mk_entries_type = U.mk_type_app entries_type_name oo mk_type_args + +val mk_entries_type_sig = U.mk_type_abstract o mk_entries_type mk_poly_type_a + +val mk_option_type = U.mk_type_app "option" + +fun mk_entries_type_struct ks = + let val record = map_index (apfst (mk_option_type o mk_poly_type_a) #> swap) ks |> U.mk_record ":" + in U.mk_type (mk_entries_type mk_poly_type_a ks) record end + +val empty_entries_name = "empty_entries" + +fun mk_empty_entries_sig ks = U.mk_val_sig + empty_entries_name (U.mk_fun_type ["unit", mk_entries_type mk_poly_type_a ks]) + +fun mk_empty_entries_struct ks = + let val record = map_index (apfst (K "NONE") #> swap) ks |> U.mk_record "=" + in U.mk_fun empty_entries_name "_" record end + +val set_entry_name = "set_entry" + +fun mk_set_entry_sig ks = + let val entries_type = mk_entries_type mk_poly_type_a ks + in + U.mk_val_sig set_entry_name + (U.mk_fun_type [mk_entry_type mk_poly_type_a ks, entries_type, entries_type]) + end + +fun mk_set_entry_struct_entry ks (i, k) = + let + val prefix_other = U.internal_name + fun record_in_data (j, k) = (k, if i = j then "_" else prefix_other k) + val record_in = map_index record_in_data ks |> U.mk_record "=" + val vparam = internal_param "v" + fun record_out_data (j, k) = (k, if i = j then U.mk_app ["SOME", vparam] else prefix_other k) + val record_out = map_index record_out_data ks |> U.mk_record "=" + in (U.spaces [U.mk_app_atomic [k, vparam], record_in], record_out) end + +fun mk_set_entry_struct ks = + U.mk_fun_cases set_entry_name (mk_set_entry_struct_entry ks) ks + +val merge_entries_name = "merge_entries" + +fun mk_merge_entries_sig ks = + let val entries_type = mk_entries_type mk_poly_type_a ks + in U.mk_val_sig merge_entries_name (U.mk_fun_type [entries_type, entries_type, entries_type]) end + +fun mk_merge_entries_struct ks = + let + fun prefix i = suffix (string_of_int i) o U.internal_name + fun record_in_data i k = (k, prefix i k) + fun record_in i = map (record_in_data i) ks |> U.mk_record "=" + fun record_out_data k = (k, U.mk_app ["merge_options", U.mk_tuple [prefix 1 k, prefix 2 k]]) + val record_out = map record_out_data ks |> U.mk_record "=" + in U.mk_fun merge_entries_name (U.spaces [record_in 1, record_in 2]) record_out end + +val mk_map_name = prefix "map_" +val mk_map_name_safe = suffix "_safe" o mk_map_name + +fun mk_map_type mk_arg_type ks i = U.mk_fun_type [ + U.mk_fun_type_atomic [mk_arg_type (mk_poly_type_a i), U.mk_poly_type_index "b" i], + mk_entries_type mk_poly_type_a ks, + mk_entries_type (fn j => if i = j then U.mk_poly_type_index "b" j else mk_poly_type_a j) ks + ] + +fun mk_map_safe_sig ks (i, k) = U.mk_val_sig (mk_map_name_safe k) (mk_map_type mk_option_type ks i) +fun mk_map_safes_sig ks = map_index (mk_map_safe_sig ks) ks |> U.lines + +fun mk_map_sig ks (i, k) = U.mk_val_sig (mk_map_name k) (mk_map_type I ks i) +fun mk_maps_sig ks = map_index (mk_map_sig ks) ks |> U.lines + +fun mk_map_safe_struct ks (i, k) = + let + val prefix = U.internal_name + fun record_in_data k = (k, prefix k) + val record_in = map record_in_data ks |> U.mk_record "=" + val fparam = internal_param "f" + fun record_out_data (j, k) = (k, if i = j + then U.mk_app ["SOME", U.mk_app_atomic [fparam, prefix k]] + else prefix k) + val record_out = map_index record_out_data ks |> U.mk_record "=" + in U.mk_fun (mk_map_name_safe k) (U.spaces [fparam, record_in]) record_out end + +fun mk_map_safes_struct ks = map_index (mk_map_safe_struct ks) ks |> U.lines + +val the_value_name = internal_param "the_value" +val mk_the_value_struct = + let val [kparam, vparam] = map internal_param ["k", "v"] + in + U.mk_fun the_value_name (U.spaces ["_", U.mk_app_atomic ["SOME", vparam]]) vparam + |> U.add_fun_case the_value_name (U.spaces [kparam, "NONE"]) + (U.spaces ["error (\"No data for key \\\"\" ^", key_to_string_name, kparam, "^ \"\\\".\")"]) + end + +fun mk_map_struct k = + let val fparam = internal_param "f" + in + U.mk_fun (mk_map_name k) fparam (U.mk_app [ + mk_map_name_safe k, + U.mk_app_atomic [fparam, "o", + U.mk_app [the_value_name, U.mk_app_atomic [key_name, k]]]]) + end + +val mk_maps_struct = map mk_map_struct #> U.lines + +val mk_get_name = prefix "get_" +val mk_get_name_safe = suffix "_safe" o mk_get_name + +fun mk_get_type mk_res_type ks i = U.mk_fun_type + [mk_entries_type mk_poly_type_a ks, mk_res_type (mk_poly_type_a i)] + +fun mk_get_safe_sig ks (i, k) = U.mk_val_sig (mk_get_name_safe k) (mk_get_type mk_option_type ks i) +fun mk_get_safes_sig ks = map_index (mk_get_safe_sig ks) ks |> U.lines + +fun mk_get_sig ks (i, k) = U.mk_val_sig (mk_get_name k) (mk_get_type I ks i) +fun mk_gets_sig ks = map_index (mk_get_sig ks) ks |> U.lines + +fun mk_get_safe_struct ks k = + let + val entries = internal_param "entries" + val arg = (U.mk_type_annot entries (mk_entries_type mk_poly_type_a ks)) + in U.mk_fun (mk_get_name_safe k) arg (U.spaces [U.mk_record_sel k, entries]) end + +fun mk_get_safes_struct ks = map (mk_get_safe_struct ks) ks |> U.lines + +fun mk_get_struct k = + let val entries = internal_param "entries" + in + U.mk_fun (mk_get_name k) entries (U.mk_app [the_value_name, + U.mk_app_atomic [key_name, k], + U.mk_app_atomic [mk_get_name_safe k, entries]]) + end + +val mk_gets_struct = map mk_get_struct #> U.lines + +val mk_list_type = U.mk_type_app "list" +fun mk_parse_key_value_entry_type mk_poly_type ks = U.mk_type_app (Parse_Key_Value_op "entry") + (U.mk_type_args [mk_key_type, mk_entry_type mk_poly_type ks]) +fun mk_parse_key_value_entries_type mk_poly_type ks = U.mk_type_app (Parse_Key_Value_op "entries") + (U.mk_type_args [mk_key_type, mk_entry_type mk_poly_type ks]) + +val key_entry_entry_type_name = "key_entry_entry" +val mk_key_entry_entry_type = + U.mk_type_app key_entry_entry_type_name oo mk_type_args + +fun mk_key_entry_entry_type_sig ks = U.mk_type (mk_key_entry_entry_type mk_poly_type_a ks) + (mk_parse_key_value_entry_type mk_poly_type_a ks) + +val mk_key_entry_entry_type_struct = mk_key_entry_entry_type_sig + +val key_entry_entries_type_name = "key_entry_entries" +val mk_key_entry_entries_type = U.mk_type_app key_entry_entries_type_name oo mk_type_args + +fun mk_key_entry_entries_type_sig ks = U.mk_type (mk_key_entry_entries_type mk_poly_type_a ks) + (mk_parse_key_value_entries_type mk_poly_type_a ks) + +val mk_key_entry_entries_type_struct = mk_key_entry_entries_type_sig + +val key_entry_entries_from_entries_name = "key_entry_entries_from_entries" + +fun mk_key_entry_entries_from_entries_sig ks = U.mk_val_sig key_entry_entries_from_entries_name + (U.mk_fun_type [mk_entries_type mk_poly_type_a ks, mk_key_entry_entries_type mk_poly_type_a ks]) + +fun mk_key_entry_entries_from_entries_struct ks = + let + val prefix_key = U.internal_name + fun record_in_data k = (k, prefix_key k) + val record_in = map record_in_data ks |> U.mk_record "=" + fun list_out_data k = U.mk_app [ + U.mk_struct_access "Option" "map", + U.mk_app_atomic ["pair", U.mk_app_atomic [key_name, k], "o", k], + prefix_key k] + val list_out = map list_out_data ks |> U.mk_list + in + U.mk_fun key_entry_entries_from_entries_name record_in (U.mk_app [list_out, + "|>", "map_filter", "I"]) + end + +val mk_entries_from_entry_list_name = "entries_from_entry_list" + +fun mk_entries_from_entry_list_sig ks = U.mk_val_sig mk_entries_from_entry_list_name + (U.mk_fun_type [mk_list_type (mk_entry_type mk_poly_type_a ks), mk_entries_type mk_poly_type_a ks]) + +val mk_entries_from_entry_list_struct = + let val kees = internal_param "kees" + in + U.mk_fun mk_entries_from_entry_list_name kees (U.mk_app + ["fold", set_entry_name, kees, U.mk_app_atomic [empty_entries_name, "()"]]) + end + +val mk_context_parser_type = U.mk_type_app (U.mk_struct_access "Token" "context_parser") + +local +fun mk_names name = (prefix "gen_" name, name, suffix "'" name) + +fun gen_mk_parse_key_value_result_type mk_result_type mk_type mk_poly_type = + mk_type mk_poly_type #> mk_result_type +val mk_parse_key_value_result_type = gen_mk_parse_key_value_result_type mk_parser_type +val mk_parse_key_value_result_type' = gen_mk_parse_key_value_result_type mk_context_parser_type +in + +val (gen_parse_key_entry_entries_required_name, parse_key_entry_entries_required_name, + parse_key_entry_entries_required_name') = mk_names "parse_key_entry_entries_required" + +val mk_gen_parse_key_entry_entries_required_struct = + let + val (args as [parse_entries_required, parse_list, ks, parse_entry]) = map internal_param + ["parse_entries_required", "parse_list", "ks", "parse_entry"] + in + U.mk_fun gen_parse_key_entry_entries_required_name (U.spaces args) + (U.mk_app [parse_entries_required, "(op =)", key_to_string_name, ks, + U.mk_app_atomic [parse_list, parse_entry]]) + end + +fun gen_mk_parse_entries_required_args_type mk_parse_key_value_result_type + mk_poly_type_in mk_poly_type_out ks = + let + val entry_parser = mk_parse_key_value_result_type mk_key_entry_entry_type mk_poly_type_in ks + val entries_parser = U.mk_fun_type_atomic [entry_parser, + mk_parse_key_value_result_type mk_key_entry_entries_type mk_poly_type_out ks] + in U.mk_fun_type [entries_parser, mk_list_type key_type_name, entry_parser] end + +fun gen_mk_parse_key_entry_entries_required_sig name mk_parse_key_value_result_type ks = + let val (mk_poly_type_in, mk_poly_type_out) = apply2 U.mk_poly_type_index ("a", "b") + in + U.mk_val_sig name (U.mk_fun_type [ + gen_mk_parse_entries_required_args_type mk_parse_key_value_result_type + mk_poly_type_in mk_poly_type_out ks, + mk_parse_key_value_result_type mk_key_entry_entries_type mk_poly_type_out ks]) + end + +fun mk_parse_key_entry_entries_required_sigs ks = U.lines [ + gen_mk_parse_key_entry_entries_required_sig parse_key_entry_entries_required_name + mk_parse_key_value_result_type ks, + gen_mk_parse_key_entry_entries_required_sig parse_key_entry_entries_required_name' + mk_parse_key_value_result_type' ks +] + +fun gen_mk_parse_key_entry_entries_required_struct name parse_entries_required = + let val parse_list = internal_param "parse_list" + in + U.mk_fun name parse_list (U.mk_app + [gen_parse_key_entry_entries_required_name, parse_entries_required, parse_list]) + end + +val mk_parse_key_entry_entries_required_structs = U.lines [ + gen_mk_parse_key_entry_entries_required_struct parse_key_entry_entries_required_name + (Parse_Key_Value_op "parse_entries_required"), + gen_mk_parse_key_entry_entries_required_struct parse_key_entry_entries_required_name' + (Parse_Key_Value_op "parse_entries_required'") +] + +val (gen_parse_entries_required_name, parse_entries_required_name, parse_entries_required_name') = + mk_names "parse_entries_required" + +val mk_gen_parse_entries_required_struct = + let + val (args as [parse_key_entry_entries_required, parse_list, ks, parse_entry, default_entries]) = + map internal_param ["parse_key_entry_entries_required", "parse_list", "ks", "parse_entry", + "default_entries"] + val kees = internal_param "kees" + in + U.mk_fun gen_parse_entries_required_name (U.spaces args) + (U.mk_app [parse_key_entry_entries_required, parse_list, ks, parse_entry, + ">>", U.mk_fn_atomic kees + (U.mk_app ["fold", U.mk_app_atomic [set_entry_name, "o", "snd"], kees, default_entries])]) + end + +fun gen_mk_parse_entries_required_sig name mk_parse_key_value_result_type ks = + let + val (mk_poly_type_in, mk_poly_type_out) = apply2 U.mk_poly_type_index ("a", "b") + val entries_type = mk_entries_type mk_poly_type_out ks + in + U.mk_val_sig name (U.mk_fun_type [ + gen_mk_parse_entries_required_args_type mk_parse_key_value_result_type + mk_poly_type_in mk_poly_type_out ks, + entries_type, + mk_parse_key_value_result_type mk_entries_type mk_poly_type_out ks]) + end + +fun mk_parse_entries_required_sigs ks = U.lines [ + gen_mk_parse_entries_required_sig parse_entries_required_name mk_parse_key_value_result_type ks, + gen_mk_parse_entries_required_sig parse_entries_required_name' mk_parse_key_value_result_type' ks +] +end + +fun gen_mk_parse_entries_required_struct name parse_key_entry_entries_required = + let val parse_list = internal_param "parse_list" + in + U.mk_fun name parse_list (U.mk_app + [gen_parse_entries_required_name, parse_key_entry_entries_required, parse_list]) + end + +val mk_parse_entries_required_structs = U.lines [ + gen_mk_parse_entries_required_struct parse_entries_required_name + parse_key_entry_entries_required_name, + gen_mk_parse_entries_required_struct parse_entries_required_name' + parse_key_entry_entries_required_name' +] + +fun mk_signature sig_name ks = U.mk_signature sig_name [ + mk_entry_type_sig ks, + mk_key_type_sig, + mk_key_sig ks, + mk_key_to_string_sig, + mk_parse_entry_sig ks, + mk_parse_key_sig, + mk_entries_type_sig ks, + mk_empty_entries_sig ks, + mk_set_entry_sig ks, + mk_merge_entries_sig ks, + mk_map_safes_sig ks, + mk_maps_sig ks, + mk_get_safes_sig ks, + mk_gets_sig ks, + mk_key_entry_entry_type_sig ks, + mk_key_entry_entries_type_sig ks, + mk_key_entry_entries_from_entries_sig ks, + mk_entries_from_entry_list_sig ks, + mk_parse_key_entry_entries_required_sigs ks, + mk_parse_entries_required_sigs ks + ] + +fun mk_structure struct_name optsig_name ks = U.mk_struct [ + mk_entry_type_struct ks, + mk_key_type_struct ks, + mk_key_struct, + mk_keys_struct ks, + mk_key_to_string_struct ks, + mk_keys_strings_struct, + mk_key_from_string_struct ks, + mk_parse_entry_struct ks, + mk_parse_key_struct, + mk_entries_type_struct ks, + mk_empty_entries_struct ks, + mk_set_entry_struct ks, + mk_merge_entries_struct ks, + mk_the_value_struct, + mk_map_safes_struct ks, + mk_maps_struct ks, + mk_get_safes_struct ks, + mk_gets_struct ks, + mk_key_entry_entry_type_struct ks, + mk_key_entry_entries_type_struct ks, + mk_key_entry_entries_from_entries_struct ks, + mk_entries_from_entry_list_struct, + mk_gen_parse_key_entry_entries_required_struct, + mk_parse_key_entry_entries_required_structs, + mk_gen_parse_entries_required_struct, + mk_parse_entries_required_structs + ] |> U.mk_structure struct_name optsig_name + +fun mk_all struct_name optsig_name ks = + let val sig_name = the_default (U.mk_signature_name struct_name) optsig_name + in + U.lines [ + mk_signature sig_name ks, + mk_structure struct_name (SOME sig_name) ks + ] + end + +val parse_sig_name = Parse_Util.nonempty_name (K "Signature name must not be empty.") +val parse_struct_name = Parse_Util.nonempty_name (K "Structure name must not be empty.") + +val parse_keys = + let + val parse_key = Parse_Util.nonempty_name (K "Key names must not be empty.") + val key_eq = (op =) + fun dup_msg xs _ = Pretty.block [ + Pretty.str "All keys must be distinct. Duplicates: ", + duplicates key_eq xs |> Parse_Key_Value.pretty_keys I + ] |> Pretty.string_of + in + Args.bracks (Parse.list1 parse_key) + |> Parse_Util.distinct_list key_eq dup_msg + end + +val parse_sig = parse_sig_name -- parse_keys +val parse_struct = parse_struct_name -- Scan.option parse_sig_name -- parse_keys +val parse_all = parse_struct_name -- Scan.option parse_sig_name -- parse_keys + +datatype mode = SIG | STRUCT | ALL + +fun mode_of_string "sig" = SIG + | mode_of_string "struct" = STRUCT + | mode_of_string "all" = ALL + | mode_of_string m = Scan.fail_with (fn m => fn _ => "illegal mode " ^ m) m + +val parse_mode = Parse_Key_Value.parse_key ["sig", "struct", "all"] mode_of_string + +val parse = + let + val parse_mode = Scan.optional (Parse_Util.parenths (Parse.!!! parse_mode)) ALL + fun parse_of_mode SIG = parse_sig >> uncurry mk_signature + | parse_of_mode STRUCT = parse_struct + >> (fn ((struct_name, optsig_name), ks) => mk_structure struct_name optsig_name ks) + | parse_of_mode ALL = parse_all + >> (fn ((struct_name, optsig_name), ks) => mk_all struct_name optsig_name ks) + in parse_mode :|-- parse_of_mode end + +val _ = Theory.setup (ML_Antiquotation.inline @{binding "parse_entries"} (Scan.lift parse)) + +end \ No newline at end of file diff --git a/thys/ML_Unification/ML_Utils/Parsing/parse_util.ML b/thys/ML_Unification/ML_Utils/Parsing/parse_util.ML new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/ML_Utils/Parsing/parse_util.ML @@ -0,0 +1,167 @@ +(* Title: ML_Utils/util.ML + Author: Kevin Kappelmann + +Parsing utilities. +*) +signature PARSE_UTIL = +sig + (*cut for context parsers*) + val !!!+ : 'a context_parser -> 'a context_parser + + val enum1' : string -> 'a context_parser -> 'a list context_parser + val enum' : string -> 'a context_parser -> 'a list context_parser + val and_list1' : 'a context_parser -> 'a list context_parser + val and_list' : 'a context_parser -> 'a list context_parser + + (*indexed state monad (generalising parsers and context parsers)*) + type ('i, 'j, 'a) istate = 'i -> 'a * 'j + + (*"if_ahead p1 f p2" applies f to result of p1 if p1 is successful; otherwise runs p2*) + val if_ahead : ('a, 'b, 'c) istate -> ('c -> 'a -> 'd) -> ('a, 'a, 'd) istate -> + ('a, 'a, 'd) istate + + (*if_ahead applied to end-of-file parser*) + val if_eof : (string -> Token.T list -> 'a) -> 'a parser -> 'a parser + val if_eof' : (string -> Context.generic * Token.T list -> 'a) -> 'a context_parser -> + 'a context_parser + + val option : 'a parser -> ('a option) parser + val option' : 'a context_parser -> ('a option) context_parser + val optional : 'a parser -> 'a -> 'a parser + val optional' : 'a context_parser -> 'a -> 'a context_parser + + val position : 'a parser -> ('a * Position.T) parser + val position' : 'a context_parser -> ('a * Position.T) context_parser + + val fail : ('a -> string) -> 'a -> 'b + (*scanner raising Scan.ABORT with given error message*) + val abort : (Token.T list -> string) -> 'a parser + val abort' : (Context.generic * Token.T list -> string) -> 'a context_parser + + (*"filter f err p" runs parser p to obtain result x and returns x if "f x" holds and "err x" otherwise*) + val filter : ('a -> bool) -> ('a -> ('b, 'b, 'a) istate) -> ('c, 'b, 'a) istate -> + ('c, 'b, 'a) istate + (*filter with cuts in case of failure*) + val filter_cut : ('a -> bool) -> ('a -> Token.T list -> string) -> 'a parser -> 'a parser + val filter_cut' : ('a -> bool) -> ('a -> Context.generic * Token.T list -> string) -> + 'a context_parser -> 'a context_parser + + val nonempty_list : ('b -> string) -> ('a, 'b, 'c list) istate -> ('a, 'b, 'c list) istate + val distinct_list : ('a * 'a -> bool) -> ('a list -> 'c -> string) -> ('b, 'c, 'a list) istate -> + ('b, 'c, 'a list) istate + + val nonempty_name : (Token.T list -> string) -> string parser + + val code : ML_Code_Util.code parser + val nonempty_code : (Token.T list -> string) -> ML_Code_Util.code parser + + val bool : bool parser + val eq : string parser + val parenths : 'a parser -> 'a parser + val parenths' : 'a context_parser -> 'a context_parser + + val thm : thm context_parser + val multi_thm : (thm list) context_parser + val thms : (thm list) context_parser + val nonempty_thms : (Context.generic * Token.T list -> string) -> (thm list) context_parser + + (*the following structure may be ignored; it is needed for below ML_int parser*) + structure Internal_Int_Data : GENERIC_DATA + val ML_int : (Token.T list -> string) -> int context_parser +end + +structure Parse_Util : PARSE_UTIL = +struct + +(*FIXME: move to Isabelle kernel?*) +fun get_tokens_pos [] = " (end-of-input)" + | get_tokens_pos (tok :: _) = Position.here (Token.pos_of tok); + +fun cut get_pos kind scan = + let + fun err (x, NONE) = (fn () => kind ^ get_pos x) + | err (x, SOME msg) = + (fn () => + let val s = msg () in + if String.isPrefix kind s then s + else kind ^ get_pos x ^ ": " ^ s + end); + in Scan.!! err scan end; + +fun !!!+ scan = cut (get_tokens_pos o snd) "Outer syntax error" scan + +fun enum1' sep scan = scan ::: Scan.repeat (Scan.lift (Parse.$$$ sep) |-- !!!+ scan); +fun enum' sep scan = enum1' sep scan || Scan.succeed []; +fun and_list1' scan = enum1' "and" scan; +fun and_list' scan = enum' "and" scan; + +type ('i, 'j, 'a) istate = 'i -> 'a * 'j + +fun if_ahead ahead f scan = + (Scan.ahead ahead :|-- (fn st => fn xs => (f st xs, xs))) || scan + +fun if_eof scan = if_ahead Parse.eof scan +fun if_eof' scan = if_ahead (Scan.lift Parse.eof) scan + +fun option scan = if_eof (K o K NONE) (Scan.option scan) +fun option' scan = if_eof' (K o K NONE) (Scan.option scan) +fun optional scan default = option scan >> the_default default +fun optional' scan default = option' scan >> the_default default + +fun gen_position not_eof scan = + (Scan.optional (Scan.ahead not_eof >> Token.pos_of) Position.none + >> (snd o Position.default)) + -- scan + >> swap + +fun position scan = gen_position Parse.not_eof scan +fun position' scan = gen_position (Scan.lift Parse.not_eof) scan + +fun fail msg_of = Scan.fail_with (fn x => fn _ => msg_of x) +fun gen_abort cut msg_of = fail msg_of |> cut + +fun abort msg_of = gen_abort Parse.!!! msg_of +fun abort' msg_of = gen_abort !!!+ msg_of + +fun filter p err scan = scan :|-- (fn x => if p x then pair x else err x) +fun gen_filter_cut abort p msg_of = filter p (abort o msg_of) +fun filter_cut p = gen_filter_cut abort p +fun filter_cut' p = gen_filter_cut abort' p + +fun nonempty_list msg_of = filter (not o null) (K (fail msg_of)) +fun distinct_list eq msg_of = filter (not o has_duplicates eq) (fail o msg_of) +fun nonempty_name msg_of = filter (curry (op <>) "") (K (fail msg_of)) Args.name + +val code = if_eof (K o K []) ML_Code_Util.parse_code +fun nonempty_code msg_of = nonempty_list msg_of code + +val bool = Scan.recover (Args.name >> Value.parse_bool) + (fail (K "expected \"true\" or \"false\"")) + +val eq = Parse.$$$ "\" || Parse.$$$ "=" || fail (K "expected \"=\" or \"\\"") + +fun parenths scan = Parse.$$$ "(" |-- scan --| Parse.$$$ ")" +fun parenths' scan = (Scan.lift (Parse.$$$ "(")) |-- scan --| (Scan.lift (Parse.$$$ ")")) + +(*for better error messages than Attrib.thm and friends*) +fun thm x = Attrib.thm x handle ERROR m => fail (K m) () +fun multi_thm x = Attrib.multi_thm x handle ERROR m => fail (K m) () +val thms = Scan.repeats multi_thm +fun nonempty_thms msg_of = nonempty_list msg_of thms + +structure Internal_Int_Data = Generic_Data( type T = int val empty = 0 val merge = fst) + +fun ML_int msg_of = + let fun int_from_code ((code, ctxt), pos) = + let val put_int_code = ML_Code_Util.read "Parse_Util.Internal_Int_Data.put" + @ ML_Code_Util.atomic code + in + Context.Proof ctxt + |> ML_Attribute_Util.attribute_map_context (ML_Attribute.run_map_context (put_int_code, pos)) + |> Internal_Int_Data.get + end + in + Scan.lift Parse.int + || ((Scan.lift (nonempty_code msg_of) -- Args.context |> position') >> int_from_code) + end +end diff --git a/thys/ML_Unification/ML_Utils/Priorities/ML_Priorities.thy b/thys/ML_Unification/ML_Utils/Priorities/ML_Priorities.thy new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/ML_Utils/Priorities/ML_Priorities.thy @@ -0,0 +1,11 @@ +\<^marker>\creator "Kevin Kappelmann"\ +theory ML_Priorities + imports ML_Parsing_Utils +begin + +paragraph \Summary\ +text \Priorities for ML tactics.\ + +ML_file\priority.ML\ + +end diff --git a/thys/ML_Unification/ML_Utils/Priorities/priority.ML b/thys/ML_Unification/ML_Utils/Priorities/priority.ML new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/ML_Utils/Priorities/priority.ML @@ -0,0 +1,39 @@ +signature PRIO = sig + type prio = int + + val VERY_LOW : prio + val LOW : prio + val MEDIUM : prio + val HIGH : prio + val VERY_HIGH : prio + + val eq : prio * prio -> bool + val ord : prio ord + + structure Table : TABLE + + val parse : prio context_parser + val pretty : prio -> Pretty.T + +end + +structure Prio : PRIO = +struct + +type prio = int + +val VERY_LOW = ~2000 +val LOW = ~1000 +val MEDIUM = 0 +val HIGH = 1000 +val VERY_HIGH = 2000 + +val eq = (op =) +val ord = int_ord o swap + +structure Table = Table(type key = prio; val ord = ord) + +val pretty = Pretty.str o string_of_int +val parse = Parse_Util.ML_int (K "priority may not be empty") + +end \ No newline at end of file diff --git a/thys/ML_Unification/ML_Utils/Tactics/ML_Tactic_Utils.thy b/thys/ML_Unification/ML_Utils/Tactics/ML_Tactic_Utils.thy new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/ML_Utils/Tactics/ML_Tactic_Utils.thy @@ -0,0 +1,15 @@ +\<^marker>\creator "Kevin Kappelmann"\ +section \ML Tactic Utils\ +theory ML_Tactic_Utils + imports + ML_Logger + ML_Term_Utils + ML_Conversion_Utils +begin + +paragraph \Summary\ +text \Utilities for tactics.\ + +ML_file\tactic_util.ML\ + +end \ No newline at end of file diff --git a/thys/ML_Unification/ML_Utils/Tactics/tactic_util.ML b/thys/ML_Unification/ML_Utils/Tactics/tactic_util.ML new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/ML_Utils/Tactics/tactic_util.ML @@ -0,0 +1,234 @@ +(* Title: ML_Utils/tactic_util.ML + Author: Kevin Kappelmann + +Tactic utilities. +*) +signature TACTIC_UTIL = +sig + include HAS_LOGGER + (* tactic combinators*) + val HEADGOAL : (int -> 'a) -> 'a + + val TRY' : ('a -> tactic) -> 'a -> tactic + + val EVERY_ARG : ('a -> tactic) -> 'a list -> tactic + val EVERY_ARG' : ('a -> 'b -> tactic) -> 'a list -> 'b -> tactic + val CONCAT : tactic list -> tactic + val CONCAT' : ('a -> tactic) list -> 'a -> tactic + + val FOCUS_PARAMS' : (Subgoal.focus -> int -> tactic) -> Proof.context -> int -> tactic + val FOCUS_PARAMS_CTXT : (Proof.context -> tactic) -> Proof.context -> int -> tactic + val FOCUS_PARAMS_CTXT' : (Proof.context -> int -> tactic) -> Proof.context -> int -> tactic + val FOCUS_PARAMS_FIXED' : (Subgoal.focus -> int -> tactic) -> Proof.context -> int -> tactic + val FOCUS_PREMS' : (Subgoal.focus -> int -> tactic) -> Proof.context -> int -> tactic + val FOCUS' : (Subgoal.focus -> int -> tactic) -> Proof.context -> int -> tactic + val SUBPROOF' : (Subgoal.focus -> int -> tactic) -> Proof.context -> int -> tactic + + val CSUBGOAL_DATA : (cterm -> 'a) -> ('a -> int -> tactic) -> int -> tactic + val CSUBGOAL_STRIPPED : (cterm Binders.binders * (cterm list * cterm) -> 'a) -> + ('a -> int -> tactic) -> int -> tactic + val SUBGOAL_DATA : (term -> 'a) -> ('a -> int -> tactic) -> int -> tactic + val SUBGOAL_STRIPPED : ((string * typ) list * (term list * term) -> 'a) -> + ('a -> int -> tactic) -> int -> tactic + + (* tactics *) + val insert_tac : thm list -> Proof.context -> int -> tactic + (*thin_tac n i deletes nth premise of the ith subgoal*) + val thin_tac : int -> int -> tactic + val thin_tacs : int list -> int -> tactic + + val safe_simp_tac : Proof.context -> int -> tactic + + (*nth_eq_assume_tac n i solves ith subgoal by assumption, without unification, + preferring premise n*) + val nth_eq_assume_tac : int -> int -> tactic + + (*resolution with matching and without lifting of premises and parameters*) + val no_lift_biresolve_tac : bool -> thm -> int -> Proof.context -> int -> tactic + val no_lift_resolve_tac : thm -> int -> Proof.context -> int -> tactic + val no_lift_eresolve_tac : thm -> int -> Proof.context -> int -> tactic + + (*rewrite subgoal according to the given equality theorem "lhs \ subgoal"*) + val rewrite_subgoal_tac : thm -> Proof.context -> int -> tactic + + (*creates equality theorem for a subgoal from an equality theorem for the subgoal's conclusion*) + val eq_subgoal_from_eq_concl : cterm Binders.binders -> cterm list -> thm -> Proof.context -> thm + + (*rewrite a subgoal given an equality theorem and environment for the subgoal's conclusion*) + val rewrite_subgoal_unif_concl : Envir_Normalisation.thm_normaliser -> + Envir_Normalisation.thm_normaliser -> term Binders.binders -> (Envir.env * thm) -> + Proof.context -> int -> tactic + + (*protect premises starting from (and excluding) the given index in the specified subgoal*) + val protect_tac : int -> Proof.context -> int -> tactic + (*unprotect the subgoal*) + val unprotect_tac : Proof.context -> int -> tactic + (*protect, then apply tactic, then unprotect*) + val protected_tac : int -> (int -> tactic) -> Proof.context -> int -> tactic + + (*focus on the specified subgoal, introducing only the specified list of premises + as theorems in the focus*) + val focus_prems_tac : int list -> (Subgoal.focus -> int -> tactic) -> + Proof.context -> int -> tactic + (*focus_prems_tac and then deletes all focused premises*) + val focus_delete_prems_tac : int list -> (Subgoal.focus -> int -> tactic) -> + Proof.context -> int -> tactic + + (*apply tactic to the specified subgoal, where the state is given as a cterm*) + val apply_tac : (int -> tactic) -> int -> cterm -> thm Seq.seq + +end + +structure Tactic_Util : TACTIC_UTIL = +struct + +val logger = Logger.setup_new_logger Logger.root_logger "Tactic_Util" + +(* tactic combinators *) +fun HEADGOAL f = f 1 + +fun TRY' tac = tac ORELSE' K all_tac + +fun EVERY_ARG tac = EVERY o map tac +fun EVERY_ARG' tac = EVERY' o map tac + +fun CONCAT tacs = fold_rev (curry op APPEND) tacs no_tac +fun CONCAT' tacs = fold_rev (curry op APPEND') tacs (K no_tac) + +fun FOCUS_PARAMS' tac = Subgoal.FOCUS_PARAMS (HEADGOAL o tac) +fun FOCUS_PARAMS_FIXED' tac = Subgoal.FOCUS_PARAMS_FIXED (HEADGOAL o tac) +fun FOCUS_PREMS' tac = Subgoal.FOCUS_PREMS (HEADGOAL o tac) +fun FOCUS' tac = Subgoal.FOCUS (HEADGOAL o tac) +fun SUBPROOF' tac = Subgoal.SUBPROOF (HEADGOAL o tac) + +fun FOCUS_PARAMS_CTXT tac = Subgoal.FOCUS_PARAMS (#context #> tac) +fun FOCUS_PARAMS_CTXT' tac = FOCUS_PARAMS' (#context #> tac) + +fun CSUBGOAL_DATA f tac = CSUBGOAL (uncurry tac o apfst f) +fun CSUBGOAL_STRIPPED f = CSUBGOAL_DATA (f o Term_Util.strip_csubgoal) +fun SUBGOAL_DATA f tac = SUBGOAL (uncurry tac o apfst f) +fun SUBGOAL_STRIPPED f = SUBGOAL_DATA (f o Term_Util.strip_subgoal) + +(* tactics *) +fun insert_tac thms ctxt = Method.insert_tac ctxt thms + +fun thin_tac n = if n < 1 then K no_tac + else rotate_tac (n - 1) THEN' (DETERM o eresolve0_tac [thin_rl]) THEN' rotate_tac (~n + 1) +val thin_tacs = sort int_ord + #> map_index ((op -) o swap) + #> EVERY_ARG' thin_tac + +(*some "safe" solvers create instantiations via flex-flex pairs, which we disallow here*) +fun safe_simp_tac ctxt i state = + let val eq_flexps = Thm.tpairs_of + #> pair (Thm.tpairs_of state) + #> eq_list (eq_pair (op aconv) (op aconv)) + in Simplifier.safe_simp_tac ctxt i state |> Seq.filter eq_flexps end + +fun nth_eq_assume_tac n = if n < 1 then K no_tac + else rotate_tac (n - 1) THEN' eq_assume_tac + +(*resolution*) +fun no_lift_biresolve_tac eres brule nsubgoals ctxt = + Thm.bicompose (SOME ctxt) {flatten = true, match = true, incremented = true} + (eres, brule, nsubgoals) + #> PRIMSEQ + +val no_lift_resolve_tac = no_lift_biresolve_tac false +val no_lift_eresolve_tac = no_lift_biresolve_tac true + +fun rewrite_subgoal_tac eq_thm = + let + val (lhs, rhs) = Thm.dest_equals (Thm.cconcl_of eq_thm) + val eq_elim = Thm.instantiate' [] [SOME lhs, SOME rhs] Drule.equal_elim_rule1 + val thm = Thm.implies_elim eq_elim eq_thm + in no_lift_resolve_tac thm 1 end + +fun eq_subgoal_from_eq_concl cbinders cprems eq_thm ctxt = + let + fun all_abstract ((x, T), cfree) = + let val all_const = Logic.all_const T |> Thm.cterm_of ctxt + in + Thm.abstract_rule x cfree + #> Drule.arg_cong_rule all_const + end + in + (*introduce the premises*) + fold_rev (Drule.imp_cong_rule o Thm.reflexive) cprems eq_thm + (*introduce abstractions for the fresh Frees*) + |> fold all_abstract cbinders + end + +fun rewrite_subgoal_unif_concl norm_state norm_eq_thm binders (env, eq_thm) ctxt i state = + let + fun rewrite_tac (rparams, (prems, _)) = + let + (*update binders with their normalised type*) + val binders = map2 (fn (_, T) => fn ((n, _), Free (n', _)) => ((n, T), Free (n', T))) + rparams binders + val cterm_of = Thm.cterm_of ctxt + val cprems = map (cterm_of o Binders.replace_binders binders) prems + val eq_thm = eq_subgoal_from_eq_concl (map (apsnd cterm_of) binders) cprems + (norm_eq_thm ctxt env eq_thm) ctxt + in rewrite_subgoal_tac eq_thm ctxt end + in + ((PRIMITIVE (norm_state ctxt env) |> K) + THEN' SUBGOAL_STRIPPED I rewrite_tac) i state + end + +fun protect_tac n ctxt = + let fun protect cbinders (cprems, cconcl) i = + let val nprems = length cprems + in + if nprems < n then + (@{log Logger.WARN} ctxt (fn _ => Pretty.block [ + Pretty.str "Not enough premises to protect. Given number: ", + Pretty.str (string_of_int n), + Pretty.str ". But there are only ", + Pretty.str (string_of_int nprems), + Pretty.str " premise(s) in subgoal ", + Pretty.str (string_of_int i), + Pretty.str ": ", + Logic.close_prop (map (apfst fst o apsnd Thm.term_of) cbinders) (map Thm.term_of cprems) + (Thm.term_of cconcl) |> Syntax.pretty_term ctxt + ] |> Pretty.string_of); + no_tac) + else + let + val (cprems_unprotected, cconcl_protected) = chop n cprems + ||> Drule.list_implies o rpair cconcl + val protected_eq_thm = @{thm Pure.prop_def} + |> Thm.instantiate' [] [SOME cconcl_protected] + |> (fn eq_thm => eq_subgoal_from_eq_concl cbinders cprems_unprotected eq_thm ctxt) + in rewrite_subgoal_tac protected_eq_thm ctxt i end + end + in if n < 0 then K no_tac else CSUBGOAL_STRIPPED I (uncurry protect) end + +fun unprotect_tac ctxt = match_tac ctxt [Drule.protectI] + +fun protected_tac n tac ctxt = + protect_tac n ctxt + THEN' tac + THEN_ALL_NEW (unprotect_tac ctxt) + +fun focus_prems_tac is tac ctxt = + CONVERSION (Conversion_Util.move_prems_to_front_conv is) + THEN' protect_tac (length is) ctxt + THEN' FOCUS' (fn {context=ctxt, params=params, prems=prems, + asms=asms, concl=concl, schematics=schematics} => + let val concl = Term_Util.unprotect concl + in + unprotect_tac ctxt + THEN' tac {context=ctxt, params=params, prems=prems, asms=asms, concl=concl, + schematics=schematics} + end + ) ctxt + +fun focus_delete_prems_tac is tac ctxt = + focus_prems_tac is tac ctxt + THEN' thin_tacs (1 upto length is) + +fun apply_tac tac i = Goal.init #> tac i #> Seq.map Goal.conclude + +end + diff --git a/thys/ML_Unification/ML_Utils/Terms/ML_Term_Utils.thy b/thys/ML_Unification/ML_Utils/Terms/ML_Term_Utils.thy new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/ML_Utils/Terms/ML_Term_Utils.thy @@ -0,0 +1,12 @@ +\<^marker>\creator "Kevin Kappelmann"\ +section \ML Term Utils\ +theory ML_Term_Utils + imports ML_Binders +begin + +paragraph \Summary\ +text \Utilities for terms.\ + +ML_file\term_util.ML\ + +end \ No newline at end of file diff --git a/thys/ML_Unification/ML_Utils/Terms/term_util.ML b/thys/ML_Unification/ML_Utils/Terms/term_util.ML new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/ML_Utils/Terms/term_util.ML @@ -0,0 +1,95 @@ +(* Title: ML_Utils/term_util.ML + Author: Kevin Kappelmann + +Term utilities. +*) +signature TERM_UTIL = +sig + val fresh_vars : string -> typ list -> int -> term list * int + val fresh_var : string -> typ -> int -> term * int + + (*returns outer parameters in reversed order*) + val strip_all : term -> (string * typ) list * term + val strip_call : cterm -> cterm Binders.binders * cterm + + (*returns premises and conclusion*) + val strip_imp : term -> term list * term + val strip_cimp : cterm -> cterm list * cterm + + (*returns outer parameters in reversed order, premises, and conclusion*) + val strip_subgoal : term -> (string * typ) list * (term list * term) + val strip_csubgoal : cterm -> cterm Binders.binders * (cterm list * cterm) + val strip_nth_subgoal : int -> thm -> (string * typ) list * (term list * term) + val strip_nth_csubgoal : int -> thm -> cterm Binders.binders * (cterm list * cterm) + + (*replace outer parameters by fresh meta variables*) + val forall_elim_vars : term -> int -> term * int + + val protect : cterm -> cterm + val unprotect : cterm -> cterm + + val no_dummy_pattern : term -> bool + val has_dummy_pattern : term -> bool +end + +structure Term_Util : TERM_UTIL = +struct + +fun fresh_vars name Ts idx = + let fun mk_var_inc T i = (Var ((name, i + 1), T), (i + 1)) + in fold_map mk_var_inc Ts idx end + +fun fresh_var name = yield_singleton (fresh_vars name) + +val strip_all = + let fun strip params (Const ("Pure.all", _) $ Abs (a, T, t)) = strip ((a, T) :: params) t + | strip params t = (params, t) + in strip [] end + +val strip_call = + let fun strip params ct = + ((let val (cp as (_, cx)) = Thm.dest_comb ct + in case apply2 Thm.term_of cp of + (Const ("Pure.all", _), Abs (a, T, _)) => + let val (cf, cb) = Thm.dest_abs_global cx + in strip (((a, T), cf) :: params) cb end + | _ => (params, ct) + end) + handle CTERM _ => (params, ct)) + in strip [] end + +fun strip_imp t = + (let + val (prem, t) = Logic.dest_implies t + val (prems, concl) = strip_imp t + in (prem :: prems, concl) end) + handle TERM _ => ([], t) + +fun strip_cimp ct = + (let + val (cprem, ct) = Thm.dest_implies ct + val (cprems, cconcl) = strip_cimp ct + in (cprem :: cprems, cconcl) end) + handle TERM _ => ([], ct) + +val strip_subgoal = strip_all ##> strip_imp +val strip_csubgoal = strip_call ##> strip_cimp +fun strip_nth_subgoal i = Thm.prop_of #> pair i #> Logic.nth_prem #> strip_subgoal +fun strip_nth_csubgoal i state = Thm.cprem_of state i |> strip_csubgoal + +fun forall_elim_vars t i = + let val (params, tbody) = strip_all t + in + if null params then (t, i) + else + fold_map (uncurry fresh_var) params i + |>> Term.subst_bounds o rpair tbody + end + +val protect = Drule.protect +val unprotect = Thm.dest_arg + +val no_dummy_pattern = can (Term.no_dummy_patterns) +val has_dummy_pattern = not o no_dummy_pattern + +end \ No newline at end of file diff --git a/thys/ML_Unification/ML_Utils/Theorems/ML_Theorem_Utils.thy b/thys/ML_Unification/ML_Utils/Theorems/ML_Theorem_Utils.thy new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/ML_Utils/Theorems/ML_Theorem_Utils.thy @@ -0,0 +1,12 @@ +\<^marker>\creator "Kevin Kappelmann"\ +section \ML Theorem Utils\ +theory ML_Theorem_Utils + imports Pure +begin + +paragraph \Summary\ +text \Utilities for theorems.\ + +ML_file\thm_util.ML\ + +end \ No newline at end of file diff --git a/thys/ML_Unification/ML_Utils/Theorems/thm_util.ML b/thys/ML_Unification/ML_Utils/Theorems/thm_util.ML new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/ML_Utils/Theorems/thm_util.ML @@ -0,0 +1,25 @@ +(* Title: ML_Utils/thm_util.ML + Author: Kevin Kappelmann + +Theorem utilities. +*) +signature THM_UTIL = +sig + (*"match_implies_elim prem thm" matches the first premise of thm against prem and then removes + the premise; without lifting*) + val match_implies_elim : thm -> thm -> thm + val protect : thm -> thm +end + +structure Thm_Util : THM_UTIL = +struct + +fun match_implies_elim prem thm = + Thm.instantiate (Thm.first_order_match (Thm.cprem_of thm 1, Thm.cprop_of prem)) thm + |> (fn thm => Thm.implies_elim thm prem) + +fun protect thm = Drule.protectI + |> Thm.instantiate (TVars.empty, Vars.make [((("A", 0), propT), Thm.cprop_of thm)]) + |> Thm.elim_implies thm + +end \ No newline at end of file diff --git a/thys/ML_Unification/Normalisations/ML_Normalisations.thy b/thys/ML_Unification/Normalisations/ML_Normalisations.thy new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/Normalisations/ML_Normalisations.thy @@ -0,0 +1,14 @@ +\<^marker>\creator "Kevin Kappelmann"\ +section \ML-Normalisations\ +theory ML_Normalisations + imports + Pure +begin + +paragraph \Summary\ +text \Normalisation functions for terms, types, and theorems.\ + +ML_file\term_normalisation.ML\ +ML_file\envir_normalisation.ML\ + +end \ No newline at end of file diff --git a/thys/ML_Unification/Normalisations/envir_normalisation.ML b/thys/ML_Unification/Normalisations/envir_normalisation.ML new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/Normalisations/envir_normalisation.ML @@ -0,0 +1,244 @@ +(* Title: ML_Unification/envir_normalisation.ML + Author: Kevin Kappelmann + +Normalisations with respect to an environment. Adapted and generalised from envir.ML +*) +signature ENVIR_NORMALISATION = +sig + (* types *) + type type_normaliser = Type.tyenv -> typ -> typ + val norm_type_match : type_normaliser + val norm_type_unif : type_normaliser + + (* terms *) + type term_type_normaliser = Type.tyenv -> Term_Normalisation.term_normaliser + val norm_term_types : type_normaliser -> term_type_normaliser + type term_normaliser = Envir.env -> Term_Normalisation.term_normaliser + + (** matching **) + val norm_term_types_match : term_type_normaliser + (*without beta-normalisation*) + val norm_term_match : term_normaliser + val eta_short_norm_term_match : term_normaliser + (*with beta-normalisation*) + val beta_norm_term_match : term_normaliser + val beta_eta_short_norm_term_match : term_normaliser + + (** unification **) + val norm_term_types_unif : term_type_normaliser + (*without beta-normalisation*) + val norm_term_unif : term_normaliser + val eta_short_norm_term_unif : term_normaliser + (*with beta-normalisation*) + val beta_norm_term_unif : term_normaliser + val beta_eta_short_norm_term_unif : term_normaliser + + (* theorems *) + type thm_normaliser = Proof.context -> Envir.env -> thm -> thm + type thm_type_normaliser = thm_normaliser + + val norm_thm_types : type_normaliser -> thm_normaliser + val norm_thm_types_match : thm_normaliser + val norm_thm_types_unif : thm_normaliser + val norm_thm : type_normaliser -> term_normaliser -> thm_normaliser + val norm_thm_match : thm_normaliser + val norm_thm_unif : thm_normaliser + val eta_short_norm_thm : type_normaliser -> term_normaliser -> thm_normaliser + val eta_short_norm_thm_match : thm_normaliser + val eta_short_norm_thm_unif : thm_normaliser + (*with beta-normalisation and eta-contraction*) + val beta_eta_short_norm_thm : type_normaliser -> term_normaliser -> thm_normaliser + val beta_eta_short_norm_thm_match : thm_normaliser + val beta_eta_short_norm_thm_unif : thm_normaliser + +end + +structure Envir_Normalisation : ENVIR_NORMALISATION = +struct + +(* types *) +type type_normaliser = Type.tyenv -> typ -> typ + +val norm_type_match = Envir.subst_type +val norm_type_unif = Envir.norm_type + +(* terms *) +type term_type_normaliser = Type.tyenv -> Term_Normalisation.term_normaliser +fun norm_term_types norm_type = map_types o norm_type + +type term_normaliser = Envir.env -> Term_Normalisation.term_normaliser + +(** matching **) +val norm_term_types_match = norm_term_types norm_type_match + +fun norm_abs_same2 normT norm (a, T, body) = + Abs (a, normT T, Same.commit norm body) + handle Same.SAME => Abs (a, T, norm body) + +fun norm_abs_comb_same beta_norm norm (abs_args as (_, _, body)) arg = if beta_norm + then Same.commit norm (subst_bound (arg, body)) + else let val f = Abs abs_args + in (norm f $ Same.commit norm arg handle Same.SAME => f $ norm arg) end + +fun norm_comb_same beta_norm norm f t = (case norm f of + (nf as Abs (_, _, body)) => if beta_norm + then Same.commit norm (subst_bound (t, body)) + else nf $ Same.commit norm t + | nf => nf $ Same.commit norm t) + handle Same.SAME => f $ norm t + +fun norm_term_match1 beta_norm tenv : term Same.operation = + let + fun norm (Var v) = (case Envir.lookup1 tenv v of + SOME u => u + | NONE => raise Same.SAME) + | norm (Abs (a, T, body)) = Abs (a, T, norm body) + | norm (Abs abs_args $ t) = norm_abs_comb_same beta_norm norm abs_args t + | norm (f $ t) = norm_comb_same beta_norm norm f t + | norm _ = raise Same.SAME + in norm end + +fun norm_term_match2 beta_norm (Envir.Envir {tenv, tyenv, ...}) : term Same.operation = + let + val normT = Envir.subst_type_same tyenv + fun norm (Const (a, T)) = Const (a, normT T) + | norm (Free (a, T)) = Free (a, normT T) + | norm (v as Var (xi, T)) = + let + fun lookup v = (case Envir.lookup1 tenv (dest_Var v) of + SOME u => u + | NONE => raise Same.SAME) + in + (normT T |> (fn T => Same.commit lookup (Var (xi, T)))) + handle Same.SAME => lookup v + end + | norm (Abs args) = norm_abs_same2 normT norm args + | norm (Abs abs_args $ t) = norm_abs_comb_same beta_norm norm abs_args t + | norm (f $ t) = norm_comb_same beta_norm norm f t + | norm _ = raise Same.SAME + in norm end + +fun norm_term_match_same beta_norm (envir as Envir.Envir {tenv, tyenv, ...}) = + if Vartab.is_empty tyenv then norm_term_match1 beta_norm tenv + else norm_term_match2 beta_norm envir + +fun norm_term_match env = Same.commit (norm_term_match_same false env) +fun beta_norm_term_match env = Same.commit (norm_term_match_same true env) +fun eta_short_norm_term_match env = norm_term_match env + #> Term_Normalisation.eta_short +fun beta_eta_short_norm_term_match env = beta_norm_term_match env + #> Term_Normalisation.eta_short + +(** unification **) +val norm_term_types_unif = norm_term_types norm_type_unif + +fun norm_type_unif_same tyenv : typ Same.operation = + let + fun norm (Type (a, Ts)) = Type (a, Same.map norm Ts) + | norm (TFree _) = raise Same.SAME + | norm (TVar v) = + (case Type.lookup tyenv v of + SOME U => Same.commit norm U + | NONE => raise Same.SAME) + in norm end + +fun norm_term_unif1 beta_norm tenv : term Same.operation = + let + fun norm (Var v) = (case Envir.lookup1 tenv v of + SOME u => Same.commit norm u + | NONE => raise Same.SAME) + | norm (Abs (a, T, body)) = Abs (a, T, norm body) + | norm (Abs abs_args $ t) = norm_abs_comb_same beta_norm norm abs_args t + | norm (f $ t) = norm_comb_same beta_norm norm f t + | norm _ = raise Same.SAME + in norm end + +fun norm_term_unif2 beta_norm (envir as Envir.Envir {tyenv, ...}) : term Same.operation = + let + val normT = norm_type_unif_same tyenv + fun norm (Const (a, T)) = Const (a, normT T) + | norm (Free (a, T)) = Free (a, normT T) + | norm (Var (xi, T)) = (case Envir.lookup envir (xi, T) of + SOME u => Same.commit norm u + | NONE => Var (xi, normT T)) + | norm (Abs args) = norm_abs_same2 normT norm args + | norm (Abs abs_args $ t) = norm_abs_comb_same beta_norm norm abs_args t + | norm (f $ t) = norm_comb_same beta_norm norm f t + | norm _ = raise Same.SAME + in norm end + +fun norm_term_unif_same beta_norm (envir as Envir.Envir {tenv, tyenv, ...}) = + if Vartab.is_empty tyenv then norm_term_unif1 beta_norm tenv + else norm_term_unif2 beta_norm envir + +fun norm_term_unif env = Same.commit (norm_term_unif_same false env) +fun beta_norm_term_unif env = Same.commit (norm_term_unif_same true env) +fun eta_short_norm_term_unif env = norm_term_unif env + #> Term_Normalisation.eta_short +fun beta_eta_short_norm_term_unif env = beta_norm_term_unif env + #> Term_Normalisation.eta_short + +type thm_normaliser = Proof.context -> Envir.env -> thm -> thm +type thm_type_normaliser = thm_normaliser + +(** theorems **) + +(*collect and normalise TVars of a term*) +fun collect_norm_types norm_type ctxt tyenv t = + let + val norm_type' = norm_type tyenv + fun prep_type_entry x = (x, Thm.ctyp_of ctxt (norm_type' (TVar x))) + in + fold_types (fold_atyps (fn TVar v => TVars.add (prep_type_entry v) | _ => I)) t + |> TVars.build + end + +(*collect and normalise Vars of a term*) +fun collect_norm_terms norm_type norm_term ctxt (env as Envir.Envir {tyenv,...}) t = + let + val norm_type' = norm_type tyenv + val norm_term' = norm_term env + fun prep_term_entry (x as (n, T)) = ((n, norm_type' T), Thm.cterm_of ctxt (norm_term' (Var x))) + in + fold_aterms (fn Var v => Vars.add (prep_term_entry v) | _ => I) t + |> Vars.build + end + +(*normalise types of a theorem*) +fun norm_thm_types norm_types ctxt (Envir.Envir {tyenv, ...}) thm = + let + val prop = Thm.full_prop_of thm + val type_inst = collect_norm_types norm_types ctxt tyenv prop + val inst = (type_inst, Vars.empty) + in Thm.instantiate inst thm end + +val norm_thm_types_match = norm_thm_types norm_type_match +val norm_thm_types_unif = norm_thm_types norm_type_unif + +(*normalise a theorem*) +fun norm_thm norm_types norm_terms ctxt (env as Envir.Envir {tyenv,...}) thm = + let + val prop = Thm.full_prop_of thm + val type_inst = collect_norm_types norm_types ctxt tyenv prop + val term_inst = collect_norm_terms norm_types norm_terms ctxt env prop + val inst = (type_inst, term_inst) + in Thm.instantiate inst thm end + +val norm_thm_match = norm_thm norm_type_match norm_term_match +val norm_thm_unif = norm_thm norm_type_unif norm_term_unif + +fun eta_short_norm_thm norm_types norm_terms ctxt env = + norm_thm norm_types norm_terms ctxt env + #> Conv.fconv_rule Thm.eta_conversion + +val eta_short_norm_thm_match = eta_short_norm_thm norm_type_match norm_term_match +val eta_short_norm_thm_unif = eta_short_norm_thm norm_type_unif norm_term_unif + +fun beta_eta_short_norm_thm norm_types norm_terms ctxt env = + norm_thm norm_types norm_terms ctxt env + #> Conv.fconv_rule Drule.beta_eta_conversion + +val beta_eta_short_norm_thm_match = beta_eta_short_norm_thm norm_type_match norm_term_match +val beta_eta_short_norm_thm_unif = beta_eta_short_norm_thm norm_type_unif norm_term_unif + +end diff --git a/thys/ML_Unification/Normalisations/term_normalisation.ML b/thys/ML_Unification/Normalisations/term_normalisation.ML new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/Normalisations/term_normalisation.ML @@ -0,0 +1,24 @@ +(* Title: ML_Unification/term_normalisation.ML + Author: Kevin Kappelmann + +Term normalisations. +*) +signature TERM_NORMALISATION = +sig + type term_normaliser = term -> term + val beta : term_normaliser + val eta_short : term_normaliser + val beta_eta_short : term_normaliser +end + +structure Term_Normalisation : TERM_NORMALISATION = +struct + +type term_normaliser = term -> term + +val beta = Envir.beta_norm +val eta_short = Envir.eta_contract +val beta_eta_short = Envir.beta_eta_contract + +end + diff --git a/thys/ML_Unification/README.md b/thys/ML_Unification/README.md new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/README.md @@ -0,0 +1,43 @@ +# ML-Unification + +## Content + +1. First-order and higher-order pattern +[E-unification](https://en.wikipedia.org/wiki/Unification_(computer_science%29#E-unification) +and E-matching. +While unifiers in Isabelle/ML only consider the alpha-beta-eta-equational theory of the lambda-calculus, +unifiers in this article +may take an extra background theory, in the form of an equational prover, into account. +For example, the unification problem `n + 1 = ?m + Suc 0` +may be solved by providing a prover for the background theory `forall n. n + 1 = n + Suc 0`. +2. Tactics, methods, and attributes with adjustable unifiers (e.g.\ resolution, fact, assumption, OF). +3. A generalisation of [unification hints](https://www.researchgate.net/publication/221302555_Hints_in_Unification). +Unification hints are a flexible extension for unifiers. +Among other things, they can be used for reflective tactics, +to provide canonical unification instances, +or to simply strengthen the background theory of a unifier in a controlled manner. +4. Simplifier integration for e-unifiers. +5. Practical combinations of unification algorithms, e.g. a combination of first-order and +higher-order pattern unification. +6. A hierarchical logger for Isabelle/ML, +including per logger configurations with log levels, output channels, message filters. +See `Logger/README.md` for details. + +While this entry works with every object logic, +some extra setup for Isabelle/HOL and application examples are provided. +All unifiers are tested with SpecCheck. + +See `Examples/` for some application examples. + +## Build + +Requirements: +1. The Isabelle development version +2. The AFP development version + +## Future Tasks + +1. Add higher-order unifier E-unification and matching +4. tests: + - adapt tests to work for open terms + diff --git a/thys/ML_Unification/ROOT b/thys/ML_Unification/ROOT new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/ROOT @@ -0,0 +1,56 @@ +chapter AFP + +session ML_Unification = "HOL" + + description + \Various unification utilities for Isabelle/ML, e.g. E-unification, E-matching, and unification hints.\ + options [timeout = 300] + sessions + SpecCheck + directories + "Binders" + "Examples" + "Logger" + "Logger/Data_Structures" + "ML_Utils" + "ML_Utils/Attributes" + "ML_Utils/Conversions" + "ML_Utils/Functor_Instances" + "ML_Utils/General" + "ML_Utils/Generic_Data" + "ML_Utils/Methods" + "ML_Utils/ML_Attributes" + "ML_Utils/ML_Code" + "ML_Utils/Parsing" + "ML_Utils/Priorities" + "ML_Utils/Tactics" + "ML_Utils/Terms" + "ML_Utils/Theorems" + "Normalisations" + "Simps_To" + "Term_Index" + "Tests" + "Unifiers" + "Unification_Attributes" + "Unification_Hints" + "Unification_Parsers" + "Unification_Tactics" + "Unification_Tactics/Assumption" + "Unification_Tactics/Fact" + "Unification_Tactics/Resolution" + theories + ML_Logger + ML_Logger_Examples + ML_Utils + ML_Unifiers + Unification_Tactics + Unification_Attributes + ML_Unification_Hints + ML_Unification_HOL_Setup + E_Unification_Examples + Unification_Hints_Reification_Examples + theories [document = false] + ML_Unification_Tests + document_files + "root.tex" + "root.bib" + diff --git a/thys/ML_Unification/Simps_To/Simps_To.thy b/thys/ML_Unification/Simps_To/Simps_To.thy new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/Simps_To/Simps_To.thy @@ -0,0 +1,72 @@ +\<^marker>\creator "Kevin Kappelmann"\ +section \Simps To\ +theory Simps_To + imports + ML_Tactic_Utils + ML_Theorem_Utils + ML_Unification_Base + Setup_Result_Commands +begin + +paragraph \Summary\ +text \Simple frameworks to ask for the simp-normal form of a term on the user-level.\ + +setup_result simps_to_base_logger = \Logger.new_logger Logger.root_logger "Simps_To_Base"\ + +paragraph \Using Simplification On Left Term\ + +definition "SIMPS_TO s t \ (s \ t)" + +lemma SIMPS_TO_eq: "SIMPS_TO s t \ (s \ t)" + unfolding SIMPS_TO_def by simp + +text \Prevent simplification of second/right argument\ +lemma SIMPS_TO_cong [cong]: "s \ s' \ SIMPS_TO s t \ SIMPS_TO s' t" by simp + +lemma SIMPS_TOI: "PROP SIMPS_TO s s" unfolding SIMPS_TO_eq by simp +lemma SIMPS_TOD: "PROP SIMPS_TO s t \ s \ t" unfolding SIMPS_TO_eq by simp + +ML_file\simps_to.ML\ + + +paragraph \Using Simplification On Left Term Followed By Unification\ + +definition "SIMPS_TO_UNIF s t \ (s \ t)" + +text \Prevent simplification\ +lemma SIMPS_TO_UNIF_cong [cong]: "SIMPS_TO_UNIF s t \ SIMPS_TO_UNIF s t" by simp + +lemma SIMPS_TO_UNIF_eq: "SIMPS_TO_UNIF s t \ (s \ t)" unfolding SIMPS_TO_UNIF_def by simp + +lemma SIMPS_TO_UNIFI: "PROP SIMPS_TO s s' \ s' \ t \ PROP SIMPS_TO_UNIF s t" + unfolding SIMPS_TO_UNIF_eq SIMPS_TO_eq by simp +lemma SIMPS_TO_UNIFD: "PROP SIMPS_TO_UNIF s t \ s \ t" + unfolding SIMPS_TO_UNIF_eq by simp + +ML_file\simps_to_unif.ML\ + + +paragraph \Examples\ + +experiment +begin +lemma + assumes [simp]: "P \ Q" + and [simp]: "Q \ R" + shows "PROP SIMPS_TO P Q" + apply simp \\Note: only the left-hand side is simplified.\ + ML_command\ \\obtaining the normal form theorem for a term in ML\ + Simps_To.SIMPS_TO_thm_resultsq (simp_tac @{context}) @{context} @{cterm P} + |> Seq.list_of |> map @{print} + \ + oops + +schematic_goal + assumes [simp]: "P \ Q" + and [simp]: "Q \ R" + shows "PROP SIMPS_TO P ?Q" + apply simp + by (rule SIMPS_TOI) +end + +end diff --git a/thys/ML_Unification/Simps_To/simps_to.ML b/thys/ML_Unification/Simps_To/simps_to.ML new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/Simps_To/simps_to.ML @@ -0,0 +1,63 @@ +(* Title: ML_Unification/simps_to.ML + Author: Kevin Kappelmann + +Create SIMPS_TO theorems. +*) +signature SIMPS_TO = +sig + include HAS_LOGGER + + val dest_SIMPS_TO : term -> (term * term) + val cdest_SIMPS_TO : cterm -> (cterm * cterm) + + val mk_SIMPS_TO_cprop : cterm -> cterm -> cterm + val mk_SIMPS_TO_var_cprop : Proof.context -> cterm -> cterm + + val finish_SIMPS_TO_tac : Proof.context -> int -> tactic + + val SIMPS_TO_thmsq : (int -> tactic) -> Proof.context -> cterm -> thm Seq.seq + val SIMPS_TO_thm_resultsq : (int -> tactic) -> Proof.context -> cterm -> + (thm * cterm) Seq.seq +end + +structure Simps_To : SIMPS_TO = +struct + +val logger = Logger.setup_new_logger simps_to_base_logger "Simps_To" + +structure Util = Tactic_Util + +val dest_SIMPS_TO = \<^Const_fn>\SIMPS_TO _ for lhs rhs => \(lhs, rhs)\\ +val cdest_SIMPS_TO = Thm.dest_comb #>> Thm.dest_arg + +fun mk_SIMPS_TO_cprop clhs crhs = \<^instantiate>\'a = \Thm.ctyp_of_cterm clhs\ and clhs and crhs + in cprop\PROP (SIMPS_TO clhs crhs)\ for clhs :: 'a\ + +fun mk_SIMPS_TO_var_cprop ctxt ct = + Var (("x", Thm.maxidx_of_cterm ct + 1), Thm.typ_of_cterm ct) + |> Thm.cterm_of ctxt + |> mk_SIMPS_TO_cprop ct + +fun finish_SIMPS_TO_tac ctxt = match_tac ctxt [@{thm SIMPS_TOI}] + +fun SIMPS_TO_thmsq simps_to_tac ctxt ct = + let fun inst_tac cconcl _ = + PRIMITIVE (cdest_SIMPS_TO cconcl |> swap |> Thm.match |> Thm.instantiate) + in + (@{log Logger.DEBUG} ctxt (fn _ => Pretty.block [ + Pretty.str "Creating ", + Syntax.pretty_term ctxt @{term SIMPS_TO}, + Pretty.str " theorems for ", + Syntax.pretty_term ctxt (Thm.term_of ct) + ] |> Pretty.string_of); + mk_SIMPS_TO_var_cprop ctxt ct + |> Util.HEADGOAL (Util.apply_tac ( + simps_to_tac + THEN' Tactic_Util.CSUBGOAL_STRIPPED (snd o snd) inst_tac + THEN' finish_SIMPS_TO_tac ctxt))) + end + +val SIMPS_TO_thm_resultsq = + Seq.map (fn thm => (thm, Thm.cconcl_of thm |> Thm.dest_arg)) ooo SIMPS_TO_thmsq + +end \ No newline at end of file diff --git a/thys/ML_Unification/Simps_To/simps_to_unif.ML b/thys/ML_Unification/Simps_To/simps_to_unif.ML new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/Simps_To/simps_to_unif.ML @@ -0,0 +1,71 @@ +(* Title: ML_Unification/simps_to_unif.ML + Author: Kevin Kappelmann + +Create SIMPS_TO_UNIF theorems. +*) +signature SIMPS_TO_UNIF = +sig + include HAS_LOGGER + + val dest_SIMPS_TO_UNIF : term -> (term * term) + val cdest_SIMPS_TO_UNIF : cterm -> (cterm * cterm) + + val mk_SIMPS_TO_UNIF_cprop : cterm -> cterm -> cterm + + val mk_SIMPS_TO_UNIF_thm : thm -> thm -> thm + + val SIMPS_TO_UNIF_env_thmsq : (int -> tactic) -> Unification_Base.normalisers -> + Unification_Base.unifier -> term Binders.binders -> Proof.context -> term * term -> Envir.env -> + (Envir.env * thm) Seq.seq +end + +structure Simps_To_Unif : SIMPS_TO_UNIF = +struct + +val logger = Logger.setup_new_logger simps_to_base_logger "Simps_To_Unif" + +val dest_SIMPS_TO_UNIF = \<^Const_fn>\SIMPS_TO_UNIF _ for lhs rhs => \(lhs, rhs)\\ +val cdest_SIMPS_TO_UNIF = Thm.dest_comb #>> Thm.dest_arg + +fun mk_SIMPS_TO_UNIF_cprop clhs crhs = \<^instantiate>\'a = \Thm.ctyp_of_cterm clhs\ and clhs and crhs + in cprop\PROP (SIMPS_TO_UNIF clhs crhs)\ for clhs :: 'a\ + +fun mk_SIMPS_TO_UNIF_thm SIMPS_TO_thm eq_thm = + Drule.incr_indexes2 SIMPS_TO_thm eq_thm @{thm SIMPS_TO_UNIFI} + |> Thm_Util.match_implies_elim SIMPS_TO_thm + |> Thm_Util.match_implies_elim eq_thm + +fun SIMPS_TO_UNIF_env_thmsq simps_to_tac norms unif binders ctxt (lhs, rhs) env = + let + val SIMPS_TO_res = Binders.replace_binders binders lhs + |> Thm.cterm_of ctxt + |> Simps_To.SIMPS_TO_thm_resultsq simps_to_tac ctxt + fun unif_res (SIMPS_TO_thm, clhs_norm) = + let val tp = (Binders.replace_frees binders (Thm.term_of clhs_norm), rhs) + in + (@{log Logger.TRACE} ctxt (fn _ => Pretty.block [ + Pretty.str "Result: ", + Thm.pretty_thm ctxt SIMPS_TO_thm, + Pretty.str " Now unifying ", + Unification_Util.pretty_unif_problem ctxt tp + ] |> Pretty.string_of); + unif binders ctxt tp env + |> Seq.map (pair SIMPS_TO_thm)) + end + fun SIMPS_TO_UNIF_res (SIMPS_TO_thm, (env, eq_thm)) = + let val (SIMPS_TO_thm, eq_thm) = + (#norm_thm norms ctxt env SIMPS_TO_thm, #norm_unif_thm norms ctxt env eq_thm) + in (env, mk_SIMPS_TO_UNIF_thm SIMPS_TO_thm eq_thm) end + in + (@{log Logger.DEBUG} ctxt (fn _ => Pretty.block [ + Pretty.str "Creating ", + Syntax.pretty_term ctxt @{term SIMPS_TO_UNIF}, + Pretty.str " theorems for ", + Unification_Util.pretty_terms ctxt [lhs, rhs] + ] |> Pretty.string_of); + SIMPS_TO_res + |> Seq.maps unif_res + |> Seq.map SIMPS_TO_UNIF_res) + end + +end diff --git a/thys/ML_Unification/Term_Index/ML_Term_Index.thy b/thys/ML_Unification/Term_Index/ML_Term_Index.thy new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/Term_Index/ML_Term_Index.thy @@ -0,0 +1,16 @@ +\<^marker>\creator "Kevin Kappelmann"\ +section \Term Indexing\ +theory ML_Term_Index + imports + ML_Normalisations +begin + +paragraph \Summary\ +text \Termin indexes signatures and implementations.\ + +ML_file\term_index.ML\ +ML_file\discrimination_tree.ML\ + +ML_file\term_index_data.ML\ + +end \ No newline at end of file diff --git a/thys/ML_Unification/Term_Index/discrimination_tree.ML b/thys/ML_Unification/Term_Index/discrimination_tree.ML new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/Term_Index/discrimination_tree.ML @@ -0,0 +1,179 @@ +(* Title: Term_Index/discrimination_tree.ML + Author: Lawrence C Paulson, Cambridge University Computer Laboratory, + Sebastian Willenbrink, Kevin Kappelmann, TU Munich + +Discrimination trees based on Pure/net.ML. +Requires operands to be beta-eta-short normalised. +*) +structure Discrimination_Tree : TERM_INDEX = +struct + +structure TIB = Term_Index_Base + +val norm_term = Term_Normalisation.beta_eta_short + +datatype key = CombK | VarK | AtomK of string + +(*Keys are preorder lists of symbols -- Combinations, Vars, Atoms. +Any term whose head is a Var is regarded entirely as a Var. +Abstractions are also regarded as Vars; this covers eta-conversion and "near" +eta-conversions such as %x. ?P (?f x).*) +fun add_key_of_terms (t, ks) = + let fun rands (t, ks) = case t of + (*other cases impossible due to beta-eta-short norm*) + t1 $ t2 => CombK :: rands (t1, add_key_of_terms (t2, ks)) + | Const (n, _) => AtomK n :: ks + | Free (n, _) => AtomK n :: ks + | Bound i => AtomK (Name.bound i) :: ks + in case head_of t of + Var _ => VarK :: ks + | Abs _ => VarK :: ks + | _ => rands (t, ks) + end + +(*convert a term to a list of keys*) +fun key_of_term t = add_key_of_terms (t, []) + +(*Trees indexed by key lists: each arc is labelled by a key. +Each node contains a list of values, and arcs to children. +The empty key addresses the entire tree. +Lookup functions preserve order of values stored at same level.*) +datatype 'a term_index = + Leaf of 'a list + | Tree of {comb: 'a term_index, var: 'a term_index, atoms: 'a term_index Symtab.table} + +val empty = Leaf [] + +fun is_empty (Leaf []) = true + | is_empty _ = false + +val empty_tree = Tree {comb=empty, var=empty, atoms=Symtab.empty} + +(* insert *) + +(*Adds value x to the list at the node addressed by the keys. +Creates node if not already present. +is_dup decides whether a value is a duplicate. +The empty list of keys generates a Leaf node, others a Tree node.*) +fun insert_keys is_dup x ksdtp = case ksdtp of + ([], Leaf xs) => if exists is_dup xs then raise TIB.INSERT else Leaf (x :: xs) + | (keys, Leaf []) => insert_keys is_dup x (keys, empty_tree) (*expand the tree*) + | (CombK :: keys, Tree {comb, var, atoms}) => + Tree {comb=insert_keys is_dup x (keys, comb), var=var, atoms=atoms} + | (VarK :: keys, Tree {comb, var, atoms}) => + Tree {comb=comb, var=insert_keys is_dup x (keys, var), atoms=atoms} + | (AtomK a :: keys, Tree {comb, var, atoms}) => + let val atoms' = Symtab.map_default (a, empty) (fn dt' => insert_keys is_dup x (keys, dt')) atoms + in Tree {comb=comb, var=var, atoms=atoms'} end + +fun insert is_dup (t, x) dt = insert_keys is_dup x (key_of_term t, dt) +fun insert_safe is_dup entry dt = insert is_dup entry dt handle TIB.INSERT => dt + +(* delete *) + +(*Create a new Tree node if it would be nonempty*) +fun new_tree (args as {comb, var, atoms}) = + if is_empty comb andalso is_empty var andalso Symtab.is_empty atoms then empty + else Tree args + +(*Deletes values from the list at the node addressed by the keys. +Raises DELETE if absent. Collapses the tree if possible. +sel is the selector for values to be deleted.*) +fun delete_keys sel ksdtp = case ksdtp of + ([], Leaf xs) => if exists sel xs then Leaf (filter_out sel xs) else raise TIB.DELETE + | (_, Leaf []) => raise TIB.DELETE + | (CombK :: keys, Tree {comb, var, atoms}) => + new_tree {comb=delete_keys sel (keys,comb), var=var, atoms=atoms} + | (VarK :: keys, Tree {comb, var, atoms}) => + new_tree {comb=comb, var=delete_keys sel (keys,var), atoms=atoms} + | (AtomK a :: keys, Tree {comb, var, atoms}) => + let val atoms' = (case Symtab.lookup atoms a of + NONE => raise TIB.DELETE + | SOME dt' => + (case delete_keys sel (keys, dt') of + Leaf [] => Symtab.delete a atoms + | dt'' => Symtab.update (a, dt'') atoms)) + in new_tree {comb=comb, var=var, atoms=atoms'} end + +fun delete eq t dt = delete_keys eq (key_of_term t, dt) +fun delete_safe eq t dt = delete eq t dt handle TIB.DELETE => dt + +(* retrievals *) +type 'a retrieval = 'a term_index -> term -> 'a list + +fun variants_keys dt keys = case (dt, keys) of + (Leaf xs, []) => xs + | (Leaf _, (_ :: _)) => [] (*non-empty keys and empty dt*) + | (Tree {comb, ...}, (CombK :: keys)) => variants_keys comb keys + | (Tree {var, ...}, (VarK :: keys)) => variants_keys var keys + | (Tree {atoms, ...}, (AtomK a :: keys)) => + case Symtab.lookup atoms a of + SOME dt => variants_keys dt keys + | NONE => [] + +fun variants dt t = variants_keys dt (key_of_term t) + +(*Skipping a term in a tree. Recursively skip 2 levels if a combination*) +fun tree_skip (Leaf _) dts = dts + | tree_skip (Tree {comb, var, atoms}) dts = + var :: dts + |> Symtab.fold (fn (_, dt) => fn dts => dt :: dts) atoms + (* tree_skip comb [] only skips first term, so another skip is required *) + |> fold_rev tree_skip (tree_skip comb []) + +(*conses the linked tree, if present, to dts*) +fun look1 (atoms, a) dts = case Symtab.lookup atoms a of + NONE => dts + | SOME dt => dt :: dts + +datatype query = Instances | Generalisations | Unifiables + +(*Return the nodes accessible from the term (cons them before trees) +Var in tree matches any term. +Abs or Var in object: for unification, regarded as a wildcard; otherwise only matches a variable.*) +fun query q t dt dts = + let + fun rands _ (Leaf _) dts = dts + | rands t (Tree {comb, atoms, ...}) dts = case t of + f $ t => fold_rev (query q t) (rands f comb []) dts + | Const (n, _) => look1 (atoms, n) dts + | Free (n, _) => look1 (atoms, n) dts + | Bound i => look1 (atoms, Name.bound i) dts + in case dt of + Leaf _ => dts + | Tree {var, ...} => case (head_of t, q) of + (Var _, Generalisations) => var :: dts (*only matches Var in dt*) + | (Var _, _) => tree_skip dt dts + | (Abs _, Generalisations) => var :: dts (*only a Var can match*) + (*A var instantiation in the abstraction could allow + an eta-reduction, so regard the abstraction as a wildcard.*) + | (Abs _, _) => tree_skip dt dts + | (_, Instances) => rands t dt dts + | (_, _) => rands t dt (var :: dts) (*var could also match*) + end + +fun extract_leaves ls = maps (fn Leaf xs => xs) ls + +fun generalisations dt t = query Generalisations t dt [] |> extract_leaves +fun unifiables dt t = query Unifiables t dt [] |> extract_leaves +fun instances dt t = query Instances t dt [] |> extract_leaves + +(* merge *) + +fun cons_fst x (xs, y) = (x :: xs, y) + +fun dest (Leaf xs) = map (pair []) xs + | dest (Tree {comb, var, atoms}) = flat [ + map (cons_fst CombK) (dest comb), + map (cons_fst VarK) (dest var), + maps (fn (a, net) => map (cons_fst (AtomK a)) (dest net)) (Symtab.dest atoms) + ] + +fun content dt = map #2 (dest dt) + +fun merge eq dt1 dt2 = + if pointer_eq (dt1, dt2) then dt1 + else if is_empty dt1 then dt2 + else fold (fn (ks, v) => fn dt => insert_keys (curry eq v) v (ks, dt)) (dest dt2) dt1 + +end diff --git a/thys/ML_Unification/Term_Index/term_index.ML b/thys/ML_Unification/Term_Index/term_index.ML new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/Term_Index/term_index.ML @@ -0,0 +1,46 @@ +(* Title: Term_Index/term_index.ML + Author: Kevin Kappelmann, Sebastian Willenbrink + +Signatures for term indexes. For a brief overview, see: +https://en.wikipedia.org/wiki/Term_indexing +*) +signature TERM_INDEX_BASE = +sig + exception INSERT + exception DELETE +end + +structure Term_Index_Base : TERM_INDEX_BASE = +struct + exception INSERT + exception DELETE +end + +signature TERM_INDEX = +sig + type 'a term_index + val empty : 'a term_index + val is_empty : 'a term_index -> bool + val content : 'a term_index -> 'a list + + (*puts a term into normal form as required by the index structure*) + val norm_term : Term_Normalisation.term_normaliser + (*first argument decides which values should be considered as duplicates; + raises Term_Index_Base.INSERT if the term, value pair is already inserted*) + val insert : ('a -> bool) -> (term * 'a) -> 'a term_index -> 'a term_index + val insert_safe : ('a -> bool) -> (term * 'a) -> 'a term_index -> 'a term_index + + (*first argument selects which values for the given term should be removed; + raises Term_Index_Base.DELETE if no value is deleted*) + val delete : ('a -> bool) -> term -> 'a term_index -> 'a term_index + val delete_safe : ('a -> bool) -> term -> 'a term_index -> 'a term_index + + type 'a retrieval = 'a term_index -> term -> 'a list + val variants : 'a retrieval + val generalisations : 'a retrieval + val instances : 'a retrieval + val unifiables : 'a retrieval + + (*first argument is the equality test for values*) + val merge : (('a * 'a) -> bool) -> 'a term_index -> 'a term_index -> 'a term_index +end \ No newline at end of file diff --git a/thys/ML_Unification/Term_Index/term_index_data.ML b/thys/ML_Unification/Term_Index/term_index_data.ML new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/Term_Index/term_index_data.ML @@ -0,0 +1,32 @@ +(* Title: Term_Index/term_index_data.ML + Author: Kevin Kappelmann + +Term index data stored in generic contexts. +*) +signature TERM_INDEX_DATA_ARGS = +sig + type data + structure TI : TERM_INDEX + val data_eq : data * data -> bool +end + +functor Term_Index_Generic_Data_Args(P : TERM_INDEX_DATA_ARGS) : GENERIC_DATA_ARGS = +struct +type T = P.data P.TI.term_index +val empty = P.TI.empty +fun merge (ti1, ti2)= P.TI.merge P.data_eq ti1 ti2 +end + +signature TERM_INDEX_DATA = +sig + structure TI : TERM_INDEX + structure TI_Data : GENERIC_DATA +end + +functor Term_Index_Data(P : TERM_INDEX_DATA_ARGS) : TERM_INDEX_DATA = +struct + +structure TI = P.TI +structure TI_Data = Generic_Data(Term_Index_Generic_Data_Args(P)) + +end diff --git a/thys/ML_Unification/Tests/First_Order_ML_Unification_Tests.thy b/thys/ML_Unification/Tests/First_Order_ML_Unification_Tests.thy new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/Tests/First_Order_ML_Unification_Tests.thy @@ -0,0 +1,238 @@ +\<^marker>\creator "Kevin Kappelmann"\ +\<^marker>\contributor "Paul Bachmann"\ +subsection \First-Order ML Unification Tests\ +theory First_Order_ML_Unification_Tests + imports + ML_Unification_Tests_Base +begin + +paragraph \Summary\ +text \Tests for @{ML_structure "First_Order_Unification"}.\ + +ML\ + structure Prop = SpecCheck_Property + structure UC = Unification_Combinator + open Unification_Tests_Base + structure Unif = First_Order_Unification + structure Norm = Envir_Normalisation + val match = Unif.match [] + val match_hints = + let fun match binders = + UC.add_fallback_matcher + (Unif.e_match Unification_Util.match_types) + ((fn binders => + (Hints.map_retrieval (Hints.mk_retrieval Hints.TI.generalisations |> K) + #> Hints.UH.map_concl_unifier (Higher_Order_Pattern_Unification.match |> K) + #> Hints.UH.map_normalisers (Unification_Util.beta_eta_short_norms_match |> K) + #> Hints.UH.map_prems_unifier (match |> UC.norm_matcher Norm.beta_norm_term_match |> K)) + |> Context.proof_map + #> Test_Unification_Hints.try_hints binders) + |> UC.norm_matcher (#norm_term Unif.norms_match)) + binders + in match [] end + + val unify = Unif.unify [] + val unify_hints = + let fun unif binders = + UC.add_fallback_unifier + (Unif.e_unify Unification_Util.unify_types) + ((fn binders => + (Hints.UH.map_concl_unifier (Higher_Order_Pattern_Unification.match |> K) + #> Hints.UH.map_normalisers (Unification_Util.beta_eta_short_norms_unif |> K) + #> Hints.UH.map_prems_unifier (unif |> UC.norm_unifier Norm.beta_norm_term_unif |> K)) + |> Context.proof_map + #> Test_Unification_Hints.try_hints binders) + |> UC.norm_unifier (#norm_term Unif.norms_unify)) + binders + in unif [] end +\ + +subsubsection \Matching\ +paragraph \Unit Tests\ + +ML_command\ + let + val ctxt = Proof_Context.set_mode Proof_Context.mode_schematic @{context} + |> Variable.declare_term @{term "f :: nat \ bool \ nat"} + val tests = map (apply2 (Syntax.read_term ctxt))[ + ("(?x :: nat \ ?'Z) 1", "f 1"), + ("?x :: nat", "(?x + 1 :: nat)"), + ("(g :: nat \ nat \ nat) ?x ?x", "(g :: nat \ nat \ nat) (?x + 1) (?x + 1)"), + ("\y. (?x :: nat \ ?'Z) y", "\y. f y"), + ("(g :: ?'X \ ?'Y \ ?'Z) (x :: ?'X) (y :: ?'Y)", "(g :: ?'Y \ ?'Z \ ?'Z) (x :: ?'Y) (y :: ?'Z)"), + ("(g :: ?'Z \ ?'X)", "\(x :: ?'X). (g :: ?'X \ ?'Y) x"), + ("\ (x :: nat). (0 :: nat)", "\ (x :: nat). (0 :: nat)"), + ("\ (x :: nat). x", "\ (x :: nat). x"), + ("\ (x :: nat) (y :: bool). (x, y)", "\ (x :: nat) (y :: bool). (x, y)"), + ("\ (x :: ?'A) (y :: bool). (?x :: ?'A \ bool \ ?'Z) x y", "\ (x :: nat) (y :: bool). f x y"), + ("\(x :: ?'X). (g :: ?'X \ ?'X) x", "(g :: ?'X \ ?'X)"), + ("?g ?x ?y d", "g ?y ?x d"), + ("f 0 True", "(\x y. f y x) True 0"), + ("\ (x :: ?'a) y. f y", "\(x :: ?'b). f"), + ("\y z. (?x :: nat \ bool \ nat) y z", "f"), + ("\x. (?f :: nat \ bool \ nat) 0 x", "\x. f 0 x") + ] + val check = check_unit_tests_hints_match tests true [] + in + Lecker.test_group ctxt () [ + check "match" match, + check "match_hints" match_hints + ] + end +\ + +subparagraph \Asymmetry\ + +ML_command\ + let + val ctxt = Proof_Context.set_mode Proof_Context.mode_schematic @{context} + |> Variable.declare_term @{term "f :: nat \ bool \ nat"} + val tests = map (apply2 (Syntax.read_term ctxt))[ + ("f 1", "(?x :: nat \ ?'Z) 1") + ] + val check = check_unit_tests_hints_match tests false [] + in + Lecker.test_group ctxt () [ + check "match" match, + check "match_hints" match_hints + ] + end +\ + +subparagraph \Ground Hints\ + +ML_command\ + let + val ctxt = Proof_Context.set_mode Proof_Context.mode_schematic @{context} + val hints = map (Skip_Proof.make_thm (Proof_Context.theory_of ctxt) o Syntax.read_term ctxt) [ + "?x \ 2 \ ?x + (-?x :: int) \ 0", + "?x \ 0 \ ?y \ 0 \ (?x :: int) + ?y \ 0" + ] + val tests = map (apply2 (Syntax.read_term ctxt))[ + ("(2 + -2) + (0 + 0) :: int", "0 :: int") + ] + val check_hints = check_unit_tests_hints_match tests + in + Lecker.test_group ctxt () [ + check_hints false [] "match" match, + check_hints false [] "match_hints without hints" match_hints, + check_hints true hints "match_hints with hints" match_hints + ] + end +\ + + +subsubsection \Unification\ +paragraph \Generated Tests\ + +ML_command\ + structure Test_Params = + struct + val unify = unify + val unify_hints = unify_hints + val params = { + nv = 4, + ni = 2, + max_h = 5, + max_args = 4 + } + end + structure First_Order_Tests = First_Order_Unification_Tests(Test_Params) + val _ = First_Order_Tests.tests @{context} (SpecCheck_Random.new ()) +\ + +paragraph \Unit Tests\ +subparagraph \Occurs-Check\ + +ML_command\ + let + val unit_tests = [ + ( + Var (("x", 0), TVar (("X", 0), [])), + Var (("y", 0), TVar (("X", 0), []) --> TFree ("'a", [])) $ + Var (("x", 0), TVar (("X", 0), [])) + ), + ( + Var (("y", 0), TVar (("X", 0), []) --> TFree ("'a", [])) $ + Var (("x", 0), TVar (("X", 0), [])), + Var (("x", 0), TVar (("X", 0), [])) + ), + ( + Free (("f", [TVar (("X", 0), []), TVar (("X", 0), [])] ---> TFree ("'a", []))) $ + Var (("x", 0), TVar (("X", 0), [])) $ + Var (("y", 0), TVar (("X", 0), [])), + Free (("f", [TVar (("X", 0), []), TVar (("X", 0), [])] ---> TFree ("'a", []))) $ + Var (("y", 0), TVar (("X", 0), [])) $ ( + Free (("g", TVar (("X", 0), []) --> TVar (("X", 0), []))) $ + Var (("x", 0), TVar (("X", 0), [])) + ) + ), + ( + Free (("f", [TVar (("X", 0), []), TVar (("Y", 0), [])] ---> TFree ("'a", []))) $ + Var (("x", 0), TVar (("X", 0), [])) $ + Var (("y", 0), TVar (("Y", 0), [])), + Free (("f", [TVar (("Y", 0), []), TVar (("X", 0), [])] ---> TFree ("'a", []))) $ + Var (("y", 0), TVar (("Y", 0), [])) $ ( + Free (("g", TVar (("X", 0), []) --> TVar (("X", 0), []))) $ + Var (("x", 0), TVar (("X", 0), [])) + ) + ) + ] + fun check name unif ctxt _ = + check_list unit_tests ("occurs-check for " ^ name) + (Prop.prop (not o terms_unify_thms_correct_unif ctxt unif)) ctxt + |> K () + in + Lecker.test_group @{context} () [ + check "unify" unify, + check "unify_hints" unify_hints + ] + end +\ + +subparagraph \Lambda-Abstractions\ + +ML_command\ + let + val ctxt = Proof_Context.set_mode Proof_Context.mode_schematic @{context} + |> Variable.declare_term @{term "f :: nat \ bool \ nat"} + val tests = map (apply2 (Syntax.read_term ctxt))[ + ("\ (x :: nat). (0 :: nat)", "\ (x :: nat). (0 :: nat)"), + ("\ (x :: nat). x", "\ (x :: nat). x"), + ("\ (x :: nat) (y :: bool). (x, y)", "\ (x :: nat) (y :: bool). (x, y)"), + ("\ (x :: ?'a) y. f y", "\(x :: ?'b). f"), + ("\ (x :: nat) (y :: bool). f x y", "\ (x :: nat) (y :: bool). (?x :: nat \ bool \ ?'Z) x y"), + ("\x. (?f :: nat \ bool \ nat) 0 x", "\x. f ?g x"), + ("(?f :: nat \ bool \ nat) (?n :: nat)", "?g :: bool \ nat") + ] + val check = check_unit_tests_hints_unif tests true [] + in + Lecker.test_group ctxt () [ + check "unify" unify, + check "unify_hints" unify_hints + ] + end +\ + +ML_command\ + let + val ctxt = Proof_Context.set_mode Proof_Context.mode_schematic @{context} + val hints = map (Skip_Proof.make_thm (Proof_Context.theory_of ctxt) o Syntax.read_term ctxt) [ + "?x \ (0 :: nat) \ ?y \ ?z \ ?x + ?y \ ?z" + ] + val tests = map (apply2 (Syntax.read_term ctxt))[ + ("\ (x :: nat) (y :: bool). x", "\ (x :: nat) (y :: bool). 0 + x"), + ("0 + (\ (x :: nat) (y :: bool). x) 0 ?y", "0 + (\ (x :: nat) (y :: bool). 0 + x) 0 ?z") + ] + val check_hints = check_unit_tests_hints_unif tests + in + Lecker.test_group ctxt () [ + check_hints false [] "unify" unify, + check_hints false [] "unify_hints without hints" unify_hints, + check_hints true hints "unify_hints with hints" unify_hints + ] + end +\ + + +end diff --git a/thys/ML_Unification/Tests/Higher_Order_ML_Unification_Tests.thy b/thys/ML_Unification/Tests/Higher_Order_ML_Unification_Tests.thy new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/Tests/Higher_Order_ML_Unification_Tests.thy @@ -0,0 +1,71 @@ +\<^marker>\creator "Kevin Kappelmann"\ +subsection \Higher-Order ML Unification Tests\ +theory Higher_Order_ML_Unification_Tests + imports + Higher_Order_Pattern_ML_Unification_Tests +begin + +paragraph \Summary\ +text \Tests for @{ML_structure "Higher_Order_Unification"}.\ + +ML\ + structure Prop = SpecCheck_Property + open Unification_Tests_Base + structure Unif = Higher_Order_Unification + val unify = Unif.unify [] +\ + +subsubsection \Unification\ + +paragraph \Generated Tests\ +subparagraph \First-Order\ + +ML_command\ + structure Test_Params = + struct + val unify = unify + (*TODO: there is no higher-order unification with hints as of now; + we hence use the higher-order pattern hint unifier for those tests*) + val unify_hints = unify_hints + val params = { + nv = 10, + ni = 10, + max_h = 5, + max_args = 4 + } + end + structure First_Order_Tests = First_Order_Unification_Tests(Test_Params) + val _ = First_Order_Tests.tests @{context} (SpecCheck_Random.new ()) +\ + +subparagraph \Higher-Order Patterns\ + +ML_file\higher_order_pattern_unification_tests.ML\ + +ML_command\ + val ctxt = @{context} + val tests = Higher_Order_Pattern_Unification_Tests.unit_tests_unifiable ctxt + val check_hints = check_unit_tests_hints_unif tests + val _ = Lecker.test_group ctxt () [ + check_hints true [] "unify" unify + ] +\ + +paragraph \Unit Tests\ + +ML_command\ + let + val ctxt = Proof_Context.set_mode Proof_Context.mode_schematic @{context} + val tests = map (apply2 (Syntax.read_term ctxt)) [ + ("?X ?Y", "?Y ?Z"), + ("?X ?Y (?K ?M)", "f c (?L ?N)") + ] + val check_hints = check_unit_tests_hints_unif tests + in + Lecker.test_group ctxt () [ + check_hints true [] "unify" unify + ] + end +\ + +end diff --git a/thys/ML_Unification/Tests/Higher_Order_Pattern_ML_Unification_Tests.thy b/thys/ML_Unification/Tests/Higher_Order_Pattern_ML_Unification_Tests.thy new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/Tests/Higher_Order_Pattern_ML_Unification_Tests.thy @@ -0,0 +1,186 @@ +\<^marker>\creator "Kevin Kappelmann"\ +\<^marker>\contributor "Paul Bachmann"\ +subsection \Higher-Order Pattern ML Unification Tests\ +theory Higher_Order_Pattern_ML_Unification_Tests + imports + ML_Unification_Tests_Base +begin + +paragraph \Summary\ +text \Tests for @{ML_structure "Higher_Order_Pattern_Unification"}.\ + +ML\ + structure Prop = SpecCheck_Property + structure UC = Unification_Combinator + open Unification_Tests_Base + structure Unif = Higher_Order_Pattern_Unification + val match = Unif.match [] + val match_hints = + let fun match binders = + UC.add_fallback_matcher + (fn match_theory => Unif.e_match Unification_Util.match_types match_theory match_theory) + ((fn binders => + (Hints.map_retrieval (Hints.mk_retrieval Hints.TI.generalisations |> K) + #> Hints.UH.map_concl_unifier (Unif.match |> K) + #> Hints.UH.map_normalisers (Unif.norms_match |> K) + #> Hints.UH.map_prems_unifier (match |> K)) + |> Context.proof_map + #> Test_Unification_Hints.try_hints binders) + |> UC.norm_matcher (#norm_term Unif.norms_match)) + binders + in match [] end + + val unify = Unif.unify [] + val unify_hints = + let fun unif binders = + UC.add_fallback_unifier + (fn unif_theory => Unif.e_unify Unification_Util.unify_types unif_theory unif_theory) + ((fn binders => + (Hints.UH.map_concl_unifier (Unif.match |> K) + #> Hints.UH.map_normalisers (Unif.norms_unify |> K) + #> Hints.UH.map_prems_unifier (unif |> K)) + |> Context.proof_map + #> Test_Unification_Hints.try_hints binders) + |> UC.norm_unifier (#norm_term Unif.norms_unify)) + binders + in unif [] end +\ + +subsubsection \Matching\ +paragraph \Unit Tests\ + +ML_command\ + let + val ctxt = Proof_Context.set_mode Proof_Context.mode_schematic @{context} + |> Variable.declare_term @{term "f :: nat \ bool \ nat"} + val tests = map (apply2 (Syntax.read_term ctxt))[ + ("?x :: nat", "(?x + 1 :: nat)"), + ("(g :: nat \ nat \ nat) ?x ?x", "(g :: nat \ nat \ nat) (?x + 1) (?x + 1)"), + ("\y. (?x :: nat \ ?'Z) y", "\y. f y"), + ("(g :: ?'X \ ?'Y \ ?'Z) (x :: ?'X) (y :: ?'Y)", "(g :: ?'Y \ ?'Z \ ?'Z) (x :: ?'Y) (y :: ?'Z)"), + ("(g :: ?'Z \ ?'X)", "\(x :: ?'X). (g :: ?'X \ ?'Y) x"), + ("\ (x :: nat). (0 :: nat)", "\ (x :: nat). (0 :: nat)"), + ("\ (x :: nat). x", "\ (x :: nat). x"), + ("\ (x :: nat) (y :: bool). (x, y)", "\ (x :: nat) (y :: bool). (x, y)"), + ("\ (x :: ?'A) (y :: bool). (?x :: ?'A \ bool \ ?'Z) x y", "\ (x :: nat) (y :: bool). f x y"), + ("\(x :: ?'X). (g :: ?'X \ ?'X) x", "(g :: ?'X \ ?'X)"), + ("\y. (?x :: nat \ bool \ nat) y", "\y. f y"), + ("\y z. (?x :: nat \ bool \ nat) y z", "f") + ] + val check = check_unit_tests_hints_match tests true [] + in + Lecker.test_group ctxt () [ + check "match" match, + check "match_hints" match_hints + ] + end +\ + +subparagraph \Asymmetry\ + +ML_command\ + let + val ctxt = Proof_Context.set_mode Proof_Context.mode_schematic @{context} + |> Variable.declare_term @{term "f :: nat \ bool \ nat"} + val tests = map (apply2 (Syntax.read_term ctxt))[ + ("\y. f y", "\y. (?x :: nat \ bool \ nat) y") + ] + val check = check_unit_tests_hints_match tests false [] + in + Lecker.test_group ctxt () [ + check "match" match, + check "match_hints" match_hints + ] + end +\ + +subparagraph \Ground Hints\ + +ML_command\ + let + val ctxt = Proof_Context.set_mode Proof_Context.mode_schematic @{context} + val hints = map (Skip_Proof.make_thm (Proof_Context.theory_of ctxt) o Syntax.read_term ctxt) [ + "?x \ 2 \ ?x + (-?x :: int) \ 0", + "?x \ 0 \ ?y \ 0 \ (?x :: int) + ?y \ 0" + ] + val tests = map (apply2 (Syntax.read_term ctxt))[ + ("(2 + -2) + (0 + 0) :: int", "0 :: int") + ] + val check_hints = check_unit_tests_hints_match tests + in + Lecker.test_group ctxt () [ + check_hints false [] "match" match, + check_hints false [] "match_hints without hints" match_hints, + check_hints true hints "match_hints with hints" match_hints + ] + end +\ + + +subsubsection \Unification\ + +paragraph \Generated Tests\ +subparagraph \First-Order\ + +ML_command\ + structure Test_Params = + struct + val unify = unify + val unify_hints = unify_hints + val params = { + nv = 10, + ni = 10, + max_h = 5, + max_args = 4 + } + end + structure First_Order_Tests = First_Order_Unification_Tests(Test_Params) + val _ = First_Order_Tests.tests @{context} (SpecCheck_Random.new ()) +\ + +subparagraph \Higher-Order Pattern\ + +ML_file\higher_order_pattern_unification_tests.ML\ + +ML_command\ + val ctxt = @{context} + val tests = Higher_Order_Pattern_Unification_Tests.unit_tests_unifiable ctxt + val check_hints = check_unit_tests_hints_unif tests + val _ = Lecker.test_group ctxt () [ + check_hints true [] "unify" unify, + check_hints true [] "unify_hints" unify_hints + ] +\ + +paragraph \Unit Tests\ +subparagraph \With Unification Hints\ + +ML_command\ + let + val ctxt = Proof_Context.set_mode Proof_Context.mode_schematic @{context} + |> Variable.declare_term @{term "f :: (nat \ nat) \ nat \ nat"} + |> Variable.declare_term @{term "g :: nat \ nat \ nat"} + |> Variable.declare_term @{term "h :: nat"} + val hints = map (Skip_Proof.make_thm (Proof_Context.theory_of ctxt) o Syntax.read_term ctxt) [ + "?x \ (0 :: nat) \ ?y \ ?z \ ?x + ?y \ ?z", + "(?x :: ?'X) \ (?y :: ?'X) \ id ?x \ ?y", + "(?x1 :: nat) \ ?x2 \ (?y1 :: nat) \ ?y2 \ ?x1 + ?y1 \ ?y2 + ?x2" + ] + val tests = map (apply2 (Syntax.read_term ctxt)) [ + ("\ x y. 0 + 1 :: nat", "\ x y. 1 :: nat"), + ("\ x. 0 + 0 + x :: nat", "\ x. x :: nat"), + ("\ x y. 0 + Suc ?x", "\ x y. Suc 2"), + ("\ (x :: nat) (y :: nat). y + (0 + x)", "\ (x :: nat) (y :: nat). x + (0 + y)"), + ("f (\ u. g (?x, h), h)", "id (f (\ u. ?y, 0 + h))") + ] + val check_hints = check_unit_tests_hints_unif tests + in + Lecker.test_group ctxt () [ + check_hints false [] "unify" unify, + check_hints false [] "unify_hints without hints" unify_hints, + check_hints true hints "unify_hints with hints" unify_hints + ] + end +\ + +end diff --git a/thys/ML_Unification/Tests/ML_Unification_Tests.thy b/thys/ML_Unification/Tests/ML_Unification_Tests.thy new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/Tests/ML_Unification_Tests.thy @@ -0,0 +1,13 @@ +\<^marker>\creator "Kevin Kappelmann"\ +section \ML-Unification Tests\ +theory ML_Unification_Tests + imports + First_Order_ML_Unification_Tests + Higher_Order_ML_Unification_Tests + Higher_Order_Pattern_ML_Unification_Tests +begin + +paragraph \Summary\ +text \Tests for ML unifiers.\ + +end \ No newline at end of file diff --git a/thys/ML_Unification/Tests/ML_Unification_Tests_Base.thy b/thys/ML_Unification/Tests/ML_Unification_Tests_Base.thy new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/Tests/ML_Unification_Tests_Base.thy @@ -0,0 +1,33 @@ +\<^marker>\creator "Kevin Kappelmann"\ +subsection \Test Setup\ +theory ML_Unification_Tests_Base + imports + ML_Unification_Hints + SpecCheck.SpecCheck + Main +begin + +paragraph \Summary\ +text \Shared setup for unification tests. We use \<^cite>\speccheck\ to generate +tests and create unit tests.\ + +ML\ + @{functor_instance struct_name = Test_Unification_Hints + and functor_name = Term_Index_Unification_Hints + and id = \"test"\ + and more_args = \ + structure TI = Discrimination_Tree + val init_args = { + concl_unifier = SOME Higher_Order_Pattern_Unification.match, + normalisers = SOME Unification_Util.beta_eta_short_norms_unif, + prems_unifier = SOME Higher_Order_Pattern_Unification.unify, + retrieval = SOME (Term_Index_Unification_Hints_Args.mk_sym_retrieval + TI.norm_term TI.generalisations), + hint_preprocessor = SOME (K I) + }\} +\ + +ML_file \tests_base.ML\ +ML_file \first_order_unification_tests.ML\ + +end \ No newline at end of file diff --git a/thys/ML_Unification/Tests/first_order_unification_tests.ML b/thys/ML_Unification/Tests/first_order_unification_tests.ML new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/Tests/first_order_unification_tests.ML @@ -0,0 +1,369 @@ +(* Title: ML_Unification/first_order_unification_tests.ML + Author: Kevin Kappelmann + +First-order unification tests. +*) +signature FIRST_ORDER_UNIFICATION_TESTS = +sig + val tests_identical : Proof.context Lecker.test + val tests_identical_env : Proof.context Lecker.test + val tests_thm : Proof.context Lecker.test + val tests_replaced : Proof.context Lecker.test + val tests_symmetry : Proof.context Lecker.test + val generated_tests : Proof.context Lecker.test + + val unit_tests_unifiable : Proof.context Lecker.test + val unit_tests_non_unifiable : Proof.context Lecker.test + + val unit_tests_hints_non_recursive : Proof.context Lecker.test + val unit_tests_multiple_matching_hints : Proof.context Lecker.test + val unit_tests_hints_recursive : Proof.context Lecker.test + val unit_tests_multiple_successful_hints : Proof.context Lecker.test + val unit_tests_hints : Proof.context Lecker.test + + val unit_tests : Proof.context Lecker.test + + val tests : Proof.context Lecker.test +end + +functor First_Order_Unification_Tests(P : sig + val unify : Unification_Base.closed_unifier + val unify_hints : Unification_Base.closed_unifier + val params : { + nv : int, + ni : int, + max_h : int, + max_args : int + } + end) : FIRST_ORDER_UNIFICATION_TESTS = +struct + +structure Prop = SpecCheck_Property +structure Gen = SpecCheck_Generator +structure UUtil = Unification_Util +open Unification_Tests_Base + +val weights = (1, 1, 1, 0) + +fun term_gen ctxt = + let val {nv, ni, max_h, max_args} = P.params + in term_gen' ctxt nv ni weights max_h max_args end + +fun term_pair_gen ctxt = + let val {nv, ni, max_h, max_args} = P.params + in Unification_Tests_Base.term_pair_gen' ctxt nv ni weights max_h max_args end + +(* standard unification *) +(** generated tests **) +(*** unification of identical terms ***) +fun tests_identical ctxt r = + let fun check_identical name unif ctxt = + let val term_gen' = term_gen ctxt + in + check_thm_unif (Gen.map (fn t => (t, t)) term_gen') + ("Unifying identical terms with " ^ name) unif ctxt + end + in + Lecker.test_group ctxt r [ + check_identical "unify" P.unify, + check_identical "unify_hints" P.unify_hints + ] + end + +fun tests_identical_env ctxt r = + let + fun check_identical_env name unif ctxt = + let val term_gen' = term_gen ctxt + in + check (Gen.map (fn t => (t, t)) term_gen') + ("Unifying identical terms does not change environment with " ^ name) + (Prop.==> (terms_unify ctxt unif, + (fn tp => unif ctxt tp (UUtil.empty_envir tp) + |> Seq.list_of + |> List.all (Envir.is_empty o fst)))) + ctxt + end + in + Lecker.test_group ctxt r [ + check_identical_env "unify" P.unify, + check_identical_env "unify_hints" P.unify_hints + ] + end + +(*** unification of randomly generated terms ***) +fun tests_thm ctxt r = + let + fun check_thm' name unif ctxt = + let val term_pair_gen' = term_pair_gen ctxt + in check_thm_unif term_pair_gen' name unif ctxt end + val ctxt' = Config.put SpecCheck_Configuration.max_discard_ratio 100 ctxt + in + Lecker.test_group ctxt' r [ + check_thm' "unify" P.unify, + check_thm' "unify_hints" P.unify_hints + ] + end + +fun tests_replaced ctxt r = + let fun check_replaced name unif ctxt = + let + val term_gen' = term_gen ctxt + val unvarify = map_aterms (fn (Var ((n, i), T)) => Free (n ^ Int.toString i, T) | T => T) + in + check_thm_unif (Gen.map (fn t => (t, unvarify t)) term_gen') + ("Unifying terms with Var replaced by Free for " ^ name) unif ctxt + end + in + Lecker.test_group ctxt r [ + check_replaced "unify" P.unify, + check_replaced "unify_hints" P.unify_hints + ] + end + +(**** symmetry of success ****) +fun tests_symmetry ctxt r = + let fun check_thm_symmetry name unif ctxt = + let + val term_pair_gen' = term_pair_gen ctxt + val terms_unify_thms_correct' = terms_unify_thms_correct_unif ctxt unif + in + check term_pair_gen' ("Symmetry of " ^ name) + (Prop.prop (fn tp => + terms_unify_thms_correct' tp = terms_unify_thms_correct' (swap tp))) + ctxt + end + in + Lecker.test_group ctxt r [ + check_thm_symmetry "unify" P.unify, + check_thm_symmetry "unify_hints" P.unify_hints + ] + end + +fun generated_tests ctxt r = + Lecker.test_group ctxt r [ + tests_identical, + tests_identical_env, + tests_thm, + tests_replaced, + tests_symmetry + ] + +(** Unit tests **) +(*** unifiable ***) +fun unit_tests_unifiable ctxt r = + let + val tests = [ + (Var (("x", 0), TVar (("X", 0), [])), Var (("x", 0), TVar (("Y", 0), []))), + (Var (("x", 0), TVar (("X", 0), [])), Var (("x", 0), TFree ("'a", []))), + (Var (("x", 0), TVar (("X", 0), [])), Var (("y", 0), TVar (("X", 0), []))), + (Var (("x", 0), TVar (("X", 0), [])), Free ("c", TFree ("'a", []))), + (Var (("x", 0), TFree ("'a", [])), Free ("f", TVar (("X", 0), []))), + (Var (("x", 0), TVar (("X", 0), [])), Free ("f", TFree ("'a", []))), + (Var (("x", 0), TFree ("'a", [])), Free ("f", TFree ("'a", []))), + ( + Free ("x", TFree ("'a", []) --> TVar (("X", 0), [])) $ + Free ("x", TVar (("Y", 0), [])), + Free ("x", TVar (("Y", 0), []) --> TFree ("'b", [])) $ + Free ("x", TFree ("'a", [])) + ), + ( + Free ("f", TVar (("X", 0), []) --> TVar (("Y", 0), [])) $ + Free ("c", TVar (("X", 0), [])), + Free ("f", TVar (("X", 0), []) --> TVar (("Y", 0), [])) $ + Var (("x", 0), TVar (("X", 1), [])) + ), + ( + Free ("f", [TVar (("A", 0), []), TVar (("B", 0), [])] ---> TFree ("c", [])) $ + Free ("a", TVar (("A", 0), [])) $ + Free ("b", TVar (("B", 0), [])), + Free ("f", [TVar (("B", 0), []), TVar (("A", 0), [])] ---> TFree ("c", [])) $ + Free ("a", TVar (("B", 0), [])) $ + Free ("b", TVar (("A", 0), [])) + ) + ] + fun check name unif ctxt r = + check_list tests ("unifiable unit tests for " ^ name) + (Prop.prop (terms_unify_thms_correct_unif ctxt unif)) ctxt + |> K r + in + Lecker.test_group ctxt r [ + check "unify" P.unify, + check "unify_hints" P.unify_hints + ] + end + +(*** non-unifiable ***) +fun unit_tests_non_unifiable ctxt r = + let + val unit_tests = [ + (Free("f", TFree("'a",[])), Free("f", TFree("'b",[]))), + (Free("f", TFree("'a",[])), Free("g", TFree("'a",[]))) + ] + fun check name unif ctxt r = + check_list unit_tests ("non-unifiable unit tests for " ^ name) + (Prop.prop (not o terms_unify ctxt unif)) ctxt + |> K r + in + Lecker.test_group ctxt r [ + check "unify" P.unify, + check "unify_hints" P.unify_hints + ] + end + +(* hint tests *) +(** non-recursive tests **) +fun unit_tests_hints_non_recursive ctxt r = + let + val ctxt' = Proof_Context.set_mode Proof_Context.mode_schematic ctxt + val hints = map (Skip_Proof.make_thm (Proof_Context.theory_of ctxt') o Syntax.read_term ctxt') [ + "?z \ ?x \ ?y \ (0 :: nat) \ ?z \ ?x + ?y", + "?x \ 1 \ ?y \ 1 \ (1 :: nat) \ ?x * ?y" + ] + val tests = map (apply2 (Syntax.read_term ctxt')) [ + ("1 :: nat", "1 + 0 :: nat"), + ("1 :: nat", "?x + 0 :: nat"), + ("1 :: nat", "?a * ?b :: nat"), + ("\(x :: nat) y. x", "\(x :: nat) y. x + 0") + ] + fun check_hints should_succeed hints name = + check_unit_tests_hints_unif tests should_succeed hints ("non-recursive hint unit tests for " ^ name) + in + Lecker.test_group ctxt' r [ + check_hints false [] "unify" P.unify, + check_hints false [] "unify_hints without hints" P.unify_hints, + check_hints true hints "unify_hints with hints" P.unify_hints + ] + end + +(** multiple matching hints **) +fun unit_tests_multiple_matching_hints ctxt r = + let + val ctxt' = Proof_Context.set_mode Proof_Context.mode_schematic ctxt + val hints = map (Skip_Proof.make_thm (Proof_Context.theory_of ctxt') o Syntax.read_term ctxt') [ + "Suc ?x \ 1", + "?x \ 0 \ Suc ?x \ 1" + ] + val tests = map (apply2 (Syntax.read_term ctxt')) [ + ("Suc x :: nat", "1 :: nat") + ] + fun check_hints should_succeed hints name = + check_unit_tests_hints_unif tests should_succeed hints + ("multiple matching hints unit tests for " ^ name) + in + Lecker.test_group ctxt' r [ + check_hints false [] "unify" P.unify, + check_hints false [] "unify_hints without hints" P.unify_hints, + check_hints false (tl hints) "unify_hints with wrong hints" P.unify_hints, + check_hints true [hd hints] "unify_hints with correct hints" P.unify_hints, + check_hints true hints "unify_hints with all hints" P.unify_hints, + check_hints true (rev hints) "unify_hints with reversed hints order" P.unify_hints + ] + end + +(** recursive hints **) +fun unit_tests_hints_recursive ctxt r = + let + val ctxt' = + Proof_Context.set_mode Proof_Context.mode_schematic ctxt + |> Variable.declare_term @{term "f :: nat => nat"} + |> Variable.declare_term @{term "g :: nat => nat"} + val hints = map (Skip_Proof.make_thm (Proof_Context.theory_of ctxt') o Syntax.read_term ctxt') [ + "?x \ ?z \ ?y \ (0 :: nat) \ ?x + ?y \ ?z", + "?y \ ?x + ?z \ ?x + (Suc ?z) \ Suc ?y", + "?x \ f (g 0) \ ?y \ g (f 0) \ f (g ?x) \ g (f ?y)", + "?y \ f ?x \ ?x \ f (g 0) \ f (g ?x) \ g (f ?y)" + ] + val tests = map (apply2 (Syntax.read_term ctxt')) [ + ("1 :: nat", "(?x + 0) + 0 :: nat"), + ("1 :: nat", "?x + (0 + 0) :: nat"), + ("x + (Suc 0) :: nat", "Suc x :: nat"), + ("f (g (f (g 0)))", "g (f (g (f 0)))"), + ("f (g (f ((g 0) + 0)))", "g (f (f (f (g 0))))") + ] + fun check_hints should_succeed hints name = + check_unit_tests_hints_unif tests should_succeed hints ("recursive hints unit tests for " ^ name) + in + Lecker.test_group ctxt' r [ + check_hints false [] "unify" P.unify, + check_hints false [] "unify_hints without hints" P.unify_hints, + check_hints true hints "unify_hints with hints" P.unify_hints + ] + end + +(** disabling symmetric application of hints **) +fun unit_tests_hints_try_symmetric ctxt r = + let + open Unification_Tests_Base + val ctxt' = Proof_Context.set_mode Proof_Context.mode_schematic ctxt + |> Context.proof_map (Hints.map_retrieval + (Hints.mk_retrieval Hints.TI.generalisations |> K)) + val hints = map (Skip_Proof.make_thm (Proof_Context.theory_of ctxt') o Syntax.read_term ctxt') [ + "?x \ ?z \ ?y \ (0 :: nat) \ ?x + ?y \ ?z" + ] + val (tests_correct, tests_incorrect) = map (apply2 (apply2 (Syntax.read_term ctxt'))) [ + (("?x + 0 :: nat", "1 :: nat"), ("1 :: nat", "?x + 0 :: nat")) + ] |> split_list + fun check_hints should_succeed tests name = + check_unit_tests_hints_unif tests should_succeed hints ("no symmetric hints for " ^ name) + in + Lecker.test_group ctxt' r [ + check_hints true tests_correct "matching problem" P.unify_hints, + check_hints false tests_incorrect "symmetric problem" P.unify_hints + ] + end + +(** multiple successful hints **) +fun unit_tests_multiple_successful_hints ctxt r = + let + val ctxt' = Proof_Context.set_mode Proof_Context.mode_schematic ctxt + val hints = map (Skip_Proof.make_thm (Proof_Context.theory_of ctxt') o Syntax.read_term ctxt') [ + "?x \ ?z \ ?y \ (0 :: nat) \ ?x + ?y \ ?z", + "?x \ (?z + 0) \ ?y \ (0 :: nat) \ ?x + ?y \ ?z" + ] + val tests = map (apply2 (Syntax.read_term ctxt')) [ + ("(?x + 0) + 0 :: nat", "1 :: nat") + ] + fun check_num_success num_success hints name unif ctxt r = + let val ctxt' = add_hints hints ctxt + in + check_list tests ("multiple successful hints unit tests for " ^ name) + (SpecCheck_Property.prop (fn tp => + (UUtil.empty_envir tp + |> unif ctxt' tp + |> Seq.list_of + |> length) + = num_success)) + ctxt' + |> K r + end + in + Lecker.test_group ctxt' r [ + check_num_success 0 [] "unify" P.unify, + check_num_success 0 [] "unify_hints without hints" P.unify_hints, + check_num_success 3 hints "unify_hints with hints" P.unify_hints + ] + end + +fun unit_tests_hints ctxt r = + Lecker.test_group ctxt r [ + unit_tests_hints_non_recursive, + unit_tests_multiple_matching_hints, + unit_tests_hints_try_symmetric, + unit_tests_hints_recursive, + unit_tests_multiple_successful_hints + ] + +fun unit_tests ctxt r = + Lecker.test_group ctxt r [ + unit_tests_unifiable, + unit_tests_non_unifiable, + unit_tests_hints + ] + +fun tests ctxt r = + Lecker.test_group ctxt r [ + unit_tests, + generated_tests + ] + +end diff --git a/thys/ML_Unification/Tests/higher_order_pattern_unification_tests.ML b/thys/ML_Unification/Tests/higher_order_pattern_unification_tests.ML new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/Tests/higher_order_pattern_unification_tests.ML @@ -0,0 +1,29 @@ +(* Title: ML_Unification/higher_order_pattern_unification_tests.ML + Author: Kevin Kappelmann + +Higher-order pattern unification tests. +*) +signature HIGHER_ORDER_PATTERN_UNIFICATION_TESTS = +sig + val unit_tests_unifiable : Proof.context -> (term * term) list +end + +structure Higher_Order_Pattern_Unification_Tests : HIGHER_ORDER_PATTERN_UNIFICATION_TESTS = +struct + +fun unit_tests_unifiable ctxt = + let val ctxt = Proof_Context.set_mode Proof_Context.mode_schematic ctxt + in map (apply2 (Syntax.read_term ctxt)) [ + ("\ x. f x", "\ x. f x"), + ("\ (x :: ?'X). (f :: ?'X \ ?'Y) x", "\ (x :: ?'X1). (?y :: ?'X1 \ ?'Y1) x"), + ("\ x. r x ?X", "\ x. r x ?Y"), + ("\ x. (x, (\ y. (y, \ z. ?x)))", "\ x. (x, (\ y. (y, \ z. g)))"), + ("?f :: ?'Z \ ?'X \ ?'Y \ ?'R", "\ (x :: ?'Z). (?f :: ?'Z \ ?'X1 \ ?'Y1 \ ?'R1) x"), + ( + "(?x :: ?'X, ?y :: ?'Y, ?z :: ?'Z)", + "((f :: ?'Y \ ?'X) (?y :: ?'Y), (g :: ?'Z \ ?'Y) (?z :: ?'Z), c :: ?'C)" + ) + ] + end + +end diff --git a/thys/ML_Unification/Tests/tests_base.ML b/thys/ML_Unification/Tests/tests_base.ML new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/Tests/tests_base.ML @@ -0,0 +1,192 @@ +(* Title: ML_Unification/tests_base.ML + Author: Kevin Kappelmann + +Basic setup for ML unifier tests. + +TODO: generalise for unifiers with open terms; currently, all tests work with closed terms. +*) +signature UNIFICATION_TESTS_BASE = +sig + + structure Hints : TERM_INDEX_UNIFICATION_HINTS + val add_hint : Unification_Hints_Base.unif_hint -> Context.generic -> Context.generic + val add_hints : thm list -> Proof.context -> Proof.context + + val term_gen : Proof.context -> int -> int -> int * int * int * int -> + (int -> int -> int SpecCheck_Generator.gen) -> term SpecCheck_Generator.gen + val term_pair_gen : Proof.context -> int -> int -> int * int * int * int -> + (int -> int -> int SpecCheck_Generator.gen) -> + (term * term) SpecCheck_Generator.gen + + val term_gen' : Proof.context -> int -> int -> int * int * int * int -> int -> + int -> term SpecCheck_Generator.gen + val term_pair_gen' : Proof.context -> int -> int -> int * int * int * int -> + int -> int -> (term * term) SpecCheck_Generator.gen + + val thms_correct : (Envir.env -> (term * term) -> (term * term)) -> + Envir_Normalisation.thm_normaliser -> Proof.context -> Unification_Base.closed_unifier -> + (term * term) SpecCheck_Property.pred + + val terms_unify : Proof.context -> Unification_Base.closed_unifier -> + (term * term) SpecCheck_Property.pred + + val terms_unify_thms_correct : (Envir.env -> (term * term) -> (term * term)) -> + Envir_Normalisation.thm_normaliser -> Proof.context -> Unification_Base.closed_unifier -> + (term * term) SpecCheck_Property.pred + + val terms_unify_thms_correct_unif : Proof.context -> Unification_Base.closed_unifier -> + (term * term) SpecCheck_Property.pred + + val terms_unify_thms_correct_match : Proof.context -> Unification_Base.closed_unifier -> + (term * term) SpecCheck_Property.pred + + val check : (term * term, 'a) SpecCheck_Generator.gen_state -> string -> + (term * term) SpecCheck_Property.prop -> + (Proof.context, 'a) Lecker.test_state + + val check_thm : (Envir.env -> (term * term) -> (term * term)) -> Envir_Normalisation.thm_normaliser -> + (term * term, 'a) SpecCheck_Generator.gen_state -> string -> + Unification_Base.closed_unifier -> (Proof.context, 'a) Lecker.test_state + + val check_thm_unif : (term * term, 'a) SpecCheck_Generator.gen_state -> string -> + Unification_Base.closed_unifier -> (Proof.context, 'a) Lecker.test_state + + val check_thm_match : (term * term, 'a) SpecCheck_Generator.gen_state -> string -> + Unification_Base.closed_unifier -> (Proof.context, 'a) Lecker.test_state + + val check_list : (term * term) list -> string -> + (term * term) SpecCheck_Property.prop -> Proof.context -> + (term * term) Seq.seq + + val check_unit_tests_hints : (Envir.env -> (term * term) -> (term * term)) -> + Envir_Normalisation.thm_normaliser -> (term * term) list -> bool -> thm list -> string -> + Unification_Base.closed_unifier -> (Proof.context, 'a) Lecker.test_state + + val check_unit_tests_hints_unif : (term * term) list -> bool -> thm list -> string -> + Unification_Base.closed_unifier -> (Proof.context, 'a) Lecker.test_state + + val check_unit_tests_hints_match : (term * term) list -> bool -> thm list -> string -> + Unification_Base.closed_unifier -> (Proof.context, 'a) Lecker.test_state + +end + +structure Unification_Tests_Base : UNIFICATION_TESTS_BASE = +struct + +structure Hints = Test_Unification_Hints + +val add_hint = Test_Unification_Hints.add_hint_prio o rpair Prio.MEDIUM +fun add_hints thms ctxt = fold add_hint thms (Context.Proof ctxt) |> Context.proof_of + +structure Gen = SpecCheck_Generator +structure UUtil = Unification_Util +structure Norm = Envir_Normalisation + +fun term_num_args_gen nv ni weights num_args_gen h i = + Gen.zip (Gen.aterm' (Gen.nonneg nv) (Gen.nonneg ni) weights) (num_args_gen h i) + +fun term_gen ctxt nv ni weights num_args_gen = + let val ctxt' = Proof_Context.set_mode Proof_Context.mode_schematic ctxt + in + Gen.term_tree (term_num_args_gen nv ni weights num_args_gen) + |> Gen.map (try (singleton (Variable.polymorphic ctxt') o Syntax.check_term ctxt')) + |> Gen.filter is_some + |> Gen.map the + end + +fun term_pair_gen ctxt nv ni weights num_args_gen = + let + val ctxt' = Proof_Context.set_mode Proof_Context.mode_schematic ctxt + val term_gen = Gen.term_tree (term_num_args_gen nv ni weights num_args_gen) + in + Gen.zip term_gen term_gen + |> Gen.map (fn (s, t) => try (Variable.polymorphic ctxt' o Syntax.check_terms ctxt') [s, t]) + |> Gen.filter is_some + |> Gen.map (fn SOME [s, t] => (s, t)) + end + +fun num_args_gen max_h max_args h _ = if h > max_h then Gen.return 0 else Gen.nonneg max_args + +fun term_gen' ctxt nv ni weights max_h max_args = + term_gen ctxt nv ni weights (num_args_gen max_h max_args) + +fun term_pair_gen' ctxt nv ni weights max_h max_args = + term_pair_gen ctxt nv ni weights (num_args_gen max_h max_args) + +fun show_termpair ctxt = + let val pretty_term = Syntax.pretty_term ctxt + in SpecCheck_Show.zip pretty_term pretty_term end + +val shrink_termpair = SpecCheck_Shrink.product SpecCheck_Shrink.term SpecCheck_Shrink.none + +fun thms_correct norm_terms norm_thm ctxt unif tp = + let + val env_thmq = unif ctxt tp (UUtil.empty_envir tp) + fun thm_correct (env, thm) = + let + val (t1, t2) = norm_terms env tp + val thm' = norm_thm ctxt env thm + val (lhs, rhs) = Thm.concl_of thm' |> Logic.dest_equals + in List.all Envir.aeconv [(t1, lhs), (t2, rhs)] andalso Thm.no_prems thm' end + in List.all thm_correct (Seq.list_of env_thmq) end + +fun terms_unify ctxt unif tp = + UUtil.empty_envir tp + |> unif ctxt tp + |> General_Util.seq_is_empty + |> not o fst + +fun terms_unify_thms_correct norm_terms norm_thm ctxt unif tp = + terms_unify ctxt unif tp andalso thms_correct norm_terms norm_thm ctxt unif tp + +val terms_unify_thms_correct_unif = terms_unify_thms_correct + (apply2 o Norm.beta_eta_short_norm_term_unif) + (Norm.beta_eta_short_norm_thm Norm.norm_type_unif Norm.beta_eta_short_norm_term_unif) + +val terms_unify_thms_correct_match = terms_unify_thms_correct + (apfst o Norm.beta_norm_term_match) + (*types are already normalised by matcher*) + (Norm.beta_eta_short_norm_thm (K I) Norm.beta_norm_term_match) + +fun check gen name prop ctxt = + SpecCheck.check_shrink (show_termpair ctxt) shrink_termpair + gen name prop ctxt + +fun check_thm norm_terms norm_thm gen name unif ctxt = check gen ("Theorem correctness: " ^ name) + (SpecCheck_Property.==> (terms_unify ctxt unif, thms_correct norm_terms norm_thm ctxt unif)) + ctxt + +fun check_thm_unif norm_terms = check_thm + (apply2 o Norm.beta_eta_short_norm_term_unif) + (Norm.beta_eta_short_norm_thm Norm.norm_type_unif Norm.beta_eta_short_norm_term_unif) + norm_terms + +fun check_thm_match norm_terms = check_thm + (apfst o Norm.beta_norm_term_match) + (Norm.beta_eta_short_norm_thm (K I) Norm.beta_norm_term_match) + norm_terms + +fun check_list tests name prop ctxt = + SpecCheck.check_list (show_termpair ctxt) tests name prop ctxt + +fun check_unit_tests_hints norm_terms norm_thm tests should_succeed hints name unif ctxt s = + let val ctxt' = add_hints hints ctxt + in + check_list tests name + (SpecCheck_Property.prop + (fn tp => terms_unify_thms_correct norm_terms norm_thm ctxt' unif tp = should_succeed)) + ctxt' + |> K s + end + +fun check_unit_tests_hints_unif norm_terms = check_unit_tests_hints + (apply2 o Norm.beta_eta_short_norm_term_unif) + (Norm.beta_eta_short_norm_thm Norm.norm_type_unif Norm.beta_eta_short_norm_term_unif) + norm_terms + +fun check_unit_tests_hints_match norm_terms = check_unit_tests_hints + (apfst o Norm.beta_norm_term_match) + (Norm.beta_eta_short_norm_thm (K I) (K I)) + norm_terms + +end diff --git a/thys/ML_Unification/Unification_Attributes/Unification_Attributes.thy b/thys/ML_Unification/Unification_Attributes/Unification_Attributes.thy new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/Unification_Attributes/Unification_Attributes.thy @@ -0,0 +1,72 @@ +\<^marker>\creator "Kevin Kappelmann"\ +section \Unification Attributes\ +theory Unification_Attributes + imports Unify_Resolve_Tactics +begin + +paragraph \Summary\ +text \OF attribute with adjustable unifier.\ + +ML_file\unify_of_base.ML\ +ML_file\unify_of.ML\ + +ML\ + @{functor_instance struct_name = Standard_Unify_OF + and functor_name = Unify_OF + and id = \""\ + and more_args = \val init_args = { + normalisers = SOME Standard_Mixed_Unification.norms_first_higherp_first_comb_higher_unify, + unifier = SOME Standard_Mixed_Unification.first_higherp_first_comb_higher_unify, + mode = SOME (Unify_OF_Args.PM.key Unify_OF_Args.PM.fact) + }\} +\ +local_setup \Standard_Unify_OF.setup_attribute NONE\ + +paragraph \Examples\ + +experiment +begin +lemma + assumes h1: "(PROP A \ PROP D) \ PROP E \ PROP C" + assumes h2: "PROP B \ PROP D" + and h3: "PROP F \ PROP E" + shows "(PROP A \ PROP B) \ PROP F \ PROP C" + by (fact h1[uOF h2 h3 where mode = resolve]) \\the line below is equivalent\ + (* by (fact h1[OF h2 h3]) *) + +lemma + assumes h1: "(PROP A \ PROP A)" + assumes h2: "(PROP A \ PROP A) \ PROP B" + shows "PROP B" + by (fact h2[uOF h1]) \\the line below is equivalent\ + (* by (fact h2[uOF h1 where mode = fact]) *) + \\Note: @{attribute OF} would not work in this case:\ + (* thm h2[OF h1] *) + +lemma + assumes h1: "\x y z. PROP P x y \ PROP P y y \ (PROP A \ PROP A) \ + (PROP A \ PROP B) \ PROP C" + and h2: "\x y. PROP P x y" + and h3 : "PROP A \ PROP A" + and h4 : "PROP D \ PROP B" + shows "(PROP A \ PROP D) \ PROP C" + by (fact h1[uOF h2 h2 h3, uOF h4 where mode = resolve]) + +lemma + assumes h1: "\P x. PROP P x \ PROP E P x" + and h2: "PROP P x" + shows "PROP E P x" + by (fact h1[uOF h2]) \\the following line does not work (multiple unifiers error)\ + (* by (fact h1[OF h2]) *) + +text\We can also specify the unifier to be used:\ + +lemma + assumes h1: "\P. PROP P \ PROP E" + and h2: "\P. PROP P" + shows "PROP E" + by (fact h1[uOF h2 where unifier = First_Order_Unification.unify]) \\the line below is equivalent\ + (* using [[uOF unifier = First_Order_Unification.unify]] by (fact h1[uOF h2]) *) +end + +end diff --git a/thys/ML_Unification/Unification_Attributes/unify_of.ML b/thys/ML_Unification/Unification_Attributes/unify_of.ML new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/Unification_Attributes/unify_of.ML @@ -0,0 +1,196 @@ +(* Title: ML_Unification/unify_of.ML + Author: Kevin Kappelmann + +OF attribute with arguments from context. +*) +@{parse_entries (sig) PARSE_UNIFY_OF_ARGS + [normalisers, unifier, mode, rules]} +@{parse_entries (sig) PARSE_UNIFY_OF_CONTEXT_ARGS + [normalisers, unifier, mode]} + +@{parse_entries (sig) PARSE_UNIFY_OF_MODES [resolve, fact]} + +signature UNIFY_OF_ARGS = +sig + structure PA : PARSE_UNIFY_OF_ARGS + structure PCA : PARSE_UNIFY_OF_CONTEXT_ARGS + + val PCA_entries_from_PA_entries : ('a, 'b, 'c, 'd) PA.entries -> ('a, 'b, 'c) PCA.entries + val PA_entries_from_PCA_entries : ('a, 'b, 'c) PCA.entries -> 'd -> ('a, 'b, 'c, 'd) PA.entries + + structure PM : PARSE_UNIFY_OF_MODES + type mode = PM.key + val parse_mode : mode parser + + type args = (Unification_Base.normalisers, Unification_Base.unifier, mode, thm list) + PA.entries + type context_args = (Unification_Base.normalisers, Unification_Base.unifier, mode) + PCA.entries + + val unify_OF : args -> Proof.context -> thm -> thm option + val unify_OF_context_args : context_args -> thm list -> Proof.context -> thm -> thm option + + val unify_OF_attribute : args -> attribute + val unify_OF_context_args_attribute : context_args -> thm list -> attribute + + val arg_parsers : (ML_Code_Util.code parser, ML_Code_Util.code parser, mode parser, + thm list context_parser) PA.entries +end + +structure Unify_OF_Args : UNIFY_OF_ARGS = +struct + +structure URB = Unify_Resolve_Base + +@{parse_entries (struct) PA + [normalisers, unifier, mode, rules]} +@{parse_entries (struct) PCA + [normalisers, unifier, mode]} + +fun PCA_entries_from_PA_entries {normalisers = normalisers, unifier = unifier, mode = mode,...} = + {normalisers = normalisers, unifier = unifier, mode = mode} +fun PA_entries_from_PCA_entries {normalisers = normalisers, unifier = unifier, mode = mode} rules = + {normalisers = normalisers, unifier = unifier, mode = mode, rules = SOME rules} + +@{parse_entries (struct) PM [resolve, fact]} +type mode = PM.key +val parse_mode = PM.parse_key + +type args = (Unification_Base.normalisers, Unification_Base.unifier, mode, thm list) + PA.entries +type context_args = (Unification_Base.normalisers, Unification_Base.unifier, mode) + PCA.entries + +fun gen_unify_of unify_OF_fn args = + let val resolve_fn = case PA.get_mode args of + PM.resolve _ => URB.unify_resolve_tac + | PM.fact _ => URB.unify_resolve_atomic_tac + in unify_OF_fn resolve_fn (PA.get_normalisers args) (PA.get_unifier args) (PA.get_rules args) end + +val unify_OF = gen_unify_of Unify_OF_Base.unify_OF +val unify_OF_context_args = unify_OF oo PA_entries_from_PCA_entries + +val unify_OF_attribute = gen_unify_of Unify_OF_Base.unify_OF_attribute +val unify_OF_context_args_attribute = unify_OF_attribute oo PA_entries_from_PCA_entries + +val arg_parsers = { + normalisers = SOME Unification_Parser.parse_normalisers, + unifier = SOME Unification_Parser.parse_unifier, + mode = SOME parse_mode, + rules = SOME Parse_Util.thms +} + +end + +signature UNIFY_OF = +sig + structure Data : GENERIC_DATA + + val get_args : Context.generic -> Unify_OF_Args.context_args + val map_args : (Unify_OF_Args.context_args -> Unify_OF_Args.context_args) -> + Context.generic -> Context.generic + + val get_normalisers : Context.generic -> Unification_Base.normalisers + val map_normalisers : (Unification_Base.normalisers -> Unification_Base.normalisers) -> + Context.generic -> Context.generic + + val get_unifier : Context.generic -> Unification_Base.unifier + val map_unifier : (Unification_Base.unifier -> Unification_Base.unifier) -> + Context.generic -> Context.generic + + val get_mode : Context.generic -> Unify_OF_Args.mode + val map_mode : (Unify_OF_Args.mode -> Unify_OF_Args.mode) -> + Context.generic -> Context.generic + + val unify_OF : thm list -> Proof.context -> thm -> thm option + val unify_OF_attribute : thm list -> attribute + + val binding : binding + + val set_args_attribute : (ML_Code_Util.code, ML_Code_Util.code, Unify_OF_Args.mode) + Unify_OF_Args.PCA.entries * Position.T -> attribute + val parse_set_args_attribute : attribute parser + + val parse_attribute : attribute context_parser + val setup_attribute : string option -> local_theory -> local_theory +end + +functor Unify_OF(A : sig + structure FIA : FUNCTOR_INSTANCE_ARGS + val init_args : Unify_OF_Args.context_args + end) : UNIFY_OF = +struct + +structure UOA = Unify_OF_Args +structure PCA = Unify_OF_Args.PCA +structure PA = Unify_OF_Args.PA +structure FIU = Functor_Instance_Util(A.FIA) +structure MCU = ML_Code_Util + +structure Data = Generic_Data(struct + type T = UOA.context_args + val empty = A.init_args + val merge = fst +end) + +val get_args = Data.get +val map_args = Data.map + +val get_normalisers = PCA.get_normalisers o get_args +val map_normalisers = map_args o PCA.map_normalisers + +val get_unifier = PCA.get_unifier o get_args +val map_unifier = map_args o PCA.map_unifier + +val get_mode = PCA.get_mode o get_args +val map_mode = map_args o PCA.map_mode + +fun unify_OF rules context = + UOA.unify_OF_context_args (get_args (Context.Proof context)) rules context +fun unify_OF_attribute rules (context, thm) = + UOA.unify_OF_context_args_attribute (get_args context) rules (context, thm) + +val binding = FIU.mk_binding_id_prefix "uOF" + +val parse_arg_entries = + let + val parsers = UOA.arg_parsers |> UOA.PCA_entries_from_PA_entries + val parse_value = PCA.parse_entry (PCA.get_normalisers parsers) + (PCA.get_unifier parsers) (PCA.get_mode parsers) + val parse_entry = Parse_Key_Value.parse_entry PCA.parse_key Parse_Util.eq parse_value + in PCA.parse_entries_required Parse.and_list1 [] parse_entry (PCA.empty_entries ()) end + +fun set_args_attribute (entries, pos) = + let + fun UOA_substructure_op substructure operation = + MCU.flat_read ["Unify_OF_Args.", substructure, ".", operation] + val (code_PM_op, code_PCA_op) = apply2 UOA_substructure_op ("PM", "PCA") + val code_from_key = code_PCA_op o PCA.key_to_string + fun code_from_entry (PCA.normalisers c) = c + | code_from_entry (PCA.unifier c) = c + | code_from_entry (PCA.mode m) = code_PM_op "key" @ code_PM_op (UOA.PM.key_to_string m) + val code_entries = PCA.key_entry_entries_from_entries entries + |> map (fn (k, v) => code_from_key k @ MCU.atomic (code_from_entry v)) + |> MCU.list + val code = + FIU.code_struct_op "map_args" @ MCU.atomic (code_PCA_op "merge_entries" @ + MCU.atomic (code_PCA_op "entries_from_entry_list" @ code_entries)) + in ML_Attribute.run_map_context (code, pos) end + +val parse_set_args_attribute = (parse_arg_entries |> Parse_Util.position) >> set_args_attribute + +val parse_attribute = + PA.get_rules UOA.arg_parsers + :-- (fn [] => (parse_set_args_attribute |> Parse.!!!) >> SOME |> Scan.lift + | _ => (Parse.where_ |-- parse_set_args_attribute) |> Parse.!!! |> Parse_Util.option |> Scan.lift) + (*if there are no rules, we update the arguments in the context*) + >> (fn ([], SOME attr) => attr + (*otherwise, we use the arguments but do not store them in the context*) + | (rules, opt_attr) => + apfst (the_default I (Option.map ML_Attribute_Util.attribute_map_context opt_attr)) + #> unify_OF_attribute rules) + +val setup_attribute = Attrib.local_setup binding parse_attribute o + the_default ("apply theorem to arguments (" ^ FIU.FIA.full_name ^ ")") + +end \ No newline at end of file diff --git a/thys/ML_Unification/Unification_Attributes/unify_of_base.ML b/thys/ML_Unification/Unification_Attributes/unify_of_base.ML new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/Unification_Attributes/unify_of_base.ML @@ -0,0 +1,25 @@ +(* Title: ML_Unification/unify_of_base.ML + Author: Kevin Kappelmann + +OF attribute with adjustable unifier. +*) +signature UNIFY_OF_BASE = +sig + val unify_OF : ('a -> 'b -> thm -> Proof.context -> int -> tactic) -> + 'a -> 'b -> thm list -> Proof.context -> thm -> thm option + val unify_OF_attribute : ('a -> 'b -> thm -> Proof.context -> int -> tactic) -> + 'a -> 'b -> thm list -> attribute +end + +structure Unify_OF_Base : UNIFY_OF_BASE = +struct + +fun unify_OF f norms unify thms ctxt = + map (fn thm => f norms unify thm ctxt) thms |> RANGE |> HEADGOAL |> SINGLE + +fun unify_OF_attribute f norms unify thms = Thm.rule_attribute [] + (fn context => fn thm => + unify_OF f norms unify thms (Context.proof_of context) thm + |> (fn SOME thm => thm | _ => raise THM ("unify_OF: no unifiers", 0, thm :: thms))) + +end \ No newline at end of file diff --git a/thys/ML_Unification/Unification_Hints/ML_Unification_Hints.thy b/thys/ML_Unification/Unification_Hints/ML_Unification_Hints.thy new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/Unification_Hints/ML_Unification_Hints.thy @@ -0,0 +1,55 @@ +\<^marker>\creator "Kevin Kappelmann"\ +section \Unification Hints\ +theory ML_Unification_Hints + imports + ML_Generic_Data_Utils + ML_Term_Index + ML_Unifiers + ML_Unification_Parsers +begin + +paragraph \Summary\ +text \A generalisation of unification hints, originally introduced in \<^cite>\"unif-hints"\. +We support a generalisation that +\<^enum> allows additional universal variables in premises +\<^enum> allows non-atomic left-hand sides for premises +\<^enum> allows arbitrary functions to perform the matching/unification of a hint with a disagreement pair. + +General shape of a hint: +\\y1...yn. (\x1...xn1. lhs1 \ rhs1) \ ... \ (\x1...xnk. lhsk \ rhsk) \ lhs \ rhs\ +\ + +ML_file\unification_hints_base.ML\ +ML_file\unification_hints.ML\ +ML_file\term_index_unification_hints.ML\ + +ML\ + @{functor_instance struct_name = Standard_Unification_Hints + and functor_name = Term_Index_Unification_Hints + and id = \""\ + and more_args = \ + structure TI = Discrimination_Tree + val init_args = { + concl_unifier = SOME Higher_Ordern_Pattern_First_Decomp_Unification.unify, + normalisers = SOME Standard_Mixed_Unification.norms_first_higherp_first_comb_higher_unify, + prems_unifier = SOME (Standard_Mixed_Unification.first_higherp_first_comb_higher_unify + |> Unification_Combinator.norm_unifier Envir_Normalisation.beta_norm_term_unif), + retrieval = SOME (Term_Index_Unification_Hints_Args.mk_sym_retrieval + TI.norm_term TI.unifiables), + hint_preprocessor = SOME (K I) + }\} +\ +local_setup \Standard_Unification_Hints.setup_attribute NONE\ + +text\Standard unification hints are accessible via @{attribute uhint}.\ + +declare [[ucombine add = \Standard_Unification_Combine.eunif_data + (Standard_Unification_Hints.try_hints + |> Unification_Combinator.norm_unifier + (#norm_term Standard_Mixed_Unification.norms_first_higherp_first_comb_higher_unify) + |> K) + (Standard_Unification_Combine.default_metadata Standard_Unification_Hints.binding)\]] + +text\Examples see @{dir "../Examples"}.\ + +end diff --git a/thys/ML_Unification/Unification_Hints/named_theorems_unification_hints.ML b/thys/ML_Unification/Unification_Hints/named_theorems_unification_hints.ML new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/Unification_Hints/named_theorems_unification_hints.ML @@ -0,0 +1,39 @@ +(* Title: ML_Unification/named_theorems_unification_hints.ML + Author: Kevin Kappelmann + +Unification hints stored in a named theorem collection (NAMED_THMS). +*) +@{parse_entries (sig) PARSE_NAMED_THEOREMS_UNIFICATION_HINTS_ARGS + [concl_unifier, normaliser, prems_unifier, try_sym, add, del, set]} + +signature NAMED_THEOREMS_UNIFICATION_HINTS = +sig + structure Hints : NAMED_THMS + structure UHI : UNIFICATION_HINTS_INDEXED +end + +functor Named_Theorems_Unification_Hints(A : sig structure FIA : FUNCTOR_INSTANCE_ARGS end) + : NAMED_THEOREMS_UNIFICATION_HINTS = +struct + +structure FIU = Functor_Instance_Util(A.FIA) + +structure Hints = Named_Thms( + val name = FIU.mk_binding_id_prefix "unif_hint" + val description = "Unification Hints (" ^ FIU.FIA.full_name ^ ")" +) +val _ = Theory.setup Hints.setup + +structure UHI : UNIFICATION_HINT_INDEX = +struct +fun get_hints get_symmetric ctxt _ = Hints.get ctxt |> Seq.of_list |> + get_symmetric ? + (fn hintsq => Seq.interleave (hintsq, Seq.map Unification_Hints.symmetric_hint hintsq)) +end + +structure UHI = Unification_Hints_Indexed(struct + structure FIA = FIU.FIA + structure UHI = UHI +end) + +end diff --git a/thys/ML_Unification/Unification_Hints/term_index_unification_hints.ML b/thys/ML_Unification/Unification_Hints/term_index_unification_hints.ML new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/Unification_Hints/term_index_unification_hints.ML @@ -0,0 +1,381 @@ +(* Title: ML_Unification/term_index_unification_hints.ML + Author: Kevin Kappelmann + +Unification hints stored in a term index. +*) +@{parse_entries (sig) PARSE_TERM_INDEX_UNIFICATION_HINTS_MODE + [add, del]} + +@{parse_entries (sig) PARSE_TERM_INDEX_UNIFICATION_HINTS_ARGS + [concl_unifier, normalisers, prems_unifier, retrieval, hint_preprocessor, prio]} + +@{parse_entries (sig) PARSE_TERM_INDEX_UNIFICATION_HINTS_CONTEXT_ARGS + [concl_unifier, normalisers, prems_unifier, retrieval, hint_preprocessor]} + +signature TERM_INDEX_UNIFICATION_HINTS_ARGS = +sig + structure PA : PARSE_TERM_INDEX_UNIFICATION_HINTS_ARGS + structure PCA : PARSE_TERM_INDEX_UNIFICATION_HINTS_CONTEXT_ARGS + + val PCA_entries_from_PA_entries : ('a, 'b, 'c, 'd, 'e, 'f) PA.entries -> + ('a, 'b, 'c, 'd, 'e) PCA.entries + val PA_entries_from_PCA_entries : ('a, 'b, 'c, 'd, 'e) PCA.entries -> 'f -> + ('a, 'b, 'c, 'd, 'e, 'f) PA.entries + val UHA_PA_entries_from_PCA_entries : ('a, 'b, 'c, 'd, 'e) PCA.entries -> + ('a, 'b, 'c) Unification_Hints_Args.PA.entries + + structure PM : PARSE_TERM_INDEX_UNIFICATION_HINTS_MODE + type mode = PM.key + val parse_mode : mode parser + + type hint_prio = Unification_Hints_Base.unif_hint * Prio.prio + val pretty_hint_prio : Proof.context -> hint_prio -> Pretty.T + + val sort_hint_prios : hint_prio list -> hint_prio list + + val mk_retrieval : Term_Normalisation.term_normaliser -> ('ti -> term -> hint_prio list) -> + 'ti -> Unification_Hints_Base.hint_retrieval + val mk_sym_retrieval : Term_Normalisation.term_normaliser -> ('ti -> term -> hint_prio list) -> + 'ti -> Unification_Hints_Base.hint_retrieval + + type 'ti args = (Unification_Base.unifier, Unification_Base.normalisers, Unification_Base.unifier, + 'ti -> Unification_Hints_Base.hint_retrieval, Unification_Hints_Base.hint_preprocessor, + Prio.prio) PA.entries + type 'ti context_args = (Unification_Base.unifier, Unification_Base.normalisers, + Unification_Base.unifier, 'ti -> Unification_Hints_Base.hint_retrieval, + Unification_Hints_Base.hint_preprocessor) PCA.entries + + val arg_parsers : (ML_Code_Util.code parser, ML_Code_Util.code parser, ML_Code_Util.code parser, + ML_Code_Util.code parser, ML_Code_Util.code parser, Prio.prio context_parser) PA.entries +end + +structure Term_Index_Unification_Hints_Args : TERM_INDEX_UNIFICATION_HINTS_ARGS = +struct + +structure UB = Unification_Base +structure EN = Envir_Normalisation +structure UHB = Unification_Hints_Base +structure UHA = Unification_Hints_Args +structure TUHP = Prio + +@{parse_entries (struct) PA + [concl_unifier, normalisers, prems_unifier, retrieval, hint_preprocessor, prio]} + +@{parse_entries (struct) PCA + [concl_unifier, normalisers, prems_unifier, retrieval, hint_preprocessor]} + +fun PCA_entries_from_PA_entries {concl_unifier = concl_unifier, normalisers = normalisers, + prems_unifier = prems_unifier, retrieval = retrieval, hint_preprocessor = hint_preprocessor,...} = + {concl_unifier = concl_unifier, normalisers = normalisers, prems_unifier = prems_unifier, + retrieval = retrieval, hint_preprocessor = hint_preprocessor} +fun PA_entries_from_PCA_entries {concl_unifier = concl_unifier, normalisers = normalisers, + prems_unifier = prems_unifier, retrieval = retrieval, hint_preprocessor = hint_preprocessor} prio = + {concl_unifier = concl_unifier, normalisers = normalisers, prems_unifier = prems_unifier, + retrieval = retrieval, hint_preprocessor = hint_preprocessor, prio = SOME prio} + +fun UHA_PA_entries_from_PCA_entries {concl_unifier = concl_unifier, normalisers = normalisers, + prems_unifier = prems_unifier,...} = + {concl_unifier = concl_unifier, normalisers = normalisers, prems_unifier = prems_unifier} + +@{parse_entries (struct) PM + [add, del]} +type mode = PM.key +val parse_mode = PM.parse_key + +type hint_prio = UHB.unif_hint * TUHP.prio + +fun pretty_hint_prio ctxt (hint, prio) = Pretty.block [ + UHB.pretty_hint ctxt hint, + Pretty.enclose " (" ")" [Pretty.str "priority: ", Prio.pretty prio] + ] + +(*FIXME: use Prio.Table instead of sorted lists*) +val sort_hint_prios = sort (Prio.ord o apply2 snd) + +val hint_seq_of_hint_prios = + sort_hint_prios + #> Seq.of_list + #> Seq.map fst + +fun mk_retrieval norm_term retrieve ti _ _ (t, _) _ = + norm_term t + |> retrieve ti + |> hint_seq_of_hint_prios + +fun interleave [] ys = ys + | interleave xs [] = xs + | interleave (x :: xs) (y :: ys) = x :: y :: interleave xs ys + +fun mk_sym_retrieval norm_term retrieve ti _ _ (t1, t2) _ = + interleave + (norm_term t1 |> retrieve ti) + (norm_term t2 |> retrieve ti |> map (apfst Unification_Hints_Base.symmetric_hint)) + |> hint_seq_of_hint_prios + +type 'ti args = (UB.unifier, UB.normalisers, UB.unifier, 'ti -> UHB.hint_retrieval, + UHB.hint_preprocessor, TUHP.prio) PA.entries +type 'ti context_args = (UB.unifier, UB.normalisers, UB.unifier, 'ti -> UHB.hint_retrieval, + UHB.hint_preprocessor) PCA.entries + +val arg_parsers = { + concl_unifier = UHA.PA.get_concl_unifier_safe UHA.arg_parsers, + normalisers = UHA.PA.get_normalisers_safe UHA.arg_parsers, + prems_unifier = UHA.PA.get_prems_unifier_safe UHA.arg_parsers, + retrieval = SOME (Parse_Util.nonempty_code (K "retrieval function must not be empty")), + hint_preprocessor = SOME (Parse_Util.nonempty_code (K "hint preprocessor must not be empty")), + prio = SOME Prio.parse + } + +end + +signature TERM_INDEX_UNIFICATION_HINTS = +sig + include HAS_LOGGER + + structure UH : UNIFICATION_HINTS + + (*underlying term index*) + structure TI : TERM_INDEX + structure Data : GENERIC_DATA + + val get_retrieval : Context.generic -> Term_Index_Unification_Hints_Args.hint_prio TI.term_index -> + Unification_Hints_Base.hint_retrieval + val map_retrieval : ( + (Term_Index_Unification_Hints_Args.hint_prio TI.term_index -> Unification_Hints_Base.hint_retrieval) -> + Term_Index_Unification_Hints_Args.hint_prio TI.term_index -> Unification_Hints_Base.hint_retrieval) -> + Context.generic -> Context.generic + + val get_hint_preprocessor : Context.generic -> Unification_Hints_Base.hint_preprocessor + val map_hint_preprocessor : (Unification_Hints_Base.hint_preprocessor -> + Unification_Hints_Base.hint_preprocessor) -> Context.generic -> Context.generic + + val get_index : Context.generic -> Term_Index_Unification_Hints_Args.hint_prio TI.term_index + val map_index : (Term_Index_Unification_Hints_Args.hint_prio TI.term_index -> + Term_Index_Unification_Hints_Args.hint_prio TI.term_index) -> Context.generic -> Context.generic + val pretty_index : Proof.context -> Pretty.T + + val add_hint_prio : Term_Index_Unification_Hints_Args.hint_prio -> Context.generic -> + Context.generic + + val del_hint : Unification_Hints_Base.unif_hint -> Context.generic -> Context.generic + + val mk_retrieval : Term_Index_Unification_Hints_Args.hint_prio TI.retrieval -> + Term_Index_Unification_Hints_Args.hint_prio TI.term_index -> + Unification_Hints_Base.hint_retrieval + val mk_sym_retrieval : Term_Index_Unification_Hints_Args.hint_prio TI.retrieval -> + Term_Index_Unification_Hints_Args.hint_prio TI.term_index -> + Unification_Hints_Base.hint_retrieval + + val try_hints : Unification_Base.unifier + + val binding : binding + + val attribute : (Term_Index_Unification_Hints_Args.mode * + ((ML_Code_Util.code, ML_Code_Util.code, ML_Code_Util.code, + ML_Code_Util.code, ML_Code_Util.code, Prio.prio) + Term_Index_Unification_Hints_Args.PA.entries)) * + Position.T -> attribute + val parse_attribute : attribute context_parser + val setup_attribute : string option -> local_theory -> local_theory +end + +functor Term_Index_Unification_Hints(A : sig + structure FIA : FUNCTOR_INSTANCE_ARGS + structure TI : TERM_INDEX + val init_args : (Term_Index_Unification_Hints_Args.hint_prio TI.term_index) + Term_Index_Unification_Hints_Args.context_args + end) : TERM_INDEX_UNIFICATION_HINTS = +struct + +structure UHB = Unification_Hints_Base +structure TUHA = Term_Index_Unification_Hints_Args +structure TUHP = Prio +structure PA = TUHA.PA +structure PCA = TUHA.PCA +structure PM = TUHA.PM + +structure FIU = Functor_Instance_Util(A.FIA) +structure TI = A.TI + +structure MCU = ML_Code_Util +structure PU = Parse_Util + +val logger = Logger.setup_new_logger Unification_Hints_Base.logger FIU.FIA.full_name + +@{functor_instance struct_name = UH + and functor_name = Unification_Hints + and accessor = FIU.accessor + and id = FIU.FIA.id + and more_args = \val init_args = TUHA.UHA_PA_entries_from_PCA_entries A.init_args\} + +fun are_hint_variants ctxt hp = + let + val tp = apply2 Thm.prop_of hp + val env = Unification_Util.empty_envir tp + fun match tp = First_Order_Unification.match [] ctxt tp env + in + match tp + |> Seq.maps (fn _ => match (swap tp)) + |> not o fst o General_Util.seq_is_empty + end + +(* fun are_hint_prio_variants ctxt ((h1, p1), (h2, p2)) = + p1 = p2 andalso are_hint_variants ctxt (h1, h2) *) + +structure Data = Generic_Data(Pair_Generic_Data_Args( +struct + structure Data1 = Pair_Generic_Data_Args( + struct + structure Data1 = + struct + type T = TUHA.hint_prio TI.term_index -> Unification_Hints_Base.hint_retrieval + val empty = PCA.get_retrieval A.init_args + val merge = fst + end + structure Data2 = + struct + type T = UHB.hint_preprocessor + val empty = PCA.get_hint_preprocessor A.init_args + val merge = fst + end + end) + structure Data2 = Term_Index_Generic_Data_Args( + struct + type data = TUHA.hint_prio + structure TI = TI + fun data_eq ((h1, _), (h2, _))= are_hint_variants (Context.the_local_context ()) (h1, h2) + end) +end)) + +val get_retrieval = fst o fst o Data.get +val map_retrieval = Data.map o apfst o apfst + +val get_hint_preprocessor = snd o fst o Data.get +val map_hint_preprocessor = Data.map o apfst o apsnd + +val get_index = snd o Data.get +val map_index = Data.map o apsnd +fun pretty_index ctxt = get_index (Context.Proof ctxt) + |> TI.content + |> TUHA.sort_hint_prios + |> map (TUHA.pretty_hint_prio ctxt) + |> Pretty.fbreaks + |> Pretty.block + +val term_index_key_from_hint = UHB.cdest_hint_concl #> fst #> Thm.term_of #> TI.norm_term + +fun msg_illegal_hint_format ctxt hint = Pretty.block [ + Pretty.str "Illegal hint format for ", + UHB.pretty_hint ctxt hint + ] |> Pretty.string_of + +fun add_hint_prio (hint, prio) context = + let + val ctxt = Context.proof_of context + val hint = hint |> get_hint_preprocessor context ctxt + val is_hint_variant = curry (are_hint_variants ctxt) hint o fst + in + (@{log Logger.DEBUG} ctxt (fn _ => Pretty.block [ + Pretty.str "Adding hint ", + UHB.pretty_hint ctxt hint + ] |> Pretty.string_of); + TI.insert is_hint_variant (term_index_key_from_hint hint, (hint, prio)) (get_index context) + |> (fn ti => map_index (K ti) context)) + handle Term_Index_Base.INSERT => + (@{log Logger.WARN} ctxt + (fn _ => Pretty.block [ + Pretty.str "Hint ", + UHB.pretty_hint ctxt hint, + Pretty.str " already added." + ] |> Pretty.string_of); + context) + | TERM _ => (@{log Logger.WARN} ctxt (fn _ => msg_illegal_hint_format ctxt hint); + context) + end + +fun del_hint hint context = + let + val ctxt = Context.proof_of context + val hint = hint |> get_hint_preprocessor context ctxt + val is_hint_variant = curry (are_hint_variants ctxt) hint o fst + in + (@{log Logger.DEBUG} ctxt (fn _ => Pretty.block [ + Pretty.str "Deleting hint ", + UHB.pretty_hint ctxt hint + ] |> Pretty.string_of); + TI.delete is_hint_variant (term_index_key_from_hint hint) (get_index context) + |> (fn ti => map_index (K ti) context)) + handle Term_Index_Base.DELETE => + (@{log Logger.WARN} ctxt + (fn _ => Pretty.block [ + Pretty.str "Hint ", + UHB.pretty_hint ctxt hint, + Pretty.str " not found." + ] |> Pretty.string_of); + context) + | TERM _ => (@{log Logger.WARN} ctxt (fn _ => msg_illegal_hint_format ctxt hint); + context) + end + +val mk_retrieval = TUHA.mk_retrieval TI.norm_term +val mk_sym_retrieval = TUHA.mk_sym_retrieval TI.norm_term + +fun try_hints binders ctxt = + let val context = Context.Proof ctxt + in UH.try_hints (get_retrieval context (get_index context)) binders ctxt end + +val binding = FIU.mk_binding_id_prefix "uhint" + +val default_entries = PA.empty_entries () |> fold PA.set_entry [PA.prio TUHP.MEDIUM] + +fun parse_arg_entries mode = + let + val parsers = TUHA.arg_parsers + val parse_value = PA.parse_entry (PA.get_concl_unifier parsers |> Scan.lift) + (PA.get_normalisers parsers |> Scan.lift) (PA.get_prems_unifier parsers |> Scan.lift) + (PA.get_retrieval parsers |> Scan.lift) (PA.get_hint_preprocessor parsers |> Scan.lift) + (PA.get_prio parsers) + val parse_entry = Parse_Key_Value.parse_entry' (Scan.lift PA.parse_key) + (Scan.lift PU.eq) parse_value + in + PA.parse_entries_required' PU.and_list1' [] parse_entry default_entries + |> mode = PM.key PM.del ? + PU.filter_cut' (is_none o PA.get_prio_safe) (K o K "Priority may not be specified for deletion.") + end + +fun attribute ((mode, entries), pos) = + let + fun update_index thm = Term_Util.no_dummy_pattern (Thm.prop_of thm) ? ( + case mode of + PM.add _ => add_hint_prio (thm, PA.get_prio entries) + | PM.del _ => del_hint thm) + val PCA_entries = TUHA.PCA_entries_from_PA_entries entries + val UHA_PA_entries = TUHA.UHA_PA_entries_from_PCA_entries PCA_entries + val run_code = ML_Attribute.run_map_context o rpair pos + fun default_code (context, _) = (SOME context, NONE) + val map_retrieval = case PCA.get_retrieval_safe PCA_entries of + SOME c => FIU.code_struct_op "map_retrieval" @ MCU.atomic (MCU.read "K" @ MCU.atomic c) + |> run_code + | NONE => default_code + val map_hint_preprocessor = case PCA.get_hint_preprocessor_safe PCA_entries of + SOME c => FIU.code_struct_op "map_hint_preprocessor" @ MCU.atomic (MCU.read "K" @ MCU.atomic c) + |> run_code + | NONE => default_code + in + ML_Attribute_Util.apply_attribute (Thm.declaration_attribute update_index) + #> ML_Attribute_Util.apply_attribute (UH.attribute (UHA_PA_entries, pos)) + #> ML_Attribute_Util.apply_attribute map_retrieval + #> map_hint_preprocessor + end + +val parse_attribute = + (Scan.lift (Scan.optional TUHA.parse_mode (PM.key PM.add)) + :-- (fn mode => PU.optional' ((Scan.lift Parse.where_ |-- parse_arg_entries mode) |> PU.!!!+) + default_entries) + |> PU.position') + >> attribute + +val setup_attribute = Attrib.local_setup binding (PU.!!!+ parse_attribute) o + the_default ("set unification hints arguments (" ^ FIU.FIA.full_name ^ ")") + +end diff --git a/thys/ML_Unification/Unification_Hints/unification_hints.ML b/thys/ML_Unification/Unification_Hints/unification_hints.ML new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/Unification_Hints/unification_hints.ML @@ -0,0 +1,119 @@ +(* Title: ML_Unification/unification_hints.ML + Author: Kevin Kappelmann + +Unification hints with arguments from context. +*) +@{parse_entries (sig) PARSE_UNIFICATION_HINTS_ARGS + [concl_unifier, normalisers, prems_unifier]} + +signature UNIFICATION_HINTS_ARGS = +sig + structure PA : PARSE_UNIFICATION_HINTS_ARGS + + type args = (Unification_Base.unifier, Unification_Base.normalisers, Unification_Base.unifier) + PA.entries + + val try_hints : args -> Unification_Hints_Base.hint_retrieval -> Unification_Base.unifier + + val arg_parsers : (ML_Code_Util.code parser, ML_Code_Util.code parser, ML_Code_Util.code parser) + PA.entries +end + +structure Unification_Hints_Args : UNIFICATION_HINTS_ARGS = +struct + +structure UB = Unification_Base +structure EN = Envir_Normalisation +structure UHB = Unification_Hints_Base + +@{parse_entries (struct) PA + [concl_unifier, normalisers, prems_unifier]} + +type args = (UB.unifier, UB.normalisers, UB.unifier) PA.entries + +fun try_hints args retrieval = + UHB.try_hints retrieval (PA.get_concl_unifier args) (PA.get_normalisers args) + (PA.get_prems_unifier args) + +val arg_parsers = { + concl_unifier = SOME Unification_Parser.parse_unifier, + normalisers = SOME Unification_Parser.parse_normalisers, + prems_unifier = SOME Unification_Parser.parse_unifier +} + +end + +signature UNIFICATION_HINTS = +sig + structure Data : GENERIC_DATA + + val get_args : Context.generic -> Unification_Hints_Args.args + val map_args : (Unification_Hints_Args.args -> Unification_Hints_Args.args) -> + Context.generic -> Context.generic + + val get_concl_unifier : Context.generic -> Unification_Base.unifier + val map_concl_unifier : (Unification_Base.unifier -> Unification_Base.unifier) -> + Context.generic -> Context.generic + + val get_normalisers : Context.generic -> Unification_Base.normalisers + val map_normalisers : (Unification_Base.normalisers -> Unification_Base.normalisers) -> + Context.generic -> Context.generic + + val get_prems_unifier : Context.generic -> Unification_Base.unifier + val map_prems_unifier : (Unification_Base.unifier -> Unification_Base.unifier) -> + Context.generic -> Context.generic + + val try_hints : Unification_Hints_Base.hint_retrieval -> Unification_Base.unifier + + val attribute : (ML_Code_Util.code, ML_Code_Util.code, ML_Code_Util.code) + Unification_Hints_Args.PA.entries * Position.T -> attribute +end + +functor Unification_Hints(A : sig + structure FIA : FUNCTOR_INSTANCE_ARGS + val init_args : Unification_Hints_Args.args + end) : UNIFICATION_HINTS = +struct + +structure UHA = Unification_Hints_Args +structure PA = UHA.PA +structure FIU = Functor_Instance_Util(A.FIA) +structure MCU = ML_Code_Util + +structure Data = Generic_Data(struct + type T = UHA.args + val empty = A.init_args + val merge = fst +end) + +val get_args = Data.get +val map_args = Data.map + +val get_concl_unifier = PA.get_concl_unifier o get_args +val map_concl_unifier = map_args o PA.map_concl_unifier + +val get_normalisers = PA.get_normalisers o get_args +val map_normalisers = map_args o PA.map_normalisers + +val get_prems_unifier = PA.get_prems_unifier o get_args +val map_prems_unifier = map_args o PA.map_prems_unifier + +fun try_hints retrieval binders ctxt = + UHA.try_hints (get_args (Context.Proof ctxt)) retrieval binders ctxt + +fun attribute (entries, pos) = + let + fun code_PA_op operation = MCU.flat_read ["Unification_Hints_Args.PA.", operation] + val code_from_key = code_PA_op o PA.key_to_string + fun code_from_entry (PA.concl_unifier c) = c + | code_from_entry (PA.normalisers c) = c + | code_from_entry (PA.prems_unifier c) = c + val code_entries = PA.key_entry_entries_from_entries entries + |> map (fn (k, v) => code_from_key k @ MCU.atomic (code_from_entry v)) + |> MCU.list + val code = + FIU.code_struct_op "map_args" @ MCU.atomic (code_PA_op "merge_entries" @ + MCU.atomic (code_PA_op "entries_from_entry_list" @ code_entries)) + in ML_Attribute.run_map_context (code, pos) end + +end diff --git a/thys/ML_Unification/Unification_Hints/unification_hints_base.ML b/thys/ML_Unification/Unification_Hints/unification_hints_base.ML new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/Unification_Hints/unification_hints_base.ML @@ -0,0 +1,145 @@ +(* Title: ML_Unification/unification_hints_base.ML + Author: Kevin Kappelmann, Paul Bachmann + +Unification hints (introduced in "Hints in unification" by Asperti et al, 2009). +We support a generalisation that +1. allows additional universal variables in premises +2. allows non-atomic left-hand sides for premises +3. allows arbitrary functions to perform the matching/unification of a hint with + a disagreement pair. + +General shape of a hint: +\\y1...yn. (\x1...xn1. lhs1 \ rhs1) \ ... \ (\x1...xnk. lhsk \ rhsk) \ lhs \ rhs\ +*) +signature UNIFICATION_HINTS_BASE = +sig + include HAS_LOGGER + + type unif_hint = thm + val cdest_hint : unif_hint -> (cterm * cterm) list * (cterm * cterm) + val cdest_hint_concl : unif_hint -> cterm * cterm + val symmetric_hint : unif_hint -> unif_hint + val pretty_hint : Proof.context -> unif_hint -> Pretty.T + + type hint_preprocessor = Proof.context -> thm -> unif_hint + val obj_logic_hint_preprocessor : thm -> conv -> hint_preprocessor + + val try_hint : Unification_Base.matcher -> Unification_Base.normalisers -> unif_hint -> + Unification_Base.e_unifier + + type hint_retrieval = term Binders.binders -> Proof.context -> term * term -> Envir.env -> + unif_hint Seq.seq + + val try_hints : hint_retrieval -> Unification_Base.matcher -> Unification_Base.normalisers -> + Unification_Base.e_unifier +end + +structure Unification_Hints_Base : UNIFICATION_HINTS_BASE = +struct + +val logger = Logger.setup_new_logger Unification_Base.logger "Unification_Hints_Base" + +structure GUtil = General_Util +structure TUtil = Term_Util +structure CUtil = Conversion_Util + +type unif_hint = thm + +val cdest_hint = Thm.cprop_of #> TUtil.strip_cimp #>> map Thm.dest_equals ##> Thm.dest_equals +val cdest_hint_concl = Thm.cconcl_of #> Thm.dest_equals + +(*only flips the conclusion since unification calls for the premises should be symmetric*) +val symmetric_hint = Conv.concl_conv ~1 CUtil.symmetric_conv |> CUtil.thm_conv +val pretty_hint = Thm.pretty_thm + +type hint_preprocessor = Proof.context -> thm -> unif_hint +fun obj_logic_hint_preprocessor eq_eq_meta_eq default_conv ctxt = + let + fun eq_conv ct = (if can Thm.dest_equals ct then Conv.all_conv + else Conv.rewr_conv eq_eq_meta_eq else_conv default_conv) ct + val forall_eq_conv = Conversion_Util.repeat_forall_conv (K o K eq_conv) + in Conversion_Util.imp_conv (forall_eq_conv ctxt) eq_conv |> Conversion_Util.thm_conv end + +(*Tries to apply a hint to solve E-unification of (t1 \? t2). +Vars in hint are lifted wrt. the passed binders. +Unifies the hint's conclusion with (t1, t2) using match. +Unifies resulting unification problems using unify. +Normalises the terms and theorems after unify with norms. +Returns a sequence of (env, thm) pairs.*) +fun try_hint match norms hint unify binders ctxt (t1, t2) + (Envir.Envir {maxidx, tenv, tyenv}) = + let + val dest_all_equals = TUtil.strip_all ##> Logic.dest_equals + val rev_binders = rev binders + val _ = @{log Logger.TRACE} ctxt (fn _ => Pretty.block [ + Pretty.str "Trying hint: ", pretty_hint ctxt hint + ] |> Pretty.string_of) + (*lift hint to include bound variables and increase indexes*) + val lifted_hint = + (*"P" is just some dummy proposition variable with appropriate index*) + Logic.list_all (map fst rev_binders, Var (("P", maxidx + 1), propT)) + |> Thm.cterm_of ctxt + (*lift bound variables to the hint*) + |> GUtil.flip Thm.lift_rule hint + val (hint_prems, hint_concl) = Thm.prop_of lifted_hint |> TUtil.strip_subgoal |> snd + + (*match hint with unification pair*) + val (P, Q) = dest_all_equals hint_concl |> snd + val match = match binders ctxt + val (no_hint_match, match_env_concl_thmpq) = + Envir.Envir {maxidx = Thm.maxidx_of lifted_hint, tenv = tenv, tyenv = tyenv} + |> match (P, t1) + |> Seq.maps (fn (env, thm1) => match (Q, t2) env |> Seq.map (apsnd (pair thm1))) + |> GUtil.seq_is_empty + in + if no_hint_match then (@{log Logger.TRACE} ctxt (K "Hint does not match."); Seq.empty) + else + let + val _ = @{log Logger.DEBUG} ctxt (fn _ => Pretty.block [ + Pretty.str "Hint ", + pretty_hint ctxt hint, + Pretty.str ((length hint_prems > 0 ? suffix " Unifying premises...") " matched.") + ] |> Pretty.string_of) + (*unify each hint premise and collect the theorems while iteratively extending the environment*) + fun unif_prem prem = + let + val (params, PQ_prem) = dest_all_equals prem + val (binders, ctxt) = Binders.fix_binders params ctxt + fun unif_add_result (env, thms) = unify binders ctxt PQ_prem env + |> Seq.map (apsnd (fn thm => (binders, thm) :: thms)) + in Seq.maps unif_add_result end + fun unify_prems (match_env, concl_thmp) = + fold unif_prem hint_prems (Seq.single (match_env, [])) + |> Seq.map (rpair concl_thmp) + fun inst_discharge ((env_res, prem_thms), concl_thmp) = + let + (*instantiate the theorems*) + val lifted_hint_inst = #norm_thm norms ctxt env_res lifted_hint + val norm_term = #norm_term norms env_res + val mk_norm_cbinders = map (Thm.cterm_of ctxt o norm_term o snd) + val norm_unif_thm = #norm_unif_thm norms ctxt env_res + fun forall_intr binders = fold Thm.forall_intr (mk_norm_cbinders binders) o norm_unif_thm + val prem_thms_inst = map (uncurry forall_intr) prem_thms |> rev + val (concl_thmL, concl_thmR) = apply2 norm_unif_thm concl_thmp + (*discharge the hint premises*) + val thm_res = Drule.implies_elim_list lifted_hint_inst prem_thms_inst + |> Drule.forall_elim_list (mk_norm_cbinders rev_binders) + |> GUtil.flip Thm.transitive concl_thmR + |> Thm.transitive (Unification_Base.symmetric concl_thmL) + in (env_res, thm_res) end + in Seq.maps unify_prems match_env_concl_thmpq |> Seq.map inst_discharge end + end + +type hint_retrieval = term Binders.binders -> Proof.context -> term * term -> Envir.env -> + unif_hint Seq.seq + +fun try_hints get_hints match norms unif binders ctxt (t1, t2) env = + (@{log Logger.DEBUG} ctxt (fn _ => + Pretty.block [ + Pretty.str "Trying unification hints for ", + Unification_Util.pretty_unif_problem ctxt (t1, t2) + ] |> Pretty.string_of); + get_hints binders ctxt (t1, t2) env + |> Seq.maps (fn hint => try_hint match norms hint unif binders ctxt (t1, t2) env)) + +end \ No newline at end of file diff --git a/thys/ML_Unification/Unification_Parsers/ML_Unification_Parsers.thy b/thys/ML_Unification/Unification_Parsers/ML_Unification_Parsers.thy new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/Unification_Parsers/ML_Unification_Parsers.thy @@ -0,0 +1,13 @@ +\<^marker>\creator "Kevin Kappelmann"\ +section \Unification Parsers\ +theory ML_Unification_Parsers + imports + ML_Parsing_Utils +begin + +paragraph \Summary\ +text \Common parsers needed for unification attributes, tactics, methods.\ + +ML_file\unification_parser.ML\ + +end diff --git a/thys/ML_Unification/Unification_Parsers/unification_parser.ML b/thys/ML_Unification/Unification_Parsers/unification_parser.ML new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/Unification_Parsers/unification_parser.ML @@ -0,0 +1,18 @@ +(* Title: ML_Unification/unification_parser.ML + Author: Kevin Kappelmann + +Common parsers needed for unification attributes, tactics, methods. +*) +signature UNIFICATION_PARSER = +sig + val parse_normalisers : ML_Code_Util.code parser + val parse_unifier : ML_Code_Util.code parser +end + +structure Unification_Parser : UNIFICATION_PARSER = +struct + +val parse_normalisers = Parse_Util.nonempty_code (K "normalisers must not be empty") +val parse_unifier = Parse_Util.nonempty_code (K "unifier must not be empty") + +end \ No newline at end of file diff --git a/thys/ML_Unification/Unification_Tactics/Assumption/Unify_Assumption_Tactic.thy b/thys/ML_Unification/Unification_Tactics/Assumption/Unify_Assumption_Tactic.thy new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/Unification_Tactics/Assumption/Unify_Assumption_Tactic.thy @@ -0,0 +1,70 @@ +\<^marker>\creator "Kevin Kappelmann"\ +subsection \Assumption Tactic\ +theory Unify_Assumption_Tactic + imports + ML_Functor_Instances + ML_Unifiers + ML_Unification_Parsers +begin + +paragraph \Summary\ +text \Assumption tactic and method with adjustable unifier.\ + +ML_file\unify_assumption_base.ML\ +ML_file\unify_assumption.ML\ + +ML\ + @{functor_instance struct_name = Standard_Unify_Assumption + and functor_name = Unify_Assumption + and id = \""\ + and more_args = \val init_args = { + normalisers = SOME Standard_Mixed_Unification.norms_first_higherp_first_comb_higher_unify, + unifier = SOME Standard_Mixed_Unification.first_higherp_first_comb_higher_unify + }\} +\ +local_setup \Standard_Unify_Assumption.setup_attribute NONE\ +local_setup \Standard_Unify_Assumption.setup_method NONE\ + + +paragraph \Examples\ + +experiment +begin + +lemma "PROP P \ PROP P" + by uassm + +lemma + assumes h: "\P. PROP P" + shows "PROP P x" + using h by uassm + +schematic_goal "\x. PROP P (c :: 'a) \ PROP ?Y (x :: 'a)" + by uassm + +schematic_goal a: "PROP ?P (y :: 'a) \ PROP ?P (?x :: 'a)" + by uassm \\compare the result with following call\ + (* by assumption *) + +schematic_goal + "PROP ?P (x :: 'a) \ PROP P (?x :: 'a)" + by uassm \\compare the result with following call\ + (* by assumption *) + +schematic_goal + "\x. PROP D \ (\P y. PROP P y x) \ PROP C \ PROP P x" + by (uassm unifier = Higher_Order_Unification.unify) \\the line below is equivalent\ + (* using [[uassm unifier = Higher_Order_Unification.unify]] by uassm *) + +text \Unlike @{method assumption}, @{method uassm} will not close the goal if the order of premises +of the assumption and the goal are different. Compare the following two examples:\ + +lemma "\x. PROP D \ (\y. PROP A y \ PROP B x) \ PROP C \ PROP A x \ PROP B x" + by uassm + +lemma "\x. PROP D \ (\y. PROP A y \ PROP B x) \ PROP A x \ PROP C \ PROP B x" + by assumption + (* by uassm *) +end + +end diff --git a/thys/ML_Unification/Unification_Tactics/Assumption/unify_assumption.ML b/thys/ML_Unification/Unification_Tactics/Assumption/unify_assumption.ML new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/Unification_Tactics/Assumption/unify_assumption.ML @@ -0,0 +1,132 @@ +(* Title: ML_Unification/unify_assumption.ML + Author: Kevin Kappelmann + +Assumption tactic with arguments from context. +*) +@{parse_entries (sig) PARSE_UNIFY_ASSUMPTION_ARGS [normalisers, unifier]} + +signature UNIFY_ASSUMPTION_ARGS = +sig + structure PA : PARSE_UNIFY_ASSUMPTION_ARGS + + type args = (Unification_Base.normalisers, Unification_Base.unifier) PA.entries + val unify_assumption_tac : args -> Proof.context -> int -> tactic + + val arg_parsers : (ML_Code_Util.code parser, ML_Code_Util.code parser) PA.entries +end + +structure Unify_Assumption_Args : UNIFY_ASSUMPTION_ARGS = +struct + +@{parse_entries (struct) PA [normalisers, unifier]} + +type args = (Unification_Base.normalisers, Unification_Base.unifier) PA.entries + +fun unify_assumption_tac args ctxt = + Unify_Assumption_Base.unify_assumption_tac (PA.get_normalisers args) (PA.get_unifier args) ctxt + +val arg_parsers = { + normalisers = SOME Unification_Parser.parse_normalisers, + unifier = SOME Unification_Parser.parse_unifier +} + +end + +signature UNIFY_ASSUMPTION = +sig + structure Data : GENERIC_DATA + + val get_args : Context.generic -> Unify_Assumption_Args.args + val map_args : (Unify_Assumption_Args.args -> Unify_Assumption_Args.args) -> + Context.generic -> Context.generic + + val get_normalisers : Context.generic -> Unification_Base.normalisers + val map_normalisers : (Unification_Base.normalisers -> Unification_Base.normalisers) -> + Context.generic -> Context.generic + + val get_unifier : Context.generic -> Unification_Base.unifier + val map_unifier : (Unification_Base.unifier -> Unification_Base.unifier) -> + Context.generic -> Context.generic + + val unify_assumption_tac : Proof.context -> int -> tactic + + val binding : binding + + val attribute : (ML_Code_Util.code, ML_Code_Util.code) Unify_Assumption_Args.PA.entries * + Position.T -> attribute + val parse_attribute : attribute parser + val setup_attribute : string option -> local_theory -> local_theory + + val parse_method : (Proof.context -> Method.method) context_parser + val setup_method : string option -> local_theory -> local_theory +end + +functor Unify_Assumption(A : sig + structure FIA : FUNCTOR_INSTANCE_ARGS + val init_args : Unify_Assumption_Args.args + end) : UNIFY_ASSUMPTION = +struct + +structure UAA = Unify_Assumption_Args +structure PA = UAA.PA + +structure FIU = Functor_Instance_Util(A.FIA) +structure MCU = ML_Code_Util + +structure Data = Generic_Data(struct + type T = UAA.args + val empty = A.init_args + val merge = fst +end) + +val get_args = Data.get +val map_args = Data.map + +val get_normalisers = PA.get_normalisers o get_args +val map_normalisers = map_args o PA.map_normalisers + +val get_unifier = PA.get_unifier o get_args +val map_unifier = map_args o PA.map_unifier + +fun unify_assumption_tac ctxt = UAA.unify_assumption_tac (get_args (Context.Proof ctxt)) ctxt + +val binding = FIU.mk_binding_id_prefix "uassm" + +val parse_arg_entries = + let + val parse_value = PA.parse_entry (PA.get_normalisers UAA.arg_parsers) + (PA.get_unifier UAA.arg_parsers) + val parse_entry = Parse_Key_Value.parse_entry PA.parse_key Parse_Util.eq parse_value + in PA.parse_entries_required Parse.and_list1 [] parse_entry (PA.empty_entries ()) end + +fun attribute (entries, pos) = + let + fun code_PA_op operation = MCU.flat_read ["Unify_Assumption_Args.PA.", operation] + val code_from_key = code_PA_op o PA.key_to_string + fun code_from_entry (PA.normalisers c) = MCU.atomic c + | code_from_entry (PA.unifier c) = MCU.atomic c + val code_entries = PA.key_entry_entries_from_entries entries + |> map (fn (k, v) => code_from_key k @ code_from_entry v) + |> MCU.list + val code = FIU.code_struct_op "map_args" @ MCU.atomic (code_PA_op "merge_entries" @ + MCU.atomic (code_PA_op "entries_from_entry_list" @ code_entries)) + in ML_Attribute.run_map_context (code, pos) end + +val parse_attribute = (parse_arg_entries |> Parse_Util.position) >> attribute + +val setup_attribute = Attrib.local_setup binding + (Parse.!!! parse_attribute |> Scan.lift) o + the_default ("set unify-assumption arguments (" ^ FIU.FIA.full_name ^ ")") + +val parse_method = + Parse_Util.option (Parse.!!! parse_attribute) + >> (fn opt_attr => + the_default I (Option.map ML_Attribute_Util.attribute_map_ctxt opt_attr) + #> SIMPLE_METHOD' o unify_assumption_tac) + |> Scan.lift + +val setup_method = Method.local_setup binding parse_method o + the_default ("solve first subgoal by assumption; with adjustable unifier (" + ^ FIU.FIA.full_name ^ ")") + +end \ No newline at end of file diff --git a/thys/ML_Unification/Unification_Tactics/Assumption/unify_assumption_base.ML b/thys/ML_Unification/Unification_Tactics/Assumption/unify_assumption_base.ML new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/Unification_Tactics/Assumption/unify_assumption_base.ML @@ -0,0 +1,96 @@ +(* Title: ML_Unification/unify_assumption_base.ML + Author: Kevin Kappelmann + +Assumption tactic with adjustable unifier. +*) +signature UNIFY_ASSUMPTION_BASE = +sig + include HAS_LOGGER + + val unify_assumption_tac_index : Unification_Base.normalisers -> Unification_Base.unifier -> + Proof.context -> int -> thm -> (int * thm) Seq.seq + val unify_assumption_tac : Unification_Base.normalisers -> Unification_Base.unifier -> + Proof.context -> int -> tactic +end + +structure Unify_Assumption_Base : UNIFY_ASSUMPTION_BASE = +struct + +val logger = Logger.setup_new_logger Unification_Base.logger "Unify_Assumption_Base" + +structure UBase = Unification_Base + +local +(*returns unification result and updated state where the unified assumption is moved to the front*) +fun unify_assumption unify ctxt i state = + (let + val strip_subgoal = Term_Util.strip_nth_subgoal i + val (rparams, (prems, _)) = strip_subgoal state + val nprems = length prems + in + if nprems < 1 then Seq.empty + else + let + val (binders, unif_ctxt) = Binders.fix_binders rparams ctxt + fun prepare_nth_prem n prem = + let val (params, nprems_prem) = Term_Util.strip_all prem ||> Logic.count_prems + in + if nprems - n < nprems_prem then Seq.empty (*too many subpremises in given premise*) + else + let val nprems_new = nprems - nprems_prem + in + (*protect conclusion*) + Tactic_Util.protect_tac nprems_new ctxt i state + (*move premise to the front*) + |> Seq.maps (rotate_tac (n - 1) i) + (*remove all other premises*) + |> Seq.maps (Tactic_Util.thin_tacs (2 upto nprems_new) i) + (*remove outermost parameters of premise*) + |> Seq.maps (REPEAT_DETERM_N (length params) (dresolve_tac ctxt [@{thm Pure.meta_spec}] i)) + (*protect premise*) + |> Seq.maps (dmatch_tac ctxt [Drule.protectI] i) + end + end + fun mk_result prem_index state = + (*unified assumption must be the first premise in the passed state*) + (let val (_, (prem :: _, concl)) = strip_subgoal state + in + unify binders unif_ctxt (prem, concl) (Envir.empty (Thm.maxidx_of state)) + |> Seq.map (pair (prem_index, binders) o rpair state) + end) + handle TERM _ => Seq.empty + fun unify_prem (n, prem) = let val prem_index = n + 1 + in + prepare_nth_prem prem_index prem + |> Seq.maps (mk_result prem_index) + |> Seq.append + end + in General_Util.fold_rev_index unify_prem prems Seq.empty end + end) + handle TERM _ => Seq.empty +in + +fun unify_assumption_tac_index norms unify ctxt i state = + let + val _ = @{log Logger.DEBUG} ctxt (fn _ => Pretty.block [ + Pretty.str "Trying proof by assumption for subgoal ", + Pretty.str (Int.toString i), + Pretty.str ": ", + (Thm.prop_of state |> Logic.nth_prem o pair i |> Syntax.pretty_term ctxt) + handle TERM _ => Pretty.str "" + ] |> Pretty.string_of) + fun rewrite_subgoal ((n, binders), (res, state)) = + (*rewrite conclusion to match the premise*) + Tactic_Util.rewrite_subgoal_unif_concl (#norm_thm norms) (#norm_unif_thm norms) + binders res ctxt i state + (*unified premise is moved to the front; now close by equality assumption*) + |> Seq.maps (DETERM (Tactic_Util.nth_eq_assume_tac 1 i)) + |> Seq.map (pair n) + in unify_assumption unify ctxt i state |> Seq.maps rewrite_subgoal end + +fun unify_assumption_tac norms = + Seq.map snd oooo unify_assumption_tac_index norms + +end + +end \ No newline at end of file diff --git a/thys/ML_Unification/Unification_Tactics/Fact/Unify_Fact_Tactic.thy b/thys/ML_Unification/Unification_Tactics/Fact/Unify_Fact_Tactic.thy new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/Unification_Tactics/Fact/Unify_Fact_Tactic.thy @@ -0,0 +1,48 @@ +\<^marker>\creator "Kevin Kappelmann"\ +subsection \Fact Tactic\ +theory Unify_Fact_Tactic + imports + Unify_Resolve_Tactics +begin + +paragraph \Summary\ +text \Fact tactic with adjustable unifier.\ + +ML_file\unify_fact_base.ML\ +ML_file\unify_fact.ML\ + +ML\ + @{functor_instance struct_name = Standard_Unify_Fact + and functor_name = Unify_Fact + and id = \""\ + and more_args = \val init_args = { + normalisers = SOME Standard_Mixed_Unification.norms_first_higherp_first_comb_higher_unify, + unifier = SOME Standard_Mixed_Unification.first_higherp_first_comb_higher_unify + }\} +\ +local_setup \Standard_Unify_Fact.setup_attribute NONE\ +local_setup \Standard_Unify_Fact.setup_method NONE\ + + +paragraph \Examples\ + +experiment +begin +lemma + assumes h: "\x y. PROP P x y" + shows "PROP P x y" + by (ufact h) + +lemma + assumes "\P y. PROP P y x" + shows "PROP P x" + by (ufact assms where unifier = Higher_Order_Unification.unify) \\the line below is equivalent\ + (* using [[ufact unifier = Higher_Order_Unification.unify]] by (ufact assms) *) + +lemma + assumes "\x y. PROP A x \ PROP B x \ PROP P x" + shows "\x y. PROP A x \ PROP B x \ PROP P x" + using assms by ufact +end + +end diff --git a/thys/ML_Unification/Unification_Tactics/Fact/unify_fact.ML b/thys/ML_Unification/Unification_Tactics/Fact/unify_fact.ML new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/Unification_Tactics/Fact/unify_fact.ML @@ -0,0 +1,157 @@ +(* Title: ML_Unification/unify_fact.ML + Author: Kevin Kappelmann + +Fact tactic tactic with arguments from context. +*) +@{parse_entries (sig) PARSE_UNIFY_FACT_ARGS + [normalisers, unifier, facts]} +@{parse_entries (sig) PARSE_UNIFY_FACT_CONTEXT_ARGS + [normalisers, unifier]} + +@{parse_entries (sig) UNIFY_RESOLVE_CHAINED_MODES [insert, resolve, fact]} + +signature UNIFY_FACT_ARGS = +sig + structure PA : PARSE_UNIFY_FACT_ARGS + structure PCA : PARSE_UNIFY_FACT_CONTEXT_ARGS + + val PCA_entries_from_PA_entries : ('a, 'b, 'c) PA.entries -> ('a, 'b) PCA.entries + val PA_entries_from_PCA_entries : ('a, 'b) PCA.entries -> 'c -> ('a, 'b, 'c) PA.entries + + type args = (Unification_Base.normalisers, Unification_Base.unifier, thm list) PA.entries + type context_args = (Unification_Base.normalisers, Unification_Base.unifier) PCA.entries + + val unify_fact_tac : args -> Proof.context -> int -> tactic + val unify_fact_context_args_tac : context_args -> thm list -> Proof.context -> int -> tactic + + val arg_parsers : (ML_Code_Util.code parser, ML_Code_Util.code parser, thm list context_parser) + PA.entries +end + +structure Unify_Fact_Args : UNIFY_FACT_ARGS = +struct + +@{parse_entries (struct) PA [normalisers, unifier, facts]} +@{parse_entries (struct) PCA [normalisers, unifier]} + +fun PCA_entries_from_PA_entries {normalisers = normalisers, unifier = unifier,...} = + {normalisers = normalisers, unifier = unifier} +fun PA_entries_from_PCA_entries {normalisers = normalisers, unifier = unifier} facts = + {normalisers = normalisers, unifier = unifier, facts = SOME facts} + +type args = (Unification_Base.normalisers, Unification_Base.unifier, thm list) PA.entries +type context_args = (Unification_Base.normalisers, Unification_Base.unifier) PCA.entries + +fun unify_fact_tac args = + Unify_Fact_Base.unify_fact_tac (PA.get_normalisers args) (PA.get_unifier args) (PA.get_facts args) + +val unify_fact_context_args_tac = unify_fact_tac oo PA_entries_from_PCA_entries + +val arg_parsers = { + normalisers = SOME Unification_Parser.parse_normalisers, + unifier = SOME Unification_Parser.parse_unifier, + facts = SOME Parse_Util.thms +} + +end + +signature UNIFY_FACT = +sig + structure Data : GENERIC_DATA + + val get_args : Context.generic -> Unify_Fact_Args.context_args + val map_args : (Unify_Fact_Args.context_args -> Unify_Fact_Args.context_args) -> + Context.generic -> Context.generic + + val get_normalisers : Context.generic -> Unification_Base.normalisers + val map_normalisers : (Unification_Base.normalisers -> Unification_Base.normalisers) -> + Context.generic -> Context.generic + + val get_unifier : Context.generic -> Unification_Base.unifier + val map_unifier : (Unification_Base.unifier -> Unification_Base.unifier) -> + Context.generic -> Context.generic + + val unify_fact_tac : thm list -> Proof.context -> int -> tactic + + val binding : binding + + val attribute : (ML_Code_Util.code, ML_Code_Util.code) Unify_Fact_Args.PCA.entries * + Position.T -> attribute + val parse_attribute : attribute parser + val setup_attribute : string option -> local_theory -> local_theory + + val parse_method : (Proof.context -> Method.method) context_parser + val setup_method : string option -> local_theory -> local_theory +end + +functor Unify_Fact(A : sig + structure FIA : FUNCTOR_INSTANCE_ARGS + val init_args : Unify_Fact_Args.context_args + end) : UNIFY_FACT = +struct + +structure UFA = Unify_Fact_Args +structure PCA = UFA.PCA +structure PA = UFA.PA + +structure FIU = Functor_Instance_Util(A.FIA) +structure MCU = ML_Code_Util + +structure Data = Generic_Data(struct + type T = UFA.context_args + val empty = A.init_args + val merge = fst +end) + +val get_args = Data.get +val map_args = Data.map + +val get_normalisers = PCA.get_normalisers o get_args +val map_normalisers = map_args o PCA.map_normalisers + +val get_unifier = PCA.get_unifier o get_args +val map_unifier = map_args o PCA.map_unifier + +fun unify_fact_tac facts ctxt = + UFA.unify_fact_context_args_tac (get_args (Context.Proof ctxt)) facts ctxt + +val binding = FIU.mk_binding_id_prefix "ufact" + +val parse_arg_entries = + let + val parsers = UFA.arg_parsers |> UFA.PCA_entries_from_PA_entries + val parse_value = PCA.parse_entry (PCA.get_normalisers parsers) (PCA.get_unifier parsers) + val parse_entry = Parse_Key_Value.parse_entry PCA.parse_key Parse_Util.eq parse_value + in PCA.parse_entries_required Parse.and_list1 [] parse_entry (PCA.empty_entries ()) end + +fun attribute (entries, pos) = + let + fun code_PCA_op operation = MCU.flat_read ["Unify_Fact_Args.PCA.", operation] + val code_from_key = code_PCA_op o PCA.key_to_string + fun code_from_entry (PCA.normalisers c) = MCU.atomic c + | code_from_entry (PCA.unifier c) = MCU.atomic c + val code_entries = PCA.key_entry_entries_from_entries entries + |> map (fn (k, v) => code_from_key k @ code_from_entry v) + |> MCU.list + val code = FIU.code_struct_op "map_args" @ MCU.atomic (code_PCA_op "merge_entries" @ + MCU.atomic (code_PCA_op "entries_from_entry_list" @ code_entries)) + in ML_Attribute.run_map_context (code, pos) end + +val parse_attribute = (parse_arg_entries |> Parse_Util.position) >> attribute + +val setup_attribute = Attrib.local_setup binding + (Parse.!!! parse_attribute |> Scan.lift) o + the_default ("set unify-fact arguments (" ^ FIU.FIA.full_name ^ ")") + +val parse_method = + PA.get_facts UFA.arg_parsers + -- ((Parse.where_ |-- parse_attribute) |> Parse.!!! |> Parse_Util.option |> Scan.lift) + >> (fn (facts, opt_attr) => + the_default I (Option.map ML_Attribute_Util.attribute_map_ctxt opt_attr) + #> Method_Util.METHOD_CONTEXT' (append facts #> unify_fact_tac)) + +val setup_method = Method.local_setup binding parse_method o + the_default ("solve first subgoal by unification with a given fact; with adjustable unifier (" + ^ FIU.FIA.full_name ^ ")") + +end \ No newline at end of file diff --git a/thys/ML_Unification/Unification_Tactics/Fact/unify_fact_base.ML b/thys/ML_Unification/Unification_Tactics/Fact/unify_fact_base.ML new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/Unification_Tactics/Fact/unify_fact_base.ML @@ -0,0 +1,17 @@ +(* Title: ML_Unification/unify_fact_base.ML + Author: Kevin Kappelmann + +Fact tactic with adjustable unifier. +*) +signature UNIFY_FACT_BASE = +sig + val unify_fact_tac : Unification_Base.normalisers -> Unification_Base.unifier -> + thm list -> Proof.context -> int -> tactic +end + +structure Unify_Fact_Base : UNIFY_FACT_BASE = +struct + +val unify_fact_tac = Unify_Resolve_Base.unify_resolve_atomic_any_tac + +end \ No newline at end of file diff --git a/thys/ML_Unification/Unification_Tactics/Resolution/Unify_Resolve_Tactics.thy b/thys/ML_Unification/Unification_Tactics/Resolution/Unify_Resolve_Tactics.thy new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/Unification_Tactics/Resolution/Unify_Resolve_Tactics.thy @@ -0,0 +1,92 @@ +\<^marker>\creator "Kevin Kappelmann"\ +subsection \Resolution Tactics\ +theory Unify_Resolve_Tactics + imports + Unify_Assumption_Tactic + ML_Method_Utils +begin + +paragraph \Summary\ +text \Resolution tactics and methods with adjustable unifier.\ + +ML_file\unify_resolve_base.ML\ +ML_file\unify_resolve.ML\ +ML\ + @{functor_instance struct_name = Standard_Unify_Resolve + and functor_name = Unify_Resolve + and id = \""\ + and more_args = \val init_args = { + normalisers = SOME Standard_Mixed_Unification.norms_first_higherp_first_comb_higher_unify, + unifier = SOME Standard_Mixed_Unification.first_higherp_first_comb_higher_unify, + mode = SOME (Unify_Resolve_Args.PM.key Unify_Resolve_Args.PM.any), + chained = SOME (Unify_Resolve_Args.PCM.key Unify_Resolve_Args.PCM.resolve) + }\} +\ +local_setup \Standard_Unify_Resolve.setup_attribute NONE\ +local_setup \Standard_Unify_Resolve.setup_method NONE\ + + +paragraph \Examples\ + +experiment +begin + +lemma + assumes h: "\x. PROP D x \ PROP C x" + shows "\x. PROP A x \ PROP B x \ PROP C x" + apply (urule h) \\the line below is equivalent\ + (* apply (rule h) *) + oops + +lemma + assumes h: "PROP C x" + shows "PROP C x" + by (urule h where unifier = First_Order_Unification.unify) \\the line below is equivalent\ + (* using [[urule unifier = First_Order_Unification.unify]] by (urule h) *) + +lemma + assumes h: "\x. PROP A x \ PROP D x" + shows "\x. PROP A x \ PROP B x \ PROP C x" + \\use (r,e,d,f) to specify the resolution mode (resolution, elim, dest, forward)\ + apply (urule (d) h) \\the line below is equivalent\ + (* apply (drule h) *) + oops + +text\You can specify how chained facts should be used. By default, @{method urule} works like +@{method rule}: it uses chained facts to resolve against the premises of the passed rules.\ + +lemma + assumes h1: "\x. (PROP F x \ PROP E x) \ PROP C x" + and h2: "\x. PROP F x \ PROP E x" + shows "\x. PROP A x \ PROP B x \ PROP C x" + \\Compare all of the following calls:\ + (* apply (rule h1) *) + (* apply (urule h1) *) + (* using h2 apply (rule h1) *) + (* using h2 apply (urule h1) *) + using h2 apply (urule h1 where chained = fact) + (* using h2 apply (urule h1 where chained = insert) *) + done + +text\You can specify whether any or every rule must resolve against the goal:\ + +lemma + assumes h1: "\x y. PROP C y \ PROP D x \ PROP C x" + and h2: "\x y. PROP C x \ PROP D x" + and h3: "\x y. PROP C x" + shows "\x. PROP A x \ PROP B x \ PROP C x" + using h3 apply (urule h1 h2 where mode = every) + (* using h3 apply (urule h1 h2) *) + done + +lemma + assumes h1: "\x y. PROP C y \ PROP A x \ PROP C x" + and h2: "\x y. PROP C x \ PROP B x \ PROP D x" + and h3: "\x y. PROP C x" + shows "\x. PROP A x \ PROP B x \ PROP C x" + using h3 apply (urule (d) h1 h2 where mode = every) + oops + +end + +end diff --git a/thys/ML_Unification/Unification_Tactics/Resolution/unify_resolve.ML b/thys/ML_Unification/Unification_Tactics/Resolution/unify_resolve.ML new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/Unification_Tactics/Resolution/unify_resolve.ML @@ -0,0 +1,277 @@ +(* Title: ML_Unification/unify_resolve.ML + Author: Kevin Kappelmann + +Unification-resolution tactics with arguments from context. +*) +@{parse_entries (sig) PARSE_UNIFY_RESOLVE_ARGS + [normalisers, unifier, mode, chained, resolve_mode, rules]} + +@{parse_entries (sig) PARSE_UNIFY_RESOLVE_CONTEXT_ARGS + [normalisers, unifier, mode, chained]} + +@{parse_entries (sig) PARSE_UNIFY_RESOLVE_MODES [any, every]} + +@{parse_entries (sig) PARSE_UNIFY_RESOLVE_CHAINED_MODES [insert, resolve, fact]} + +@{parse_entries (sig) PARSE_UNIFY_RESOLVE_RESOLVE_MODES [r, e, d, f]} + +signature UNIFY_RESOLVE_ARGS = +sig + include HAS_LOGGER + + structure PA : PARSE_UNIFY_RESOLVE_ARGS + structure PCA : PARSE_UNIFY_RESOLVE_CONTEXT_ARGS + + val PCA_entries_from_PA_entries : ('a, 'b, 'c, 'd, 'e, 'f) PA.entries -> ('a, 'b, 'c, 'd) PCA.entries + val PA_entries_from_PCA_entries : ('a, 'b, 'c, 'd) PCA.entries -> 'e -> 'f -> + ('a, 'b, 'c, 'd, 'e, 'f) PA.entries + + structure PM : PARSE_UNIFY_RESOLVE_MODES + type mode = PM.key + val parse_mode : mode parser + + structure PCM : PARSE_UNIFY_RESOLVE_CHAINED_MODES + type chained_mode = PCM.key + val parse_chained_mode : chained_mode parser + + structure PRM : PARSE_UNIFY_RESOLVE_RESOLVE_MODES + type resolve_mode = PRM.key + val parse_resolve_mode : resolve_mode parser + + type args = (Unification_Base.normalisers, Unification_Base.unifier, mode, chained_mode, + resolve_mode, thm list) PA.entries + type context_args = (Unification_Base.normalisers, Unification_Base.unifier, mode, + chained_mode) PCA.entries + + val unify_resolve_tac : args -> thm list -> Proof.context -> int -> tactic + val unify_resolve_context_args_tac : context_args -> resolve_mode -> thm list -> thm list -> + Proof.context -> int -> tactic + + val arg_parsers : (ML_Code_Util.code parser, ML_Code_Util.code parser, mode parser, + chained_mode parser, resolve_mode parser, thm list context_parser) PA.entries +end + +structure Unify_Resolve_Args : UNIFY_RESOLVE_ARGS = +struct + +structure URB = Unify_Resolve_Base + +val logger = Logger.setup_new_logger URB.logger "Unify_Resolve_Args" + +@{parse_entries (struct) PA + [normalisers, unifier, mode, chained, resolve_mode, rules]} +@{parse_entries (struct) PCA + [normalisers, unifier, mode, chained]} + +fun PCA_entries_from_PA_entries {normalisers = normalisers, unifier = unifier, mode = mode, + chained = chained,...} = + {normalisers = normalisers, unifier = unifier, mode = mode, chained = chained} +fun PA_entries_from_PCA_entries {normalisers = normalisers, unifier = unifier, mode = mode, + chained = chained} resolve__mode rules = {normalisers = normalisers, unifier = unifier, + mode = mode, chained = chained, resolve_mode = SOME resolve__mode, rules = SOME rules} + +@{parse_entries (struct) PM [any, every]} +type mode = PM.key +val parse_mode = PM.parse_key + +@{parse_entries (struct) PCM [insert, resolve, fact]} +type chained_mode = PCM.key +val parse_chained_mode = PCM.parse_key + +@{parse_entries (struct) PRM [r, e, d, f]} +type resolve_mode = PRM.key +val parse_resolve_mode = PRM.parse_key + +type args = (Unification_Base.normalisers, Unification_Base.unifier, mode, chained_mode, + resolve_mode, thm list) PA.entries +type context_args = (Unification_Base.normalisers, Unification_Base.unifier, mode, + chained_mode) PCA.entries + +fun unify_resolve_tac args chained_facts ctxt i = + let + val norms = PA.get_normalisers args + val unify = PA.get_unifier args + val mode = PA.get_mode args + val resolve_fn = (case (PA.get_resolve_mode args, mode) of + (PRM.r _, PM.any _) => URB.unify_resolve_any_tac + | (PRM.r _, PM.every _) => URB.unify_resolve_every_tac + (*note: for e-resolution, we use the passed unifier to unify the conclusion + as well as the premise of the elim-rule*) + | (PRM.e _, PM.any _) => URB.unify_eresolve_any_tac norms unify + | (PRM.e _, PM.every _) => URB.unify_eresolve_every_tac norms unify + | (PRM.d _, PM.any _) => URB.unify_dresolve_any_tac + | (PRM.d _, PM.every _) => URB.unify_dresolve_every_tac + | (PRM.f _, PM.any _) => URB.unify_fresolve_any_tac + | (PRM.f _, PM.every _) => URB.unify_fresolve_every_tac) + norms unify + val rules = PA.get_rules args + in case PA.get_chained args of + PCM.insert _ => (*insert chained facts*) + (Tactic_Util.insert_tac chained_facts ctxt + THEN' resolve_fn rules ctxt) i + | cm => (*pre-process rules with chained facts*) + let + val process_rule_tac = (case cm of + PCM.resolve _ => Unify_Resolve_Base.unify_resolve_every_tac + | PCM.fact _ => Unify_Resolve_Base.unify_resolve_atomic_every_tac) + norms unify chained_facts ctxt + |> HEADGOAL + fun add_rulesq (is_empty, rulesq) = case (is_empty, mode) of + (true, PM.any _) => I + | (true, PM.every _) => Seq.fail + | _ => Seq.maps (fn rules => Seq.map (fn rule => rule :: rules) rulesq) + fun process_rule rule = process_rule_tac rule + |> General_Util.seq_is_empty + |> tap (fn (true, _) => @{log Logger.WARN} ctxt (fn _ => Pretty.block [ + Pretty.str "Could not use chained facts for rule ", + Thm.pretty_thm ctxt rule + ] |> Pretty.string_of) + | _ => ()) + |> add_rulesq + in + fold_rev process_rule rules (Seq.single []) + |> Seq.lifts (fn rules => resolve_fn rules ctxt i) + end + end + +val unify_resolve_context_args_tac = unify_resolve_tac ooo PA_entries_from_PCA_entries + +val arg_parsers = { + normalisers = SOME Unification_Parser.parse_normalisers, + unifier = SOME Unification_Parser.parse_unifier, + mode = SOME parse_mode, + chained = SOME parse_chained_mode, + resolve_mode = SOME parse_resolve_mode, + rules = SOME (Parse_Util.nonempty_thms (K "must provide at least one rule to resolve")) +} + +end + +signature UNIFY_RESOLVE = +sig + structure Data : GENERIC_DATA + + val get_args : Context.generic -> Unify_Resolve_Args.context_args + val map_args : (Unify_Resolve_Args.context_args -> Unify_Resolve_Args.context_args) -> + Context.generic -> Context.generic + + val get_normalisers : Context.generic -> Unification_Base.normalisers + val map_normalisers : (Unification_Base.normalisers -> Unification_Base.normalisers) -> + Context.generic -> Context.generic + + val get_unifier : Context.generic -> Unification_Base.unifier + val map_unifier : (Unification_Base.unifier -> Unification_Base.unifier) -> + Context.generic -> Context.generic + + val get_mode : Context.generic -> Unify_Resolve_Args.mode + val map_mode : (Unify_Resolve_Args.mode -> Unify_Resolve_Args.mode) -> + Context.generic -> Context.generic + + val get_chained : Context.generic -> Unify_Resolve_Args.chained_mode + val map_chained : (Unify_Resolve_Args.chained_mode -> Unify_Resolve_Args.chained_mode) -> + Context.generic -> Context.generic + + val unify_resolve_tac : Unify_Resolve_Args.resolve_mode -> thm list -> thm list -> Proof.context -> + int -> tactic + + val binding : binding + + val attribute : (ML_Code_Util.code, ML_Code_Util.code, Unify_Resolve_Args.mode, + Unify_Resolve_Args.chained_mode) Unify_Resolve_Args.PCA.entries * Position.T -> attribute + val parse_attribute : attribute parser + val setup_attribute : string option -> local_theory -> local_theory + + val parse_method : (Proof.context -> Method.method) context_parser + val setup_method : string option -> local_theory -> local_theory +end + +functor Unify_Resolve(A : sig + structure FIA : FUNCTOR_INSTANCE_ARGS + val init_args : Unify_Resolve_Args.context_args + end) : UNIFY_RESOLVE = +struct + +structure URA = Unify_Resolve_Args +structure PCA = URA.PCA +structure PA = URA.PA +structure PRM = URA.PRM +structure FIU = Functor_Instance_Util(A.FIA) +structure MCU = ML_Code_Util + +structure Data = Generic_Data(struct + type T = URA.context_args + val empty = A.init_args + val merge = fst +end) + +val get_args = Data.get +val map_args = Data.map + +val get_normalisers = PCA.get_normalisers o get_args +val map_normalisers = map_args o PCA.map_normalisers + +val get_unifier = PCA.get_unifier o get_args +val map_unifier = map_args o PCA.map_unifier + +val get_mode = PCA.get_mode o get_args +val map_mode = map_args o PCA.map_mode + +val get_chained = PCA.get_chained o get_args +val map_chained = map_args o PCA.map_chained + +fun unify_resolve_tac resolve_mode rules facts ctxt = + URA.unify_resolve_context_args_tac (get_args (Context.Proof ctxt)) resolve_mode rules facts + ctxt + +val binding = FIU.mk_binding_id_prefix "urule" + +val parse_arg_entries = + let + val parsers = URA.arg_parsers |> URA.PCA_entries_from_PA_entries + val parse_value = PCA.parse_entry (PCA.get_normalisers parsers) + (PCA.get_unifier parsers) (PCA.get_mode parsers) (PCA.get_chained parsers) + val parse_entry = Parse_Key_Value.parse_entry PCA.parse_key Parse_Util.eq parse_value + in PCA.parse_entries_required Parse.and_list1 [] parse_entry (PCA.empty_entries ()) end + +fun attribute (entries, pos) = + let + fun URA_substructure_op substructure operation = + MCU.flat_read ["Unify_Resolve_Args.", substructure, ".", operation] + val [code_PM_op, code_PCM_op, code_PCA_op] = map URA_substructure_op ["PM", "PCM", "PCA"] + val code_from_key = code_PCA_op o PCA.key_to_string + fun code_from_entry (PCA.normalisers c) = c + | code_from_entry (PCA.unifier c) = c + | code_from_entry (PCA.mode m) = code_PM_op "key" @ code_PM_op (URA.PM.key_to_string m) + | code_from_entry (PCA.chained cm) = code_PCM_op "key" @ code_PCM_op (URA.PCM.key_to_string cm) + val code_entries = PCA.key_entry_entries_from_entries entries + |> map (fn (k, v) => code_from_key k @ MCU.atomic (code_from_entry v)) + |> MCU.list + val code = + FIU.code_struct_op "map_args" @ MCU.atomic (code_PCA_op "merge_entries" @ + MCU.atomic (code_PCA_op "entries_from_entry_list" @ code_entries)) + in ML_Attribute.run_map_context (code, pos) end + +val parse_attribute = (parse_arg_entries |> Parse_Util.position) >> attribute + +val setup_attribute = Attrib.local_setup binding + (Parse.!!! parse_attribute |> Scan.lift) o + the_default ("set unify-resolution arguments (" ^ FIU.FIA.full_name ^ ")") + +val parse_method = + let val parse_resolve_mode = Scan.optional + (Parse_Util.parenths (PA.get_resolve_mode URA.arg_parsers |> Parse.!!!)) + (PRM.key PRM.r) + in + Scan.lift parse_resolve_mode + -- PA.get_rules URA.arg_parsers + -- ((Parse.where_ |-- parse_attribute) |> Parse.!!! |> Parse_Util.option |> Scan.lift) + >> (fn ((resolve_mode, rules), opt_attr) => + the_default I (Option.map ML_Attribute_Util.attribute_map_ctxt opt_attr) + #> Method_Util.METHOD_CONTEXT' (unify_resolve_tac resolve_mode rules)) + end + +val setup_method = Method.local_setup binding parse_method o + the_default ("r/e/d/f-resolve rule(s) with first subgoal; with adjustable unifier (" + ^ FIU.FIA.full_name ^ ")") + +end \ No newline at end of file diff --git a/thys/ML_Unification/Unification_Tactics/Resolution/unify_resolve_base.ML b/thys/ML_Unification/Unification_Tactics/Resolution/unify_resolve_base.ML new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/Unification_Tactics/Resolution/unify_resolve_base.ML @@ -0,0 +1,201 @@ +(* Title: ML_Unification/unify_resolve_base.ML + Author: Kevin Kappelmann + +Resolution with adjustable unifier. +*) +signature UNIFY_RESOLVE_BASE = +sig + include HAS_LOGGER + + (*unify the conclusions of the rule and the subgoal and resolve*) + val unify_resolve_tac : Unification_Base.normalisers -> + Unification_Base.unifier -> thm -> Proof.context -> int -> tactic + val unify_resolve_every_tac : Unification_Base.normalisers -> + Unification_Base.unifier -> thm list -> Proof.context -> int -> tactic + val unify_resolve_any_tac : Unification_Base.normalisers -> + Unification_Base.unifier -> thm list -> Proof.context -> int -> tactic + (*treats the rule and subgoal as atomic propositions; example: + resolving rule A ==> B and subgoal A ==> B leads to A ==> A while + resolving rule A ==> B and subgoal A ==> B atomically solves the subgoal*) + val unify_resolve_atomic_tac : Unification_Base.normalisers -> + Unification_Base.unifier -> thm -> Proof.context -> int -> tactic + val unify_resolve_atomic_every_tac : Unification_Base.normalisers -> + Unification_Base.unifier -> thm list -> Proof.context -> int -> tactic + val unify_resolve_atomic_any_tac : Unification_Base.normalisers -> + Unification_Base.unifier -> thm list -> Proof.context -> int -> tactic + + val unify_eresolve_tac_index : Unification_Base.normalisers -> Unification_Base.unifier -> + Unification_Base.normalisers -> Unification_Base.unifier -> thm -> Proof.context -> + int -> thm -> (int * thm) Seq.seq + val unify_eresolve_tac : Unification_Base.normalisers -> Unification_Base.unifier -> + Unification_Base.normalisers -> Unification_Base.unifier -> thm -> + Proof.context -> int -> tactic + val unify_eresolve_every_tac : Unification_Base.normalisers -> Unification_Base.unifier -> + Unification_Base.normalisers -> Unification_Base.unifier -> thm list -> + Proof.context -> int -> tactic + val unify_eresolve_any_tac : Unification_Base.normalisers -> Unification_Base.unifier -> + Unification_Base.normalisers -> Unification_Base.unifier -> thm list -> + Proof.context -> int -> tactic + + val unify_dresolve_tac : Unification_Base.normalisers -> Unification_Base.unifier -> thm -> + Proof.context -> int -> tactic + val unify_dresolve_every_tac : Unification_Base.normalisers -> Unification_Base.unifier -> + thm list -> Proof.context -> int -> tactic + val unify_dresolve_any_tac : Unification_Base.normalisers -> Unification_Base.unifier -> + thm list -> Proof.context -> int -> tactic + + val unify_fresolve_tac : Unification_Base.normalisers -> Unification_Base.unifier -> thm -> + Proof.context -> int -> tactic + val unify_fresolve_every_tac : Unification_Base.normalisers -> Unification_Base.unifier -> + thm list -> Proof.context -> int -> tactic + val unify_fresolve_any_tac : Unification_Base.normalisers -> Unification_Base.unifier -> + thm list -> Proof.context -> int -> tactic + +end + +structure Unify_Resolve_Base : UNIFY_RESOLVE_BASE = +struct + +val logger = Logger.setup_new_logger Unification_Base.logger "Unify_Resolve_Base" + +structure Util = Unification_Util +structure HOPU = Higher_Order_Pattern_Unification + +fun unify_resolve_tac norms unify rule ctxt = + let fun tac cgoal i state = + let + val _ = @{log Logger.TRACE} ctxt (fn _ => + Pretty.block [ + Pretty.str ("Calling unification resolution tactic on subgoal "), + Pretty.str (Int.toString i), + Pretty.str ": ", + Util.pretty_terms ctxt [Thm.term_of cgoal], + Pretty.str (" and rule: "), + Thm.pretty_thm ctxt rule + ] + |> Pretty.string_of + ) + val lifted_rule = Thm.lift_rule cgoal (Drule.incr_indexes state rule) + val nprems_rule = Thm.nprems_of rule + val rparams = Term_Util.strip_all (Thm.term_of cgoal) |> fst + val (binders, ctxt) = Binders.fix_binders rparams ctxt + val conclp = (*only unify the conclusion of the rule and the goal*) + apply2 Logic.strip_assums_concl (Thm.prop_of lifted_rule, Thm.term_of cgoal) + val env_thmq = unify binders ctxt conclp (Envir.empty (Thm.maxidx_of lifted_rule)) + val norm_thm = #norm_thm norms + fun rewrite_tac (env, unif_thm) = + Tactic_Util.rewrite_subgoal_unif_concl norm_thm (#norm_unif_thm norms) + binders (env, unif_thm) ctxt + THEN' Tactic_Util.no_lift_resolve_tac (norm_thm ctxt env lifted_rule) nprems_rule ctxt + in Seq.maps (fn res => rewrite_tac res i state) env_thmq end + in Tactic_Util.CSUBGOAL_DATA I tac end + +fun unify_resolve_tacs norms unify rules ctxt = + map (fn rule => unify_resolve_tac norms unify rule ctxt) rules + +fun unify_resolve_every_tac norms unify rules = + unify_resolve_tacs norms unify rules #> EVERY' + +fun unify_resolve_any_tac norms unify rules = + unify_resolve_tacs norms unify rules #> Tactic_Util.CONCAT' + +fun unify_resolve_atomic_tac norms unify rule ctxt = + Tactic_Util.protect_tac 0 ctxt + THEN' unify_resolve_tac norms unify (Thm_Util.protect rule) ctxt + +fun unify_resolve_atomic_tacs norms unify rules ctxt = + map (fn rule => unify_resolve_atomic_tac norms unify rule ctxt) rules + +fun unify_resolve_atomic_every_tac norms unify rules = + unify_resolve_atomic_tacs norms unify rules #> EVERY' + +fun unify_resolve_atomic_any_tac norms unify rules = + unify_resolve_atomic_tacs norms unify rules #> Tactic_Util.CONCAT' + +fun unify_eresolve_tac_index norms_resolve unify_resolve norms_assume unify_assume + rule ctxt i state = + unify_resolve_tac norms_resolve unify_resolve rule ctxt i state + |> Seq.maps (Unify_Assumption_Base.unify_assumption_tac_index norms_assume unify_assume ctxt i) + +fun unify_eresolve_tac norms_resolve unify_resolve norms_assume unify_assume + rule ctxt i state = + let val nprems = Thm.nprems_of rule + in + if nprems = 0 then + (@{log Logger.WARN} ctxt (fn _ => + Pretty.block [ + Pretty.str "elim-resolution rule ", + Thm.pretty_thm ctxt rule, + Pretty.str " has no premises." + ] |> Pretty.string_of + ); + Seq.empty) + else + unify_eresolve_tac_index norms_resolve unify_resolve norms_assume unify_assume rule ctxt i state + |> Seq.maps (fn (n, state) => + Tactic_Util.EVERY_ARG (Tactic_Util.thin_tac n) (i upto i + nprems - 2) state) + end + +fun unify_eresolve_tacs norms_resolve unify_resolve norms_assume unify_assume rules ctxt = + map (fn rule => unify_eresolve_tac norms_resolve unify_resolve norms_assume unify_assume rule ctxt) + rules + +fun unify_eresolve_every_tac norms_resolve unify_resolve norms_assume unify_assume rules = + unify_eresolve_tacs norms_resolve unify_resolve norms_assume unify_assume rules + #> EVERY' + +fun unify_eresolve_any_tac norms_resolve unify_resolve norms_assume unify_assume rules = + unify_eresolve_tacs norms_resolve unify_resolve norms_assume unify_assume rules + #> Tactic_Util.CONCAT' + +fun unify_dresolve_tac norms unify rule ctxt = + let val nprems = Thm.nprems_of rule + in + if nprems = 0 then + (@{log Logger.WARN} ctxt (fn _ => + Pretty.block [ + Pretty.str "destruct-resolution rule ", + Thm.pretty_thm ctxt rule, + Pretty.str " has no premises." + ] |> Pretty.string_of + ); + K no_tac) + else unify_eresolve_tac HOPU.norms_unify HOPU.unify norms unify (Tactic.make_elim rule) ctxt + end + +fun unify_dresolve_tacs norms unify rules ctxt = + map (fn rule => unify_dresolve_tac norms unify rule ctxt) rules + +fun unify_dresolve_every_tac norms unify rules = unify_dresolve_tacs norms unify rules #> EVERY' + +fun unify_dresolve_any_tac norms unify rules = + unify_dresolve_tacs norms unify rules #> Tactic_Util.CONCAT' + +fun unify_fresolve_tac norms unify rule ctxt i state = + let val nprems = Thm.nprems_of rule + in + if nprems = 0 then + (@{log Logger.WARN} ctxt (fn _ => + Pretty.block [ + Pretty.str "forward-resolution rule ", + Thm.pretty_thm ctxt rule, + Pretty.str " has no premises." + ] |> Pretty.string_of + ); + Seq.empty) + else + unify_eresolve_tac_index HOPU.norms_unify HOPU.unify norms unify (Tactic.make_elim rule) + ctxt i state + |> Seq.map snd + end + +fun unify_fresolve_tacs norms unify rules ctxt = + map (fn rule => unify_fresolve_tac norms unify rule ctxt) rules + +fun unify_fresolve_every_tac norms unify rules = + unify_fresolve_tacs norms unify rules #> EVERY' + +fun unify_fresolve_any_tac norms unify rules = + unify_fresolve_tacs norms unify rules #> Tactic_Util.CONCAT' + +end diff --git a/thys/ML_Unification/Unification_Tactics/Unification_Tactics.thy b/thys/ML_Unification/Unification_Tactics/Unification_Tactics.thy new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/Unification_Tactics/Unification_Tactics.thy @@ -0,0 +1,13 @@ +\<^marker>\creator "Kevin Kappelmann"\ +section \Unification Tactics\ +theory Unification_Tactics + imports + Unify_Assumption_Tactic + Unify_Resolve_Tactics + Unify_Fact_Tactic +begin + +paragraph \Summary\ +text \Tactics with adjustable unifiers.\ + +end diff --git a/thys/ML_Unification/Unifiers/ML_Unifiers.thy b/thys/ML_Unification/Unifiers/ML_Unifiers.thy new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/Unifiers/ML_Unifiers.thy @@ -0,0 +1,68 @@ +\<^marker>\creator "Kevin Kappelmann"\ +section \ML Unifiers\ +theory ML_Unifiers + imports + ML_Unification_Base + ML_Functor_Instances + ML_Priorities + Simps_To +begin + +paragraph \Summary\ +text \Unification modulo equations and combinators for unifiers.\ + +paragraph \Combinators\ + +ML_file\unification_combinator.ML\ + +ML_file\unification_combine.ML\ +ML\ + @{functor_instance struct_name = Standard_Unification_Combine + and functor_name = Unification_Combine + and id = \""\} +\ +local_setup \Standard_Unification_Combine.setup_attribute NONE\ + + +paragraph \Standard Unifiers\ + +(*the following is a copy of unify.ML from the Isabelle distribution that includes the +patch 3c57995c255c (since it required here but not included in Isabelle2023). +FIXME: delete for next release version*) +ML_file\unify_copy.ML\ + +ML_file\higher_order_unification.ML\ +ML_file\higher_order_pattern_unification.ML\ +ML_file\first_order_unification.ML\ + + +paragraph \Unification via Tactics\ + +ML_file\tactic_unification.ML\ + +subparagraph \Unification via Simplification\ + +ML_file\simplifier_unification.ML\ + + +paragraph \Mixture of Unifiers\ + +ML_file\higher_ordern_pattern_first_decomp_unification.ML\ + +ML_file\mixed_unification.ML\ +ML\ + @{functor_instance struct_name = Standard_Mixed_Unification + and functor_name = Mixed_Unification + and id = \""\ + and more_args = \structure UC = Standard_Unification_Combine\} +\ + +declare [[ucombine add = \Standard_Unification_Combine.eunif_data + (Simplifier_Unification.simp_unify + |> Unification_Combinator.norm_closed_unifier + (#norm_term Standard_Mixed_Unification.norms_first_higherp_first_comb_higher_unify) + |> Unification_Combinator.unifier_from_closed_unifier + |> K) + (Standard_Unification_Combine.default_metadata \<^binding>\simp_unif\)\]] + +end diff --git a/thys/ML_Unification/Unifiers/first_order_unification.ML b/thys/ML_Unification/Unifiers/first_order_unification.ML new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/Unifiers/first_order_unification.ML @@ -0,0 +1,235 @@ +(* Title: ML_Unification/first_order_unification.ML + Author: Kevin Kappelmann, Paul Bachmann + +First-order E-unification. +*) +signature FIRST_ORDER_UNIFICATION = +sig + include HAS_LOGGER + + (* e-matching *) + val e_match : Unification_Base.type_matcher -> Unification_Base.e_matcher + val match : Unification_Base.matcher + val norms_match : Unification_Base.normalisers + + (* e-unification *) + val e_unify : Unification_Base.type_unifier -> Unification_Base.e_unifier + val unify : Unification_Base.unifier + val norms_unify : Unification_Base.normalisers +end + +structure First_Order_Unification : FIRST_ORDER_UNIFICATION = +struct + +val logger = Logger.setup_new_logger Unification_Base.logger "First_Order_Unification" + +(* shared utils *) +structure Util = Unification_Util +structure Norm = Envir_Normalisation + +exception FALLBACK + +(*check if sequence is empty or raise FALLBACK exception*) +fun seq_try sq = General_Util.seq_try FALLBACK sq + + +(* e-matching *) + +val norms_match = Util.eta_short_norms_match + +(*Note: by E-matching we mean "matching modulo equalities" in the general sense, +i.e. when matching p \? t, t may contain variables.*) +fun e_match match_types match_theory binders ctxt pt env = + let + val norm_pt = apfst o (#norm_term norms_match) + fun rigid_rigid ctxt p n env = (*normalise the types in rigid-rigid cases*) + Util.rigid_rigid Norm.norm_term_types_match match_types ctxt p n env |> seq_try + handle Unification_Base.UNIF => Seq.empty (*types do not match*) + (*standard first-order matcher that calls match_theory on failure; + generated theorem is already normalised wrt. the resulting environment*) + fun match binders ctxt (pt as (p, t)) (env as Envir.Envir {tenv, tyenv,...}) = + (case pt of + (Var (x, Tx), _) => + (case Envir.lookup1 tenv (x, Norm.norm_type_match tyenv Tx) of + NONE => + if Term.is_open t then + (@{log Logger.DEBUG} ctxt (fn _ => Pretty.block [ + Pretty.str "Failed to match (open term) ", + Util.pretty_unif_problem ctxt (norm_pt env pt) + ] |> Pretty.string_of); + (*directly return empty sequence because the theory unifier + cannot do anything meaningful at this point*) + raise FALLBACK) + else + let val Tt = fastype_of1 (Binders.binder_types binders, t) + in + ((match_types ctxt (Tx, Tt) env, Unification_Base.reflexive_term ctxt t) + (*insert new substitution x \ t*) + |>> Envir.update ((x, Tt), t) + |> Seq.single) + handle Unification_Base.UNIF => Seq.empty (*types do not match*) + end + | SOME p' => + if Envir.aeconv (apply2 Envir.beta_norm (p', t)) then + Seq.single (env, Unification_Base.reflexive_term ctxt t) + else raise FALLBACK) + | (Abs (np, Tp, tp), Abs (nt, Tt, tt)) => + ((let val name = if np = "" then nt else np + in + match_types ctxt (Tp, Tt) env + |> Util.abstract_abstract (K I) match binders ctxt name Tt (tp, tt) + |> seq_try + end) + handle Unification_Base.UNIF => Seq.empty) (*types do not match*) + (*eta-expand on the fly*) + | (Abs (np, Tp, tp), _) => + ((let val Ttarg = fastype_of1 (Binders.binder_types binders, t) |> Term.dest_funT |> fst + in + match_types ctxt (Tp, Ttarg) env + |> Util.abstract_abstract (K I) match binders ctxt np Ttarg + (tp, incr_boundvars 1 t $ Bound 0) + |> seq_try + end) + handle TYPE _ => (@{log Logger.DEBUG} ctxt (fn _ => Pretty.block [ + Pretty.str "First-order matching failed. Term is not of function type ", + Util.pretty_unif_problem ctxt (norm_pt env pt) + ] + |> Pretty.string_of); + Seq.empty) + | Unification_Base.UNIF => Seq.empty) (*types do not match*) + | (_, Abs (nt, Tt, tt)) => + ((let val Tp = Envir.fastype env (Binders.binder_types binders) p + in + match_types ctxt (Tp, fastype_of1 (Binders.binder_types binders, t)) env + |> Util.abstract_abstract (K I) match binders ctxt nt Tt + (incr_boundvars 1 p $ Bound 0, tt) + |> seq_try + end) + handle Unification_Base.UNIF => Seq.empty) (*types do not match*) + | (Bound i, Bound j) => Util.bound_bound binders ctxt (i, j) |> Seq.map (pair env) + | (Free _, Free g) => rigid_rigid ctxt p g env + | (Const _, Const d) => rigid_rigid ctxt p d env + | (f $ x, g $ y) => + (*Note: types of recursive theorems are already normalised ==> we have to + pass the identity type normaliser*) + Util.comb_comb (K o K I) (match binders) ctxt (f, x) (g, y) env |> seq_try + | _ => raise FALLBACK) + handle FALLBACK => + (@{log Logger.DEBUG} ctxt (fn _ => Pretty.block [ + Pretty.str "First-order matching failed. Calling theory matcher for ", + Util.pretty_unif_problem ctxt (norm_pt env pt) + ] |> Pretty.string_of); + match_theory binders ctxt pt env) + in + (@{log Logger.DEBUG} ctxt (fn _ => + Pretty.block [ + Pretty.str "First-order matching ", + Util.pretty_unif_problem ctxt (norm_pt env pt) + ] |> Pretty.string_of); + match binders ctxt pt env) + end + +(*standard first-order matching*) +val match = e_match Util.match_types Unification_Combinator.fail_match + +(* e-unification *) + +(*occurs check*) +fun occurs v (Var (x, _)) = x = v + | occurs v (s $ t) = occurs v s orelse occurs v t + | occurs v (Abs (_, _, t)) = occurs v t + | occurs _ _ = false + +val norms_unify = Util.eta_short_norms_unif + +fun e_unify unify_types unify_theory binders ctxt st env = + let + val norm_st = apply2 o (#norm_term norms_unify) + fun rigid_rigid ctxt t n env = (*we do not have to normalise types in rigid-rigid cases*) + Util.rigid_rigid (K I) unify_types ctxt t n env |> seq_try + handle Unification_Base.UNIF => Seq.empty (*types do not unify*) + (*standard first-order unifier that calls unify_theory on failure*) + fun unif binders ctxt (st as (s, t)) env = + (case st of + (Var (x, Tx), _) => + let val unif' = unif binders ctxt + in + (unif' (Envir.norm_term_same env s, t) env + handle Same.SAME => + (unif' (s, Envir.norm_term_same env t) env + handle Same.SAME => + if Term.is_open t then + (@{log Logger.DEBUG} ctxt (fn _ => Pretty.block [ + Pretty.str "Failed to unify (open term) ", + Util.pretty_unif_problem ctxt (norm_st env st) + ] + |> Pretty.string_of); + raise FALLBACK) + else + let + val vars_eq = is_Var t andalso x = fst (dest_Var t) + fun update_env env = + (*unifying x=x ==> no new term substitution necessary*) + if vars_eq then env + (*insert new substitution x \ t*) + else Envir.vupdate ((x, Tx), t) env + in + if not vars_eq andalso occurs x t then + (@{log Logger.DEBUG} ctxt (fn _ => Pretty.block [ + Pretty.str "Failed to unify (occurs check) ", + Util.pretty_unif_problem ctxt (norm_st env st) + ] |> Pretty.string_of); + raise FALLBACK) + else + ((unify_types ctxt (Tx, Envir.fastype env (Binders.binder_types binders) t) env, + Unification_Base.reflexive_term ctxt s) + |>> update_env + |> Seq.single) + handle Unification_Base.UNIF => Seq.empty (*types do not unify*) + end)) + end + | (_, Var _) => unif binders ctxt (t, s) env |> Seq.map (apsnd Unification_Base.symmetric) + | (Abs (ns, Ts, ts), Abs (nt, Tt, tt)) => + ((let val name = if ns = "" then nt else ns + in + unify_types ctxt (Ts, Tt) env + |> Util.abstract_abstract Norm.norm_term_types_unif unif binders ctxt name Ts (ts, tt) + |> seq_try + end) + handle Unification_Base.UNIF => Seq.empty) (*types do not unify*) + (*eta-expand on the fly*) + | (Abs (ns, Ts, ts), _) => + ((let val Tp = apply2 (Envir.fastype env (Binders.binder_types binders)) (s, t) + in + unify_types ctxt Tp env + |> Util.abstract_abstract Norm.norm_term_types_unif unif binders ctxt ns Ts + (ts, incr_boundvars 1 t $ Bound 0) + |> seq_try + end) + handle Unification_Base.UNIF => Seq.empty) (*types do not unify*) + | (_, Abs _) => unif binders ctxt (t, s) env |> Seq.map (apsnd Unification_Base.symmetric) + | (Bound i, Bound j) => Util.bound_bound binders ctxt (i, j) |> Seq.map (pair env) + | (Free _, Free g) => rigid_rigid ctxt s g env + | (Const _, Const d) => rigid_rigid ctxt s d env + (*we have to normalise types in comb cases*) + | (f $ x, g $ y) => + Util.comb_comb Norm.norm_thm_types_unif (unif binders) ctxt (f, x) (g, y) env |> seq_try + | _ => raise FALLBACK) + handle FALLBACK => + (@{log Logger.DEBUG} ctxt (fn _ => Pretty.block [ + Pretty.str "First-order unification failed. Calling theory unifier for ", + Util.pretty_unif_problem ctxt (norm_st env st) + ] |> Pretty.string_of); + unify_theory binders ctxt st env) + in + (@{log Logger.DEBUG} ctxt (fn _ => Pretty.block [ + Pretty.str "First-order unifying ", + Util.pretty_unif_problem ctxt (norm_st env st) + ] |> Pretty.string_of); + unif binders ctxt st env) + end + +(*standard first-order unification*) +val unify = e_unify Util.unify_types Unification_Combinator.fail_unify + +end diff --git a/thys/ML_Unification/Unifiers/higher_order_pattern_unification.ML b/thys/ML_Unification/Unifiers/higher_order_pattern_unification.ML new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/Unifiers/higher_order_pattern_unification.ML @@ -0,0 +1,482 @@ +(* Title: ML_Unification/higher_order_pattern_unification.ML + Author: Tobias Nipkow, Christine Heinzelmann, Stefan Berghofer, and + Kevin Kappelmann TU Muenchen + +Higher-Order Patterns due to Tobias Nipkow. +E-Unification and theorem construction due to Kevin Kappelmann. + +See also: +Tobias Nipkow. Functional Unification of Higher-Order Patterns. +In Proceedings of the 8th IEEE Symposium Logic in Computer Science, 1993. + +TODO: optimize red by special-casing it +*) +signature HIGHER_ORDER_PATTERN_UNIFICATION = +sig + include HAS_LOGGER + + (* e-matching *) + val e_match : Unification_Base.type_matcher -> Unification_Base.matcher -> + Unification_Base.e_matcher + val match : Unification_Base.matcher + val norms_match : Unification_Base.normalisers + + (* e-unification *) + val e_unify : Unification_Base.type_unifier -> Unification_Base.unifier -> + Unification_Base.e_unifier + val unify : Unification_Base.unifier + val norms_unify : Unification_Base.normalisers +end + +structure Higher_Order_Pattern_Unification : HIGHER_ORDER_PATTERN_UNIFICATION = +struct + +val logger = Logger.setup_new_logger Unification_Base.logger "Higher_Order_Pattern_Unification" + +(* shared utils *) +structure Util = Unification_Util +structure Norm = Envir_Normalisation + +exception FALLBACK + +(*check if sequence is empty or raise FALLBACK exception*) +fun seq_try sq = General_Util.seq_try FALLBACK sq + +(*predicate: downto0 (is, n) <=> is = [n, n - 1, ..., 0]*) +fun downto0 ([], n) = n = ~1 + | downto0 (i :: is, n) = i = n andalso downto0 (is, n - 1) + +fun mkabs (binders, is, t) = + let fun abstract i acc = + let val ((x, T), _) = nth binders i + in Abs (x, T, acc) end + in fold_rev abstract is t end + +exception IDX + +fun idx [] _ = raise IDX + | idx (i :: is) j = if i = j then length is else idx is j + +fun ints_of bs = + (*collects arguments and checks if they are all distinct, bound variables*) + let fun dest_check (Bound i) acc = + if member (op =) acc i then raise Unification_Base.PATTERN else i :: acc + | dest_check _ _ = raise Unification_Base.PATTERN + in fold_rev dest_check bs [] end + +fun app (s, []) = s + | app (s, (i :: is)) = app (s $ Bound i, is) + +(* matching *) + +fun mapbnd f = + let fun mpb d (Bound i) = if i < d then Bound i else Bound (f (i - d) + d) + | mpb d (Abs (s, T, t)) = Abs (s, T, mpb (d + 1) t) + | mpb d (u1 $ u2) = (mpb d u1) $ (mpb d u2) + | mpb _ atom = atom + in mpb 0 end + +fun red (Abs (_, _, s)) (i :: is) js = red s is (i :: js) + | red t [] [] = t + | red t is jn = app (mapbnd (nth jn) t, is) + +exception OPEN_TERM + +fun match_bind (tenv, binders, ixn, T, is, t) = + let val js = loose_bnos t + in + if null is then + if null js then Vartab.update_new (ixn, (T, t)) tenv + else raise OPEN_TERM + else if subset (op =) (js, is) then + let val t' = if downto0 (is, length binders - 1) then t else mapbnd (idx is) t + in Vartab.update_new (ixn, (T, mkabs (binders, is, t'))) tenv end + else raise OPEN_TERM + end + + +(** higher-order pattern E-unification **) +val norms_match = Util.beta_eta_short_norms_match + +fun e_match match_types match_theory_pattern match_theory_fail binders ctxt (pt as (p, t)) env = + let + val norm_pt = apfst o (#norm_term norms_match) + fun rigid_rigid ctxt p n env = (*normalise the types in rigid-rigid cases*) + Util.rigid_rigid Norm.norm_term_types_match match_types ctxt p n env |> seq_try + handle Unification_Base.UNIF => Seq.empty (*types do not match*) + (*generated theorem is already normalised wrt. the resulting environment*) + fun match binders ctxt (pt as (p, t)) (env as Envir.Envir {maxidx, tenv, tyenv}) = + let + (*call match_theory on failure*) + fun handle_failure match_theory failure_msg = + (@{log Logger.DEBUG} ctxt (fn _ => Pretty.breaks [ + failure_msg (), + Pretty.block [ + Pretty.str "Calling theory matcher for ", + Util.pretty_unif_problem ctxt (norm_pt env pt) + ] + ] |> Pretty.block |> Pretty.string_of); + match_theory binders ctxt pt env) + in + (case pt of + (Abs (np, _, tp), Abs (nt, Tt, tt)) => + let val name = if np = "" then nt else np + in + Util.abstract_abstract (K I) match binders ctxt name Tt (tp, tt) env + |> seq_try + end + (*eta-expand on the fly*) + | (Abs (np, Tp, tp), _) => + let val Tp' = Envir.subst_type tyenv Tp + in + Util.abstract_abstract (K I) match binders ctxt np Tp' + (tp, incr_boundvars 1 t $ Bound 0) env + |> seq_try + end + | (_, Abs (nt, Tt, tt)) => + Util.abstract_abstract (K I) match binders ctxt nt Tt + (incr_boundvars 1 p $ Bound 0, tt) env + |> seq_try + | (Bound i, Bound j) => Util.bound_bound binders ctxt (i, j) |> Seq.map (pair env) + | (Free _, Free g) => rigid_rigid ctxt p g env + | (Const _, Const d) => rigid_rigid ctxt p d env + | _ => case strip_comb p of + (Var (x, T), ps) => + let + val is = ints_of ps + val T' = Norm.norm_type_match tyenv T + in + case Envir.lookup1 tenv (x, T') of + NONE => + ((let + val tenv' = match_bind (tenv, binders, x, T', is, t) + val env' = Envir.Envir {maxidx=maxidx, tenv=tenv', tyenv=tyenv} + val t' = Binders.replace_binders binders t + in Seq.single (env', Unification_Base.reflexive_term ctxt t') end) + handle OPEN_TERM => (@{log Logger.DEBUG} ctxt (fn _ => Pretty.block [ + Pretty.str "Failed to unify (open term) ", + Util.pretty_unif_problem ctxt (norm_pt env pt) + ] |> Pretty.string_of); + Seq.empty)) + | SOME ph' => + if Envir.aeconv (red ph' is [], t) then + Seq.single (env, Unification_Base.reflexive_term ctxt + (Binders.replace_binders binders t)) + else raise FALLBACK + end + | (ph, ps) => + let val (th, ts) = strip_comb t + in case (ph, th) of + (Abs _, _) => raise Unification_Base.PATTERN + | (_, Abs _) => raise Unification_Base.PATTERN + | _ => if null ps orelse null ts then raise FALLBACK + else (*Note: types of theorems are already normalised + ==> we pass the identity type normaliser*) + Util.strip_comb_strip_comb (K o K I) match binders ctxt (ph, th) (ps, ts) env + |> seq_try + end) + handle FALLBACK => handle_failure match_theory_fail + (fn _ => Pretty.str "Higher-order pattern matching failed.") + | Unification_Base.PATTERN => handle_failure match_theory_pattern + (fn _ => Pretty.str "Terms not in Pattern fragment.") + end + val binder_types = Binders.binder_types binders + in + (@{log Logger.DEBUG} ctxt (fn _ => Pretty.block [ + Pretty.str "Higher-order pattern matching ", + Util.pretty_unif_problem ctxt (norm_pt env pt) + ] |> Pretty.string_of); + match_types ctxt (Envir.fastype env binder_types p, fastype_of1 (binder_types, t)) env + |> match binders ctxt pt) + handle Unification_Base.UNIF => Seq.empty (*types do not match*) + end + + +(*standard higher-order pattern matching*) +val match = e_match Util.match_types Unification_Combinator.fail_match + Unification_Combinator.fail_match + +(*unification*) + +fun string_of_term ctxt env binders t = + (map (Free o fst) binders, t) + |> subst_bounds + |> Norm.norm_term_unif env + |> Syntax.string_of_term ctxt + +fun bname binders = fst o fst o nth binders +fun bnames binders is = map (bname binders) is |> space_implode " " + +fun proj_fail ctxt (env, binders, F, _, is, t) reason = + @{log Logger.DEBUG} ctxt (fn _ => + let + val f = Term.string_of_vname F + val xs = bnames binders is + val u = string_of_term ctxt env binders t + in + cat_lines [ + "Cannot unify variable " ^ f ^ " (depending on bound variables [" ^ xs ^ "])", + "with term " ^ u, + reason () + ] + end) + +fun proj_fail_unif ctxt (params as (_, binders, _, _, is, t)) = + let + fun reason () = + let val ys = bnames binders (subtract (op =) is (loose_bnos t)) + in "Term contains additional bound variable(s) " ^ ys end + in proj_fail ctxt params reason end + +fun proj_fail_type_app ctxt (params as (env, binders, _, _, _, _)) var_app = + let + fun reason () = + let val var_app = string_of_term ctxt env binders var_app + in "Projection " ^ var_app ^ " is not well-typed." end + in proj_fail ctxt params reason end + +fun ocheck_fail ctxt (F, t, binders, env) = + @{log Logger.DEBUG} ctxt (fn _ => cat_lines [ + "Variable " ^ Term.string_of_vname F ^ " occurs in term", + string_of_term ctxt env binders t, + "Cannot unify!" + ]) + +fun occurs (F, t, env) = + let fun occ (Var (G, T)) = (case Envir.lookup env (G, T) of + SOME t => occ t + | NONE => F = G) + | occ (t1 $ t2) = occ t1 orelse occ t2 + | occ (Abs (_, _, t)) = occ t + | occ _ = false + in occ t end + +fun ints_of' env ts = ints_of (map (Envir.head_norm env) ts) + +(*split_type ([T1,....,Tn]---> T) n = ([Tn,...,T1], T)*) +fun split_type t n = + let fun split (T, 0, Ts) = (Ts, T) + | split (Type ("fun", [T1, T2]), n, Ts) = split (T2, n - 1, T1 :: Ts) + | split _ = raise Fail "split_type" + in split (t, n, []) end + +fun type_of_G (Envir.Envir {tyenv,...}) (T, n, is) = + let val (Ts, U) = split_type (Norm.norm_type_unif tyenv T) n + in map (nth Ts) is ---> U end + +fun mk_hnf (binders, is, G, js) = mkabs (binders, is, app (G, js)) + +fun mk_new_hnf (env, binders, is, F as (a, _), T, js) = + let val (env', G) = Envir.genvar a (env, type_of_G env (T, length is, js)) + in Envir.update ((F, T), mk_hnf (binders, is, G, js)) env' end + +(*mk_proj_list is = [ |is| - k - 1 | 0 <= k < |is| and is[k] >= 0 ]*) +fun mk_proj_list is = + let fun mk (SOME _) (acc, j) = (j :: acc, j + 1) + | mk NONE (acc, j) = (acc, j + 1) + in fold_rev mk is ([], 0) |> fst end + +fun proj ctxt (s, env, binders, is) = + let + fun trans d i = if i < d then i else idx is (i - d) + d + fun pr (s, env, d, binders) = (case Envir.head_norm env s of + Abs (a, T, t) => + let + val (binder, _) = Binders.fix_binder (a, T) ctxt + val (t', env') = pr (t, env, d + 1, binder :: binders) + in (Abs (a, T, t'), env') end + | t => (case strip_comb t of + (c as Const _, ts) => + let val (ts', env') = prs (ts, env, d, binders) + in (list_comb (c, ts'), env') end + | (f as Free _, ts) => + let val (ts', env') = prs (ts, env, d, binders) + in (list_comb (f, ts'), env') end + | (Bound i, ts) => + let + val j = trans d i + val (ts', env') = prs (ts, env, d, binders) + in (list_comb (Bound j, ts'), env') end + | (Var (F as (a, _), Fty), ts) => + let + val js = ints_of' env ts + val js' = map (try (trans d)) js + val ks = mk_proj_list js' + val ls = map_filter I js' + val Hty = type_of_G env (Fty, length js, ks) + val (env', H) = Envir.genvar a (env, Hty) + val env'' = Envir.update ((F, Fty), mk_hnf (binders, js, H, ks)) env' + in (app (H, ls), env'') end + | _ => raise Unification_Base.PATTERN)) + and prs (s :: ss, env, d, binders) = + let + val (s', env1) = pr (s, env, d, binders) + val (ss', env2) = prs (ss, env1, d, binders) + in (s' :: ss', env2) end + | prs ([], env, _, _) = ([], env) + in + if downto0 (is, length binders - 1) then (s, env) + else pr (s, env, 0, binders) + end + +(*mk_ff_list (is, js) = [ |is| - k - 1 | 0 <= k < |is| and is[k] = js[k] ]*) +fun mk_ff_list (is,js) = + let fun mk ([], [], _) = [] + | mk (i :: is, j :: js, k) = + if i = j then k :: mk (is, js, k - 1) + else mk (is, js, k - 1) + | mk _ = raise Fail "mk_ff_list" + in mk (is, js, length is - 1) end; + +fun app_free (Envir.Envir {tyenv,...}) binders n T is = + let val norm_type = Norm.norm_type_unif tyenv + in list_comb (Var (n, norm_type T), map (map_types norm_type o Binders.nth_binder_data binders) is) end + +fun flexflex1 unify_types ctxt (env, binders, F, Fty, Fty', is, js) = + let + val env' = unify_types ctxt (Fty, Fty') env + val thm = app_free env binders F Fty is |> Unification_Base.reflexive_term ctxt + in + if is = js then (env', thm) + else + let + val ks = mk_ff_list (is, js) + val env'' = mk_new_hnf (env', binders, is, F, Fty, ks) + in (env'', thm) end + end + +fun flexflex2 unify_types ctxt (env, binders, F, Fty, is, G, Gty, js) = + let + val var_app = app_free env binders F Fty is + val binder_types = Binders.binder_types binders + val env' = unify_types ctxt + (apply2 (Envir.fastype env binder_types) (var_app, app (Var (G, Gty), js))) env + val thm = Unification_Base.reflexive_term ctxt var_app + fun ff (F, Fty, is, G as (a, _), Gty, js) = + if subset (op =) (js, is) then + let + val t = mkabs (binders, is, app (Var (G, Gty), map (idx is) js)) + val env'' = Envir.update ((F, Fty), t) env' + in (env'', thm) end + else + let + val ks = inter (op =) js is + val Hty = type_of_G env' (Fty, length is, map (idx is) ks) + val (env'', H) = Envir.genvar a (env', Hty) + fun lam is = mkabs (binders, is, app (H, map (idx is) ks)) + val env''' = + Envir.update ((F, Fty), lam is) env'' + |> Envir.update ((G, Gty), lam js) + in (env''', thm) end + in + if is_less (Term_Ord.indexname_ord (G, F)) then ff (F, Fty, is, G, Gty, js) + else ff (G, Gty, js, F, Fty, is) + end + +fun flexrigid unify_types ctxt (params as (env, binders, F, Fty, is, t)) = + if occurs (F, t, env) then (ocheck_fail ctxt (F, t, binders, env); raise FALLBACK) + else + let val var_app = app_free env binders F Fty is + in + let + val binder_types = Binders.binder_types binders + val env' = unify_types ctxt + (apply2 (Envir.fastype env binder_types) (var_app, t)) env + val (u, env'') = proj ctxt (t, env', binders, is) + val env''' = Envir.update ((F, Fty), mkabs (binders, is, u)) env'' + val thm = Unification_Base.reflexive_term ctxt var_app + in (env''', thm) end + handle IDX => (proj_fail_unif ctxt params; raise OPEN_TERM) + | TYPE _ => (proj_fail_type_app ctxt params var_app; raise Unification_Base.UNIF) + end + +(** higher-order pattern E-unification **) +val norms_unify = Util.beta_eta_short_norms_unif + +fun e_unify unify_types unify_theory_pattern unify_theory_fail binders ctxt st env = + let + fun unif binders ctxt st env = + let + val (st' as (s', t')) = apply2 (Envir.head_norm env) st + fun rigid_rigid ctxt t n env = (*we do not normalise types in base cases*) + Util.rigid_rigid (K I) unify_types ctxt t n env |> seq_try + handle Unification_Base.UNIF => Seq.empty (*types do not match*) + (*call unify_theory on failure*) + fun handle_failure unify_theory failure_msg = + (@{log Logger.DEBUG} ctxt (fn _ => Pretty.breaks [ + failure_msg (), + Pretty.block [ + Pretty.str "Calling theory unifier for ", + Util.pretty_unif_problem ctxt st' + ] + ] |> Pretty.block |> Pretty.string_of); + (*Note: st' is already normalised*) + unify_theory binders ctxt st' env) + in + (case st' of + (Abs (ns, Ts, ts), Abs (nt, Tt, tt)) => + ((let val name = if ns = "" then nt else ns + in + unify_types ctxt (Ts, Tt) env + |> Util.abstract_abstract Norm.norm_term_types_unif unif binders ctxt name Ts (ts, tt) + |> seq_try + end) + handle Unification_Base.UNIF => Seq.empty) (*types do not unify*) + (*eta-expand on the fly*) + | (Abs (ns, Ts, ts), _) => + ((let val st = apply2 (Envir.fastype env (Binders.binder_types binders)) st' + in + unify_types ctxt st env + |> Util.abstract_abstract Norm.norm_term_types_unif unif binders ctxt ns Ts + (ts, incr_boundvars 1 t' $ Bound 0) + |> seq_try + end) + handle Unification_Base.UNIF => Seq.empty) (*types do not unify*) + | (_, Abs _) => unif binders ctxt (t', s') env |> Seq.map (apsnd Unification_Base.symmetric) + | (Bound i, Bound j) => Util.bound_bound binders ctxt (i, j) |> Seq.map (pair env) + | (Free _, Free g) => rigid_rigid ctxt s' g env + | (Const _, Const d) => rigid_rigid ctxt s' d env + (*case distinctions on head term*) + | _ => case apply2 strip_comb st' of + ((Var (F, Fty), ss), (Var (G, Gty), ts)) => + (((if F = G then + flexflex1 unify_types ctxt (env, binders, F, Fty, Gty, ints_of' env ss, ints_of' env ts) + else + flexflex2 unify_types ctxt (env, binders, F, Fty, ints_of' env ss, G, Gty, ints_of' env ts)) + |> Seq.single) + handle Unification_Base.UNIF => Seq.empty) (*types do not unify*) + | ((Var (F, Fty), ss), _) => + (flexrigid unify_types ctxt (env, binders, F, Fty, ints_of' env ss, t')|> Seq.single + handle OPEN_TERM => Seq.empty + | Unification_Base.UNIF => Seq.empty) (*types do not unify*) + | (_, (Var (F, Fty), ts)) => + (flexrigid unify_types ctxt (env, binders, F, Fty, ints_of' env ts, s') |> Seq.single + handle OPEN_TERM => Seq.empty + | Unification_Base.UNIF => Seq.empty) (*types do not unify*) + | ((sh, ss), (th, ts)) => if null ss orelse null ts then raise FALLBACK + else (*but we have to normalise in comb cases*) + Util.strip_comb_strip_comb Norm.norm_thm_types_unif unif binders ctxt + (sh, th) (ss, ts) env + |> seq_try) + handle FALLBACK => handle_failure unify_theory_fail + (fn _ => Pretty.str "Higher-order pattern unification failed.") + | Unification_Base.PATTERN => handle_failure unify_theory_pattern + (fn _ => Pretty.str "Terms not in Pattern fragment.") + end + val binder_types = Binders.binder_types binders + in + (@{log Logger.DEBUG} ctxt (fn _ => + Pretty.block [ + Pretty.str "Higher-order pattern unifying ", + Util.pretty_unif_problem ctxt (apply2 (Norm.norm_term_unif env) st) + ] + |> Pretty.string_of); + unif binders ctxt st (unify_types ctxt (apply2 (Envir.fastype env binder_types) st) env)) + handle Unification_Base.UNIF => Seq.empty (*types do not unify*) + end + +(*standard higher-order pattern unification*) +val unify = e_unify Util.unify_types Unification_Combinator.fail_unify + Unification_Combinator.fail_unify + +end diff --git a/thys/ML_Unification/Unifiers/higher_order_unification.ML b/thys/ML_Unification/Unifiers/higher_order_unification.ML new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/Unifiers/higher_order_unification.ML @@ -0,0 +1,61 @@ +(* Title: ML_Unification/higher_order_unification.ML + Author: Kevin Kappelmann TU Muenchen + +Higher-order unification from the Pure kernel adjusted to fit the +Unification_Base.unifier type, i.e. moving flex-flex pairs into the theorem +and support of open term unification. +*) + +signature HIGHER_ORDER_UNIFICATION = +sig + include HAS_LOGGER + + val unify : Unification_Base.unifier + val norms : Unification_Base.normalisers +end + +structure Higher_Order_Unification : HIGHER_ORDER_UNIFICATION = +struct + +val logger = Logger.setup_new_logger Unification_Base.logger "Higher_Order_Unification" + +structure Util = Unification_Util + +fun unify binders ctxt tp env = + let + fun unif env = + let + val init_goal = Logic.mk_equals #> Thm.cterm_of ctxt #> Goal.init + fun create_env_thmq env = + (*replace binders and apply the computed unifier*) + apply2 (Envir_Normalisation.beta_eta_short_norm_term_unif env o Binders.replace_binders binders) tp + (*create a goal such that we can create a theorem including the flex-flex pairs*) + |> init_goal + (*kind of a hack: resolving against the equality theorem will add the + flex-flex pairs to the theorem*) + |> HEADGOAL (match_tac ctxt @{thms Pure.reflexive}) + (*conclude the goal and pair with the environment*) + |> Seq.map (pair env o Goal.conclude) + in + (*find the unifiers*) + Unify_Copy.hounifiers (map fst binders) (Context.Proof ctxt, env, [tp]) + (*FIXME: replace above line by the following one for the next Isabelle release*) + (*Unify.hounifiers (map fst binders) (Context.Proof ctxt, env, [tp])*) + (*add the flex-flex pairs to the theorem*) + |> Seq.maps (create_env_thmq o fst) + end + in + ((@{log Logger.DEBUG} ctxt (fn _ => + Pretty.block [ + Pretty.str "Higher-order unifying ", + Util.pretty_unif_problem ctxt (apply2 (Envir.norm_term env) tp) + ] + |> Pretty.string_of); + unif (Util.unify_types ctxt (apply2 Term.fastype_of tp) env)) + handle Unification_Base.UNIF => Seq.empty) (*types do not match*) + end + +val norms = Unification_Util.beta_eta_short_norms_unif + +end + diff --git a/thys/ML_Unification/Unifiers/higher_ordern_pattern_first_decomp_unification.ML b/thys/ML_Unification/Unifiers/higher_ordern_pattern_first_decomp_unification.ML new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/Unifiers/higher_ordern_pattern_first_decomp_unification.ML @@ -0,0 +1,150 @@ +(* Title: ML_Unification/higher_ordern_pattern_first_decomp_unification.ML + Author: Kevin Kappelmann + +Higher-order pattern matching/unification with decomposition of non-pattern terms into a +higher-order pattern prefix and a list of remaining arguments. This is essentially +a combination of higher-order pattern and first-order matching/unification. + +Example: Let 0, 1 be bound variables and f, g, c be constants. +The unification problem ?x 0 1 (?y c 0) \\<^sup>? f (g c 0) +is decomposed into ?x 0 1 \\<^sup>? f and (?y c 0) \\<^sup>? (g c 0). +The latter, in turn, is decomposed into ?y \\<^sup>? g and c \\<^sup>? c and 0 \\<^sup>? 0. +*) +signature HIGHER_ORDERN_PATTERN_FIRST_DECOMP_UNIFICATION = +sig + include HAS_LOGGER + + val e_match : Unification_Base.type_matcher -> Unification_Base.matcher -> + Unification_Base.e_matcher + val match : Unification_Base.matcher + val norms_match : Unification_Base.normalisers + + val e_unify : Unification_Base.type_unifier -> Unification_Base.unifier -> + Unification_Base.e_unifier + val unify : Unification_Base.unifier + val norms_unify : Unification_Base.normalisers +end + +structure Higher_Ordern_Pattern_First_Decomp_Unification : + HIGHER_ORDERN_PATTERN_FIRST_DECOMP_UNIFICATION = +struct + +val logger = Logger.setup_new_logger Unification_Base.logger + "Higher_Ordern_Pattern_First_Decomp_Unification" + +structure UUtil = Unification_Util +structure HOP = Higher_Order_Pattern_Unification +structure UC = Unification_Combinator + +(*compute pattern prefixes*) +fun bounds_prefix bounds [] = (rev bounds, []) + | bounds_prefix bounds (t :: ts) = + if is_Bound t andalso not (member (op =) bounds t) + then bounds_prefix (t :: bounds) ts + else (rev bounds, t :: ts) +val bounds_prefix = bounds_prefix [] + +datatype ('a, 'b) either = Left of 'a | Right of 'b + +fun bounds_prefix2 (bounds1, _) ([], _) = Left (rev bounds1, []) + | bounds_prefix2 (_, bounds2) (_, []) = Right (rev bounds2, []) + | bounds_prefix2 (bounds1, bounds2) (s :: ss, t :: ts) = + if is_Bound s andalso not (member (op =) bounds1 s) + then if is_Bound t andalso not (member (op =) bounds2 t) + then bounds_prefix2 (s :: bounds1, t :: bounds2) (ss, ts) + else Right (rev bounds2, t :: ts) + else Left (rev bounds1, s :: ss) +val bounds_prefix2 = bounds_prefix2 ([], []) + +fun mk_decomp _ (_, []) = NONE (*no decomposition took place*) + | mk_decomp h (args, rem) = SOME (list_comb (h, args), rem) + +fun mk_bounds_decomp h args = bounds_prefix args |> mk_decomp h + +fun pattern_prefix_match p = case strip_comb p of + (v as Var _, args) => mk_bounds_decomp v args + | _ => NONE + +fun pattern_prefix_unif tp = case apply2 strip_comb tp of + ((vl as Var _, argsl), (vr as Var _, argsr)) => (case bounds_prefix2 (argsl, argsr) of + Left res => mk_decomp vl res |> Option.map Left + | Right res => mk_decomp vr res |> Option.map Right) + | ((v as Var _, args), _) => mk_bounds_decomp v args |> Option.map Left + | (_, (v as Var _, args)) => mk_bounds_decomp v args |> Option.map Right + | _ => NONE + +fun mk_decomp_other ss t = + let + val (th, ts) = strip_comb t + val n = length ts - length ss + in if n < 0 then NONE else chop n ts |> mk_decomp th end + +fun decompose_msg ctxt tp = Pretty.block [ + Pretty.str "Decomposing ", + UUtil.pretty_unif_problem ctxt tp, + Pretty.str " into higher-order pattern prefix and list of remaining arguments." + ] |> Pretty.string_of + +fun decomposed_msg ctxt (ph, th) (ps, ts) = Pretty.block [ + Pretty.str "Decomposition result ", + map (UUtil.pretty_unif_problem ctxt) ((ph, th) :: (ps ~~ ts)) + |> Pretty.separate "," |> Pretty.block + ] |> Pretty.string_of + +fun e_match match_types match_pattern match_theory = + let fun fallback_pattern binders ctxt (pt as (p, t)) = + let fun failure () = + (@{log Logger.DEBUG} ctxt (fn _ => "Decomposition failed. Calling theory matcher."); + match_theory binders ctxt pt) + in + (@{log Logger.DEBUG} ctxt (fn _ => decompose_msg ctxt pt); + case pattern_prefix_match p of + SOME (ph, ps) => (case mk_decomp_other ps t of + SOME (th, ts) => + (@{log Logger.TRACE} ctxt (fn _ => decomposed_msg ctxt (ph, th) (ps, ts)); + UUtil.strip_comb_strip_comb (K o K I) match_pattern + binders ctxt (ph, th) (ps, ts)) + | NONE => failure ()) + | NONE => failure ()) + end + in + HOP.e_match match_types (UC.norm_matcher (#norm_term HOP.norms_match) fallback_pattern) + match_theory + end + +val match = e_match UUtil.match_types UC.fail_match UC.fail_match +val norms_match = HOP.norms_match + +fun e_unify unif_types unif_pattern unif_theory = + let + val unif_patterns = + UUtil.strip_comb_strip_comb Envir_Normalisation.norm_thm_types_unif unif_pattern + fun fallback_pattern binders ctxt (tp as (s, t)) = + let fun failure () = + (@{log Logger.DEBUG} ctxt (fn _ => "Decomposition failed. Calling theory unifier."); + unif_theory binders ctxt tp) + in + (@{log Logger.DEBUG} ctxt (fn _ => decompose_msg ctxt tp); + case pattern_prefix_unif tp of + SOME (Left (sh, ss)) => (case mk_decomp_other ss t of + SOME (th, ts) => + (@{log Logger.TRACE} ctxt (fn _ => decomposed_msg ctxt (sh, th) (ss, ts)); + unif_patterns binders ctxt (sh, th) (ss, ts)) + | NONE => failure ()) + | SOME (Right (th, ts)) => (case mk_decomp_other ts s of + SOME (sh, ss) => + (@{log Logger.TRACE} ctxt (fn _ => decomposed_msg ctxt (sh, th) (ss, ts)); + unif_patterns binders ctxt (sh, th) (ss, ts)) + | NONE => failure ()) + | NONE => failure ()) + end + in + HOP.e_unify unif_types (UC.norm_unifier (#norm_term HOP.norms_unify) fallback_pattern) + unif_theory + end + +val unify = e_unify UUtil.unify_types UC.fail_unify UC.fail_unify + +val norms_unify = HOP.norms_unify + +end diff --git a/thys/ML_Unification/Unifiers/mixed_unification.ML b/thys/ML_Unification/Unifiers/mixed_unification.ML new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/Unifiers/mixed_unification.ML @@ -0,0 +1,52 @@ +(* Title: ML_Unification/mixed_unification.ML + Author: Kevin Kappelmann + +Mixture of unification algorithms. +*) +signature MIXED_UNIFICATION = +sig + include HAS_LOGGER + + structure UC : UNIFICATION_COMBINE + + (*first-order, then higher-order pattern with first-order decomposition, then UC.e_unify, + then higher-order unification fallback*) + val first_higherp_first_comb_higher_unify : Unification_Base.unifier + val norms_first_higherp_first_comb_higher_unify : Unification_Base.normalisers +end + +functor Mixed_Unification (A : sig + structure FIA : FUNCTOR_INSTANCE_ARGS + structure UC : UNIFICATION_COMBINE + end) : MIXED_UNIFICATION = +struct + +val logger = Logger.setup_new_logger Unification_Base.logger A.FIA.full_name + +structure UUtil = Unification_Util +structure UCO = Unification_Combinator +structure UC = A.UC + +fun first_higherp_first_comb_higher_unify binders ctxt tp env = + let + val unify_types = UUtil.unify_types + val comb_higher = UCO.add_fallback_unifier UC.e_unify Higher_Order_Unification.unify + val higherp_comb_higher = UCO.add_fallback_unifier + (Higher_Ordern_Pattern_First_Decomp_Unification.e_unify unify_types + first_higherp_first_comb_higher_unify) + comb_higher + val fo_higherp_comb_higher = UCO.add_fallback_unifier + (First_Order_Unification.e_unify unify_types) higherp_comb_higher + in + (@{log Logger.DEBUG} ctxt (fn _ => Pretty.block [ + Pretty.str "First-order with higher-order pattern with first-order decomposition with ", + Binding.pretty UC.binding, + Pretty.str " with higher-order fallback unifying ", + UUtil.pretty_unif_problem ctxt (apply2 (Envir_Normalisation.norm_term_unif env) tp)] + |> Pretty.string_of); + fo_higherp_comb_higher binders ctxt tp env) + end + +val norms_first_higherp_first_comb_higher_unify = UUtil.beta_eta_short_norms_unif + +end diff --git a/thys/ML_Unification/Unifiers/simplifier_unification.ML b/thys/ML_Unification/Unifiers/simplifier_unification.ML new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/Unifiers/simplifier_unification.ML @@ -0,0 +1,74 @@ +(* Title: ML_Unification/simplifier_unification.ML + Author: Kevin Kappelmann + +Solving equations for unification problems with the simplifier. +*) +signature SIMPLIFIER_UNIFICATION = +sig + include HAS_LOGGER + (*solves "s \\<^sup>? t" via simplification*) + val simp_unify : Unification_Base.closed_unifier + (*solves "SIMPS_TO s t \ rhs" via simplification of s when given a matching theorem + "SIMPS_TO s t \ SIMPS_TO s t \ rhs"*) + val SIMPS_TO_unify : thm -> Unification_Base.closed_unifier + (*solves "SIMPS_TO_UNIF s t \ rhs" via simplification of s to s', followed by unification of + "s' \\<^sup>? t", when given a matching theorem "SIMPS_TO_UNIF s t \ SIMPS_TO_UNIF s t \ rhs"*) + val SIMPS_TO_UNIF_unify : thm -> Unification_Base.normalisers -> Unification_Base.e_unifier +end + +structure Simplifier_Unification : SIMPLIFIER_UNIFICATION = +struct + +val logger = Logger.setup_new_logger Unification_Base.logger "Simplifier_Unification" + +(*some "safe" solvers create instantiations via flex-flex pairs, which we disallow*) +val safe_simp_tac = Tactic_Util.safe_simp_tac + +fun simp_unify ctxt tp = + (@{log Logger.DEBUG} ctxt (fn _ => Pretty.block [ + Pretty.str "Solving unification problem via simplification ", + Unification_Util.pretty_unif_problem ctxt tp + ] |> Pretty.string_of); + Tactic_Unification.unify (Tactic_Unification.env_tac_from_no_inst_tac (safe_simp_tac ctxt) |> K) + ctxt tp) + +fun preprocess_tac ctxt = match_tac ctxt o single + +fun SIMPS_TO_unify preprocess_rule ctxt = + let fun tac (tp as (lhs, _)) = + if can Simps_To.dest_SIMPS_TO lhs + then + (@{log Logger.DEBUG} ctxt (fn _ => Pretty.block [ + Pretty.str "Solving ", + Syntax.pretty_term ctxt @{term SIMPS_TO}, + Pretty.str " unification problem ", + Unification_Util.pretty_unif_problem ctxt tp + ] |> Pretty.string_of); + preprocess_tac ctxt preprocess_rule + THEN' safe_simp_tac ctxt + THEN' Simps_To.finish_SIMPS_TO_tac ctxt) + else K no_tac + in Tactic_Unification.unify (Tactic_Unification.env_tac_from_no_inst_tac o tac) ctxt end + +fun SIMPS_TO_UNIF_unify preprocess_rule norms unif binders ctxt = + let fun tac ctxt (tp as (lhs, _)) i (env, state) = + (let + val simps_to_tp = Simps_To_Unif.dest_SIMPS_TO_UNIF lhs + fun solve_tac state = + Simps_To_Unif.SIMPS_TO_UNIF_env_thmsq (safe_simp_tac ctxt) norms unif binders ctxt + simps_to_tp env + |> Seq.map (fn (env, thm) => (env, #norm_thm norms ctxt env state |> Thm.elim_implies thm)) + in + (@{log Logger.DEBUG} ctxt (fn _ => Pretty.block [ + Pretty.str "Solving ", + Syntax.pretty_term ctxt @{term SIMPS_TO_UNIF}, + Pretty.str " unification problem ", + Unification_Util.pretty_unif_problem ctxt tp + ] |> Pretty.string_of); + preprocess_tac ctxt preprocess_rule i state + |> Seq.maps solve_tac) + end) + handle TERM _ => Seq.empty + in Tactic_Unification.unify (tac ctxt) ctxt end + +end diff --git a/thys/ML_Unification/Unifiers/tactic_unification.ML b/thys/ML_Unification/Unifiers/tactic_unification.ML new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/Unifiers/tactic_unification.ML @@ -0,0 +1,29 @@ +(* Title: ML_Unification/tactic_unification.ML + Author: Kevin Kappelmann + +Solving equations for unification problems with environment-aware tactics. +*) +signature TACTIC_UNIFICATION = +sig + type env_tactic = Envir.env * thm -> (Envir.env * thm) Seq.seq + (*create environment-aware tactic from a tactic that *DOES NOT* instantiate meta variables*) + val env_tac_from_no_inst_tac : (int -> tactic) -> int -> env_tactic + + val unify : (term * term -> int -> env_tactic) -> Unification_Base.closed_unifier +end + +structure Tactic_Unification : TACTIC_UNIFICATION = +struct + +type env_tactic = Envir.env * thm -> (Envir.env * thm) Seq.seq +fun env_tac_from_no_inst_tac tac i (env, state) = tac i state |> Seq.map (pair env) + +fun unify tac ctxt tp env = + (Logic.mk_equals tp + |> Thm.cterm_of ctxt + |> Goal.init + |> Tactic_Util.HEADGOAL (tac tp) o pair env + |> Seq.map_filter (try (apsnd (Goal.finish ctxt)))) + handle TYPE _ => Seq.empty + +end diff --git a/thys/ML_Unification/Unifiers/test.ML b/thys/ML_Unification/Unifiers/test.ML new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/Unifiers/test.ML @@ -0,0 +1,658 @@ +(* Title: Pure/unify.ML + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright Cambridge University 1992 + +Higher-Order Unification. + +Types as well as terms are unified. The outermost functions assume +the terms to be unified already have the same type. In resolution, +this is assured because both have type "prop". +*) + +signature UNIFY = +sig + val trace_bound: int Config.T + val search_bound: int Config.T + val trace_simp: bool Config.T + val trace_types: bool Config.T + val hounifiers: (string * typ) list -> Context.generic * Envir.env * ((term * term) list) -> + (Envir.env * (term * term) list) Seq.seq + val unifiers: Context.generic * Envir.env * ((term * term) list) -> + (Envir.env * (term * term) list) Seq.seq + val smash_unifiers: Context.generic -> (term * term) list -> Envir.env -> Envir.env Seq.seq +end + +structure Unify : UNIFY = +struct + +(*Unification options*) + +(*tracing starts above this depth, 0 for full*) +val trace_bound = Config.declare_int ("unify_trace_bound", \<^here>) (K 50); + +(*unification quits above this depth*) +val search_bound = Config.declare_int ("unify_search_bound", \<^here>) (K 60); + +(*print dpairs before calling SIMPL*) +val trace_simp = Config.declare_bool ("unify_trace_simp", \<^here>) (K false); + +(*announce potential incompleteness of type unification*) +val trace_types = Config.declare_bool ("unify_trace_types", \<^here>) (K false); + + +type binderlist = (string * typ) list; + +type dpair = binderlist * term * term; + +fun fastype env (Ts, t) = Envir.fastype env (map snd Ts) t; + + +(* eta normal form *) + +fun eta_norm env = + let + val tyenv = Envir.type_env env; + fun etif (Type ("fun", [T, U]), t) = + Abs ("", T, etif (U, incr_boundvars 1 t $ Bound 0)) + | etif (TVar v, t) = + (case Type.lookup tyenv v of + NONE => t + | SOME T => etif (T, t)) + | etif (_, t) = t; + fun eta_nm (rbinder, Abs (a, T, body)) = + Abs (a, T, eta_nm ((a, T) :: rbinder, body)) + | eta_nm (rbinder, t) = etif (fastype env (rbinder, t), t); + in eta_nm end; + + +(*OCCURS CHECK + Does the uvar occur in the term t? + two forms of search, for whether there is a rigid path to the current term. + "seen" is list of variables passed thru, is a memo variable for sharing. + This version searches for nonrigid occurrence, returns true if found. + Since terms may contain variables with same name and different types, + the occurs check must ignore the types of variables. This avoids + that ?x::?'a is unified with f(?x::T), which may lead to a cyclic + substitution when ?'a is instantiated with T later. *) +fun occurs_terms (seen: indexname list Unsynchronized.ref, + env: Envir.env, v: indexname, ts: term list): bool = + let + fun occurs [] = false + | occurs (t :: ts) = occur t orelse occurs ts + and occur (Const _) = false + | occur (Bound _) = false + | occur (Free _) = false + | occur (Var (w, T)) = + if member (op =) (!seen) w then false + else if Term.eq_ix (v, w) then true + (*no need to lookup: v has no assignment*) + else + (seen := w :: !seen; + case Envir.lookup env (w, T) of + NONE => false + | SOME t => occur t) + | occur (Abs (_, _, body)) = occur body + | occur (f $ t) = occur t orelse occur f; + in occurs ts end; + + +(* f a1 ... an ----> f using the assignments*) +fun head_of_in env t = + (case t of + f $ _ => head_of_in env f + | Var vT => + (case Envir.lookup env vT of + SOME u => head_of_in env u + | NONE => t) + | _ => t); + + +datatype occ = NoOcc | Nonrigid | Rigid; + +(* Rigid occur check +Returns Rigid if it finds a rigid occurrence of the variable, + Nonrigid if it finds a nonrigid path to the variable. + NoOcc otherwise. + Continues searching for a rigid occurrence even if it finds a nonrigid one. + +Condition for detecting non-unifable terms: [ section 5.3 of Huet (1975) ] + a rigid path to the variable, appearing with no arguments. +Here completeness is sacrificed in order to reduce danger of divergence: + reject ALL rigid paths to the variable. +Could check for rigid paths to bound variables that are out of scope. +Not necessary because the assignment test looks at variable's ENTIRE rbinder. + +Treatment of head(arg1,...,argn): +If head is a variable then no rigid path, switch to nonrigid search +for arg1,...,argn. +If head is an abstraction then possibly no rigid path (head could be a + constant function) so again use nonrigid search. Happens only if + term is not in normal form. + +Warning: finds a rigid occurrence of ?f in ?f(t). + Should NOT be called in this case: there is a flex-flex unifier +*) +fun rigid_occurs_term (seen: indexname list Unsynchronized.ref, env, v: indexname, t) = + let + fun nonrigid t = + if occurs_terms (seen, env, v, [t]) then Nonrigid + else NoOcc + fun occurs [] = NoOcc + | occurs (t :: ts) = + (case occur t of + Rigid => Rigid + | oc => (case occurs ts of NoOcc => oc | oc2 => oc2)) + and occomb (f $ t) = + (case occur t of + Rigid => Rigid + | oc => (case occomb f of NoOcc => oc | oc2 => oc2)) + | occomb t = occur t + and occur (Const _) = NoOcc + | occur (Bound _) = NoOcc + | occur (Free _) = NoOcc + | occur (Var (w, T)) = + if member (op =) (!seen) w then NoOcc + else if Term.eq_ix (v, w) then Rigid + else + (seen := w :: !seen; + case Envir.lookup env (w, T) of + NONE => NoOcc + | SOME t => occur t) + | occur (Abs (_, _, body)) = occur body + | occur (t as f $ _) = (*switch to nonrigid search?*) + (case head_of_in env f of + Var (w,_) => (*w is not assigned*) + if Term.eq_ix (v, w) then Rigid + else nonrigid t + | Abs _ => nonrigid t (*not in normal form*) + | _ => occomb t) + in occur t end; + + +exception CANTUNIFY; (*Signals non-unifiability. Does not signal errors!*) +exception ASSIGN; (*Raised if not an assignment*) + + +fun unify_types context TU env = + Pattern.unify_types context TU env handle Pattern.Unif => raise CANTUNIFY; + +fun test_unify_types context (T, U) env = + let + fun trace () = + if Context_Position.is_visible_generic context then + let val str_of = Syntax.string_of_typ (Context.proof_of context) + in tracing ("Potential loss of completeness: " ^ str_of U ^ " = " ^ str_of T) end + else (); + val env' = unify_types context (T, U) env; + in if is_TVar T orelse is_TVar U then trace () else (); env' end; + +(*Is the term eta-convertible to a single variable with the given rbinder? + Examples: ?a ?f(B.0) ?g(B.1,B.0) + Result is var a for use in SIMPL. *) +fun get_eta_var ([], _, Var vT) = vT + | get_eta_var (_::rbinder, n, f $ Bound i) = + if n = i then get_eta_var (rbinder, n + 1, f) + else raise ASSIGN + | get_eta_var _ = raise ASSIGN; + + +(*Solve v=u by assignment -- "fixedpoint" to Huet -- if v not in u. + If v occurs rigidly then nonunifiable. + If v occurs nonrigidly then must use full algorithm. *) +fun assignment context (rbinder, t, u) env = + let val vT as (v,T) = get_eta_var (rbinder, 0, t) in + (case rigid_occurs_term (Unsynchronized.ref [], env, v, u) of + NoOcc => + let val env = unify_types context (Envir.body_type env T, fastype env (rbinder, u)) env + in Envir.update (vT, Logic.rlist_abs (rbinder, u)) env end + | Nonrigid => raise ASSIGN + | Rigid => raise CANTUNIFY) + end; + + +(*Extends an rbinder with a new disagreement pair, if both are abstractions. + Tries to unify types of the bound variables! + Checks that binders have same length, since terms should be eta-normal; + if not, raises TERM, probably indicating type mismatch. + Uses variable a (unless the null string) to preserve user's naming.*) +fun new_dpair context (rbinder, Abs (a, T, body1), Abs (b, U, body2)) env = + let + val env' = unify_types context (T, U) env; + val c = if a = "" then b else a; + in new_dpair context ((c,T) :: rbinder, body1, body2) env' end + | new_dpair _ (_, Abs _, _) _ = raise TERM ("new_dpair", []) + | new_dpair _ (_, _, Abs _) _ = raise TERM ("new_dpair", []) + | new_dpair _ (rbinder, t1, t2) env = ((rbinder, t1, t2), env); + + +fun head_norm_dpair context (env, (rbinder, t, u)) : dpair * Envir.env = + new_dpair context (rbinder, + eta_norm env (rbinder, Envir.head_norm env t), + eta_norm env (rbinder, Envir.head_norm env u)) env; + + + +(*flexflex: the flex-flex pairs, flexrigid: the flex-rigid pairs + Does not perform assignments for flex-flex pairs: + may create nonrigid paths, which prevent other assignments. + Does not even identify Vars in dpairs such as ?a \\<^sup>? ?b; an attempt to + do so caused numerous problems with no compensating advantage. +*) +fun SIMPL0 context dp0 (env,flexflex,flexrigid) : Envir.env * dpair list * dpair list = + let + val (dp as (rbinder, t, u), env) = head_norm_dpair context (env, dp0); + fun SIMRANDS (f $ t, g $ u, env) = + SIMPL0 context (rbinder, t, u) (SIMRANDS (f, g, env)) + | SIMRANDS (t as _$_, _, _) = + raise TERM ("SIMPL: operands mismatch", [t, u]) + | SIMRANDS (t, u as _ $ _, _) = + raise TERM ("SIMPL: operands mismatch", [t, u]) + | SIMRANDS (_, _, env) = (env, flexflex, flexrigid); + in + (case (head_of t, head_of u) of + (Var (_, T), Var (_, U)) => + let + val T' = Envir.body_type env T and U' = Envir.body_type env U; + val env = unify_types context (T', U') env; + in (env, dp :: flexflex, flexrigid) end + | (Var _, _) => + ((assignment context (rbinder,t,u) env, flexflex, flexrigid) + handle ASSIGN => (env, flexflex, dp :: flexrigid)) + | (_, Var _) => + ((assignment context (rbinder, u, t) env, flexflex, flexrigid) + handle ASSIGN => (env, flexflex, (rbinder, u, t) :: flexrigid)) + | (Const (a, T), Const (b, U)) => + if a = b then SIMRANDS (t, u, unify_types context (T, U) env) + else raise CANTUNIFY + | (Bound i, Bound j) => + if i = j then SIMRANDS (t, u, env) else raise CANTUNIFY + | (Free (a, T), Free (b, U)) => + if a = b then SIMRANDS (t, u, unify_types context (T, U) env) + else raise CANTUNIFY + | _ => raise CANTUNIFY) + end; + + +(* changed(env,t) checks whether the head of t is a variable assigned in env*) +fun changed env (f $ _) = changed env f + | changed env (Var v) = (case Envir.lookup env v of NONE => false | _ => true) + | changed _ _ = false; + + +(*Recursion needed if any of the 'head variables' have been updated + Clever would be to re-do just the affected dpairs*) +fun SIMPL context (env,dpairs) : Envir.env * dpair list * dpair list = + let + val all as (env', flexflex, flexrigid) = fold_rev (SIMPL0 context) dpairs (env, [], []); + val dps = flexrigid @ flexflex; + in + if exists (fn (_, t, u) => changed env' t orelse changed env' u) dps + then SIMPL context (env', dps) else all + end; + + +(*Makes the terms E1,...,Em, where Ts = [T...Tm]. + Each Ei is ?Gi(B.(n-1),...,B.0), and has type Ti + The B.j are bound vars of binder. + The terms are not made in eta-normal-form, SIMPL does that later. + If done here, eta-expansion must be recursive in the arguments! *) +fun make_args _ (_, env, []) = (env, []) (*frequent case*) + | make_args name (binder: typ list, env, Ts) : Envir.env * term list = + let + fun funtype T = binder ---> T; + val (env', vars) = Envir.genvars name (env, map funtype Ts); + in (env', map (fn var => Logic.combound (var, 0, length binder)) vars) end; + + +(*Abstraction over a list of types*) +fun types_abs ([], u) = u + | types_abs (T :: Ts, u) = Abs ("", T, types_abs (Ts, u)); + +(*Abstraction over the binder of a type*) +fun type_abs (env, T, t) = types_abs (Envir.binder_types env T, t); + + +(*MATCH taking "big steps". + Copies u into the Var v, using projection on targs or imitation. + A projection is allowed unless SIMPL raises an exception. + Allocates new variables in projection on a higher-order argument, + or if u is a variable (flex-flex dpair). + Returns long sequence of every way of copying u, for backtracking + For example, projection in ?b'(?a) may be wrong if other dpairs constrain ?a. + The order for trying projections is crucial in ?b'(?a) + NB "vname" is only used in the call to make_args!! *) +fun matchcopy context vname = + let + fun mc (rbinder, targs, u, ed as (env, dpairs)) : (term * (Envir.env * dpair list)) Seq.seq = + let + val trace_types = Config.get_generic context trace_types; + (*Produce copies of uarg and cons them in front of uargs*) + fun copycons uarg (uargs, (env, dpairs)) = + Seq.map (fn (uarg', ed') => (uarg' :: uargs, ed')) + (mc (rbinder, targs,eta_norm env (rbinder, Envir.head_norm env uarg), + (env, dpairs))); + (*Produce sequence of all possible ways of copying the arg list*) + fun copyargs [] = Seq.cons ([], ed) Seq.empty + | copyargs (uarg :: uargs) = Seq.maps (copycons uarg) (copyargs uargs); + val (uhead, uargs) = strip_comb u; + val base = Envir.body_type env (fastype env (rbinder, uhead)); + fun joinargs (uargs', ed') = (list_comb (uhead, uargs'), ed'); + (*attempt projection on argument with given typ*) + val Ts = map (curry (fastype env) rbinder) targs; + fun projenv (head, (Us, bary), targ, tail) = + let + val env = + if trace_types then test_unify_types context (base, bary) env + else unify_types context (base, bary) env + in + Seq.make (fn () => + let + val (env', args) = make_args vname (Ts, env, Us); + (*higher-order projection: plug in targs for bound vars*) + fun plugin arg = list_comb (head_of arg, targs); + val dp = (rbinder, list_comb (targ, map plugin args), u); + val (env2, frigid, fflex) = SIMPL context (env', dp :: dpairs); + (*may raise exception CANTUNIFY*) + in + SOME ((list_comb (head, args), (env2, frigid @ fflex)), tail) + end handle CANTUNIFY => Seq.pull tail) + end handle CANTUNIFY => tail; + (*make a list of projections*) + fun make_projs (T::Ts, targ::targs) = + (Bound(length Ts), T, targ) :: make_projs (Ts,targs) + | make_projs ([],[]) = [] + | make_projs _ = raise TERM ("make_projs", u::targs); + (*try projections and imitation*) + fun matchfun ((bvar,T,targ)::projs) = + (projenv(bvar, Envir.strip_type env T, targ, matchfun projs)) + | matchfun [] = (*imitation last of all*) + (case uhead of + Const _ => Seq.map joinargs (copyargs uargs) + | Free _ => Seq.map joinargs (copyargs uargs) + | _ => Seq.empty) (*if Var, would be a loop!*) + in + (case uhead of + Abs (a, T, body) => + Seq.map (fn (body', ed') => (Abs (a, T, body'), ed')) + (mc ((a, T) :: rbinder, (map (incr_boundvars 1) targs) @ [Bound 0], body, ed)) + | Var (w, _) => + (*a flex-flex dpair: make variable for t*) + let + val (env', newhd) = Envir.genvar (#1 w) (env, Ts ---> base); + val tabs = Logic.combound (newhd, 0, length Ts); + val tsub = list_comb (newhd, targs); + in Seq.single (tabs, (env', (rbinder, tsub, u) :: dpairs)) end + | _ => matchfun (rev (make_projs (Ts, targs)))) + end; + in mc end; + + +(*Call matchcopy to produce assignments to the variable in the dpair*) +fun MATCH context (env, (rbinder, t, u), dpairs) : (Envir.env * dpair list) Seq.seq = + let + val (Var (vT as (v, T)), targs) = strip_comb t; + val Ts = Envir.binder_types env T; + fun new_dset (u', (env', dpairs')) = + (*if v was updated to s, must unify s with u' *) + (case Envir.lookup env' vT of + NONE => (Envir.update (vT, types_abs (Ts, u')) env', dpairs') + | SOME s => (env', ([], s, types_abs (Ts, u')) :: dpairs')); + in + Seq.map new_dset (matchcopy context (#1 v) (rbinder, targs, u, (env, dpairs))) + end; + + + +(**** Flex-flex processing ****) + +(*At end of unification, do flex-flex assignments like ?a -> ?f(?b) + Attempts to update t with u, raising ASSIGN if impossible*) +fun ff_assign context (env, rbinder, t, u) : Envir.env = + let val vT as (v, T) = get_eta_var (rbinder, 0, t) in + if occurs_terms (Unsynchronized.ref [], env, v, [u]) then raise ASSIGN + else + let val env = unify_types context (Envir.body_type env T, fastype env (rbinder, u)) env + in Envir.vupdate (vT, Logic.rlist_abs (rbinder, u)) env end + end; + + +(*If an argument contains a banned Bound, then it should be deleted. + But if the only path is flexible, this is difficult; the code gives up! + In \x y. ?a x \\<^sup>? \x y. ?b (?c y) should we instantiate ?b or ?c *) +exception CHANGE_FAIL; (*flexible occurrence of banned variable, or other reason to quit*) + + +(*Flex argument: a term, its type, and the index that refers to it.*) +type flarg = {t: term, T: typ, j: int}; + +(*Form the arguments into records for deletion/sorting.*) +fun flexargs ([], [], []) = [] : flarg list + | flexargs (j :: js, t :: ts, T :: Ts) = {j = j, t = t, T = T} :: flexargs (js, ts, Ts) + | flexargs _ = raise CHANGE_FAIL; +(*We give up if we see a variable of function type not applied to a full list of + arguments (remember, this code assumes that terms are fully eta-expanded). This situation + can occur if a type variable is instantiated with a function type. +*) + +(*Check whether the 'banned' bound var indices occur rigidly in t*) +fun rigid_bound (lev, banned) t = + let val (head,args) = strip_comb t in + (case head of + Bound i => + member (op =) banned (i - lev) orelse exists (rigid_bound (lev, banned)) args + | Var _ => false (*no rigid occurrences here!*) + | Abs (_, _, u) => + rigid_bound (lev + 1, banned) u orelse + exists (rigid_bound (lev, banned)) args + | _ => exists (rigid_bound (lev, banned)) args) + end; + +(*Squash down indices at level >=lev to delete the banned from a term.*) +fun change_bnos banned = + let + fun change lev (Bound i) = + if i < lev then Bound i + else if member (op =) banned (i - lev) then + raise CHANGE_FAIL (**flexible occurrence: give up**) + else Bound (i - length (filter (fn j => j < i - lev) banned)) + | change lev (Abs (a, T, t)) = Abs (a, T, change(lev + 1) t) + | change lev (t $ u) = change lev t $ change lev u + | change lev t = t; + in change 0 end; + +(*Change indices, delete the argument if it contains a banned Bound*) +fun change_arg banned {j, t, T} args : flarg list = + if rigid_bound (0, banned) t then args (*delete argument!*) + else {j = j, t = change_bnos banned t, T = T} :: args; + + +(*Sort the arguments to create assignments if possible: + create eta-terms like ?g B.1 B.0*) +local + fun less_arg ({t = Bound i1, ...}, {t = Bound i2, ...}) = (i2 < i1) + | less_arg (_: flarg, _: flarg) = false; + + fun ins_arg x [] = [x] + | ins_arg x (y :: ys) = + if less_arg (y, x) then y :: ins_arg x ys else x :: y :: ys; +in + fun sort_args [] = [] + | sort_args (x :: xs) = ins_arg x (sort_args xs); +end; + +(*Test whether the new term would be eta-equivalent to a variable -- + if so then there is no point in creating a new variable*) +fun decreasing n ([]: flarg list) = (n = 0) + | decreasing n ({j, ...} :: args) = j = n - 1 andalso decreasing (n - 1) args; + +(*Delete banned indices in the term, simplifying it. + Force an assignment, if possible, by sorting the arguments. + Update its head; squash indices in arguments. *) +fun clean_term banned (env,t) = + let + val (Var (v, T), ts) = strip_comb t; + val (Ts, U) = Envir.strip_type env T + and js = length ts - 1 downto 0; + val args = sort_args (fold_rev (change_arg banned) (flexargs (js, ts, Ts)) []) + val ts' = map #t args; + in + if decreasing (length Ts) args then (env, (list_comb (Var (v, T), ts'))) + else + let + val (env', v') = Envir.genvar (#1 v) (env, map #T args ---> U); + val body = list_comb (v', map (Bound o #j) args); + val env2 = Envir.vupdate ((v, T), types_abs (Ts, body)) env'; + (*the vupdate affects ts' if they contain v*) + in (env2, Envir.norm_term env2 (list_comb (v', ts'))) end + end; + + +(*Add tpair if not trivial or already there. + Should check for swapped pairs??*) +fun add_tpair (rbinder, (t0, u0), tpairs) : (term * term) list = + if t0 aconv u0 then tpairs + else + let + val t = Logic.rlist_abs (rbinder, t0) + and u = Logic.rlist_abs (rbinder, u0); + fun same (t', u') = (t aconv t') andalso (u aconv u') + in if exists same tpairs then tpairs else (t, u) :: tpairs end; + + +(*Simplify both terms and check for assignments. + Bound vars in the binder are "banned" unless used in both t AND u *) +fun clean_ffpair context ((rbinder, t, u), (env, tpairs)) = + let + val loot = loose_bnos t and loou = loose_bnos u + fun add_index (j, (a, T)) (bnos, newbinder) = + if member (op =) loot j andalso member (op =) loou j + then (bnos, (a, T) :: newbinder) (*needed by both: keep*) + else (j :: bnos, newbinder); (*remove*) + val (banned, rbin') = fold_rev add_index ((0 upto (length rbinder - 1)) ~~ rbinder) ([], []); + val (env', t') = clean_term banned (env, t); + val (env'',u') = clean_term banned (env',u); + in + (ff_assign context (env'', rbin', t', u'), tpairs) + handle ASSIGN => + (ff_assign context (env'', rbin', u', t'), tpairs) + handle ASSIGN => (env'', add_tpair (rbin', (t', u'), tpairs)) + end + handle CHANGE_FAIL => (env, add_tpair (rbinder, (t, u), tpairs)); + + +(*IF the flex-flex dpair is an assignment THEN do it ELSE put in tpairs + eliminates trivial tpairs like t=t, as well as repeated ones + trivial tpairs can easily escape SIMPL: ?A=t, ?A=?B, ?B=t gives t=t + Resulting tpairs MAY NOT be in normal form: assignments may occur here.*) +fun add_ffpair context (rbinder,t0,u0) (env,tpairs) : Envir.env * (term * term) list = + let + val t = Envir.norm_term env t0 + and u = Envir.norm_term env u0; + in + (case (head_of t, head_of u) of + (Var (v, T), Var (w, U)) => (*Check for identical variables...*) + if Term.eq_ix (v, w) then (*...occur check would falsely return true!*) + if T = U then (env, add_tpair (rbinder, (t, u), tpairs)) + else raise TERM ("add_ffpair: Var name confusion", [t, u]) + else if Term_Ord.indexname_ord (v, w) = LESS then (*prefer to update the LARGER variable*) + clean_ffpair context ((rbinder, u, t), (env, tpairs)) + else clean_ffpair context ((rbinder, t, u), (env, tpairs)) + | _ => raise TERM ("add_ffpair: Vars expected", [t, u])) + end; + + +(*Print a tracing message + list of dpairs. + In t \ u print u first because it may be rigid or flexible -- + t is always flexible.*) +fun print_dpairs context msg (env, dpairs) = + if Context_Position.is_visible_generic context then + let + fun pdp (rbinder, t, u) = + let + val ctxt = Context.proof_of context; + fun termT t = + Syntax.pretty_term ctxt (Envir.norm_term env (Logic.rlist_abs (rbinder, t))); + val prt = Pretty.blk (0, [termT u, Pretty.str " \\<^sup>?", Pretty.brk 1, termT t]); + in tracing (Pretty.string_of prt) end; + in tracing msg; List.app pdp dpairs end + else (); + + +(*Unify the dpairs in the environment. + Returns flex-flex disagreement pairs NOT IN normal form. + SIMPL may raise exception CANTUNIFY. *) +fun hounifiers binders (context, env, tus : (term * term) list) + : (Envir.env * (term * term) list) Seq.seq = + let + val trace_bound = Config.get_generic context trace_bound; + val search_bound = Config.get_generic context search_bound; + val trace_simp = Config.get_generic context trace_simp; + fun add_unify tdepth ((env, dpairs), reseq) = + Seq.make (fn () => + let + val (env', flexflex, flexrigid) = + (if tdepth > trace_bound andalso trace_simp + then print_dpairs context "Enter SIMPL" (env, dpairs) else (); + SIMPL context (env, dpairs)); + in + (case flexrigid of + [] => SOME (fold_rev (add_ffpair context) flexflex (env', []), reseq) + | dp :: frigid' => + if tdepth > search_bound then + (if Context_Position.is_visible_generic context + then warning "Unification bound exceeded" else (); Seq.pull reseq) + else + (if tdepth > trace_bound then + print_dpairs context "Enter MATCH" (env',flexrigid@flexflex) + else (); + Seq.pull (Seq.it_right + (add_unify (tdepth + 1)) (MATCH context (env',dp, frigid'@flexflex), reseq)))) + end + handle CANTUNIFY => + (if tdepth > trace_bound andalso Context_Position.is_visible_generic context + then tracing "Failure node" + else (); Seq.pull reseq)); + val dps = map (fn (t, u) => ([], t, u)) tus; + in add_unify 1 ((env, dps), Seq.empty) end; + +fun unifiers (params as (context, env, tus)) = + Seq.cons (fold (Pattern.unify context) tus env, []) Seq.empty + handle Pattern.Unif => Seq.empty + | Pattern.Pattern => hounifiers [] params; + + +(*For smash_flexflex1*) +fun var_head_of (env,t) : indexname * typ = + (case head_of (strip_abs_body (Envir.norm_term env t)) of + Var (v, T) => (v, T) + | _ => raise CANTUNIFY); (*not flexible, cannot use trivial substitution*) + + +(*Eliminate a flex-flex pair by the trivial substitution, see Huet (1975) + Unifies ?f t1 ... rm with ?g u1 ... un by ?f -> \x1...xm. ?a, ?g -> \x1...xn. ?a + Unfortunately, unifies ?f t u with ?g t u by ?f, ?g -> \x y. ?a, + though just ?g->?f is a more general unifier. + Unlike Huet (1975), does not smash together all variables of same type -- + requires more work yet gives a less general unifier (fewer variables). + Handles ?f t1 ... rm with ?f u1 ... um to avoid multiple updates. *) +fun smash_flexflex1 (t, u) env : Envir.env = + let + val vT as (v, T) = var_head_of (env, t) + and wU as (w, U) = var_head_of (env, u); + val (env', var) = Envir.genvar (#1 v) (env, Envir.body_type env T); + val env'' = Envir.vupdate (wU, type_abs (env', U, var)) env'; + in + if vT = wU then env'' (*the other update would be identical*) + else Envir.vupdate (vT, type_abs (env', T, var)) env'' + end; + + +(*Smash all flex-flexpairs. Should allow selection of pairs by a predicate?*) +fun smash_flexflex (env, tpairs) : Envir.env = + fold_rev smash_flexflex1 tpairs env; + +(*Returns unifiers with no remaining disagreement pairs*) +fun smash_unifiers context tus env = + Seq.map smash_flexflex (unifiers (context, env, tus)); + +end; diff --git a/thys/ML_Unification/Unifiers/unification_combinator.ML b/thys/ML_Unification/Unifiers/unification_combinator.ML new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/Unifiers/unification_combinator.ML @@ -0,0 +1,108 @@ +(* Title: ML_Unification/unification_combinator.ML + Author: Kevin Kappelmann + +Unification combinators. +*) +signature UNIFICATION_COMBINATOR = +sig + (* failures *) + val fail_closed_unify : Unification_Base.closed_unifier + val fail_unify : Unification_Base.unifier + val fail_e_unify : Unification_Base.e_unifier + + val fail_closed_match : Unification_Base.closed_matcher + val fail_match : Unification_Base.matcher + val fail_e_match : Unification_Base.e_matcher + + (* fallbacks *) + (*"add_fallback_e_unifier u1 u2 u" creates a unifier that calls u1 with fallback unifier "u2 u"*) + val add_fallback_e_unifier : Unification_Base.e_unifier -> Unification_Base.e_unifier -> + Unification_Base.e_unifier + (*"add_fallback_unifier u1 u2" creates a unifier that calls u1 with fallback unifier u2*) + val add_fallback_unifier : Unification_Base.e_unifier -> Unification_Base.unifier -> + Unification_Base.unifier + + val add_fallback_e_matcher : Unification_Base.e_matcher -> Unification_Base.e_matcher -> + Unification_Base.e_matcher + val add_fallback_matcher : Unification_Base.e_matcher -> Unification_Base.matcher -> + Unification_Base.matcher + + (* sequential *) + (*try all unifiers in sequence*) + val concat_closed_unifiers : Unification_Base.closed_unifier list -> + Unification_Base.closed_unifier + val concat_closed_matchers : Unification_Base.closed_matcher list -> + Unification_Base.closed_matcher + val concat_unifiers : Unification_Base.unifier list -> Unification_Base.unifier + val concat_matchers : Unification_Base.matcher list -> Unification_Base.matcher + val concat_e_unifiers : Unification_Base.e_unifier list -> Unification_Base.e_unifier + val concat_e_matchers : Unification_Base.e_matcher list -> Unification_Base.e_matcher + + (* normalisation *) + (*apply normaliser before calling the matcher/unifier*) + val norm_closed_unifier : Envir_Normalisation.term_normaliser -> Unification_Base.closed_unifier -> + Unification_Base.closed_unifier + val norm_closed_matcher : Envir_Normalisation.term_normaliser -> Unification_Base.closed_matcher -> + Unification_Base.closed_matcher + val norm_unifier : Envir_Normalisation.term_normaliser -> Unification_Base.unifier -> + Unification_Base.unifier + val norm_matcher : Envir_Normalisation.term_normaliser -> Unification_Base.matcher -> + Unification_Base.matcher + + (* miscellaneous *) + (*approximates a unifier (working with bound variables) from a closed unifier + by replacing all bound variables with their associated free variables*) + val unifier_from_closed_unifier : Unification_Base.closed_unifier -> Unification_Base.unifier + val matcher_from_closed_matcher : Unification_Base.closed_matcher -> Unification_Base.matcher + +end + +structure Unification_Combinator : UNIFICATION_COMBINATOR = +struct + +structure UB = Unification_Base + +(* failures *) +val fail_closed_unify = K o K o K Seq.empty +val fail_unify = K fail_closed_unify +val fail_e_unify = K fail_unify + +val fail_closed_match = fail_closed_unify +val fail_match = fail_unify +val fail_e_match = fail_e_unify + +(* fallbacks *) +fun add_fallback_e_unifier e_unif1 e_unif2 = e_unif1 o e_unif2 +fun add_fallback_unifier e_unif1 e_unif2 = add_fallback_e_unifier e_unif1 (K e_unif2) fail_unify + +val add_fallback_e_matcher = add_fallback_e_unifier +fun add_fallback_matcher e_match1 e_match2 = add_fallback_e_matcher e_match1 (K e_match2) fail_match + +(* sequential *) +fun concat_closed_unifiers unifiers ctxt tp = + fold_rev (fn unify => curry Seq.APPEND (unify ctxt tp)) unifiers Seq.fail +val concat_closed_matchers = concat_closed_unifiers + +fun concat_unifiers unifiers binders ctxt tp = + fold_rev (fn unify => curry Seq.APPEND (unify binders ctxt tp)) unifiers Seq.fail +val concat_matchers = concat_unifiers + +fun concat_e_unifiers unifiers unif binders ctxt tp = + fold_rev (fn unify => curry Seq.APPEND (unify unif binders ctxt tp)) unifiers Seq.fail +val concat_e_matchers = concat_e_unifiers + +(* normalisation *) +fun norm_closed_unifier norm_t unif ctxt tp env = unif ctxt (apply2 (norm_t env) tp) env +fun norm_closed_matcher norm_p match ctxt (p, t) env = match ctxt (norm_p env p, t) env + +fun norm_unifier norm_t unif binders ctxt tp env = + norm_closed_unifier norm_t (unif (Binders.norm_binders (norm_t env) binders)) ctxt tp env +fun norm_matcher norm_p match binders ctxt pt env = + norm_closed_matcher norm_p (match (Binders.norm_binders (norm_p env) binders)) ctxt pt env + +(* miscellaneous *) +fun unifier_from_closed_unifier unify binders ctxt = + unify ctxt o apply2 (Binders.replace_binders binders) +val matcher_from_closed_matcher = unifier_from_closed_unifier + +end diff --git a/thys/ML_Unification/Unifiers/unification_combine.ML b/thys/ML_Unification/Unifiers/unification_combine.ML new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/Unifiers/unification_combine.ML @@ -0,0 +1,193 @@ +(* Title: ML_Unification/unification_combine.ML + Author: Kevin Kappelmann + +Combine a list of e-unifiers. +*) +@{parse_entries (sig) PARSE_UNIFICATION_COMBINE_ARGS + [add, del]} + +signature UNIFICATION_COMBINE = +sig + include HAS_LOGGER + + structure PA : PARSE_UNIFICATION_COMBINE_ARGS + + type binding = binding + val eq_binding : binding * binding -> bool + val pretty_binding : binding -> Pretty.T + + type metadata + val metadata : binding -> Prio.prio -> metadata + val dest_metadata : metadata -> binding * Prio.prio + val metadata_binding : metadata -> binding + val metadata_prio : metadata -> Prio.prio + val eq_metadata : metadata * metadata -> bool + val pretty_metadata : metadata -> Pretty.T + val default_metadata : binding -> metadata + + type eunif_data + val eunif_data : Unification_Base.e_unifier -> metadata -> eunif_data + val dest_eunif_data : eunif_data -> Unification_Base.e_unifier * metadata + val eunif_data_e_unifier : eunif_data -> Unification_Base.e_unifier + val eunif_data_metadata : eunif_data -> metadata + val pretty_eunif_data : eunif_data -> Pretty.T + + type eunif_datas = eunif_data list Prio.Table.table + (*storing the e-unifiers and metadata*) + structure Data : GENERIC_DATA + + val get_eunif_datas : Context.generic -> eunif_datas + val map_eunif_datas : (eunif_datas -> eunif_datas) -> Context.generic -> Context.generic + val insert_eunif_data : eunif_data -> Context.generic -> Context.generic + val delete_eunif_data : metadata -> Context.generic -> Context.generic + + val pretty_data : Proof.context -> Pretty.T + + val e_unify : Unification_Base.e_unifier + + val binding : Binding.binding + + val arg_parsers : (ML_Code_Util.code parser, ML_Code_Util.code parser) PA.entries + val attribute : (ML_Code_Util.code, ML_Code_Util.code) PA.entries * Position.T -> attribute + val parse_attribute : attribute parser + val setup_attribute : string option -> local_theory -> local_theory +end + +functor Unification_Combine(A : sig + structure FIA : FUNCTOR_INSTANCE_ARGS + end) : UNIFICATION_COMBINE = +struct + +structure UCB = Unification_Combinator +structure UB = Unification_Base +structure FIU = Functor_Instance_Util(A.FIA) + +val logger = Logger.setup_new_logger Unification_Base.logger FIU.FIA.full_name + +@{parse_entries (struct) PA + [add, del]} + +type binding = binding + +val eq_binding = Binding.eq_name +val pretty_binding = Binding.pretty + +datatype metadata = Metadata of binding * Prio.prio + +fun metadata binding prio = Metadata (binding, prio) +fun dest_metadata (Metadata meta) = meta +fun metadata_binding (Metadata (binding, _)) = binding +fun metadata_prio (Metadata (_, prio)) = prio + +val eq_metadata = eq_pair eq_binding Prio.eq o apply2 dest_metadata +fun pretty_metadata (Metadata (binding, prio))= Pretty.block [ + Pretty.str "Binding: ", + pretty_binding binding, + Pretty.str ", Priority: ", + Prio.pretty prio + ] + +fun default_metadata binding = metadata binding Prio.MEDIUM + +datatype eunif_data = EUnif_Data of UB.e_unifier * metadata + +fun eunif_data unify meta = EUnif_Data (unify, meta) +fun dest_eunif_data (EUnif_Data data) = data +fun eunif_data_e_unifier (EUnif_Data (unify, _)) = unify +fun eunif_data_metadata (EUnif_Data (_, meta)) = meta +val eunif_data_prio = eunif_data_metadata #> metadata_prio + +val eq_eunif_data = eq_metadata o apply2 eunif_data_metadata +fun pretty_eunif_data (EUnif_Data (_, meta)) = pretty_metadata meta + +type eunif_datas = eunif_data list Prio.Table.table + +(*invariant: data is sorted with respect to the priorities*) +structure Data = Generic_Data(struct + type T = eunif_datas + val empty = Prio.Table.empty + val merge = Prio.Table.merge_list eq_eunif_data +end) + +val get_eunif_datas = Data.get +val get_e_unifiers = map (eunif_data_e_unifier o snd) o Prio.Table.dest_list o get_eunif_datas +val map_eunif_datas = Data.map + +fun insert_eunif_data data context = + let val prio = eunif_data_prio data + in + Data.map (fn datas => + if member eq_eunif_data (Prio.Table.lookup_list datas prio) data + then (@{log Logger.WARN} (Context.proof_of context) (fn _ => + Pretty.block [ + Pretty.str "eunif data (", + pretty_eunif_data data, + Pretty.str ") already added." + ] |> Pretty.string_of); + datas) + else Prio.Table.insert_list eq_eunif_data (prio, data) datas + ) context + end + +fun delete_eunif_data metadata context = + let val prio = metadata_prio metadata + in + Data.map (fn datas => + if member eq_metadata (Prio.Table.lookup_list datas prio |> map eunif_data_metadata) metadata + then Prio.Table.remove_list (eq_metadata o apsnd eunif_data_metadata) (prio, metadata) datas + else (@{log Logger.WARN} (Context.proof_of context) (fn _ => + Pretty.block [ + Pretty.str "metadata (", + pretty_metadata metadata, + Pretty.str ") not found." + ] |> Pretty.string_of); + datas) + ) context + end + +fun pretty_data ctxt = get_eunif_datas (Context.Proof ctxt) + |> Prio.Table.dest_list + |> map (pretty_eunif_data o snd) + |> Pretty.fbreaks + |> Pretty.block + +fun e_unify unif binders ctxt tp env = + (@{log Logger.DEBUG} ctxt (fn _ => + Pretty.block [ + Pretty.str "Combining e-unifiers for ", + Unification_Util.pretty_unif_problem ctxt (apply2 (Envir_Normalisation.norm_term_unif env) tp) + ] |> Pretty.string_of); + UCB.concat_e_unifiers (get_e_unifiers (Context.Proof ctxt)) unif binders ctxt tp env) + +val binding = FIU.mk_binding_id_prefix "ucombine" + +val arg_parsers = { + add = SOME (Parse_Util.nonempty_code (K "eunif_data to add must not be empty")), + del = SOME (Parse_Util.nonempty_code (K "eunif_data to delete must not be empty")) +} + +val parse_arg_entries = + let + val parse_value = PA.parse_entry (PA.get_add arg_parsers) (PA.get_del arg_parsers) + val parse_entry = Parse_Key_Value.parse_entry PA.parse_key Parse_Util.eq parse_value + in PA.parse_entries_required Parse.and_list1 [] parse_entry (PA.empty_entries ()) end + +fun attribute (entries, pos) = + let + fun default (context, _) = (SOME context, NONE) + val run_code = ML_Attribute.run_map_context o rpair pos + val add = case PA.get_add_safe entries of + SOME c => FIU.code_struct_op "insert_eunif_data" @ ML_Code_Util.atomic c |> run_code + | NONE => default + val del = case PA.get_del_safe entries of + SOME c => FIU.code_struct_op "delete_eunif_data" @ ML_Code_Util.atomic c |> run_code + | NONE => default + in ML_Attribute_Util.apply_attribute del #> add end + +val parse_attribute = (parse_arg_entries |> Parse_Util.position) >> attribute + +val setup_attribute = Attrib.local_setup binding + (Parse.!!! parse_attribute |> Scan.lift) o + the_default ("set unification-combine arguments (" ^ FIU.FIA.full_name ^ ")") + +end diff --git a/thys/ML_Unification/Unifiers/unify_copy.ML b/thys/ML_Unification/Unifiers/unify_copy.ML new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/Unifiers/unify_copy.ML @@ -0,0 +1,664 @@ +(* +NOTE: This is just a copy of unify.ML from the Isabelle distribution that includes the +patch 3c57995c255c (since it required here but not included in Isabelle2023). +FIXME: delete for next release version +*) + +(* Title: Pure/unify.ML + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright Cambridge University 1992 + +Higher-Order Unification. + +Types as well as terms are unified. The outermost functions assume +the terms to be unified already have the same type. In resolution, +this is assured because both have type "prop". +*) + +signature UNIFY_COPY = +sig + val trace_bound: int Config.T + val search_bound: int Config.T + val trace_simp: bool Config.T + val trace_types: bool Config.T + val hounifiers: (string * typ) list -> Context.generic * Envir.env * ((term * term) list) -> + (Envir.env * (term * term) list) Seq.seq + val unifiers: Context.generic * Envir.env * ((term * term) list) -> + (Envir.env * (term * term) list) Seq.seq + val smash_unifiers: Context.generic -> (term * term) list -> Envir.env -> Envir.env Seq.seq +end + +structure Unify_Copy : UNIFY_COPY = +struct + +(*Unification options*) + +(*tracing starts above this depth, 0 for full*) +val trace_bound = Unify.trace_bound; + +(*unification quits above this depth*) +val search_bound = Unify.search_bound; + +(*print dpairs before calling SIMPL*) +val trace_simp = Unify.trace_simp; + +(*announce potential incompleteness of type unification*) +val trace_types = Unify.trace_types; + + +type binderlist = (string * typ) list; + +type dpair = binderlist * term * term; + +fun fastype env (Ts, t) = Envir.fastype env (map snd Ts) t; + + +(* eta normal form *) + +fun eta_norm env = + let + val tyenv = Envir.type_env env; + fun etif (Type ("fun", [T, U]), t) = + Abs ("", T, etif (U, incr_boundvars 1 t $ Bound 0)) + | etif (TVar v, t) = + (case Type.lookup tyenv v of + NONE => t + | SOME T => etif (T, t)) + | etif (_, t) = t; + fun eta_nm (rbinder, Abs (a, T, body)) = + Abs (a, T, eta_nm ((a, T) :: rbinder, body)) + | eta_nm (rbinder, t) = etif (fastype env (rbinder, t), t); + in eta_nm end; + + +(*OCCURS CHECK + Does the uvar occur in the term t? + two forms of search, for whether there is a rigid path to the current term. + "seen" is list of variables passed thru, is a memo variable for sharing. + This version searches for nonrigid occurrence, returns true if found. + Since terms may contain variables with same name and different types, + the occurs check must ignore the types of variables. This avoids + that ?x::?'a is unified with f(?x::T), which may lead to a cyclic + substitution when ?'a is instantiated with T later. *) +fun occurs_terms (seen: indexname list Unsynchronized.ref, + env: Envir.env, v: indexname, ts: term list): bool = + let + fun occurs [] = false + | occurs (t :: ts) = occur t orelse occurs ts + and occur (Const _) = false + | occur (Bound _) = false + | occur (Free _) = false + | occur (Var (w, T)) = + if member (op =) (!seen) w then false + else if Term.eq_ix (v, w) then true + (*no need to lookup: v has no assignment*) + else + (seen := w :: !seen; + case Envir.lookup env (w, T) of + NONE => false + | SOME t => occur t) + | occur (Abs (_, _, body)) = occur body + | occur (f $ t) = occur t orelse occur f; + in occurs ts end; + + +(* f a1 ... an ----> f using the assignments*) +fun head_of_in env t = + (case t of + f $ _ => head_of_in env f + | Var vT => + (case Envir.lookup env vT of + SOME u => head_of_in env u + | NONE => t) + | _ => t); + + +datatype occ = NoOcc | Nonrigid | Rigid; + +(* Rigid occur check +Returns Rigid if it finds a rigid occurrence of the variable, + Nonrigid if it finds a nonrigid path to the variable. + NoOcc otherwise. + Continues searching for a rigid occurrence even if it finds a nonrigid one. + +Condition for detecting non-unifable terms: [ section 5.3 of Huet (1975) ] + a rigid path to the variable, appearing with no arguments. +Here completeness is sacrificed in order to reduce danger of divergence: + reject ALL rigid paths to the variable. +Could check for rigid paths to bound variables that are out of scope. +Not necessary because the assignment test looks at variable's ENTIRE rbinder. + +Treatment of head(arg1,...,argn): +If head is a variable then no rigid path, switch to nonrigid search +for arg1,...,argn. +If head is an abstraction then possibly no rigid path (head could be a + constant function) so again use nonrigid search. Happens only if + term is not in normal form. + +Warning: finds a rigid occurrence of ?f in ?f(t). + Should NOT be called in this case: there is a flex-flex unifier +*) +fun rigid_occurs_term (seen: indexname list Unsynchronized.ref, env, v: indexname, t) = + let + fun nonrigid t = + if occurs_terms (seen, env, v, [t]) then Nonrigid + else NoOcc + fun occurs [] = NoOcc + | occurs (t :: ts) = + (case occur t of + Rigid => Rigid + | oc => (case occurs ts of NoOcc => oc | oc2 => oc2)) + and occomb (f $ t) = + (case occur t of + Rigid => Rigid + | oc => (case occomb f of NoOcc => oc | oc2 => oc2)) + | occomb t = occur t + and occur (Const _) = NoOcc + | occur (Bound _) = NoOcc + | occur (Free _) = NoOcc + | occur (Var (w, T)) = + if member (op =) (!seen) w then NoOcc + else if Term.eq_ix (v, w) then Rigid + else + (seen := w :: !seen; + case Envir.lookup env (w, T) of + NONE => NoOcc + | SOME t => occur t) + | occur (Abs (_, _, body)) = occur body + | occur (t as f $ _) = (*switch to nonrigid search?*) + (case head_of_in env f of + Var (w,_) => (*w is not assigned*) + if Term.eq_ix (v, w) then Rigid + else nonrigid t + | Abs _ => nonrigid t (*not in normal form*) + | _ => occomb t) + in occur t end; + + +exception CANTUNIFY; (*Signals non-unifiability. Does not signal errors!*) +exception ASSIGN; (*Raised if not an assignment*) + + +fun unify_types context TU env = + Pattern.unify_types context TU env handle Pattern.Unif => raise CANTUNIFY; + +fun test_unify_types context (T, U) env = + let + fun trace () = + if Context_Position.is_visible_generic context then + let val str_of = Syntax.string_of_typ (Context.proof_of context) + in tracing ("Potential loss of completeness: " ^ str_of U ^ " = " ^ str_of T) end + else (); + val env' = unify_types context (T, U) env; + in if is_TVar T orelse is_TVar U then trace () else (); env' end; + +(*Is the term eta-convertible to a single variable with the given rbinder? + Examples: ?a ?f(B.0) ?g(B.1,B.0) + Result is var a for use in SIMPL. *) +fun get_eta_var ([], _, Var vT) = vT + | get_eta_var (_::rbinder, n, f $ Bound i) = + if n = i then get_eta_var (rbinder, n + 1, f) + else raise ASSIGN + | get_eta_var _ = raise ASSIGN; + + +(*Solve v=u by assignment -- "fixedpoint" to Huet -- if v not in u. + If v occurs rigidly then nonunifiable. + If v occurs nonrigidly then must use full algorithm. *) +fun assignment context (rbinder, t, u) env = + let val vT as (v,T) = get_eta_var (rbinder, 0, t) in + (case rigid_occurs_term (Unsynchronized.ref [], env, v, u) of + NoOcc => + let val env = unify_types context (Envir.body_type env T, fastype env (rbinder, u)) env + in Envir.update (vT, Logic.rlist_abs (rbinder, u)) env end + | Nonrigid => raise ASSIGN + | Rigid => raise CANTUNIFY) + end; + + +(*Extends an rbinder with a new disagreement pair, if both are abstractions. + Tries to unify types of the bound variables! + Checks that binders have same length, since terms should be eta-normal; + if not, raises TERM, probably indicating type mismatch. + Uses variable a (unless the null string) to preserve user's naming.*) +fun new_dpair context (rbinder, Abs (a, T, body1), Abs (b, U, body2)) env = + let + val env' = unify_types context (T, U) env; + val c = if a = "" then b else a; + in new_dpair context ((c,T) :: rbinder, body1, body2) env' end + | new_dpair _ (_, Abs _, _) _ = raise TERM ("new_dpair", []) + | new_dpair _ (_, _, Abs _) _ = raise TERM ("new_dpair", []) + | new_dpair _ (rbinder, t1, t2) env = ((rbinder, t1, t2), env); + + +fun head_norm_dpair context (env, (rbinder, t, u)) : dpair * Envir.env = + new_dpair context (rbinder, + eta_norm env (rbinder, Envir.head_norm env t), + eta_norm env (rbinder, Envir.head_norm env u)) env; + + + +(*flexflex: the flex-flex pairs, flexrigid: the flex-rigid pairs + Does not perform assignments for flex-flex pairs: + may create nonrigid paths, which prevent other assignments. + Does not even identify Vars in dpairs such as ?a \\<^sup>? ?b; an attempt to + do so caused numerous problems with no compensating advantage. +*) +fun SIMPL0 context dp0 (env,flexflex,flexrigid) : Envir.env * dpair list * dpair list = + let + val (dp as (rbinder, t, u), env) = head_norm_dpair context (env, dp0); + fun SIMRANDS (f $ t, g $ u, env) = + SIMPL0 context (rbinder, t, u) (SIMRANDS (f, g, env)) + | SIMRANDS (t as _$_, _, _) = + raise TERM ("SIMPL: operands mismatch", [t, u]) + | SIMRANDS (t, u as _ $ _, _) = + raise TERM ("SIMPL: operands mismatch", [t, u]) + | SIMRANDS (_, _, env) = (env, flexflex, flexrigid); + in + (case (head_of t, head_of u) of + (Var (_, T), Var (_, U)) => + let + val T' = Envir.body_type env T and U' = Envir.body_type env U; + val env = unify_types context (T', U') env; + in (env, dp :: flexflex, flexrigid) end + | (Var _, _) => + ((assignment context (rbinder,t,u) env, flexflex, flexrigid) + handle ASSIGN => (env, flexflex, dp :: flexrigid)) + | (_, Var _) => + ((assignment context (rbinder, u, t) env, flexflex, flexrigid) + handle ASSIGN => (env, flexflex, (rbinder, u, t) :: flexrigid)) + | (Const (a, T), Const (b, U)) => + if a = b then SIMRANDS (t, u, unify_types context (T, U) env) + else raise CANTUNIFY + | (Bound i, Bound j) => + if i = j then SIMRANDS (t, u, env) else raise CANTUNIFY + | (Free (a, T), Free (b, U)) => + if a = b then SIMRANDS (t, u, unify_types context (T, U) env) + else raise CANTUNIFY + | _ => raise CANTUNIFY) + end; + + +(* changed(env,t) checks whether the head of t is a variable assigned in env*) +fun changed env (f $ _) = changed env f + | changed env (Var v) = (case Envir.lookup env v of NONE => false | _ => true) + | changed _ _ = false; + + +(*Recursion needed if any of the 'head variables' have been updated + Clever would be to re-do just the affected dpairs*) +fun SIMPL context (env,dpairs) : Envir.env * dpair list * dpair list = + let + val all as (env', flexflex, flexrigid) = fold_rev (SIMPL0 context) dpairs (env, [], []); + val dps = flexrigid @ flexflex; + in + if exists (fn (_, t, u) => changed env' t orelse changed env' u) dps + then SIMPL context (env', dps) else all + end; + + +(*Makes the terms E1,...,Em, where Ts = [T...Tm]. + Each Ei is ?Gi(B.(n-1),...,B.0), and has type Ti + The B.j are bound vars of binder. + The terms are not made in eta-normal-form, SIMPL does that later. + If done here, eta-expansion must be recursive in the arguments! *) +fun make_args _ (_, env, []) = (env, []) (*frequent case*) + | make_args name (binder: typ list, env, Ts) : Envir.env * term list = + let + fun funtype T = binder ---> T; + val (env', vars) = Envir.genvars name (env, map funtype Ts); + in (env', map (fn var => Logic.combound (var, 0, length binder)) vars) end; + + +(*Abstraction over a list of types*) +fun types_abs ([], u) = u + | types_abs (T :: Ts, u) = Abs ("", T, types_abs (Ts, u)); + +(*Abstraction over the binder of a type*) +fun type_abs (env, T, t) = types_abs (Envir.binder_types env T, t); + + +(*MATCH taking "big steps". + Copies u into the Var v, using projection on targs or imitation. + A projection is allowed unless SIMPL raises an exception. + Allocates new variables in projection on a higher-order argument, + or if u is a variable (flex-flex dpair). + Returns long sequence of every way of copying u, for backtracking + For example, projection in ?b'(?a) may be wrong if other dpairs constrain ?a. + The order for trying projections is crucial in ?b'(?a) + NB "vname" is only used in the call to make_args!! *) +fun matchcopy context vname = + let + fun mc (rbinder, targs, u, ed as (env, dpairs)) : (term * (Envir.env * dpair list)) Seq.seq = + let + val trace_types = Config.get_generic context trace_types; + (*Produce copies of uarg and cons them in front of uargs*) + fun copycons uarg (uargs, (env, dpairs)) = + Seq.map (fn (uarg', ed') => (uarg' :: uargs, ed')) + (mc (rbinder, targs,eta_norm env (rbinder, Envir.head_norm env uarg), + (env, dpairs))); + (*Produce sequence of all possible ways of copying the arg list*) + fun copyargs [] = Seq.cons ([], ed) Seq.empty + | copyargs (uarg :: uargs) = Seq.maps (copycons uarg) (copyargs uargs); + val (uhead, uargs) = strip_comb u; + val base = Envir.body_type env (fastype env (rbinder, uhead)); + fun joinargs (uargs', ed') = (list_comb (uhead, uargs'), ed'); + (*attempt projection on argument with given typ*) + val Ts = map (curry (fastype env) rbinder) targs; + fun projenv (head, (Us, bary), targ, tail) = + let + val env = + if trace_types then test_unify_types context (base, bary) env + else unify_types context (base, bary) env + in + Seq.make (fn () => + let + val (env', args) = make_args vname (Ts, env, Us); + (*higher-order projection: plug in targs for bound vars*) + fun plugin arg = list_comb (head_of arg, targs); + val dp = (rbinder, list_comb (targ, map plugin args), u); + val (env2, frigid, fflex) = SIMPL context (env', dp :: dpairs); + (*may raise exception CANTUNIFY*) + in + SOME ((list_comb (head, args), (env2, frigid @ fflex)), tail) + end handle CANTUNIFY => Seq.pull tail) + end handle CANTUNIFY => tail; + (*make a list of projections*) + fun make_projs (T::Ts, targ::targs) = + (Bound(length Ts), T, targ) :: make_projs (Ts,targs) + | make_projs ([],[]) = [] + | make_projs _ = raise TERM ("make_projs", u::targs); + (*try projections and imitation*) + fun matchfun ((bvar,T,targ)::projs) = + (projenv(bvar, Envir.strip_type env T, targ, matchfun projs)) + | matchfun [] = (*imitation last of all*) + (case uhead of + Const _ => Seq.map joinargs (copyargs uargs) + | Free _ => Seq.map joinargs (copyargs uargs) + | _ => Seq.empty) (*if Var, would be a loop!*) + in + (case uhead of + Abs (a, T, body) => + Seq.map (fn (body', ed') => (Abs (a, T, body'), ed')) + (mc ((a, T) :: rbinder, (map (incr_boundvars 1) targs) @ [Bound 0], body, ed)) + | Var (w, _) => + (*a flex-flex dpair: make variable for t*) + let + val (env', newhd) = Envir.genvar (#1 w) (env, Ts ---> base); + val tabs = Logic.combound (newhd, 0, length Ts); + val tsub = list_comb (newhd, targs); + in Seq.single (tabs, (env', (rbinder, tsub, u) :: dpairs)) end + | _ => matchfun (rev (make_projs (Ts, targs)))) + end; + in mc end; + + +(*Call matchcopy to produce assignments to the variable in the dpair*) +fun MATCH context (env, (rbinder, t, u), dpairs) : (Envir.env * dpair list) Seq.seq = + let + val (Var (vT as (v, T)), targs) = strip_comb t; + val Ts = Envir.binder_types env T; + fun new_dset (u', (env', dpairs')) = + (*if v was updated to s, must unify s with u' *) + (case Envir.lookup env' vT of + NONE => (Envir.update (vT, types_abs (Ts, u')) env', dpairs') + | SOME s => (env', ([], s, types_abs (Ts, u')) :: dpairs')); + in + Seq.map new_dset (matchcopy context (#1 v) (rbinder, targs, u, (env, dpairs))) + end; + + + +(**** Flex-flex processing ****) + +(*At end of unification, do flex-flex assignments like ?a -> ?f(?b) + Attempts to update t with u, raising ASSIGN if impossible*) +fun ff_assign context (env, rbinder, t, u) : Envir.env = + let val vT as (v, T) = get_eta_var (rbinder, 0, t) in + if occurs_terms (Unsynchronized.ref [], env, v, [u]) then raise ASSIGN + else + let val env = unify_types context (Envir.body_type env T, fastype env (rbinder, u)) env + in Envir.vupdate (vT, Logic.rlist_abs (rbinder, u)) env end + end; + + +(*If an argument contains a banned Bound, then it should be deleted. + But if the only path is flexible, this is difficult; the code gives up! + In \x y. ?a x \\<^sup>? \x y. ?b (?c y) should we instantiate ?b or ?c *) +exception CHANGE_FAIL; (*flexible occurrence of banned variable, or other reason to quit*) + + +(*Flex argument: a term, its type, and the index that refers to it.*) +type flarg = {t: term, T: typ, j: int}; + +(*Form the arguments into records for deletion/sorting.*) +fun flexargs ([], [], []) = [] : flarg list + | flexargs (j :: js, t :: ts, T :: Ts) = {j = j, t = t, T = T} :: flexargs (js, ts, Ts) + | flexargs _ = raise CHANGE_FAIL; +(*We give up if we see a variable of function type not applied to a full list of + arguments (remember, this code assumes that terms are fully eta-expanded). This situation + can occur if a type variable is instantiated with a function type. +*) + +(*Check whether the 'banned' bound var indices occur rigidly in t*) +fun rigid_bound (lev, banned) t = + let val (head,args) = strip_comb t in + (case head of + Bound i => + member (op =) banned (i - lev) orelse exists (rigid_bound (lev, banned)) args + | Var _ => false (*no rigid occurrences here!*) + | Abs (_, _, u) => + rigid_bound (lev + 1, banned) u orelse + exists (rigid_bound (lev, banned)) args + | _ => exists (rigid_bound (lev, banned)) args) + end; + +(*Squash down indices at level >=lev to delete the banned from a term.*) +fun change_bnos banned = + let + fun change lev (Bound i) = + if i < lev then Bound i + else if member (op =) banned (i - lev) then + raise CHANGE_FAIL (**flexible occurrence: give up**) + else Bound (i - length (filter (fn j => j < i - lev) banned)) + | change lev (Abs (a, T, t)) = Abs (a, T, change(lev + 1) t) + | change lev (t $ u) = change lev t $ change lev u + | change lev t = t; + in change 0 end; + +(*Change indices, delete the argument if it contains a banned Bound*) +fun change_arg banned {j, t, T} args : flarg list = + if rigid_bound (0, banned) t then args (*delete argument!*) + else {j = j, t = change_bnos banned t, T = T} :: args; + + +(*Sort the arguments to create assignments if possible: + create eta-terms like ?g B.1 B.0*) +local + fun less_arg ({t = Bound i1, ...}, {t = Bound i2, ...}) = (i2 < i1) + | less_arg (_: flarg, _: flarg) = false; + + fun ins_arg x [] = [x] + | ins_arg x (y :: ys) = + if less_arg (y, x) then y :: ins_arg x ys else x :: y :: ys; +in + fun sort_args [] = [] + | sort_args (x :: xs) = ins_arg x (sort_args xs); +end; + +(*Test whether the new term would be eta-equivalent to a variable -- + if so then there is no point in creating a new variable*) +fun decreasing n ([]: flarg list) = (n = 0) + | decreasing n ({j, ...} :: args) = j = n - 1 andalso decreasing (n - 1) args; + +(*Delete banned indices in the term, simplifying it. + Force an assignment, if possible, by sorting the arguments. + Update its head; squash indices in arguments. *) +fun clean_term banned (env,t) = + let + val (Var (v, T), ts) = strip_comb t; + val (Ts, U) = Envir.strip_type env T + and js = length ts - 1 downto 0; + val args = sort_args (fold_rev (change_arg banned) (flexargs (js, ts, Ts)) []) + val ts' = map #t args; + in + if decreasing (length Ts) args then (env, (list_comb (Var (v, T), ts'))) + else + let + val (env', v') = Envir.genvar (#1 v) (env, map #T args ---> U); + val body = list_comb (v', map (Bound o #j) args); + val env2 = Envir.vupdate ((v, T), types_abs (Ts, body)) env'; + (*the vupdate affects ts' if they contain v*) + in (env2, Envir.norm_term env2 (list_comb (v', ts'))) end + end; + + +(*Add tpair if not trivial or already there. + Should check for swapped pairs??*) +fun add_tpair (rbinder, (t0, u0), tpairs) : (term * term) list = + if t0 aconv u0 then tpairs + else + let + val t = Logic.rlist_abs (rbinder, t0) + and u = Logic.rlist_abs (rbinder, u0); + fun same (t', u') = (t aconv t') andalso (u aconv u') + in if exists same tpairs then tpairs else (t, u) :: tpairs end; + + +(*Simplify both terms and check for assignments. + Bound vars in the binder are "banned" unless used in both t AND u *) +fun clean_ffpair context ((rbinder, t, u), (env, tpairs)) = + let + val loot = loose_bnos t and loou = loose_bnos u + fun add_index (j, (a, T)) (bnos, newbinder) = + if member (op =) loot j andalso member (op =) loou j + then (bnos, (a, T) :: newbinder) (*needed by both: keep*) + else (j :: bnos, newbinder); (*remove*) + val (banned, rbin') = fold_rev add_index ((0 upto (length rbinder - 1)) ~~ rbinder) ([], []); + val (env', t') = clean_term banned (env, t); + val (env'',u') = clean_term banned (env',u); + in + (ff_assign context (env'', rbin', t', u'), tpairs) + handle ASSIGN => + (ff_assign context (env'', rbin', u', t'), tpairs) + handle ASSIGN => (env'', add_tpair (rbin', (t', u'), tpairs)) + end + handle CHANGE_FAIL => (env, add_tpair (rbinder, (t, u), tpairs)); + + +(*IF the flex-flex dpair is an assignment THEN do it ELSE put in tpairs + eliminates trivial tpairs like t=t, as well as repeated ones + trivial tpairs can easily escape SIMPL: ?A=t, ?A=?B, ?B=t gives t=t + Resulting tpairs MAY NOT be in normal form: assignments may occur here.*) +fun add_ffpair context (rbinder,t0,u0) (env,tpairs) : Envir.env * (term * term) list = + let + val t = Envir.norm_term env t0 + and u = Envir.norm_term env u0; + in + (case (head_of t, head_of u) of + (Var (v, T), Var (w, U)) => (*Check for identical variables...*) + if Term.eq_ix (v, w) then (*...occur check would falsely return true!*) + if T = U then (env, add_tpair (rbinder, (t, u), tpairs)) + else raise TERM ("add_ffpair: Var name confusion", [t, u]) + else if Term_Ord.indexname_ord (v, w) = LESS then (*prefer to update the LARGER variable*) + clean_ffpair context ((rbinder, u, t), (env, tpairs)) + else clean_ffpair context ((rbinder, t, u), (env, tpairs)) + | _ => raise TERM ("add_ffpair: Vars expected", [t, u])) + end; + + +(*Print a tracing message + list of dpairs. + In t \ u print u first because it may be rigid or flexible -- + t is always flexible.*) +fun print_dpairs context msg (env, dpairs) = + if Context_Position.is_visible_generic context then + let + fun pdp (rbinder, t, u) = + let + val ctxt = Context.proof_of context; + fun termT t = + Syntax.pretty_term ctxt (Envir.norm_term env (Logic.rlist_abs (rbinder, t))); + val prt = Pretty.blk (0, [termT u, Pretty.str " \\<^sup>?", Pretty.brk 1, termT t]); + in tracing (Pretty.string_of prt) end; + in tracing msg; List.app pdp dpairs end + else (); + + +(*Unify the dpairs in the environment. + Returns flex-flex disagreement pairs NOT IN normal form. + SIMPL may raise exception CANTUNIFY. *) +fun hounifiers binders (context, env, tus : (term * term) list) + : (Envir.env * (term * term) list) Seq.seq = + let + val trace_bound = Config.get_generic context trace_bound; + val search_bound = Config.get_generic context search_bound; + val trace_simp = Config.get_generic context trace_simp; + fun add_unify tdepth ((env, dpairs), reseq) = + Seq.make (fn () => + let + val (env', flexflex, flexrigid) = + (if tdepth > trace_bound andalso trace_simp + then print_dpairs context "Enter SIMPL" (env, dpairs) else (); + SIMPL context (env, dpairs)); + in + (case flexrigid of + [] => SOME (fold_rev (add_ffpair context) flexflex (env', []), reseq) + | dp :: frigid' => + if tdepth > search_bound then + (if Context_Position.is_visible_generic context + then warning "Unification bound exceeded" else (); Seq.pull reseq) + else + (if tdepth > trace_bound then + print_dpairs context "Enter MATCH" (env',flexrigid@flexflex) + else (); + Seq.pull (Seq.it_right + (add_unify (tdepth + 1)) (MATCH context (env',dp, frigid'@flexflex), reseq)))) + end + handle CANTUNIFY => + (if tdepth > trace_bound andalso Context_Position.is_visible_generic context + then tracing "Failure node" + else (); Seq.pull reseq)); + val dps = map (fn (t, u) => (binders, t, u)) tus; + in add_unify 1 ((env, dps), Seq.empty) end; + +fun unifiers (params as (context, env, tus)) = + Seq.cons (fold (Pattern.unify context) tus env, []) Seq.empty + handle Pattern.Unif => Seq.empty + | Pattern.Pattern => hounifiers [] params; + + +(*For smash_flexflex1*) +fun var_head_of (env,t) : indexname * typ = + (case head_of (strip_abs_body (Envir.norm_term env t)) of + Var (v, T) => (v, T) + | _ => raise CANTUNIFY); (*not flexible, cannot use trivial substitution*) + + +(*Eliminate a flex-flex pair by the trivial substitution, see Huet (1975) + Unifies ?f t1 ... rm with ?g u1 ... un by ?f -> \x1...xm. ?a, ?g -> \x1...xn. ?a + Unfortunately, unifies ?f t u with ?g t u by ?f, ?g -> \x y. ?a, + though just ?g->?f is a more general unifier. + Unlike Huet (1975), does not smash together all variables of same type -- + requires more work yet gives a less general unifier (fewer variables). + Handles ?f t1 ... rm with ?f u1 ... um to avoid multiple updates. *) +fun smash_flexflex1 (t, u) env : Envir.env = + let + val vT as (v, T) = var_head_of (env, t) + and wU as (w, U) = var_head_of (env, u); + val (env', var) = Envir.genvar (#1 v) (env, Envir.body_type env T); + val env'' = Envir.vupdate (wU, type_abs (env', U, var)) env'; + in + if vT = wU then env'' (*the other update would be identical*) + else Envir.vupdate (vT, type_abs (env', T, var)) env'' + end; + + +(*Smash all flex-flexpairs. Should allow selection of pairs by a predicate?*) +fun smash_flexflex (env, tpairs) : Envir.env = + fold_rev smash_flexflex1 tpairs env; + +(*Returns unifiers with no remaining disagreement pairs*) +fun smash_unifiers context tus env = + Seq.map smash_flexflex (unifiers (context, env, tus)); + +end; diff --git a/thys/ML_Unification/document/root.bib b/thys/ML_Unification/document/root.bib new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/document/root.bib @@ -0,0 +1,29 @@ +@article{speccheck, + author = {Kevin Kappelmann and Lukas Bulwahn and Sebastian Willenbrink}, + title = {SpecCheck - Specification-Based Testing for Isabelle/ML}, + journal = {Archive of Formal Proofs}, + month = {July}, + year = {2021}, + note = {\url{https://isa-afp.org/entries/SpecCheck.html}, + Formal proof development}, + ISSN = {2150-914x}, +} + +@InProceedings{unif-hints, + author = "Asperti, Andrea + and Ricciotti, Wilmer + and Sacerdoti Coen, Claudio + and Tassi, Enrico", + editor = "Berghofer, Stefan + and Nipkow, Tobias + and Urban, Christian + and Wenzel, Makarius", + title = "Hints in Unification", + booktitle = "Theorem Proving in Higher Order Logics", + year = "2009", + publisher = "Springer Berlin Heidelberg", + address = "Berlin, Heidelberg", + pages = "84--98", + isbn = "978-3-642-03359-9", + doi = "10.1007/978-3-642-03359-9_8" +} diff --git a/thys/ML_Unification/document/root.tex b/thys/ML_Unification/document/root.tex new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/document/root.tex @@ -0,0 +1,54 @@ +\documentclass[11pt,a4paper]{article} +\usepackage[T1]{fontenc} +\usepackage{isabelle,isabellesym} + +% this should be the last package used +\usepackage{pdfsetup} + +% urls in roman style, theory text in math-similar italics +\urlstyle{rm} +\isabellestyle{it} + +\begin{document} + +\title{Unification Utilities for Isabelle/ML} +\author{Kevin Kappelmann} +\maketitle + +\begin{abstract} +This article provides various unification utilities for Isabelle/ML, most prominently: +\begin{enumerate} +\item First-order and higher-order pattern +\href{https://en.wikipedia.org/wiki/Unification_(computer_science)#E-unification}{E-unification} +and E-matching. +While unifiers in Isabelle/ML only consider the $\alpha\beta\eta$-equational theory of the $\lambda$-calculus, +unifiers in this article +may take an extra background theory, in the form of an equational prover, into account. +For example, the unification problem $n + 1 \equiv {}?m + Suc\; 0$ +may be solved by providing a prover for the background theory $\forall n.\ n + 1 \equiv n + Suc\; 0$. +\item Tactics, methods, and attributes with adjustable unifiers (e.g.\ resolution, fact, assumption, OF). +\item A generalisation of unification hints~\cite{unif-hints}. +Unification hints are a flexible extension for unifiers. +Among other things, they can be used for reflective tactics, +to provide canonical unification instances, +or to simply strengthen the background theory of a unifier in a controlled manner. +\item Simplifier integration for e-unifiers. +\item Practical combinations of unification algorithms, e.g. a combination of first-order and +higher-order pattern unification. +\item A hierarchical logger for Isabelle/ML, +including per logger configurations with log levels, output channels, message filters. +\end{enumerate} +While this entry works with every object logic, +some extra setup for Isabelle/HOL and application examples are provided. +All unifiers are tested with SpecCheck~\cite{speccheck}. +\end{abstract} + +\tableofcontents + +% include generated text of all theories +\input{session} + +\bibliographystyle{abbrv} +\bibliography{root} + +\end{document} diff --git a/thys/ML_Unification/unification_base.ML b/thys/ML_Unification/unification_base.ML new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/unification_base.ML @@ -0,0 +1,74 @@ +(* Title: ML_Unification/unification_base.ML + Author: Kevin Kappelmann + +Basic definitions for unifiers. +*) +signature UNIFICATION_BASE = +sig + include HAS_LOGGER + + val reflexive : cterm -> thm + val combination : thm -> thm -> thm + val symmetric : thm -> thm + val abstract_rule : string -> cterm -> thm -> thm + val reflexive_term : Proof.context -> term -> thm + + (*raised on unsupported input*) + exception PATTERN + + (*raised on unification failure for non-sequence outputs (e.g. type unification)*) + exception UNIF + + type type_unifier = Proof.context -> typ * typ -> Envir.env -> Envir.env + type type_matcher = type_unifier + + type closed_unifier = Proof.context -> term * term -> Envir.env -> (Envir.env * thm) Seq.seq + (*term binders stores fresh free variables associated to each loose bound variable*) + type unifier = term Binders.binders -> closed_unifier + type e_unifier = unifier -> unifier + + type closed_matcher = closed_unifier + type matcher = unifier + type e_matcher = matcher -> matcher + + (* normalisers for matchers/unifiers *) + type normalisers = { + norm_unif_thm : Envir_Normalisation.thm_normaliser, + norm_thm : Envir_Normalisation.thm_normaliser, + norm_term : Envir_Normalisation.term_normaliser + } + +end + +structure Unification_Base : UNIFICATION_BASE = +struct + +val logger = Logger.setup_new_logger Logger.root_logger "Unification_Base" + +val reflexive = Thm.reflexive +val combination = Thm.combination +val symmetric = Thm.symmetric +val abstract_rule = Thm.abstract_rule + +val reflexive_term = reflexive oo Thm.cterm_of + +exception PATTERN = Pattern.Pattern + +exception UNIF = Pattern.Unif +type type_unifier = Proof.context -> typ * typ -> Envir.env -> Envir.env +type type_matcher = type_unifier +type closed_unifier = Proof.context -> term * term -> Envir.env -> + (Envir.env * thm) Seq.seq +type unifier = term Binders.binders -> closed_unifier +type e_unifier = unifier -> unifier +type closed_matcher = closed_unifier +type matcher = unifier +type e_matcher = matcher -> matcher + +type normalisers = { + norm_unif_thm : Envir_Normalisation.thm_normaliser, + norm_thm : Envir_Normalisation.thm_normaliser, + norm_term : Envir_Normalisation.term_normaliser +} + +end diff --git a/thys/ML_Unification/unification_util.ML b/thys/ML_Unification/unification_util.ML new file mode 100644 --- /dev/null +++ b/thys/ML_Unification/unification_util.ML @@ -0,0 +1,253 @@ +(* Title: ML_Unification/unification_util.ML + Author: Kevin Kappelmann + +Utilities used for e-unifications. +*) +signature UNIFICATION_UTIL = +sig + include HAS_LOGGER + + (* pretty-printing *) + val pretty_types : Proof.context -> typ list -> Pretty.T + val pretty_terms : Proof.context -> term list -> Pretty.T + + val pretty_tyenv : Proof.context -> Type.tyenv -> Pretty.T + val pretty_tenv : Proof.context -> Envir.tenv -> Pretty.T + val pretty_env : Proof.context -> Envir.env -> Pretty.T + + val pretty_unif_problem : Proof.context -> (term * term) -> Pretty.T + + (* terms and environments *) + val maxidx_of_terms : term list -> int + (*returns empty environment with maxidx set to maximum of given terms*) + val empty_envir : term * term -> Envir.env + + (* type unification *) + + (*raises Unification_Base.UNIF on failure*) + val unify_types : Unification_Base.type_unifier + (*raises Unification_Base.UNIF on failure*) + val match_types : Unification_Base.type_matcher + + (* normalisers *) + val eta_short_norms_match : Unification_Base.normalisers + val beta_eta_short_norms_match : Unification_Base.normalisers + val eta_short_norms_unif : Unification_Base.normalisers + val beta_eta_short_norms_unif : Unification_Base.normalisers + + (* shared standard cases for unifiers *) + val abstract_abstract : Envir_Normalisation.term_type_normaliser -> Unification_Base.unifier -> + term Binders.binders -> Proof.context -> string -> typ -> + (term * term) -> Envir.env -> (Envir.env * thm) Seq.seq + (*raises UNIF if types do not unify*) + val rigid_rigid : Envir_Normalisation.term_type_normaliser -> Unification_Base.type_unifier -> + Proof.context -> term -> (string * typ) -> Envir.env -> (Envir.env * thm) Seq.seq + val bound_bound : term Binders.binders -> Proof.context -> (int * int) -> + thm Seq.seq + val comb_comb : Envir_Normalisation.thm_type_normaliser -> Unification_Base.closed_unifier -> + Proof.context -> (term * term) -> (term * term) -> Envir.env -> (Envir.env * thm) Seq.seq + val args_args : Envir_Normalisation.thm_type_normaliser -> Unification_Base.closed_unifier -> + Proof.context -> (term list * term list) -> (Envir.env * thm) Seq.seq -> (Envir.env * thm) Seq.seq + val strip_comb_strip_comb : Envir_Normalisation.thm_type_normaliser -> Unification_Base.unifier -> + term Binders.binders -> Proof.context -> (term * term) -> + (term list * term list) -> Envir.env -> (Envir.env * thm) Seq.seq + + (* logging *) + val log_unif_result : Proof.context -> term * term -> Envir.env * thm -> unit + val log_unif_results : Proof.context -> term * term -> Unification_Base.closed_unifier -> unit + val log_unif_results' : int -> Proof.context -> term * term -> Unification_Base.closed_unifier -> unit + +end + +structure Unification_Util : UNIFICATION_UTIL = +struct + +val logger = Logger.setup_new_logger Unification_Base.logger "Util" + +structure Norm = Envir_Normalisation + +(* pretty-printing *) +local +val pretty_term = Syntax.pretty_term +val pretty_type = Syntax.pretty_typ + +fun pretty_env_aux show_entry = + Vartab.dest + #> map show_entry + #> Pretty.list "[" "]" + +fun pretty_env_entry show (s, t) = Pretty.block [show s, Pretty.str " := ", show t] + +val pretty_record = + map (fn (key, entry) => Pretty.block [Pretty.str key, Pretty.str "=", entry]) + #> Pretty.enum "," "{" "}" +in + +fun pretty_types ctxt = Pretty.block o Pretty.commas o map (pretty_type ctxt) +fun pretty_terms ctxt = Pretty.block o Pretty.commas o map (pretty_term ctxt) + +fun pretty_tyenv ctxt = + let + val show_entry = pretty_env_entry (pretty_type ctxt) + fun get_typs (v, (s, T)) = (TVar (v, s), T) + in pretty_env_aux (show_entry o get_typs) end + +fun pretty_tenv ctxt = + let + val show_entry = pretty_env_entry (pretty_term ctxt) + fun get_trms (v, (T, t)) = (Var (v, T), t) + in pretty_env_aux (show_entry o get_trms) end + +fun pretty_env ctxt (Envir.Envir {maxidx, tyenv, tenv}) = + pretty_record [ + ("maxidx", Pretty.str (string_of_int maxidx)), + ("tyenv", pretty_tyenv ctxt tyenv), + ("tenv", pretty_tenv ctxt tenv) + ] + +end + +fun pretty_unif_problem ctxt (t1, t2) = + Pretty.block [pretty_terms ctxt [t1], Pretty.str " \\<^sup>? " , pretty_terms ctxt [t2]] + +(* terms and environments *) +fun maxidx_of_terms ts = fold (Integer.max o maxidx_of_term) ts ~1 + +fun empty_envir (t1, t2) = Envir.empty (maxidx_of_terms [t1, t2]) + +(* type unification *) + +fun match_types ctxt (T, U) (env as Envir.Envir {maxidx, tenv, tyenv}) = + (if pointer_eq (T, U) then env + else + let val tyenv' = Sign.typ_match (Proof_Context.theory_of ctxt) (T, U) tyenv + in Envir.Envir {maxidx = maxidx, tenv = tenv, tyenv = tyenv'} end) + handle Type.TYPE_MATCH => + (@{log Logger.DEBUG} ctxt + (fn _ => Pretty.block [ + Pretty.str "Failed to match types ", + pretty_types ctxt [Envir_Normalisation.norm_type_match tyenv T, U] + ] |> Pretty.string_of); + raise Unification_Base.UNIF) + +fun unify_types ctxt (T, U) (env as Envir.Envir {maxidx, tenv, tyenv}) = + (if pointer_eq (T, U) then env + else + let val (tyenv', maxidx') = + Sign.typ_unify (Proof_Context.theory_of ctxt) (T, U) (tyenv, maxidx) + in Envir.Envir {maxidx = maxidx', tenv = tenv, tyenv = tyenv'} end) + handle Type.TUNIFY => + (@{log Logger.DEBUG} ctxt + (fn _ => Pretty.block [ + Pretty.str "Failed to unify types ", + pretty_types ctxt (map (Envir_Normalisation.norm_type_unif tyenv) [T, U]) + ] |> Pretty.string_of); + raise Unification_Base.UNIF) + +(* normalisers *) +val eta_short_norms_match = { + (*Precondition: the matcher must already normalise the types in its theorem!*) + norm_unif_thm = Norm.eta_short_norm_thm (K I) Norm.eta_short_norm_term_match, + norm_thm = Norm.eta_short_norm_thm_match, + norm_term = Norm.eta_short_norm_term_match +} +val beta_eta_short_norms_match = { + (*Precondition: the matcher must already normalise the types in its theorem!*) + norm_unif_thm = Norm.beta_eta_short_norm_thm (K I) Norm.beta_eta_short_norm_term_match, + norm_thm = Norm.beta_eta_short_norm_thm_match, + norm_term = Norm.beta_eta_short_norm_term_match +} +val eta_short_norms_unif = { + norm_unif_thm = Norm.eta_short_norm_thm_unif, + norm_thm = Norm.eta_short_norm_thm_unif, + norm_term = Norm.eta_short_norm_term_unif +} +val beta_eta_short_norms_unif = { + norm_unif_thm = Norm.beta_eta_short_norm_thm_unif, + norm_thm = Norm.beta_eta_short_norm_thm_unif, + norm_term = Norm.beta_eta_short_norm_term_unif +} + +(* shared standard cases for unifiers *) + +fun abstract_abstract norm_term_type unify binders ctxt name T tbp = + let + val (binder as (_, bvar), ctxt') = Binders.fix_binder (name, T) ctxt + fun norm_abstract (Envir.Envir {tyenv,...}) thm = + let val bvar' = norm_term_type tyenv bvar + in Unification_Base.abstract_rule name (Thm.cterm_of ctxt' bvar') thm end + in + unify (binder :: binders) ctxt' tbp + #> Seq.map (fn (env, thm) => (env, norm_abstract env thm)) + end + +fun rigid_rigid norm_term_type unify_types ctxt s (nt, Tt) env = + let val (ns, Ts) = (if is_Const s then dest_Const else dest_Free) s + in + if ns = nt then + let val (env' as Envir.Envir {tyenv,...}) = unify_types ctxt (Ts, Tt) env + in + (env', Unification_Base.reflexive_term ctxt (norm_term_type tyenv s)) + |> Seq.single + end + else Seq.empty + end + +fun bound_bound binders ctxt (i, j) = + if i = j + then Binders.nth_binder_data binders i |> Unification_Base.reflexive_term ctxt |> Seq.single + else Seq.empty + +fun comb_comb norm_thm_types unify ctxt (f, x) (g, y) env = + let + val unif' = unify ctxt + val env_thmq = unif' (f, g) env + |> Seq.maps (fn (env, thm_fg) => unif' (x, y) env |> Seq.map (apsnd (pair thm_fg))) + fun combine (env, thmp) = + (*normalise types for the combination theorem to succeed*) + apply2 (norm_thm_types ctxt env) thmp + |> uncurry Unification_Base.combination + |> pair env + in Seq.map combine env_thmq end + +fun args_args norm_thm_types unify ctxt (ss, ts) env_thmhq = + (let + val args = ss ~~ ts + fun unif_arg tp = Seq.maps + (fn (env, thms) => unify ctxt tp env |> Seq.map (apsnd (fn thm => thm :: thms))) + fun unif_args env = fold unif_arg args (Seq.single (env, [])) + (*combine the theorems*) + fun combine thmh (env_res, arg_thms) = + let + (*normalise types for the combination theorem to succeed*) + val norm_thm_types' = norm_thm_types ctxt env_res + fun norm_combine thm_arg thm = norm_thm_types' thm_arg |> Unification_Base.combination thm + in (env_res, fold_rev norm_combine arg_thms (norm_thm_types' thmh)) end + in Seq.maps (fn (env, thmh) => unif_args env |> Seq.map (combine thmh)) env_thmhq end) + handle ListPair.UnequalLengths => Seq.empty + +fun strip_comb_strip_comb norm_thm_types unify binders ctxt (sh, th) (ss, ts) = + unify binders ctxt (sh, th) + #> args_args norm_thm_types (unify binders) ctxt (ss, ts) + +(* logging *) + +fun log_unif_result ctxt _ (env, thm) = + @{log Logger.INFO} ctxt (fn _ => Pretty.block [ + Pretty.str "Environment: ", + pretty_env ctxt env, + Pretty.fbrk, + Pretty.str "Theorem: ", + Thm.pretty_thm ctxt thm + ] |> Pretty.string_of) + +fun log_unif_results ctxt tp unify = + let val res = unify ctxt tp (empty_envir tp) |> Seq.list_of + in fold (log_unif_result ctxt tp #> K) res () end + +fun log_unif_results' limit ctxt tp unify = + let val res = unify ctxt tp (empty_envir tp) |> Seq.take limit |> Seq.list_of + in fold (log_unif_result ctxt tp #> K) res () end + +end + diff --git a/thys/ROOTS b/thys/ROOTS --- a/thys/ROOTS +++ b/thys/ROOTS @@ -1,767 +1,768 @@ ABY3_Protocols ADS_Functor AI_Planning_Languages_Semantics AODV AOT AVL-Trees AWN Abortable_Linearizable_Modules Abs_Int_ITP2012 Abstract-Hoare-Logics Abstract-Rewriting Abstract_Completeness Abstract_Soundness Ackermanns_not_PR Actuarial_Mathematics Adaptive_State_Counting Affine_Arithmetic Aggregation_Algebras Akra_Bazzi Algebraic_Numbers Algebraic_VCs Allen_Calculus Amicable_Numbers Amortized_Complexity AnselmGod Applicative_Lifting Approximation_Algorithms Architectural_Design_Patterns Aristotles_Assertoric_Syllogistic Arith_Prog_Rel_Primes ArrowImpossibilityGS Attack_Trees Auto2_HOL Auto2_Imperative_HOL AutoFocus-Stream Automated_Stateful_Protocol_Verification Automatic_Refinement AxiomaticCategoryTheory BDD BD_Security_Compositional BNF_CC BNF_Operations BTree Balog_Szemeredi_Gowers Banach_Steinhaus Belief_Revision Bell_Numbers_Spivey BenOr_Kozen_Reif Berlekamp_Zassenhaus Bernoulli Bertrands_Postulate Bicategory BinarySearchTree Binary_Code_Imprimitive Binding_Syntax_Theory Binomial-Heaps Binomial-Queues BirdKMP Birkhoff_Finite_Distributive_Lattices Blue_Eyes Bondy Boolean_Expression_Checkers Boolos_Curious_Inference Boolos_Curious_Inference_Automated Bounded_Deducibility_Security Buchi_Complementation Budan_Fourier Buffons_Needle Buildings BytecodeLogicJmlTypes C2KA_DistributedSystems CAVA_Automata CAVA_LTL_Modelchecker CCS CHERI-C_Memory_Model CISC-Kernel CRDT CRYSTALS-Kyber CSP_RefTK CVP_Hardness CYK CZH_Elementary_Categories CZH_Foundations CZH_Universal_Constructions CakeML CakeML_Codegen Call_Arity Card_Equiv_Relations Card_Multisets Card_Number_Partitions Card_Partitions Cartan_FP Case_Labeling Catalan_Numbers Category Category2 Category3 Catoids Cauchy Cayley_Hamilton Certification_Monads Ceva Chandy_Lamport Chord_Segments Circus Clean Clique_and_Monotone_Circuits ClockSynchInst Closest_Pair_Points CoCon CoSMeDis CoSMed CofGroups Coinductive Coinductive_Languages Collections Combinable_Wands Combinatorial_Enumeration_Algorithms Combinatorics_Words Combinatorics_Words_Graph_Lemma Combinatorics_Words_Lyndon CommCSL Commuting_Hermitian Comparison_Sort_Lower_Bound Compiling-Exceptions-Correctly Complete_Non_Orders Completeness Complex_Bounded_Operators Complex_Geometry Complx ComponentDependencies ConcurrentGC ConcurrentIMP Concurrent_Ref_Alg Concurrent_Revisions Conditional_Simplification Conditional_Transfer_Rule Consensus_Refined Constructive_Cryptography Constructive_Cryptography_CM Constructor_Funs Containers Cook_Levin CoreC++ Core_DOM Core_SC_DOM Correctness_Algebras Cotangent_PFD_Formula Count_Complex_Roots CryptHOL CryptoBasedCompositionalProperties Crypto_Standards Cubic_Quartic_Equations DCR-ExecutionEquivalence DFS_Framework DOM_Components DPRM_Theorem DPT-SAT-Solver DataRefinementIBP Datatype_Order_Generator Decl_Sem_Fun_PL Decreasing-Diagrams Decreasing-Diagrams-II Dedekind_Real Deep_Learning Delta_System_Lemma Density_Compiler Dependent_SIFUM_Refinement Dependent_SIFUM_Type_Systems Depth-First-Search Derangements Deriving Descartes_Sign_Rule Design_Theory Dict_Construction Differential_Dynamic_Logic Differential_Game_Logic DigitsInBase Digit_Expansions Dijkstra_Shortest_Path Diophantine_Eqns_Lin_Hom Directed_Sets Dirichlet_L Dirichlet_Series DiscretePricing Discrete_Summation DiskPaxos Distributed_Distinct_Elements Dominance_CHK DynamicArchitectures Dynamic_Tables E_Transcendental Earley_Parser Echelon_Form EdmondsKarp_Maxflow Edwards_Elliptic_Curves_Group Efficient-Mergesort Efficient_Weighted_Path_Order Elliptic_Curves_Group_Law Encodability_Process_Calculi Epistemic_Logic Equivalence_Relation_Enumeration Ergodic_Theory Error_Function Euler_MacLaurin Euler_Partition Euler_Polyhedron_Formula Eval_FO Example-Submission Executable_Randomized_Algorithms Expander_Graphs Extended_Finite_State_Machine_Inference Extended_Finite_State_Machines FFT FLP FOL-Fitting FOL_Axiomatic FOL_Harrison FOL_Seq_Calc1 FOL_Seq_Calc2 FOL_Seq_Calc3 FO_Theory_Rewriting FSM_Tests Factor_Algebraic_Polynomial Factored_Transition_System_Bounding Falling_Factorial_Sum Farkas FeatherweightJava Featherweight_OCL Fermat3_4 FileRefinement FinFun Finger-Trees Finite-Map-Extras Finite_Automata_HF Finite_Fields Finitely_Generated_Abelian_Groups First_Order_Terms First_Welfare_Theorem Fishburn_Impossibility Fisher_Yates Fishers_Inequality Fixed_Length_Vector Flow_Networks Floyd_Warshall Flyspeck-Tame FocusStreamsCaseStudies Forcing Formal_Puiseux_Series Formal_SSA Formula_Derivatives Foundation_of_geometry Fourier Free-Boolean-Algebra Free-Groups Frequency_Moments Fresh_Identifiers FunWithFunctions FunWithTilings Functional-Automata Functional_Ordered_Resolution_Prover Furstenberg_Topology GPU_Kernel_PL Gabow_SCC GaleStewart_Games Gale_Shapley Game_Based_Crypto Gauss-Jordan-Elim-Fun Gauss_Jordan Gauss_Sums Gaussian_Integers GenClock General-Triangle Generalized_Counting_Sort Generic_Deriving Generic_Join GewirthPGCProof Girth_Chromatic Given_Clause_Loops GoedelGod Goedel_HFSet_Semantic Goedel_HFSet_Semanticless Goedel_Incompleteness Goodstein_Lambda GraphMarkingIBP Graph_Saturation Graph_Theory Gray_Codes Green Groebner_Bases Groebner_Macaulay Gromov_Hyperbolicity Grothendieck_Schemes Group-Ring-Module HOL-CSP HOLCF-Prelude HRB-Slicing Hahn_Jordan_Decomposition Hales_Jewett Heard_Of Hello_World HereditarilyFinite Hermite Hermite_Lindemann Hidden_Markov_Models Higher_Order_Terms HoareForDivergence Hoare_Time Hood_Melville_Queue HotelKeyCards Huffman Hybrid_Logic Hybrid_Multi_Lane_Spatial_Logic Hybrid_Systems_VCs HyperCTL HyperHoareLogic Hyperdual Hypergraph_Basics IEEE_Floating_Point IFC_Tracking IMAP-CRDT IMO2019 IMP2 IMP2_Binary_Heap IMP_Compiler IMP_Compiler_Reuse IP_Addresses Imperative_Insertion_Sort Implicational_Logic Impossible_Geometry Incompleteness Incredible_Proof_Machine Independence_CH Inductive_Confidentiality Inductive_Inference InfPathElimination InformationFlowSlicing InformationFlowSlicing_Inter Integration Interpolation_Polynomials_HOL_Algebra Interpreter_Optimizations Interval_Arithmetic_Word32 Intro_Dest_Elim Involutions2Squares Iptables_Semantics Irrational_Series_Erdos_Straus Irrationality_J_Hancl Irrationals_From_THEBOOK IsaGeoCoq IsaNet Isabelle_C Isabelle_Marries_Dirac Isabelle_Meta_Model Jacobson_Basic_Algebra Jinja JinjaDCI JinjaThreads JiveDataStoreModel Jordan_Hoelder Jordan_Normal_Form KAD KAT_and_DRA KBPs KD_Tree Key_Agreement_Strong_Adversaries Khovanskii_Theorem Kleene_Algebra Kneser_Cauchy_Davenport Knights_Tour Knot_Theory Knuth_Bendix_Order Knuth_Morris_Pratt Koenigsberg_Friendship Kruskal Kuratowski_Closure_Complement LLL_Basis_Reduction LLL_Factorization LOFT LP_Duality LTL LTL_Master_Theorem LTL_Normal_Form LTL_to_DRA LTL_to_GBA Lam-ml-Normalization LambdaAuth LambdaMu Lambda_Free_EPO Lambda_Free_KBOs Lambda_Free_RPOs Lambert_W Landau_Symbols Laplace_Transform Latin_Square LatticeProperties Launchbury Laws_of_Large_Numbers Lazy-Lists-II Lazy_Case Lehmer Lifting_Definition_Option Lifting_the_Exponent LightweightJava LinearQuantifierElim Linear_Inequalities Linear_Programming Linear_Recurrences Liouville_Numbers List-Index List-Infinite List_Interleaving List_Inversions List_Update LocalLexing Localization_Ring Locally-Nameless-Sigma Logging_Independent_Anonymity Lovasz_Local Lowe_Ontological_Argument Lower_Semicontinuous Lp Lucas_Theorem MDP-Algorithms MDP-Rewards MFMC_Countable MFODL_Monitor_Optimized MFOTL_Monitor MSO_Regex_Equivalence Markov_Models Marriage Mason_Stothers Matrices_for_ODEs Matrix Matrix_Tensor Matroids Max-Card-Matching Maximum_Segment_Sum Median_Method Median_Of_Medians_Selection Menger Mereology Mersenne_Primes Metalogic_ProofChecker MHComputation MiniML MiniSail Minimal_SSA Minkowskis_Theorem Minsky_Machines MLSS_Decision_Proc +ML_Unification Modal_Logics_for_NTS Modular_Assembly_Kit_Security Modular_arithmetic_LLL_and_HNF_algorithms Monad_Memo_DP Monad_Normalisation MonoBoolTranAlgebra MonoidalCategory Monomorphic_Monad MuchAdoAboutTwo Multi_Party_Computation Multirelations Multirelations_Heterogeneous Multiset_Ordering_NPC Multitape_To_Singletape_TM Myhill-Nerode Name_Carrying_Type_Inference Nano_JSON Nash_Williams Nat-Interval-Logic Native_Word Nested_Multisets_Ordinals Network_Security_Policy_Verification Neumann_Morgenstern_Utility No_FTL_observers No_FTL_observers_Gen_Rel Nominal2 Noninterference_CSP Noninterference_Concurrent_Composition Noninterference_Generic_Unwinding Noninterference_Inductive_Unwinding Noninterference_Ipurge_Unwinding Noninterference_Sequential_Composition NormByEval Nullstellensatz Number_Theoretic_Transform Octonions OpSets Open_Induction Optics Optimal_BST Orbit_Stabiliser Order_Lattice_Props Ordered_Resolution_Prover Ordinal Ordinal_Partitions Ordinals_and_Cardinals Ordinary_Differential_Equations PAC_Checker PAL PAPP_Impossibility PCF PLM POPLmark-deBruijn PSemigroupsConvolution Package_logic Padic_Field Padic_Ints Pairing_Heap Paraconsistency Parity_Game Partial_Function_MR Partial_Order_Reduction Password_Authentication_Protocol Pell Perfect-Number-Thm Perron_Frobenius Physical_Quantities Pi_Calculus Pi_Transcendental Planarity_Certificates Pluennecke_Ruzsa_Inequality Poincare_Bendixson Poincare_Disc Polygonal_Number_Theorem Polynomial_Factorization Polynomial_Interpolation Polynomials Pop_Refinement Posix-Lexing Possibilistic_Noninterference Power_Sum_Polynomials Pratt_Certificate Prefix_Free_Code_Combinators Presburger-Automata Prim_Dijkstra_Simple Prime_Distribution_Elementary Prime_Harmonic_Series Prime_Number_Theorem Priority_Queue_Braun Priority_Search_Trees Probabilistic_Noninterference Probabilistic_Prime_Tests Probabilistic_System_Zoo Probabilistic_Timed_Automata Probabilistic_While Probability_Inequality_Completeness Program-Conflict-Analysis Progress_Tracking Projective_Geometry Projective_Measurements Promela Proof_Strategy_Language PropResPI Propositional_Logic_Class Propositional_Proof_Systems Prpu_Maxflow PseudoHoops Psi_Calculi Ptolemys_Theorem Public_Announcement_Logic QHLProver QR_Decomposition Quantales Quantales_Converse Quantifier_Elimination_Hybrid Quasi_Borel_Spaces Quaternions Query_Optimization Quick_Sort_Cost RIPEMD-160-SPARK ROBDD RSAPSS Ramsey-Infinite Random_BSTs Random_Graph_Subgraph_Threshold Randomised_BSTs Randomised_Social_Choice Rank_Nullity_Theorem Real_Impl Real_Power Real_Time_Deque Recursion-Addition Recursion-Theory-I Refine_Imperative_HOL Refine_Monadic RefinementReactive Regex_Equivalence Registers Regression_Test_Selection Regular-Sets Regular_Algebras Regular_Tree_Relations Relation_Algebra Relational-Incorrectness-Logic Relational_Cardinality Relational_Disjoint_Set_Forests Relational_Forests Relational_Method Relational_Minimum_Spanning_Trees Relational_Paths Rensets Rep_Fin_Groups ResiduatedTransitionSystem Residuated_Lattices Resolution_FOL Rewrite_Properties_Reduction Rewriting_Z Ribbon_Proofs Risk_Free_Lending Robbins-Conjecture Robinson_Arithmetic Root_Balanced_Tree Roth_Arithmetic_Progressions Routing Roy_Floyd_Warshall SATSolverVerification SCC_Bloemen_Sequential SC_DOM_Components SDS_Impossibility SIFPL SIFUM_Type_Systems SPARCv8 Safe_Distance Safe_OCL Safe_Range_RC Saturation_Framework Saturation_Framework_Extensions Sauer_Shelah_Lemma Schutz_Spacetime Schwartz_Zippel Secondary_Sylow Security_Protocol_Refinement Selection_Heap_Sort SenSocialChoice Separata Separation_Algebra Separation_Logic_Imperative_HOL Separation_Logic_Unbounded SequentInvertibility Shadow_DOM Shadow_SC_DOM Shivers-CFA ShortestPath Show Sigma_Commit_Crypto Signature_Groebner Simpl Simple_Clause_Learning Simple_Firewall Simplex Simplicial_complexes_and_boolean_functions SimplifiedOntologicalArgument Skew_Heap Skip_Lists Slicing Sliding_Window_Algorithm Smith_Normal_Form Smooth_Manifolds Solidity Sophomores_Dream Sort_Encodings Source_Coding_Theorem SpecCheck Special_Function_Bounds Splay_Tree Sqrt_Babylonian Stable_Matching Stalnaker_Logic Statecharts Stateful_Protocol_Composition_and_Typing Stellar_Quorums Stern_Brocot Stewart_Apollonius Stirling_Formula Stochastic_Matrices Stone_Algebras Stone_Kleene_Relation_Algebras Stone_Relation_Algebras Store_Buffer_Reduction Stream-Fusion Stream_Fusion_Code StrictOmegaCategories Strong_Security Sturm_Sequences Sturm_Tarski Stuttering_Equivalence Subresultants Subset_Boolean_Algebras SumSquares Sunflowers SuperCalc Suppes_Theorem Surprise_Paradox Symmetric_Polynomials Syntax_Independent_Logic Synthetic_Completeness Szemeredi_Regularity Szpilrajn TESL_Language TLA Tail_Recursive_Functions Tarskis_Geometry Taylor_Models Three_Circles Three_Squares Timed_Automata Topological_Semantics Topology TortoiseHare TsirelsonBound Transcendence_Series_Hancl_Rucki Transformer_Semantics Transition_Systems_and_Automata Transitive-Closure Transitive-Closure-II Transitive_Models Treaps Tree-Automata Tree_Decomposition Tree_Enumeration Triangle Trie Turans_Graph_Theorem Twelvefold_Way Two_Generated_Word_Monoids_Intersection Tycon Types_Tableaus_and_Goedels_God Types_To_Sets_Extension UPF UPF_Firewall UTP Undirected_Graph_Theory Universal_Hash_Families Universal_Turing_Machine UpDown_Scheme VYDRA_MDL Valuation Van_Emde_Boas_Trees Van_der_Waerden VectorSpace VeriComp Verified-Prover Verified_SAT_Based_AI_Planning VerifyThis2018 VerifyThis2019 Vickrey_Clarke_Groves Virtual_Substitution VolpanoSmith WHATandWHERE_Security WOOT_Strong_Eventual_Consistency WebAssembly Weight_Balanced_Trees Weighted_Arithmetic_Geometric_Mean Weighted_Path_Order Well_Quasi_Orders Wetzels_Problem Winding_Number_Eval Word_Lib WorkerWrapper X86_Semantics XML Youngs_Inequality ZFC_in_HOL Zeckendorf Zeta_3_Irrational Zeta_Function pGCL