| author | nipkow | 
| Mon, 02 Nov 1998 18:02:53 +0100 | |
| changeset 5789 | 7d4ac02677a6 | 
| parent 5278 | a903b66822e2 | 
| permissions | -rw-r--r-- | 
| 3268 
012c43174664
Mostly cosmetic changes: updated headers, ID lines, etc.
 paulson parents: 
3192diff
changeset | 1 | (* Title: HOL/Subst/Unifier.ML | 
| 1266 | 2 | ID: $Id$ | 
| 1465 | 3 | Author: Martin Coen, Cambridge University Computer Laboratory | 
| 968 | 4 | Copyright 1993 University of Cambridge | 
| 5 | ||
| 6 | Properties of unifiers. | |
| 7 | *) | |
| 8 | ||
| 9 | open Unifier; | |
| 10 | ||
| 3192 
a75558a4ed37
New version, modified by Konrad Slind and LCP for TFL
 paulson parents: 
2088diff
changeset | 11 | val unify_defs = [Unifier_def, MoreGeneral_def, MGUnifier_def]; | 
| 
a75558a4ed37
New version, modified by Konrad Slind and LCP for TFL
 paulson parents: 
2088diff
changeset | 12 | |
| 
a75558a4ed37
New version, modified by Konrad Slind and LCP for TFL
 paulson parents: 
2088diff
changeset | 13 | (*--------------------------------------------------------------------------- | 
| 
a75558a4ed37
New version, modified by Konrad Slind and LCP for TFL
 paulson parents: 
2088diff
changeset | 14 | * Unifiers | 
| 
a75558a4ed37
New version, modified by Konrad Slind and LCP for TFL
 paulson parents: 
2088diff
changeset | 15 | *---------------------------------------------------------------------------*) | 
| 968 | 16 | |
| 5278 | 17 | Goal "Unifier s (Comb t u) (Comb v w) = (Unifier s t v & Unifier s u w)"; | 
| 4089 | 18 | by (simp_tac (simpset() addsimps [Unifier_def]) 1); | 
| 3192 
a75558a4ed37
New version, modified by Konrad Slind and LCP for TFL
 paulson parents: 
2088diff
changeset | 19 | qed "Unifier_Comb"; | 
| 968 | 20 | |
| 3192 
a75558a4ed37
New version, modified by Konrad Slind and LCP for TFL
 paulson parents: 
2088diff
changeset | 21 | AddIffs [Unifier_Comb]; | 
| 968 | 22 | |
| 5278 | 23 | Goal "[| v ~: vars_of t; v ~: vars_of u; Unifier s t u |] ==> \ | 
| 5119 | 24 | \ Unifier ((v,r)#s) t u"; | 
| 4089 | 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: 
2088diff
changeset | 26 | qed "Cons_Unifier"; | 
| 
a75558a4ed37
New version, modified by Konrad Slind and LCP for TFL
 paulson parents: 
2088diff
changeset | 27 | |
| 968 | 28 | |
| 3192 
a75558a4ed37
New version, modified by Konrad Slind and LCP for TFL
 paulson parents: 
2088diff
changeset | 29 | (*--------------------------------------------------------------------------- | 
| 
a75558a4ed37
New version, modified by Konrad Slind and LCP for TFL
 paulson parents: 
2088diff
changeset | 30 | * Most General Unifiers | 
| 
a75558a4ed37
New version, modified by Konrad Slind and LCP for TFL
 paulson parents: 
2088diff
changeset | 31 | *---------------------------------------------------------------------------*) | 
| 968 | 32 | |
| 5069 | 33 | Goalw unify_defs "MGUnifier s t u = MGUnifier s u t"; | 
| 4089 | 34 | by (blast_tac (claset() addIs [sym]) 1); | 
| 3192 
a75558a4ed37
New version, modified by Konrad Slind and LCP for TFL
 paulson parents: 
2088diff
changeset | 35 | qed "mgu_sym"; | 
| 968 | 36 | |
| 37 | ||
| 5069 | 38 | Goal "[] >> s"; | 
| 4089 | 39 | by (simp_tac (simpset() addsimps [MoreGeneral_def]) 1); | 
| 3192 
a75558a4ed37
New version, modified by Konrad Slind and LCP for TFL
 paulson parents: 
2088diff
changeset | 40 | by (Blast_tac 1); | 
| 968 | 41 | qed "MoreGen_Nil"; | 
| 42 | ||
| 3192 
a75558a4ed37
New version, modified by Konrad Slind and LCP for TFL
 paulson parents: 
2088diff
changeset | 43 | AddIffs [MoreGen_Nil]; | 
| 968 | 44 | |
| 5069 | 45 | Goalw unify_defs | 
| 3192 
a75558a4ed37
New version, modified by Konrad Slind and LCP for TFL
 paulson parents: 
2088diff
changeset | 46 | "MGUnifier s t u = (ALL r. Unifier r t u = s >> r)"; | 
| 4089 | 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: 
2088diff
changeset | 48 | qed "MGU_iff"; | 
| 2088 | 49 | |
| 50 | ||
| 5278 | 51 | Goal "~ Var v <: t ==> MGUnifier [(v,t)] (Var v) t"; | 
| 4089 | 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: 
2088diff
changeset | 53 | delsimps [subst_Var]) 1); | 
| 3724 | 54 | by Safe_tac; | 
| 3457 | 55 | by (rtac exI 1); | 
| 3192 
a75558a4ed37
New version, modified by Konrad Slind and LCP for TFL
 paulson parents: 
2088diff
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: 
2088diff
changeset | 57 | by (etac ssubst_subst2 1); | 
| 4089 | 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: 
2088diff
changeset | 59 | qed "MGUnifier_Var"; | 
| 968 | 60 | |
| 3192 
a75558a4ed37
New version, modified by Konrad Slind and LCP for TFL
 paulson parents: 
2088diff
changeset | 61 | AddSIs [MGUnifier_Var]; | 
| 968 | 62 | |
| 63 | ||
| 3192 
a75558a4ed37
New version, modified by Konrad Slind and LCP for TFL
 paulson parents: 
2088diff
changeset | 64 | (*--------------------------------------------------------------------------- | 
| 
a75558a4ed37
New version, modified by Konrad Slind and LCP for TFL
 paulson parents: 
2088diff
changeset | 65 | * Idempotence. | 
| 
a75558a4ed37
New version, modified by Konrad Slind and LCP for TFL
 paulson parents: 
2088diff
changeset | 66 | *---------------------------------------------------------------------------*) | 
| 968 | 67 | |
| 5069 | 68 | Goalw [Idem_def] "Idem([])"; | 
| 3192 
a75558a4ed37
New version, modified by Konrad Slind and LCP for TFL
 paulson parents: 
2088diff
changeset | 69 | by (Simp_tac 1); | 
| 
a75558a4ed37
New version, modified by Konrad Slind and LCP for TFL
 paulson parents: 
2088diff
changeset | 70 | qed "Idem_Nil"; | 
| 
a75558a4ed37
New version, modified by Konrad Slind and LCP for TFL
 paulson parents: 
2088diff
changeset | 71 | |
| 
a75558a4ed37
New version, modified by Konrad Slind and LCP for TFL
 paulson parents: 
2088diff
changeset | 72 | AddIffs [Idem_Nil]; | 
| 
a75558a4ed37
New version, modified by Konrad Slind and LCP for TFL
 paulson parents: 
2088diff
changeset | 73 | |
| 5069 | 74 | Goalw [Idem_def] "Idem(s) = (sdom(s) Int srange(s) = {})";
 | 
| 4089 | 75 | by (simp_tac (simpset() addsimps [subst_eq_iff, invariance, | 
| 3192 
a75558a4ed37
New version, modified by Konrad Slind and LCP for TFL
 paulson parents: 
2088diff
changeset | 76 | dom_range_disjoint]) 1); | 
| 
a75558a4ed37
New version, modified by Konrad Slind and LCP for TFL
 paulson parents: 
2088diff
changeset | 77 | qed "Idem_iff"; | 
| 968 | 78 | |
| 5069 | 79 | Goal "~ (Var(v) <: t) --> Idem([(v,t)])"; | 
| 4089 | 80 | by (simp_tac (simpset() addsimps [vars_iff_occseq, Idem_iff, srange_iff, | 
| 4719 | 81 | empty_iff_all_not]) 1); | 
| 3192 
a75558a4ed37
New version, modified by Konrad Slind and LCP for TFL
 paulson parents: 
2088diff
changeset | 82 | qed_spec_mp "Var_Idem"; | 
| 
a75558a4ed37
New version, modified by Konrad Slind and LCP for TFL
 paulson parents: 
2088diff
changeset | 83 | |
| 
a75558a4ed37
New version, modified by Konrad Slind and LCP for TFL
 paulson parents: 
2088diff
changeset | 84 | AddSIs [Var_Idem]; | 
| 968 | 85 | |
| 5069 | 86 | Goalw [Idem_def] | 
| 5119 | 87 | "[| Idem(r); Unifier s (t<|r) (u<|r) |] \ | 
| 88 | \ ==> Unifier (r <> s) (t <| r) (u <| r)"; | |
| 4089 | 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: 
2088diff
changeset | 90 | qed "Unifier_Idem_subst"; | 
| 
a75558a4ed37
New version, modified by Konrad Slind and LCP for TFL
 paulson parents: 
2088diff
changeset | 91 | |
| 
a75558a4ed37
New version, modified by Konrad Slind and LCP for TFL
 paulson parents: 
2088diff
changeset | 92 | val [idemr,unifier,minor] = goal Unifier.thy | 
| 
a75558a4ed37
New version, modified by Konrad Slind and LCP for TFL
 paulson parents: 
2088diff
changeset | 93 | "[| Idem(r); Unifier s (t <| r) (u <| r); \ | 
| 
a75558a4ed37
New version, modified by Konrad Slind and LCP for TFL
 paulson parents: 
2088diff
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: 
2088diff
changeset | 95 | \ |] ==> Idem(r <> s)"; | 
| 
a75558a4ed37
New version, modified by Konrad Slind and LCP for TFL
 paulson parents: 
2088diff
changeset | 96 | by (cut_facts_tac [idemr, | 
| 
a75558a4ed37
New version, modified by Konrad Slind and LCP for TFL
 paulson parents: 
2088diff
changeset | 97 | unifier RS (idemr RS Unifier_Idem_subst RS minor)] 1); | 
| 4089 | 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: 
2088diff
changeset | 99 | qed "Idem_comp"; |