src/HOL/Subst/Unifier.ML
author paulson
Thu, 10 Oct 1996 11:59:01 +0200
changeset 2088 e814e03bbbb2
parent 1673 d22110ddd0af
child 3192 a75558a4ed37
permissions -rw-r--r--
Added comments describing better proofs
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1266
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
For unifier.thy.
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
Cases for partial correctness of algorithm and conditions for termination.
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
     9
*)
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    10
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    11
open Unifier;
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    12
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    13
val unify_defs =
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    14
    [Idem_def,Unifier_def,MoreGeneral_def,MGUnifier_def,MGIUnifier_def];
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    15
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    16
(**** Unifiers ****)
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    17
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    18
goalw Unifier.thy [Unifier_def] "Unifier s t u = (t <| s = u <| s)";
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    19
by (rtac refl 1);
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    20
qed "Unifier_iff";
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    21
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    22
goal Unifier.thy
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    23
    "Unifier s (Comb t u) (Comb v w) --> Unifier s t v & Unifier s u w";
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    24
by (simp_tac (subst_ss addsimps [Unifier_iff]) 1);
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    25
val Unifier_Comb  = store_thm("Unifier_Comb", result() RS mp RS conjE);
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    26
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    27
goal Unifier.thy
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    28
  "~v : vars_of(t) --> ~v : vars_of(u) -->Unifier s t u --> \
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 968
diff changeset
    29
\  Unifier ((v,r)#s) t u";
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    30
by (simp_tac (subst_ss addsimps [Unifier_iff,repl_invariance]) 1);
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    31
val Cons_Unifier  = store_thm("Cons_Unifier", result() RS mp RS mp RS mp);
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    32
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    33
(**** Most General Unifiers ****)
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    34
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    35
goalw Unifier.thy [MoreGeneral_def]  "r >> s = (EX q. s =s= r <> q)";
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    36
by (rtac refl 1);
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    37
qed "MoreGen_iff";
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    38
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    39
goal Unifier.thy  "[] >> s";
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    40
by (simp_tac (subst_ss addsimps [MoreGen_iff]) 1);
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    41
by (fast_tac (set_cs addIs [refl RS subst_refl]) 1);
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    42
qed "MoreGen_Nil";
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    43
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    44
goalw Unifier.thy unify_defs
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    45
    "MGUnifier s t u = (ALL r.Unifier r t u = s >> r)";
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    46
by (REPEAT (ares_tac [iffI,allI] 1 ORELSE 
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    47
            eresolve_tac [conjE,allE,mp,exE,ssubst_subst2] 1));
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    48
by (asm_simp_tac (subst_ss addsimps [subst_comp]) 1);
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    49
by (fast_tac (set_cs addIs [comp_Nil RS sym RS subst_refl]) 1);
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    50
qed "MGU_iff";
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    51
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    52
val [prem] = goal Unifier.thy
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 968
diff changeset
    53
     "~ Var(v) <: t ==> MGUnifier [(v,t)] (Var v) t";
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    54
by (simp_tac (subst_ss addsimps [MGU_iff,MoreGen_iff,Unifier_iff]) 1);
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    55
by (REPEAT_SOME (step_tac set_cs));
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    56
by (etac subst 1);
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    57
by (etac ssubst_subst2 2);
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    58
by (rtac (Cons_trivial RS subst_sym) 1);
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    59
by (simp_tac (subst_ss addsimps [prem RS Var_not_occs,Var_subst]) 1);
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    60
qed "MGUnifier_Var";
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    61
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    62
(**** Most General Idempotent Unifiers ****)
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    63
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    64
goal Unifier.thy "r <> r =s= r --> s =s= r <> q --> r <> s =s= s";
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    65
by (simp_tac (subst_ss addsimps [subst_eq_iff,subst_comp]) 1);
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    66
val MGIU_iff_lemma  = store_thm("MGIU_iff_lemma", result() RS mp RS mp);
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    67
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    68
goalw Unifier.thy unify_defs
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    69
 "MGIUnifier s t u = \
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    70
\  (Idem(s) & Unifier s t u & (ALL r.Unifier r t u --> s<>r=s=r))";
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    71
by (fast_tac (set_cs addEs [subst_sym,MGIU_iff_lemma]) 1);
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    72
qed "MGIU_iff";
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    73
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    74
(**** Idempotence ****)
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    75
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    76
goalw Unifier.thy unify_defs "Idem(s) = (s <> s =s= s)";
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    77
by (rtac refl 1);
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    78
qed "raw_Idem_iff";
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    79
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    80
goal Unifier.thy "Idem(s) = (sdom(s) Int srange(s) = {})";
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    81
by (simp_tac (subst_ss addsimps [raw_Idem_iff,subst_eq_iff,subst_comp,
1673
d22110ddd0af repaired critical proofs depending on the order inside non-confluent SimpSets
oheimb
parents: 1465
diff changeset
    82
                                invariance,dom_range_disjoint]
d22110ddd0af repaired critical proofs depending on the order inside non-confluent SimpSets
oheimb
parents: 1465
diff changeset
    83
		       delsimps (empty_iff::mem_simps)) 1);
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    84
qed "Idem_iff";
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    85
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    86
goal Unifier.thy "Idem([])";
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    87
by (simp_tac (subst_ss addsimps [raw_Idem_iff,refl RS subst_refl]) 1);
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    88
qed "Idem_Nil";
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    89
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 968
diff changeset
    90
goal Unifier.thy "~ (Var(v) <: t) --> Idem([(v,t)])";
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    91
by (simp_tac (subst_ss addsimps [Var_subst,vars_iff_occseq,Idem_iff,srange_iff]
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    92
                       setloop (split_tac [expand_if])) 1);
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    93
by (fast_tac set_cs 1);
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    94
val Var_Idem  = store_thm("Var_Idem", result() RS mp);
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    95
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    96
val [prem] = goalw Unifier.thy [Idem_def]
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    97
     "Idem(r) ==>  Unifier s (t <| r) (u <| r) --> Unifier (r <> s) (t <| r) (u <| r)";
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    98
by (simp_tac (subst_ss addsimps 
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1266
diff changeset
    99
              [Unifier_iff,subst_comp,prem RS comp_subst_subst]) 1);
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   100
val Unifier_Idem_subst  = store_thm("Unifier_Idem_subst", result() RS mp);
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   101
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   102
val [prem] = goal Unifier.thy 
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   103
     "r <> s =s= s ==>  Unifier s t u --> Unifier s (t <| r) (u <| r)";
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   104
by (simp_tac (subst_ss addsimps 
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1266
diff changeset
   105
              [Unifier_iff,subst_comp,prem RS comp_subst_subst]) 1);
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   106
val Unifier_comp_subst  = store_thm("Unifier_comp_subst", result() RS mp);
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   107
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   108
(*** The domain of a MGIU is a subset of the variables in the terms ***)
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   109
(***      NB this and one for range are only needed for termination ***)
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   110
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   111
val [prem] = goal Unifier.thy
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   112
    "~ vars_of(Var(x) <| r) = vars_of(Var(x) <| s) ==> ~r =s= s";
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   113
by (rtac (prem RS contrapos) 1);
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   114
by (fast_tac (set_cs addEs [subst_subst2]) 1);
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   115
qed "lemma_lemma";
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   116
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   117
val prems = goal Unifier.thy 
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   118
    "x : sdom(s) -->  ~x : srange(s) --> \
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 968
diff changeset
   119
\   ~vars_of(Var(x) <| s<> (x,Var(x))#s) = \
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 968
diff changeset
   120
\      vars_of(Var(x) <| (x,Var(x))#s)";
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   121
by (simp_tac (subst_ss addsimps [not_equal_iff]) 1);
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   122
by (REPEAT (resolve_tac [impI,disjI2] 1));
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   123
by(res_inst_tac [("x","x")] exI 1);
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1266
diff changeset
   124
by (rtac conjI 1);
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   125
by (asm_simp_tac (subst_ss addsimps [Var_elim,subst_comp,repl_invariance]) 1);
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   126
by (asm_simp_tac (subst_ss addsimps [Var_subst]) 1);
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   127
val MGIU_sdom_lemma = store_thm("MGIU_sdom_lemma", result() RS mp RS mp RS lemma_lemma RS notE);
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   128
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   129
goal Unifier.thy "MGIUnifier s t u --> sdom(s) <= vars_of(t) Un vars_of(u)";
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   130
by (subgoal_tac "! P Q.(P | Q) = (~( ~P & ~Q))" 1);
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   131
by (asm_simp_tac (subst_ss addsimps [MGIU_iff,Idem_iff,subset_iff]) 1);
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   132
by (safe_tac set_cs);
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   133
by (eresolve_tac ([spec] RL [impE]) 1);
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   134
by (rtac Cons_Unifier 1);
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   135
by (ALLGOALS (fast_tac (set_cs addIs [Cons_Unifier,MGIU_sdom_lemma])));
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   136
val MGIU_sdom  = store_thm("MGIU_sdom", result() RS mp);
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   137
2088
e814e03bbbb2 Added comments describing better proofs
paulson
parents: 1673
diff changeset
   138
e814e03bbbb2 Added comments describing better proofs
paulson
parents: 1673
diff changeset
   139
(** COULD REPLACE THE TWO THEOREMS ABOVE BY THE FOLLOWING, IN THE PRESENCE
e814e03bbbb2 Added comments describing better proofs
paulson
parents: 1673
diff changeset
   140
    OF DEMORGAN LAWS.  DON'T KNOW WHAT TO DO WITH THE SIMILAR PROOF BELOW.
e814e03bbbb2 Added comments describing better proofs
paulson
parents: 1673
diff changeset
   141
val prems = goal Unifier.thy 
e814e03bbbb2 Added comments describing better proofs
paulson
parents: 1673
diff changeset
   142
    "x : sdom(s) -->  ~x : srange(s) --> \
e814e03bbbb2 Added comments describing better proofs
paulson
parents: 1673
diff changeset
   143
\   ~vars_of(Var(x) <| s<> (x,Var(x))#s) = \
e814e03bbbb2 Added comments describing better proofs
paulson
parents: 1673
diff changeset
   144
\      vars_of(Var(x) <| (x,Var(x))#s)";
e814e03bbbb2 Added comments describing better proofs
paulson
parents: 1673
diff changeset
   145
by (simp_tac (subst_ss addsimps [not_equal_iff]) 1);
e814e03bbbb2 Added comments describing better proofs
paulson
parents: 1673
diff changeset
   146
by (REPEAT (resolve_tac [impI,disjI2] 1));
e814e03bbbb2 Added comments describing better proofs
paulson
parents: 1673
diff changeset
   147
by(res_inst_tac [("x","x")] exI 1);
e814e03bbbb2 Added comments describing better proofs
paulson
parents: 1673
diff changeset
   148
by (rtac conjI 1);
e814e03bbbb2 Added comments describing better proofs
paulson
parents: 1673
diff changeset
   149
by (asm_simp_tac (subst_ss addsimps [Var_elim,subst_comp,repl_invariance]) 1);
e814e03bbbb2 Added comments describing better proofs
paulson
parents: 1673
diff changeset
   150
by (asm_simp_tac (subst_ss addsimps [Var_subst]) 1);
e814e03bbbb2 Added comments describing better proofs
paulson
parents: 1673
diff changeset
   151
val MGIU_sdom_lemma = result() RS mp RS mp RS lemma_lemma RSN (2, rev_notE);
e814e03bbbb2 Added comments describing better proofs
paulson
parents: 1673
diff changeset
   152
e814e03bbbb2 Added comments describing better proofs
paulson
parents: 1673
diff changeset
   153
goal Unifier.thy "MGIUnifier s t u --> sdom(s) <= vars_of(t) Un vars_of(u)";
e814e03bbbb2 Added comments describing better proofs
paulson
parents: 1673
diff changeset
   154
by (asm_simp_tac (subst_ss addsimps [MGIU_iff,Idem_iff,subset_iff]) 1);
e814e03bbbb2 Added comments describing better proofs
paulson
parents: 1673
diff changeset
   155
by (deepen_tac (set_cs addIs [Cons_Unifier] addEs [MGIU_sdom_lemma]) 3 1);
e814e03bbbb2 Added comments describing better proofs
paulson
parents: 1673
diff changeset
   156
val MGIU_sdom  = store_thm("MGIU_sdom", result() RS mp);
e814e03bbbb2 Added comments describing better proofs
paulson
parents: 1673
diff changeset
   157
**)
e814e03bbbb2 Added comments describing better proofs
paulson
parents: 1673
diff changeset
   158
e814e03bbbb2 Added comments describing better proofs
paulson
parents: 1673
diff changeset
   159
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   160
(*** The range of a MGIU is a subset of the variables in the terms ***)
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   161
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   162
val prems = goal HOL.thy  "P = Q ==> (~P) = (~Q)";
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 972
diff changeset
   163
by (simp_tac (subst_ss addsimps prems) 1);
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   164
qed "not_cong";
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   165
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   166
val prems = goal Unifier.thy 
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   167
   "~w=x --> x : vars_of(Var(w) <| s) --> w : sdom(s) --> ~w : srange(s) --> \
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 968
diff changeset
   168
\   ~vars_of(Var(w) <| s<> (x,Var(w))#s) = \
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 968
diff changeset
   169
\   vars_of(Var(w) <| (x,Var(w))#s)";
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   170
by (simp_tac (subst_ss addsimps [not_equal_iff]) 1);
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   171
by (REPEAT (resolve_tac [impI,disjI1] 1));
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   172
by(res_inst_tac [("x","w")] exI 1);
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   173
by (ALLGOALS (asm_simp_tac (subst_ss addsimps  [Var_elim,subst_comp,
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   174
                vars_var_iff RS not_cong RS iffD2 RS repl_invariance]) ));
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   175
by (fast_tac (set_cs addIs [Var_in_subst]) 1);
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   176
val MGIU_srange_lemma  = store_thm("MGIU_srange_lemma", result() RS mp RS mp RS mp RS mp RS lemma_lemma RS notE);
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   177
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   178
goal Unifier.thy "MGIUnifier s t u --> srange(s) <= vars_of(t) Un vars_of(u)";
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   179
by (subgoal_tac "! P Q.(P | Q) = (~( ~P & ~Q))" 1);
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   180
by (asm_simp_tac (subst_ss addsimps [MGIU_iff,srange_iff,subset_iff]) 1);
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   181
by (simp_tac (subst_ss addsimps [Idem_iff]) 1);
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   182
by (safe_tac set_cs);
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   183
by (eresolve_tac ([spec] RL [impE]) 1);
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   184
by (rtac Cons_Unifier 1);
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   185
by (imp_excluded_middle_tac "w=ta" 4);
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   186
by (fast_tac (set_cs addEs [MGIU_srange_lemma]) 5);
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   187
by (ALLGOALS (fast_tac (set_cs addIs [Var_elim2])));
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   188
val MGIU_srange = store_thm("MGIU_srange", result() RS mp);
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   189
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   190
(*************** Correctness of a simple unification algorithm ***************)
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   191
(*                                                                           *)
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   192
(*  fun unify Const(m) Const(n) = if m=n then Nil else Fail                  *)
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   193
(*    | unify Const(m) _        = Fail                                       *)
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 968
diff changeset
   194
(*    | unify Var(v)   t        = if Var(v)<:t then Fail else (v,t)#Nil      *)
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   195
(*    | unify Comb(t,u) Const(n) = Fail                                      *)
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   196
(*    | unify Comb(t,u) Var(v)  = if Var(v) <: Comb(t,u) then Fail           *)
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 968
diff changeset
   197
(*                                               else (v,Comb(t,u)#Nil       *)
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   198
(*    | unify Comb(t,u) Comb(v,w) = let s = unify t v                        *)
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   199
(*                                  in if s=Fail then Fail                   *)
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   200
(*                                               else unify (u<|s) (w<|s);   *)
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   201
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   202
(**** Cases for the partial correctness of the algorithm ****)
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   203
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   204
goalw Unifier.thy unify_defs "MGIUnifier s t u = MGIUnifier s u t";
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   205
by (safe_tac (HOL_cs addSEs ([sym]@([spec] RL [mp]))));
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   206
qed "Unify_comm";
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   207
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   208
goal Unifier.thy "MGIUnifier [] (Const n) (Const n)";
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   209
by (simp_tac (subst_ss addsimps
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1266
diff changeset
   210
              [MGIU_iff,MGU_iff,Unifier_iff,subst_eq_iff,Idem_Nil]) 1);
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   211
qed "Unify1";
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   212
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   213
goal Unifier.thy "~m=n --> (ALL l.~Unifier l (Const m) (Const n))";
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   214
by (simp_tac (subst_ss addsimps[Unifier_iff]) 1);
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   215
val Unify2 = store_thm("Unify2", result() RS mp);
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   216
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   217
val [prem] = goalw Unifier.thy [MGIUnifier_def] 
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 968
diff changeset
   218
 "~Var(v) <: t ==> MGIUnifier [(v,t)] (Var v) t";
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   219
by (fast_tac (HOL_cs addSIs [prem RS MGUnifier_Var,prem RS Var_Idem]) 1);
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   220
qed "Unify3";
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   221
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   222
val [prem] = goal Unifier.thy "Var(v) <: t ==> (ALL l.~Unifier l (Var v) t)";
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   223
by (simp_tac (subst_ss addsimps
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1266
diff changeset
   224
              [Unifier_iff,prem RS subst_mono RS occs_irrefl2]) 1);
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   225
qed "Unify4";
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   226
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   227
goal Unifier.thy "ALL l.~Unifier l (Const m) (Comb t u)";
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   228
by (simp_tac (subst_ss addsimps [Unifier_iff]) 1);
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   229
qed "Unify5";
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   230
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   231
goal Unifier.thy
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   232
    "(ALL l.~Unifier l t v) --> (ALL l.~Unifier l (Comb t u) (Comb v w))";
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   233
by (simp_tac (subst_ss addsimps [Unifier_iff]) 1);
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   234
val Unify6 = store_thm("Unify6", result() RS mp);
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   235
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   236
goal Unifier.thy "MGIUnifier s t v --> (ALL l.~Unifier l (u <| s) (w <| s)) \
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   237
\                                --> (ALL l.~Unifier l (Comb t u) (Comb v w))";
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   238
by (simp_tac (subst_ss addsimps [MGIU_iff]) 1);
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   239
by (fast_tac (set_cs addIs [Unifier_comp_subst] addSEs [Unifier_Comb]) 1);
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   240
val Unify7 = store_thm("Unify7", result() RS mp RS mp);
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   241
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   242
val [p1,p2,p3] = goal Unifier.thy
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   243
     "[| Idem(r); Unifier s (t <| r) (u <| r); \
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   244
\     (! q.Unifier q (t <| r) (u <| r) --> s <> q =s= q) |] ==> \
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   245
\     Idem(r <> s)";
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   246
by (cut_facts_tac [p1,
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1266
diff changeset
   247
                   p2 RS (p1 RS Unifier_Idem_subst RS (p3 RS spec RS mp))] 1);
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   248
by (REPEAT_SOME (etac rev_mp));
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   249
by (simp_tac (subst_ss addsimps [raw_Idem_iff,subst_eq_iff,subst_comp]) 1);
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   250
qed "Unify8_lemma1";
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   251
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   252
val [p1,p2,p3,p4] = goal Unifier.thy
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   253
   "[| Unifier q t v; Unifier q u w; (! q.Unifier q t v --> r <> q =s= q); \
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   254
\      (! q.Unifier q (u <| r) (w <| r) --> s <> q =s= q) |] ==> \
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   255
\   r <> s <> q =s= q";
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   256
val pp = p1 RS (p3 RS spec RS mp);
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   257
by (cut_facts_tac [pp,
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1266
diff changeset
   258
                   p2 RS (pp RS Unifier_comp_subst) RS (p4 RS spec RS mp)] 1);
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   259
by (REPEAT_SOME (etac rev_mp));
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   260
by (simp_tac (subst_ss addsimps [subst_eq_iff,subst_comp]) 1);
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   261
qed "Unify8_lemma2";
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   262
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   263
goal Unifier.thy  "MGIUnifier r t v -->  MGIUnifier s (u <| r) (w <| r) --> \
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   264
\                MGIUnifier (r <> s) (Comb t u) (Comb v w)";
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   265
by (simp_tac (subst_ss addsimps [MGIU_iff,subst_comp,comp_assoc]) 1);
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   266
by (safe_tac HOL_cs);
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   267
by (REPEAT (etac rev_mp 2));
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   268
by (simp_tac (subst_ss addsimps 
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1266
diff changeset
   269
              [Unifier_iff,MGIU_iff,subst_comp,comp_assoc]) 2);
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   270
by (ALLGOALS (fast_tac (set_cs addEs 
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1266
diff changeset
   271
                        [Unifier_Comb,Unify8_lemma1,Unify8_lemma2])));
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   272
qed "Unify8";
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   273
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   274
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   275
(********************** Termination of the algorithm *************************)
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   276
(*                                                                           *)
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   277
(*UWFD is a well-founded relation that orders the 2 recursive calls in unify *)
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   278
(*                   NB well-foundedness of UWFD isn't proved                *)
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   279
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   280
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   281
goalw Unifier.thy [UWFD_def]  "UWFD t t' (Comb t u) (Comb t' u')";
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   282
by (simp_tac subst_ss 1);
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   283
by (fast_tac set_cs 1);
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   284
qed "UnifyWFD1";
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   285
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   286
val [prem] = goal Unifier.thy 
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   287
     "MGIUnifier s t t' ==> vars_of(u <| s) Un vars_of(u' <| s) <= \
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   288
\                           vars_of (Comb t u) Un vars_of (Comb t' u')";
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   289
by (subgoal_tac "vars_of(u <| s) Un vars_of(u' <| s) <= \
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   290
\                srange(s) Un vars_of(u) Un srange(s) Un vars_of(u')" 1);
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   291
by (etac subset_trans 1);
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   292
by (ALLGOALS (simp_tac (subst_ss addsimps [Var_intro,subset_iff])));
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   293
by (ALLGOALS (fast_tac (set_cs addDs 
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1266
diff changeset
   294
                        [Var_intro,prem RS MGIU_srange RS subsetD])));
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   295
qed "UWFD2_lemma1";
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   296
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   297
val [major,minor] = goal Unifier.thy 
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   298
     "[| MGIUnifier s t t';  ~ u <| s = u |] ==> \
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   299
\     ~ vars_of(u <| s) Un vars_of(u' <| s) =  \
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   300
\         (vars_of(t) Un vars_of(u)) Un (vars_of(t') Un vars_of(u'))";
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   301
by (cut_facts_tac 
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   302
    [major RS (MGIU_iff RS iffD1) RS conjunct1 RS (Idem_iff RS iffD1)] 1);
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   303
by (rtac (minor RS subst_not_empty RS exE) 1);
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   304
by (rtac (make_elim ((major RS MGIU_sdom) RS subsetD)) 1 THEN assume_tac 1);
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   305
by (rtac (disjI2 RS (not_equal_iff RS iffD2)) 1);
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   306
by (REPEAT (etac rev_mp 1));
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   307
by (asm_simp_tac subst_ss 1);
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   308
by (fast_tac (set_cs addIs [Var_elim2]) 1);
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   309
qed "UWFD2_lemma2";
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   310
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   311
val [prem] = goalw Unifier.thy [UWFD_def]  
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   312
  "MGIUnifier s t t' ==> UWFD (u <| s) (u' <| s) (Comb t u) (Comb t' u')";
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   313
by (cut_facts_tac 
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   314
    [prem RS UWFD2_lemma1 RS (subseteq_iff_subset_eq RS iffD1)] 1);
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   315
by (imp_excluded_middle_tac "u <| s = u" 1);
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 972
diff changeset
   316
by (simp_tac (subst_ss delsimps (ssubset_iff :: utlemmas_rews)
3ae9fe3c0f68 added local simpsets
clasohm
parents: 972
diff changeset
   317
                       addsimps [occs_Comb2]) 1);
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   318
by (rtac impI 1 THEN etac subst 1 THEN assume_tac 1);
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   319
by (rtac impI 1);
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   320
by (rtac (conjI RS (ssubset_iff RS iffD2) RS disjI1) 1);
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 972
diff changeset
   321
by (asm_simp_tac (subst_ss delsimps (ssubset_iff :: utlemmas_rews) addsimps [subseteq_iff_subset_eq]) 1);
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   322
by (asm_simp_tac subst_ss 1);
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   323
by (fast_tac (set_cs addDs [prem RS UWFD2_lemma2]) 1);
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   324
qed "UnifyWFD2";