author | haftmann |
Sun, 27 Jul 2025 17:52:06 +0200 | |
changeset 82909 | e4fae2227594 |
parent 67613 | ce654b0e6d69 |
permissions | -rw-r--r-- |
42151 | 1 |
(* Title: HOL/HOLCF/IOA/NTP/Lemmas.thy |
3073
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 |
||
62002 | 9 |
subsubsection \<open>Logic\<close> |
19739 | 10 |
|
67613 | 11 |
lemma neg_flip: "(X = (\<not> Y)) = ((\<not>X) = Y)" |
19739 | 12 |
by blast |
13 |
||
14 |
||
62002 | 15 |
subsection \<open>Sets\<close> |
19739 | 16 |
|
17 |
lemma set_lemmas: |
|
67613 | 18 |
"f(x) \<in> (\<Union>x. {f(x)})" |
19 |
"f x y \<in> (\<Union>x y. {f x y})" |
|
20 |
"\<And>a. (\<forall>x. a \<noteq> f(x)) \<Longrightarrow> a \<notin> (\<Union>x. {f(x)})" |
|
21 |
"\<And>a. (\<forall>x y. a \<noteq> f x y) \<Longrightarrow> a \<notin> (\<Union>x y. {f x y})" |
|
19739 | 22 |
by auto |
23 |
||
24 |
||
62002 | 25 |
subsection \<open>Arithmetic\<close> |
19739 | 26 |
|
27 |
lemma pred_suc: "0<x ==> (x - 1 = y) = (x = Suc(y))" |
|
63648 | 28 |
by (simp add: diff_Suc split: nat.split) |
19739 | 29 |
|
30 |
lemmas [simp] = hd_append set_lemmas |
|
31 |
||
17244 | 32 |
end |