| author | blanchet | 
| Wed, 14 Apr 2010 18:23:51 +0200 | |
| changeset 36142 | f5e15e9aae10 | 
| parent 35174 | e15040ae75d7 | 
| permissions | -rw-r--r-- | 
| 3073 
88366253a09a
Old NTP files now running under the IOA meta theory based on HOLCF;
 mueller parents: diff
changeset | 1 | (* Title: HOL/IOA/NTP/Lemmas.thy | 
| 
88366253a09a
Old NTP files now running under the IOA meta theory based on HOLCF;
 mueller parents: diff
changeset | 2 | Author: Tobias Nipkow & Konrad Slind | 
| 
88366253a09a
Old NTP files now running under the IOA meta theory based on HOLCF;
 mueller parents: diff
changeset | 3 | *) | 
| 
88366253a09a
Old NTP files now running under the IOA meta theory based on HOLCF;
 mueller parents: diff
changeset | 4 | |
| 17244 | 5 | theory Lemmas | 
| 6 | imports Main | |
| 7 | begin | |
| 8 | ||
| 19739 | 9 | subsubsection {* Logic *}
 | 
| 10 | ||
| 11 | lemma neg_flip: "(X = (~ Y)) = ((~X) = Y)" | |
| 12 | by blast | |
| 13 | ||
| 14 | ||
| 15 | subsection {* Sets *}
 | |
| 16 | ||
| 17 | lemma set_lemmas: | |
| 18 |   "f(x) : (UN x. {f(x)})"
 | |
| 19 |   "f x y : (UN x y. {f x y})"
 | |
| 20 |   "!!a. (!x. a ~= f(x)) ==> a ~: (UN x. {f(x)})"
 | |
| 21 |   "!!a. (!x y. a ~= f x y) ==> a ~: (UN x y. {f x y})"
 | |
| 22 | by auto | |
| 23 | ||
| 24 | ||
| 25 | subsection {* Arithmetic *}
 | |
| 26 | ||
| 27 | lemma pred_suc: "0<x ==> (x - 1 = y) = (x = Suc(y))" | |
| 28 | by (simp add: diff_Suc split add: nat.split) | |
| 29 | ||
| 30 | lemmas [simp] = hd_append set_lemmas | |
| 31 | ||
| 17244 | 32 | end |