Subst/utlemmas.ML
changeset 0 7949f97df77a
equal deleted inserted replaced
-1:000000000000 0:7949f97df77a
       
     1 (*  Title: 	Substitutions/utermlemmas.ML
       
     2     Author: 	Martin Coen, Cambridge University Computer Laboratory
       
     3     Copyright   1993  University of Cambridge
       
     4 
       
     5 For utermlemmas.thy.  
       
     6 *)
       
     7 
       
     8 open UTLemmas;
       
     9 
       
    10 (***********)
       
    11 
       
    12 val utlemmas_defs = [vars_of_def,occs_def];
       
    13 
       
    14 local fun mk_thm s = prove_goalw UTLemmas.thy utlemmas_defs s 
       
    15                                  (fn _ => [simp_tac uterm_ss 1])
       
    16 in val utlemmas_rews = map mk_thm 
       
    17       ["vars_of(Const(c)) = {}",
       
    18        "vars_of(Var(x)) = {x}",
       
    19        "vars_of(Comb(t,u)) = vars_of(t) Un vars_of(u)",
       
    20        "t <: Const(c) = False",
       
    21        "t <: Var(x) = False",
       
    22        "t <: Comb(u,v) = (t=u | t=v | t <: u | t <: v)"];
       
    23 end;
       
    24 
       
    25 val utlemmas_ss = pair_ss addsimps (setplus_rews @ uterm_rews @ utlemmas_rews);
       
    26 
       
    27 (****  occs irrefl ****)
       
    28 
       
    29 goal UTLemmas.thy  "t <: u & u <: v --> t <: v";
       
    30 by (uterm_ind_tac "v" 1);
       
    31 by (ALLGOALS (simp_tac utlemmas_ss));
       
    32 by (fast_tac HOL_cs 1);
       
    33 val occs_trans = conjI RS (result() RS mp);
       
    34 
       
    35 goal UTLemmas.thy   "~ t <: v --> ~ t <: u | ~ u <: v";
       
    36 by (fast_tac (HOL_cs addIs [occs_trans]) 1);
       
    37 val contr_occs_trans = result() RS mp;
       
    38 
       
    39 goal UTLemmas.thy   "t <: Comb(t,u)";
       
    40 by (simp_tac utlemmas_ss 1);
       
    41 val occs_Comb1 = result();
       
    42 
       
    43 goal UTLemmas.thy  "u <: Comb(t,u)";
       
    44 by (simp_tac utlemmas_ss 1);
       
    45 val occs_Comb2 = result();
       
    46 
       
    47 goal HOL.thy  "(~(P|Q)) = (~P & ~Q)";
       
    48 by (fast_tac HOL_cs 1);
       
    49 val demorgan_disj = result();
       
    50 
       
    51 goal UTLemmas.thy  "~ t <: t";
       
    52 by (uterm_ind_tac "t" 1);
       
    53 by (ALLGOALS (simp_tac (utlemmas_ss addsimps [demorgan_disj])));
       
    54 by (REPEAT (resolve_tac [impI,conjI] 1 ORELSE
       
    55             (etac contrapos 1 THEN etac subst 1 THEN 
       
    56              resolve_tac [occs_Comb1,occs_Comb2] 1) ORELSE
       
    57             (etac (contr_occs_trans RS disjE) 1 THEN assume_tac 2) ORELSE
       
    58             eresolve_tac ([occs_Comb1,occs_Comb2] RLN(2,[notE])) 1));
       
    59 val occs_irrefl = result();
       
    60 
       
    61 goal UTLemmas.thy  "t <: u --> ~t=u";
       
    62 by (fast_tac (HOL_cs addEs [occs_irrefl RS notE]) 1);
       
    63 val occs_irrefl2 = result() RS mp;
       
    64 
       
    65 
       
    66 (**** vars_of lemmas  ****)
       
    67 
       
    68 goal UTLemmas.thy "(v : vars_of(Var(w))) = (w=v)";
       
    69 by (simp_tac utlemmas_ss 1);
       
    70 by (fast_tac HOL_cs 1);
       
    71 val vars_var_iff = result();
       
    72 
       
    73 goal UTLemmas.thy  "(x : vars_of(t)) = (Var(x) <: t | Var(x) = t)";
       
    74 by (uterm_ind_tac "t" 1);
       
    75 by (ALLGOALS (simp_tac utlemmas_ss));
       
    76 by (fast_tac HOL_cs 1);
       
    77 val vars_iff_occseq = result();