src/HOL/Subst/Subst.ML
author ballarin
Mon, 02 Aug 2004 10:12:02 +0200
changeset 15096 be1d3b8cfbd5
parent 5278 a903b66822e2
permissions -rw-r--r--
Documentation added; minor improvements.
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/Subst.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
3268
012c43174664 Mostly cosmetic changes: updated headers, ID lines, etc.
paulson
parents: 3192
diff changeset
     6
Substitutions on uterms
968
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 Subst;
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    10
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    11
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    12
(**** Substitutions ****)
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    13
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4686
diff changeset
    14
Goal "t <| [] = t";
3192
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2087
diff changeset
    15
by (induct_tac "t" 1);
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2087
diff changeset
    16
by (ALLGOALS Asm_simp_tac);
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    17
qed "subst_Nil";
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    18
3192
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2087
diff changeset
    19
Addsimps [subst_Nil];
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2087
diff changeset
    20
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4686
diff changeset
    21
Goal "t <: u --> t <| s <: u <| s";
3192
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2087
diff changeset
    22
by (induct_tac "u" 1);
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2087
diff changeset
    23
by (ALLGOALS Asm_simp_tac);
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2087
diff changeset
    24
qed_spec_mp "subst_mono";
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    25
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4686
diff changeset
    26
Goal  "~ (Var(v) <: t) --> t <| (v,t <| s) # s = t <| s";
3192
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2087
diff changeset
    27
by (case_tac "t = Var(v)" 1);
3457
a8ab7c64817c Ran expandshort
paulson
parents: 3268
diff changeset
    28
by (etac rev_mp 2);
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    29
by (res_inst_tac [("P",
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 968
diff changeset
    30
    "%x.~x=Var(v) --> ~(Var(v) <: x) --> x <| (v,t<|s)#s=x<|s")]
3192
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2087
diff changeset
    31
    uterm.induct 2);
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2087
diff changeset
    32
by (ALLGOALS Asm_simp_tac);
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2087
diff changeset
    33
by (Blast_tac 1);
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2087
diff changeset
    34
qed_spec_mp "Var_not_occs";
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    35
5278
a903b66822e2 even more tidying of Goal commands
paulson
parents: 5119
diff changeset
    36
Goal "(t <|r = t <|s) = (! v. v : vars_of(t) --> Var(v) <|r = Var(v) <|s)";
3192
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2087
diff changeset
    37
by (induct_tac "t" 1);
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2087
diff changeset
    38
by (ALLGOALS Asm_full_simp_tac);
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2087
diff changeset
    39
by (ALLGOALS Blast_tac);
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    40
qed "agreement";
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    41
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4686
diff changeset
    42
Goal   "~ v: vars_of(t) --> t <| (v,u)#s = t <| s";
4686
74a12e86b20b Removed `addsplits [expand_if]'
nipkow
parents: 4477
diff changeset
    43
by (simp_tac (simpset() addsimps [agreement]) 1);
3192
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2087
diff changeset
    44
qed_spec_mp"repl_invariance";
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    45
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    46
val asms = goal Subst.thy 
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 968
diff changeset
    47
     "v : vars_of(t) --> w : vars_of(t <| (v,Var(w))#s)";
3192
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2087
diff changeset
    48
by (induct_tac "t" 1);
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2087
diff changeset
    49
by (ALLGOALS Asm_simp_tac);
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2087
diff changeset
    50
qed_spec_mp"Var_in_subst";
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2087
diff changeset
    51
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    52
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    53
(**** Equality between Substitutions ****)
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    54
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4686
diff changeset
    55
Goalw [subst_eq_def] "r =$= s = (! t. t <| r = t <| s)";
3192
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2087
diff changeset
    56
by (Simp_tac 1);
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    57
qed "subst_eq_iff";
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    58
3192
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2087
diff changeset
    59
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2087
diff changeset
    60
local fun prove s = prove_goal Subst.thy s
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    61
                  (fn prems => [cut_facts_tac prems 1,
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    62
                                REPEAT (etac rev_mp 1),
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 4060
diff changeset
    63
                                simp_tac (simpset() addsimps [subst_eq_iff]) 1])
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    64
in 
3192
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2087
diff changeset
    65
  val subst_refl      = prove "r =$= r";
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2087
diff changeset
    66
  val subst_sym       = prove "r =$= s ==> s =$= r";
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2087
diff changeset
    67
  val subst_trans     = prove "[| q =$= r; r =$= s |] ==> q =$= s";
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    68
end;
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    69
3192
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2087
diff changeset
    70
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2087
diff changeset
    71
AddIffs [subst_refl];
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2087
diff changeset
    72
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2087
diff changeset
    73
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    74
val eq::prems = goalw Subst.thy [subst_eq_def] 
3192
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2087
diff changeset
    75
    "[| r =$= s; P (t <| r) (u <| r) |] ==> P (t <| s) (u <| s)";
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    76
by (resolve_tac [eq RS spec RS subst] 1);
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    77
by (resolve_tac (prems RL [eq RS spec RS subst]) 1);
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    78
qed "subst_subst2";
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    79
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    80
val ssubst_subst2 = subst_sym RS subst_subst2;
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    81
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    82
(**** Composition of Substitutions ****)
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    83
3192
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2087
diff changeset
    84
let fun prove s = 
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2087
diff changeset
    85
 prove_goalw Subst.thy [comp_def,sdom_def] s (fn _ => [Simp_tac 1])
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2087
diff changeset
    86
in 
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2087
diff changeset
    87
Addsimps
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2087
diff changeset
    88
 (
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2087
diff changeset
    89
   map prove 
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2087
diff changeset
    90
   [ "[] <> bl = bl",
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2087
diff changeset
    91
     "((a,b)#al) <> bl = (a,b <| bl) # (al <> bl)",
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2087
diff changeset
    92
     "sdom([]) = {}",
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2087
diff changeset
    93
     "sdom((a,b)#al) = (if Var(a)=b then (sdom al) - {a} else sdom al Un {a})"]
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2087
diff changeset
    94
 )
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2087
diff changeset
    95
end;
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2087
diff changeset
    96
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2087
diff changeset
    97
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4686
diff changeset
    98
Goal "s <> [] = s";
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    99
by (alist_ind_tac "s" 1);
3192
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2087
diff changeset
   100
by (ALLGOALS Asm_simp_tac);
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   101
qed "comp_Nil";
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   102
3192
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2087
diff changeset
   103
Addsimps [comp_Nil];
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2087
diff changeset
   104
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4686
diff changeset
   105
Goal "s =$= s <> []";
3192
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2087
diff changeset
   106
by (Simp_tac 1);
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2087
diff changeset
   107
qed "subst_comp_Nil";
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2087
diff changeset
   108
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4686
diff changeset
   109
Goal "(t <| r <> s) = (t <| r <| s)";
3192
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2087
diff changeset
   110
by (induct_tac "t" 1);
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2087
diff changeset
   111
by (ALLGOALS Asm_simp_tac);
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   112
by (alist_ind_tac "r" 1);
4686
74a12e86b20b Removed `addsplits [expand_if]'
nipkow
parents: 4477
diff changeset
   113
by (ALLGOALS Asm_simp_tac);
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   114
qed "subst_comp";
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   115
3192
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2087
diff changeset
   116
Addsimps [subst_comp];
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2087
diff changeset
   117
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4686
diff changeset
   118
Goal "(q <> r) <> s =$= q <> (r <> s)";
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 4060
diff changeset
   119
by (simp_tac (simpset() addsimps [subst_eq_iff]) 1);
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   120
qed "comp_assoc";
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   121
5119
58d267fc8630 Removed leading !! in goals
nipkow
parents: 5069
diff changeset
   122
Goal "[| theta =$= theta1; sigma =$= sigma1|] ==> \
58d267fc8630 Removed leading !! in goals
nipkow
parents: 5069
diff changeset
   123
\     (theta <> sigma) =$= (theta1 <> sigma1)";
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 4060
diff changeset
   124
by (asm_full_simp_tac (simpset() addsimps [subst_eq_def]) 1);
3192
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2087
diff changeset
   125
qed "subst_cong";
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2087
diff changeset
   126
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2087
diff changeset
   127
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4686
diff changeset
   128
Goal "(w, Var(w) <| s) # s =$= s"; 
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 4060
diff changeset
   129
by (simp_tac (simpset() addsimps [subst_eq_iff]) 1);
3192
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2087
diff changeset
   130
by (rtac allI 1);
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2087
diff changeset
   131
by (induct_tac "t" 1);
4686
74a12e86b20b Removed `addsplits [expand_if]'
nipkow
parents: 4477
diff changeset
   132
by (ALLGOALS Asm_full_simp_tac);
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   133
qed "Cons_trivial";
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   134
3192
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2087
diff changeset
   135
5119
58d267fc8630 Removed leading !! in goals
nipkow
parents: 5069
diff changeset
   136
Goal "q <> r =$= s ==>  t <| q <| r = t <| s";
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 4060
diff changeset
   137
by (asm_full_simp_tac (simpset() addsimps [subst_eq_iff]) 1);
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   138
qed "comp_subst_subst";
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   139
3192
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2087
diff changeset
   140
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   141
(****  Domain and range of Substitutions ****)
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   142
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4686
diff changeset
   143
Goal  "(v : sdom(s)) = (Var(v) <| s ~= Var(v))";
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   144
by (alist_ind_tac "s" 1);
4686
74a12e86b20b Removed `addsplits [expand_if]'
nipkow
parents: 4477
diff changeset
   145
by (ALLGOALS Asm_simp_tac);
3192
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2087
diff changeset
   146
by (Blast_tac 1);
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   147
qed "sdom_iff";
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   148
3192
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2087
diff changeset
   149
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4686
diff changeset
   150
Goalw [srange_def]  
3842
b55686a7b22c fixed dots;
wenzelm
parents: 3724
diff changeset
   151
   "v : srange(s) = (? w. w : sdom(s) & v : vars_of(Var(w) <| s))";
3192
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2087
diff changeset
   152
by (Blast_tac 1);
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   153
qed "srange_iff";
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   154
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4686
diff changeset
   155
Goalw [empty_def] "(A = {}) = (ALL a.~ a:A)";
3192
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2087
diff changeset
   156
by (Blast_tac 1);
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2087
diff changeset
   157
qed "empty_iff_all_not";
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2087
diff changeset
   158
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4686
diff changeset
   159
Goal  "(t <| s = t) = (sdom(s) Int vars_of(t) = {})";
3192
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2087
diff changeset
   160
by (induct_tac "t" 1);
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2087
diff changeset
   161
by (ALLGOALS
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 4060
diff changeset
   162
    (asm_full_simp_tac (simpset() addsimps [empty_iff_all_not, sdom_iff])));
3192
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2087
diff changeset
   163
by (ALLGOALS Blast_tac);
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   164
qed "invariance";
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   165
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4686
diff changeset
   166
Goal  "v : sdom(s) -->  v : vars_of(t <| s) --> v : srange(s)";
3192
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2087
diff changeset
   167
by (induct_tac "t" 1);
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2087
diff changeset
   168
by (case_tac "a : sdom(s)" 1);
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 4060
diff changeset
   169
by (ALLGOALS (asm_full_simp_tac (simpset() addsimps [sdom_iff, srange_iff])));
3192
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2087
diff changeset
   170
by (ALLGOALS Blast_tac);
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2087
diff changeset
   171
qed_spec_mp "Var_in_srange";
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   172
5278
a903b66822e2 even more tidying of Goal commands
paulson
parents: 5119
diff changeset
   173
Goal "[| v : sdom(s); v ~: srange(s) |] ==>  v ~: vars_of(t <| s)";
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 4060
diff changeset
   174
by (blast_tac (claset() addIs [Var_in_srange]) 1);
3192
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2087
diff changeset
   175
qed "Var_elim";
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   176
5278
a903b66822e2 even more tidying of Goal commands
paulson
parents: 5119
diff changeset
   177
Goal "v : vars_of(t <| s) --> v : srange(s) | v : vars_of(t)";
3192
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2087
diff changeset
   178
by (induct_tac "t" 1);
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 4060
diff changeset
   179
by (ALLGOALS (asm_full_simp_tac (simpset() addsimps [sdom_iff,srange_iff])));
3192
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2087
diff changeset
   180
by (Blast_tac 2);
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 4060
diff changeset
   181
by (safe_tac (claset() addSIs [exI, vars_var_iff RS iffD1 RS sym]));
4477
b3e5857d8d99 New Auto_tac (by Oheimb), and new syntax (without parens), and expandshort
paulson
parents: 4089
diff changeset
   182
by Auto_tac;
3192
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2087
diff changeset
   183
qed_spec_mp "Var_intro";
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   184
5278
a903b66822e2 even more tidying of Goal commands
paulson
parents: 5119
diff changeset
   185
Goal "v : srange(s) --> (? w. w : sdom(s) & v : vars_of(Var(w) <| s))";
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 4060
diff changeset
   186
by (simp_tac (simpset() addsimps [srange_iff]) 1);
3192
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2087
diff changeset
   187
qed_spec_mp "srangeD";
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   188
5278
a903b66822e2 even more tidying of Goal commands
paulson
parents: 5119
diff changeset
   189
Goal "sdom(s) Int srange(s) = {} = (! t. sdom(s) Int vars_of(t <| s) = {})";
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 4060
diff changeset
   190
by (simp_tac (simpset() addsimps [empty_iff_all_not]) 1);
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 4060
diff changeset
   191
by (fast_tac (claset() addIs [Var_in_srange] addDs [srangeD]) 1);
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   192
qed "dom_range_disjoint";
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
   193
5119
58d267fc8630 Removed leading !! in goals
nipkow
parents: 5069
diff changeset
   194
Goal "~ u <| s = u ==> (? x. x : sdom(s))";
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 4060
diff changeset
   195
by (full_simp_tac (simpset() addsimps [empty_iff_all_not, invariance]) 1);
3192
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2087
diff changeset
   196
by (Blast_tac 1);
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2087
diff changeset
   197
qed "subst_not_empty";
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2087
diff changeset
   198
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2087
diff changeset
   199
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4686
diff changeset
   200
Goal "(M <| [(x, Var x)]) = M";
3192
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2087
diff changeset
   201
by (induct_tac "M" 1);
4686
74a12e86b20b Removed `addsplits [expand_if]'
nipkow
parents: 4477
diff changeset
   202
by (ALLGOALS Asm_simp_tac);
3192
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2087
diff changeset
   203
qed "id_subst_lemma";
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2087
diff changeset
   204
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 2087
diff changeset
   205
Addsimps [id_subst_lemma];