src/ZF/Induct/ListN.thy
author huffman
Tue, 02 Mar 2010 09:54:50 -0800
changeset 35512 d1ef88d7de5a
parent 18415 eb68dc98bda2
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/ListN.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 {* Lists of n elements *}
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 ListN imports Main begin
12088
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    10
12610
8b9845807f77 tuned document sources;
wenzelm
parents: 12560
diff changeset
    11
text {*
8b9845807f77 tuned document sources;
wenzelm
parents: 12560
diff changeset
    12
  Inductive definition of lists of @{text n} elements; see
8b9845807f77 tuned document sources;
wenzelm
parents: 12560
diff changeset
    13
  \cite{paulin-tlca}.
8b9845807f77 tuned document sources;
wenzelm
parents: 12560
diff changeset
    14
*}
8b9845807f77 tuned document sources;
wenzelm
parents: 12560
diff changeset
    15
8b9845807f77 tuned document sources;
wenzelm
parents: 12560
diff changeset
    16
consts listn :: "i=>i"
12088
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    17
inductive
12610
8b9845807f77 tuned document sources;
wenzelm
parents: 12560
diff changeset
    18
  domains "listn(A)" \<subseteq> "nat \<times> list(A)"
12560
5820841f21fd converted some ZF/Induct examples to Isar
paulson
parents: 12088
diff changeset
    19
  intros
12610
8b9845807f77 tuned document sources;
wenzelm
parents: 12560
diff changeset
    20
    NilI: "<0,Nil> \<in> listn(A)"
12560
5820841f21fd converted some ZF/Induct examples to Isar
paulson
parents: 12088
diff changeset
    21
    ConsI: "[| a \<in> A; <n,l> \<in> listn(A) |] ==> <succ(n), Cons(a,l)> \<in> listn(A)"
12610
8b9845807f77 tuned document sources;
wenzelm
parents: 12560
diff changeset
    22
  type_intros nat_typechecks list.intros
12560
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
5820841f21fd converted some ZF/Induct examples to Isar
paulson
parents: 12088
diff changeset
    25
lemma list_into_listn: "l \<in> list(A) ==> <length(l),l> \<in> listn(A)"
18415
eb68dc98bda2 improved proofs;
wenzelm
parents: 16417
diff changeset
    26
  by (induct set: list) (simp_all add: listn.intros)
12560
5820841f21fd converted some ZF/Induct examples to Isar
paulson
parents: 12088
diff changeset
    27
5820841f21fd converted some ZF/Induct examples to Isar
paulson
parents: 12088
diff changeset
    28
lemma listn_iff: "<n,l> \<in> listn(A) <-> l \<in> list(A) & length(l)=n"
12610
8b9845807f77 tuned document sources;
wenzelm
parents: 12560
diff changeset
    29
  apply (rule iffI)
8b9845807f77 tuned document sources;
wenzelm
parents: 12560
diff changeset
    30
   apply (erule listn.induct)
8b9845807f77 tuned document sources;
wenzelm
parents: 12560
diff changeset
    31
    apply auto
8b9845807f77 tuned document sources;
wenzelm
parents: 12560
diff changeset
    32
  apply (blast intro: list_into_listn)
8b9845807f77 tuned document sources;
wenzelm
parents: 12560
diff changeset
    33
  done
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
lemma listn_image_eq: "listn(A)``{n} = {l \<in> list(A). length(l)=n}"
12610
8b9845807f77 tuned document sources;
wenzelm
parents: 12560
diff changeset
    36
  apply (rule equality_iffI)
8b9845807f77 tuned document sources;
wenzelm
parents: 12560
diff changeset
    37
  apply (simp add: listn_iff separation image_singleton_iff)
8b9845807f77 tuned document sources;
wenzelm
parents: 12560
diff changeset
    38
  done
12560
5820841f21fd converted some ZF/Induct examples to Isar
paulson
parents: 12088
diff changeset
    39
5820841f21fd converted some ZF/Induct examples to Isar
paulson
parents: 12088
diff changeset
    40
lemma listn_mono: "A \<subseteq> B ==> listn(A) \<subseteq> listn(B)"
12610
8b9845807f77 tuned document sources;
wenzelm
parents: 12560
diff changeset
    41
  apply (unfold listn.defs)
8b9845807f77 tuned document sources;
wenzelm
parents: 12560
diff changeset
    42
  apply (rule lfp_mono)
8b9845807f77 tuned document sources;
wenzelm
parents: 12560
diff changeset
    43
    apply (rule listn.bnd_mono)+
8b9845807f77 tuned document sources;
wenzelm
parents: 12560
diff changeset
    44
  apply (assumption | rule univ_mono Sigma_mono list_mono basic_monos)+
8b9845807f77 tuned document sources;
wenzelm
parents: 12560
diff changeset
    45
  done
12560
5820841f21fd converted some ZF/Induct examples to Isar
paulson
parents: 12088
diff changeset
    46
5820841f21fd converted some ZF/Induct examples to Isar
paulson
parents: 12088
diff changeset
    47
lemma listn_append:
12610
8b9845807f77 tuned document sources;
wenzelm
parents: 12560
diff changeset
    48
    "[| <n,l> \<in> listn(A); <n',l'> \<in> listn(A) |] ==> <n#+n', l@l'> \<in> listn(A)"
8b9845807f77 tuned document sources;
wenzelm
parents: 12560
diff changeset
    49
  apply (erule listn.induct)
8b9845807f77 tuned document sources;
wenzelm
parents: 12560
diff changeset
    50
   apply (frule listn.dom_subset [THEN subsetD])
8b9845807f77 tuned document sources;
wenzelm
parents: 12560
diff changeset
    51
   apply (simp_all add: listn.intros)
8b9845807f77 tuned document sources;
wenzelm
parents: 12560
diff changeset
    52
  done
12560
5820841f21fd converted some ZF/Induct examples to Isar
paulson
parents: 12088
diff changeset
    53
12610
8b9845807f77 tuned document sources;
wenzelm
parents: 12560
diff changeset
    54
inductive_cases
8b9845807f77 tuned document sources;
wenzelm
parents: 12560
diff changeset
    55
      Nil_listn_case: "<i,Nil> \<in> listn(A)"
8b9845807f77 tuned document sources;
wenzelm
parents: 12560
diff changeset
    56
  and Cons_listn_case: "<i,Cons(x,l)> \<in> listn(A)"
12560
5820841f21fd converted some ZF/Induct examples to Isar
paulson
parents: 12088
diff changeset
    57
12610
8b9845807f77 tuned document sources;
wenzelm
parents: 12560
diff changeset
    58
inductive_cases
8b9845807f77 tuned document sources;
wenzelm
parents: 12560
diff changeset
    59
      zero_listn_case: "<0,l> \<in> listn(A)"
8b9845807f77 tuned document sources;
wenzelm
parents: 12560
diff changeset
    60
  and succ_listn_case: "<succ(i),l> \<in> listn(A)"
12088
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    61
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    62
end