src/ZF/Induct/Rmap.thy
author huffman
Tue, 02 Mar 2010 09:54:50 -0800
changeset 35512 d1ef88d7de5a
parent 16417 9bc16273c2d4
child 35762 af3ff2ba4c54
permissions -rw-r--r--
remove dead code
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
12610
8b9845807f77 tuned document sources;
wenzelm
parents: 12560
diff changeset
     1
(*  Title:      ZF/Induct/Rmap.thy
12088
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
12610
8b9845807f77 tuned document sources;
wenzelm
parents: 12560
diff changeset
     5
*)
12088
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
     6
12610
8b9845807f77 tuned document sources;
wenzelm
parents: 12560
diff changeset
     7
header {* An operator to ``map'' a relation over a list *}
12088
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
     8
16417
9bc16273c2d4 migrated theory headers to new format
haftmann
parents: 12610
diff changeset
     9
theory Rmap imports Main begin
12088
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    10
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    11
consts
12560
5820841f21fd converted some ZF/Induct examples to Isar
paulson
parents: 12088
diff changeset
    12
  rmap :: "i=>i"
12088
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    13
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    14
inductive
12610
8b9845807f77 tuned document sources;
wenzelm
parents: 12560
diff changeset
    15
  domains "rmap(r)" \<subseteq> "list(domain(r)) \<times> list(range(r))"
12560
5820841f21fd converted some ZF/Induct examples to Isar
paulson
parents: 12088
diff changeset
    16
  intros
5820841f21fd converted some ZF/Induct examples to Isar
paulson
parents: 12088
diff changeset
    17
    NilI:  "<Nil,Nil> \<in> rmap(r)"
5820841f21fd converted some ZF/Induct examples to Isar
paulson
parents: 12088
diff changeset
    18
12610
8b9845807f77 tuned document sources;
wenzelm
parents: 12560
diff changeset
    19
    ConsI: "[| <x,y>: r;  <xs,ys> \<in> rmap(r) |]
12560
5820841f21fd converted some ZF/Induct examples to Isar
paulson
parents: 12088
diff changeset
    20
            ==> <Cons(x,xs), Cons(y,ys)> \<in> rmap(r)"
5820841f21fd converted some ZF/Induct examples to Isar
paulson
parents: 12088
diff changeset
    21
5820841f21fd converted some ZF/Induct examples to Isar
paulson
parents: 12088
diff changeset
    22
  type_intros domainI rangeI list.intros
5820841f21fd converted some ZF/Induct examples to Isar
paulson
parents: 12088
diff changeset
    23
5820841f21fd converted some ZF/Induct examples to Isar
paulson
parents: 12088
diff changeset
    24
lemma rmap_mono: "r \<subseteq> s ==> rmap(r) \<subseteq> rmap(s)"
12610
8b9845807f77 tuned document sources;
wenzelm
parents: 12560
diff changeset
    25
  apply (unfold rmap.defs)
8b9845807f77 tuned document sources;
wenzelm
parents: 12560
diff changeset
    26
  apply (rule lfp_mono)
8b9845807f77 tuned document sources;
wenzelm
parents: 12560
diff changeset
    27
    apply (rule rmap.bnd_mono)+
8b9845807f77 tuned document sources;
wenzelm
parents: 12560
diff changeset
    28
  apply (assumption | rule Sigma_mono list_mono domain_mono range_mono basic_monos)+
8b9845807f77 tuned document sources;
wenzelm
parents: 12560
diff changeset
    29
  done
12560
5820841f21fd converted some ZF/Induct examples to Isar
paulson
parents: 12088
diff changeset
    30
12610
8b9845807f77 tuned document sources;
wenzelm
parents: 12560
diff changeset
    31
inductive_cases
8b9845807f77 tuned document sources;
wenzelm
parents: 12560
diff changeset
    32
      Nil_rmap_case [elim!]: "<Nil,zs> \<in> rmap(r)"
8b9845807f77 tuned document sources;
wenzelm
parents: 12560
diff changeset
    33
  and Cons_rmap_case [elim!]: "<Cons(x,xs),zs> \<in> rmap(r)"
12560
5820841f21fd converted some ZF/Induct examples to Isar
paulson
parents: 12088
diff changeset
    34
5820841f21fd converted some ZF/Induct examples to Isar
paulson
parents: 12088
diff changeset
    35
declare rmap.intros [intro]
5820841f21fd converted some ZF/Induct examples to Isar
paulson
parents: 12088
diff changeset
    36
12610
8b9845807f77 tuned document sources;
wenzelm
parents: 12560
diff changeset
    37
lemma rmap_rel_type: "r \<subseteq> A \<times> B ==> rmap(r) \<subseteq> list(A) \<times> list(B)"
8b9845807f77 tuned document sources;
wenzelm
parents: 12560
diff changeset
    38
  apply (rule rmap.dom_subset [THEN subset_trans])
8b9845807f77 tuned document sources;
wenzelm
parents: 12560
diff changeset
    39
  apply (assumption |
8b9845807f77 tuned document sources;
wenzelm
parents: 12560
diff changeset
    40
    rule domain_rel_subset range_rel_subset Sigma_mono list_mono)+
8b9845807f77 tuned document sources;
wenzelm
parents: 12560
diff changeset
    41
  done
12088
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    42
12560
5820841f21fd converted some ZF/Induct examples to Isar
paulson
parents: 12088
diff changeset
    43
lemma rmap_total: "A \<subseteq> domain(r) ==> list(A) \<subseteq> domain(rmap(r))"
12610
8b9845807f77 tuned document sources;
wenzelm
parents: 12560
diff changeset
    44
  apply (rule subsetI)
8b9845807f77 tuned document sources;
wenzelm
parents: 12560
diff changeset
    45
  apply (erule list.induct)
8b9845807f77 tuned document sources;
wenzelm
parents: 12560
diff changeset
    46
   apply blast+
8b9845807f77 tuned document sources;
wenzelm
parents: 12560
diff changeset
    47
  done
12560
5820841f21fd converted some ZF/Induct examples to Isar
paulson
parents: 12088
diff changeset
    48
5820841f21fd converted some ZF/Induct examples to Isar
paulson
parents: 12088
diff changeset
    49
lemma rmap_functional: "function(r) ==> function(rmap(r))"
12610
8b9845807f77 tuned document sources;
wenzelm
parents: 12560
diff changeset
    50
  apply (unfold function_def)
8b9845807f77 tuned document sources;
wenzelm
parents: 12560
diff changeset
    51
  apply (rule impI [THEN allI, THEN allI])
8b9845807f77 tuned document sources;
wenzelm
parents: 12560
diff changeset
    52
  apply (erule rmap.induct)
8b9845807f77 tuned document sources;
wenzelm
parents: 12560
diff changeset
    53
   apply blast+
8b9845807f77 tuned document sources;
wenzelm
parents: 12560
diff changeset
    54
  done
12088
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    55
12610
8b9845807f77 tuned document sources;
wenzelm
parents: 12560
diff changeset
    56
text {*
8b9845807f77 tuned document sources;
wenzelm
parents: 12560
diff changeset
    57
  \medskip If @{text f} is a function then @{text "rmap(f)"} behaves
8b9845807f77 tuned document sources;
wenzelm
parents: 12560
diff changeset
    58
  as expected.
8b9845807f77 tuned document sources;
wenzelm
parents: 12560
diff changeset
    59
*}
12560
5820841f21fd converted some ZF/Induct examples to Isar
paulson
parents: 12088
diff changeset
    60
5820841f21fd converted some ZF/Induct examples to Isar
paulson
parents: 12088
diff changeset
    61
lemma rmap_fun_type: "f \<in> A->B ==> rmap(f): list(A)->list(B)"
12610
8b9845807f77 tuned document sources;
wenzelm
parents: 12560
diff changeset
    62
  by (simp add: Pi_iff rmap_rel_type rmap_functional rmap_total)
12560
5820841f21fd converted some ZF/Induct examples to Isar
paulson
parents: 12088
diff changeset
    63
5820841f21fd converted some ZF/Induct examples to Isar
paulson
parents: 12088
diff changeset
    64
lemma rmap_Nil: "rmap(f)`Nil = Nil"
12610
8b9845807f77 tuned document sources;
wenzelm
parents: 12560
diff changeset
    65
  by (unfold apply_def) blast
12560
5820841f21fd converted some ZF/Induct examples to Isar
paulson
parents: 12088
diff changeset
    66
12610
8b9845807f77 tuned document sources;
wenzelm
parents: 12560
diff changeset
    67
lemma rmap_Cons: "[| f \<in> A->B;  x \<in> A;  xs: list(A) |]
12560
5820841f21fd converted some ZF/Induct examples to Isar
paulson
parents: 12088
diff changeset
    68
      ==> rmap(f) ` Cons(x,xs) = Cons(f`x, rmap(f)`xs)"
12610
8b9845807f77 tuned document sources;
wenzelm
parents: 12560
diff changeset
    69
  by (blast intro: apply_equality apply_Pair rmap_fun_type rmap.intros)
12088
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    70
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    71
end