TFL/examples/Subst/Unify1.ML
author wenzelm
Mon, 06 Jan 1997 17:02:09 +0100
changeset 2471 09634c9cbf3c
parent 2113 21266526ac42
permissions -rw-r--r--
added stamp util;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
2113
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
     1
(*---------------------------------------------------------------------------
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
     2
 * This file defines a nested unification algorithm, then proves that it 
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
     3
 * terminates, then proves 2 correctness theorems: that when the algorithm
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
     4
 * succeeds, it 1) returns an MGU; and 2) returns an idempotent substitution.
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
     5
 * Although the proofs may seem long, they are actually quite direct, in that
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
     6
 * the correctness and termination properties are not mingled as much as in 
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
     7
 * previous proofs of this algorithm. 
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
     8
 *
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
     9
 * Our approach for nested recursive functions is as follows: 
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
    10
 *
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
    11
 *    0. Prove the wellfoundedness of the termination relation.
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
    12
 *    1. Prove the non-nested termination conditions.
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
    13
 *    2. Eliminate (0) and (1) from the recursion equations and the 
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
    14
 *       induction theorem.
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
    15
 *    3. Prove the nested termination conditions by using the induction 
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
    16
 *       theorem from (2) and by using the recursion equations from (2). 
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
    17
 *       These are constrained by the nested termination conditions, but 
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
    18
 *       things work out magically (by wellfoundedness of the termination 
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
    19
 *       relation).
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
    20
 *    4. Eliminate the nested TCs from the results of (2).
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
    21
 *    5. Prove further correctness properties using the results of (4).
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
    22
 *
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
    23
 * Deeper nestings require iteration of steps (3) and (4).
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
    24
 *---------------------------------------------------------------------------*)
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
    25
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
    26
Thry.add_datatype_facts
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
    27
   (UTerm.thy, ("uterm",["Var", "Const", "Comb"]), uterm.induct_tac);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
    28
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
    29
Thry.add_datatype_facts 
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
    30
   (Unify1.thy, ("subst",["Fail", "Subst"]), Unify1.subst.induct_tac);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
    31
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
    32
(* This is just a wrapper for the definition mechanism. *)
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
    33
local fun cread thy s = read_cterm (sign_of thy) (s, (TVar(("DUMMY",0),[])));
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
    34
in
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
    35
fun Rfunc thy R eqs =
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
    36
   let val _ = reset_count()
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
    37
       val _ =  tych_counting true
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
    38
       val read = term_of o cread thy;
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
    39
       val {induction,rules,theory,tcs} = Tfl.Rfunction thy (read R) (read eqs)
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
    40
   in {induction=induction, rules=rules, theory=theory, 
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
    41
       typechecks = count(), tcs = tcs}
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
    42
   end
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
    43
end;
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
    44
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
    45
(*---------------------------------------------------------------------------
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
    46
 * The algorithm.
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
    47
 *---------------------------------------------------------------------------*)
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
    48
val {theory,induction,rules,tcs,typechecks} =
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
    49
Rfunc Unify1.thy "R"
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
    50
  "(Unify(Const m, Const n)  = (if (m=n) then Subst[] else Fail))    & \
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
    51
\  (Unify(Const m, Comb M N) = Fail)                                 & \
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
    52
\  (Unify(Const m, Var v)    = Subst[(v,Const m)])                   & \
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
    53
\  (Unify(Var v, M) = (if (Var v <: M) then Fail else Subst[(v,M)])) & \
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
    54
\  (Unify(Comb M N, Const x) = Fail)                                 & \
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
    55
\  (Unify(Comb M N, Var v) = (if (Var v <: Comb M N) then Fail  \
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
    56
\                             else Subst[(v,Comb M N)]))             & \
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
    57
\  (Unify(Comb M1 N1, Comb M2 N2) =  \
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
    58
\     (case Unify(M1,M2) \
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
    59
\       of Fail => Fail \
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
    60
\        | Subst theta => (case Unify(N1 <| theta, N2 <| theta) \
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
    61
\                           of Fail => Fail \
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
    62
\                             | Subst sigma => Subst (theta <> sigma))))";
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
    63
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
    64
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
    65
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
    66
open Unify1;
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
    67
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
    68
(*---------------------------------------------------------------------------
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
    69
 * A slightly augmented strip_tac. 
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
    70
 *---------------------------------------------------------------------------*)
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
    71
(* Needs a correct "CHANGED" to work! This one taken from Carsten's version. *)
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
    72
(*Returns all changed states*)
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
    73
fun CHANGED tac st = 
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
    74
    let fun diff st' = not (eq_thm(st,st'))
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
    75
    in  Sequence.filters diff (tac st)  end;
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
    76
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
    77
fun my_strip_tac i = 
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
    78
   CHANGED (strip_tac i 
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
    79
            THEN REPEAT ((etac exE ORELSE' etac conjE) i)
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
    80
            THEN TRY (hyp_subst_tac i));
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
    81
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
    82
(*---------------------------------------------------------------------------
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
    83
 * A slightly augmented fast_tac for sets. It handles the case where the 
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
    84
 * top connective is "=".
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
    85
 *---------------------------------------------------------------------------*)
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
    86
fun my_fast_set_tac i = (TRY(rtac set_ext i) THEN fast_tac set_cs i);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
    87
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
    88
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
    89
(*---------------------------------------------------------------------------
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
    90
 * Wellfoundedness of proper subset on finite sets.
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
    91
 *---------------------------------------------------------------------------*)
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
    92
goalw Unify1.thy [R0_def] "wf(R0)";
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
    93
by (rtac ((wf_subset RS mp) RS mp) 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
    94
by (rtac wf_measure 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
    95
by(simp_tac(!simpset addsimps[measure_def,inv_image_def,symmetric less_def])1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
    96
by (my_strip_tac 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
    97
by (forward_tac[ssubset_card] 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
    98
by (fast_tac set_cs 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
    99
val wf_R0 = result();
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   100
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   101
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   102
(*---------------------------------------------------------------------------
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   103
 * Tactic for selecting and working on the first projection of R.
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   104
 *---------------------------------------------------------------------------*)
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   105
fun R0_tac thms i =
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   106
  (simp_tac (!simpset addsimps (thms@[R_def,lex_prod_def,
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   107
               measure_def,inv_image_def,point_to_prod_def])) i THEN
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   108
   REPEAT (rtac exI i) THEN
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   109
   REPEAT ((rtac conjI THEN' rtac refl) i) THEN
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   110
   rtac disjI1 i THEN
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   111
   simp_tac (!simpset addsimps [R0_def,finite_vars_of]) i);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   112
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   113
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   114
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   115
(*---------------------------------------------------------------------------
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   116
 * Tactic for selecting and working on the second projection of R.
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   117
 *---------------------------------------------------------------------------*)
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   118
fun R1_tac thms i = 
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   119
   (simp_tac (!simpset addsimps (thms@[R_def,lex_prod_def,
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   120
                 measure_def,inv_image_def,point_to_prod_def])) i THEN 
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   121
    REPEAT (rtac exI i) THEN 
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   122
    REPEAT ((rtac conjI THEN' rtac refl) i) THEN
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   123
    rtac disjI2 i THEN
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   124
    asm_simp_tac (!simpset addsimps [R1_def,rprod_def]) i);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   125
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   126
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   127
(*---------------------------------------------------------------------------
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   128
 * The non-nested TC plus the wellfoundedness of R.
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   129
 *---------------------------------------------------------------------------*)
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   130
tgoalw Unify1.thy [] rules;
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   131
by (rtac conjI 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   132
(* TC *)
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   133
by (my_strip_tac 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   134
by (cut_facts_tac [monotone_vars_of] 1); 
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   135
by (asm_full_simp_tac(!simpset addsimps [subseteq_iff_subset_eq]) 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   136
by (etac disjE 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   137
by (R0_tac[] 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   138
by (R1_tac[] 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   139
by (REPEAT (rtac exI 1) THEN REPEAT ((rtac conjI THEN' rtac refl) 1));
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   140
by (simp_tac
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   141
     (!simpset addsimps [measure_def,inv_image_def,less_eq,less_add_Suc1]) 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   142
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   143
(* Wellfoundedness of R *)
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   144
by (simp_tac (!simpset addsimps [Unify1.R_def,Unify1.R1_def]) 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   145
by (REPEAT (resolve_tac [wf_inv_image,wf_lex_prod,wf_R0,
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   146
                         wf_rel_prod, wf_measure] 1));
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   147
val tc0 = result();
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   148
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   149
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   150
(*---------------------------------------------------------------------------
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   151
 * Eliminate tc0 from the recursion equations and the induction theorem.
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   152
 *---------------------------------------------------------------------------*)
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   153
val [tc,wfr] = Prim.Rules.CONJUNCTS tc0;
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   154
val rules1 = implies_intr_hyps rules;
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   155
val rules2 = wfr RS rules1;
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   156
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   157
val [a,b,c,d,e,f,g] = Prim.Rules.CONJUNCTS rules2;
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   158
val g' = tc RS (g RS mp);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   159
val rules4 = standard (Prim.Rules.LIST_CONJ[a,b,c,d,e,f,g']);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   160
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   161
val induction1 = implies_intr_hyps induction;
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   162
val induction2 = wfr RS induction1;
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   163
val induction3 = tc RS induction2;
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   164
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   165
val induction4 = standard
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   166
 (rewrite_rule[fst_conv RS eq_reflection, snd_conv RS eq_reflection]
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   167
   (induction3 RS (read_instantiate_sg (sign_of theory)
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   168
      [("x","%p. Phi (fst p) (snd p)")] spec)));
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   169
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   170
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   171
(*---------------------------------------------------------------------------
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   172
 * Some theorems about transitivity of WF combinators. Only the last
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   173
 * (transR) is used, in the proof of termination. The others are generic and
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   174
 * should maybe go somewhere.
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   175
 *---------------------------------------------------------------------------*)
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   176
goalw WF1.thy [trans_def,lex_prod_def,mem_Collect_eq RS eq_reflection]
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   177
           "trans R1 & trans R2 --> trans (R1 ** R2)";
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   178
by (my_strip_tac 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   179
by (res_inst_tac [("x","a")] exI 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   180
by (res_inst_tac [("x","a'a")] exI 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   181
by (res_inst_tac [("x","b")] exI 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   182
by (res_inst_tac [("x","b'a")] exI 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   183
by (REPEAT (rewrite_tac [Pair_eq RS eq_reflection] THEN my_strip_tac 1));
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   184
by (Simp_tac 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   185
by (REPEAT (etac disjE 1));
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   186
by (rtac disjI1 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   187
by (ALLGOALS (fast_tac set_cs));
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   188
val trans_lex_prod = result() RS mp;
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   189
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   190
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   191
goalw WF1.thy [trans_def,rprod_def,mem_Collect_eq RS eq_reflection]
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   192
           "trans R1 & trans R2 --> trans (rprod R1  R2)";
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   193
by (my_strip_tac 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   194
by (res_inst_tac [("x","a")] exI 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   195
by (res_inst_tac [("x","a'a")] exI 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   196
by (res_inst_tac [("x","b")] exI 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   197
by (res_inst_tac [("x","b'a")] exI 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   198
by (REPEAT (rewrite_tac [Pair_eq RS eq_reflection] THEN my_strip_tac 1));
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   199
by (Simp_tac 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   200
by (fast_tac set_cs 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   201
val trans_rprod = result() RS mp;
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   202
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   203
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   204
goalw Unify1.thy [trans_def,inv_image_def,mem_Collect_eq RS eq_reflection]
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   205
 "trans r --> trans (inv_image r f)";
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   206
by (rewrite_tac [fst_conv RS eq_reflection, snd_conv RS eq_reflection]);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   207
by (fast_tac set_cs 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   208
val trans_inv_image = result() RS mp;
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   209
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   210
goalw Unify1.thy [R0_def, trans_def, mem_Collect_eq RS eq_reflection]
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   211
 "trans R0";
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   212
by (rewrite_tac [fst_conv RS eq_reflection,snd_conv RS eq_reflection,
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   213
                 ssubset_def, set_eq_subset RS eq_reflection]);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   214
by (fast_tac set_cs 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   215
val trans_R0 = result();
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   216
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   217
goalw Unify1.thy [R_def,R1_def,measure_def] "trans R";
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   218
by (REPEAT (resolve_tac[trans_inv_image,trans_lex_prod,conjI, trans_R0,
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   219
                        trans_rprod, trans_inv_image, trans_trancl] 1));
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   220
val transR = result();
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   221
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   222
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   223
(*---------------------------------------------------------------------------
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   224
 * The following lemma is used in the last step of the termination proof for 
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   225
 * the nested call in Unify. Loosely, it says that R doesn't care so much
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   226
 * about term structure.
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   227
 *---------------------------------------------------------------------------*)
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   228
goalw Unify1.thy [R_def,lex_prod_def, inv_image_def,point_to_prod_def]
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   229
     "((X,Y), (Comb A (Comb B C), Comb D (Comb E F))) : R --> \
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   230
    \ ((X,Y), (Comb (Comb A B) C, Comb (Comb D E) F)) : R";
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   231
by (Simp_tac 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   232
by (my_strip_tac 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   233
by (REPEAT (rtac exI 1) THEN REPEAT ((rtac conjI THEN' rtac refl) 1));
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   234
by (etac disjE 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   235
by (rtac disjI1 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   236
by (subgoal_tac "(vars_of A Un vars_of B Un vars_of C Un \
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   237
                \  (vars_of D Un vars_of E Un vars_of F)) = \
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   238
                \ (vars_of A Un (vars_of B Un vars_of C) Un \
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   239
                \  (vars_of D Un (vars_of E Un vars_of F)))" 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   240
by (my_fast_set_tac 2);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   241
by (Asm_simp_tac 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   242
by (rtac disjI2 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   243
by (etac conjE 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   244
by (Asm_simp_tac 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   245
by (rtac conjI 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   246
by (my_fast_set_tac 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   247
by (asm_full_simp_tac (!simpset addsimps [R1_def, measure_def, rprod_def,
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   248
                                          less_eq, inv_image_def]) 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   249
by (my_strip_tac 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   250
by (REPEAT (rtac exI 1) THEN REPEAT ((rtac conjI THEN' rtac refl) 1));
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   251
by (asm_full_simp_tac (HOL_ss addsimps [uterm_size_Comb,
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   252
                                        add_Suc_right,add_Suc,add_assoc]) 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   253
val Rassoc = result() RS mp;
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   254
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   255
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   256
(*---------------------------------------------------------------------------
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   257
 * Rewriting support.
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   258
 *---------------------------------------------------------------------------*)
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   259
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   260
val termin_ss = (!simpset addsimps (srange_iff::(subst_rews@al_rews)));
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   261
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   262
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   263
(*---------------------------------------------------------------------------
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   264
 * This lemma proves the nested termination condition for the base cases 
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   265
 * 3, 4, and 6. It's a clumsy formulation (requiring two conjuncts, each with
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   266
 * exactly the same proof) of a more general theorem.
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   267
 *---------------------------------------------------------------------------*)
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   268
goal theory "(~(Var x <: M)) --> [(x, M)] = theta -->       \
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   269
\ (! N1 N2. (((N1 <| theta, N2 <| theta), (Comb M N1, Comb (Var x) N2)) : R) \
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   270
\       &   (((N1 <| theta, N2 <| theta), (Comb(Var x) N1, Comb M N2)) : R))";
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   271
by (my_strip_tac 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   272
by (case_tac "Var x = M" 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   273
by (hyp_subst_tac 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   274
by (case_tac "x:(vars_of N1 Un vars_of N2)" 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   275
let val case1 = 
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   276
   EVERY1[R1_tac[id_subst_lemma], rtac conjI, my_fast_set_tac,
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   277
          REPEAT o (rtac exI), REPEAT o (rtac conjI THEN' rtac refl),
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   278
          simp_tac (!simpset addsimps [measure_def,inv_image_def,less_eq])];
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   279
in by (rtac conjI 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   280
   by case1;
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   281
   by case1
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   282
end;
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   283
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   284
let val case2 = 
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   285
   EVERY1[R0_tac[id_subst_lemma],
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   286
          simp_tac (!simpset addsimps [ssubset_def,set_eq_subset]),
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   287
          fast_tac set_cs]
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   288
in by (rtac conjI 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   289
   by case2;
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   290
   by case2
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   291
end;
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   292
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   293
let val case3 =  
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   294
 EVERY1 [R0_tac[],
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   295
        cut_inst_tac [("s2","[(x, M)]"), ("v2", "x"), ("t2","N1")] Var_elim] 
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   296
 THEN ALLGOALS(asm_simp_tac(termin_ss addsimps [vars_iff_occseq]))
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   297
 THEN cut_inst_tac [("s2","[(x, M)]"),("v2", "x"), ("t2","N2")] Var_elim 1
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   298
 THEN ALLGOALS(asm_simp_tac(termin_ss addsimps [vars_iff_occseq]))
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   299
 THEN EVERY1 [simp_tac (HOL_ss addsimps [ssubset_def]),
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   300
             rtac conjI, simp_tac (HOL_ss addsimps [subset_iff]),
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   301
             my_strip_tac, etac UnE, dtac Var_intro] 
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   302
 THEN dtac Var_intro 2
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   303
 THEN ALLGOALS (asm_full_simp_tac (termin_ss addsimps [set_eq_subset])) 
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   304
 THEN TRYALL (fast_tac set_cs)
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   305
in 
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   306
  by (rtac conjI 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   307
  by case3;
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   308
  by case3
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   309
end;
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   310
val var_elimR = result() RS mp RS mp RS spec RS spec;
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   311
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   312
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   313
val Some{nchotomy = subst_nchotomy,...} = 
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   314
    assoc(Thry.datatype_facts theory,"subst");
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   315
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   316
(*---------------------------------------------------------------------------
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   317
 * Do a case analysis on something of type 'a subst. 
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   318
 *---------------------------------------------------------------------------*)
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   319
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   320
fun Subst_case_tac theta =
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   321
(cut_inst_tac theta (standard (Prim.Rules.SPEC_ALL subst_nchotomy)) 1 
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   322
  THEN etac disjE 1 
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   323
  THEN rotate_tac ~1 1 
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   324
  THEN Asm_full_simp_tac 1 
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   325
  THEN etac exE 1
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   326
  THEN rotate_tac ~1 1 
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   327
  THEN Asm_full_simp_tac 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   328
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   329
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   330
goals_limit := 1;
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   331
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   332
(*---------------------------------------------------------------------------
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   333
 * The nested TC. Proved by recursion induction.
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   334
 *---------------------------------------------------------------------------*)
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   335
goalw_cterm [] 
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   336
     (hd(tl(tl(map (cterm_of (sign_of theory) o USyntax.mk_prop) tcs))));
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   337
(*---------------------------------------------------------------------------
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   338
 * The extracted TC needs the scope of its quantifiers adjusted, so our 
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   339
 * first step is to restrict the scopes of N1 and N2.
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   340
 *---------------------------------------------------------------------------*)
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   341
by (subgoal_tac "!M1 M2 theta.  \
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   342
 \     Unify (M1, M2) = Subst theta --> \
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   343
 \    (!N1 N2. ((N1 <| theta, N2 <| theta), Comb M1 N1, Comb M2 N2) : R)" 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   344
by (fast_tac HOL_cs 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   345
by (rtac allI 1); 
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   346
by (rtac allI 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   347
(* Apply induction *)
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   348
by (res_inst_tac [("xa","M1"),("x","M2")] 
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   349
                 (standard (induction4 RS mp RS spec RS spec)) 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   350
by (simp_tac (!simpset addsimps (rules4::(subst_rews@al_rews))
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   351
                       setloop (split_tac [expand_if])) 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   352
(* 1 *)
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   353
by (rtac conjI 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   354
by (my_strip_tac 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   355
by (R1_tac[subst_Nil] 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   356
by (REPEAT (rtac exI 1) THEN REPEAT ((rtac conjI THEN' rtac refl) 1));
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   357
by (simp_tac (!simpset addsimps [measure_def,inv_image_def,less_eq]) 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   358
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   359
(* 3 *)
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   360
by (rtac conjI 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   361
by (my_strip_tac 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   362
by (rtac (Prim.Rules.CONJUNCT1 var_elimR) 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   363
by (Simp_tac 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   364
by (rtac refl 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   365
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   366
(* 4 *)
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   367
by (rtac conjI 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   368
by (strip_tac 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   369
by (rtac (Prim.Rules.CONJUNCT2 var_elimR) 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   370
by (assume_tac 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   371
by (assume_tac 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   372
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   373
(* 6 *)
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   374
by (rtac conjI 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   375
by (rewrite_tac [symmetric (occs_Comb RS eq_reflection)]);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   376
by (my_strip_tac 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   377
by (rtac (Prim.Rules.CONJUNCT1 var_elimR) 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   378
by (assume_tac 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   379
by (rtac refl 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   380
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   381
(* 7 *)
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   382
by (REPEAT (rtac allI 1));
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   383
by (rtac impI 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   384
by (etac conjE 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   385
by (Subst_case_tac [("v","Unify(M1a, M2a)")]);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   386
by (REPEAT (eres_inst_tac [("x","list")] allE 1));
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   387
by (asm_full_simp_tac  HOL_ss 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   388
by (subgoal_tac "((N1 <| list, N2 <| list), Comb M1a N1, Comb M2a N2) : R" 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   389
by (asm_simp_tac HOL_ss 2);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   390
by (dtac mp 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   391
by (assume_tac 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   392
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   393
by (Subst_case_tac [("v","Unify(N1 <| list, N2 <| list)")]);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   394
by (eres_inst_tac [("x","lista")] allE 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   395
by (asm_full_simp_tac  HOL_ss 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   396
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   397
by (rtac allI 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   398
by (rtac impI 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   399
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   400
by (hyp_subst_tac 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   401
by (REPEAT (rtac allI 1));
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   402
by (rename_tac "foo bar M1 N1 M2 N2 theta sigma gamma P1 P2" 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   403
by (simp_tac (HOL_ss addsimps [subst_comp]) 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   404
by(rtac(rewrite_rule[trans_def] transR RS spec RS spec RS spec RS mp RS mp) 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   405
by (fast_tac HOL_cs 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   406
by (simp_tac (HOL_ss addsimps [symmetric (subst_Comb RS eq_reflection)]) 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   407
by (subgoal_tac "((Comb N1 P1 <| theta, Comb N2 P2 <| theta), \
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   408
                \ (Comb M1 (Comb N1 P1), Comb M2 (Comb N2 P2))) :R" 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   409
by (asm_simp_tac HOL_ss 2);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   410
by (rtac Rassoc 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   411
by (assume_tac 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   412
val Unify_TC2 = result();
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   413
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   414
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   415
(*---------------------------------------------------------------------------
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   416
 * Now for elimination of nested TC from rules and induction. This step 
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   417
 * would be easier if "rewrite_rule" used context.
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   418
 *---------------------------------------------------------------------------*)
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   419
goal theory 
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   420
 "(Unify (Comb M1 N1, Comb M2 N2) =  \
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   421
\   (case Unify (M1, M2) of Fail => Fail \
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   422
\    | Subst theta => \
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   423
\        (case if ((N1 <| theta, N2 <| theta), Comb M1 N1, Comb M2 N2) : R \
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   424
\              then Unify (N1 <| theta, N2 <| theta) else @ z. True of \
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   425
\        Fail => Fail | Subst sigma => Subst (theta <> sigma)))) \
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   426
\  = \
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   427
\ (Unify (Comb M1 N1, Comb M2 N2) = \
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   428
\   (case Unify (M1, M2)  \
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   429
\      of Fail => Fail \
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   430
\      | Subst theta => (case Unify (N1 <| theta, N2 <| theta) \
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   431
\                          of Fail => Fail  \
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   432
\                           | Subst sigma => Subst (theta <> sigma))))";
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   433
by (cut_inst_tac [("v","Unify(M1, M2)")]
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   434
                 (standard (Prim.Rules.SPEC_ALL subst_nchotomy)) 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   435
by (etac disjE 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   436
by (Asm_simp_tac 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   437
by (etac exE 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   438
by (Asm_simp_tac 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   439
by (cut_inst_tac 
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   440
     [("x","list"), ("xb","N1"), ("xa","N2"),("xc","M2"), ("xd","M1")]
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   441
     (standard(Unify_TC2 RS spec RS spec RS spec RS spec RS spec)) 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   442
by (Asm_full_simp_tac 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   443
val Unify_rec_simpl = result() RS eq_reflection;
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   444
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   445
val Unify_rules = rewrite_rule[Unify_rec_simpl] rules4;
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   446
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   447
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   448
goal theory 
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   449
 "(! M1 N1 M2 N2.  \
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   450
\       (! theta.  \
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   451
\           Unify (M1, M2) = Subst theta -->  \
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   452
\           ((N1 <| theta, N2 <| theta), Comb M1 N1, Comb M2 N2) : R -->  \
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   453
\           ?Phi (N1 <| theta) (N2 <| theta)) & ?Phi M1 M2 -->  \
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   454
\       ?Phi (Comb M1 N1) (Comb M2 N2))  \
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   455
\    =  \
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   456
\ (! M1 N1 M2 N2.  \
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   457
\       (! theta.  \
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   458
\           Unify (M1, M2) = Subst theta -->  \
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   459
\           ?Phi (N1 <| theta) (N2 <| theta)) & ?Phi M1 M2 -->  \
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   460
\       ?Phi (Comb M1 N1) (Comb M2 N2))";
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   461
by (simp_tac (HOL_ss addsimps [Unify_TC2]) 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   462
val Unify_induction = rewrite_rule[result() RS eq_reflection] induction4;
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   463
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   464
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   465
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   466
(*---------------------------------------------------------------------------
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   467
 * Correctness. Notice that idempotence is not needed to prove that the 
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   468
 * algorithm terminates and is not needed to prove the algorithm correct, 
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   469
 * if you are only interested in an MGU. This is in contrast to the
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   470
 * approach of M&W, who used idempotence and MGU-ness in the termination proof.
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   471
 *---------------------------------------------------------------------------*)
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   472
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   473
goal theory "!theta. Unify (P,Q) = Subst theta --> MGUnifier theta P Q";
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   474
by (res_inst_tac [("xa","P"),("x","Q")] 
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   475
                 (standard (Unify_induction RS mp RS spec RS spec)) 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   476
by (simp_tac (!simpset addsimps [Unify_rules] 
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   477
                       setloop (split_tac [expand_if])) 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   478
(*1*)
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   479
by (rtac conjI 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   480
by (REPEAT (rtac allI 1));
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   481
by (simp_tac (!simpset addsimps [MGUnifier_def,Unifier_def]) 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   482
by (my_strip_tac 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   483
by (rtac MoreGen_Nil 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   484
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   485
(*3*)
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   486
by (rtac conjI 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   487
by (my_strip_tac 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   488
by (rtac (mgu_sym RS iffD1) 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   489
by (rtac MGUnifier_Var 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   490
by (Simp_tac 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   491
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   492
(*4*)
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   493
by (rtac conjI 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   494
by (my_strip_tac 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   495
by (rtac MGUnifier_Var 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   496
by (assume_tac 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   497
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   498
(*6*)
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   499
by (rtac conjI 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   500
by (rewrite_tac NNF_rews);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   501
by (my_strip_tac 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   502
by (rtac (mgu_sym RS iffD1) 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   503
by (rtac MGUnifier_Var 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   504
by (Asm_simp_tac 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   505
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   506
(*7*)
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   507
by (safe_tac HOL_cs);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   508
by (Subst_case_tac [("v","Unify(M1, M2)")]);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   509
by (REPEAT (eres_inst_tac[("x","list")] allE 1));
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   510
by (Subst_case_tac [("v","Unify(N1 <| list, N2 <| list)")]);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   511
by (eres_inst_tac[("x","lista")] allE 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   512
by (Asm_full_simp_tac 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   513
by (hyp_subst_tac 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   514
by prune_params_tac;
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   515
by (rename_tac "M1 N1 M2 N2 theta sigma" 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   516
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   517
by (asm_full_simp_tac(HOL_ss addsimps [MGUnifier_def,Unifier_def])1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   518
by (rtac conjI 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   519
by (asm_simp_tac (!simpset addsimps [subst_comp]) 1); (* It's a unifier.*)
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   520
by (Simp_tac 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   521
by (safe_tac HOL_cs);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   522
by (rename_tac "M1 N1 M2 N2 theta sigma gamma" 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   523
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   524
by (rewrite_tac [MoreGeneral_def]);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   525
by (eres_inst_tac [("x","gamma")] allE 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   526
by (Asm_full_simp_tac 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   527
by (etac exE 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   528
by (rename_tac "M1 N1 M2 N2 theta sigma gamma theta1" 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   529
by (eres_inst_tac [("x","theta1")] allE 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   530
by (subgoal_tac "N1 <| theta <| theta1 = N2 <| theta <| theta1" 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   531
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   532
by (simp_tac (HOL_ss addsimps [subst_comp RS sym]) 2);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   533
by (rtac subst_subst2 2);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   534
by (assume_tac 3);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   535
by (assume_tac 2);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   536
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   537
by (dtac mp 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   538
by (assume_tac 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   539
by (etac exE 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   540
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   541
by (rename_tac "M1 N1 M2 N2 theta sigma gamma theta1 sigma1" 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   542
by (res_inst_tac [("x","sigma1")] exI 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   543
by (rtac subst_trans 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   544
by (assume_tac 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   545
by (rtac subst_trans 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   546
by (rtac subst_sym 2);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   547
by (rtac comp_assoc 2);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   548
by (rtac subst_cong 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   549
by (assume_tac 2);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   550
by (simp_tac (HOL_ss addsimps [subst_eq_def]) 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   551
val Unify_gives_MGU = standard(result() RS spec RS mp);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   552
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   553
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   554
(*---------------------------------------------------------------------------
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   555
 * Unify returns idempotent substitutions, when it succeeds.
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   556
 *---------------------------------------------------------------------------*)
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   557
goal theory "!theta. Unify (P,Q) = Subst theta --> Idem theta";
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   558
by (res_inst_tac [("xa","P"),("x","Q")] 
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   559
                 (standard (Unify_induction RS mp RS spec RS spec)) 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   560
by (simp_tac (!simpset addsimps [Unify_rules] 
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   561
                       setloop (split_tac [expand_if])) 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   562
(*1*)
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   563
by (rtac conjI 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   564
by (my_strip_tac 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   565
by (rtac Idem_Nil 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   566
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   567
(*3*)
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   568
by (rtac conjI 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   569
by (my_strip_tac 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   570
by (rtac Var_Idem 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   571
by (Simp_tac 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   572
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   573
(*4*)
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   574
by (rtac conjI 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   575
by (my_strip_tac 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   576
by (rtac Var_Idem 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   577
by (atac 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   578
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   579
(*6*)
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   580
by (rtac conjI 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   581
by (rewrite_tac [symmetric (occs_Comb RS eq_reflection)]);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   582
by (my_strip_tac 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   583
by (rtac Var_Idem 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   584
by (atac 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   585
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   586
(*7*)
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   587
by (safe_tac HOL_cs);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   588
by (Subst_case_tac [("v","Unify(M1, M2)")]);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   589
by (REPEAT (eres_inst_tac[("x","list")] allE 1));
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   590
by (Subst_case_tac [("v","Unify(N1 <| list, N2 <| list)")]);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   591
by (eres_inst_tac[("x","lista")] allE 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   592
by (Asm_full_simp_tac 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   593
by (hyp_subst_tac 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   594
by prune_params_tac;
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   595
by (rename_tac "M1 N1 M2 N2 theta sigma" 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   596
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   597
by (dtac Unify_gives_MGU 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   598
by (dtac Unify_gives_MGU 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   599
by (rewrite_tac [MGUnifier_def]);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   600
by (my_strip_tac 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   601
by (rtac Idem_comp 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   602
by (atac 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   603
by (atac 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   604
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   605
by (rewrite_tac [MGUnifier_def]);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   606
by (my_strip_tac 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   607
by (eres_inst_tac [("x","q")] allE 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   608
by (Asm_full_simp_tac 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   609
by (rewrite_tac [MoreGeneral_def]);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   610
by (my_strip_tac 1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   611
by (asm_full_simp_tac(termin_ss addsimps [subst_eq_iff,subst_comp,Idem_def])1);
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   612
val Unify_gives_Idem = result();
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   613
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   614
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   615
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   616
(*---------------------------------------------------------------------------
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   617
 * Exercise. The given algorithm is a bit inelegant. What about the
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   618
 * following "improvement", which adds a few recursive calls in former
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   619
 * base cases? It seems that the termination relation needs another
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   620
 * case in the lexico. product.
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   621
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   622
val {theory,induction,rules,tcs,typechecks} =
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   623
Rfunc Unify.thy ??
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   624
  `(Unify(Const m, Const n)  = (if (m=n) then Subst[] else Fail))    &
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   625
   (Unify(Const m, Comb M N) = Fail)                                 &
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   626
   (Unify(Const m, Var v)    = Unify(Var v, Const m))                &
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   627
   (Unify(Var v, M) = (if (Var v <: M) then Fail else Subst[(v,M)])) &
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   628
   (Unify(Comb M N, Const x) = Fail)                                 &
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   629
   (Unify(Comb M N, Var v) = Unify(Var v, Comb M N))                 &
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   630
   (Unify(Comb M1 N1, Comb M2 N2) = 
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   631
      (case Unify(M1,M2)
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   632
        of Fail => Fail
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   633
         | Subst theta => (case Unify(N1 <| theta, N2 <| theta)
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   634
                            of Fail => Fail
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   635
                             | Subst sigma => Subst (theta <> sigma))))`;
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   636
21266526ac42 Subst as modified by Konrad Slind
paulson
parents:
diff changeset
   637
 *---------------------------------------------------------------------------*)