src/HOL/ex/Unification.thy
author wenzelm
Fri, 29 Nov 2024 17:40:15 +0100
changeset 81507 08574da77b4a
parent 80914 d97fdabd9e2b
permissions -rw-r--r--
clarified signature: shorten common cases;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
44372
f9825056dbab more precise authors and comments;
krauss
parents: 44371
diff changeset
     1
(*  Title:      HOL/ex/Unification.thy
f9825056dbab more precise authors and comments;
krauss
parents: 44371
diff changeset
     2
    Author:     Martin Coen, Cambridge University Computer Laboratory
f9825056dbab more precise authors and comments;
krauss
parents: 44371
diff changeset
     3
    Author:     Konrad Slind, TUM & Cambridge University Computer Laboratory
f9825056dbab more precise authors and comments;
krauss
parents: 44371
diff changeset
     4
    Author:     Alexander Krauss, TUM
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
     5
*)
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
     6
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
     7
section \<open>Substitution and Unification\<close>
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
     8
23219
87ad6e8a5f2c tuned document;
wenzelm
parents: 23024
diff changeset
     9
theory Unification
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
    10
imports Main
23219
87ad6e8a5f2c tuned document;
wenzelm
parents: 23024
diff changeset
    11
begin
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
    12
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
    13
text \<open>
44428
ccb8998f70b7 fixed document;
wenzelm
parents: 44373
diff changeset
    14
  Implements Manna \& Waldinger's formalization, with Paulson's
44372
f9825056dbab more precise authors and comments;
krauss
parents: 44371
diff changeset
    15
  simplifications, and some new simplifications by Slind and Krauss.
f9825056dbab more precise authors and comments;
krauss
parents: 44371
diff changeset
    16
44428
ccb8998f70b7 fixed document;
wenzelm
parents: 44373
diff changeset
    17
  Z Manna \& R Waldinger, Deductive Synthesis of the Unification
44372
f9825056dbab more precise authors and comments;
krauss
parents: 44371
diff changeset
    18
  Algorithm.  SCP 1 (1981), 5-48
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
    19
44372
f9825056dbab more precise authors and comments;
krauss
parents: 44371
diff changeset
    20
  L C Paulson, Verifying the Unification Algorithm in LCF. SCP 5
f9825056dbab more precise authors and comments;
krauss
parents: 44371
diff changeset
    21
  (1985), 143-170
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
    22
44372
f9825056dbab more precise authors and comments;
krauss
parents: 44371
diff changeset
    23
  K Slind, Reasoning about Terminating Functional Programs,
f9825056dbab more precise authors and comments;
krauss
parents: 44371
diff changeset
    24
  Ph.D. thesis, TUM, 1999, Sect. 5.8
f9825056dbab more precise authors and comments;
krauss
parents: 44371
diff changeset
    25
f9825056dbab more precise authors and comments;
krauss
parents: 44371
diff changeset
    26
  A Krauss, Partial and Nested Recursive Function Definitions in
56790
f54097170704 prefer plain ASCII / latex over not-so-universal Unicode;
wenzelm
parents: 44428
diff changeset
    27
  Higher-Order Logic, JAR 44(4):303-336, 2010. Sect. 6.3
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
    28
\<close>
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
    29
23219
87ad6e8a5f2c tuned document;
wenzelm
parents: 23024
diff changeset
    30
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
    31
subsection \<open>Terms\<close>
44368
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
    32
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
    33
text \<open>Binary trees with leaves that are constants or variables.\<close>
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
    34
58310
91ea607a34d8 updated news
blanchet
parents: 58249
diff changeset
    35
datatype 'a trm = 
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
    36
  Var 'a 
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
    37
  | Const 'a
80914
d97fdabd9e2b standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents: 79558
diff changeset
    38
  | Comb "'a trm" "'a trm" (infix \<open>\<cdot>\<close> 60)
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
    39
44368
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
    40
primrec vars_of :: "'a trm \<Rightarrow> 'a set"
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
    41
where
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
    42
  "vars_of (Var v) = {v}"
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
    43
| "vars_of (Const c) = {}"
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
    44
| "vars_of (M \<cdot> N) = vars_of M \<union> vars_of N"
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
    45
80914
d97fdabd9e2b standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents: 79558
diff changeset
    46
fun occs :: "'a trm \<Rightarrow> 'a trm \<Rightarrow> bool" (infixl \<open>\<prec>\<close> 54) 
44368
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
    47
where
44369
02e13192a053 tuned notation
krauss
parents: 44368
diff changeset
    48
  "u \<prec> Var v \<longleftrightarrow> False"
02e13192a053 tuned notation
krauss
parents: 44368
diff changeset
    49
| "u \<prec> Const c \<longleftrightarrow> False"
02e13192a053 tuned notation
krauss
parents: 44368
diff changeset
    50
| "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
    51
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
    52
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
    53
lemma finite_vars_of[intro]: "finite (vars_of t)"
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
    54
  by (induct t) simp_all
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
    55
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
    56
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
    57
  by (induct t) auto
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
    58
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
    59
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
    60
  by (induct N) auto
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
    61
76643
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
    62
lemma size_less_size_if_occs: "t \<prec> u \<Longrightarrow> size t < size u"
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
    63
proof (induction u arbitrary: t)
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
    64
  case (Comb u1 u2)
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
    65
  thus ?case by fastforce
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
    66
qed simp_all
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
    67
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
    68
corollary neq_if_occs: "t \<prec> u \<Longrightarrow> t \<noteq> u"
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
    69
  using size_less_size_if_occs by auto
44368
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
    70
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
    71
subsection \<open>Substitutions\<close>
44368
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
    72
42463
f270e3e18be5 modernized specifications;
wenzelm
parents: 41460
diff changeset
    73
type_synonym 'a subst = "('a \<times> 'a trm) list"
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
    74
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
    75
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
    76
where
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
    77
  "assoc x d [] = d"
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
    78
| "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
    79
80914
d97fdabd9e2b standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents: 79558
diff changeset
    80
primrec subst :: "'a trm \<Rightarrow> 'a subst \<Rightarrow> 'a trm" (infixl \<open>\<lhd>\<close> 55)
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
    81
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
    82
  "(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
    83
| "(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
    84
| "(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
    85
80914
d97fdabd9e2b standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents: 79558
diff changeset
    86
definition subst_eq (infixr \<open>\<doteq>\<close> 52)
44368
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
    87
where
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
    88
  "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
    89
80914
d97fdabd9e2b standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents: 79558
diff changeset
    90
fun comp :: "'a subst \<Rightarrow> 'a subst \<Rightarrow> 'a subst" (infixl \<open>\<lozenge>\<close> 56)
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
    91
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
    92
  "[] \<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
    93
| "((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
    94
44368
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
    95
lemma subst_Nil[simp]: "t \<lhd> [] = t"
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
    96
by (induct t) auto
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
    97
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
    98
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
    99
by (induct u) auto
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
   100
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
   101
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
   102
by (induct t) auto
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
   103
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
   104
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
   105
by (simp add: agreement)
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   106
44370
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   107
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
   108
by (induct t) simp_all
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   109
44368
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
   110
lemma subst_refl[iff]: "s \<doteq> s"
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
   111
  by (auto simp:subst_eq_def)
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
   112
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
   113
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
   114
  by (auto simp:subst_eq_def)
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
   115
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
   116
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
   117
  by (auto simp:subst_eq_def)
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
   118
44371
3a10392fb8c3 added proof of idempotence, roughly after HOL/Subst/Unify.thy
krauss
parents: 44370
diff changeset
   119
lemma subst_no_occs: "\<not> Var v \<prec> t \<Longrightarrow> Var v \<noteq> t
3a10392fb8c3 added proof of idempotence, roughly after HOL/Subst/Unify.thy
krauss
parents: 44370
diff changeset
   120
  \<Longrightarrow> t \<lhd> [(v,s)] = t"
3a10392fb8c3 added proof of idempotence, roughly after HOL/Subst/Unify.thy
krauss
parents: 44370
diff changeset
   121
by (induct t) auto
44370
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   122
44368
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
   123
lemma comp_Nil[simp]: "\<sigma> \<lozenge> [] = \<sigma>"
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   124
by (induct \<sigma>) auto
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   125
44368
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
   126
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
   127
proof (induct t)
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   128
  case (Var v) thus ?case
44368
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
   129
    by (induct r) auto
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
   130
qed auto
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   131
44368
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
   132
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
   133
  by (auto simp:subst_eq_def)
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   134
44368
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
   135
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
   136
  by (auto simp:subst_eq_def)
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   137
44368
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
   138
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
   139
  by auto
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   140
44368
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
   141
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
   142
  by (auto simp: subst_eq_def)
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   143
44370
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   144
lemma var_self: "[(v, Var v)] \<doteq> []"
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   145
proof
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   146
  fix t show "t \<lhd> [(v, Var v)] = t \<lhd> []"
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   147
    by (induct t) simp_all
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   148
qed
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   149
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   150
lemma var_same[simp]: "[(v, t)] \<doteq> [] \<longleftrightarrow> t = Var v"
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   151
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
   152
75637
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   153
lemma vars_of_subst_conv_Union: "vars_of (t \<lhd> \<eta>) = \<Union>(vars_of ` (\<lambda>x. Var x \<lhd> \<eta>) ` vars_of t)"
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   154
  by (induction t) simp_all
23219
87ad6e8a5f2c tuned document;
wenzelm
parents: 23024
diff changeset
   155
75638
aaa22adef039 added lemmas domain_comp and unify_gives_minimal_domain
desharna
parents: 75637
diff changeset
   156
lemma domain_comp: "fst ` set (\<sigma> \<lozenge> \<theta>) = fst ` (set \<sigma> \<union> set \<theta>)"
aaa22adef039 added lemmas domain_comp and unify_gives_minimal_domain
desharna
parents: 75637
diff changeset
   157
  by (induction \<sigma> \<theta> rule: comp.induct) auto
aaa22adef039 added lemmas domain_comp and unify_gives_minimal_domain
desharna
parents: 75637
diff changeset
   158
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   159
subsection \<open>Unifiers and Most General Unifiers\<close>
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   160
44368
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
   161
definition Unifier :: "'a subst \<Rightarrow> 'a trm \<Rightarrow> 'a trm \<Rightarrow> bool"
76643
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   162
  where "Unifier \<sigma> t u \<longleftrightarrow> (t \<lhd> \<sigma> = u \<lhd> \<sigma>)"
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   163
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   164
lemma not_occs_if_Unifier:
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   165
  assumes "Unifier \<sigma> t u"
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   166
  shows "\<not> (t \<prec> u) \<and> \<not> (u \<prec> t)"
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   167
proof -
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   168
  from assms have "t \<lhd> \<sigma> = u \<lhd> \<sigma>"
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   169
    unfolding Unifier_def by simp
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   170
  thus ?thesis
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   171
    using neq_if_occs subst_mono by metis
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   172
qed
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   173
44368
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
   174
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
   175
  "MGU \<sigma> t u \<longleftrightarrow> 
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
   176
   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
   177
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   178
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
   179
  "\<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
   180
  \<Longrightarrow> MGU \<sigma> t u"
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   181
  by (simp only:Unifier_def MGU_def, auto)
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   182
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   183
lemma MGU_sym[sym]:
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   184
  "MGU \<sigma> s t \<Longrightarrow> MGU \<sigma> t s"
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   185
  by (auto simp:MGU_def Unifier_def)
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   186
44371
3a10392fb8c3 added proof of idempotence, roughly after HOL/Subst/Unify.thy
krauss
parents: 44370
diff changeset
   187
lemma MGU_is_Unifier: "MGU \<sigma> t u \<Longrightarrow> Unifier \<sigma> t u"
3a10392fb8c3 added proof of idempotence, roughly after HOL/Subst/Unify.thy
krauss
parents: 44370
diff changeset
   188
unfolding MGU_def by (rule conjunct1)
3a10392fb8c3 added proof of idempotence, roughly after HOL/Subst/Unify.thy
krauss
parents: 44370
diff changeset
   189
44370
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   190
lemma MGU_Var: 
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   191
  assumes "\<not> Var v \<prec> t"
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   192
  shows "MGU [(v,t)] (Var v) t"
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   193
proof (intro MGUI exI)
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   194
  show "Var v \<lhd> [(v,t)] = t \<lhd> [(v,t)]" using assms
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   195
    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
   196
next
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   197
  fix \<theta> assume th: "Var v \<lhd> \<theta> = t \<lhd> \<theta>" 
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   198
  show "\<theta> \<doteq> [(v,t)] \<lozenge> \<theta>" 
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   199
  proof
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   200
    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
   201
      by (induct s) auto
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   202
  qed
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   203
qed
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   204
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   205
lemma MGU_Const: "MGU [] (Const c) (Const d) \<longleftrightarrow> c = d"
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   206
  by (auto simp: MGU_def Unifier_def)
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   207
  
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   208
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   209
subsection \<open>The unification algorithm\<close>
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   210
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   211
function unify :: "'a trm \<Rightarrow> 'a trm \<Rightarrow> 'a subst option"
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   212
where
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   213
  "unify (Const c) (M \<cdot> N)   = None"
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   214
| "unify (M \<cdot> N)   (Const c) = None"
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   215
| "unify (Const c) (Var v)   = Some [(v, Const c)]"
44369
02e13192a053 tuned notation
krauss
parents: 44368
diff changeset
   216
| "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
   217
                                        then None
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   218
                                        else Some [(v, M \<cdot> N)])"
44369
02e13192a053 tuned notation
krauss
parents: 44368
diff changeset
   219
| "unify (Var v)   M         = (if Var v \<prec> M
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   220
                                        then None
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   221
                                        else Some [(v, M)])"
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   222
| "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
   223
| "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
   224
                                    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
   225
                                    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
   226
                                      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
   227
                                         Some \<sigma> \<Rightarrow> Some (\<theta> \<lozenge> \<sigma>)))"
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   228
  by pat_completeness auto
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   229
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   230
subsection \<open>Properties used in termination proof\<close>
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   231
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   232
text \<open>Elimination of variables by a substitution:\<close>
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   233
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   234
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
   235
  "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
   236
44367
74c08021ab2e changed constant names and notation to match HOL/Subst/*.thy, from which this theory is a clone.
krauss
parents: 42463
diff changeset
   237
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
   238
  by (auto simp:elim_def)
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   239
44367
74c08021ab2e changed constant names and notation to match HOL/Subst/*.thy, from which this theory is a clone.
krauss
parents: 42463
diff changeset
   240
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
   241
  by (auto simp:elim_def)
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   242
44370
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   243
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
   244
  by (auto simp:elim_def subst_eq_def)
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   245
44369
02e13192a053 tuned notation
krauss
parents: 44368
diff changeset
   246
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
   247
  \<Longrightarrow> elim [(v,t)] v \<or> [(v,t)] \<doteq> []"
44370
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   248
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
   249
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   250
text \<open>The result of a unification never introduces new variables:\<close>
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   251
44370
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   252
declare unify.psimps[simp]
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   253
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   254
lemma unify_vars: 
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   255
  assumes "unify_dom (M, N)"
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   256
  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
   257
  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
   258
  (is "?P M N \<sigma> t")
24444
448978b61556 induct: proper separation of initial and terminal step;
wenzelm
parents: 23777
diff changeset
   259
using assms
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   260
proof (induct M N arbitrary:\<sigma> t)
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   261
  case (3 c v) 
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   262
  hence "\<sigma> = [(v, Const c)]" by simp
24444
448978b61556 induct: proper separation of initial and terminal step;
wenzelm
parents: 23777
diff changeset
   263
  thus ?case by (induct t) auto
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   264
next
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   265
  case (4 M N v) 
44369
02e13192a053 tuned notation
krauss
parents: 44368
diff changeset
   266
  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
   267
  with 4 have "\<sigma> = [(v, M\<cdot>N)]" by simp
448978b61556 induct: proper separation of initial and terminal step;
wenzelm
parents: 23777
diff changeset
   268
  thus ?case by (induct t) auto
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   269
next
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   270
  case (5 v M)
44369
02e13192a053 tuned notation
krauss
parents: 44368
diff changeset
   271
  hence "\<not> Var v \<prec> M" by auto
24444
448978b61556 induct: proper separation of initial and terminal step;
wenzelm
parents: 23777
diff changeset
   272
  with 5 have "\<sigma> = [(v, M)]" by simp
448978b61556 induct: proper separation of initial and terminal step;
wenzelm
parents: 23777
diff changeset
   273
  thus ?case by (induct t) auto
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   274
next
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   275
  case (7 M N M' N' \<sigma>)
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   276
  then obtain \<theta>1 \<theta>2 
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   277
    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
   278
    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
   279
    and \<sigma>: "\<sigma> = \<theta>1 \<lozenge> \<theta>2"
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   280
    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
   281
    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
   282
    by (auto split:option.split_asm)
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   283
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   284
  show ?case
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   285
  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
   286
    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
   287
    
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   288
    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
   289
    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
   290
        \<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
   291
      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
   292
      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
   293
        by auto
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   294
      
44367
74c08021ab2e changed constant names and notation to match HOL/Subst/*.thy, from which this theory is a clone.
krauss
parents: 42463
diff changeset
   295
      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
   296
      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
   297
        \<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
   298
        by auto
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   299
      hence "v \<in> vars_of t"
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   300
      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
   301
        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
   302
        with True show ?thesis by (auto dest:l)
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   303
      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
   304
        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
   305
        thus ?thesis by (rule l)
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   306
      qed
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   307
      
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   308
      thus ?thesis by auto
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   309
    qed auto
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   310
  qed
62390
842917225d56 more canonical names
nipkow
parents: 61933
diff changeset
   311
qed (auto split: if_split_asm)
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   312
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   313
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   314
text \<open>The result of a unification is either the identity
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   315
substitution or it eliminates a variable from one of the terms:\<close>
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   316
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   317
lemma unify_eliminates: 
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   318
  assumes "unify_dom (M, N)"
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   319
  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
   320
  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
   321
  (is "?P M N \<sigma>")
24444
448978b61556 induct: proper separation of initial and terminal step;
wenzelm
parents: 23777
diff changeset
   322
using assms
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   323
proof (induct M N arbitrary:\<sigma>)
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   324
  case 1 thus ?case by simp
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 2 thus ?case by simp
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   327
next
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   328
  case (3 c v)
44369
02e13192a053 tuned notation
krauss
parents: 44368
diff changeset
   329
  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
   330
  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
   331
  with occs_elim[OF no_occs]
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   332
  show ?case by auto
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   333
next
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   334
  case (4 M N v)
44369
02e13192a053 tuned notation
krauss
parents: 44368
diff changeset
   335
  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
   336
  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
   337
  with occs_elim[OF no_occs]
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   338
  show ?case by auto 
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   339
next
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   340
  case (5 v M) 
44369
02e13192a053 tuned notation
krauss
parents: 44368
diff changeset
   341
  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
   342
  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
   343
  with occs_elim[OF no_occs]
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   344
  show ?case by auto 
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   345
next 
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   346
  case (6 c d) thus ?case
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   347
    by (cases "c = d") auto
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   348
next
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   349
  case (7 M N M' N' \<sigma>)
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   350
  then obtain \<theta>1 \<theta>2 
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   351
    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
   352
    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
   353
    and \<sigma>: "\<sigma> = \<theta>1 \<lozenge> \<theta>2"
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   354
    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
   355
    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
   356
    by (auto split:option.split_asm)
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   357
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   358
  from \<open>unify_dom (M \<cdot> N, M' \<cdot> N')\<close>
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   359
  have "unify_dom (M, M')"
23777
60b7800338d5 Renamed accessible part for predicates to accp.
berghofe
parents: 23373
diff changeset
   360
    by (rule accp_downward) (rule unify_rel.intros)
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   361
  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
   362
    "\<And>t. vars_of (t \<lhd> \<theta>1) \<subseteq> vars_of M \<union> vars_of M' \<union> vars_of t"
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   363
    by (rule unify_vars) (rule \<open>unify M M' = Some \<theta>1\<close>)
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   364
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   365
  from ih2 show ?case 
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   366
  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
   367
    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
   368
    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
   369
      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
   370
      and el: "elim \<theta>2 v" by auto
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   371
    with no_new_vars show ?thesis unfolding \<sigma> 
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   372
      by (auto simp:elim_def)
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   373
  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
   374
    assume empty[simp]: "\<theta>2 \<doteq> []"
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   375
44367
74c08021ab2e changed constant names and notation to match HOL/Subst/*.thy, from which this theory is a clone.
krauss
parents: 42463
diff changeset
   376
    have "\<sigma> \<doteq> (\<theta>1 \<lozenge> [])" unfolding \<sigma>
44368
91e8062605d5 ported some lemmas from HOL/Subst/*;
krauss
parents: 44367
diff changeset
   377
      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
   378
    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
   379
    finally have "\<sigma> \<doteq> \<theta>1" .
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   380
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   381
    from ih1 show ?thesis
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   382
    proof
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   383
      assume "\<exists>v\<in>vars_of M \<union> vars_of M'. elim \<theta>1 v"
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   384
      with elim_eq[OF \<open>\<sigma> \<doteq> \<theta>1\<close>]
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   385
      show ?thesis by auto
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   386
    next
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   387
      note \<open>\<sigma> \<doteq> \<theta>1\<close>
44367
74c08021ab2e changed constant names and notation to match HOL/Subst/*.thy, from which this theory is a clone.
krauss
parents: 42463
diff changeset
   388
      also assume "\<theta>1 \<doteq> []"
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   389
      finally show ?thesis ..
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   390
    qed
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   391
  qed
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   392
qed
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   393
44370
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   394
declare unify.psimps[simp del]
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   395
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   396
subsection \<open>Termination proof\<close>
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   397
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   398
termination unify
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   399
proof 
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   400
  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
   401
                           \<lambda>(M, N). size M]"
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   402
  show "wf ?R" by simp
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   403
44370
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   404
  fix M N M' N' :: "'a trm"
67443
3abf6a722518 standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents: 62390
diff changeset
   405
  show "((M, M'), (M \<cdot> N, M' \<cdot> N')) \<in> ?R" \<comment> \<open>Inner call\<close>
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   406
    by (rule measures_lesseq) (auto intro: card_mono)
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   407
67443
3abf6a722518 standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents: 62390
diff changeset
   408
  fix \<theta>                                   \<comment> \<open>Outer call\<close>
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   409
  assume inner: "unify_dom (M, M')"
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   410
    "unify M M' = Some \<theta>"
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   411
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   412
  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
   413
  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
   414
  proof
61933
cf58b5b794b2 isabelle update_cartouches -c -t;
wenzelm
parents: 61343
diff changeset
   415
    \<comment> \<open>Either a variable is eliminated \ldots\<close>
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   416
    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
   417
    then obtain v 
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30909
diff changeset
   418
      where "elim \<theta> v" 
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30909
diff changeset
   419
      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
   420
    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
   421
    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
   422
      \<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
   423
      by auto
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   424
    
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   425
    thus ?thesis
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   426
      by (auto intro!: measures_less intro: psubset_card_mono)
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   427
  next
61933
cf58b5b794b2 isabelle update_cartouches -c -t;
wenzelm
parents: 61343
diff changeset
   428
    \<comment> \<open>Or the substitution is empty\<close>
44367
74c08021ab2e changed constant names and notation to match HOL/Subst/*.thy, from which this theory is a clone.
krauss
parents: 42463
diff changeset
   429
    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
   430
    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
   431
      and "N' \<lhd> \<theta> = N'" by auto
22999
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   432
    thus ?thesis 
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   433
       by (auto intro!: measures_less intro: psubset_card_mono)
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   434
  qed
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   435
qed
c1ce129e6f9c Added unification case study (using new function package)
krauss
parents:
diff changeset
   436
44370
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   437
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   438
subsection \<open>Unification returns a Most General Unifier\<close>
44370
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   439
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   440
lemma unify_computes_MGU:
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   441
  "unify M N = Some \<sigma> \<Longrightarrow> MGU \<sigma> M N"
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   442
proof (induct M N arbitrary: \<sigma> rule: unify.induct)
67443
3abf6a722518 standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents: 62390
diff changeset
   443
  case (7 M N M' N' \<sigma>) \<comment> \<open>The interesting case\<close>
44370
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   444
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   445
  then obtain \<theta>1 \<theta>2 
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   446
    where "unify M M' = Some \<theta>1"
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   447
    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
   448
    and \<sigma>: "\<sigma> = \<theta>1 \<lozenge> \<theta>2"
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   449
    and MGU_inner: "MGU \<theta>1 M M'" 
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   450
    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
   451
    by (auto split:option.split_asm)
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   452
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   453
  show ?case
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   454
  proof
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   455
    from MGU_inner and MGU_outer
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   456
    have "M \<lhd> \<theta>1 = M' \<lhd> \<theta>1" 
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   457
      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
   458
      unfolding MGU_def Unifier_def
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   459
      by auto
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   460
    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
   461
      by simp
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   462
  next
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   463
    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
   464
    hence "M \<lhd> \<sigma>' = M' \<lhd> \<sigma>'"
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   465
      and Ns: "N \<lhd> \<sigma>' = N' \<lhd> \<sigma>'" by auto
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   466
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   467
    with MGU_inner obtain \<delta>
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   468
      where eqv: "\<sigma>' \<doteq> \<theta>1 \<lozenge> \<delta>"
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   469
      unfolding MGU_def Unifier_def
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   470
      by auto
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   471
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   472
    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
   473
      by (simp add:subst_eq_dest[OF eqv])
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   474
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   475
    with MGU_outer obtain \<rho>
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   476
      where eqv2: "\<delta> \<doteq> \<theta>2 \<lozenge> \<rho>"
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   477
      unfolding MGU_def Unifier_def
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   478
      by auto
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   479
    
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   480
    have "\<sigma>' \<doteq> \<sigma> \<lozenge> \<rho>" unfolding \<sigma>
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   481
      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
   482
    thus "\<exists>\<gamma>. \<sigma>' \<doteq> \<sigma> \<lozenge> \<gamma>" ..
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   483
  qed
62390
842917225d56 more canonical names
nipkow
parents: 61933
diff changeset
   484
qed (auto simp: MGU_Const intro: MGU_Var MGU_Var[symmetric] split: if_split_asm)
44370
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   485
44372
f9825056dbab more precise authors and comments;
krauss
parents: 44371
diff changeset
   486
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   487
subsection \<open>Unification returns Idempotent Substitution\<close>
44372
f9825056dbab more precise authors and comments;
krauss
parents: 44371
diff changeset
   488
f9825056dbab more precise authors and comments;
krauss
parents: 44371
diff changeset
   489
definition Idem :: "'a subst \<Rightarrow> bool"
f9825056dbab more precise authors and comments;
krauss
parents: 44371
diff changeset
   490
where "Idem s \<longleftrightarrow> (s \<lozenge> s) \<doteq> s"
f9825056dbab more precise authors and comments;
krauss
parents: 44371
diff changeset
   491
44371
3a10392fb8c3 added proof of idempotence, roughly after HOL/Subst/Unify.thy
krauss
parents: 44370
diff changeset
   492
lemma Idem_Nil [iff]: "Idem []"
3a10392fb8c3 added proof of idempotence, roughly after HOL/Subst/Unify.thy
krauss
parents: 44370
diff changeset
   493
  by (simp add: Idem_def)
44370
03d91bfad83b tuned proofs, sledgehammering overly verbose parts
krauss
parents: 44369
diff changeset
   494
44371
3a10392fb8c3 added proof of idempotence, roughly after HOL/Subst/Unify.thy
krauss
parents: 44370
diff changeset
   495
lemma Var_Idem: 
3a10392fb8c3 added proof of idempotence, roughly after HOL/Subst/Unify.thy
krauss
parents: 44370
diff changeset
   496
  assumes "~ (Var v \<prec> t)" shows "Idem [(v,t)]"
3a10392fb8c3 added proof of idempotence, roughly after HOL/Subst/Unify.thy
krauss
parents: 44370
diff changeset
   497
  unfolding Idem_def
3a10392fb8c3 added proof of idempotence, roughly after HOL/Subst/Unify.thy
krauss
parents: 44370
diff changeset
   498
proof
3a10392fb8c3 added proof of idempotence, roughly after HOL/Subst/Unify.thy
krauss
parents: 44370
diff changeset
   499
  from assms have [simp]: "t \<lhd> [(v, t)] = t"
3a10392fb8c3 added proof of idempotence, roughly after HOL/Subst/Unify.thy
krauss
parents: 44370
diff changeset
   500
    by (metis assoc.simps(2) subst.simps(1) subst_no_occs)
3a10392fb8c3 added proof of idempotence, roughly after HOL/Subst/Unify.thy
krauss
parents: 44370
diff changeset
   501
3a10392fb8c3 added proof of idempotence, roughly after HOL/Subst/Unify.thy
krauss
parents: 44370
diff changeset
   502
  fix s show "s \<lhd> [(v, t)] \<lozenge> [(v, t)] = s \<lhd> [(v, t)]"
3a10392fb8c3 added proof of idempotence, roughly after HOL/Subst/Unify.thy
krauss
parents: 44370
diff changeset
   503
    by (induct s) auto
3a10392fb8c3 added proof of idempotence, roughly after HOL/Subst/Unify.thy
krauss
parents: 44370
diff changeset
   504
qed
3a10392fb8c3 added proof of idempotence, roughly after HOL/Subst/Unify.thy
krauss
parents: 44370
diff changeset
   505
3a10392fb8c3 added proof of idempotence, roughly after HOL/Subst/Unify.thy
krauss
parents: 44370
diff changeset
   506
lemma Unifier_Idem_subst: 
3a10392fb8c3 added proof of idempotence, roughly after HOL/Subst/Unify.thy
krauss
parents: 44370
diff changeset
   507
  "Idem(r) \<Longrightarrow> Unifier s (t \<lhd> r) (u \<lhd> r) \<Longrightarrow>
3a10392fb8c3 added proof of idempotence, roughly after HOL/Subst/Unify.thy
krauss
parents: 44370
diff changeset
   508
    Unifier (r \<lozenge> s) (t \<lhd> r) (u \<lhd> r)"
3a10392fb8c3 added proof of idempotence, roughly after HOL/Subst/Unify.thy
krauss
parents: 44370
diff changeset
   509
by (simp add: Idem_def Unifier_def subst_eq_def)
3a10392fb8c3 added proof of idempotence, roughly after HOL/Subst/Unify.thy
krauss
parents: 44370
diff changeset
   510
3a10392fb8c3 added proof of idempotence, roughly after HOL/Subst/Unify.thy
krauss
parents: 44370
diff changeset
   511
lemma Idem_comp:
3a10392fb8c3 added proof of idempotence, roughly after HOL/Subst/Unify.thy
krauss
parents: 44370
diff changeset
   512
  "Idem r \<Longrightarrow> Unifier s (t \<lhd> r) (u \<lhd> r) \<Longrightarrow>
3a10392fb8c3 added proof of idempotence, roughly after HOL/Subst/Unify.thy
krauss
parents: 44370
diff changeset
   513
      (!!q. Unifier q (t \<lhd> r) (u \<lhd> r) \<Longrightarrow> s \<lozenge> q \<doteq> q) \<Longrightarrow>
3a10392fb8c3 added proof of idempotence, roughly after HOL/Subst/Unify.thy
krauss
parents: 44370
diff changeset
   514
    Idem (r \<lozenge> s)"
3a10392fb8c3 added proof of idempotence, roughly after HOL/Subst/Unify.thy
krauss
parents: 44370
diff changeset
   515
  apply (frule Unifier_Idem_subst, blast) 
3a10392fb8c3 added proof of idempotence, roughly after HOL/Subst/Unify.thy
krauss
parents: 44370
diff changeset
   516
  apply (force simp add: Idem_def subst_eq_def)
3a10392fb8c3 added proof of idempotence, roughly after HOL/Subst/Unify.thy
krauss
parents: 44370
diff changeset
   517
  done
3a10392fb8c3 added proof of idempotence, roughly after HOL/Subst/Unify.thy
krauss
parents: 44370
diff changeset
   518
3a10392fb8c3 added proof of idempotence, roughly after HOL/Subst/Unify.thy
krauss
parents: 44370
diff changeset
   519
theorem unify_gives_Idem:
3a10392fb8c3 added proof of idempotence, roughly after HOL/Subst/Unify.thy
krauss
parents: 44370
diff changeset
   520
  "unify M N  = Some \<sigma> \<Longrightarrow> Idem \<sigma>"
3a10392fb8c3 added proof of idempotence, roughly after HOL/Subst/Unify.thy
krauss
parents: 44370
diff changeset
   521
proof (induct M N arbitrary: \<sigma> rule: unify.induct)
3a10392fb8c3 added proof of idempotence, roughly after HOL/Subst/Unify.thy
krauss
parents: 44370
diff changeset
   522
  case (7 M M' N N' \<sigma>)
3a10392fb8c3 added proof of idempotence, roughly after HOL/Subst/Unify.thy
krauss
parents: 44370
diff changeset
   523
3a10392fb8c3 added proof of idempotence, roughly after HOL/Subst/Unify.thy
krauss
parents: 44370
diff changeset
   524
  then obtain \<theta>1 \<theta>2 
3a10392fb8c3 added proof of idempotence, roughly after HOL/Subst/Unify.thy
krauss
parents: 44370
diff changeset
   525
    where "unify M N = Some \<theta>1"
3a10392fb8c3 added proof of idempotence, roughly after HOL/Subst/Unify.thy
krauss
parents: 44370
diff changeset
   526
    and \<theta>2: "unify (M' \<lhd> \<theta>1) (N' \<lhd> \<theta>1) = Some \<theta>2"
3a10392fb8c3 added proof of idempotence, roughly after HOL/Subst/Unify.thy
krauss
parents: 44370
diff changeset
   527
    and \<sigma>: "\<sigma> = \<theta>1 \<lozenge> \<theta>2"
3a10392fb8c3 added proof of idempotence, roughly after HOL/Subst/Unify.thy
krauss
parents: 44370
diff changeset
   528
    and "Idem \<theta>1" 
3a10392fb8c3 added proof of idempotence, roughly after HOL/Subst/Unify.thy
krauss
parents: 44370
diff changeset
   529
    and "Idem \<theta>2"
3a10392fb8c3 added proof of idempotence, roughly after HOL/Subst/Unify.thy
krauss
parents: 44370
diff changeset
   530
    by (auto split: option.split_asm)
3a10392fb8c3 added proof of idempotence, roughly after HOL/Subst/Unify.thy
krauss
parents: 44370
diff changeset
   531
3a10392fb8c3 added proof of idempotence, roughly after HOL/Subst/Unify.thy
krauss
parents: 44370
diff changeset
   532
  from \<theta>2 have "Unifier \<theta>2 (M' \<lhd> \<theta>1) (N' \<lhd> \<theta>1)"
3a10392fb8c3 added proof of idempotence, roughly after HOL/Subst/Unify.thy
krauss
parents: 44370
diff changeset
   533
    by (rule unify_computes_MGU[THEN MGU_is_Unifier])
3a10392fb8c3 added proof of idempotence, roughly after HOL/Subst/Unify.thy
krauss
parents: 44370
diff changeset
   534
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   535
  with \<open>Idem \<theta>1\<close>
44371
3a10392fb8c3 added proof of idempotence, roughly after HOL/Subst/Unify.thy
krauss
parents: 44370
diff changeset
   536
  show "Idem \<sigma>" unfolding \<sigma>
3a10392fb8c3 added proof of idempotence, roughly after HOL/Subst/Unify.thy
krauss
parents: 44370
diff changeset
   537
  proof (rule Idem_comp)
3a10392fb8c3 added proof of idempotence, roughly after HOL/Subst/Unify.thy
krauss
parents: 44370
diff changeset
   538
    fix \<sigma> assume "Unifier \<sigma> (M' \<lhd> \<theta>1) (N' \<lhd> \<theta>1)"
3a10392fb8c3 added proof of idempotence, roughly after HOL/Subst/Unify.thy
krauss
parents: 44370
diff changeset
   539
    with \<theta>2 obtain \<gamma> where \<sigma>: "\<sigma> \<doteq> \<theta>2 \<lozenge> \<gamma>"
3a10392fb8c3 added proof of idempotence, roughly after HOL/Subst/Unify.thy
krauss
parents: 44370
diff changeset
   540
      using unify_computes_MGU MGU_def by blast
3a10392fb8c3 added proof of idempotence, roughly after HOL/Subst/Unify.thy
krauss
parents: 44370
diff changeset
   541
3a10392fb8c3 added proof of idempotence, roughly after HOL/Subst/Unify.thy
krauss
parents: 44370
diff changeset
   542
    have "\<theta>2 \<lozenge> \<sigma> \<doteq> \<theta>2 \<lozenge> (\<theta>2 \<lozenge> \<gamma>)" by (rule subst_cong) (auto simp: \<sigma>)
3a10392fb8c3 added proof of idempotence, roughly after HOL/Subst/Unify.thy
krauss
parents: 44370
diff changeset
   543
    also have "... \<doteq> (\<theta>2 \<lozenge> \<theta>2) \<lozenge> \<gamma>" by (rule comp_assoc[symmetric])
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   544
    also have "... \<doteq> \<theta>2 \<lozenge> \<gamma>" by (rule subst_cong) (auto simp: \<open>Idem \<theta>2\<close>[unfolded Idem_def])
44371
3a10392fb8c3 added proof of idempotence, roughly after HOL/Subst/Unify.thy
krauss
parents: 44370
diff changeset
   545
    also have "... \<doteq> \<sigma>" by (rule \<sigma>[symmetric])
3a10392fb8c3 added proof of idempotence, roughly after HOL/Subst/Unify.thy
krauss
parents: 44370
diff changeset
   546
    finally show "\<theta>2 \<lozenge> \<sigma> \<doteq> \<sigma>" .
3a10392fb8c3 added proof of idempotence, roughly after HOL/Subst/Unify.thy
krauss
parents: 44370
diff changeset
   547
  qed
3a10392fb8c3 added proof of idempotence, roughly after HOL/Subst/Unify.thy
krauss
parents: 44370
diff changeset
   548
qed (auto intro!: Var_Idem split: option.splits if_splits)
39754
150f831ce4a3 no longer declare .psimps rules as [simp].
krauss
parents: 32960
diff changeset
   549
75635
3ba38a119739 added definition IMGU and lemmas IMGU_iff_Idem_and_MGU and unify_computes_IMGU
desharna
parents: 67443
diff changeset
   550
75641
a7e0129b8b2d tuned proofs
desharna
parents: 75639
diff changeset
   551
subsection \<open>Unification Returns Substitution With Minimal Domain And Range\<close>
75637
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   552
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   553
definition range_vars where
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   554
  "range_vars \<sigma> = \<Union> {vars_of (Var x \<lhd> \<sigma>) |x. Var x \<lhd> \<sigma> \<noteq> Var x}"
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   555
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   556
lemma vars_of_subst_subset: "vars_of (N \<lhd> \<sigma>) \<subseteq> vars_of N \<union> range_vars \<sigma>"
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   557
proof (rule subsetI)
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   558
  fix x assume "x \<in> vars_of (N \<lhd> \<sigma>)"
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   559
  thus "x \<in> vars_of N \<union> range_vars \<sigma>"
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   560
  proof (induction N)
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   561
    case (Var y)
75641
a7e0129b8b2d tuned proofs
desharna
parents: 75639
diff changeset
   562
    thus ?case
a7e0129b8b2d tuned proofs
desharna
parents: 75639
diff changeset
   563
      unfolding range_vars_def vars_of.simps by force
75637
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   564
  next
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   565
    case (Const y)
75641
a7e0129b8b2d tuned proofs
desharna
parents: 75639
diff changeset
   566
    thus ?case
a7e0129b8b2d tuned proofs
desharna
parents: 75639
diff changeset
   567
      by simp
75637
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   568
  next
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   569
    case (Comb N1 N2)
75641
a7e0129b8b2d tuned proofs
desharna
parents: 75639
diff changeset
   570
    thus ?case
75637
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   571
      by auto
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   572
  qed
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   573
qed
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   574
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   575
lemma range_vars_comp_subset: "range_vars (\<sigma>\<^sub>1 \<lozenge> \<sigma>\<^sub>2) \<subseteq> range_vars \<sigma>\<^sub>1 \<union> range_vars \<sigma>\<^sub>2"
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   576
proof (rule subsetI)
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   577
  fix x assume "x \<in> range_vars (\<sigma>\<^sub>1 \<lozenge> \<sigma>\<^sub>2)"
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   578
  then obtain x' where
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   579
    x'_\<sigma>\<^sub>1_\<sigma>\<^sub>2: "Var x' \<lhd> \<sigma>\<^sub>1 \<lhd> \<sigma>\<^sub>2 \<noteq> Var x'" and x_in: "x \<in> vars_of (Var x' \<lhd> \<sigma>\<^sub>1 \<lhd> \<sigma>\<^sub>2)"
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   580
    unfolding range_vars_def by auto
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   581
  
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   582
  show "x \<in> range_vars \<sigma>\<^sub>1 \<union> range_vars \<sigma>\<^sub>2"
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   583
  proof (cases "Var x' \<lhd> \<sigma>\<^sub>1 = Var x'")
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   584
    case True
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   585
    with x'_\<sigma>\<^sub>1_\<sigma>\<^sub>2 x_in show ?thesis
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   586
      unfolding range_vars_def by auto
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   587
  next
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   588
    case x'_\<sigma>\<^sub>1_neq: False
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   589
    show ?thesis
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   590
    proof (cases "Var x' \<lhd> \<sigma>\<^sub>1 \<lhd> \<sigma>\<^sub>2 = Var x' \<lhd> \<sigma>\<^sub>1")
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   591
      case True
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   592
      with x'_\<sigma>\<^sub>1_\<sigma>\<^sub>2 x_in x'_\<sigma>\<^sub>1_neq show ?thesis
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   593
        unfolding range_vars_def by auto
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   594
    next
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   595
      case False
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   596
      with x_in obtain y where "y \<in> vars_of (Var x' \<lhd> \<sigma>\<^sub>1)" and "x \<in> vars_of (Var y \<lhd> \<sigma>\<^sub>2)"
79558
58e974ef0625 tuned proofs --- avoid smt with external prover, which is somewhat unstable on arm64-linux;
wenzelm
parents: 76643
diff changeset
   597
        by (metis (no_types, lifting) UN_E UN_simps(10) vars_of_subst_conv_Union)
75637
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   598
      with x'_\<sigma>\<^sub>1_neq show ?thesis
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   599
        unfolding range_vars_def by force
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   600
    qed
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   601
  qed
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   602
qed
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   603
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   604
theorem unify_gives_minimal_range:
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   605
  "unify M N  = Some \<sigma> \<Longrightarrow> range_vars \<sigma> \<subseteq> vars_of M \<union> vars_of N"
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   606
proof (induct M N arbitrary: \<sigma> rule: unify.induct)
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   607
  case (1 c M N)
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   608
  thus ?case by simp
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   609
next
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   610
  case (2 M N c)
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   611
  thus ?case by simp
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   612
next
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   613
  case (3 c v)
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   614
  hence "\<sigma> = [(v, Const c)]"
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   615
    by simp
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   616
  thus ?case
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   617
    by (simp add: range_vars_def)
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   618
next
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   619
  case (4 M N v)
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   620
  hence "\<sigma> = [(v, M \<cdot> N)]"
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   621
    by (metis option.discI option.sel unify.simps(4))
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   622
  thus ?case
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   623
    by (auto simp: range_vars_def)
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   624
next
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   625
  case (5 v M)
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   626
  hence "\<sigma> = [(v, M)]"
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   627
    by (metis option.discI option.inject unify.simps(5))
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   628
  thus ?case
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   629
    by (auto simp: range_vars_def)
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   630
next
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   631
  case (6 c d)
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   632
  hence "\<sigma> = []"
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   633
    by (metis option.distinct(1) option.sel unify.simps(6))
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   634
  thus ?case
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   635
    by (simp add: range_vars_def)
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   636
next
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   637
  case (7 M N M' N')
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   638
  from "7.prems" obtain \<theta>\<^sub>1 \<theta>\<^sub>2 where
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   639
    "unify M M' = Some \<theta>\<^sub>1" and "unify (N \<lhd> \<theta>\<^sub>1) (N' \<lhd> \<theta>\<^sub>1) = Some \<theta>\<^sub>2" and "\<sigma> = \<theta>\<^sub>1 \<lozenge> \<theta>\<^sub>2"
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   640
    apply simp
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   641
    by (metis (no_types, lifting) option.case_eq_if option.collapse option.discI option.sel)
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   642
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   643
  from "7.hyps"(1) have range_\<theta>\<^sub>1: "range_vars \<theta>\<^sub>1 \<subseteq> vars_of M \<union> vars_of M'"
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   644
    using \<open>unify M M' = Some \<theta>\<^sub>1\<close> by simp
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   645
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   646
  from "7.hyps"(2) have "range_vars \<theta>\<^sub>2 \<subseteq> vars_of (N \<lhd> \<theta>\<^sub>1) \<union> vars_of (N' \<lhd> \<theta>\<^sub>1)"
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   647
    using \<open>unify M M' = Some \<theta>\<^sub>1\<close> \<open>unify (N \<lhd> \<theta>\<^sub>1) (N' \<lhd> \<theta>\<^sub>1) = Some \<theta>\<^sub>2\<close> by simp
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   648
  hence range_\<theta>\<^sub>2: "range_vars \<theta>\<^sub>2 \<subseteq> vars_of N \<union> vars_of N' \<union> range_vars \<theta>\<^sub>1"
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   649
    using vars_of_subst_subset[of _ \<theta>\<^sub>1] by auto
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   650
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   651
  have "range_vars \<sigma> = range_vars (\<theta>\<^sub>1 \<lozenge> \<theta>\<^sub>2)"
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   652
    unfolding \<open>\<sigma> = \<theta>\<^sub>1 \<lozenge> \<theta>\<^sub>2\<close> by simp
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   653
  also have "... \<subseteq> range_vars \<theta>\<^sub>1 \<union> range_vars \<theta>\<^sub>2"
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   654
    by (rule range_vars_comp_subset)
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   655
  also have "... \<subseteq> range_vars \<theta>\<^sub>1 \<union> vars_of N \<union> vars_of N'"
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   656
    using range_\<theta>\<^sub>2 by auto
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   657
  also have "... \<subseteq> vars_of M \<union> vars_of M' \<union> vars_of N \<union> vars_of N'"
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   658
    using range_\<theta>\<^sub>1 by auto
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   659
  finally show ?case
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   660
    by auto
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   661
qed
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   662
75638
aaa22adef039 added lemmas domain_comp and unify_gives_minimal_domain
desharna
parents: 75637
diff changeset
   663
theorem unify_gives_minimal_domain:
aaa22adef039 added lemmas domain_comp and unify_gives_minimal_domain
desharna
parents: 75637
diff changeset
   664
  "unify M N  = Some \<sigma> \<Longrightarrow> fst ` set \<sigma> \<subseteq> vars_of M \<union> vars_of N"
aaa22adef039 added lemmas domain_comp and unify_gives_minimal_domain
desharna
parents: 75637
diff changeset
   665
proof (induct M N arbitrary: \<sigma> rule: unify.induct)
aaa22adef039 added lemmas domain_comp and unify_gives_minimal_domain
desharna
parents: 75637
diff changeset
   666
  case (1 c M N)
aaa22adef039 added lemmas domain_comp and unify_gives_minimal_domain
desharna
parents: 75637
diff changeset
   667
  thus ?case
aaa22adef039 added lemmas domain_comp and unify_gives_minimal_domain
desharna
parents: 75637
diff changeset
   668
    by simp
aaa22adef039 added lemmas domain_comp and unify_gives_minimal_domain
desharna
parents: 75637
diff changeset
   669
next
aaa22adef039 added lemmas domain_comp and unify_gives_minimal_domain
desharna
parents: 75637
diff changeset
   670
  case (2 M N c)
aaa22adef039 added lemmas domain_comp and unify_gives_minimal_domain
desharna
parents: 75637
diff changeset
   671
  thus ?case
aaa22adef039 added lemmas domain_comp and unify_gives_minimal_domain
desharna
parents: 75637
diff changeset
   672
    by simp
aaa22adef039 added lemmas domain_comp and unify_gives_minimal_domain
desharna
parents: 75637
diff changeset
   673
next
aaa22adef039 added lemmas domain_comp and unify_gives_minimal_domain
desharna
parents: 75637
diff changeset
   674
  case (3 c v)
aaa22adef039 added lemmas domain_comp and unify_gives_minimal_domain
desharna
parents: 75637
diff changeset
   675
  hence "\<sigma> = [(v, Const c)]"
aaa22adef039 added lemmas domain_comp and unify_gives_minimal_domain
desharna
parents: 75637
diff changeset
   676
    by simp
aaa22adef039 added lemmas domain_comp and unify_gives_minimal_domain
desharna
parents: 75637
diff changeset
   677
  thus ?case
aaa22adef039 added lemmas domain_comp and unify_gives_minimal_domain
desharna
parents: 75637
diff changeset
   678
    by (simp add: dom_def)
aaa22adef039 added lemmas domain_comp and unify_gives_minimal_domain
desharna
parents: 75637
diff changeset
   679
next
aaa22adef039 added lemmas domain_comp and unify_gives_minimal_domain
desharna
parents: 75637
diff changeset
   680
  case (4 M N v)
aaa22adef039 added lemmas domain_comp and unify_gives_minimal_domain
desharna
parents: 75637
diff changeset
   681
  hence "\<sigma> = [(v, M \<cdot> N)]"
aaa22adef039 added lemmas domain_comp and unify_gives_minimal_domain
desharna
parents: 75637
diff changeset
   682
    by (metis option.distinct(1) option.inject unify.simps(4))
aaa22adef039 added lemmas domain_comp and unify_gives_minimal_domain
desharna
parents: 75637
diff changeset
   683
  thus ?case
aaa22adef039 added lemmas domain_comp and unify_gives_minimal_domain
desharna
parents: 75637
diff changeset
   684
    by (simp add: dom_def)
aaa22adef039 added lemmas domain_comp and unify_gives_minimal_domain
desharna
parents: 75637
diff changeset
   685
next
aaa22adef039 added lemmas domain_comp and unify_gives_minimal_domain
desharna
parents: 75637
diff changeset
   686
  case (5 v M)
aaa22adef039 added lemmas domain_comp and unify_gives_minimal_domain
desharna
parents: 75637
diff changeset
   687
  hence "\<sigma> = [(v, M)]"
aaa22adef039 added lemmas domain_comp and unify_gives_minimal_domain
desharna
parents: 75637
diff changeset
   688
    by (metis option.distinct(1) option.inject unify.simps(5))
aaa22adef039 added lemmas domain_comp and unify_gives_minimal_domain
desharna
parents: 75637
diff changeset
   689
  thus ?case
aaa22adef039 added lemmas domain_comp and unify_gives_minimal_domain
desharna
parents: 75637
diff changeset
   690
    by (simp add: dom_def)
aaa22adef039 added lemmas domain_comp and unify_gives_minimal_domain
desharna
parents: 75637
diff changeset
   691
next
aaa22adef039 added lemmas domain_comp and unify_gives_minimal_domain
desharna
parents: 75637
diff changeset
   692
  case (6 c d)
75641
a7e0129b8b2d tuned proofs
desharna
parents: 75639
diff changeset
   693
  hence "\<sigma> = []"
a7e0129b8b2d tuned proofs
desharna
parents: 75639
diff changeset
   694
    by (metis option.distinct(1) option.sel unify.simps(6))
a7e0129b8b2d tuned proofs
desharna
parents: 75639
diff changeset
   695
  thus ?case
a7e0129b8b2d tuned proofs
desharna
parents: 75639
diff changeset
   696
    by simp
75638
aaa22adef039 added lemmas domain_comp and unify_gives_minimal_domain
desharna
parents: 75637
diff changeset
   697
next
aaa22adef039 added lemmas domain_comp and unify_gives_minimal_domain
desharna
parents: 75637
diff changeset
   698
  case (7 M N M' N')
aaa22adef039 added lemmas domain_comp and unify_gives_minimal_domain
desharna
parents: 75637
diff changeset
   699
  from "7.prems" obtain \<theta>\<^sub>1 \<theta>\<^sub>2 where
aaa22adef039 added lemmas domain_comp and unify_gives_minimal_domain
desharna
parents: 75637
diff changeset
   700
    "unify M M' = Some \<theta>\<^sub>1" and "unify (N \<lhd> \<theta>\<^sub>1) (N' \<lhd> \<theta>\<^sub>1) = Some \<theta>\<^sub>2" and "\<sigma> = \<theta>\<^sub>1 \<lozenge> \<theta>\<^sub>2"
aaa22adef039 added lemmas domain_comp and unify_gives_minimal_domain
desharna
parents: 75637
diff changeset
   701
    apply simp
aaa22adef039 added lemmas domain_comp and unify_gives_minimal_domain
desharna
parents: 75637
diff changeset
   702
    by (metis (no_types, lifting) option.case_eq_if option.collapse option.discI option.sel)
aaa22adef039 added lemmas domain_comp and unify_gives_minimal_domain
desharna
parents: 75637
diff changeset
   703
aaa22adef039 added lemmas domain_comp and unify_gives_minimal_domain
desharna
parents: 75637
diff changeset
   704
  from "7.hyps"(1) have dom_\<theta>\<^sub>1: "fst ` set \<theta>\<^sub>1 \<subseteq> vars_of M \<union> vars_of M'"
aaa22adef039 added lemmas domain_comp and unify_gives_minimal_domain
desharna
parents: 75637
diff changeset
   705
    using \<open>unify M M' = Some \<theta>\<^sub>1\<close> by simp
aaa22adef039 added lemmas domain_comp and unify_gives_minimal_domain
desharna
parents: 75637
diff changeset
   706
aaa22adef039 added lemmas domain_comp and unify_gives_minimal_domain
desharna
parents: 75637
diff changeset
   707
  from "7.hyps"(2) have "fst ` set \<theta>\<^sub>2 \<subseteq> vars_of (N \<lhd> \<theta>\<^sub>1) \<union> vars_of (N' \<lhd> \<theta>\<^sub>1)"
aaa22adef039 added lemmas domain_comp and unify_gives_minimal_domain
desharna
parents: 75637
diff changeset
   708
    using \<open>unify M M' = Some \<theta>\<^sub>1\<close> \<open>unify (N \<lhd> \<theta>\<^sub>1) (N' \<lhd> \<theta>\<^sub>1) = Some \<theta>\<^sub>2\<close> by simp
75639
b8bd01897578 tuned proof
desharna
parents: 75638
diff changeset
   709
  hence dom_\<theta>\<^sub>2: "fst ` set \<theta>\<^sub>2 \<subseteq> vars_of N \<union> vars_of N' \<union> range_vars \<theta>\<^sub>1"
75638
aaa22adef039 added lemmas domain_comp and unify_gives_minimal_domain
desharna
parents: 75637
diff changeset
   710
    using vars_of_subst_subset[of _ \<theta>\<^sub>1] by auto
aaa22adef039 added lemmas domain_comp and unify_gives_minimal_domain
desharna
parents: 75637
diff changeset
   711
aaa22adef039 added lemmas domain_comp and unify_gives_minimal_domain
desharna
parents: 75637
diff changeset
   712
  have "fst ` set \<sigma> = fst ` set (\<theta>\<^sub>1 \<lozenge> \<theta>\<^sub>2)"
aaa22adef039 added lemmas domain_comp and unify_gives_minimal_domain
desharna
parents: 75637
diff changeset
   713
    unfolding \<open>\<sigma> = \<theta>\<^sub>1 \<lozenge> \<theta>\<^sub>2\<close> by simp
aaa22adef039 added lemmas domain_comp and unify_gives_minimal_domain
desharna
parents: 75637
diff changeset
   714
  also have "... = fst ` set \<theta>\<^sub>1 \<union> fst ` set \<theta>\<^sub>2"
aaa22adef039 added lemmas domain_comp and unify_gives_minimal_domain
desharna
parents: 75637
diff changeset
   715
    by (auto simp: domain_comp)
aaa22adef039 added lemmas domain_comp and unify_gives_minimal_domain
desharna
parents: 75637
diff changeset
   716
  also have "... \<subseteq> vars_of M \<union> vars_of M' \<union> fst ` set \<theta>\<^sub>2"
75641
a7e0129b8b2d tuned proofs
desharna
parents: 75639
diff changeset
   717
    using dom_\<theta>\<^sub>1 by auto
75638
aaa22adef039 added lemmas domain_comp and unify_gives_minimal_domain
desharna
parents: 75637
diff changeset
   718
  also have "... \<subseteq> vars_of M \<union> vars_of M' \<union> vars_of N \<union> vars_of N' \<union> range_vars \<theta>\<^sub>1"
75641
a7e0129b8b2d tuned proofs
desharna
parents: 75639
diff changeset
   719
    using dom_\<theta>\<^sub>2 by auto
75638
aaa22adef039 added lemmas domain_comp and unify_gives_minimal_domain
desharna
parents: 75637
diff changeset
   720
  also have "... \<subseteq> vars_of M \<union> vars_of M' \<union> vars_of N \<union> vars_of N'"
aaa22adef039 added lemmas domain_comp and unify_gives_minimal_domain
desharna
parents: 75637
diff changeset
   721
    using unify_gives_minimal_range[OF \<open>unify M M' = Some \<theta>\<^sub>1\<close>] by auto
aaa22adef039 added lemmas domain_comp and unify_gives_minimal_domain
desharna
parents: 75637
diff changeset
   722
  finally show ?case
aaa22adef039 added lemmas domain_comp and unify_gives_minimal_domain
desharna
parents: 75637
diff changeset
   723
    by auto
aaa22adef039 added lemmas domain_comp and unify_gives_minimal_domain
desharna
parents: 75637
diff changeset
   724
qed
aaa22adef039 added lemmas domain_comp and unify_gives_minimal_domain
desharna
parents: 75637
diff changeset
   725
75637
66a9aa769d63 added definition range_vars and lemmas vars_of_subst_conv_Union, vars_of_subst_subset, range_vars_comp_subset, and unify_gives_minimal_range
desharna
parents: 75635
diff changeset
   726
75635
3ba38a119739 added definition IMGU and lemmas IMGU_iff_Idem_and_MGU and unify_computes_IMGU
desharna
parents: 67443
diff changeset
   727
subsection \<open>Idempotent Most General Unifier\<close>
3ba38a119739 added definition IMGU and lemmas IMGU_iff_Idem_and_MGU and unify_computes_IMGU
desharna
parents: 67443
diff changeset
   728
3ba38a119739 added definition IMGU and lemmas IMGU_iff_Idem_and_MGU and unify_computes_IMGU
desharna
parents: 67443
diff changeset
   729
definition IMGU :: "'a subst \<Rightarrow> 'a trm \<Rightarrow> 'a trm \<Rightarrow> bool" where
3ba38a119739 added definition IMGU and lemmas IMGU_iff_Idem_and_MGU and unify_computes_IMGU
desharna
parents: 67443
diff changeset
   730
  "IMGU \<mu> t u \<longleftrightarrow> Unifier \<mu> t u \<and> (\<forall>\<theta>. Unifier \<theta> t u \<longrightarrow> \<theta> \<doteq> \<mu> \<lozenge> \<theta>)"
3ba38a119739 added definition IMGU and lemmas IMGU_iff_Idem_and_MGU and unify_computes_IMGU
desharna
parents: 67443
diff changeset
   731
3ba38a119739 added definition IMGU and lemmas IMGU_iff_Idem_and_MGU and unify_computes_IMGU
desharna
parents: 67443
diff changeset
   732
lemma IMGU_iff_Idem_and_MGU: "IMGU \<mu> t u \<longleftrightarrow> Idem \<mu> \<and> MGU \<mu> t u"
3ba38a119739 added definition IMGU and lemmas IMGU_iff_Idem_and_MGU and unify_computes_IMGU
desharna
parents: 67443
diff changeset
   733
  unfolding IMGU_def Idem_def MGU_def
79558
58e974ef0625 tuned proofs --- avoid smt with external prover, which is somewhat unstable on arm64-linux;
wenzelm
parents: 76643
diff changeset
   734
  by (meson Unification.comp_assoc subst_cong subst_refl subst_sym subst_trans)
75635
3ba38a119739 added definition IMGU and lemmas IMGU_iff_Idem_and_MGU and unify_computes_IMGU
desharna
parents: 67443
diff changeset
   735
3ba38a119739 added definition IMGU and lemmas IMGU_iff_Idem_and_MGU and unify_computes_IMGU
desharna
parents: 67443
diff changeset
   736
lemma unify_computes_IMGU:
3ba38a119739 added definition IMGU and lemmas IMGU_iff_Idem_and_MGU and unify_computes_IMGU
desharna
parents: 67443
diff changeset
   737
  "unify M N = Some \<sigma> \<Longrightarrow> IMGU \<sigma> M N"
3ba38a119739 added definition IMGU and lemmas IMGU_iff_Idem_and_MGU and unify_computes_IMGU
desharna
parents: 67443
diff changeset
   738
  by (simp add: IMGU_iff_Idem_and_MGU unify_computes_MGU unify_gives_Idem)
3ba38a119739 added definition IMGU and lemmas IMGU_iff_Idem_and_MGU and unify_computes_IMGU
desharna
parents: 67443
diff changeset
   739
76643
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   740
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   741
subsection \<open>Unification is complete\<close>
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   742
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   743
lemma unify_eq_Some_if_Unifier:
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   744
  assumes "Unifier \<sigma> t u"
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   745
  shows "\<exists>\<tau>. unify t u = Some \<tau>"
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   746
  using assms
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   747
proof (induction t u rule: unify.induct)
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   748
  case (1 c M N)
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   749
  thus ?case
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   750
    by (simp add: Unifier_def)
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   751
next
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   752
  case (2 M N c)
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   753
  thus ?case
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   754
    by (simp add: Unifier_def)
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   755
next
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   756
  case (3 c v)
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   757
  thus ?case
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   758
    by simp
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   759
next
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   760
  case (4 M N v)
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   761
  hence "\<not> (Var v \<prec> M \<cdot> N)"
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   762
    by (auto dest: not_occs_if_Unifier)
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   763
  thus ?case
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   764
    by simp
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   765
next
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   766
  case (5 v M)
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   767
  thus ?case
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   768
    by (auto dest: not_occs_if_Unifier)
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   769
next
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   770
  case (6 c d)
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   771
  thus ?case
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   772
    by (simp add: Unifier_def)
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   773
next
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   774
  case (7 M N M' N')
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   775
  from "7.prems" have "Unifier \<sigma> M M'"
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   776
    by (simp add: Unifier_def)
79558
58e974ef0625 tuned proofs --- avoid smt with external prover, which is somewhat unstable on arm64-linux;
wenzelm
parents: 76643
diff changeset
   777
  with "7.IH"(1) obtain \<tau> where \<tau>: "unify M M' = Some \<tau>"
76643
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   778
    by auto
79558
58e974ef0625 tuned proofs --- avoid smt with external prover, which is somewhat unstable on arm64-linux;
wenzelm
parents: 76643
diff changeset
   779
  then have "Unifier \<sigma> (N \<lhd> \<tau>) (N' \<lhd> \<tau>)"
58e974ef0625 tuned proofs --- avoid smt with external prover, which is somewhat unstable on arm64-linux;
wenzelm
parents: 76643
diff changeset
   780
    unfolding Unifier_def
58e974ef0625 tuned proofs --- avoid smt with external prover, which is somewhat unstable on arm64-linux;
wenzelm
parents: 76643
diff changeset
   781
    by (metis "7.prems" IMGU_def Unifier_def subst.simps(3) subst_comp subst_eq_def trm.distinct(3) trm.distinct(5) trm.exhaust trm.inject(3) unify_computes_IMGU)
58e974ef0625 tuned proofs --- avoid smt with external prover, which is somewhat unstable on arm64-linux;
wenzelm
parents: 76643
diff changeset
   782
  with \<tau> show ?case
76643
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   783
    using "7.IH"(2) by auto
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   784
qed
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   785
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   786
definition subst_domain where
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   787
  "subst_domain \<sigma> = {x. Var x \<lhd> \<sigma> \<noteq> Var x}"
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   788
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   789
lemma subst_domain_subset_list_domain: "subst_domain \<sigma> \<subseteq> fst ` set \<sigma>"
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   790
proof (rule Set.subsetI)
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   791
  fix x assume "x \<in> subst_domain \<sigma>"
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   792
  hence "Var x \<lhd> \<sigma> \<noteq> Var x"
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   793
    by (simp add: subst_domain_def)
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   794
  then show "x \<in> fst ` set \<sigma>"
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   795
  proof (induction \<sigma>)
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   796
    case Nil
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   797
    thus ?case
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   798
      by simp
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   799
  next
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   800
    case (Cons p \<sigma>)
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   801
    show ?case
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   802
    proof (cases "x = fst p")
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   803
      case True
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   804
      thus ?thesis
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   805
        by simp
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   806
    next
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   807
      case False
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   808
      with Cons.IH Cons.prems show ?thesis
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   809
        by (cases p) simp
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   810
    qed
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   811
  qed
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   812
qed
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   813
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   814
lemma subst_apply_eq_Var:
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   815
  assumes "s \<lhd> \<sigma> = Var x"
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   816
  obtains y where "s = Var y" and "Var y \<lhd> \<sigma> = Var x"
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   817
  using assms by (induct s) auto
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   818
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   819
lemma mem_range_varsI:
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   820
  assumes "Var x \<lhd> \<sigma> = Var y" and "x \<noteq> y"
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   821
  shows "y \<in> range_vars \<sigma>"
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   822
  using assms unfolding range_vars_def
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   823
  by (metis (mono_tags, lifting) UnionI mem_Collect_eq trm.inject(1) vars_iff_occseq)
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   824
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   825
lemma IMGU_subst_domain_subset:
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   826
  assumes "IMGU \<mu> t u"
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   827
  shows "subst_domain \<mu> \<subseteq> vars_of t \<union> vars_of u"
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   828
proof (rule Set.subsetI)
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   829
  from assms have "Unifier \<mu> t u"
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   830
    by (simp add: IMGU_def)
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   831
  then obtain \<upsilon> where "unify t u = Some \<upsilon>"
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   832
    using unify_eq_Some_if_Unifier by metis
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   833
  hence "Unifier \<upsilon> t u"
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   834
    using MGU_def unify_computes_MGU by blast
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   835
  with assms have "\<upsilon> \<doteq> \<mu> \<lozenge> \<upsilon>"
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   836
    by (simp add: IMGU_def)
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   837
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   838
  fix x assume "x \<in> subst_domain \<mu>"
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   839
  hence "Var x \<lhd> \<mu> \<noteq> Var x"
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   840
    by (simp add: subst_domain_def)
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   841
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   842
  show "x \<in> vars_of t \<union> vars_of u"
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   843
  proof (cases "x \<in> subst_domain \<upsilon>")
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   844
    case True
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   845
    hence "x \<in> fst ` set \<upsilon>"
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   846
      using subst_domain_subset_list_domain by fast
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   847
    thus ?thesis
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   848
      using unify_gives_minimal_domain[OF \<open>unify t u = Some \<upsilon>\<close>] by auto
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   849
  next
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   850
    case False
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   851
    hence "Var x \<lhd> \<upsilon> = Var x"
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   852
      by (simp add: subst_domain_def)
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   853
    hence "Var x \<lhd> \<mu> \<lhd> \<upsilon> = Var x"
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   854
      using \<open>\<upsilon> \<doteq> \<mu> \<lozenge> \<upsilon>\<close>
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   855
      by (metis subst_comp subst_eq_dest)
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   856
    then show ?thesis
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   857
      apply (rule subst_apply_eq_Var)
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   858
      using \<open>Var x \<lhd> \<mu> \<noteq> Var x\<close>
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   859
      using unify_gives_minimal_range[OF \<open>unify t u = Some \<upsilon>\<close>]
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   860
      using mem_range_varsI
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   861
      by force
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   862
  qed
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   863
qed
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   864
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   865
lemma IMGU_range_vars_subset:
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   866
  assumes "IMGU \<mu> t u"
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   867
  shows "range_vars \<mu> \<subseteq> vars_of t \<union> vars_of u"
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   868
proof (rule Set.subsetI)
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   869
  from assms have "Unifier \<mu> t u"
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   870
    by (simp add: IMGU_def)
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   871
  then obtain \<upsilon> where "unify t u = Some \<upsilon>"
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   872
    using unify_eq_Some_if_Unifier by metis
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   873
  hence "Unifier \<upsilon> t u"
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   874
    using MGU_def unify_computes_MGU by blast
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   875
  with assms have "\<upsilon> \<doteq> \<mu> \<lozenge> \<upsilon>"
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   876
    by (simp add: IMGU_def)
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   877
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   878
  have ball_subst_dom: "\<forall>x \<in> subst_domain \<upsilon>. vars_of (Var x \<lhd> \<upsilon>) \<subseteq> vars_of t \<union> vars_of u"
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   879
    using unify_gives_minimal_range[OF \<open>unify t u = Some \<upsilon>\<close>]
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   880
    using range_vars_def subst_domain_def by fastforce
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   881
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   882
  fix x assume "x \<in> range_vars \<mu>"
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   883
  then obtain y where "x \<in> vars_of (Var y \<lhd> \<mu>)" and "Var y \<lhd> \<mu> \<noteq> Var y"
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   884
    by (auto simp: range_vars_def)
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   885
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   886
  have vars_of_y_\<upsilon>: "vars_of (Var y \<lhd> \<upsilon>) \<subseteq> vars_of t \<union> vars_of u"
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   887
    using ball_subst_dom
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   888
    by (metis (mono_tags, lifting) IMGU_subst_domain_subset \<open>Var y \<lhd> \<mu> \<noteq> Var y\<close> assms empty_iff
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   889
        insert_iff mem_Collect_eq subset_eq subst_domain_def vars_of.simps(1))
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   890
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   891
  show "x \<in> vars_of t \<union> vars_of u"
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   892
  proof (rule ccontr)
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   893
    assume "x \<notin> vars_of t \<union> vars_of u"
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   894
    hence "x \<notin> vars_of (Var y \<lhd> \<upsilon>)"
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   895
      using vars_of_y_\<upsilon> by blast
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   896
    moreover have "x \<in> vars_of (Var y \<lhd> \<mu> \<lhd> \<upsilon>)"
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   897
    proof -
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   898
      have "Var x \<lhd> \<upsilon> = Var x"
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   899
        using \<open>x \<notin> vars_of t \<union> vars_of u\<close>
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   900
        using IMGU_subst_domain_subset \<open>unify t u = Some \<upsilon>\<close> subst_domain_def unify_computes_IMGU
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   901
        by fastforce
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   902
      thus ?thesis
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   903
        by (metis \<open>x \<in> vars_of (Var y \<lhd> \<mu>)\<close> subst_mono vars_iff_occseq)
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   904
    qed
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   905
    ultimately show False
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   906
      using \<open>\<upsilon> \<doteq> \<mu> \<lozenge> \<upsilon>\<close>
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   907
      by (metis subst_comp subst_eq_dest)
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   908
  qed
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   909
qed
f8826fc8c419 added lemmas IMGU_subst_domain_subset and IMGU_range_vars_subset
desharna
parents: 75641
diff changeset
   910
23219
87ad6e8a5f2c tuned document;
wenzelm
parents: 23024
diff changeset
   911
end