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