Unification/Unification.thy
changeset 107 5c816239deaa
equal deleted inserted replaced
106:ed54ec416bb3 107:5c816239deaa
       
     1 
       
     2 
       
     3 theory Unification = Main + Terms + Fresh + Equ + Substs + Mgu:
       
     4 
       
     5 (* problems to which no reduction applies *)
       
     6 
       
     7 consts stuck :: "problem_type set"
       
     8 defs
       
     9   stuck_def: "stuck \<equiv> { P1. \<not>(\<exists>P2 nabla s. P1 \<Turnstile>(nabla,s)\<Rightarrow>P2)}"
       
    10 
       
    11 (* all problems which are stuck and have no unifier *)
       
    12 
       
    13 consts fail :: "problem_type set"
       
    14 inductive fail
       
    15 intros
       
    16 [intro!]: "\<lbrakk>occurs X t\<rbrakk>\<Longrightarrow>(Susp pi X\<approx>?Abst a t#xs,ys)\<in>fail"
       
    17 [intro!]: "\<lbrakk>occurs X t\<rbrakk>\<Longrightarrow>(Susp pi X\<approx>?Func F t#xs,ys)\<in>fail"
       
    18 [intro!]: "\<lbrakk>occurs X t1\<or>occurs X t2\<rbrakk>\<Longrightarrow>(Susp pi X\<approx>?Paar t1 t2#xs,ys)\<in>fail"
       
    19 [intro!]: "\<lbrakk>occurs X t\<rbrakk>\<Longrightarrow>(Abst a t\<approx>?Susp pi X#xs,ys)\<in>fail"
       
    20 [intro!]: "\<lbrakk>occurs X t\<rbrakk>\<Longrightarrow>(Func F t\<approx>?Susp pi X#xs,ys)\<in>fail"
       
    21 [intro!]: "\<lbrakk>occurs X t1\<or>occurs X t2\<rbrakk>\<Longrightarrow>(Paar t1 t2\<approx>?Susp pi X#xs,ys)\<in>fail"
       
    22 [intro!,dest!]: "([],a\<sharp>? Atom a#ys)\<in>fail"
       
    23 [intro!]: "a\<noteq>b\<Longrightarrow>(Atom a\<approx>? Atom b#xs,ys)\<in>fail"
       
    24 [intro!,dest!]: "(Abst a t\<approx>?Unit#xs,ys)\<in>fail"
       
    25 [intro!,dest!]: "(Unit\<approx>?Abst a t#xs,ys)\<in>fail"
       
    26 [intro!,dest!]: "(Abst a t\<approx>?Atom b#xs,ys)\<in>fail"
       
    27 [intro!,dest!]: "(Atom b\<approx>?Abst a t#xs,ys)\<in>fail"
       
    28 [intro!,dest!]: "(Abst a t\<approx>?Paar t1 t2#xs,ys)\<in>fail"
       
    29 [intro!,dest!]: "(Paar t1 t2\<approx>?Abst a t#xs,ys)\<in>fail"
       
    30 [intro!,dest!]: "(Abst a t\<approx>?Func F t1#xs,ys)\<in>fail"
       
    31 [intro!,dest!]: "(Func F t1\<approx>?Abst a t#xs,ys)\<in>fail"
       
    32 [intro!,dest!]: "(Unit\<approx>?Atom b#xs,ys)\<in>fail"
       
    33 [intro!,dest!]: "(Atom b\<approx>?Unit#xs,ys)\<in>fail"
       
    34 [intro!,dest!]: "(Unit\<approx>?Paar t1 t2#xs,ys)\<in>fail"
       
    35 [intro!,dest!]: "(Paar t1 t2\<approx>?Unit#xs,ys)\<in>fail"
       
    36 [intro!,dest!]: "(Unit\<approx>?Func F t1#xs,ys)\<in>fail"
       
    37 [intro!,dest!]: "(Func F t1\<approx>?Unit#xs,ys)\<in>fail"
       
    38 [intro!,dest!]: "(Atom a\<approx>?Paar t1 t2#xs,ys)\<in>fail"
       
    39 [intro!,dest!]: "(Paar t1 t2\<approx>?Atom a#xs,ys)\<in>fail"
       
    40 [intro!,dest!]: "(Atom a\<approx>?Func F t1#xs,ys)\<in>fail"
       
    41 [intro!,dest!]: "(Func F t1\<approx>?Atom a#xs,ys)\<in>fail"
       
    42 [intro!,dest!]: "(Func F t\<approx>?Paar t1 t2#xs,ys)\<in>fail"
       
    43 [intro!,dest!]: "(Paar t1 t2\<approx>?Func F t#xs,ys)\<in>fail"
       
    44 [intro!]: "\<lbrakk>F1\<noteq>F2\<rbrakk>\<Longrightarrow>(Func F1 t1\<approx>?Func F2 t2#xs,ys)\<in>fail"
       
    45 
       
    46 (* the results that are interesting are the stuck ones *)
       
    47 
       
    48 consts 
       
    49   results :: "problem_type \<Rightarrow> problem_type set"
       
    50 defs
       
    51   results_def: 
       
    52   "results P1 \<equiv> if P1\<in>stuck then {P1} else {P2. \<exists>nabla s. P1\<Turnstile>(nabla,s)\<Rightarrow>P2 \<and> P2\<in>stuck}"
       
    53 
       
    54 (* a "failed" problem has no unifier *)
       
    55 
       
    56 lemma fail_then_empty: 
       
    57   "(P1\<in>fail) \<Longrightarrow> (U P1={})"
       
    58 apply(erule fail.cases)
       
    59 apply(simp add: all_solutions_def)
       
    60 apply(rule allI)+
       
    61 apply(rule impI)
       
    62 apply(drule_tac nabla1="aa" and s1="b" and "pi1.1"="pi" in occurs_sub_trm_equ[THEN mp])
       
    63 apply(subgoal_tac "\<forall>t1\<in>psub_trms (Abst a (subst b t)).\<not>(\<exists>pi2. aa\<turnstile>Abst a (subst b t)\<approx>swap pi2 t1)")
       
    64 apply(simp)
       
    65 apply(drule equ_sym)
       
    66 apply(clarify)
       
    67 apply(drule_tac "t1.0"="Abst a (subst b t)" and 
       
    68                 "t2.0"="subst b (Susp pi X)" and 
       
    69                 "t3.0"="swap pia t2" in equ_trans)
       
    70 apply(simp)
       
    71 apply(best)
       
    72 apply(rule psub_trm_not_equ)
       
    73 apply(simp add: all_solutions_def)
       
    74 apply(rule allI)+
       
    75 apply(rule impI)
       
    76 apply(drule_tac nabla1="a" and s1="b" and "pi1.1"="pi" in occurs_sub_trm_equ[THEN mp])
       
    77 apply(subgoal_tac "\<forall>t1\<in>psub_trms (Func F (subst b t)).\<not>(\<exists>pi2. a\<turnstile>Func F (subst b t)\<approx>swap pi2 t1)")
       
    78 apply(simp)
       
    79 apply(drule equ_sym)
       
    80 apply(clarify)
       
    81 apply(drule_tac "t1.0"="Func F (subst b t)" and 
       
    82                 "t2.0"="subst b (Susp pi X)" and 
       
    83                 "t3.0"="swap pia t2" in equ_trans)
       
    84 apply(simp)
       
    85 apply(best)
       
    86 apply(rule psub_trm_not_equ)
       
    87 apply(simp add: all_solutions_def)
       
    88 apply(rule allI)+
       
    89 apply(rule impI)
       
    90 apply(erule disjE)
       
    91 apply(drule_tac nabla1="a" and s1="b" and "pi1.1"="pi" in occurs_sub_trm_equ[THEN mp])
       
    92 apply(subgoal_tac "\<forall>t3\<in>psub_trms (Paar (subst b t1) (subst b t2)).
       
    93                                 \<not>(\<exists>pi2. a\<turnstile>Paar (subst b t1) (subst b t2)\<approx>swap pi2 t3)")
       
    94 apply(simp)
       
    95 apply(drule equ_sym)
       
    96 apply(clarify)
       
    97 apply(drule_tac "t1.0"="Paar (subst b t1) (subst b t2)" and 
       
    98                 "t2.0"="subst b (Susp pi X)" and 
       
    99                 "t3.0"="swap pia t2a" in equ_trans)
       
   100 apply(simp)
       
   101 apply(best)
       
   102 apply(rule psub_trm_not_equ)
       
   103 apply(drule_tac nabla1="a" and s1="b" and "pi1.1"="pi" in occurs_sub_trm_equ[THEN mp])
       
   104 apply(subgoal_tac "\<forall>t3\<in>psub_trms (Paar (subst b t1) (subst b t2)).
       
   105                                 \<not>(\<exists>pi2. a\<turnstile>Paar (subst b t1) (subst b t2)\<approx>swap pi2 t3)")
       
   106 apply(simp)
       
   107 apply(drule equ_sym)
       
   108 apply(clarify)
       
   109 apply(drule_tac "t1.0"="Paar (subst b t1) (subst b t2)" and 
       
   110                 "t2.0"="subst b (Susp pi X)" and 
       
   111                 "t3.0"="swap pia t2a" in equ_trans)
       
   112 apply(simp)
       
   113 apply(best)
       
   114 apply(rule psub_trm_not_equ)
       
   115 apply(simp add: all_solutions_def)
       
   116 apply(rule allI)+
       
   117 apply(rule impI)
       
   118 apply(drule_tac nabla1="aa" and s1="b" and "pi1.1"="pi" in occurs_sub_trm_equ[THEN mp])
       
   119 apply(subgoal_tac "\<forall>t3\<in>psub_trms (Abst a (subst b t)).
       
   120                                 \<not>(\<exists>pi2. aa\<turnstile>Abst a (subst b t)\<approx>swap pi2 t3)")
       
   121 apply(simp)
       
   122 apply(clarify)
       
   123 apply(drule_tac "t1.0"="Abst a (subst b t)" and 
       
   124                 "t2.0"="subst b (Susp pi X)" and 
       
   125                 "t3.0"="swap pia t2" in equ_trans)
       
   126 apply(simp)
       
   127 apply(best)
       
   128 apply(rule psub_trm_not_equ)
       
   129 apply(simp add: all_solutions_def)
       
   130 apply(rule allI)+
       
   131 apply(rule impI)
       
   132 apply(drule_tac nabla1="a" and s1="b" and "pi1.1"="pi" in occurs_sub_trm_equ[THEN mp])
       
   133 apply(subgoal_tac "\<forall>t1\<in>psub_trms (Func F (subst b t)).\<not>(\<exists>pi2. a\<turnstile>Func F (subst b t)\<approx>swap pi2 t1)")
       
   134 apply(simp)
       
   135 apply(clarify)
       
   136 apply(drule_tac "t1.0"="Func F (subst b t)" and 
       
   137                 "t2.0"="subst b (Susp pi X)" and 
       
   138                 "t3.0"="swap pia t2" in equ_trans)
       
   139 apply(simp)
       
   140 apply(best)
       
   141 apply(rule psub_trm_not_equ)
       
   142 apply(simp add: all_solutions_def)
       
   143 apply(rule allI)+
       
   144 apply(rule impI)
       
   145 apply(erule disjE)
       
   146 apply(drule_tac nabla1="a" and s1="b" and "pi1.1"="pi" in occurs_sub_trm_equ[THEN mp])
       
   147 apply(subgoal_tac "\<forall>t3\<in>psub_trms (Paar (subst b t1) (subst b t2)).
       
   148                                 \<not>(\<exists>pi2. a\<turnstile>Paar (subst b t1) (subst b t2)\<approx>swap pi2 t3)")
       
   149 apply(simp)
       
   150 apply(clarify)
       
   151 apply(drule_tac "t1.0"="Paar (subst b t1) (subst b t2)" and 
       
   152                 "t2.0"="subst b (Susp pi X)" and 
       
   153                 "t3.0"="swap pia t2a" in equ_trans)
       
   154 apply(simp)
       
   155 apply(best)
       
   156 apply(rule psub_trm_not_equ)
       
   157 apply(drule_tac nabla1="a" and s1="b" and "pi1.1"="pi" in occurs_sub_trm_equ[THEN mp])
       
   158 apply(subgoal_tac "\<forall>t3\<in>psub_trms (Paar (subst b t1) (subst b t2)).
       
   159                                 \<not>(\<exists>pi2. a\<turnstile>Paar (subst b t1) (subst b t2)\<approx>swap pi2 t3)")
       
   160 apply(simp)
       
   161 apply(clarify)
       
   162 apply(drule_tac "t1.0"="Paar (subst b t1) (subst b t2)" and 
       
   163                 "t2.0"="subst b (Susp pi X)" and 
       
   164                 "t3.0"="swap pia t2a" in equ_trans)
       
   165 apply(simp)
       
   166 apply(best)
       
   167 apply(rule psub_trm_not_equ)
       
   168 apply(simp add: all_solutions_def, fast dest!: fresh.cases)
       
   169 apply(simp add: all_solutions_def, fast dest!: equ.cases)+
       
   170 done
       
   171 
       
   172 (* the only stuck problems are the "failed" problems and the empty problem *)
       
   173 
       
   174 lemma stuck_equiv: "stuck = {([],[])}\<union>fail"
       
   175 apply(subgoal_tac "([],[])\<in>stuck")
       
   176 apply(subgoal_tac "\<forall>P\<in>fail. P\<in>stuck")
       
   177 apply(subgoal_tac "\<forall>P\<in>stuck. P=([],[]) \<or> P\<in>fail")
       
   178 apply(force)
       
   179 apply(rule ballI)
       
   180 apply(thin_tac "([], []) \<in> stuck")
       
   181 apply(thin_tac "\<forall>P\<in>fail. P \<in> stuck")
       
   182 apply(simp add: stuck_def)
       
   183 apply(clarify)
       
   184 apply(case_tac a)
       
   185 apply(simp)
       
   186 apply(case_tac b)
       
   187 apply(simp)
       
   188 apply(simp)
       
   189 apply(case_tac aa)
       
   190 apply(simp)
       
   191 apply(case_tac ba)
       
   192 apply(simp_all)
       
   193 apply(case_tac "ab=lista")
       
   194 apply(force)
       
   195 apply(force)
       
   196 apply(force)
       
   197 apply(force)
       
   198 apply(force)
       
   199 apply(force)
       
   200 apply(force)
       
   201 apply(case_tac aa)
       
   202 apply(simp)
       
   203 apply(case_tac ab)
       
   204 apply(simp_all)
       
   205 apply(case_tac ba)
       
   206 apply(simp_all)
       
   207 apply(case_tac "lista=listb")
       
   208 apply(force)
       
   209 apply(force)
       
   210 apply(case_tac "occurs list2 (Abst lista trm)")
       
   211 apply(force)
       
   212 apply(drule_tac x="fst (apply_subst [(list2,swap (rev list1) (Abst lista trm))] (list,b))" in spec)
       
   213 apply(drule_tac x="snd (apply_subst [(list2,swap (rev list1) (Abst lista trm))] (list,b))" in spec)
       
   214 apply(drule_tac x="{}" in spec)
       
   215 apply(drule_tac x="[(list2, swap (rev list1) (Abst lista trm))]" in spec)
       
   216 apply(simp only: surjective_pairing[THEN sym])
       
   217 apply(force)
       
   218 apply(force)
       
   219 apply(force)
       
   220 apply(force)
       
   221 apply(force)
       
   222 apply(case_tac ba)
       
   223 apply(simp_all)
       
   224 apply(case_tac "occurs list2 (Abst lista trm)")
       
   225 apply(force)
       
   226 apply(drule_tac x="fst (apply_subst [(list2,swap (rev list1) (Abst lista trm))] (list,b))" in spec)
       
   227 apply(drule_tac x="snd (apply_subst [(list2,swap (rev list1) (Abst lista trm))] (list,b))" in spec)
       
   228 apply(drule_tac x="{}" in spec)
       
   229 apply(drule_tac x="[(list2, swap (rev list1) (Abst lista trm))]" in spec)
       
   230 apply(simp only: surjective_pairing[THEN sym])
       
   231 apply(force)
       
   232 apply(case_tac "list2=list2a")
       
   233 apply(force)
       
   234 apply(case_tac "occurs list2 (Susp list1a list2a)")
       
   235 apply(simp)
       
   236 apply(drule_tac 
       
   237     x="fst (apply_subst [(list2,swap (rev list1) (Susp list1a list2a))] (list,b))" in spec)
       
   238 apply(drule_tac 
       
   239     x="snd (apply_subst [(list2,swap (rev list1) (Susp list1a list2a))] (list,b))" in spec)
       
   240 apply(drule_tac x="{}" in spec)
       
   241 apply(drule_tac x="[(list2, swap (rev list1) (Susp list1a list2a))]" in spec)
       
   242 apply(simp only: surjective_pairing[THEN sym])
       
   243 apply(force)
       
   244 apply(case_tac "occurs list2 Unit")
       
   245 apply(simp)
       
   246 apply(drule_tac x="fst (apply_subst [(list2,swap (rev list1) Unit)] (list,b))" in spec)
       
   247 apply(drule_tac x="snd (apply_subst [(list2,swap (rev list1) Unit)] (list,b))" in spec)
       
   248 apply(drule_tac x="{}" in spec)
       
   249 apply(drule_tac x="[(list2, swap (rev list1) Unit)]" in spec)
       
   250 apply(simp only: surjective_pairing[THEN sym])
       
   251 apply(force)
       
   252 apply(case_tac "occurs list2 (Atom lista)")
       
   253 apply(simp)
       
   254 apply(drule_tac x="fst (apply_subst [(list2,swap (rev list1) (Atom lista))] (list,b))" in spec)
       
   255 apply(drule_tac x="snd (apply_subst [(list2,swap (rev list1) (Atom lista))] (list,b))" in spec)
       
   256 apply(drule_tac x="{}" in spec)
       
   257 apply(drule_tac x="[(list2, swap (rev list1) (Atom lista))]" in spec)
       
   258 apply(simp only: surjective_pairing[THEN sym])
       
   259 apply(force)
       
   260 apply(case_tac "occurs list2 (Paar trm1 trm2)")
       
   261 apply(force)
       
   262 apply(drule_tac x="fst (apply_subst [(list2,swap (rev list1) (Paar trm1 trm2))] (list,b))" in spec)
       
   263 apply(drule_tac x="snd (apply_subst [(list2,swap (rev list1) (Paar trm1 trm2))] (list,b))" in spec)
       
   264 apply(drule_tac x="{}" in spec)
       
   265 apply(drule_tac x="[(list2, swap (rev list1) (Paar trm1 trm2))]" in spec)
       
   266 apply(simp only: surjective_pairing[THEN sym])
       
   267 apply(force)
       
   268 apply(case_tac "occurs list2 (Func lista trm)")
       
   269 apply(force)
       
   270 apply(drule_tac x="fst (apply_subst [(list2,swap (rev list1) (Func lista trm))] (list,b))" in spec)
       
   271 apply(drule_tac x="snd (apply_subst [(list2,swap (rev list1) (Func lista trm))] (list,b))" in spec)
       
   272 apply(drule_tac x="{}" in spec)
       
   273 apply(drule_tac x="[(list2, swap (rev list1) (Func lista trm))]" in spec)
       
   274 apply(simp only: surjective_pairing[THEN sym])
       
   275 apply(force)
       
   276 apply(case_tac ba)
       
   277 apply(simp_all)
       
   278 apply(force)
       
   279 apply(case_tac "occurs list2 Unit")
       
   280 apply(simp)
       
   281 apply(drule_tac x="fst (apply_subst [(list2,swap (rev list1) Unit)] (list,b))" in spec)
       
   282 apply(drule_tac x="snd (apply_subst [(list2,swap (rev list1) Unit)] (list,b))" in spec)
       
   283 apply(drule_tac x="{}" in spec)
       
   284 apply(drule_tac x="[(list2, swap (rev list1) Unit)]" in spec)
       
   285 apply(simp only: surjective_pairing[THEN sym])
       
   286 apply(force)
       
   287 apply(force)
       
   288 apply(force)
       
   289 apply(force)
       
   290 apply(force)
       
   291 apply(case_tac ba)
       
   292 apply(simp_all)
       
   293 apply(force)
       
   294 apply(case_tac "occurs list2 (Atom lista)")
       
   295 apply(simp)
       
   296 apply(drule_tac x="fst (apply_subst [(list2,swap (rev list1) (Atom lista))] (list,b))" in spec)
       
   297 apply(drule_tac x="snd (apply_subst [(list2,swap (rev list1) (Atom lista))] (list,b))" in spec)
       
   298 apply(drule_tac x="{}" in spec)
       
   299 apply(drule_tac x="[(list2, swap (rev list1) (Atom lista))]" in spec)
       
   300 apply(simp only: surjective_pairing[THEN sym])
       
   301 apply(force)
       
   302 apply(force)
       
   303 apply(force)
       
   304 apply(force)
       
   305 apply(force)
       
   306 apply(case_tac ba)
       
   307 apply(simp_all)
       
   308 apply(force)
       
   309 apply(case_tac "occurs list2 (Paar trm1 trm2)")
       
   310 apply(force)
       
   311 apply(drule_tac x="fst (apply_subst [(list2,swap (rev list1) (Paar trm1 trm2))] (list,b))" in spec)
       
   312 apply(drule_tac x="snd (apply_subst [(list2,swap (rev list1) (Paar trm1 trm2))] (list,b))" in spec)
       
   313 apply(drule_tac x="{}" in spec)
       
   314 apply(drule_tac x="[(list2, swap (rev list1) (Paar trm1 trm2))]" in spec)
       
   315 apply(simp only: surjective_pairing[THEN sym])
       
   316 apply(force)
       
   317 apply(force)
       
   318 apply(force)
       
   319 apply(force)
       
   320 apply(force)
       
   321 apply(case_tac ba)
       
   322 apply(simp_all)
       
   323 apply(force)
       
   324 apply(case_tac "occurs list2 (Func lista trm)")
       
   325 apply(force)
       
   326 apply(drule_tac x="fst (apply_subst [(list2,swap (rev list1) (Func lista trm))] (list,b))" in spec)
       
   327 apply(drule_tac x="snd (apply_subst [(list2,swap (rev list1) (Func lista trm))] (list,b))" in spec)
       
   328 apply(drule_tac x="{}" in spec)
       
   329 apply(drule_tac x="[(list2, swap (rev list1) (Func lista trm))]" in spec)
       
   330 apply(simp only: surjective_pairing[THEN sym])
       
   331 apply(force)
       
   332 apply(force)
       
   333 apply(force)
       
   334 apply(force)
       
   335 apply(force)
       
   336 apply(rule ballI)
       
   337 apply(thin_tac "([], []) \<in> stuck")
       
   338 apply(simp add: stuck_def)
       
   339 apply(clarify)
       
   340 apply(ind_cases "((a, b), (nabla, s), aa, ba) \<in> red_plus")
       
   341 apply(ind_cases "((a, b), s, aa, ba) \<in> s_red")
       
   342 apply(simp_all)
       
   343 apply(ind_cases "((Unit, Unit) # aa, b) \<in> fail")
       
   344 apply(ind_cases "((Paar t1 t2, Paar s1 s2) # xs, b) \<in> fail")
       
   345 apply(ind_cases "((Func F t1, Func F t2) # xs, b) \<in> fail")
       
   346 apply(simp)
       
   347 apply(ind_cases "((Abst ab t1, Abst ab t2) # xs, b) \<in> fail")
       
   348 apply(ind_cases "((Abst ab t1, Abst bb t2) # xs, b) \<in> fail")
       
   349 apply(ind_cases "((Atom ab, Atom ab) # aa, b) \<in> fail")
       
   350 apply(simp)
       
   351 apply(ind_cases "((Susp pi1 X, Susp pi2 X) # aa, b) \<in> fail")
       
   352 apply(ind_cases "((Susp pi X, t) # xs, b) \<in> fail")
       
   353 apply(simp add: apply_subst_def)+
       
   354 apply(case_tac "occurs X t1")
       
   355 apply(simp)
       
   356 apply(simp)
       
   357 apply(ind_cases "((t, Susp pi X) # xs, b) \<in> fail")
       
   358 apply(simp add: apply_subst_def)
       
   359 apply(case_tac "occurs X t1")
       
   360 apply(simp)
       
   361 apply(simp)
       
   362 apply(case_tac "occurs X t1")
       
   363 apply(simp)
       
   364 apply(simp)
       
   365 apply(ind_cases "((a, b), nabla, aa, ba) \<in> c_red")
       
   366 apply(simp_all)
       
   367 apply(ind_cases "([], (ab, Unit) # ba) \<in> fail")
       
   368 apply(ind_cases "([], (ab, Paar t1 t2) # xs) \<in> fail")
       
   369 apply(ind_cases "([], (ab, Func F t) # xs) \<in> fail")
       
   370 apply(ind_cases "([], (ab, Abst ab t) # ba) \<in> fail")
       
   371 apply(ind_cases "([], (ab, Abst bb t) # xs) \<in> fail")
       
   372 apply(ind_cases "([], (ab, Atom bb) # ba) \<in> fail")
       
   373 apply(simp)
       
   374 apply(ind_cases "([], (ab, Susp pi X) # ba) \<in> fail")
       
   375 apply(ind_cases "(a, b) \<turnstile> s1 \<leadsto> P2")
       
   376 apply(simp_all)
       
   377 apply(ind_cases "((Unit, Unit) # xs, b) \<in> fail")
       
   378 apply(ind_cases "((Paar t1 t2, Paar s1 s2) # xs, b) \<in> fail")
       
   379 apply(ind_cases "((Func F t1, Func F t2) # xs, b) \<in> fail")
       
   380 apply(simp)
       
   381 apply(ind_cases "((Abst ab t1, Abst ab t2) # xs, b) \<in> fail")
       
   382 apply(ind_cases "((Abst ab t1, Abst bb t2) # xs, b) \<in> fail")
       
   383 apply(ind_cases "((Atom ab, Atom ab) # aa, b) \<in> fail")
       
   384 apply(simp)
       
   385 apply(ind_cases "((Susp pi1 X, Susp pi2 X) # aa, b) \<in> fail")
       
   386 apply(ind_cases "((Susp pi X, t) # xs, b) \<in> fail")
       
   387 apply(simp add: apply_subst_def)+
       
   388 apply(case_tac "occurs X t1")
       
   389 apply(simp)
       
   390 apply(simp)
       
   391 apply(ind_cases "((t, Susp pi X) # xs, b) \<in> fail")
       
   392 apply(simp add: apply_subst_def)
       
   393 apply(case_tac "occurs X t1")
       
   394 apply(simp)
       
   395 apply(simp)
       
   396 apply(case_tac "occurs X t1")
       
   397 apply(simp)
       
   398 apply(simp)
       
   399 apply(ind_cases "(a, b) \<turnstile> nabla1 \<rightarrow> P2")
       
   400 apply(simp_all)
       
   401 apply(ind_cases "([], (ab, Unit) # ba) \<in> fail")
       
   402 apply(ind_cases "([], (ab, Paar t1 t2) # xs) \<in> fail")
       
   403 apply(ind_cases "([], (ab, Func F t) # xs) \<in> fail")
       
   404 apply(ind_cases "([], (ab, Abst ab t) # ba) \<in> fail")
       
   405 apply(ind_cases "([], (ab, Abst bb t) # xs) \<in> fail")
       
   406 apply(ind_cases "([], (ab, Atom bb) # ba) \<in> fail")
       
   407 apply(simp)
       
   408 apply(ind_cases "([], (ab, Susp pi X) # ba) \<in> fail")
       
   409 apply(simp add: stuck_def)
       
   410 apply(rule allI)+
       
   411 apply(clarify)
       
   412 apply(ind_cases "(([], []), (nabla, s), a, b) \<in> red_plus")
       
   413 apply(ind_cases "(([], []), s, a, b) \<in> s_red")
       
   414 apply(ind_cases "(([], []), nabla, a, b) \<in> c_red")
       
   415 apply(ind_cases "([], []) \<turnstile> s1 \<leadsto> P2")
       
   416 apply(ind_cases "([], []) \<turnstile> nabla1 \<rightarrow> P2")
       
   417 done
       
   418 
       
   419 lemma u_empty_sred: 
       
   420   "P1\<turnstile>s\<leadsto>P2 \<longrightarrow> U P2 ={} \<longrightarrow> U P1={}"
       
   421 apply(rule impI)
       
   422 apply(ind_cases "P1 \<turnstile> s \<leadsto> P2")
       
   423 apply(rule impI, simp add: all_solutions_def)
       
   424 apply(rule impI, simp add: all_solutions_def)
       
   425 apply(fast dest!: equ_paar_elim)
       
   426 apply(rule impI, simp add: all_solutions_def)
       
   427 apply(fast dest!: equ_func_elim)
       
   428 apply(rule impI, simp add: all_solutions_def)
       
   429 apply(fast dest!: equ_abst_aa_elim)
       
   430 apply(rule impI, simp add: all_solutions_def)
       
   431 apply(force dest!: equ_abst_ab_elim simp add: subst_swap_comm[THEN sym])
       
   432 apply(rule impI, simp add: all_solutions_def)
       
   433 apply(rule impI, simp add: all_solutions_def)
       
   434 apply(simp add: ds_list_equ_ds)
       
   435 apply(rule allI)+
       
   436 apply(rule impI)
       
   437 apply(drule_tac x="a" in spec)
       
   438 apply(drule_tac x="b" in spec)
       
   439 apply(erule disjE)
       
   440 apply(force)
       
   441 apply(simp add: subst_susp)
       
   442 apply(drule equ_pi1_pi2_dec[THEN mp])
       
   443 apply(force simp add: subst_susp)
       
   444 apply(auto)
       
   445 apply(simp add: all_solutions_def)
       
   446 apply(simp_all add: apply_subst_def)
       
   447 apply(auto)
       
   448 apply(drule_tac x="a" in spec)
       
   449 apply(drule_tac x="b" in spec)
       
   450 apply(drule unif_1)
       
   451 apply(auto)
       
   452 apply(drule_tac x="(aa,ba)" in bspec)
       
   453 apply(assumption)
       
   454 apply(simp)
       
   455 apply(drule_tac "t1.0"="aa" and "t2.0"="ba" in unif_2a)
       
   456 apply(simp add: subst_comp_expand)
       
   457 apply(drule_tac x="(aa,ba)" in bspec)
       
   458 apply(assumption)
       
   459 apply(simp)
       
   460 apply(drule_tac a="aa" and "t"="ba" in unif_2b)
       
   461 apply(simp add: subst_comp_expand)
       
   462 apply(simp add: all_solutions_def)
       
   463 apply(auto)
       
   464 apply(drule_tac x="a" in spec)
       
   465 apply(drule_tac x="b" in spec)
       
   466 apply(drule equ_sym)
       
   467 apply(drule unif_1)
       
   468 apply(auto)
       
   469 apply(drule_tac x="(aa,ba)" in bspec)
       
   470 apply(assumption)
       
   471 apply(simp)
       
   472 apply(drule_tac "t1.0"="aa" and "t2.0"="ba" in unif_2a)
       
   473 apply(simp add: subst_comp_expand)
       
   474 apply(drule_tac x="(aa,ba)" in bspec)
       
   475 apply(assumption)
       
   476 apply(simp)
       
   477 apply(drule_tac a="aa" and "t"="ba" in unif_2b)
       
   478 apply(simp add: subst_comp_expand)
       
   479 done
       
   480 
       
   481 lemma u_empty_cred: 
       
   482   "P1\<turnstile>nabla\<rightarrow>P2 \<longrightarrow> U P2 ={} \<longrightarrow> U P1={}"
       
   483 apply(rule impI)
       
   484 apply(ind_cases "P1 \<turnstile>nabla\<rightarrow>P2")
       
   485 apply(rule impI, simp add: all_solutions_def)
       
   486 apply(rule impI, simp add: all_solutions_def)
       
   487 apply(fast dest!: fresh_paar_elim)
       
   488 apply(rule impI, simp add: all_solutions_def)
       
   489 apply(fast dest!: fresh_func_elim)
       
   490 apply(rule impI, simp add: all_solutions_def)
       
   491 apply(rule impI, simp add: all_solutions_def)
       
   492 apply(force dest!: fresh_abst_ab_elim)
       
   493 apply(rule impI, simp add: all_solutions_def)
       
   494 apply(rule impI, simp add: all_solutions_def)
       
   495 done
       
   496 
       
   497 lemma u_empty_red_plus: 
       
   498   "P1\<Turnstile>(nabla,s)\<Rightarrow>P2 \<longrightarrow> U P2 ={} \<longrightarrow> U P1={}"
       
   499 apply(rule impI)
       
   500 apply(erule red_plus.induct)
       
   501 apply(drule u_empty_sred[THEN mp], assumption)
       
   502 apply(drule u_empty_cred[THEN mp], assumption)
       
   503 apply(drule u_empty_sred[THEN mp], force)
       
   504 apply(drule u_empty_cred[THEN mp], force)
       
   505 done
       
   506 
       
   507 (* all problems that cannot be solved produce "failed" problems only *)
       
   508 
       
   509 lemma empty_then_fail: "U P1={} \<longrightarrow> (\<forall>P\<in>results P1. P\<in>fail)"
       
   510 apply(simp add: results_def)
       
   511 apply(rule conjI)
       
   512 apply(rule impI)
       
   513 apply(rule impI)
       
   514 apply(simp add: stuck_equiv)
       
   515 apply(erule disjE)
       
   516 apply(subgoal_tac "({},[])\<in>U ([],[])")
       
   517 apply(simp)
       
   518 apply(simp add: all_solutions_def)
       
   519 apply(assumption)
       
   520 apply(rule impI)+
       
   521 apply(rule allI)+
       
   522 apply(rule impI)
       
   523 apply(erule conjE)
       
   524 apply(simp add: stuck_equiv)
       
   525 apply(auto)
       
   526 apply(subgoal_tac "({},[])\<in>U ([],[])")
       
   527 apply(drule_tac "nabla3.0"="nabla" and "nabla1.0"="{}" and "s1.0"="[]" in P1_from_P2_red_plus)
       
   528 apply(simp add: ext_subst_def)
       
   529 apply(auto)
       
   530 apply(simp add: all_solutions_def)
       
   531 done
       
   532 
       
   533 (* if a problem can be solved then no "failed" problem is produced *)
       
   534 
       
   535 lemma not_empty_then_not_fail: "U P1\<noteq>{} \<longrightarrow> \<not>(\<exists>P\<in>results P1. P\<in>fail)"
       
   536 apply(rule impI)
       
   537 apply(simp)
       
   538 apply(rule ballI)
       
   539 apply(clarify)
       
   540 apply(simp add: results_def)
       
   541 apply(case_tac "P1\<in>stuck")
       
   542 apply(simp_all)
       
   543 apply(drule fail_then_empty)
       
   544 apply(simp)
       
   545 apply(drule fail_then_empty)
       
   546 apply(erule conjE)
       
   547 apply(clarify)
       
   548 apply(drule u_empty_red_plus[THEN mp])
       
   549 apply(simp)
       
   550 done
       
   551 
       
   552 end
       
   553 
       
   554 
       
   555 
       
   556 
       
   557 
       
   558 
       
   559 
       
   560 
       
   561 
       
   562 
       
   563 
       
   564 
       
   565 
       
   566 
       
   567 
       
   568 
       
   569 
       
   570 
       
   571 
       
   572 
       
   573 
       
   574 
       
   575 
       
   576 
       
   577 
       
   578 
       
   579 
       
   580 
       
   581 
       
   582 
       
   583 
       
   584 
       
   585 
       
   586 
       
   587 
       
   588 
       
   589 
       
   590 
       
   591 
       
   592 
       
   593 
       
   594 
       
   595 
       
   596 
       
   597 
       
   598 
       
   599 
       
   600 
       
   601 
       
   602 
       
   603 
       
   604 
       
   605 
       
   606 
       
   607 
       
   608 
       
   609 
       
   610 
       
   611 
       
   612 
       
   613 
       
   614 
       
   615 
       
   616 
       
   617 
       
   618 
       
   619 
       
   620 
       
   621 
       
   622 
       
   623 
       
   624 
       
   625