| author | wenzelm | 
| Sun, 11 Sep 2011 15:20:09 +0200 | |
| changeset 44878 | 1cbe20966cdb | 
| parent 44740 | a2940bc24bad | 
| permissions | -rw-r--r-- | 
| 14620 
1be590fd2422
Minor cleanup of headers and some speedup of the HOL4 import.
 skalberg parents: 
14516diff
changeset | 1 | (* Title: HOL/Import/HOL4Compat.thy | 
| 
1be590fd2422
Minor cleanup of headers and some speedup of the HOL4 import.
 skalberg parents: 
14516diff
changeset | 2 | Author: Sebastian Skalberg (TU Muenchen) | 
| 
1be590fd2422
Minor cleanup of headers and some speedup of the HOL4 import.
 skalberg parents: 
14516diff
changeset | 3 | *) | 
| 
1be590fd2422
Minor cleanup of headers and some speedup of the HOL4 import.
 skalberg parents: 
14516diff
changeset | 4 | |
| 30660 | 5 | theory HOL4Compat | 
| 41413 
64cd30d6b0b8
explicit file specifications -- avoid secondary load path;
 wenzelm parents: 
40607diff
changeset | 6 | imports | 
| 
64cd30d6b0b8
explicit file specifications -- avoid secondary load path;
 wenzelm parents: 
40607diff
changeset | 7 | HOL4Setup | 
| 
64cd30d6b0b8
explicit file specifications -- avoid secondary load path;
 wenzelm parents: 
40607diff
changeset | 8 | Complex_Main | 
| 
64cd30d6b0b8
explicit file specifications -- avoid secondary load path;
 wenzelm parents: 
40607diff
changeset | 9 | "~~/src/HOL/Old_Number_Theory/Primes" | 
| 
64cd30d6b0b8
explicit file specifications -- avoid secondary load path;
 wenzelm parents: 
40607diff
changeset | 10 | "~~/src/HOL/Library/ContNotDenum" | 
| 19064 | 11 | begin | 
| 14516 | 12 | |
| 37596 | 13 | abbreviation (input) mem (infixl "mem" 55) where "x mem xs \<equiv> List.member xs x" | 
| 30660 | 14 | no_notation differentiable (infixl "differentiable" 60) | 
| 15 | no_notation sums (infixr "sums" 80) | |
| 16 | ||
| 14516 | 17 | lemma EXISTS_UNIQUE_DEF: "(Ex1 P) = (Ex P & (ALL x y. P x & P y --> (x = y)))" | 
| 18 | by auto | |
| 19 | ||
| 20 | lemma COND_DEF:"(If b t f) = (@x. ((b = True) --> (x = t)) & ((b = False) --> (x = f)))" | |
| 21 | by auto | |
| 22 | ||
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
32960diff
changeset | 23 | definition LET :: "['a \<Rightarrow> 'b,'a] \<Rightarrow> 'b" where | 
| 14516 | 24 | "LET f s == f s" | 
| 25 | ||
| 26 | lemma [hol4rew]: "LET f s = Let s f" | |
| 27 | by (simp add: LET_def Let_def) | |
| 28 | ||
| 29 | lemmas [hol4rew] = ONE_ONE_rew | |
| 30 | ||
| 31 | lemma bool_case_DEF: "(bool_case x y b) = (if b then x else y)" | |
| 30660 | 32 | by simp | 
| 14516 | 33 | |
| 34 | lemma INR_INL_11: "(ALL y x. (Inl x = Inl y) = (x = y)) & (ALL y x. (Inr x = Inr y) = (x = y))" | |
| 35 | by safe | |
| 36 | ||
| 17188 | 37 | (*lemma INL_neq_INR: "ALL v1 v2. Sum_Type.Inr v2 ~= Sum_Type.Inl v1" | 
| 38 | by simp*) | |
| 39 | ||
| 39246 | 40 | primrec ISL :: "'a + 'b => bool" where | 
| 14516 | 41 | "ISL (Inl x) = True" | 
| 39246 | 42 | | "ISL (Inr x) = False" | 
| 14516 | 43 | |
| 39246 | 44 | primrec ISR :: "'a + 'b => bool" where | 
| 14516 | 45 | "ISR (Inl x) = False" | 
| 39246 | 46 | | "ISR (Inr x) = True" | 
| 14516 | 47 | |
| 48 | lemma ISL: "(ALL x. ISL (Inl x)) & (ALL y. ~ISL (Inr y))" | |
| 49 | by simp | |
| 50 | ||
| 51 | lemma ISR: "(ALL x. ISR (Inr x)) & (ALL y. ~ISR (Inl y))" | |
| 52 | by simp | |
| 53 | ||
| 39246 | 54 | primrec OUTL :: "'a + 'b => 'a" where | 
| 14516 | 55 | "OUTL (Inl x) = x" | 
| 56 | ||
| 39246 | 57 | primrec OUTR :: "'a + 'b => 'b" where | 
| 14516 | 58 | "OUTR (Inr x) = x" | 
| 59 | ||
| 60 | lemma OUTL: "OUTL (Inl x) = x" | |
| 61 | by simp | |
| 62 | ||
| 63 | lemma OUTR: "OUTR (Inr x) = x" | |
| 64 | by simp | |
| 65 | ||
| 44740 
a2940bc24bad
HOL/Import: Make HOL4 Import work with current Isabelle. Updated constant maps, added bool type map, and tuned compat theorem.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
44690diff
changeset | 66 | lemma sum_axiom: "EX! h. h o Inl = f & h o Inr = g" | 
| 
a2940bc24bad
HOL/Import: Make HOL4 Import work with current Isabelle. Updated constant maps, added bool type map, and tuned compat theorem.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
44690diff
changeset | 67 | apply (intro allI ex1I[of _ "sum_case f g"] conjI) | 
| 
a2940bc24bad
HOL/Import: Make HOL4 Import work with current Isabelle. Updated constant maps, added bool type map, and tuned compat theorem.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
44690diff
changeset | 68 | apply (simp_all add: o_def fun_eq_iff) | 
| 
a2940bc24bad
HOL/Import: Make HOL4 Import work with current Isabelle. Updated constant maps, added bool type map, and tuned compat theorem.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
44690diff
changeset | 69 | apply (rule) | 
| 
a2940bc24bad
HOL/Import: Make HOL4 Import work with current Isabelle. Updated constant maps, added bool type map, and tuned compat theorem.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
44690diff
changeset | 70 | apply (induct_tac x) | 
| 
a2940bc24bad
HOL/Import: Make HOL4 Import work with current Isabelle. Updated constant maps, added bool type map, and tuned compat theorem.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
44690diff
changeset | 71 | apply simp_all | 
| 
a2940bc24bad
HOL/Import: Make HOL4 Import work with current Isabelle. Updated constant maps, added bool type map, and tuned compat theorem.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
44690diff
changeset | 72 | done | 
| 
a2940bc24bad
HOL/Import: Make HOL4 Import work with current Isabelle. Updated constant maps, added bool type map, and tuned compat theorem.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
44690diff
changeset | 73 | |
| 14516 | 74 | lemma sum_case_def: "(ALL f g x. sum_case f g (Inl x) = f x) & (ALL f g y. sum_case f g (Inr y) = g y)" | 
| 41550 | 75 | by simp | 
| 14516 | 76 | |
| 77 | lemma one: "ALL v. v = ()" | |
| 41550 | 78 | by simp | 
| 14516 | 79 | |
| 80 | lemma option_case_def: "(!u f. option_case u f None = u) & (!u f x. option_case u f (Some x) = f x)" | |
| 81 | by simp | |
| 82 | ||
| 30235 
58d147683393
Made Option a separate theory and renamed option_map to Option.map
 nipkow parents: 
29044diff
changeset | 83 | lemma OPTION_MAP_DEF: "(!f x. Option.map f (Some x) = Some (f x)) & (!f. Option.map f None = None)" | 
| 14516 | 84 | by simp | 
| 85 | ||
| 39246 | 86 | primrec IS_SOME :: "'a option => bool" where | 
| 14516 | 87 | "IS_SOME (Some x) = True" | 
| 39246 | 88 | | "IS_SOME None = False" | 
| 14516 | 89 | |
| 39246 | 90 | primrec IS_NONE :: "'a option => bool" where | 
| 14516 | 91 | "IS_NONE (Some x) = False" | 
| 39246 | 92 | | "IS_NONE None = True" | 
| 14516 | 93 | |
| 94 | lemma IS_NONE_DEF: "(!x. IS_NONE (Some x) = False) & (IS_NONE None = True)" | |
| 95 | by simp | |
| 96 | ||
| 97 | lemma IS_SOME_DEF: "(!x. IS_SOME (Some x) = True) & (IS_SOME None = False)" | |
| 98 | by simp | |
| 99 | ||
| 39246 | 100 | primrec OPTION_JOIN :: "'a option option => 'a option" where | 
| 14516 | 101 | "OPTION_JOIN None = None" | 
| 39246 | 102 | | "OPTION_JOIN (Some x) = x" | 
| 14516 | 103 | |
| 104 | lemma OPTION_JOIN_DEF: "(OPTION_JOIN None = None) & (ALL x. OPTION_JOIN (Some x) = x)" | |
| 39246 | 105 | by simp | 
| 14516 | 106 | |
| 107 | lemma PAIR: "(fst x,snd x) = x" | |
| 108 | by simp | |
| 109 | ||
| 40607 | 110 | lemma PAIR_MAP: "map_pair f g p = (f (fst p),g (snd p))" | 
| 111 | by (simp add: map_pair_def split_def) | |
| 14516 | 112 | |
| 113 | lemma pair_case_def: "split = split" | |
| 41550 | 114 | .. | 
| 14516 | 115 | |
| 116 | lemma LESS_OR_EQ: "m <= (n::nat) = (m < n | m = n)" | |
| 117 | by auto | |
| 118 | ||
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
32960diff
changeset | 119 | definition nat_gt :: "nat => nat => bool" where | 
| 14516 | 120 | "nat_gt == %m n. n < m" | 
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
32960diff
changeset | 121 | |
| 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
32960diff
changeset | 122 | definition nat_ge :: "nat => nat => bool" where | 
| 14516 | 123 | "nat_ge == %m n. nat_gt m n | m = n" | 
| 124 | ||
| 125 | lemma [hol4rew]: "nat_gt m n = (n < m)" | |
| 126 | by (simp add: nat_gt_def) | |
| 127 | ||
| 128 | lemma [hol4rew]: "nat_ge m n = (n <= m)" | |
| 129 | by (auto simp add: nat_ge_def nat_gt_def) | |
| 130 | ||
| 131 | lemma GREATER_DEF: "ALL m n. (n < m) = (n < m)" | |
| 132 | by simp | |
| 133 | ||
| 134 | lemma GREATER_OR_EQ: "ALL m n. n <= (m::nat) = (n < m | m = n)" | |
| 135 | by auto | |
| 136 | ||
| 137 | lemma LESS_DEF: "m < n = (? P. (!n. P (Suc n) --> P n) & P m & ~P n)" | |
| 138 | proof safe | |
| 41550 | 139 | assume 1: "m < n" | 
| 14516 | 140 | def P == "%n. n <= m" | 
| 141 | have "(!n. P (Suc n) \<longrightarrow> P n) & P m & ~P n" | |
| 142 | proof (auto simp add: P_def) | |
| 143 | assume "n <= m" | |
| 41550 | 144 | with 1 | 
| 14516 | 145 | show False | 
| 146 | by auto | |
| 147 | qed | |
| 148 | thus "? P. (!n. P (Suc n) \<longrightarrow> P n) & P m & ~P n" | |
| 149 | by auto | |
| 150 | next | |
| 151 | fix P | |
| 152 | assume alln: "!n. P (Suc n) \<longrightarrow> P n" | |
| 153 | assume pm: "P m" | |
| 154 | assume npn: "~P n" | |
| 155 | have "!k q. q + k = m \<longrightarrow> P q" | |
| 156 | proof | |
| 157 | fix k | |
| 158 | show "!q. q + k = m \<longrightarrow> P q" | |
| 159 | proof (induct k,simp_all) | |
| 23389 | 160 | show "P m" by fact | 
| 14516 | 161 | next | 
| 162 | fix k | |
| 163 | assume ind: "!q. q + k = m \<longrightarrow> P q" | |
| 164 | show "!q. Suc (q + k) = m \<longrightarrow> P q" | |
| 165 | proof (rule+) | |
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32479diff
changeset | 166 | fix q | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32479diff
changeset | 167 | assume "Suc (q + k) = m" | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32479diff
changeset | 168 | hence "(Suc q) + k = m" | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32479diff
changeset | 169 | by simp | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32479diff
changeset | 170 | with ind | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32479diff
changeset | 171 | have psq: "P (Suc q)" | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32479diff
changeset | 172 | by simp | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32479diff
changeset | 173 | from alln | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32479diff
changeset | 174 | have "P (Suc q) --> P q" | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32479diff
changeset | 175 | .. | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32479diff
changeset | 176 | with psq | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32479diff
changeset | 177 | show "P q" | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32479diff
changeset | 178 | by simp | 
| 14516 | 179 | qed | 
| 180 | qed | |
| 181 | qed | |
| 182 | hence "!q. q + (m - n) = m \<longrightarrow> P q" | |
| 183 | .. | |
| 184 | hence hehe: "n + (m - n) = m \<longrightarrow> P n" | |
| 185 | .. | |
| 186 | show "m < n" | |
| 187 | proof (rule classical) | |
| 188 | assume "~(m<n)" | |
| 189 | hence "n <= m" | |
| 190 | by simp | |
| 191 | with hehe | |
| 192 | have "P n" | |
| 193 | by simp | |
| 194 | with npn | |
| 195 | show "m < n" | |
| 196 | .. | |
| 197 | qed | |
| 41550 | 198 | qed | 
| 14516 | 199 | |
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
32960diff
changeset | 200 | definition FUNPOW :: "('a => 'a) => nat => 'a => 'a" where
 | 
| 30971 | 201 | "FUNPOW f n == f ^^ n" | 
| 14516 | 202 | |
| 30971 | 203 | lemma FUNPOW: "(ALL f x. (f ^^ 0) x = x) & | 
| 204 | (ALL f n x. (f ^^ Suc n) x = (f ^^ n) (f x))" | |
| 30952 
7ab2716dd93b
power operation on functions with syntax o^; power operation on relations with syntax ^^
 haftmann parents: 
30660diff
changeset | 205 | by (simp add: funpow_swap1) | 
| 14516 | 206 | |
| 30971 | 207 | lemma [hol4rew]: "FUNPOW f n = f ^^ n" | 
| 14516 | 208 | by (simp add: FUNPOW_def) | 
| 209 | ||
| 210 | lemma ADD: "(!n. (0::nat) + n = n) & (!m n. Suc m + n = Suc (m + n))" | |
| 211 | by simp | |
| 212 | ||
| 213 | lemma MULT: "(!n. (0::nat) * n = 0) & (!m n. Suc m * n = m * n + n)" | |
| 214 | by simp | |
| 215 | ||
| 216 | lemma SUB: "(!m. (0::nat) - m = 0) & (!m n. (Suc m) - n = (if m < n then 0 else Suc (m - n)))" | |
| 30952 
7ab2716dd93b
power operation on functions with syntax o^; power operation on relations with syntax ^^
 haftmann parents: 
30660diff
changeset | 217 | by (simp) arith | 
| 14516 | 218 | |
| 219 | lemma MAX_DEF: "max (m::nat) n = (if m < n then n else m)" | |
| 220 | by (simp add: max_def) | |
| 221 | ||
| 222 | lemma MIN_DEF: "min (m::nat) n = (if m < n then m else n)" | |
| 223 | by (simp add: min_def) | |
| 224 | ||
| 225 | lemma DIVISION: "(0::nat) < n --> (!k. (k = k div n * n + k mod n) & k mod n < n)" | |
| 226 | by simp | |
| 227 | ||
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
32960diff
changeset | 228 | definition ALT_ZERO :: nat where | 
| 14516 | 229 | "ALT_ZERO == 0" | 
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
32960diff
changeset | 230 | |
| 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
32960diff
changeset | 231 | definition NUMERAL_BIT1 :: "nat \<Rightarrow> nat" where | 
| 14516 | 232 | "NUMERAL_BIT1 n == n + (n + Suc 0)" | 
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
32960diff
changeset | 233 | |
| 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
32960diff
changeset | 234 | definition NUMERAL_BIT2 :: "nat \<Rightarrow> nat" where | 
| 14516 | 235 | "NUMERAL_BIT2 n == n + (n + Suc (Suc 0))" | 
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
32960diff
changeset | 236 | |
| 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
32960diff
changeset | 237 | definition NUMERAL :: "nat \<Rightarrow> nat" where | 
| 14516 | 238 | "NUMERAL x == x" | 
| 239 | ||
| 240 | lemma [hol4rew]: "NUMERAL ALT_ZERO = 0" | |
| 241 | by (simp add: ALT_ZERO_def NUMERAL_def) | |
| 242 | ||
| 243 | lemma [hol4rew]: "NUMERAL (NUMERAL_BIT1 ALT_ZERO) = 1" | |
| 244 | by (simp add: ALT_ZERO_def NUMERAL_BIT1_def NUMERAL_def) | |
| 245 | ||
| 246 | lemma [hol4rew]: "NUMERAL (NUMERAL_BIT2 ALT_ZERO) = 2" | |
| 247 | by (simp add: ALT_ZERO_def NUMERAL_BIT2_def NUMERAL_def) | |
| 248 | ||
| 249 | lemma EXP: "(!m. m ^ 0 = (1::nat)) & (!m n. m ^ Suc n = m * (m::nat) ^ n)" | |
| 250 | by auto | |
| 251 | ||
| 252 | lemma num_case_def: "(!b f. nat_case b f 0 = b) & (!b f n. nat_case b f (Suc n) = f n)" | |
| 41550 | 253 | by simp | 
| 14516 | 254 | |
| 255 | lemma divides_def: "(a::nat) dvd b = (? q. b = q * a)" | |
| 41550 | 256 | by (auto simp add: dvd_def) | 
| 14516 | 257 | |
| 258 | lemma list_case_def: "(!v f. list_case v f [] = v) & (!v f a0 a1. list_case v f (a0#a1) = f a0 a1)" | |
| 259 | by simp | |
| 260 | ||
| 39246 | 261 | primrec list_size :: "('a \<Rightarrow> nat) \<Rightarrow> 'a list \<Rightarrow> nat" where
 | 
| 262 | "list_size f [] = 0" | |
| 263 | | "list_size f (a0#a1) = 1 + (f a0 + list_size f a1)" | |
| 14516 | 264 | |
| 39246 | 265 | lemma list_size_def': "(!f. list_size f [] = 0) & | 
| 14516 | 266 | (!f a0 a1. list_size f (a0#a1) = 1 + (f a0 + list_size f a1))" | 
| 267 | by simp | |
| 268 | ||
| 269 | lemma list_case_cong: "! M M' v f. M = M' & (M' = [] \<longrightarrow> v = v') & | |
| 270 | (!a0 a1. (M' = a0#a1) \<longrightarrow> (f a0 a1 = f' a0 a1)) --> | |
| 271 | (list_case v f M = list_case v' f' M')" | |
| 272 | proof clarify | |
| 273 | fix M M' v f | |
| 41550 | 274 | assume 1: "M' = [] \<longrightarrow> v = v'" | 
| 275 | and 2: "!a0 a1. M' = a0 # a1 \<longrightarrow> f a0 a1 = f' a0 a1" | |
| 14516 | 276 | show "list_case v f M' = list_case v' f' M'" | 
| 277 | proof (rule List.list.case_cong) | |
| 278 | show "M' = M'" | |
| 279 | .. | |
| 280 | next | |
| 281 | assume "M' = []" | |
| 41550 | 282 | with 1 2 | 
| 14516 | 283 | show "v = v'" | 
| 284 | by auto | |
| 285 | next | |
| 286 | fix a0 a1 | |
| 287 | assume "M' = a0 # a1" | |
| 41550 | 288 | with 1 2 | 
| 14516 | 289 | show "f a0 a1 = f' a0 a1" | 
| 290 | by auto | |
| 291 | qed | |
| 292 | qed | |
| 293 | ||
| 294 | lemma list_Axiom: "ALL f0 f1. EX fn. (fn [] = f0) & (ALL a0 a1. fn (a0#a1) = f1 a0 a1 (fn a1))" | |
| 295 | proof safe | |
| 296 | fix f0 f1 | |
| 297 | def fn == "list_rec f0 f1" | |
| 298 | have "fn [] = f0 & (ALL a0 a1. fn (a0 # a1) = f1 a0 a1 (fn a1))" | |
| 299 | by (simp add: fn_def) | |
| 300 | thus "EX fn. fn [] = f0 & (ALL a0 a1. fn (a0 # a1) = f1 a0 a1 (fn a1))" | |
| 301 | by auto | |
| 302 | qed | |
| 303 | ||
| 304 | lemma list_Axiom_old: "EX! fn. (fn [] = x) & (ALL h t. fn (h#t) = f (fn t) h t)" | |
| 305 | proof safe | |
| 306 | def fn == "list_rec x (%h t r. f r h t)" | |
| 307 | have "fn [] = x & (ALL h t. fn (h # t) = f (fn t) h t)" | |
| 308 | by (simp add: fn_def) | |
| 309 | thus "EX fn. fn [] = x & (ALL h t. fn (h # t) = f (fn t) h t)" | |
| 310 | by auto | |
| 311 | next | |
| 312 | fix fn1 fn2 | |
| 41550 | 313 | assume 1: "ALL h t. fn1 (h # t) = f (fn1 t) h t" | 
| 314 | assume 2: "ALL h t. fn2 (h # t) = f (fn2 t) h t" | |
| 315 | assume 3: "fn2 [] = fn1 []" | |
| 14516 | 316 | show "fn1 = fn2" | 
| 317 | proof | |
| 318 | fix xs | |
| 319 | show "fn1 xs = fn2 xs" | |
| 41550 | 320 | by (induct xs) (simp_all add: 1 2 3) | 
| 14516 | 321 | qed | 
| 322 | qed | |
| 323 | ||
| 37596 | 324 | lemma NULL_DEF: "(List.null [] = True) & (!h t. List.null (h # t) = False)" | 
| 325 | by (simp add: null_def) | |
| 14516 | 326 | |
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
32960diff
changeset | 327 | definition sum :: "nat list \<Rightarrow> nat" where | 
| 14516 | 328 | "sum l == foldr (op +) l 0" | 
| 329 | ||
| 330 | lemma SUM: "(sum [] = 0) & (!h t. sum (h#t) = h + sum t)" | |
| 331 | by (simp add: sum_def) | |
| 332 | ||
| 333 | lemma APPEND: "(!l. [] @ l = l) & (!l1 l2 h. (h#l1) @ l2 = h# l1 @ l2)" | |
| 334 | by simp | |
| 335 | ||
| 336 | lemma FLAT: "(concat [] = []) & (!h t. concat (h#t) = h @ (concat t))" | |
| 337 | by simp | |
| 338 | ||
| 339 | lemma LENGTH: "(length [] = 0) & (!h t. length (h#t) = Suc (length t))" | |
| 340 | by simp | |
| 341 | ||
| 342 | lemma MAP: "(!f. map f [] = []) & (!f h t. map f (h#t) = f h#map f t)" | |
| 343 | by simp | |
| 344 | ||
| 37596 | 345 | lemma MEM: "(!x. List.member [] x = False) & (!x h t. List.member (h#t) x = ((x = h) | List.member t x))" | 
| 346 | by (simp add: member_def) | |
| 14516 | 347 | |
| 348 | lemma FILTER: "(!P. filter P [] = []) & (!P h t. | |
| 349 | filter P (h#t) = (if P h then h#filter P t else filter P t))" | |
| 350 | by simp | |
| 351 | ||
| 352 | lemma REPLICATE: "(ALL x. replicate 0 x = []) & | |
| 353 | (ALL n x. replicate (Suc n) x = x # replicate n x)" | |
| 354 | by simp | |
| 355 | ||
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
32960diff
changeset | 356 | definition FOLDR :: "[['a,'b]\<Rightarrow>'b,'b,'a list] \<Rightarrow> 'b" where | 
| 14516 | 357 | "FOLDR f e l == foldr f l e" | 
| 358 | ||
| 359 | lemma [hol4rew]: "FOLDR f e l = foldr f l e" | |
| 360 | by (simp add: FOLDR_def) | |
| 361 | ||
| 362 | lemma FOLDR: "(!f e. foldr f [] e = e) & (!f e x l. foldr f (x#l) e = f x (foldr f l e))" | |
| 363 | by simp | |
| 364 | ||
| 365 | lemma FOLDL: "(!f e. foldl f e [] = e) & (!f e x l. foldl f e (x#l) = foldl f (f e x) l)" | |
| 366 | by simp | |
| 367 | ||
| 368 | lemma EVERY_DEF: "(!P. list_all P [] = True) & (!P h t. list_all P (h#t) = (P h & list_all P t))" | |
| 369 | by simp | |
| 370 | ||
| 37596 | 371 | lemma list_exists_DEF: "(!P. list_ex P [] = False) & (!P h t. list_ex P (h#t) = (P h | list_ex P t))" | 
| 14516 | 372 | by simp | 
| 373 | ||
| 39246 | 374 | primrec map2 :: "[['a,'b]\<Rightarrow>'c,'a list,'b list] \<Rightarrow> 'c list" where | 
| 14516 | 375 | map2_Nil: "map2 f [] l2 = []" | 
| 39246 | 376 | | map2_Cons: "map2 f (x#xs) l2 = f x (hd l2) # map2 f xs (tl l2)" | 
| 14516 | 377 | |
| 378 | lemma MAP2: "(!f. map2 f [] [] = []) & (!f h1 t1 h2 t2. map2 f (h1#t1) (h2#t2) = f h1 h2#map2 f t1 t2)" | |
| 379 | by simp | |
| 380 | ||
| 381 | lemma list_INDUCT: "\<lbrakk> P [] ; !t. P t \<longrightarrow> (!h. P (h#t)) \<rbrakk> \<Longrightarrow> !l. P l" | |
| 382 | proof | |
| 383 | fix l | |
| 384 | assume "P []" | |
| 385 | assume allt: "!t. P t \<longrightarrow> (!h. P (h # t))" | |
| 386 | show "P l" | |
| 387 | proof (induct l) | |
| 23389 | 388 | show "P []" by fact | 
| 14516 | 389 | next | 
| 390 | fix h t | |
| 391 | assume "P t" | |
| 392 | with allt | |
| 393 | have "!h. P (h # t)" | |
| 394 | by auto | |
| 395 | thus "P (h # t)" | |
| 396 | .. | |
| 397 | qed | |
| 398 | qed | |
| 399 | ||
| 400 | lemma list_CASES: "(l = []) | (? t h. l = h#t)" | |
| 401 | by (induct l,auto) | |
| 402 | ||
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
32960diff
changeset | 403 | definition ZIP :: "'a list * 'b list \<Rightarrow> ('a * 'b) list" where
 | 
| 14516 | 404 | "ZIP == %(a,b). zip a b" | 
| 405 | ||
| 406 | lemma ZIP: "(zip [] [] = []) & | |
| 407 | (!x1 l1 x2 l2. zip (x1#l1) (x2#l2) = (x1,x2)#zip l1 l2)" | |
| 408 | by simp | |
| 409 | ||
| 410 | lemma [hol4rew]: "ZIP (a,b) = zip a b" | |
| 411 | by (simp add: ZIP_def) | |
| 412 | ||
| 39246 | 413 | primrec unzip :: "('a * 'b) list \<Rightarrow> 'a list * 'b list" where
 | 
| 14516 | 414 | unzip_Nil: "unzip [] = ([],[])" | 
| 39246 | 415 | | unzip_Cons: "unzip (xy#xys) = (let zs = unzip xys in (fst xy # fst zs,snd xy # snd zs))" | 
| 14516 | 416 | |
| 417 | lemma UNZIP: "(unzip [] = ([],[])) & | |
| 418 | (!x l. unzip (x#l) = (fst x#fst (unzip l),snd x#snd (unzip l)))" | |
| 419 | by (simp add: Let_def) | |
| 420 | ||
| 421 | lemma REVERSE: "(rev [] = []) & (!h t. rev (h#t) = (rev t) @ [h])" | |
| 41550 | 422 | by simp | 
| 14516 | 423 | |
| 424 | lemma REAL_SUP_ALLPOS: "\<lbrakk> ALL x. P (x::real) \<longrightarrow> 0 < x ; EX x. P x; EX z. ALL x. P x \<longrightarrow> x < z \<rbrakk> \<Longrightarrow> EX s. ALL y. (EX x. P x & y < x) = (y < s)" | |
| 425 | proof safe | |
| 426 | fix x z | |
| 427 | assume allx: "ALL x. P x \<longrightarrow> 0 < x" | |
| 428 | assume px: "P x" | |
| 429 | assume allx': "ALL x. P x \<longrightarrow> x < z" | |
| 430 | have "EX s. ALL y. (EX x : Collect P. y < x) = (y < s)" | |
| 431 | proof (rule posreal_complete) | |
| 432 | from px | |
| 433 | show "EX x. x : Collect P" | |
| 434 | by auto | |
| 435 | next | |
| 436 | from allx' | |
| 437 | show "EX y. ALL x : Collect P. x < y" | |
| 438 | apply simp | |
| 439 | .. | |
| 440 | qed | |
| 441 | thus "EX s. ALL y. (EX x. P x & y < x) = (y < s)" | |
| 442 | by simp | |
| 443 | qed | |
| 444 | ||
| 445 | lemma REAL_10: "~((1::real) = 0)" | |
| 446 | by simp | |
| 447 | ||
| 448 | lemma REAL_ADD_ASSOC: "(x::real) + (y + z) = x + y + z" | |
| 449 | by simp | |
| 450 | ||
| 451 | lemma REAL_MUL_ASSOC: "(x::real) * (y * z) = x * y * z" | |
| 452 | by simp | |
| 453 | ||
| 454 | lemma REAL_ADD_LINV: "-x + x = (0::real)" | |
| 455 | by simp | |
| 456 | ||
| 457 | lemma REAL_MUL_LINV: "x ~= (0::real) ==> inverse x * x = 1" | |
| 458 | by simp | |
| 459 | ||
| 460 | lemma REAL_LT_TOTAL: "((x::real) = y) | x < y | y < x" | |
| 41550 | 461 | by auto | 
| 14516 | 462 | |
| 463 | lemma [hol4rew]: "real (0::nat) = 0" | |
| 464 | by simp | |
| 465 | ||
| 466 | lemma [hol4rew]: "real (1::nat) = 1" | |
| 467 | by simp | |
| 468 | ||
| 469 | lemma [hol4rew]: "real (2::nat) = 2" | |
| 470 | by simp | |
| 471 | ||
| 472 | lemma real_lte: "((x::real) <= y) = (~(y < x))" | |
| 473 | by auto | |
| 474 | ||
| 475 | lemma real_of_num: "((0::real) = 0) & (!n. real (Suc n) = real n + 1)" | |
| 476 | by (simp add: real_of_nat_Suc) | |
| 477 | ||
| 478 | lemma abs: "abs (x::real) = (if 0 <= x then x else -x)" | |
| 15003 | 479 | by (simp add: abs_if) | 
| 14516 | 480 | |
| 481 | lemma pow: "(!x::real. x ^ 0 = 1) & (!x::real. ALL n. x ^ (Suc n) = x * x ^ n)" | |
| 15003 | 482 | by simp | 
| 14516 | 483 | |
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
32960diff
changeset | 484 | definition real_gt :: "real => real => bool" where | 
| 14516 | 485 | "real_gt == %x y. y < x" | 
| 486 | ||
| 487 | lemma [hol4rew]: "real_gt x y = (y < x)" | |
| 488 | by (simp add: real_gt_def) | |
| 489 | ||
| 490 | lemma real_gt: "ALL x (y::real). (y < x) = (y < x)" | |
| 491 | by simp | |
| 492 | ||
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
32960diff
changeset | 493 | definition real_ge :: "real => real => bool" where | 
| 14516 | 494 | "real_ge x y == y <= x" | 
| 495 | ||
| 496 | lemma [hol4rew]: "real_ge x y = (y <= x)" | |
| 497 | by (simp add: real_ge_def) | |
| 498 | ||
| 499 | lemma real_ge: "ALL x y. (y <= x) = (y <= x)" | |
| 500 | by simp | |
| 501 | ||
| 44740 
a2940bc24bad
HOL/Import: Make HOL4 Import work with current Isabelle. Updated constant maps, added bool type map, and tuned compat theorem.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
44690diff
changeset | 502 | definition [hol4rew]: "list_mem x xs = List.member xs x" | 
| 
a2940bc24bad
HOL/Import: Make HOL4 Import work with current Isabelle. Updated constant maps, added bool type map, and tuned compat theorem.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
44690diff
changeset | 503 | |
| 14516 | 504 | end |