12610
|
1 |
(* Title: ZF/Induct/Rmap.thy
|
12088
|
2 |
ID: $Id$
|
|
3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory
|
|
4 |
Copyright 1994 University of Cambridge
|
12610
|
5 |
*)
|
12088
|
6 |
|
12610
|
7 |
header {* An operator to ``map'' a relation over a list *}
|
12088
|
8 |
|
16417
|
9 |
theory Rmap imports Main begin
|
12088
|
10 |
|
|
11 |
consts
|
12560
|
12 |
rmap :: "i=>i"
|
12088
|
13 |
|
|
14 |
inductive
|
12610
|
15 |
domains "rmap(r)" \<subseteq> "list(domain(r)) \<times> list(range(r))"
|
12560
|
16 |
intros
|
|
17 |
NilI: "<Nil,Nil> \<in> rmap(r)"
|
|
18 |
|
12610
|
19 |
ConsI: "[| <x,y>: r; <xs,ys> \<in> rmap(r) |]
|
12560
|
20 |
==> <Cons(x,xs), Cons(y,ys)> \<in> rmap(r)"
|
|
21 |
|
|
22 |
type_intros domainI rangeI list.intros
|
|
23 |
|
|
24 |
lemma rmap_mono: "r \<subseteq> s ==> rmap(r) \<subseteq> rmap(s)"
|
12610
|
25 |
apply (unfold rmap.defs)
|
|
26 |
apply (rule lfp_mono)
|
|
27 |
apply (rule rmap.bnd_mono)+
|
|
28 |
apply (assumption | rule Sigma_mono list_mono domain_mono range_mono basic_monos)+
|
|
29 |
done
|
12560
|
30 |
|
12610
|
31 |
inductive_cases
|
|
32 |
Nil_rmap_case [elim!]: "<Nil,zs> \<in> rmap(r)"
|
|
33 |
and Cons_rmap_case [elim!]: "<Cons(x,xs),zs> \<in> rmap(r)"
|
12560
|
34 |
|
|
35 |
declare rmap.intros [intro]
|
|
36 |
|
12610
|
37 |
lemma rmap_rel_type: "r \<subseteq> A \<times> B ==> rmap(r) \<subseteq> list(A) \<times> list(B)"
|
|
38 |
apply (rule rmap.dom_subset [THEN subset_trans])
|
|
39 |
apply (assumption |
|
|
40 |
rule domain_rel_subset range_rel_subset Sigma_mono list_mono)+
|
|
41 |
done
|
12088
|
42 |
|
12560
|
43 |
lemma rmap_total: "A \<subseteq> domain(r) ==> list(A) \<subseteq> domain(rmap(r))"
|
12610
|
44 |
apply (rule subsetI)
|
|
45 |
apply (erule list.induct)
|
|
46 |
apply blast+
|
|
47 |
done
|
12560
|
48 |
|
|
49 |
lemma rmap_functional: "function(r) ==> function(rmap(r))"
|
12610
|
50 |
apply (unfold function_def)
|
|
51 |
apply (rule impI [THEN allI, THEN allI])
|
|
52 |
apply (erule rmap.induct)
|
|
53 |
apply blast+
|
|
54 |
done
|
12088
|
55 |
|
12610
|
56 |
text {*
|
|
57 |
\medskip If @{text f} is a function then @{text "rmap(f)"} behaves
|
|
58 |
as expected.
|
|
59 |
*}
|
12560
|
60 |
|
|
61 |
lemma rmap_fun_type: "f \<in> A->B ==> rmap(f): list(A)->list(B)"
|
12610
|
62 |
by (simp add: Pi_iff rmap_rel_type rmap_functional rmap_total)
|
12560
|
63 |
|
|
64 |
lemma rmap_Nil: "rmap(f)`Nil = Nil"
|
12610
|
65 |
by (unfold apply_def) blast
|
12560
|
66 |
|
12610
|
67 |
lemma rmap_Cons: "[| f \<in> A->B; x \<in> A; xs: list(A) |]
|
12560
|
68 |
==> rmap(f) ` Cons(x,xs) = Cons(f`x, rmap(f)`xs)"
|
12610
|
69 |
by (blast intro: apply_equality apply_Pair rmap_fun_type rmap.intros)
|
12088
|
70 |
|
|
71 |
end
|