| 9907 |      1 | (*  Title:      ZF/Update.ML
 | 
| 5157 |      2 |     ID:         $Id$
 | 
|  |      3 |     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
 | 
|  |      4 |     Copyright   1998  University of Cambridge
 | 
|  |      5 | 
 | 
|  |      6 | Function updates: like theory Map, but for ordinary functions
 | 
|  |      7 | *)
 | 
|  |      8 | 
 | 
| 6068 |      9 | Goal "f(x:=y) ` z = (if z=x then y else f`z)";
 | 
| 5157 |     10 | by (simp_tac (simpset() addsimps [update_def]) 1);
 | 
|  |     11 | by (case_tac "z : domain(f)" 1);
 | 
|  |     12 | by (Asm_simp_tac 1);
 | 
|  |     13 | by (asm_simp_tac (simpset() addsimps [apply_0]) 1);
 | 
|  |     14 | qed "update_apply";
 | 
|  |     15 | Addsimps [update_apply];
 | 
|  |     16 | 
 | 
|  |     17 | Goalw [update_def] "[| f`x = y;  f: Pi(A,B);  x: A |] ==> f(x:=y) = f";
 | 
|  |     18 | by (asm_simp_tac (simpset() addsimps [domain_of_fun, cons_absorb]) 1);
 | 
|  |     19 | by (rtac fun_extension 1);
 | 
|  |     20 | by (best_tac (claset() addIs [apply_type, if_type, lam_type]) 1);
 | 
| 5168 |     21 | by (assume_tac 1);
 | 
| 5157 |     22 | by (Asm_simp_tac 1);
 | 
|  |     23 | qed "update_idem";
 | 
|  |     24 | 
 | 
|  |     25 | 
 | 
|  |     26 | (* [| f: Pi(A, B); x:A |] ==> f(x := f`x) = f *)
 | 
|  |     27 | Addsimps [refl RS update_idem];
 | 
|  |     28 | 
 | 
|  |     29 | Goalw [update_def] "domain(f(x:=y)) = cons(x, domain(f))";
 | 
|  |     30 | by (Asm_simp_tac 1);
 | 
|  |     31 | qed "domain_update";
 | 
|  |     32 | Addsimps [domain_update];
 | 
|  |     33 | 
 | 
| 6048 |     34 | Goalw [update_def] "[| f: A -> B;  x : A;  y: B |] ==> f(x:=y) : A -> B";
 | 
|  |     35 | by (asm_simp_tac (simpset() addsimps [domain_of_fun, cons_absorb, 
 | 
|  |     36 | 				      apply_funtype, lam_type]) 1);
 | 
|  |     37 | qed "update_type";
 | 
|  |     38 | 
 | 
|  |     39 | 
 |