Integ/Relation.ML
author clasohm
Tue, 24 Oct 1995 14:59:17 +0100
changeset 251 f04b33ce250f
parent 217 b6c0407f203e
permissions -rw-r--r--
added calls of init_html and make_chart
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
217
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
     1
(*  Title: 	Relation.ML
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
     2
    ID:         $Id$
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
     3
    Authors: 	Riccardo Mattolini, Dip. Sistemi e Informatica
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
     4
        	Lawrence C Paulson, Cambridge University Computer Laboratory
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
     5
    Copyright   1994 Universita' di Firenze
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
     6
    Copyright   1993  University of Cambridge
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
     7
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
     8
Functions represented as relations in HOL Set Theory 
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
     9
*)
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    10
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    11
val RSLIST = curry (op MRS);
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    12
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    13
open Relation;
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    14
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    15
goalw Relation.thy [converse_def] "!!a b r. <a,b>:r ==> <b,a>:converse(r)";
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    16
by (simp_tac prod_ss 1);
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    17
by (fast_tac set_cs 1);
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    18
qed "converseI";
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    19
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    20
goalw Relation.thy [converse_def] "!!a b r. <a,b> : converse(r) ==> <b,a> : r";
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    21
by (fast_tac comp_cs 1);
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    22
qed "converseD";
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    23
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    24
qed_goalw "converseE" Relation.thy [converse_def]
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    25
    "[| yx : converse(r);  \
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    26
\       !!x y. [| yx=<y,x>;  <x,y>:r |] ==> P \
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    27
\    |] ==> P"
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    28
 (fn [major,minor]=>
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    29
  [ (rtac (major RS CollectE) 1),
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    30
    (REPEAT (eresolve_tac [bexE,exE, conjE, minor] 1)),
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    31
    (hyp_subst_tac 1),
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    32
    (assume_tac 1) ]);
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    33
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    34
val converse_cs = comp_cs addSIs [converseI] 
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    35
			  addSEs [converseD,converseE];
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    36
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    37
qed_goalw "Domain_iff" Relation.thy [Domain_def]
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    38
    "a: Domain(r) = (EX y. <a,y>: r)"
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    39
 (fn _=> [ (fast_tac comp_cs 1) ]);
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    40
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    41
qed_goal "DomainI" Relation.thy "!!a b r. <a,b>: r ==> a: Domain(r)"
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    42
 (fn _ => [ (etac (exI RS (Domain_iff RS iffD2)) 1) ]);
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    43
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    44
qed_goal "DomainE" Relation.thy
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    45
    "[| a : Domain(r);  !!y. <a,y>: r ==> P |] ==> P"
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    46
 (fn prems=>
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    47
  [ (rtac (Domain_iff RS iffD1 RS exE) 1),
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    48
    (REPEAT (ares_tac prems 1)) ]);
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    49
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    50
qed_goalw "RangeI" Relation.thy [Range_def] "!!a b r.<a,b>: r ==> b : Range(r)"
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    51
 (fn _ => [ (etac (converseI RS DomainI) 1) ]);
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    52
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    53
qed_goalw "RangeE" Relation.thy [Range_def]
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    54
    "[| b : Range(r);  !!x. <x,b>: r ==> P |] ==> P"
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    55
 (fn major::prems=>
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    56
  [ (rtac (major RS DomainE) 1),
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    57
    (resolve_tac prems 1),
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    58
    (etac converseD 1) ]);
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    59
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    60
(*** Image of a set under a function/relation ***)
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    61
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    62
qed_goalw "Image_iff" Relation.thy [Image_def]
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    63
    "b : r^^A = (? x:A. <x,b>:r)"
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    64
 (fn _ => [ fast_tac (comp_cs addIs [RangeI]) 1 ]);
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    65
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    66
qed_goal "Image_singleton_iff" Relation.thy
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    67
    "(b : r^^{a}) = (<a,b>:r)"
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    68
 (fn _ => [ rtac (Image_iff RS trans) 1,
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    69
	    fast_tac comp_cs 1 ]);
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    70
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    71
qed_goalw "ImageI" Relation.thy [Image_def]
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    72
    "!!a b r. [| <a,b>: r;  a:A |] ==> b : r^^A"
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    73
 (fn _ => [ (REPEAT (ares_tac [CollectI,RangeI,bexI] 1)),
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    74
            (resolve_tac [conjI ] 1),
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    75
            (resolve_tac [RangeI] 1),
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    76
            (REPEAT (fast_tac set_cs 1))]);
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    77
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    78
qed_goalw "ImageE" Relation.thy [Image_def]
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    79
    "[| b: r^^A;  !!x.[| <x,b>: r;  x:A |] ==> P |] ==> P"
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    80
 (fn major::prems=>
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    81
  [ (rtac (major RS CollectE) 1),
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    82
    (safe_tac set_cs),
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    83
    (etac RangeE 1),
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    84
    (rtac (hd prems) 1),
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    85
    (REPEAT (etac bexE 1 ORELSE ares_tac prems 1)) ]);
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    86
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    87
qed_goal "Image_subset" Relation.thy
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    88
    "!!A B r. r <= Sigma(A,%x.B) ==> r^^C <= B"
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    89
 (fn _ =>
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    90
  [ (rtac subsetI 1),
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    91
    (REPEAT (eresolve_tac [asm_rl, ImageE, subsetD RS SigmaD2] 1)) ]);
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    92
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    93
val rel_cs = converse_cs addSIs [converseI] 
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    94
                         addIs  [ImageI, DomainI, RangeI]
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    95
                         addSEs [ImageE, DomainE, RangeE];
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    96
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    97
val rel_eq_cs = rel_cs addSIs [equalityI];
b6c0407f203e Installation of Integ (ported from ZF by Mattolini)
lcp
parents:
diff changeset
    98