src/HOL/Subst/Unifier.ML
author kleing
Mon, 19 Apr 2004 09:31:00 +0200
changeset 14626 dfb8d2977263
parent 5278 a903b66822e2
permissions -rw-r--r--
renamed HOL-Import-HOL to HOL4, added to images target
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
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
     6
Properties of unifiers.
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
     7
*)
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
     8
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
     9
open Unifier;
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    10
3192
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    11
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
    12
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
 * Unifiers 
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    15
 *---------------------------------------------------------------------------*)
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    16
5278
a903b66822e2 even more tidying of Goal commands
paulson
parents: 5119
diff changeset
    17
Goal "Unifier s (Comb t u) (Comb v w) = (Unifier s t v & Unifier s u w)";
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 3919
diff changeset
    18
by (simp_tac (simpset() addsimps [Unifier_def]) 1);
3192
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    19
qed "Unifier_Comb";
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    20
3192
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    21
AddIffs [Unifier_Comb];
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    22
5278
a903b66822e2 even more tidying of Goal commands
paulson
parents: 5119
diff changeset
    23
Goal "[| v ~: vars_of t; v ~: vars_of u; Unifier s t u |] ==> \
5119
58d267fc8630 Removed leading !! in goals
nipkow
parents: 5069
diff changeset
    24
\  Unifier ((v,r)#s) t u";
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 3919
diff changeset
    25
by (asm_full_simp_tac (simpset() addsimps [Unifier_def, repl_invariance]) 1);
3192
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    26
qed "Cons_Unifier";
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    27
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    28
3192
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    29
(*---------------------------------------------------------------------------
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    30
 * Most General Unifiers 
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    31
 *---------------------------------------------------------------------------*)
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    32
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4719
diff changeset
    33
Goalw unify_defs "MGUnifier s t u = MGUnifier s u t";
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 3919
diff changeset
    34
by (blast_tac (claset() addIs [sym]) 1);
3192
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    35
qed "mgu_sym";
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    36
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    37
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4719
diff changeset
    38
Goal  "[] >> s";
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 3919
diff changeset
    39
by (simp_tac (simpset() addsimps [MoreGeneral_def]) 1);
3192
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    40
by (Blast_tac 1);
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    41
qed "MoreGen_Nil";
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    42
3192
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    43
AddIffs [MoreGen_Nil];
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    44
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4719
diff changeset
    45
Goalw unify_defs
3192
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    46
    "MGUnifier s t u = (ALL r. Unifier r t u = s >> r)";
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 3919
diff changeset
    47
by (auto_tac (claset() addIs [ssubst_subst2, subst_comp_Nil], simpset()));
3192
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    48
qed "MGU_iff";
2088
e814e03bbbb2 Added comments describing better proofs
paulson
parents: 1673
diff changeset
    49
e814e03bbbb2 Added comments describing better proofs
paulson
parents: 1673
diff changeset
    50
5278
a903b66822e2 even more tidying of Goal commands
paulson
parents: 5119
diff changeset
    51
Goal "~ Var v <: t ==> MGUnifier [(v,t)] (Var v) t";
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 3919
diff changeset
    52
by (simp_tac(simpset() addsimps [MGU_iff, Unifier_def, MoreGeneral_def]
3192
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    53
	              delsimps [subst_Var]) 1);
3724
f33e301a89f5 Step_tac -> Safe_tac
paulson
parents: 3457
diff changeset
    54
by Safe_tac;
3457
a8ab7c64817c Ran expandshort
paulson
parents: 3268
diff changeset
    55
by (rtac exI 1);
3192
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    56
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
    57
by (etac ssubst_subst2 1);
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 3919
diff changeset
    58
by (asm_simp_tac (simpset() addsimps [Var_not_occs]) 1);
3192
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    59
qed "MGUnifier_Var";
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    60
3192
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    61
AddSIs [MGUnifier_Var];
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    62
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    63
3192
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    64
(*---------------------------------------------------------------------------
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    65
 * Idempotence.
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    66
 *---------------------------------------------------------------------------*)
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    67
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4719
diff changeset
    68
Goalw [Idem_def] "Idem([])";
3192
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    69
by (Simp_tac 1);
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    70
qed "Idem_Nil";
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    71
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    72
AddIffs [Idem_Nil];
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    73
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4719
diff changeset
    74
Goalw [Idem_def] "Idem(s) = (sdom(s) Int srange(s) = {})";
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 3919
diff changeset
    75
by (simp_tac (simpset() addsimps [subst_eq_iff, invariance, 
3192
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    76
				 dom_range_disjoint]) 1);
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    77
qed "Idem_iff";
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    78
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4719
diff changeset
    79
Goal "~ (Var(v) <: t) --> Idem([(v,t)])";
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 3919
diff changeset
    80
by (simp_tac (simpset() addsimps [vars_iff_occseq, Idem_iff, srange_iff, 
4719
21af5c0be0c9 added not1_or and if_eq_cancel to simpset()
oheimb
parents: 4686
diff changeset
    81
				  empty_iff_all_not]) 1);
3192
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    82
qed_spec_mp "Var_Idem";
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    83
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    84
AddSIs [Var_Idem];
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    85
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4719
diff changeset
    86
Goalw [Idem_def]
5119
58d267fc8630 Removed leading !! in goals
nipkow
parents: 5069
diff changeset
    87
  "[| Idem(r); Unifier s (t<|r) (u<|r) |] \
58d267fc8630 Removed leading !! in goals
nipkow
parents: 5069
diff changeset
    88
\  ==> Unifier (r <> s) (t <| r) (u <| r)";
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 3919
diff changeset
    89
by (asm_full_simp_tac (simpset() addsimps [Unifier_def, comp_subst_subst]) 1);
3192
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    90
qed "Unifier_Idem_subst";
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    91
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    92
val [idemr,unifier,minor] = goal Unifier.thy
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    93
     "[| 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
\        !!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
    95
\     |] ==> Idem(r <> s)";
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    96
by (cut_facts_tac [idemr,
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    97
                   unifier RS (idemr RS Unifier_Idem_subst RS minor)] 1);
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 3919
diff changeset
    98
by (asm_full_simp_tac (simpset() addsimps [Idem_def, subst_eq_iff]) 1);
3192
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2088
diff changeset
    99
qed "Idem_comp";