src/HOL/Vimage.ML
author paulson
Fri, 27 Feb 1998 11:06:28 +0100
changeset 4660 63f0b2601792
parent 4648 f04da668581c
child 4674 248b876c2fa8
permissions -rw-r--r--
New vimage laws
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
4648
f04da668581c New theory of the inverse image of a function
paulson
parents:
diff changeset
     1
(*  Title:      HOL/Vimage
f04da668581c New theory of the inverse image of a function
paulson
parents:
diff changeset
     2
    ID:         $Id$
f04da668581c New theory of the inverse image of a function
paulson
parents:
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
f04da668581c New theory of the inverse image of a function
paulson
parents:
diff changeset
     4
    Copyright   1998  University of Cambridge
f04da668581c New theory of the inverse image of a function
paulson
parents:
diff changeset
     5
f04da668581c New theory of the inverse image of a function
paulson
parents:
diff changeset
     6
Inverse image of a function
f04da668581c New theory of the inverse image of a function
paulson
parents:
diff changeset
     7
*)
f04da668581c New theory of the inverse image of a function
paulson
parents:
diff changeset
     8
f04da668581c New theory of the inverse image of a function
paulson
parents:
diff changeset
     9
open Vimage;
f04da668581c New theory of the inverse image of a function
paulson
parents:
diff changeset
    10
f04da668581c New theory of the inverse image of a function
paulson
parents:
diff changeset
    11
(** Basic rules **)
f04da668581c New theory of the inverse image of a function
paulson
parents:
diff changeset
    12
f04da668581c New theory of the inverse image of a function
paulson
parents:
diff changeset
    13
qed_goalw "vimage_eq" thy [vimage_def]
f04da668581c New theory of the inverse image of a function
paulson
parents:
diff changeset
    14
    "(a : f-``B) = (f a : B)"
f04da668581c New theory of the inverse image of a function
paulson
parents:
diff changeset
    15
 (fn _ => [ (Blast_tac 1) ]);
f04da668581c New theory of the inverse image of a function
paulson
parents:
diff changeset
    16
f04da668581c New theory of the inverse image of a function
paulson
parents:
diff changeset
    17
qed_goal "vimage_singleton_eq" thy
f04da668581c New theory of the inverse image of a function
paulson
parents:
diff changeset
    18
    "(a : f-``{b}) = (f a = b)"
f04da668581c New theory of the inverse image of a function
paulson
parents:
diff changeset
    19
 (fn _ => [ simp_tac (simpset() addsimps [vimage_eq]) 1 ]);
f04da668581c New theory of the inverse image of a function
paulson
parents:
diff changeset
    20
f04da668581c New theory of the inverse image of a function
paulson
parents:
diff changeset
    21
qed_goalw "vimageI" thy [vimage_def]
f04da668581c New theory of the inverse image of a function
paulson
parents:
diff changeset
    22
    "!!A B f. [| f a = b;  b:B |] ==> a : f-``B"
f04da668581c New theory of the inverse image of a function
paulson
parents:
diff changeset
    23
 (fn _ => [ (Blast_tac 1) ]);
f04da668581c New theory of the inverse image of a function
paulson
parents:
diff changeset
    24
f04da668581c New theory of the inverse image of a function
paulson
parents:
diff changeset
    25
qed_goalw "vimageE" thy [vimage_def]
f04da668581c New theory of the inverse image of a function
paulson
parents:
diff changeset
    26
    "[| a: f-``B;  !!x.[| f a = x;  x:B |] ==> P |] ==> P"
f04da668581c New theory of the inverse image of a function
paulson
parents:
diff changeset
    27
 (fn major::prems=>
f04da668581c New theory of the inverse image of a function
paulson
parents:
diff changeset
    28
  [ (rtac (major RS CollectE) 1),
f04da668581c New theory of the inverse image of a function
paulson
parents:
diff changeset
    29
    (blast_tac (claset() addIs prems) 1) ]);
f04da668581c New theory of the inverse image of a function
paulson
parents:
diff changeset
    30
f04da668581c New theory of the inverse image of a function
paulson
parents:
diff changeset
    31
AddIs  [vimageI];
f04da668581c New theory of the inverse image of a function
paulson
parents:
diff changeset
    32
AddSEs [vimageE];
f04da668581c New theory of the inverse image of a function
paulson
parents:
diff changeset
    33
f04da668581c New theory of the inverse image of a function
paulson
parents:
diff changeset
    34
f04da668581c New theory of the inverse image of a function
paulson
parents:
diff changeset
    35
(*** Simple equalities ***)
f04da668581c New theory of the inverse image of a function
paulson
parents:
diff changeset
    36
f04da668581c New theory of the inverse image of a function
paulson
parents:
diff changeset
    37
goal thy "f-``{} = {}";
f04da668581c New theory of the inverse image of a function
paulson
parents:
diff changeset
    38
by (Blast_tac 1);
f04da668581c New theory of the inverse image of a function
paulson
parents:
diff changeset
    39
qed "vimage_empty";
f04da668581c New theory of the inverse image of a function
paulson
parents:
diff changeset
    40
4660
63f0b2601792 New vimage laws
paulson
parents: 4648
diff changeset
    41
goal thy "f-``(Compl A) = Compl (f-``A)";
63f0b2601792 New vimage laws
paulson
parents: 4648
diff changeset
    42
by (Blast_tac 1);
63f0b2601792 New vimage laws
paulson
parents: 4648
diff changeset
    43
qed "vimage_Compl";
63f0b2601792 New vimage laws
paulson
parents: 4648
diff changeset
    44
4648
f04da668581c New theory of the inverse image of a function
paulson
parents:
diff changeset
    45
goal thy "f-``(A Un B) = (f-``A) Un (f-``B)";
f04da668581c New theory of the inverse image of a function
paulson
parents:
diff changeset
    46
by (Blast_tac 1);
f04da668581c New theory of the inverse image of a function
paulson
parents:
diff changeset
    47
qed "vimage_Un";
f04da668581c New theory of the inverse image of a function
paulson
parents:
diff changeset
    48
4660
63f0b2601792 New vimage laws
paulson
parents: 4648
diff changeset
    49
Addsimps [vimage_empty, vimage_Compl, vimage_Un];
63f0b2601792 New vimage laws
paulson
parents: 4648
diff changeset
    50
4648
f04da668581c New theory of the inverse image of a function
paulson
parents:
diff changeset
    51
(*NOT suitable for rewriting because of the recurrence of {a}*)
f04da668581c New theory of the inverse image of a function
paulson
parents:
diff changeset
    52
goal thy "f-``(insert a B) = (f-``{a}) Un (f-``B)";
f04da668581c New theory of the inverse image of a function
paulson
parents:
diff changeset
    53
by (Blast_tac 1);
f04da668581c New theory of the inverse image of a function
paulson
parents:
diff changeset
    54
qed "vimage_insert";
f04da668581c New theory of the inverse image of a function
paulson
parents:
diff changeset
    55
f04da668581c New theory of the inverse image of a function
paulson
parents:
diff changeset
    56
goal thy "f-``(A Int B) <= (f-``A) Int (f-``B)";
f04da668581c New theory of the inverse image of a function
paulson
parents:
diff changeset
    57
by (Blast_tac 1);
f04da668581c New theory of the inverse image of a function
paulson
parents:
diff changeset
    58
qed "vimage_Int_subset";
f04da668581c New theory of the inverse image of a function
paulson
parents:
diff changeset
    59
f04da668581c New theory of the inverse image of a function
paulson
parents:
diff changeset
    60
goal thy "f-``UNIV = UNIV";
f04da668581c New theory of the inverse image of a function
paulson
parents:
diff changeset
    61
by (Blast_tac 1);
f04da668581c New theory of the inverse image of a function
paulson
parents:
diff changeset
    62
qed "vimage_UNIV";
f04da668581c New theory of the inverse image of a function
paulson
parents:
diff changeset
    63
Addsimps [vimage_UNIV];
f04da668581c New theory of the inverse image of a function
paulson
parents:
diff changeset
    64
f04da668581c New theory of the inverse image of a function
paulson
parents:
diff changeset
    65
goal thy "(UN x:A. f -`` B x) = f -`` (UN x:A. B x)";
f04da668581c New theory of the inverse image of a function
paulson
parents:
diff changeset
    66
by (Blast_tac 1);
f04da668581c New theory of the inverse image of a function
paulson
parents:
diff changeset
    67
qed "UN_vimage";
f04da668581c New theory of the inverse image of a function
paulson
parents:
diff changeset
    68
Addsimps [UN_vimage];
f04da668581c New theory of the inverse image of a function
paulson
parents:
diff changeset
    69
f04da668581c New theory of the inverse image of a function
paulson
parents:
diff changeset
    70
f04da668581c New theory of the inverse image of a function
paulson
parents:
diff changeset
    71
(** monotonicity **)
f04da668581c New theory of the inverse image of a function
paulson
parents:
diff changeset
    72
f04da668581c New theory of the inverse image of a function
paulson
parents:
diff changeset
    73
goal thy "!!f. A<=B ==> f-``A <= f-``B";
f04da668581c New theory of the inverse image of a function
paulson
parents:
diff changeset
    74
by (Blast_tac 1);
f04da668581c New theory of the inverse image of a function
paulson
parents:
diff changeset
    75
qed "vimage_mono";