src/HOL/ex/Unification.thy
author krauss
Sun, 21 Aug 2011 22:13:04 +0200
changeset 44370 03d91bfad83b
parent 44369 02e13192a053
child 44371 3a10392fb8c3
permissions -rw-r--r--
tuned proofs, sledgehammering overly verbose parts
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
41460
ea56b98aee83 removing obselete Id comments from HOL/ex theories
bulwahn
parents: 39754
diff changeset
     1
(*
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
     2
    Author:     Alexander Krauss, Technische Universitaet Muenchen
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
     3
*)
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
     4
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
     5
header {* Case study: Unification Algorithm *}
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
     6
23219
87ad6e8a5f2c tuned document;
wenzelm
parents: 23024
diff changeset
     7
theory Unification
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
     8
imports Main
23219
87ad6e8a5f2c tuned document;
wenzelm
parents: 23024
diff changeset
     9
begin
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
    10
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
    11
text {* 
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
    12
  This is a formalization of a first-order unification
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
    13
  algorithm. It uses the new "function" package to define recursive
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
    14
  functions, which allows a better treatment of nested recursion. 
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
    15
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
    16
  This is basically a modernized version of a previous formalization
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
    17
  by Konrad Slind (see: HOL/Subst/Unify.thy), which itself builds on
23219
87ad6e8a5f2c tuned document;
wenzelm
parents: 23024
diff changeset
    18
  previous work by Paulson and Manna \& Waldinger (for details, see
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
    19
  there).
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
    20
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
    21
  Unlike that formalization, where the proofs of termination and
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
    22
  some partial correctness properties are intertwined, we can prove
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
    23
  partial correctness and termination separately.
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
    24
*}
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
    25
23219
87ad6e8a5f2c tuned document;
wenzelm
parents: 23024
diff changeset
    26
44368
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
    27
subsection {* Terms *}
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
    28
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
    29
text {* Binary trees with leaves that are constants or variables. *}
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
    30
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
    31
datatype 'a trm = 
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
    32
  Var 'a 
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
    33
  | Const 'a
44367
74c08021ab2e changed constant names and notation to match HOL/Subst/*.thy, from which this theory is a clone.
krauss
parents: 42463
diff changeset
    34
  | Comb "'a trm" "'a trm" (infix "\<cdot>" 60)
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
    35
44368
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
    36
primrec vars_of :: "'a trm \<Rightarrow> 'a set"
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
    37
where
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
    38
  "vars_of (Var v) = {v}"
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
    39
| "vars_of (Const c) = {}"
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
    40
| "vars_of (M \<cdot> N) = vars_of M \<union> vars_of N"
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
    41
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
    42
fun occs :: "'a trm \<Rightarrow> 'a trm \<Rightarrow> bool" (infixl "\<prec>" 54) 
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
    43
where
44369
02e13192a053 tuned notation
krauss
parents: 44368
diff changeset
    44
  "u \<prec> Var v \<longleftrightarrow> False"
02e13192a053 tuned notation
krauss
parents: 44368
diff changeset
    45
| "u \<prec> Const c \<longleftrightarrow> False"
02e13192a053 tuned notation
krauss
parents: 44368
diff changeset
    46
| "u \<prec> M \<cdot> N \<longleftrightarrow> u = M \<or> u = N \<or> u \<prec> M \<or> u \<prec> N"
44368
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
    47
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
    48
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
    49
lemma finite_vars_of[intro]: "finite (vars_of t)"
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
    50
  by (induct t) simp_all
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
    51
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
    52
lemma vars_var_iff: "v \<in> vars_of (Var w) \<longleftrightarrow> w = v"
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
    53
  by auto
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
    54
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
    55
lemma vars_iff_occseq: "x \<in> vars_of t \<longleftrightarrow> Var x \<prec> t \<or> Var x = t"
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
    56
  by (induct t) auto
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
    57
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
    58
lemma occs_vars_subset: "M \<prec> N \<Longrightarrow> vars_of M \<subseteq> vars_of N"
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
    59
  by (induct N) auto
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
    60
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
    61
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
    62
subsection {* Substitutions *}
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
    63
42463
f270e3e18be5 modernized specifications;
wenzelm
parents: 41460
diff changeset
    64
type_synonym 'a subst = "('a \<times> 'a trm) list"
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
    65
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
    66
text {* Applying a substitution to a variable: *}
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
    67
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
    68
fun assoc :: "'a \<Rightarrow> 'b \<Rightarrow> ('a \<times> 'b) list \<Rightarrow> 'b"
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
    69
where
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
    70
  "assoc x d [] = d"
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
    71
| "assoc x d ((p,q)#t) = (if x = p then q else assoc x d t)"
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
    72
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
    73
text {* Applying a substitution to a term: *}
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
    74
44367
74c08021ab2e changed constant names and notation to match HOL/Subst/*.thy, from which this theory is a clone.
krauss
parents: 42463
diff changeset
    75
primrec subst :: "'a trm \<Rightarrow> 'a subst \<Rightarrow> 'a trm" (infixl "\<lhd>" 55)
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
    76
where
44367
74c08021ab2e changed constant names and notation to match HOL/Subst/*.thy, from which this theory is a clone.
krauss
parents: 42463
diff changeset
    77
  "(Var v) \<lhd> s = assoc v (Var v) s"
74c08021ab2e changed constant names and notation to match HOL/Subst/*.thy, from which this theory is a clone.
krauss
parents: 42463
diff changeset
    78
| "(Const c) \<lhd> s = (Const c)"
74c08021ab2e changed constant names and notation to match HOL/Subst/*.thy, from which this theory is a clone.
krauss
parents: 42463
diff changeset
    79
| "(M \<cdot> N) \<lhd> s = (M \<lhd> s) \<cdot> (N \<lhd> s)"
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
    80
44368
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
    81
definition subst_eq (infixr "\<doteq>" 52)
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
    82
where
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
    83
  "s1 \<doteq> s2 \<longleftrightarrow> (\<forall>t. t \<lhd> s1 = t \<lhd> s2)" 
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
    84
44368
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
    85
fun comp :: "'a subst \<Rightarrow> 'a subst \<Rightarrow> 'a subst" (infixl "\<lozenge>" 56)
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
    86
where
44367
74c08021ab2e changed constant names and notation to match HOL/Subst/*.thy, from which this theory is a clone.
krauss
parents: 42463
diff changeset
    87
  "[] \<lozenge> bl = bl"
74c08021ab2e changed constant names and notation to match HOL/Subst/*.thy, from which this theory is a clone.
krauss
parents: 42463
diff changeset
    88
| "((a,b) # al) \<lozenge> bl = (a, b \<lhd> bl) # (al \<lozenge> bl)"
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
    89
44368
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
    90
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
    91
subsection {* Basic Laws *}
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
    92
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
    93
lemma subst_Nil[simp]: "t \<lhd> [] = t"
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
    94
by (induct t) auto
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
    95
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
    96
lemma subst_mono: "t \<prec> u \<Longrightarrow> t \<lhd> s \<prec> u \<lhd> s"
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
    97
by (induct u) auto
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
    98
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
    99
lemma agreement: "(t \<lhd> r = t \<lhd> s) \<longleftrightarrow> (\<forall>v \<in> vars_of t. Var v \<lhd> r = Var v \<lhd> s)"
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
   100
by (induct t) auto
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
   101
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
   102
lemma repl_invariance: "v \<notin> vars_of t \<Longrightarrow> t \<lhd> (v,u) # s = t \<lhd> s"
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
   103
by (simp add: agreement)
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   104
44368
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
   105
lemma Var_in_subst:
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
   106
  "v \<in> vars_of t \<Longrightarrow> w \<in> vars_of (t \<lhd> (v, Var(w)) # s)"
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
   107
by (induct t) auto
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
   108
44370
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   109
lemma remove_var: "v \<notin> vars_of s \<Longrightarrow> v \<notin> vars_of (t \<lhd> [(v, s)])"
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   110
by (induct t) simp_all
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   111
44368
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
   112
lemma subst_refl[iff]: "s \<doteq> s"
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
   113
  by (auto simp:subst_eq_def)
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
   114
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
   115
lemma subst_sym[sym]: "\<lbrakk>s1 \<doteq> s2\<rbrakk> \<Longrightarrow> s2 \<doteq> s1"
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
   116
  by (auto simp:subst_eq_def)
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
   117
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
   118
lemma subst_trans[trans]: "\<lbrakk>s1 \<doteq> s2; s2 \<doteq> s3\<rbrakk> \<Longrightarrow> s1 \<doteq> s3"
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
   119
  by (auto simp:subst_eq_def)
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
   120
44370
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   121
44368
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
   122
text {* Composition of Substitutions *}
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   123
44368
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
   124
lemma comp_Nil[simp]: "\<sigma> \<lozenge> [] = \<sigma>"
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   125
by (induct \<sigma>) auto
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   126
44368
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
   127
lemma subst_comp[simp]: "t \<lhd> (r \<lozenge> s) = t \<lhd> r \<lhd> s"
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   128
proof (induct t)
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   129
  case (Var v) thus ?case
44368
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
   130
    by (induct r) auto
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
   131
qed auto
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   132
44368
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
   133
lemma subst_eq_intro[intro]: "(\<And>t. t \<lhd> \<sigma> = t \<lhd> \<theta>) \<Longrightarrow> \<sigma> \<doteq> \<theta>"
44367
74c08021ab2e changed constant names and notation to match HOL/Subst/*.thy, from which this theory is a clone.
krauss
parents: 42463
diff changeset
   134
  by (auto simp:subst_eq_def)
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   135
44368
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
   136
lemma subst_eq_dest[dest]: "s1 \<doteq> s2 \<Longrightarrow> t \<lhd> s1 = t \<lhd> s2"
44367
74c08021ab2e changed constant names and notation to match HOL/Subst/*.thy, from which this theory is a clone.
krauss
parents: 42463
diff changeset
   137
  by (auto simp:subst_eq_def)
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   138
44368
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
   139
lemma comp_assoc: "(a \<lozenge> b) \<lozenge> c \<doteq> a \<lozenge> (b \<lozenge> c)"
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
   140
  by auto
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   141
44368
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
   142
lemma subst_cong: "\<lbrakk>\<sigma> \<doteq> \<sigma>'; \<theta> \<doteq> \<theta>'\<rbrakk> \<Longrightarrow> (\<sigma> \<lozenge> \<theta>) \<doteq> (\<sigma>' \<lozenge> \<theta>')"
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
   143
  by (auto simp: subst_eq_def)
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   144
44370
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   145
lemma var_self: "[(v, Var v)] \<doteq> []"
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   146
proof
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   147
  fix t show "t \<lhd> [(v, Var v)] = t \<lhd> []"
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   148
    by (induct t) simp_all
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   149
qed
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   150
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   151
lemma var_same[simp]: "[(v, t)] \<doteq> [] \<longleftrightarrow> t = Var v"
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   152
by (metis assoc.simps(2) subst.simps(1) subst_eq_def var_self)
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   153
23219
87ad6e8a5f2c tuned document;
wenzelm
parents: 23024
diff changeset
   154
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   155
subsection {* Specification: Most general unifiers *}
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   156
44368
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
   157
definition Unifier :: "'a subst \<Rightarrow> 'a trm \<Rightarrow> 'a trm \<Rightarrow> bool"
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
   158
where "Unifier \<sigma> t u \<longleftrightarrow> (t \<lhd> \<sigma> = u \<lhd> \<sigma>)"
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   159
44368
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
   160
definition MGU :: "'a subst \<Rightarrow> 'a trm \<Rightarrow> 'a trm \<Rightarrow> bool" where
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
   161
  "MGU \<sigma> t u \<longleftrightarrow> 
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
   162
   Unifier \<sigma> t u \<and> (\<forall>\<theta>. Unifier \<theta> t u \<longrightarrow> (\<exists>\<gamma>. \<theta> \<doteq> \<sigma> \<lozenge> \<gamma>))"
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   163
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   164
lemma MGUI[intro]:
44367
74c08021ab2e changed constant names and notation to match HOL/Subst/*.thy, from which this theory is a clone.
krauss
parents: 42463
diff changeset
   165
  "\<lbrakk>t \<lhd> \<sigma> = u \<lhd> \<sigma>; \<And>\<theta>. t \<lhd> \<theta> = u \<lhd> \<theta> \<Longrightarrow> \<exists>\<gamma>. \<theta> \<doteq> \<sigma> \<lozenge> \<gamma>\<rbrakk>
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   166
  \<Longrightarrow> MGU \<sigma> t u"
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   167
  by (simp only:Unifier_def MGU_def, auto)
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   168
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   169
lemma MGU_sym[sym]:
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   170
  "MGU \<sigma> s t \<Longrightarrow> MGU \<sigma> t s"
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   171
  by (auto simp:MGU_def Unifier_def)
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   172
44370
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   173
lemma MGU_Var: 
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   174
  assumes "\<not> Var v \<prec> t"
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   175
  shows "MGU [(v,t)] (Var v) t"
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   176
proof (intro MGUI exI)
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   177
  show "Var v \<lhd> [(v,t)] = t \<lhd> [(v,t)]" using assms
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   178
    by (metis assoc.simps(2) repl_invariance subst.simps(1) subst_Nil vars_iff_occseq)
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   179
next
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   180
  fix \<theta> assume th: "Var v \<lhd> \<theta> = t \<lhd> \<theta>" 
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   181
  show "\<theta> \<doteq> [(v,t)] \<lozenge> \<theta>" 
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   182
  proof
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   183
    fix s show "s \<lhd> \<theta> = s \<lhd> [(v,t)] \<lozenge> \<theta>" using th 
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   184
      by (induct s) auto
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   185
  qed
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   186
qed
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   187
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   188
lemma MGU_Const: "MGU [] (Const c) (Const d) \<longleftrightarrow> c = d"
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   189
  by (auto simp: MGU_def Unifier_def)
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   190
  
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   191
44368
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
   192
definition Idem :: "'a subst \<Rightarrow> bool"
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
   193
where "Idem s \<longleftrightarrow> (s \<lozenge> s) \<doteq> s"
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
   194
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
   195
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
   196
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   197
subsection {* The unification algorithm *}
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   198
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   199
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   200
text {* The unification algorithm: *}
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   201
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   202
function unify :: "'a trm \<Rightarrow> 'a trm \<Rightarrow> 'a subst option"
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   203
where
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   204
  "unify (Const c) (M \<cdot> N)   = None"
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   205
| "unify (M \<cdot> N)   (Const c) = None"
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   206
| "unify (Const c) (Var v)   = Some [(v, Const c)]"
44369
02e13192a053 tuned notation
krauss
parents: 44368
diff changeset
   207
| "unify (M \<cdot> N)   (Var v)   = (if Var v \<prec> M \<cdot> N 
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   208
                                        then None
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   209
                                        else Some [(v, M \<cdot> N)])"
44369
02e13192a053 tuned notation
krauss
parents: 44368
diff changeset
   210
| "unify (Var v)   M         = (if Var v \<prec> M
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   211
                                        then None
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   212
                                        else Some [(v, M)])"
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   213
| "unify (Const c) (Const d) = (if c=d then Some [] else None)"
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   214
| "unify (M \<cdot> N) (M' \<cdot> N') = (case unify M M' of
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   215
                                    None \<Rightarrow> None |
44367
74c08021ab2e changed constant names and notation to match HOL/Subst/*.thy, from which this theory is a clone.
krauss
parents: 42463
diff changeset
   216
                                    Some \<theta> \<Rightarrow> (case unify (N \<lhd> \<theta>) (N' \<lhd> \<theta>)
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   217
                                      of None \<Rightarrow> None |
44367
74c08021ab2e changed constant names and notation to match HOL/Subst/*.thy, from which this theory is a clone.
krauss
parents: 42463
diff changeset
   218
                                         Some \<sigma> \<Rightarrow> Some (\<theta> \<lozenge> \<sigma>)))"
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   219
  by pat_completeness auto
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   220
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   221
subsection {* Properties used in termination proof *}
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   222
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   223
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   224
text {* Elimination of variables by a substitution: *}
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   225
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   226
definition
44367
74c08021ab2e changed constant names and notation to match HOL/Subst/*.thy, from which this theory is a clone.
krauss
parents: 42463
diff changeset
   227
  "elim \<sigma> v \<equiv> \<forall>t. v \<notin> vars_of (t \<lhd> \<sigma>)"
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   228
44367
74c08021ab2e changed constant names and notation to match HOL/Subst/*.thy, from which this theory is a clone.
krauss
parents: 42463
diff changeset
   229
lemma elim_intro[intro]: "(\<And>t. v \<notin> vars_of (t \<lhd> \<sigma>)) \<Longrightarrow> elim \<sigma> v"
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   230
  by (auto simp:elim_def)
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   231
44367
74c08021ab2e changed constant names and notation to match HOL/Subst/*.thy, from which this theory is a clone.
krauss
parents: 42463
diff changeset
   232
lemma elim_dest[dest]: "elim \<sigma> v \<Longrightarrow> v \<notin> vars_of (t \<lhd> \<sigma>)"
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   233
  by (auto simp:elim_def)
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   234
44370
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   235
lemma elim_eq: "\<sigma> \<doteq> \<theta> \<Longrightarrow> elim \<sigma> x = elim \<theta> x"
44367
74c08021ab2e changed constant names and notation to match HOL/Subst/*.thy, from which this theory is a clone.
krauss
parents: 42463
diff changeset
   236
  by (auto simp:elim_def subst_eq_def)
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   237
44369
02e13192a053 tuned notation
krauss
parents: 44368
diff changeset
   238
lemma occs_elim: "\<not> Var v \<prec> t 
44367
74c08021ab2e changed constant names and notation to match HOL/Subst/*.thy, from which this theory is a clone.
krauss
parents: 42463
diff changeset
   239
  \<Longrightarrow> elim [(v,t)] v \<or> [(v,t)] \<doteq> []"
44370
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   240
by (metis elim_intro remove_var var_same vars_iff_occseq)
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   241
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   242
text {* The result of a unification never introduces new variables: *}
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   243
44370
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   244
declare unify.psimps[simp]
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   245
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   246
lemma unify_vars: 
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   247
  assumes "unify_dom (M, N)"
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   248
  assumes "unify M N = Some \<sigma>"
44367
74c08021ab2e changed constant names and notation to match HOL/Subst/*.thy, from which this theory is a clone.
krauss
parents: 42463
diff changeset
   249
  shows "vars_of (t \<lhd> \<sigma>) \<subseteq> vars_of M \<union> vars_of N \<union> vars_of t"
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   250
  (is "?P M N \<sigma> t")
24444
448978b61556 induct: proper separation of initial and terminal step;
wenzelm
parents: 23777
diff changeset
   251
using assms
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   252
proof (induct M N arbitrary:\<sigma> t)
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   253
  case (3 c v) 
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   254
  hence "\<sigma> = [(v, Const c)]" by simp
24444
448978b61556 induct: proper separation of initial and terminal step;
wenzelm
parents: 23777
diff changeset
   255
  thus ?case by (induct t) auto
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   256
next
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   257
  case (4 M N v) 
44369
02e13192a053 tuned notation
krauss
parents: 44368
diff changeset
   258
  hence "\<not> Var v \<prec> M \<cdot> N" by auto
24444
448978b61556 induct: proper separation of initial and terminal step;
wenzelm
parents: 23777
diff changeset
   259
  with 4 have "\<sigma> = [(v, M\<cdot>N)]" by simp
448978b61556 induct: proper separation of initial and terminal step;
wenzelm
parents: 23777
diff changeset
   260
  thus ?case by (induct t) auto
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   261
next
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   262
  case (5 v M)
44369
02e13192a053 tuned notation
krauss
parents: 44368
diff changeset
   263
  hence "\<not> Var v \<prec> M" by auto
24444
448978b61556 induct: proper separation of initial and terminal step;
wenzelm
parents: 23777
diff changeset
   264
  with 5 have "\<sigma> = [(v, M)]" by simp
448978b61556 induct: proper separation of initial and terminal step;
wenzelm
parents: 23777
diff changeset
   265
  thus ?case by (induct t) auto
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   266
next
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   267
  case (7 M N M' N' \<sigma>)
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   268
  then obtain \<theta>1 \<theta>2 
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   269
    where "unify M M' = Some \<theta>1"
44367
74c08021ab2e changed constant names and notation to match HOL/Subst/*.thy, from which this theory is a clone.
krauss
parents: 42463
diff changeset
   270
    and "unify (N \<lhd> \<theta>1) (N' \<lhd> \<theta>1) = Some \<theta>2"
74c08021ab2e changed constant names and notation to match HOL/Subst/*.thy, from which this theory is a clone.
krauss
parents: 42463
diff changeset
   271
    and \<sigma>: "\<sigma> = \<theta>1 \<lozenge> \<theta>2"
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   272
    and ih1: "\<And>t. ?P M M' \<theta>1 t"
44367
74c08021ab2e changed constant names and notation to match HOL/Subst/*.thy, from which this theory is a clone.
krauss
parents: 42463
diff changeset
   273
    and ih2: "\<And>t. ?P (N\<lhd>\<theta>1) (N'\<lhd>\<theta>1) \<theta>2 t"
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   274
    by (auto split:option.split_asm)
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   275
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   276
  show ?case
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   277
  proof
44367
74c08021ab2e changed constant names and notation to match HOL/Subst/*.thy, from which this theory is a clone.
krauss
parents: 42463
diff changeset
   278
    fix v assume a: "v \<in> vars_of (t \<lhd> \<sigma>)"
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   279
    
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   280
    show "v \<in> vars_of (M \<cdot> N) \<union> vars_of (M' \<cdot> N') \<union> vars_of t"
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   281
    proof (cases "v \<notin> vars_of M \<and> v \<notin> vars_of M'
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30909
diff changeset
   282
        \<and> v \<notin> vars_of N \<and> v \<notin> vars_of N'")
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   283
      case True
44367
74c08021ab2e changed constant names and notation to match HOL/Subst/*.thy, from which this theory is a clone.
krauss
parents: 42463
diff changeset
   284
      with ih1 have l:"\<And>t. v \<in> vars_of (t \<lhd> \<theta>1) \<Longrightarrow> v \<in> vars_of t"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30909
diff changeset
   285
        by auto
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   286
      
44367
74c08021ab2e changed constant names and notation to match HOL/Subst/*.thy, from which this theory is a clone.
krauss
parents: 42463
diff changeset
   287
      from a and ih2[where t="t \<lhd> \<theta>1"]
74c08021ab2e changed constant names and notation to match HOL/Subst/*.thy, from which this theory is a clone.
krauss
parents: 42463
diff changeset
   288
      have "v \<in> vars_of (N \<lhd> \<theta>1) \<union> vars_of (N' \<lhd> \<theta>1) 
74c08021ab2e changed constant names and notation to match HOL/Subst/*.thy, from which this theory is a clone.
krauss
parents: 42463
diff changeset
   289
        \<or> v \<in> vars_of (t \<lhd> \<theta>1)" unfolding \<sigma>
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30909
diff changeset
   290
        by auto
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   291
      hence "v \<in> vars_of t"
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   292
      proof
44367
74c08021ab2e changed constant names and notation to match HOL/Subst/*.thy, from which this theory is a clone.
krauss
parents: 42463
diff changeset
   293
        assume "v \<in> vars_of (N \<lhd> \<theta>1) \<union> vars_of (N' \<lhd> \<theta>1)"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30909
diff changeset
   294
        with True show ?thesis by (auto dest:l)
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   295
      next
44367
74c08021ab2e changed constant names and notation to match HOL/Subst/*.thy, from which this theory is a clone.
krauss
parents: 42463
diff changeset
   296
        assume "v \<in> vars_of (t \<lhd> \<theta>1)" 
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30909
diff changeset
   297
        thus ?thesis by (rule l)
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   298
      qed
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   299
      
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   300
      thus ?thesis by auto
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   301
    qed auto
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   302
  qed
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   303
qed (auto split: split_if_asm)
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   304
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   305
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   306
text {* The result of a unification is either the identity
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   307
substitution or it eliminates a variable from one of the terms: *}
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   308
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   309
lemma unify_eliminates: 
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   310
  assumes "unify_dom (M, N)"
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   311
  assumes "unify M N = Some \<sigma>"
44367
74c08021ab2e changed constant names and notation to match HOL/Subst/*.thy, from which this theory is a clone.
krauss
parents: 42463
diff changeset
   312
  shows "(\<exists>v\<in>vars_of M \<union> vars_of N. elim \<sigma> v) \<or> \<sigma> \<doteq> []"
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   313
  (is "?P M N \<sigma>")
24444
448978b61556 induct: proper separation of initial and terminal step;
wenzelm
parents: 23777
diff changeset
   314
using assms
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   315
proof (induct M N arbitrary:\<sigma>)
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   316
  case 1 thus ?case by simp
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   317
next
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   318
  case 2 thus ?case by simp
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   319
next
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   320
  case (3 c v)
44369
02e13192a053 tuned notation
krauss
parents: 44368
diff changeset
   321
  have no_occs: "\<not> Var v \<prec> Const c" by simp
24444
448978b61556 induct: proper separation of initial and terminal step;
wenzelm
parents: 23777
diff changeset
   322
  with 3 have "\<sigma> = [(v, Const c)]" by simp
44367
74c08021ab2e changed constant names and notation to match HOL/Subst/*.thy, from which this theory is a clone.
krauss
parents: 42463
diff changeset
   323
  with occs_elim[OF no_occs]
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   324
  show ?case by auto
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   325
next
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   326
  case (4 M N v)
44369
02e13192a053 tuned notation
krauss
parents: 44368
diff changeset
   327
  hence no_occs: "\<not> Var v \<prec> M \<cdot> N" by auto
24444
448978b61556 induct: proper separation of initial and terminal step;
wenzelm
parents: 23777
diff changeset
   328
  with 4 have "\<sigma> = [(v, M\<cdot>N)]" by simp
44367
74c08021ab2e changed constant names and notation to match HOL/Subst/*.thy, from which this theory is a clone.
krauss
parents: 42463
diff changeset
   329
  with occs_elim[OF no_occs]
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   330
  show ?case by auto 
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   331
next
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   332
  case (5 v M) 
44369
02e13192a053 tuned notation
krauss
parents: 44368
diff changeset
   333
  hence no_occs: "\<not> Var v \<prec> M" by auto
24444
448978b61556 induct: proper separation of initial and terminal step;
wenzelm
parents: 23777
diff changeset
   334
  with 5 have "\<sigma> = [(v, M)]" by simp
44367
74c08021ab2e changed constant names and notation to match HOL/Subst/*.thy, from which this theory is a clone.
krauss
parents: 42463
diff changeset
   335
  with occs_elim[OF no_occs]
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   336
  show ?case by auto 
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   337
next 
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   338
  case (6 c d) thus ?case
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   339
    by (cases "c = d") auto
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   340
next
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   341
  case (7 M N M' N' \<sigma>)
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   342
  then obtain \<theta>1 \<theta>2 
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   343
    where "unify M M' = Some \<theta>1"
44367
74c08021ab2e changed constant names and notation to match HOL/Subst/*.thy, from which this theory is a clone.
krauss
parents: 42463
diff changeset
   344
    and "unify (N \<lhd> \<theta>1) (N' \<lhd> \<theta>1) = Some \<theta>2"
74c08021ab2e changed constant names and notation to match HOL/Subst/*.thy, from which this theory is a clone.
krauss
parents: 42463
diff changeset
   345
    and \<sigma>: "\<sigma> = \<theta>1 \<lozenge> \<theta>2"
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   346
    and ih1: "?P M M' \<theta>1"
44367
74c08021ab2e changed constant names and notation to match HOL/Subst/*.thy, from which this theory is a clone.
krauss
parents: 42463
diff changeset
   347
    and ih2: "?P (N\<lhd>\<theta>1) (N'\<lhd>\<theta>1) \<theta>2"
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   348
    by (auto split:option.split_asm)
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   349
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   350
  from `unify_dom (M \<cdot> N, M' \<cdot> N')`
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   351
  have "unify_dom (M, M')"
23777
60b7800338d5 Renamed accessible part for predicates to accp.
berghofe
parents: 23373
diff changeset
   352
    by (rule accp_downward) (rule unify_rel.intros)
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   353
  hence no_new_vars: 
44367
74c08021ab2e changed constant names and notation to match HOL/Subst/*.thy, from which this theory is a clone.
krauss
parents: 42463
diff changeset
   354
    "\<And>t. vars_of (t \<lhd> \<theta>1) \<subseteq> vars_of M \<union> vars_of M' \<union> vars_of t"
23373
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 23219
diff changeset
   355
    by (rule unify_vars) (rule `unify M M' = Some \<theta>1`)
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   356
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   357
  from ih2 show ?case 
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   358
  proof 
44367
74c08021ab2e changed constant names and notation to match HOL/Subst/*.thy, from which this theory is a clone.
krauss
parents: 42463
diff changeset
   359
    assume "\<exists>v\<in>vars_of (N \<lhd> \<theta>1) \<union> vars_of (N' \<lhd> \<theta>1). elim \<theta>2 v"
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   360
    then obtain v 
44367
74c08021ab2e changed constant names and notation to match HOL/Subst/*.thy, from which this theory is a clone.
krauss
parents: 42463
diff changeset
   361
      where "v\<in>vars_of (N \<lhd> \<theta>1) \<union> vars_of (N' \<lhd> \<theta>1)"
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   362
      and el: "elim \<theta>2 v" by auto
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   363
    with no_new_vars show ?thesis unfolding \<sigma> 
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   364
      by (auto simp:elim_def)
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   365
  next
44367
74c08021ab2e changed constant names and notation to match HOL/Subst/*.thy, from which this theory is a clone.
krauss
parents: 42463
diff changeset
   366
    assume empty[simp]: "\<theta>2 \<doteq> []"
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   367
44367
74c08021ab2e changed constant names and notation to match HOL/Subst/*.thy, from which this theory is a clone.
krauss
parents: 42463
diff changeset
   368
    have "\<sigma> \<doteq> (\<theta>1 \<lozenge> [])" unfolding \<sigma>
44368
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
   369
      by (rule subst_cong) auto
44367
74c08021ab2e changed constant names and notation to match HOL/Subst/*.thy, from which this theory is a clone.
krauss
parents: 42463
diff changeset
   370
    also have "\<dots> \<doteq> \<theta>1" by auto
74c08021ab2e changed constant names and notation to match HOL/Subst/*.thy, from which this theory is a clone.
krauss
parents: 42463
diff changeset
   371
    finally have "\<sigma> \<doteq> \<theta>1" .
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   372
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   373
    from ih1 show ?thesis
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   374
    proof
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   375
      assume "\<exists>v\<in>vars_of M \<union> vars_of M'. elim \<theta>1 v"
44370
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   376
      with elim_eq[OF `\<sigma> \<doteq> \<theta>1`]
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   377
      show ?thesis by auto
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   378
    next
44367
74c08021ab2e changed constant names and notation to match HOL/Subst/*.thy, from which this theory is a clone.
krauss
parents: 42463
diff changeset
   379
      note `\<sigma> \<doteq> \<theta>1`
74c08021ab2e changed constant names and notation to match HOL/Subst/*.thy, from which this theory is a clone.
krauss
parents: 42463
diff changeset
   380
      also assume "\<theta>1 \<doteq> []"
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   381
      finally show ?thesis ..
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   382
    qed
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   383
  qed
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   384
qed
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   385
44370
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   386
declare unify.psimps[simp del]
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   387
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   388
subsection {* Termination proof *}
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   389
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   390
termination unify
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   391
proof 
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   392
  let ?R = "measures [\<lambda>(M,N). card (vars_of M \<union> vars_of N),
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   393
                           \<lambda>(M, N). size M]"
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   394
  show "wf ?R" by simp
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   395
44370
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   396
  fix M N M' N' :: "'a trm"
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   397
  show "((M, M'), (M \<cdot> N, M' \<cdot> N')) \<in> ?R" -- "Inner call"
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   398
    by (rule measures_lesseq) (auto intro: card_mono)
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   399
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   400
  fix \<theta>                                   -- "Outer call"
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   401
  assume inner: "unify_dom (M, M')"
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   402
    "unify M M' = Some \<theta>"
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   403
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   404
  from unify_eliminates[OF inner]
44367
74c08021ab2e changed constant names and notation to match HOL/Subst/*.thy, from which this theory is a clone.
krauss
parents: 42463
diff changeset
   405
  show "((N \<lhd> \<theta>, N' \<lhd> \<theta>), (M \<cdot> N, M' \<cdot> N')) \<in>?R"
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   406
  proof
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   407
    -- {* Either a variable is eliminated \ldots *}
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   408
    assume "(\<exists>v\<in>vars_of M \<union> vars_of M'. elim \<theta> v)"
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   409
    then obtain v 
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30909
diff changeset
   410
      where "elim \<theta> v" 
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30909
diff changeset
   411
      and "v\<in>vars_of M \<union> vars_of M'" by auto
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   412
    with unify_vars[OF inner]
44367
74c08021ab2e changed constant names and notation to match HOL/Subst/*.thy, from which this theory is a clone.
krauss
parents: 42463
diff changeset
   413
    have "vars_of (N\<lhd>\<theta>) \<union> vars_of (N'\<lhd>\<theta>)
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30909
diff changeset
   414
      \<subset> vars_of (M\<cdot>N) \<union> vars_of (M'\<cdot>N')"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30909
diff changeset
   415
      by auto
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   416
    
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   417
    thus ?thesis
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   418
      by (auto intro!: measures_less intro: psubset_card_mono)
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   419
  next
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   420
    -- {* Or the substitution is empty *}
44367
74c08021ab2e changed constant names and notation to match HOL/Subst/*.thy, from which this theory is a clone.
krauss
parents: 42463
diff changeset
   421
    assume "\<theta> \<doteq> []"
74c08021ab2e changed constant names and notation to match HOL/Subst/*.thy, from which this theory is a clone.
krauss
parents: 42463
diff changeset
   422
    hence "N \<lhd> \<theta> = N" 
74c08021ab2e changed constant names and notation to match HOL/Subst/*.thy, from which this theory is a clone.
krauss
parents: 42463
diff changeset
   423
      and "N' \<lhd> \<theta> = N'" by auto
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   424
    thus ?thesis 
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   425
       by (auto intro!: measures_less intro: psubset_card_mono)
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   426
  qed
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   427
qed
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   428
44370
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   429
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   430
subsection {* Other Properties *}
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   431
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   432
lemma unify_computes_MGU:
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   433
  "unify M N = Some \<sigma> \<Longrightarrow> MGU \<sigma> M N"
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   434
using assms
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   435
proof (induct M N arbitrary: \<sigma> rule: unify.induct)
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   436
  case (7 M N M' N' \<sigma>) -- "The interesting case"
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   437
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   438
  then obtain \<theta>1 \<theta>2 
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   439
    where "unify M M' = Some \<theta>1"
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   440
    and "unify (N \<lhd> \<theta>1) (N' \<lhd> \<theta>1) = Some \<theta>2"
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   441
    and \<sigma>: "\<sigma> = \<theta>1 \<lozenge> \<theta>2"
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   442
    and MGU_inner: "MGU \<theta>1 M M'" 
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   443
    and MGU_outer: "MGU \<theta>2 (N \<lhd> \<theta>1) (N' \<lhd> \<theta>1)"
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   444
    by (auto split:option.split_asm)
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   445
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   446
  show ?case
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   447
  proof
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   448
    from MGU_inner and MGU_outer
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   449
    have "M \<lhd> \<theta>1 = M' \<lhd> \<theta>1" 
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   450
      and "N \<lhd> \<theta>1 \<lhd> \<theta>2 = N' \<lhd> \<theta>1 \<lhd> \<theta>2"
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   451
      unfolding MGU_def Unifier_def
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   452
      by auto
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   453
    thus "M \<cdot> N \<lhd> \<sigma> = M' \<cdot> N' \<lhd> \<sigma>" unfolding \<sigma>
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   454
      by simp
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   455
  next
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   456
    fix \<sigma>' assume "M \<cdot> N \<lhd> \<sigma>' = M' \<cdot> N' \<lhd> \<sigma>'"
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   457
    hence "M \<lhd> \<sigma>' = M' \<lhd> \<sigma>'"
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   458
      and Ns: "N \<lhd> \<sigma>' = N' \<lhd> \<sigma>'" by auto
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   459
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   460
    with MGU_inner obtain \<delta>
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   461
      where eqv: "\<sigma>' \<doteq> \<theta>1 \<lozenge> \<delta>"
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   462
      unfolding MGU_def Unifier_def
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   463
      by auto
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   464
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   465
    from Ns have "N \<lhd> \<theta>1 \<lhd> \<delta> = N' \<lhd> \<theta>1 \<lhd> \<delta>"
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   466
      by (simp add:subst_eq_dest[OF eqv])
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   467
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   468
    with MGU_outer obtain \<rho>
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   469
      where eqv2: "\<delta> \<doteq> \<theta>2 \<lozenge> \<rho>"
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   470
      unfolding MGU_def Unifier_def
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   471
      by auto
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   472
    
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   473
    have "\<sigma>' \<doteq> \<sigma> \<lozenge> \<rho>" unfolding \<sigma>
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   474
      by (rule subst_eq_intro, auto simp:subst_eq_dest[OF eqv] subst_eq_dest[OF eqv2])
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   475
    thus "\<exists>\<gamma>. \<sigma>' \<doteq> \<sigma> \<lozenge> \<gamma>" ..
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   476
  qed
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   477
qed (auto simp: MGU_Const intro: MGU_Var MGU_Var[symmetric] split: split_if_asm)
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   478
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   479
39754
150f831ce4a3 no longer declare .psimps rules as [simp].
krauss
parents: 32960
diff changeset
   480
23219
87ad6e8a5f2c tuned document;
wenzelm
parents: 23024
diff changeset
   481
end