| author | wenzelm | 
| Fri, 04 Aug 2000 22:55:08 +0200 | |
| changeset 9526 | e20323caff47 | 
| parent 9458 | c613cd06d5cf | 
| child 9541 | d17c0b34d5c8 | 
| permissions | -rw-r--r-- | 
| 8745 | 1 | (*<*) | 
| 2 | theory let_rewr = Main:; | |
| 3 | (*>*) | |
| 4 | lemma "(let xs = [] in xs@ys@xs) = ys"; | |
| 9458 | 5 | by(simp add: Let_def); | 
| 8745 | 6 | |
| 7 | text{*
 | |
| 8 | If, in a particular context, there is no danger of a combinatorial explosion | |
| 8771 | 9 | of nested \isa{let}s one could even add \isa{Let_def} permanently:
 | 
| 8745 | 10 | *} | 
| 11 | theorems [simp] = Let_def; | |
| 12 | (*<*) | |
| 13 | end | |
| 14 | (*>*) |