author | traytel |
Tue, 03 Mar 2015 19:08:04 +0100 | |
changeset 59580 | cbc38731d42f |
parent 58871 | c399ae4b836f |
child 60770 | 240563fbf41d |
permissions | -rw-r--r-- |
12610 | 1 |
(* Title: ZF/Induct/ListN.thy |
12088 | 2 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
3 |
Copyright 1994 University of Cambridge |
|
12610 | 4 |
*) |
12088 | 5 |
|
58871 | 6 |
section {* Lists of n elements *} |
12088 | 7 |
|
16417 | 8 |
theory ListN imports Main begin |
12088 | 9 |
|
12610 | 10 |
text {* |
11 |
Inductive definition of lists of @{text n} elements; see |
|
58623 | 12 |
@{cite "paulin-tlca"}. |
12610 | 13 |
*} |
14 |
||
15 |
consts listn :: "i=>i" |
|
12088 | 16 |
inductive |
12610 | 17 |
domains "listn(A)" \<subseteq> "nat \<times> list(A)" |
12560 | 18 |
intros |
12610 | 19 |
NilI: "<0,Nil> \<in> listn(A)" |
12560 | 20 |
ConsI: "[| a \<in> A; <n,l> \<in> listn(A) |] ==> <succ(n), Cons(a,l)> \<in> listn(A)" |
12610 | 21 |
type_intros nat_typechecks list.intros |
12560 | 22 |
|
23 |
||
24 |
lemma list_into_listn: "l \<in> list(A) ==> <length(l),l> \<in> listn(A)" |
|
18415 | 25 |
by (induct set: list) (simp_all add: listn.intros) |
12560 | 26 |
|
46822
95f1e700b712
mathematical symbols for Isabelle/ZF example theories
paulson
parents:
35762
diff
changeset
|
27 |
lemma listn_iff: "<n,l> \<in> listn(A) \<longleftrightarrow> l \<in> list(A) & length(l)=n" |
12610 | 28 |
apply (rule iffI) |
29 |
apply (erule listn.induct) |
|
30 |
apply auto |
|
31 |
apply (blast intro: list_into_listn) |
|
32 |
done |
|
12560 | 33 |
|
34 |
lemma listn_image_eq: "listn(A)``{n} = {l \<in> list(A). length(l)=n}" |
|
12610 | 35 |
apply (rule equality_iffI) |
36 |
apply (simp add: listn_iff separation image_singleton_iff) |
|
37 |
done |
|
12560 | 38 |
|
39 |
lemma listn_mono: "A \<subseteq> B ==> listn(A) \<subseteq> listn(B)" |
|
12610 | 40 |
apply (unfold listn.defs) |
41 |
apply (rule lfp_mono) |
|
42 |
apply (rule listn.bnd_mono)+ |
|
43 |
apply (assumption | rule univ_mono Sigma_mono list_mono basic_monos)+ |
|
44 |
done |
|
12560 | 45 |
|
46 |
lemma listn_append: |
|
12610 | 47 |
"[| <n,l> \<in> listn(A); <n',l'> \<in> listn(A) |] ==> <n#+n', l@l'> \<in> listn(A)" |
48 |
apply (erule listn.induct) |
|
49 |
apply (frule listn.dom_subset [THEN subsetD]) |
|
50 |
apply (simp_all add: listn.intros) |
|
51 |
done |
|
12560 | 52 |
|
12610 | 53 |
inductive_cases |
54 |
Nil_listn_case: "<i,Nil> \<in> listn(A)" |
|
55 |
and Cons_listn_case: "<i,Cons(x,l)> \<in> listn(A)" |
|
12560 | 56 |
|
12610 | 57 |
inductive_cases |
58 |
zero_listn_case: "<0,l> \<in> listn(A)" |
|
59 |
and succ_listn_case: "<succ(i),l> \<in> listn(A)" |
|
12088 | 60 |
|
61 |
end |