Unification/Mgu.thy
changeset 107 5c816239deaa
equal deleted inserted replaced
106:ed54ec416bb3 107:5c816239deaa
       
     1 
       
     2 
       
     3 theory Mgu = Main + Terms + Fresh + Equ + Substs:
       
     4 
       
     5 (* unification problems *)
       
     6 
       
     7 syntax
       
     8  "_equ_prob"   :: "trm \<Rightarrow> trm \<Rightarrow> (trm\<times>trm)"        ("_ \<approx>? _" [81,81] 81)
       
     9  "_fresh_prob" :: "string \<Rightarrow> trm \<Rightarrow> (string\<times>trm)"  ("_ \<sharp>? _" [81,81] 81)
       
    10 
       
    11 translations 
       
    12   "t1 \<approx>? t2" \<rightharpoonup> "(t1,t2)"
       
    13   " a \<sharp>? t"  \<rightharpoonup> "(a,t)"
       
    14 
       
    15 (* all solutions for a unification problem *)
       
    16 
       
    17 types 
       
    18   problem_type = "((trm\<times>trm)list) \<times> ((string\<times>trm)list)"
       
    19   unifier_type = "fresh_envs \<times> substs"
       
    20 
       
    21 consts 
       
    22   U :: "problem_type \<Rightarrow> (unifier_type set)"
       
    23 defs all_solutions_def : 
       
    24   "U P  \<equiv> {(nabla,s). 
       
    25              (\<forall> (t1,t2)\<in>set (fst P). nabla \<turnstile> subst s t1 \<approx> subst s t2) \<and> 
       
    26              (\<forall>   (a,t)\<in>set (snd P). nabla \<turnstile> a \<sharp> subst s t) }"
       
    27 
       
    28 (* set of variables in unification problems *)
       
    29 
       
    30 consts
       
    31   vars_fprobs :: "((string\<times>trm) list) \<Rightarrow> (string set)"
       
    32   vars_eprobs :: "((trm\<times>trm)list) \<Rightarrow> (string set)"
       
    33   vars_probs  :: "problem_type \<Rightarrow> nat"
       
    34 primrec
       
    35   "vars_fprobs [] = {}"
       
    36   "vars_fprobs (x#xs) = (vars_trm (snd x))\<union>(vars_fprobs xs)"
       
    37 primrec
       
    38   "vars_eprobs [] = {}"
       
    39   "vars_eprobs (x#xs) = (vars_trm (snd x))\<union>(vars_trm (fst x))\<union>(vars_eprobs xs)"
       
    40 defs
       
    41   vars_probs_def: "vars_probs P \<equiv> card((vars_fprobs (snd P))\<union>(vars_eprobs (fst P)))"
       
    42 
       
    43 
       
    44 (* most general unifier *)
       
    45 
       
    46 consts 
       
    47   mgu :: "problem_type \<Rightarrow> unifier_type \<Rightarrow> bool"
       
    48 defs mgu_def:
       
    49   "mgu P unif \<equiv> 
       
    50    \<forall> (nabla,s1)\<in> U P. (\<exists> s2. (nabla\<Turnstile>(subst s2) (fst unif)) \<and> 
       
    51                               (nabla\<Turnstile>subst (s2 \<bullet>(snd unif)) \<approx> subst s1))"
       
    52 
       
    53 (* idempotency of a unifier *)
       
    54 
       
    55 consts 
       
    56   idem :: "unifier_type \<Rightarrow> bool"
       
    57 defs 
       
    58   idem_def: "idem unif \<equiv> (fst unif)\<Turnstile> subst ((snd unif)\<bullet>(snd unif)) \<approx> subst (snd unif)"
       
    59 
       
    60 (* application of a substitution to a problem *)
       
    61 
       
    62 consts 
       
    63   apply_subst :: "substs \<Rightarrow> problem_type \<Rightarrow> problem_type"
       
    64 defs apply_subst_def: 
       
    65   "apply_subst s P \<equiv> (map (\<lambda>(t1,t2). (subst s t1 \<approx>? subst s t2)) (fst P),
       
    66                       map (\<lambda>(a,t).   (a \<sharp>? (subst s t)) ) (snd P))" 
       
    67 
       
    68 (* equality reductions *)
       
    69 
       
    70 consts 
       
    71   s_red :: "(problem_type \<times> substs \<times> problem_type) set"
       
    72 syntax 
       
    73   "_s_red" :: "problem_type \<Rightarrow> substs \<Rightarrow> problem_type \<Rightarrow> bool"  ("_ \<turnstile> _ \<leadsto> _ " [80,80,80] 80)
       
    74 translations "P1 \<turnstile>sigma\<leadsto> P2" \<rightleftharpoons> "(P1,sigma,P2)\<in>s_red"
       
    75 inductive s_red
       
    76 intros
       
    77   unit_sred[intro!,dest!]:    "((Unit\<approx>?Unit)#xs,ys) \<turnstile>[]\<leadsto> (xs,ys)"
       
    78   paar_sred[intro!,dest!]:    "((Paar t1 t2\<approx>?Paar s1 s2)#xs,ys) \<turnstile>[]\<leadsto> ((t1\<approx>?s1)#(t2\<approx>?s2)#xs,ys)"
       
    79   func_sred[intro!,dest!]:    "((Func F t1\<approx>?Func F t2)#xs,ys) \<turnstile>[]\<leadsto> ((t1\<approx>?t2)#xs,ys)"
       
    80   abst_aa_sred[intro!,dest!]: "((Abst a t1\<approx>?Abst a t2)#xs,ys) \<turnstile>[]\<leadsto> ((t1\<approx>?t2)#xs,ys)"
       
    81   abst_ab_sred[intro!]: "a\<noteq>b\<Longrightarrow> 
       
    82                        ((Abst a t1\<approx>?Abst b t2)#xs,ys) \<turnstile>[]\<leadsto> ((t1\<approx>?swap [(a,b)] t2)#xs,(a\<sharp>?t2)#ys)"
       
    83   atom_sred[intro!,dest!]:    "((Atom a\<approx>?Atom a)#xs,ys) \<turnstile>[]\<leadsto> (xs,ys)"  
       
    84   susp_sred[intro!,dest!]:    "((Susp pi1 X\<approx>?Susp pi2 X)#xs,ys) 
       
    85                                 \<turnstile>[]\<leadsto> (xs,(map (\<lambda>a. a\<sharp>? Susp [] X) (ds_list pi1 pi2))@ys)"
       
    86   var_1_sred[intro!]:   "\<not>(occurs X t)\<Longrightarrow>((Susp pi X\<approx>?t)#xs,ys)
       
    87                                \<turnstile>[(X,swap (rev pi) t)]\<leadsto> apply_subst [(X,swap (rev pi) t)] (xs,ys)"
       
    88   var_2_sred[intro!]:   "\<not>(occurs X t)\<Longrightarrow>((t\<approx>?Susp pi X)#xs,ys) 
       
    89                                \<turnstile>[(X,swap (rev pi) t)]\<leadsto> apply_subst [(X,swap (rev pi) t)] (xs,ys)"
       
    90 
       
    91 (* freshness reductions *)
       
    92 
       
    93 consts 
       
    94   c_red :: "(problem_type \<times> fresh_envs \<times> problem_type) set"
       
    95 syntax 
       
    96   "_c_red" :: "problem_type \<Rightarrow> fresh_envs \<Rightarrow> problem_type \<Rightarrow> bool" ("_ \<turnstile> _ \<rightarrow> _ " [80,80,80] 80)
       
    97 translations "P1 \<turnstile>nabla\<rightarrow> P2" \<rightleftharpoons> "(P1,nabla,P2)\<in>c_red"
       
    98 inductive c_red
       
    99 intros
       
   100   unit_cred[intro!]:    "([],(a \<sharp>? Unit)#xs) \<turnstile>{}\<rightarrow> ([],xs)"
       
   101   paar_cred[intro!]:    "([],(a \<sharp>? Paar t1 t2)#xs) \<turnstile>{}\<rightarrow> ([],(a\<sharp>?t1)#(a\<sharp>?t2)#xs)"
       
   102   func_cred[intro!]:    "([],(a \<sharp>? Func F t)#xs) \<turnstile>{}\<rightarrow> ([],(a\<sharp>?t)#xs)"
       
   103   abst_aa_cred[intro!]: "([],(a \<sharp>? Abst a t)#xs) \<turnstile>{}\<rightarrow> ([],xs)"
       
   104   abst_ab_cred[intro!]: "a\<noteq>b\<Longrightarrow>([],(a \<sharp>? Abst b t)#xs) \<turnstile>{}\<rightarrow> ([],(a\<sharp>?t)#xs)"
       
   105   atom_cred[intro!]:    "a\<noteq>b\<Longrightarrow>([],(a \<sharp>? Atom b)#xs) \<turnstile>{}\<rightarrow> ([],xs)"  
       
   106   susp_cred[intro!]:    "([],(a \<sharp>? Susp pi X)#xs) \<turnstile>{((swapas (rev pi) a),X)}\<rightarrow> ([],xs)" 
       
   107 
       
   108 (* unification reduction sequence *)
       
   109 
       
   110 consts 
       
   111   red_plus :: "(problem_type \<times> unifier_type \<times> problem_type) set"
       
   112 syntax 
       
   113   red_plus :: "problem_type \<Rightarrow> unifier_type \<Rightarrow> problem_type \<Rightarrow> bool" ("_ \<Turnstile> _ \<Rightarrow> _ " [80,80,80] 80)
       
   114 
       
   115 translations "P1 \<Turnstile>(nabla,s)\<Rightarrow> P2" \<rightleftharpoons> "(P1,(nabla,s),P2)\<in>red_plus"
       
   116 inductive red_plus
       
   117 intros
       
   118   sred_single[intro!]: "\<lbrakk>P1\<turnstile>s1\<leadsto>P2\<rbrakk>\<Longrightarrow>P1\<Turnstile>({},s1)\<Rightarrow>P2"
       
   119   cred_single[intro!]: "\<lbrakk>P1\<turnstile>nabla1\<rightarrow>P2\<rbrakk>\<Longrightarrow>P1\<Turnstile>(nabla1,[])\<Rightarrow>P2"
       
   120   sred_step[intro!]:   "\<lbrakk>P1\<turnstile>s1\<leadsto>P2;P2\<Turnstile>(nabla2,s2)\<Rightarrow>P3\<rbrakk>\<Longrightarrow>P1\<Turnstile>(nabla2,(s2\<bullet>s1))\<Rightarrow>P3"
       
   121   cred_step[intro!]:   "\<lbrakk>P1\<turnstile>nabla1\<rightarrow>P2;P2\<Turnstile>(nabla2,[])\<Rightarrow>P3\<rbrakk>\<Longrightarrow>P1\<Turnstile>(nabla2\<union>nabla1,[])\<Rightarrow>P3"
       
   122 
       
   123 lemma mgu_idem: 
       
   124   "\<lbrakk>(nabla1,s1)\<in>U P; 
       
   125    \<forall>(nabla2,s2)\<in> U P. nabla2\<Turnstile>(subst s2) nabla1 \<and>  nabla2\<Turnstile>subst(s2 \<bullet> s1)\<approx>subst s2 \<rbrakk>\<Longrightarrow>
       
   126    mgu P (nabla1,s1)\<and>idem (nabla1,s1)"
       
   127 apply(rule conjI)
       
   128 apply(simp only: mgu_def)
       
   129 apply(rule ballI)
       
   130 apply(simp)
       
   131 apply(drule_tac x="x" in bspec)
       
   132 apply(assumption)
       
   133 apply(force)
       
   134 apply(drule_tac x="(nabla1,s1)" in bspec)
       
   135 apply(assumption)
       
   136 apply(simp add: idem_def)
       
   137 done
       
   138 
       
   139 lemma problem_subst_comm: "((nabla,s2)\<in>U (apply_subst s1 P)) = ((nabla,(s2\<bullet>s1))\<in>U P)"
       
   140 apply(simp add: all_solutions_def apply_subst_def)
       
   141 apply(auto)
       
   142 apply(drule_tac x="(a,b)" in bspec, assumption, simp add: subst_comp_expand)+
       
   143 done
       
   144 
       
   145 lemma P1_to_P2_sred: 
       
   146   "\<lbrakk>(nabla1,s1)\<in>U P1; P1 \<turnstile>s2\<leadsto> P2 \<rbrakk>\<Longrightarrow>((nabla1,s1)\<in>U P2) \<and> (nabla1\<Turnstile>subst (s1\<bullet>s2)\<approx>subst s1)"
       
   147 apply(ind_cases "P1 \<turnstile>s2\<leadsto> P2")
       
   148 apply(simp_all)
       
   149 --Unit
       
   150 apply(force intro!: equ_refl simp add: all_solutions_def ext_subst_def subst_equ_def subst_susp)
       
   151 --Paar
       
   152 apply(simp add: all_solutions_def ext_subst_def subst_equ_def subst_susp)
       
   153 apply(force intro!: equ_refl dest!: equ_paar_elim)
       
   154 --Func
       
   155 apply(simp add: all_solutions_def ext_subst_def subst_equ_def subst_susp)
       
   156 apply(force intro!: equ_refl dest!: equ_func_elim)
       
   157 --Abst.aa
       
   158 apply(simp add: all_solutions_def ext_subst_def subst_equ_def subst_susp)
       
   159 apply(force intro!: equ_refl dest!: equ_abst_aa_elim)
       
   160 --Abst.ab
       
   161 apply(simp add: all_solutions_def ext_subst_def subst_equ_def subst_susp)
       
   162 apply(force intro!: equ_refl dest!: equ_abst_ab_elim simp add: subst_swap_comm)
       
   163 --Atom
       
   164 apply(simp add: all_solutions_def ext_subst_def subst_equ_def subst_susp)
       
   165 apply(force intro!: equ_refl)
       
   166 --Susp
       
   167 apply(rule conjI)
       
   168 apply(auto)
       
   169 apply(simp add: all_solutions_def)
       
   170 apply(erule conjE)+
       
   171 apply(simp add: ds_list_equ_ds)
       
   172 apply(simp only: subst_susp)
       
   173 apply(drule equ_pi1_pi2_dec[THEN mp])
       
   174 apply(auto)
       
   175 apply(drule_tac x="aa" in bspec)
       
   176 apply(assumption)
       
   177 apply(simp add: subst_susp)
       
   178 apply(simp add: subst_equ_def subst_susp)
       
   179 apply(rule ballI)
       
   180 apply(rule equ_refl)
       
   181 --Var.one
       
   182 apply(drule_tac "t2.1"="swap (rev pi) t" in subst_not_occurs[THEN mp])
       
   183 apply(simp only: problem_subst_comm)
       
   184 apply(simp add: all_solutions_def ext_subst_def subst_equ_def)
       
   185 apply(rule conjI)
       
   186 apply(rule ballI)
       
   187 apply(erule conjE)+
       
   188 apply(drule unif_1)
       
   189 apply(clarify)
       
   190 apply(drule_tac x="(a,b)" in bspec)
       
   191 apply(assumption)
       
   192 apply(simp)
       
   193 apply(simp add: unif_2a)
       
   194 apply(erule conjE)+
       
   195 apply(drule unif_1)
       
   196 apply(rule ballI)
       
   197 apply(clarify)
       
   198 apply(drule_tac x="(a,b)" in bspec)
       
   199 apply(assumption)
       
   200 apply(simp)
       
   201 apply(simp add: unif_2b)
       
   202 apply(rule unif_1)
       
   203 apply(simp add: all_solutions_def)
       
   204 --Var.two
       
   205 apply(drule_tac "t2.1"="swap (rev pi) t" in subst_not_occurs[THEN mp])
       
   206 apply(simp only: problem_subst_comm)
       
   207 apply(simp add: all_solutions_def ext_subst_def subst_equ_def)
       
   208 apply(auto)
       
   209 apply(drule_tac x="(a,b)" in bspec)
       
   210 apply(assumption)
       
   211 apply(simp)
       
   212 apply(drule equ_sym)
       
   213 apply(drule unif_1)
       
   214 apply(simp add: unif_2a)
       
   215 apply(drule_tac x="(a,b)" in bspec)
       
   216 apply(assumption)
       
   217 apply(simp)
       
   218 apply(drule equ_sym)
       
   219 apply(drule unif_1)
       
   220 apply(simp add: unif_2b)
       
   221 apply(rule unif_1)
       
   222 apply(rule equ_sym)
       
   223 apply(simp add: all_solutions_def)
       
   224 done
       
   225 
       
   226 lemma P1_from_P2_sred: "\<lbrakk>(nabla1,s1)\<in>U P2; P1\<turnstile>s2\<leadsto>P2\<rbrakk>\<Longrightarrow>(nabla1,s1\<bullet>s2)\<in>U P1"
       
   227 apply(ind_cases "P1 \<turnstile>s2\<leadsto> P2")
       
   228 --Susp.Paar.Func.Abst.aa
       
   229 apply(simp add: all_solutions_def, force)
       
   230 apply(simp add: all_solutions_def, force)
       
   231 apply(simp add: all_solutions_def, force)
       
   232 apply(simp add: all_solutions_def, force)
       
   233 --Abst.ab
       
   234 apply(simp only: all_solutions_def)
       
   235 apply(force simp add: subst_swap_comm)
       
   236 --Atom
       
   237 apply(simp only: all_solutions_def, force)
       
   238 --Susp
       
   239 apply(simp)
       
   240 apply(auto)
       
   241 apply(simp add: all_solutions_def)
       
   242 apply(simp add: ds_list_equ_ds)
       
   243 apply(subgoal_tac "nabla1\<turnstile>(swap pi1 (subst s1 (Susp [] X)))\<approx>(swap pi2 (subst s1 (Susp [] X)))")
       
   244 apply(simp add: subst_susp subst_swap_comm)
       
   245 apply(simp add: subst_susp subst_swap_comm)
       
   246 apply(rule equ_pi1_pi2_add[THEN mp])
       
   247 apply(drule conjunct2)
       
   248 apply(auto)
       
   249 apply(drule_tac x="(a,Susp [] X)" in bspec)
       
   250 apply(auto)
       
   251 apply(simp add: subst_susp)
       
   252 --Var.one
       
   253 apply(simp only: problem_subst_comm)
       
   254 apply(simp only: all_solutions_def)
       
   255 apply(simp)
       
   256 apply(simp only: subst_comp_expand)
       
   257 apply(subgoal_tac "subst [(X, swap (rev pi) t)] t = t")--A
       
   258 apply(simp add: subst_susp)
       
   259 apply(simp only: subst_swap_comm)
       
   260 apply(simp only: equ_pi_to_right[THEN sym])
       
   261 apply(simp only: equ_involutive_right)
       
   262 apply(rule equ_refl)
       
   263 --A
       
   264 apply(rule subst_not_occurs[THEN mp])
       
   265 apply(assumption)
       
   266 --Var.two
       
   267 apply(simp only: problem_subst_comm)
       
   268 apply(simp only: all_solutions_def)
       
   269 apply(simp)
       
   270 apply(simp only: subst_comp_expand)
       
   271 apply(subgoal_tac "subst [(X, swap (rev pi) t)] t = t")--B
       
   272 apply(simp add: subst_susp)
       
   273 apply(simp only: subst_swap_comm)
       
   274 apply(simp only: equ_pi_to_left[THEN sym])
       
   275 apply(simp only: equ_involutive_left)
       
   276 apply(rule equ_refl)
       
   277 --B
       
   278 apply(rule subst_not_occurs[THEN mp])
       
   279 apply(assumption)
       
   280 done
       
   281 
       
   282 lemma P1_to_P2_cred: 
       
   283   "\<lbrakk>(nabla1,s1)\<in>U P1; P1 \<turnstile>nabla2\<rightarrow> P2 \<rbrakk>\<Longrightarrow>((nabla1,s1)\<in>U P2) \<and> (nabla1\<Turnstile>(subst s1) nabla2)"
       
   284 apply(ind_cases " P1\<turnstile>nabla2\<rightarrow>P2")
       
   285 apply(simp_all)
       
   286 apply(auto simp add: ext_subst_def all_solutions_def)
       
   287 apply(rule fresh_swap_left[THEN mp])
       
   288 apply(simp add: subst_swap_comm[THEN sym] subst_susp)
       
   289 done
       
   290 
       
   291 lemma P1_from_P2_cred: 
       
   292   "\<lbrakk>(nabla1,s1)\<in>U P2; P1 \<turnstile>nabla2\<rightarrow> P2; nabla3\<Turnstile>(subst s1) nabla2\<rbrakk>\<Longrightarrow>(nabla1\<union>nabla3,s1)\<in>U P1"
       
   293 apply(ind_cases "P1 \<turnstile>nabla2\<rightarrow> P2")
       
   294 apply(simp_all)
       
   295 apply(auto simp add: ext_subst_def all_solutions_def fresh_weak)
       
   296 apply(simp add: subst_susp)
       
   297 apply(rule fresh_swap_right[THEN mp])
       
   298 apply(drule_tac "nabla2.1"="nabla1" in fresh_weak[THEN mp])
       
   299 apply(subgoal_tac "nabla3 \<union> nabla1=nabla1 \<union> nabla3")--A
       
   300 apply(simp)
       
   301 apply(rule Un_commute)
       
   302 done
       
   303 
       
   304 lemma P1_to_P2_red_plus: "\<lbrakk>P1 \<Turnstile>(nabla,s)\<Rightarrow>P2\<rbrakk>\<Longrightarrow> (nabla1,s1)\<in>U P1 \<longrightarrow>
       
   305   ((nabla1,s1)\<in>U P2)\<and>(nabla1\<Turnstile>subst (s1\<bullet>s)\<approx>subst s1)\<and>(nabla1\<Turnstile>(subst s1) nabla)"
       
   306 apply(erule red_plus.induct)
       
   307 -- sred
       
   308 apply(rule impI)
       
   309 apply(drule_tac "P2.0"="P2" and "s2.0"="s1a" in P1_to_P2_sred)
       
   310 apply(force)
       
   311 apply(rule conjI, force)+
       
   312 apply(force simp add: ext_subst_def)
       
   313 -- cred
       
   314 apply(rule impI)
       
   315 apply(drule_tac "P2.0"="P2" and "nabla2.0"="nabla1a" in P1_to_P2_cred)
       
   316 apply(assumption)
       
   317 apply(force intro!: equ_refl simp add: subst_equ_def)
       
   318 -- sred
       
   319 apply(rule impI)
       
   320 apply(drule_tac "P2.0"="P2" and "s2.0"="s1a" in P1_to_P2_sred)
       
   321 apply(assumption)
       
   322 apply(erule conjE)+
       
   323 apply(rule conjI)
       
   324 apply(force)
       
   325 apply(rule conjI)
       
   326 apply(drule mp)
       
   327 apply(assumption)
       
   328 apply(erule conjE)+
       
   329 apply(rule_tac "s2.0"="((s1\<bullet>s2)\<bullet>s1a)" in subst_trans)
       
   330 apply(simp only: subst_assoc subst_equ_def)
       
   331 apply(rule ballI)
       
   332 apply(rule equ_refl)
       
   333 apply(rule_tac "s2.0"="(s1\<bullet>s1a)" in subst_trans)
       
   334 apply(rule subst_cancel_right)
       
   335 apply(assumption)
       
   336 apply(assumption)
       
   337 apply(force)
       
   338 -- cred
       
   339 apply(rule impI)
       
   340 apply(drule_tac "P2.0"="P2" and "nabla2.0"="nabla1a" in P1_to_P2_cred)
       
   341 apply(auto simp add: ext_subst_def)
       
   342 done
       
   343 
       
   344 lemma P1_from_P2_red_plus: "\<lbrakk>P1 \<Turnstile>(nabla,s)\<Rightarrow>P2\<rbrakk>\<Longrightarrow>(nabla1,s1)\<in>U P2\<longrightarrow>
       
   345         nabla3\<Turnstile>(subst s1)(nabla)\<longrightarrow>(nabla1\<union>nabla3,(s1\<bullet>s))\<in>U P1"
       
   346 apply(erule red_plus.induct)
       
   347 -- sred
       
   348 apply(rule impI)+
       
   349 apply(drule_tac "P1.0"="P1" and "s2.0"="s1a" in P1_from_P2_sred)
       
   350 apply(assumption)
       
   351 apply(force simp only: all_solutions_def equ_weak fresh_weak)
       
   352 -- cred
       
   353 apply(rule impI)+
       
   354 apply(drule_tac "P1.0"="P1" and "nabla3.0"="nabla3" and "nabla2.0"="nabla1a" in P1_from_P2_cred)
       
   355 apply(assumption)+
       
   356 apply(simp add: all_solutions_def)
       
   357 -- sred
       
   358 apply(rule impI)+
       
   359 apply(simp)
       
   360 apply(drule_tac "P1.0"="P1" and "P2.0"="P2" and "s2.0"="s1a" in P1_from_P2_sred)
       
   361 apply(assumption)
       
   362 apply(simp add: all_solutions_def subst_assoc)
       
   363 -- cred
       
   364 apply(rule impI)+
       
   365 apply(subgoal_tac "nabla3 \<Turnstile> (subst s1) nabla2")--A
       
   366 apply(simp)
       
   367 apply(drule_tac "P1.0"="P1" and "P2.0"="P2" and
       
   368                 "nabla2.0"="nabla1a" and "nabla3.0"="nabla3" in P1_from_P2_cred)
       
   369 apply(assumption)
       
   370 apply(simp)
       
   371 apply(simp add: ext_subst_def)
       
   372 apply(subgoal_tac "nabla1 \<union> nabla3 \<union> nabla3=nabla1 \<union> nabla3")--A
       
   373 apply(simp)
       
   374 --A
       
   375 apply(force)
       
   376 --B
       
   377 apply(simp add: ext_subst_def)
       
   378 done
       
   379 
       
   380 lemma mgu: "\<lbrakk>P \<Turnstile>(nabla,s)\<Rightarrow>([],[])\<rbrakk>\<Longrightarrow> mgu P (nabla,s) \<and> idem (nabla,s)"
       
   381 apply(frule_tac "nabla3.2"="nabla" and "nabla2"="nabla" and 
       
   382                 "s1.2"="[]" and "nabla1.2"="{}"
       
   383                 in P1_from_P2_red_plus[THEN mp,THEN mp])
       
   384 apply(force simp add: all_solutions_def)
       
   385 apply(force simp add: ext_subst_def)
       
   386 apply(rule mgu_idem)
       
   387 apply(simp add: all_solutions_def)
       
   388 apply(rule ballI)
       
   389 apply(clarify)
       
   390 apply(drule_tac  "nabla1.0"="a" and "s1.0"="b"in P1_to_P2_red_plus)
       
   391 apply(simp)
       
   392 done  
       
   393 
       
   394 end
       
   395 
       
   396 
       
   397 
       
   398 
       
   399 
       
   400 
       
   401 
       
   402 
       
   403