src/HOL/Subst/Unifier.ML
author paulson
Mon, 23 Jun 1997 10:42:03 +0200
changeset 3457 a8ab7c64817c
parent 3268 012c43174664
child 3724 f33e301a89f5
permissions -rw-r--r--
Ran expandshort
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
3268
012c43174664 Mostly cosmetic changes: updated headers, ID lines, etc.
paulson
parents: 3192
diff changeset
     1
(*  Title:      HOL/Subst/Unifier.ML
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 972
diff changeset
     2
    ID:         $Id$
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1266
diff changeset
     3
    Author:     Martin Coen, Cambridge University Computer Laboratory
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
     4
    Copyright   1993  University of Cambridge
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
     5
3192
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
     6
For Unifier.thy.
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
     7
Properties of unifiers.
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
     8
*)
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
     9
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    10
open Unifier;
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    11
3192
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    12
val unify_defs = [Unifier_def, MoreGeneral_def, MGUnifier_def];
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    13
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    14
(*---------------------------------------------------------------------------
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    15
 * Unifiers 
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    16
 *---------------------------------------------------------------------------*)
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    17
3192
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    18
goal Unifier.thy
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    19
    "Unifier s (Comb t u) (Comb v w) = (Unifier s t v & Unifier s u w)";
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    20
by (simp_tac (!simpset addsimps [Unifier_def]) 1);
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    21
qed "Unifier_Comb";
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    22
3192
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    23
AddIffs [Unifier_Comb];
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    24
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    25
goal Unifier.thy
3192
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    26
  "!!v. [| v ~: vars_of t; v ~: vars_of u; Unifier s t u |] ==> \
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    27
\       Unifier ((v,r)#s) t u";
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    28
by (asm_full_simp_tac (!simpset addsimps [Unifier_def, repl_invariance]) 1);
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    29
qed "Cons_Unifier";
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    30
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    31
3192
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    32
(*---------------------------------------------------------------------------
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    33
 * Most General Unifiers 
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    34
 *---------------------------------------------------------------------------*)
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    35
3192
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    36
goalw Unifier.thy unify_defs "MGUnifier s t u = MGUnifier s u t";
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    37
by (blast_tac (!claset addIs [sym]) 1);
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    38
qed "mgu_sym";
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    39
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    40
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    41
goal Unifier.thy  "[] >> s";
3192
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    42
by (simp_tac (!simpset addsimps [MoreGeneral_def]) 1);
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    43
by (Blast_tac 1);
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    44
qed "MoreGen_Nil";
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    45
3192
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    46
AddIffs [MoreGen_Nil];
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    47
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    48
goalw Unifier.thy unify_defs
3192
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    49
    "MGUnifier s t u = (ALL r. Unifier r t u = s >> r)";
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    50
by (auto_tac (!claset addIs [ssubst_subst2, subst_comp_Nil], !simpset));
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    51
qed "MGU_iff";
2088
e814e03bbbb2 Added comments describing better proofs
paulson
parents: 1673
diff changeset
    52
e814e03bbbb2 Added comments describing better proofs
paulson
parents: 1673
diff changeset
    53
3192
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    54
goal Unifier.thy
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    55
     "!!v. ~ Var v <: t ==> MGUnifier [(v,t)] (Var v) t";
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    56
by (simp_tac(!simpset addsimps [MGU_iff, Unifier_def, MoreGeneral_def]
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    57
	              delsimps [subst_Var]) 1);
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    58
by (Step_tac 1);
3457
a8ab7c64817c Ran expandshort
paulson
parents: 3268
diff changeset
    59
by (rtac exI 1);
3192
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    60
by (etac subst 1 THEN rtac (Cons_trivial RS subst_sym) 1);
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    61
by (etac ssubst_subst2 1);
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    62
by (asm_simp_tac (!simpset addsimps [Var_not_occs]) 1);
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    63
qed "MGUnifier_Var";
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    64
3192
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    65
AddSIs [MGUnifier_Var];
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    66
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    67
3192
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    68
(*---------------------------------------------------------------------------
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    69
 * Idempotence.
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    70
 *---------------------------------------------------------------------------*)
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    71
3192
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    72
goalw Unifier.thy [Idem_def] "Idem([])";
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    73
by (Simp_tac 1);
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    74
qed "Idem_Nil";
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    75
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    76
AddIffs [Idem_Nil];
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    77
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    78
goalw Unifier.thy [Idem_def] "Idem(s) = (sdom(s) Int srange(s) = {})";
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    79
by (simp_tac (!simpset addsimps [subst_eq_iff, invariance, 
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    80
				 dom_range_disjoint]) 1);
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    81
qed "Idem_iff";
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    82
3192
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    83
goal Unifier.thy "~ (Var(v) <: t) --> Idem([(v,t)])";
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    84
by (simp_tac (!simpset addsimps [vars_iff_occseq, Idem_iff, srange_iff, 
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    85
				 empty_iff_all_not]
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    86
                       setloop (split_tac [expand_if])) 1);
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    87
by (Blast_tac 1);
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    88
qed_spec_mp "Var_Idem";
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    89
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    90
AddSIs [Var_Idem];
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    91
3192
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    92
goalw Unifier.thy [Idem_def]
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    93
  "!!r. [| Idem(r); Unifier s (t<|r) (u<|r) |]       \
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    94
\       ==> Unifier (r <> s) (t <| r) (u <| r)";
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    95
by (asm_full_simp_tac (!simpset addsimps [Unifier_def, comp_subst_subst]) 1);
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    96
qed "Unifier_Idem_subst";
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    97
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    98
val [idemr,unifier,minor] = goal Unifier.thy
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    99
     "[| Idem(r);  Unifier s (t <| r) (u <| r); \
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
   100
\        !!q. Unifier q (t <| r) (u <| r) ==> s <> q =$= q \
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
   101
\     |] ==> Idem(r <> s)";
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
   102
by (cut_facts_tac [idemr,
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
   103
                   unifier RS (idemr RS Unifier_Idem_subst RS minor)] 1);
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
   104
by (asm_full_simp_tac (!simpset addsimps [Idem_def, subst_eq_iff]) 1);
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
   105
qed "Idem_comp";