src/HOL/Subst/UTLemmas.ML
author clasohm
Tue, 30 Jan 1996 15:24:36 +0100
changeset 1465 5d7a7e439cec
parent 1266 3ae9fe3c0f68
permissions -rw-r--r--
expanded tabs
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1266
diff changeset
     1
(*  Title:      HOL/Subst/UTLemmas.ML
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 968
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 UTLemmas.thy.  
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 UTLemmas;
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
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 968
diff changeset
    13
val utlemmas_defs = [vars_of_def, occs_def];
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    14
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    15
local fun mk_thm s = prove_goalw UTLemmas.thy utlemmas_defs s 
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 968
diff changeset
    16
                                 (fn _ => [Simp_tac 1])
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    17
in val utlemmas_rews = map mk_thm 
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    18
      ["vars_of(Const(c)) = {}",
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    19
       "vars_of(Var(x)) = {x}",
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    20
       "vars_of(Comb t u) = vars_of(t) Un vars_of(u)",
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    21
       "t <: Const(c) = False",
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    22
       "t <: Var(x) = False",
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    23
       "t <: Comb u v = (t=u | t=v | t <: u | t <: v)"];
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    24
end;
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    25
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 968
diff changeset
    26
Addsimps (setplus_rews @ uterm_rews @ utlemmas_rews);
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    27
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    28
(****  occs irrefl ****)
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    29
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    30
goal UTLemmas.thy  "t <: u & u <: v --> t <: v";
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    31
by (uterm_ind_tac "v" 1);
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 968
diff changeset
    32
by (ALLGOALS Simp_tac);
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    33
by (fast_tac HOL_cs 1);
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    34
val occs_trans  = store_thm("occs_trans", conjI RS (result() RS mp));
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    35
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    36
goal UTLemmas.thy   "~ t <: v --> ~ t <: u | ~ u <: v";
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    37
by (fast_tac (HOL_cs addIs [occs_trans]) 1);
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    38
val contr_occs_trans  = store_thm("contr_occs_trans", result() RS mp);
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    39
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    40
goal UTLemmas.thy   "t <: Comb t u";
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 968
diff changeset
    41
by (Simp_tac 1);
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    42
qed "occs_Comb1";
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    43
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    44
goal UTLemmas.thy  "u <: Comb t u";
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 968
diff changeset
    45
by (Simp_tac 1);
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    46
qed "occs_Comb2";
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    47
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    48
goal HOL.thy  "(~(P|Q)) = (~P & ~Q)";
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    49
by (fast_tac HOL_cs 1);
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    50
qed "demorgan_disj";
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    51
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    52
goal UTLemmas.thy  "~ t <: t";
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    53
by (uterm_ind_tac "t" 1);
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 968
diff changeset
    54
by (ALLGOALS (simp_tac (!simpset addsimps [demorgan_disj])));
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    55
by (REPEAT (resolve_tac [impI,conjI] 1 ORELSE
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    56
            (etac contrapos 1 THEN etac subst 1 THEN 
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    57
             resolve_tac [occs_Comb1,occs_Comb2] 1) ORELSE
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    58
            (etac (contr_occs_trans RS disjE) 1 THEN assume_tac 2) ORELSE
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    59
            eresolve_tac ([occs_Comb1,occs_Comb2] RLN(2,[notE])) 1));
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    60
qed "occs_irrefl";
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    61
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    62
goal UTLemmas.thy  "t <: u --> ~t=u";
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    63
by (fast_tac (HOL_cs addEs [occs_irrefl RS notE]) 1);
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    64
val occs_irrefl2  = store_thm("occs_irrefl2", result() RS mp);
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    65
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    66
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    67
(**** vars_of lemmas  ****)
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    68
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    69
goal UTLemmas.thy "(v : vars_of(Var(w))) = (w=v)";
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 968
diff changeset
    70
by (Simp_tac 1);
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    71
by (fast_tac HOL_cs 1);
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    72
qed "vars_var_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
goal UTLemmas.thy  "(x : vars_of(t)) = (Var(x) <: t | Var(x) = t)";
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    75
by (uterm_ind_tac "t" 1);
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 968
diff changeset
    76
by (ALLGOALS Simp_tac);
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    77
by (fast_tac HOL_cs 1);
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    78
qed "vars_iff_occseq";