| 2113 |      1 | (*  Title:      HOL/Subst/unifier.ML
 | 
|  |      2 |     ID:         $Id$
 | 
|  |      3 |     Author:     Martin Coen, Cambridge University Computer Laboratory
 | 
|  |      4 |     Copyright   1993  University of Cambridge
 | 
|  |      5 | 
 | 
|  |      6 | For Unifier.thy.
 | 
|  |      7 | Properties of unifiers.
 | 
|  |      8 | *)
 | 
|  |      9 | 
 | 
|  |     10 | open Unifier;
 | 
|  |     11 | 
 | 
|  |     12 | val unify_defs = [Unifier_def,MoreGeneral_def,MGUnifier_def];
 | 
|  |     13 | 
 | 
|  |     14 | val rews = subst_rews@al_rews@setplus_rews;
 | 
|  |     15 | 
 | 
|  |     16 | (*---------------------------------------------------------------------------
 | 
|  |     17 |  * Unifiers 
 | 
|  |     18 |  *---------------------------------------------------------------------------*)
 | 
|  |     19 | 
 | 
|  |     20 | goalw Unifier.thy [Unifier_def] "Unifier s t u = (t <| s = u <| s)";
 | 
|  |     21 | by (rtac refl 1);
 | 
|  |     22 | val Unifier_iff = result();
 | 
|  |     23 | 
 | 
|  |     24 | goal Unifier.thy
 | 
|  |     25 |     "Unifier s (Comb t u) (Comb v w) --> Unifier s t v & Unifier s u w";
 | 
|  |     26 | by (simp_tac (!simpset addsimps [Unifier_iff]) 1);
 | 
|  |     27 | val Unifier_Comb  = result() RS mp RS conjE;
 | 
|  |     28 | 
 | 
|  |     29 | goal Unifier.thy
 | 
|  |     30 |   "~v : vars_of(t) --> ~v : vars_of(u) --> Unifier s t u --> \
 | 
|  |     31 | \  Unifier ((v,r)#s) t u";
 | 
|  |     32 | by (simp_tac (!simpset addsimps [Unifier_iff,repl_invariance]) 1);
 | 
|  |     33 | val Cons_Unifier  = result() RS mp RS mp RS mp;
 | 
|  |     34 | 
 | 
|  |     35 | 
 | 
|  |     36 | (*---------------------------------------------------------------------------
 | 
|  |     37 |  * Most General Unifiers 
 | 
|  |     38 |  *---------------------------------------------------------------------------*)
 | 
|  |     39 | 
 | 
|  |     40 | goalw Unifier.thy unify_defs "MGUnifier s t u = MGUnifier s u t";
 | 
|  |     41 | by (safe_tac (HOL_cs addSEs ([sym]@([spec] RL [mp]))));
 | 
|  |     42 | val mgu_sym = result();
 | 
|  |     43 | 
 | 
|  |     44 | 
 | 
|  |     45 | goalw Unifier.thy [MoreGeneral_def]  "r >> s = (EX q. s =s= r <> q)";
 | 
|  |     46 | by (rtac refl 1);
 | 
|  |     47 | val MoreGen_iff = result();
 | 
|  |     48 | 
 | 
|  |     49 | 
 | 
|  |     50 | goal Unifier.thy  "[] >> s";
 | 
|  |     51 | by (simp_tac (!simpset addsimps (MoreGen_iff::subst_rews)) 1);
 | 
|  |     52 | by (fast_tac (set_cs addIs [refl RS subst_refl]) 1);
 | 
|  |     53 | val MoreGen_Nil = result();
 | 
|  |     54 | 
 | 
|  |     55 | 
 | 
|  |     56 | goalw Unifier.thy unify_defs
 | 
|  |     57 |     "MGUnifier s t u = (ALL r. Unifier r t u = s >> r)";
 | 
|  |     58 | by (REPEAT (ares_tac [iffI,allI] 1 ORELSE 
 | 
|  |     59 |             eresolve_tac [conjE,allE,mp,exE,ssubst_subst2] 1));
 | 
|  |     60 | by (asm_simp_tac (!simpset addsimps [subst_comp]) 1);
 | 
|  |     61 | by (fast_tac (set_cs addIs [comp_Nil RS sym RS subst_refl]) 1);
 | 
|  |     62 | val MGU_iff = result();
 | 
|  |     63 | 
 | 
|  |     64 | 
 | 
|  |     65 | val [prem] = goal Unifier.thy
 | 
|  |     66 |      "~ Var(v) <: t ==> MGUnifier [(v,t)] (Var v) t";
 | 
|  |     67 | by(simp_tac(HOL_ss addsimps([MGU_iff,MoreGen_iff,Unifier_iff]@rews))1);
 | 
|  |     68 | by (REPEAT_SOME (step_tac set_cs));
 | 
|  |     69 | by (etac subst 1);
 | 
|  |     70 | by (etac ssubst_subst2 2);
 | 
|  |     71 | by (rtac (Cons_trivial RS subst_sym) 1);
 | 
|  |     72 | by (simp_tac (!simpset addsimps ((prem RS Var_not_occs)::rews)) 1);
 | 
|  |     73 | val MGUnifier_Var = result();
 | 
|  |     74 | 
 | 
|  |     75 | 
 | 
|  |     76 | 
 | 
|  |     77 | (*---------------------------------------------------------------------------
 | 
|  |     78 |  * Idempotence.
 | 
|  |     79 |  *---------------------------------------------------------------------------*)
 | 
|  |     80 | 
 | 
|  |     81 | goalw Unifier.thy [Idem_def] "Idem([])";
 | 
|  |     82 | by (simp_tac (!simpset addsimps (refl RS subst_refl)::rews) 1);
 | 
|  |     83 | qed "Idem_Nil";
 | 
|  |     84 | 
 | 
|  |     85 | goalw Unifier.thy [Idem_def] "Idem(s) = (sdom(s) Int srange(s) = {})";
 | 
|  |     86 | by (simp_tac (!simpset addsimps [subst_eq_iff,subst_comp,
 | 
|  |     87 |                                  invariance,dom_range_disjoint])1);
 | 
|  |     88 | qed "Idem_iff";
 | 
|  |     89 | 
 | 
|  |     90 | val rews = subst_rews@al_rews@setplus_rews;
 | 
|  |     91 | goal Unifier.thy "~ (Var(v) <: t) --> Idem([(v,t)])";
 | 
|  |     92 | by (simp_tac (!simpset addsimps (vars_iff_occseq::Idem_iff::srange_iff::rews)
 | 
|  |     93 |                         setloop (split_tac [expand_if])) 1);
 | 
|  |     94 | by (fast_tac set_cs 1);
 | 
|  |     95 | val Var_Idem  = store_thm("Var_Idem", result() RS mp);
 | 
|  |     96 | 
 | 
|  |     97 | 
 | 
|  |     98 | val [prem] = goalw Unifier.thy [Idem_def]
 | 
|  |     99 | "Idem(r) ==>  Unifier s (t<|r) (u<|r) --> Unifier (r <> s) (t <| r) (u <| r)";
 | 
|  |    100 | by (simp_tac (!simpset addsimps 
 | 
|  |    101 |               [Unifier_iff,subst_comp,prem RS comp_subst_subst]) 1);
 | 
|  |    102 | val Unifier_Idem_subst  = store_thm("Unifier_Idem_subst", result() RS mp);
 | 
|  |    103 | 
 | 
|  |    104 | val [p1,p2,p3] = goal Unifier.thy
 | 
|  |    105 |      "[| Idem(r); Unifier s (t <| r) (u <| r); \
 | 
|  |    106 | \        (!q. Unifier q (t <| r) (u <| r) --> s <> q =s= q) \
 | 
|  |    107 | \     |] ==> Idem(r <> s)";
 | 
|  |    108 | 
 | 
|  |    109 | by (cut_facts_tac [p1,
 | 
|  |    110 |                    p2 RS (p1 RS Unifier_Idem_subst RS (p3 RS spec RS mp))] 1);
 | 
|  |    111 | by (REPEAT_SOME (etac rev_mp));
 | 
|  |    112 | by (simp_tac (!simpset addsimps [Idem_def,subst_eq_iff,subst_comp]) 1);
 | 
|  |    113 | qed "Idem_comp";
 | 
|  |    114 | 
 |