| author | nipkow | 
| Mon, 10 Jan 2000 16:06:43 +0100 | |
| changeset 8115 | c802042066e8 | 
| parent 6141 | a6922171b396 | 
| child 11316 | b4e71bd751e4 | 
| permissions | -rw-r--r-- | 
| 1461 | 1 | (* Title: ZF/ex/Rmap | 
| 16 
0b033d50ca1c
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 lcp parents: diff
changeset | 2 | ID: $Id$ | 
| 1461 | 3 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | 
| 515 | 4 | Copyright 1994 University of Cambridge | 
| 16 
0b033d50ca1c
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 lcp parents: diff
changeset | 5 | |
| 
0b033d50ca1c
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 lcp parents: diff
changeset | 6 | Inductive definition of an operator to "map" a relation over a list | 
| 
0b033d50ca1c
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 lcp parents: diff
changeset | 7 | *) | 
| 
0b033d50ca1c
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 lcp parents: diff
changeset | 8 | |
| 5137 | 9 | Goalw rmap.defs "r<=s ==> rmap(r) <= rmap(s)"; | 
| 16 
0b033d50ca1c
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 lcp parents: diff
changeset | 10 | by (rtac lfp_mono 1); | 
| 515 | 11 | by (REPEAT (rtac rmap.bnd_mono 1)); | 
| 16 
0b033d50ca1c
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 lcp parents: diff
changeset | 12 | by (REPEAT (ares_tac ([Sigma_mono, list_mono, domain_mono, range_mono] @ | 
| 1461 | 13 | basic_monos) 1)); | 
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
694diff
changeset | 14 | qed "rmap_mono"; | 
| 16 
0b033d50ca1c
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 lcp parents: diff
changeset | 15 | |
| 6141 | 16 | val Nil_rmap_case = rmap.mk_cases "<Nil,zs> : rmap(r)" | 
| 17 | and Cons_rmap_case = rmap.mk_cases "<Cons(x,xs),zs> : rmap(r)"; | |
| 16 
0b033d50ca1c
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 lcp parents: diff
changeset | 18 | |
| 2469 | 19 | AddIs rmap.intrs; | 
| 20 | AddSEs [Nil_rmap_case, Cons_rmap_case]; | |
| 16 
0b033d50ca1c
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 lcp parents: diff
changeset | 21 | |
| 5137 | 22 | Goal "r <= A*B ==> rmap(r) <= list(A)*list(B)"; | 
| 515 | 23 | by (rtac (rmap.dom_subset RS subset_trans) 1); | 
| 16 
0b033d50ca1c
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 lcp parents: diff
changeset | 24 | by (REPEAT (ares_tac [domain_rel_subset, range_rel_subset, | 
| 1461 | 25 | Sigma_mono, list_mono] 1)); | 
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
694diff
changeset | 26 | qed "rmap_rel_type"; | 
| 16 
0b033d50ca1c
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 lcp parents: diff
changeset | 27 | |
| 5147 
825877190618
More tidying and removal of "\!\!... from Goal commands
 paulson parents: 
5137diff
changeset | 28 | Goal "A <= domain(r) ==> list(A) <= domain(rmap(r))"; | 
| 1461 | 29 | by (rtac subsetI 1); | 
| 515 | 30 | by (etac list.induct 1); | 
| 2469 | 31 | by (ALLGOALS Fast_tac); | 
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
694diff
changeset | 32 | qed "rmap_total"; | 
| 16 
0b033d50ca1c
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 lcp parents: diff
changeset | 33 | |
| 5137 | 34 | Goalw [function_def] "function(r) ==> function(rmap(r))"; | 
| 694 
c7d592f6312a
ZF/ex/Ramsey,Rmap,misc.ML: modified for new definition of Pi(A,B)
 lcp parents: 
515diff
changeset | 35 | by (rtac (impI RS allI RS allI) 1); | 
| 1732 | 36 | by (etac rmap.induct 1); | 
| 2469 | 37 | by (ALLGOALS Fast_tac); | 
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
694diff
changeset | 38 | qed "rmap_functional"; | 
| 694 
c7d592f6312a
ZF/ex/Ramsey,Rmap,misc.ML: modified for new definition of Pi(A,B)
 lcp parents: 
515diff
changeset | 39 | |
| 16 
0b033d50ca1c
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 lcp parents: diff
changeset | 40 | |
| 
0b033d50ca1c
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 lcp parents: diff
changeset | 41 | (** If f is a function then rmap(f) behaves as expected. **) | 
| 
0b033d50ca1c
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 lcp parents: diff
changeset | 42 | |
| 5137 | 43 | Goal "f: A->B ==> rmap(f): list(A)->list(B)"; | 
| 6141 | 44 | by (asm_full_simp_tac (simpset() addsimps [Pi_iff, rmap_rel_type, | 
| 45 | rmap_functional, rmap_total]) 1); | |
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
694diff
changeset | 46 | qed "rmap_fun_type"; | 
| 16 
0b033d50ca1c
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 lcp parents: diff
changeset | 47 | |
| 5068 | 48 | Goalw [apply_def] "rmap(f)`Nil = Nil"; | 
| 5505 | 49 | by (Blast_tac 1); | 
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
694diff
changeset | 50 | qed "rmap_Nil"; | 
| 16 
0b033d50ca1c
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 lcp parents: diff
changeset | 51 | |
| 6141 | 52 | Goal "[| f: A->B; x: A; xs: list(A) |] \ | 
| 53 | \ ==> rmap(f) ` Cons(x,xs) = Cons(f`x, rmap(f)`xs)"; | |
| 16 
0b033d50ca1c
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 lcp parents: diff
changeset | 54 | by (rtac apply_equality 1); | 
| 515 | 55 | by (REPEAT (ares_tac ([apply_Pair, rmap_fun_type] @ rmap.intrs) 1)); | 
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
694diff
changeset | 56 | qed "rmap_Cons"; |