src/HOL/Subst/Unify.thy
author wenzelm
Wed, 03 Oct 2007 19:36:05 +0200
changeset 24823 bfb619994060
parent 24327 a207114007c6
child 26806 40b411ec05aa
permissions -rw-r--r--
modernized specifications; tuned proofs;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
15635
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
     1
(*  ID:         $Id$
3192
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents:
diff changeset
     2
    Author:     Konrad Slind, Cambridge University Computer Laboratory
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents:
diff changeset
     3
    Copyright   1997  University of Cambridge
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents:
diff changeset
     4
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents:
diff changeset
     5
*)
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents:
diff changeset
     6
15635
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
     7
header{*Unification Algorithm*}
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
     8
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
     9
theory Unify
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
    10
imports Unifier
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
    11
begin
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
    12
23000
6f158bba99e4 added pointer to new Unification theory
krauss
parents: 19769
diff changeset
    13
subsection {* Substitution and Unification in Higher-Order Logic. *}
6f158bba99e4 added pointer to new Unification theory
krauss
parents: 19769
diff changeset
    14
6f158bba99e4 added pointer to new Unification theory
krauss
parents: 19769
diff changeset
    15
text {*
6f158bba99e4 added pointer to new Unification theory
krauss
parents: 19769
diff changeset
    16
NB: This theory is already quite old.
6f158bba99e4 added pointer to new Unification theory
krauss
parents: 19769
diff changeset
    17
You might want to look at the newer Isar formalization of
6f158bba99e4 added pointer to new Unification theory
krauss
parents: 19769
diff changeset
    18
unification in HOL/ex/Unification.thy.
15635
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
    19
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
    20
Implements Manna and Waldinger's formalization, with Paulson's simplifications,
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
    21
and some new simplifications by Slind.
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
    22
15648
f6da795ee27a x-symbols and other tidying
paulson
parents: 15635
diff changeset
    23
Z Manna and R Waldinger, Deductive Synthesis of the Unification Algorithm.
15635
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
    24
SCP 1 (1981), 5-48
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
    25
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
    26
L C Paulson, Verifying the Unification Algorithm in LCF. SCP 5 (1985), 143-170
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
    27
*}
3192
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents:
diff changeset
    28
24823
bfb619994060 modernized specifications;
wenzelm
parents: 24327
diff changeset
    29
definition
bfb619994060 modernized specifications;
wenzelm
parents: 24327
diff changeset
    30
   unifyRel :: "(('a uterm * 'a uterm) * ('a uterm * 'a uterm)) set" where
bfb619994060 modernized specifications;
wenzelm
parents: 24327
diff changeset
    31
  "unifyRel = inv_image (finite_psubset <*lex*> measure uterm_size)
bfb619994060 modernized specifications;
wenzelm
parents: 24327
diff changeset
    32
      (%(M,N). (vars_of M Un vars_of N, M))"
15635
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
    33
   --{*Termination relation for the Unify function:
15648
f6da795ee27a x-symbols and other tidying
paulson
parents: 15635
diff changeset
    34
         either the set of variables decreases,
15635
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
    35
         or the first argument does (in fact, both do) *}
3192
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents:
diff changeset
    36
24823
bfb619994060 modernized specifications;
wenzelm
parents: 24327
diff changeset
    37
15635
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
    38
text{* Wellfoundedness of unifyRel *}
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
    39
lemma wf_unifyRel [iff]: "wf unifyRel"
24823
bfb619994060 modernized specifications;
wenzelm
parents: 24327
diff changeset
    40
  by (simp add: unifyRel_def wf_lex_prod wf_finite_psubset)
3192
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents:
diff changeset
    41
24823
bfb619994060 modernized specifications;
wenzelm
parents: 24327
diff changeset
    42
consts unify :: "'a uterm * 'a uterm => ('a * 'a uterm) list option"
15635
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
    43
recdef (permissive) unify "unifyRel"
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
    44
 unify_CC: "unify(Const m, Const n)  = (if (m=n) then Some[] else None)"
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
    45
 unify_CB: "unify(Const m, Comb M N) = None"
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
    46
 unify_CV: "unify(Const m, Var v)    = Some[(v,Const m)]"
15648
f6da795ee27a x-symbols and other tidying
paulson
parents: 15635
diff changeset
    47
 unify_V:  "unify(Var v, M) = (if (Var v \<prec> M) then None else Some[(v,M)])"
15635
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
    48
 unify_BC: "unify(Comb M N, Const x) = None"
15648
f6da795ee27a x-symbols and other tidying
paulson
parents: 15635
diff changeset
    49
 unify_BV: "unify(Comb M N, Var v)   = (if (Var v \<prec> Comb M N) then None
15635
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
    50
                                        else Some[(v,Comb M N)])"
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
    51
 unify_BB:
15648
f6da795ee27a x-symbols and other tidying
paulson
parents: 15635
diff changeset
    52
  "unify(Comb M1 N1, Comb M2 N2) =
f6da795ee27a x-symbols and other tidying
paulson
parents: 15635
diff changeset
    53
      (case unify(M1,M2)
f6da795ee27a x-symbols and other tidying
paulson
parents: 15635
diff changeset
    54
        of None       => None
f6da795ee27a x-symbols and other tidying
paulson
parents: 15635
diff changeset
    55
         | Some theta => (case unify(N1 \<lhd> theta, N2 \<lhd> theta)
f6da795ee27a x-symbols and other tidying
paulson
parents: 15635
diff changeset
    56
                            of None       => None
f6da795ee27a x-symbols and other tidying
paulson
parents: 15635
diff changeset
    57
                             | Some sigma => Some (theta \<lozenge> sigma)))"
15635
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
    58
  (hints recdef_wf: wf_unifyRel)
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
    59
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
    60
15648
f6da795ee27a x-symbols and other tidying
paulson
parents: 15635
diff changeset
    61
text{* This file defines a nested unification algorithm, then proves that it
f6da795ee27a x-symbols and other tidying
paulson
parents: 15635
diff changeset
    62
 terminates, then proves 2 correctness theorems: that when the algorithm
f6da795ee27a x-symbols and other tidying
paulson
parents: 15635
diff changeset
    63
 succeeds, it 1) returns an MGU; and 2) returns an idempotent substitution.
f6da795ee27a x-symbols and other tidying
paulson
parents: 15635
diff changeset
    64
 Although the proofs may seem long, they are actually quite direct, in that
f6da795ee27a x-symbols and other tidying
paulson
parents: 15635
diff changeset
    65
 the correctness and termination properties are not mingled as much as in
f6da795ee27a x-symbols and other tidying
paulson
parents: 15635
diff changeset
    66
 previous proofs of this algorithm.*}
15635
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
    67
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
    68
(*---------------------------------------------------------------------------
15648
f6da795ee27a x-symbols and other tidying
paulson
parents: 15635
diff changeset
    69
 * Our approach for nested recursive functions is as follows:
15635
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
    70
 *
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
    71
 *    0. Prove the wellfoundedness of the termination relation.
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
    72
 *    1. Prove the non-nested termination conditions.
15648
f6da795ee27a x-symbols and other tidying
paulson
parents: 15635
diff changeset
    73
 *    2. Eliminate (0) and (1) from the recursion equations and the
15635
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
    74
 *       induction theorem.
15648
f6da795ee27a x-symbols and other tidying
paulson
parents: 15635
diff changeset
    75
 *    3. Prove the nested termination conditions by using the induction
f6da795ee27a x-symbols and other tidying
paulson
parents: 15635
diff changeset
    76
 *       theorem from (2) and by using the recursion equations from (2).
f6da795ee27a x-symbols and other tidying
paulson
parents: 15635
diff changeset
    77
 *       These are constrained by the nested termination conditions, but
f6da795ee27a x-symbols and other tidying
paulson
parents: 15635
diff changeset
    78
 *       things work out magically (by wellfoundedness of the termination
15635
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
    79
 *       relation).
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
    80
 *    4. Eliminate the nested TCs from the results of (2).
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
    81
 *    5. Prove further correctness properties using the results of (4).
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
    82
 *
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
    83
 * Deeper nestings require iteration of steps (3) and (4).
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
    84
 *---------------------------------------------------------------------------*)
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
    85
15648
f6da795ee27a x-symbols and other tidying
paulson
parents: 15635
diff changeset
    86
text{*The non-nested TC (terminiation condition).*}
f6da795ee27a x-symbols and other tidying
paulson
parents: 15635
diff changeset
    87
recdef_tc unify_tc1: unify (1)
19769
c40ce2de2020 Added [simp]-lemmas "in_inv_image" and "in_lex_prod" in the spirit of "in_measure".
krauss
parents: 19623
diff changeset
    88
  by (auto simp: unifyRel_def finite_psubset_def finite_vars_of)
15635
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
    89
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
    90
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
    91
text{*Termination proof.*}
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
    92
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
    93
lemma trans_unifyRel: "trans unifyRel"
24823
bfb619994060 modernized specifications;
wenzelm
parents: 24327
diff changeset
    94
  by (simp add: unifyRel_def measure_def trans_inv_image trans_lex_prod
bfb619994060 modernized specifications;
wenzelm
parents: 24327
diff changeset
    95
    trans_finite_psubset)
15635
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
    96
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
    97
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
    98
text{*The following lemma is used in the last step of the termination proof
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
    99
for the nested call in Unify.  Loosely, it says that unifyRel doesn't care
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
   100
about term structure.*}
15648
f6da795ee27a x-symbols and other tidying
paulson
parents: 15635
diff changeset
   101
lemma Rassoc:
f6da795ee27a x-symbols and other tidying
paulson
parents: 15635
diff changeset
   102
  "((X,Y), (Comb A (Comb B C), Comb D (Comb E F))) \<in> unifyRel  ==>
24823
bfb619994060 modernized specifications;
wenzelm
parents: 24327
diff changeset
   103
    ((X,Y), (Comb (Comb A B) C, Comb (Comb D E) F)) \<in> unifyRel"
bfb619994060 modernized specifications;
wenzelm
parents: 24327
diff changeset
   104
  by (simp add: less_eq add_assoc Un_assoc unifyRel_def lex_prod_def)
15635
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
   105
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
   106
15648
f6da795ee27a x-symbols and other tidying
paulson
parents: 15635
diff changeset
   107
text{*This lemma proves the nested termination condition for the base cases
15635
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
   108
 * 3, 4, and 6.*}
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
   109
lemma var_elimR:
15648
f6da795ee27a x-symbols and other tidying
paulson
parents: 15635
diff changeset
   110
  "~(Var x \<prec> M) ==>
f6da795ee27a x-symbols and other tidying
paulson
parents: 15635
diff changeset
   111
    ((N1 \<lhd> [(x,M)], N2 \<lhd> [(x,M)]), (Comb M N1, Comb(Var x) N2)) \<in> unifyRel
f6da795ee27a x-symbols and other tidying
paulson
parents: 15635
diff changeset
   112
  & ((N1 \<lhd> [(x,M)], N2 \<lhd> [(x,M)]), (Comb(Var x) N1, Comb M N2)) \<in> unifyRel"
15635
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
   113
apply (case_tac "Var x = M", clarify, simp)
15648
f6da795ee27a x-symbols and other tidying
paulson
parents: 15635
diff changeset
   114
apply (case_tac "x \<in> (vars_of N1 Un vars_of N2)")
24823
bfb619994060 modernized specifications;
wenzelm
parents: 24327
diff changeset
   115
txt{*@{text uterm_less} case*}
19769
c40ce2de2020 Added [simp]-lemmas "in_inv_image" and "in_lex_prod" in the spirit of "in_measure".
krauss
parents: 19623
diff changeset
   116
apply (simp add: less_eq unifyRel_def lex_prod_def)
15635
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
   117
apply blast
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
   118
txt{*@{text finite_psubset} case*}
19769
c40ce2de2020 Added [simp]-lemmas "in_inv_image" and "in_lex_prod" in the spirit of "in_measure".
krauss
parents: 19623
diff changeset
   119
apply (simp add: unifyRel_def lex_prod_def)
15635
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
   120
apply (simp add: finite_psubset_def finite_vars_of psubset_def)
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
   121
apply blast
15648
f6da795ee27a x-symbols and other tidying
paulson
parents: 15635
diff changeset
   122
txt{*Final case, also @{text finite_psubset}*}
19769
c40ce2de2020 Added [simp]-lemmas "in_inv_image" and "in_lex_prod" in the spirit of "in_measure".
krauss
parents: 19623
diff changeset
   123
apply (simp add: finite_vars_of unifyRel_def finite_psubset_def lex_prod_def)
15648
f6da795ee27a x-symbols and other tidying
paulson
parents: 15635
diff changeset
   124
apply (cut_tac s = "[(x,M)]" and v = x and t = N2 in Var_elim)
f6da795ee27a x-symbols and other tidying
paulson
parents: 15635
diff changeset
   125
apply (cut_tac [3] s = "[(x,M)]" and v = x and t = N1 in Var_elim)
15635
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
   126
apply (simp_all (no_asm_simp) add: srange_iff vars_iff_occseq)
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
   127
apply (auto elim!: Var_intro [THEN disjE] simp add: srange_iff)
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
   128
done
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
   129
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
   130
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
   131
text{*Eliminate tc1 from the recursion equations and the induction theorem.*}
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
   132
15648
f6da795ee27a x-symbols and other tidying
paulson
parents: 15635
diff changeset
   133
lemmas unify_nonrec [simp] =
f6da795ee27a x-symbols and other tidying
paulson
parents: 15635
diff changeset
   134
       unify_CC unify_CB unify_CV unify_V unify_BC unify_BV
15635
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
   135
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
   136
lemmas unify_simps0 = unify_nonrec unify_BB [OF unify_tc1]
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
   137
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
   138
lemmas unify_induct0 = unify.induct [OF unify_tc1]
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
   139
15648
f6da795ee27a x-symbols and other tidying
paulson
parents: 15635
diff changeset
   140
text{*The nested TC. The (2) requests the second one.
f6da795ee27a x-symbols and other tidying
paulson
parents: 15635
diff changeset
   141
      Proved by recursion induction.*}
f6da795ee27a x-symbols and other tidying
paulson
parents: 15635
diff changeset
   142
recdef_tc unify_tc2: unify (2)
15635
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
   143
txt{*The extracted TC needs the scope of its quantifiers adjusted, so our
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
   144
 first step is to restrict the scopes of N1 and N2.*}
15648
f6da795ee27a x-symbols and other tidying
paulson
parents: 15635
diff changeset
   145
apply (subgoal_tac "\<forall>M1 M2 theta. unify (M1, M2) = Some theta -->
f6da795ee27a x-symbols and other tidying
paulson
parents: 15635
diff changeset
   146
      (\<forall>N1 N2.((N1\<lhd>theta, N2\<lhd>theta), (Comb M1 N1, Comb M2 N2)) \<in> unifyRel)")
15635
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
   147
apply blast
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
   148
apply (rule allI)
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
   149
apply (rule allI)
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
   150
txt{*Apply induction on this still-quantified formula*}
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
   151
apply (rule_tac u = M1 and v = M2 in unify_induct0)
15648
f6da795ee27a x-symbols and other tidying
paulson
parents: 15635
diff changeset
   152
      apply (simp_all (no_asm_simp) add: var_elimR unify_simps0)
f6da795ee27a x-symbols and other tidying
paulson
parents: 15635
diff changeset
   153
 txt{*Const-Const case*}
19769
c40ce2de2020 Added [simp]-lemmas "in_inv_image" and "in_lex_prod" in the spirit of "in_measure".
krauss
parents: 19623
diff changeset
   154
 apply (simp add: unifyRel_def lex_prod_def less_eq)
15635
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
   155
txt{*Comb-Comb case*}
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
   156
apply (simp (no_asm_simp) split add: option.split)
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
   157
apply (intro strip)
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
   158
apply (rule trans_unifyRel [THEN transD], blast)
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
   159
apply (simp only: subst_Comb [symmetric])
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
   160
apply (rule Rassoc, blast)
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
   161
done
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
   162
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
   163
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
   164
text{* Now for elimination of nested TC from unify.simps and induction.*}
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
   165
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
   166
text{*Desired rule, copied from the theory file.*}
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
   167
lemma unifyCombComb [simp]:
15648
f6da795ee27a x-symbols and other tidying
paulson
parents: 15635
diff changeset
   168
    "unify(Comb M1 N1, Comb M2 N2) =
f6da795ee27a x-symbols and other tidying
paulson
parents: 15635
diff changeset
   169
       (case unify(M1,M2)
f6da795ee27a x-symbols and other tidying
paulson
parents: 15635
diff changeset
   170
         of None => None
f6da795ee27a x-symbols and other tidying
paulson
parents: 15635
diff changeset
   171
          | Some theta => (case unify(N1 \<lhd> theta, N2 \<lhd> theta)
f6da795ee27a x-symbols and other tidying
paulson
parents: 15635
diff changeset
   172
                             of None => None
f6da795ee27a x-symbols and other tidying
paulson
parents: 15635
diff changeset
   173
                              | Some sigma => Some (theta \<lozenge> sigma)))"
15635
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
   174
by (simp add: unify_tc2 unify_simps0 split add: option.split)
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
   175
24327
a207114007c6 removed obsolete ML bindings;
wenzelm
parents: 23000
diff changeset
   176
lemma unify_induct:
a207114007c6 removed obsolete ML bindings;
wenzelm
parents: 23000
diff changeset
   177
  "(\<And>m n. P (Const m) (Const n)) \<Longrightarrow>
a207114007c6 removed obsolete ML bindings;
wenzelm
parents: 23000
diff changeset
   178
  (\<And>m M N. P (Const m) (Comb M N)) \<Longrightarrow>
a207114007c6 removed obsolete ML bindings;
wenzelm
parents: 23000
diff changeset
   179
  (\<And>m v. P (Const m) (Var v)) \<Longrightarrow>
a207114007c6 removed obsolete ML bindings;
wenzelm
parents: 23000
diff changeset
   180
  (\<And>v M. P (Var v) M) \<Longrightarrow>
a207114007c6 removed obsolete ML bindings;
wenzelm
parents: 23000
diff changeset
   181
  (\<And>M N x. P (Comb M N) (Const x)) \<Longrightarrow>
a207114007c6 removed obsolete ML bindings;
wenzelm
parents: 23000
diff changeset
   182
  (\<And>M N v. P (Comb M N) (Var v)) \<Longrightarrow>
a207114007c6 removed obsolete ML bindings;
wenzelm
parents: 23000
diff changeset
   183
  (\<And>M1 N1 M2 N2.
a207114007c6 removed obsolete ML bindings;
wenzelm
parents: 23000
diff changeset
   184
    \<forall>theta. unify (M1, M2) = Some theta \<longrightarrow> P (N1 \<lhd> theta) (N2 \<lhd> theta) \<Longrightarrow>
a207114007c6 removed obsolete ML bindings;
wenzelm
parents: 23000
diff changeset
   185
    P M1 M2 \<Longrightarrow> P (Comb M1 N1) (Comb M2 N2)) \<Longrightarrow>
a207114007c6 removed obsolete ML bindings;
wenzelm
parents: 23000
diff changeset
   186
  P u v"
a207114007c6 removed obsolete ML bindings;
wenzelm
parents: 23000
diff changeset
   187
by (rule unify_induct0) (simp_all add: unify_tc2)
15635
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
   188
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
   189
text{*Correctness. Notice that idempotence is not needed to prove that the
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
   190
algorithm terminates and is not needed to prove the algorithm correct, if you
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
   191
are only interested in an MGU.  This is in contrast to the approach of M&W,
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
   192
who used idempotence and MGU-ness in the termination proof.*}
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
   193
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
   194
theorem unify_gives_MGU [rule_format]:
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
   195
     "\<forall>theta. unify(M,N) = Some theta --> MGUnifier theta M N"
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
   196
apply (rule_tac u = M and v = N in unify_induct0)
15648
f6da795ee27a x-symbols and other tidying
paulson
parents: 15635
diff changeset
   197
    apply (simp_all (no_asm_simp))
f6da795ee27a x-symbols and other tidying
paulson
parents: 15635
diff changeset
   198
    txt{*Const-Const case*}
f6da795ee27a x-symbols and other tidying
paulson
parents: 15635
diff changeset
   199
    apply (simp add: MGUnifier_def Unifier_def)
f6da795ee27a x-symbols and other tidying
paulson
parents: 15635
diff changeset
   200
   txt{*Const-Var case*}
f6da795ee27a x-symbols and other tidying
paulson
parents: 15635
diff changeset
   201
   apply (subst mgu_sym)
f6da795ee27a x-symbols and other tidying
paulson
parents: 15635
diff changeset
   202
   apply (simp add: MGUnifier_Var)
f6da795ee27a x-symbols and other tidying
paulson
parents: 15635
diff changeset
   203
  txt{*Var-M case*}
f6da795ee27a x-symbols and other tidying
paulson
parents: 15635
diff changeset
   204
  apply (simp add: MGUnifier_Var)
f6da795ee27a x-symbols and other tidying
paulson
parents: 15635
diff changeset
   205
 txt{*Comb-Var case*}
f6da795ee27a x-symbols and other tidying
paulson
parents: 15635
diff changeset
   206
 apply (subst mgu_sym)
f6da795ee27a x-symbols and other tidying
paulson
parents: 15635
diff changeset
   207
 apply (simp add: MGUnifier_Var)
f6da795ee27a x-symbols and other tidying
paulson
parents: 15635
diff changeset
   208
txt{*Comb-Comb case*}
f6da795ee27a x-symbols and other tidying
paulson
parents: 15635
diff changeset
   209
apply (simp add: unify_tc2)
15635
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
   210
apply (simp (no_asm_simp) split add: option.split)
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
   211
apply (intro strip)
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
   212
apply (simp add: MGUnifier_def Unifier_def MoreGeneral_def)
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
   213
apply (safe, rename_tac theta sigma gamma)
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
   214
apply (erule_tac x = gamma in allE, erule (1) notE impE)
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
   215
apply (erule exE, rename_tac delta)
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
   216
apply (erule_tac x = delta in allE)
15648
f6da795ee27a x-symbols and other tidying
paulson
parents: 15635
diff changeset
   217
apply (subgoal_tac "N1 \<lhd> theta \<lhd> delta = N2 \<lhd> theta \<lhd> delta")
15635
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
   218
 apply (blast intro: subst_trans intro!: subst_cong comp_assoc[THEN subst_sym])
15648
f6da795ee27a x-symbols and other tidying
paulson
parents: 15635
diff changeset
   219
apply (simp add: subst_eq_iff)
15635
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
   220
done
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
   221
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
   222
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
   223
text{*Unify returns idempotent substitutions, when it succeeds.*}
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
   224
theorem unify_gives_Idem [rule_format]:
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
   225
     "\<forall>theta. unify(M,N) = Some theta --> Idem theta"
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
   226
apply (rule_tac u = M and v = N in unify_induct0)
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
   227
apply (simp_all add: Var_Idem unify_tc2 split add: option.split)
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
   228
txt{*Comb-Comb case*}
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
   229
apply safe
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
   230
apply (drule spec, erule (1) notE impE)+
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
   231
apply (safe dest!: unify_gives_MGU [unfolded MGUnifier_def])
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
   232
apply (rule Idem_comp, assumption+)
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
   233
apply (force simp add: MoreGeneral_def subst_eq_iff Idem_def)
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
   234
done
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
   235
3192
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents:
diff changeset
   236
end