src/ZF/ex/Rmap.ML
author paulson
Fri, 11 Sep 1998 18:09:54 +0200
changeset 5484 e9430ed7e8d6
parent 5147 825877190618
child 5505 b0856ff6fc69
permissions -rw-r--r--
Extra steps at end to make it run faster
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 782
diff changeset
     1
(*  Title:      ZF/ex/Rmap
16
0b033d50ca1c ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents:
diff changeset
     2
    ID:         $Id$
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 782
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 496
diff changeset
     4
    Copyright   1994  University of Cambridge
16
0b033d50ca1c ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents:
diff changeset
     5
0b033d50ca1c ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents:
diff changeset
     6
Inductive definition of an operator to "map" a relation over a list
0b033d50ca1c ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents:
diff changeset
     7
*)
0b033d50ca1c ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents:
diff changeset
     8
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 496
diff changeset
     9
open Rmap;
16
0b033d50ca1c ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents:
diff changeset
    10
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5068
diff changeset
    11
Goalw rmap.defs "r<=s ==> rmap(r) <= rmap(s)";
16
0b033d50ca1c ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents:
diff changeset
    12
by (rtac lfp_mono 1);
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 496
diff changeset
    13
by (REPEAT (rtac rmap.bnd_mono 1));
16
0b033d50ca1c ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents:
diff changeset
    14
by (REPEAT (ares_tac ([Sigma_mono, list_mono, domain_mono, range_mono] @ 
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 782
diff changeset
    15
                      basic_monos) 1));
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 694
diff changeset
    16
qed "rmap_mono";
16
0b033d50ca1c ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents:
diff changeset
    17
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 496
diff changeset
    18
val Nil_rmap_case = rmap.mk_cases list.con_defs "<Nil,zs> : rmap(r)"
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 496
diff changeset
    19
and Cons_rmap_case = rmap.mk_cases list.con_defs "<Cons(x,xs),zs> : rmap(r)";
16
0b033d50ca1c ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents:
diff changeset
    20
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1732
diff changeset
    21
AddIs  rmap.intrs;
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1732
diff changeset
    22
AddSEs [Nil_rmap_case, Cons_rmap_case];
16
0b033d50ca1c ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents:
diff changeset
    23
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5068
diff changeset
    24
Goal "r <= A*B ==> rmap(r) <= list(A)*list(B)";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 496
diff changeset
    25
by (rtac (rmap.dom_subset RS subset_trans) 1);
16
0b033d50ca1c ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents:
diff changeset
    26
by (REPEAT (ares_tac [domain_rel_subset, range_rel_subset,
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 782
diff changeset
    27
                      Sigma_mono, list_mono] 1));
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 694
diff changeset
    28
qed "rmap_rel_type";
16
0b033d50ca1c ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents:
diff changeset
    29
5147
825877190618 More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5137
diff changeset
    30
Goal "A <= domain(r) ==> list(A) <= domain(rmap(r))";
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 782
diff changeset
    31
by (rtac subsetI 1);
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 496
diff changeset
    32
by (etac list.induct 1);
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1732
diff changeset
    33
by (ALLGOALS Fast_tac);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 694
diff changeset
    34
qed "rmap_total";
16
0b033d50ca1c ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents:
diff changeset
    35
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5068
diff changeset
    36
Goalw [function_def] "function(r) ==> function(rmap(r))";
694
c7d592f6312a ZF/ex/Ramsey,Rmap,misc.ML: modified for new definition of Pi(A,B)
lcp
parents: 515
diff changeset
    37
by (rtac (impI RS allI RS allI) 1);
1732
38776e927da8 Updated for new form of induction rules
paulson
parents: 1461
diff changeset
    38
by (etac rmap.induct 1);
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1732
diff changeset
    39
by (ALLGOALS Fast_tac);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 694
diff changeset
    40
qed "rmap_functional";
694
c7d592f6312a ZF/ex/Ramsey,Rmap,misc.ML: modified for new definition of Pi(A,B)
lcp
parents: 515
diff changeset
    41
16
0b033d50ca1c ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents:
diff changeset
    42
0b033d50ca1c ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents:
diff changeset
    43
(** If f is a function then rmap(f) behaves as expected. **)
0b033d50ca1c ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents:
diff changeset
    44
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5068
diff changeset
    45
Goal "f: A->B ==> rmap(f): list(A)->list(B)";
694
c7d592f6312a ZF/ex/Ramsey,Rmap,misc.ML: modified for new definition of Pi(A,B)
lcp
parents: 515
diff changeset
    46
by (asm_full_simp_tac 
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 2469
diff changeset
    47
    (simpset() addsimps [Pi_iff, rmap_rel_type, rmap_functional, rmap_total]) 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 694
diff changeset
    48
qed "rmap_fun_type";
16
0b033d50ca1c ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents:
diff changeset
    49
5068
fb28eaa07e01 isatool fixgoal;
wenzelm
parents: 4091
diff changeset
    50
Goalw [apply_def] "rmap(f)`Nil = Nil";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 2469
diff changeset
    51
by (fast_tac (claset() addIs [the_equality]) 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 694
diff changeset
    52
qed "rmap_Nil";
16
0b033d50ca1c ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents:
diff changeset
    53
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5068
diff changeset
    54
Goal "[| f: A->B;  x: A;  xs: list(A) |] ==> \
16
0b033d50ca1c ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents:
diff changeset
    55
\                   rmap(f) ` Cons(x,xs) = Cons(f`x, rmap(f)`xs)";
0b033d50ca1c ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents:
diff changeset
    56
by (rtac apply_equality 1);
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 496
diff changeset
    57
by (REPEAT (ares_tac ([apply_Pair, rmap_fun_type] @ rmap.intrs) 1));
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 694
diff changeset
    58
qed "rmap_Cons";