src/ZF/Induct/Rmap.ML
author wenzelm
Thu, 15 Nov 2001 18:20:13 +0100
changeset 12207 4dff931b852f
parent 12088 6f463d16cbd0
permissions -rw-r--r--
added Induct/Binary_Trees.thy, Induct/Tree_Forest (converted from former ex/TF.ML ex/TF.thy ex/Term.ML ex/Term.thy);
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
12088
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
     1
(*  Title:      ZF/ex/Rmap
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
     2
    ID:         $Id$
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
     4
    Copyright   1994  University of Cambridge
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
     5
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
     6
Inductive definition of an operator to "map" a relation over a list
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
     7
*)
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
     8
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
     9
Goalw rmap.defs "r \\<subseteq> s ==> rmap(r) \\<subseteq> rmap(s)";
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    10
by (rtac lfp_mono 1);
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    11
by (REPEAT (rtac rmap.bnd_mono 1));
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    12
by (REPEAT (ares_tac ([Sigma_mono, list_mono, domain_mono, range_mono] @ 
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    13
                      basic_monos) 1));
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    14
qed "rmap_mono";
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    15
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    16
val Nil_rmap_case  = rmap.mk_cases "<Nil,zs> \\<in> rmap(r)"
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    17
and Cons_rmap_case = rmap.mk_cases "<Cons(x,xs),zs> \\<in> rmap(r)";
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    18
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    19
AddIs  rmap.intrs;
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    20
AddSEs [Nil_rmap_case, Cons_rmap_case];
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    21
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    22
Goal "r \\<subseteq> A*B ==> rmap(r) \\<subseteq> list(A)*list(B)";
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    23
by (rtac (rmap.dom_subset RS subset_trans) 1);
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    24
by (REPEAT (ares_tac [domain_rel_subset, range_rel_subset,
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    25
                      Sigma_mono, list_mono] 1));
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    26
qed "rmap_rel_type";
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    27
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    28
Goal "A \\<subseteq> domain(r) ==> list(A) \\<subseteq> domain(rmap(r))";
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    29
by (rtac subsetI 1);
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    30
by (etac list.induct 1);
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    31
by (ALLGOALS Fast_tac);
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    32
qed "rmap_total";
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    33
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    34
Goalw [function_def] "function(r) ==> function(rmap(r))";
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    35
by (rtac (impI RS allI RS allI) 1);
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    36
by (etac rmap.induct 1);
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    37
by (ALLGOALS Fast_tac);
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    38
qed "rmap_functional";
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    39
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    40
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    41
(** If f is a function then rmap(f) behaves as expected. **)
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    42
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    43
Goal "f \\<in> A->B ==> rmap(f): list(A)->list(B)";
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    44
by (asm_full_simp_tac (simpset() addsimps [Pi_iff, rmap_rel_type,
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    45
					   rmap_functional, rmap_total]) 1);
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    46
qed "rmap_fun_type";
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    47
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    48
Goalw [apply_def] "rmap(f)`Nil = Nil";
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    49
by (Blast_tac 1);
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    50
qed "rmap_Nil";
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    51
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    52
Goal "[| f \\<in> A->B;  x \\<in> A;  xs: list(A) |]  \
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    53
\     ==> rmap(f) ` Cons(x,xs) = Cons(f`x, rmap(f)`xs)";
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    54
by (rtac apply_equality 1);
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    55
by (REPEAT (ares_tac ([apply_Pair, rmap_fun_type] @ rmap.intrs) 1));
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    56
qed "rmap_Cons";