| author | wenzelm | 
| Mon, 18 Sep 2006 19:12:46 +0200 | |
| changeset 20574 | a10885a269cb | 
| parent 20553 | 7ced6152e52c | 
| child 20633 | e98f59806244 | 
| permissions | -rw-r--r-- | 
| 17296 | 1 | (* Title : HOL/Hyperreal/StarClasses.thy | 
| 2 | ID : $Id$ | |
| 3 | Author : Brian Huffman | |
| 4 | *) | |
| 5 | ||
| 6 | header {* Class Instances *}
 | |
| 7 | ||
| 8 | theory StarClasses | |
| 17429 
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
 huffman parents: 
17332diff
changeset | 9 | imports StarDef | 
| 17296 | 10 | begin | 
| 11 | ||
| 17332 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 12 | subsection {* Syntactic classes *}
 | 
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 13 | |
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 14 | instance star :: (ord) ord .. | 
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 15 | instance star :: (zero) zero .. | 
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 16 | instance star :: (one) one .. | 
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 17 | instance star :: (plus) plus .. | 
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 18 | instance star :: (times) times .. | 
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 19 | instance star :: (minus) minus .. | 
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 20 | instance star :: (inverse) inverse .. | 
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 21 | instance star :: (number) number .. | 
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 22 | instance star :: ("Divides.div") "Divides.div" ..
 | 
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 23 | instance star :: (power) power .. | 
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 24 | |
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 25 | defs (overloaded) | 
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 26 | star_zero_def: "0 \<equiv> star_of 0" | 
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 27 | star_one_def: "1 \<equiv> star_of 1" | 
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 28 | star_number_def: "number_of b \<equiv> star_of (number_of b)" | 
| 17429 
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
 huffman parents: 
17332diff
changeset | 29 | star_add_def: "(op +) \<equiv> *f2* (op +)" | 
| 
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
 huffman parents: 
17332diff
changeset | 30 | star_diff_def: "(op -) \<equiv> *f2* (op -)" | 
| 
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
 huffman parents: 
17332diff
changeset | 31 | star_minus_def: "uminus \<equiv> *f* uminus" | 
| 
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
 huffman parents: 
17332diff
changeset | 32 | star_mult_def: "(op *) \<equiv> *f2* (op *)" | 
| 
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
 huffman parents: 
17332diff
changeset | 33 | star_divide_def: "(op /) \<equiv> *f2* (op /)" | 
| 
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
 huffman parents: 
17332diff
changeset | 34 | star_inverse_def: "inverse \<equiv> *f* inverse" | 
| 
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
 huffman parents: 
17332diff
changeset | 35 | star_le_def: "(op \<le>) \<equiv> *p2* (op \<le>)" | 
| 
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
 huffman parents: 
17332diff
changeset | 36 | star_less_def: "(op <) \<equiv> *p2* (op <)" | 
| 
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
 huffman parents: 
17332diff
changeset | 37 | star_abs_def: "abs \<equiv> *f* abs" | 
| 
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
 huffman parents: 
17332diff
changeset | 38 | star_div_def: "(op div) \<equiv> *f2* (op div)" | 
| 
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
 huffman parents: 
17332diff
changeset | 39 | star_mod_def: "(op mod) \<equiv> *f2* (op mod)" | 
| 
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
 huffman parents: 
17332diff
changeset | 40 | star_power_def: "(op ^) \<equiv> \<lambda>x n. ( *f* (\<lambda>x. x ^ n)) x" | 
| 17332 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 41 | |
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 42 | lemmas star_class_defs [transfer_unfold] = | 
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 43 | star_zero_def star_one_def star_number_def | 
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 44 | star_add_def star_diff_def star_minus_def | 
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 45 | star_mult_def star_divide_def star_inverse_def | 
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 46 | star_le_def star_less_def star_abs_def | 
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 47 | star_div_def star_mod_def star_power_def | 
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 48 | |
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 49 | text {* @{term star_of} preserves class operations *}
 | 
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 50 | |
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 51 | lemma star_of_add: "star_of (x + y) = star_of x + star_of y" | 
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 52 | by transfer (rule refl) | 
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 53 | |
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 54 | lemma star_of_diff: "star_of (x - y) = star_of x - star_of y" | 
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 55 | by transfer (rule refl) | 
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 56 | |
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 57 | lemma star_of_minus: "star_of (-x) = - star_of x" | 
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 58 | by transfer (rule refl) | 
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 59 | |
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 60 | lemma star_of_mult: "star_of (x * y) = star_of x * star_of y" | 
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 61 | by transfer (rule refl) | 
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 62 | |
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 63 | lemma star_of_divide: "star_of (x / y) = star_of x / star_of y" | 
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 64 | by transfer (rule refl) | 
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 65 | |
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 66 | lemma star_of_inverse: "star_of (inverse x) = inverse (star_of x)" | 
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 67 | by transfer (rule refl) | 
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 68 | |
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 69 | lemma star_of_div: "star_of (x div y) = star_of x div star_of y" | 
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 70 | by transfer (rule refl) | 
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 71 | |
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 72 | lemma star_of_mod: "star_of (x mod y) = star_of x mod star_of y" | 
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 73 | by transfer (rule refl) | 
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 74 | |
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 75 | lemma star_of_power: "star_of (x ^ n) = star_of x ^ n" | 
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 76 | by transfer (rule refl) | 
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 77 | |
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 78 | lemma star_of_abs: "star_of (abs x) = abs (star_of x)" | 
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 79 | by transfer (rule refl) | 
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 80 | |
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 81 | text {* @{term star_of} preserves numerals *}
 | 
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 82 | |
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 83 | lemma star_of_zero: "star_of 0 = 0" | 
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 84 | by transfer (rule refl) | 
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 85 | |
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 86 | lemma star_of_one: "star_of 1 = 1" | 
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 87 | by transfer (rule refl) | 
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 88 | |
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 89 | lemma star_of_number_of: "star_of (number_of x) = number_of x" | 
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 90 | by transfer (rule refl) | 
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 91 | |
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 92 | text {* @{term star_of} preserves orderings *}
 | 
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 93 | |
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 94 | lemma star_of_less: "(star_of x < star_of y) = (x < y)" | 
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 95 | by transfer (rule refl) | 
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 96 | |
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 97 | lemma star_of_le: "(star_of x \<le> star_of y) = (x \<le> y)" | 
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 98 | by transfer (rule refl) | 
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 99 | |
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 100 | lemma star_of_eq: "(star_of x = star_of y) = (x = y)" | 
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 101 | by transfer (rule refl) | 
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 102 | |
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 103 | text{*As above, for 0*}
 | 
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 104 | |
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 105 | lemmas star_of_0_less = star_of_less [of 0, simplified star_of_zero] | 
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 106 | lemmas star_of_0_le = star_of_le [of 0, simplified star_of_zero] | 
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 107 | lemmas star_of_0_eq = star_of_eq [of 0, simplified star_of_zero] | 
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 108 | |
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 109 | lemmas star_of_less_0 = star_of_less [of _ 0, simplified star_of_zero] | 
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 110 | lemmas star_of_le_0 = star_of_le [of _ 0, simplified star_of_zero] | 
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 111 | lemmas star_of_eq_0 = star_of_eq [of _ 0, simplified star_of_zero] | 
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 112 | |
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 113 | text{*As above, for 1*}
 | 
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 114 | |
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 115 | lemmas star_of_1_less = star_of_less [of 1, simplified star_of_one] | 
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 116 | lemmas star_of_1_le = star_of_le [of 1, simplified star_of_one] | 
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 117 | lemmas star_of_1_eq = star_of_eq [of 1, simplified star_of_one] | 
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 118 | |
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 119 | lemmas star_of_less_1 = star_of_less [of _ 1, simplified star_of_one] | 
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 120 | lemmas star_of_le_1 = star_of_le [of _ 1, simplified star_of_one] | 
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 121 | lemmas star_of_eq_1 = star_of_eq [of _ 1, simplified star_of_one] | 
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 122 | |
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 123 | text{*As above, for numerals*}
 | 
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 124 | |
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 125 | lemmas star_of_number_less = | 
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 126 | star_of_less [of "number_of w", standard, simplified star_of_number_of] | 
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 127 | lemmas star_of_number_le = | 
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 128 | star_of_le [of "number_of w", standard, simplified star_of_number_of] | 
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 129 | lemmas star_of_number_eq = | 
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 130 | star_of_eq [of "number_of w", standard, simplified star_of_number_of] | 
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 131 | |
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 132 | lemmas star_of_less_number = | 
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 133 | star_of_less [of _ "number_of w", standard, simplified star_of_number_of] | 
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 134 | lemmas star_of_le_number = | 
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 135 | star_of_le [of _ "number_of w", standard, simplified star_of_number_of] | 
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 136 | lemmas star_of_eq_number = | 
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 137 | star_of_eq [of _ "number_of w", standard, simplified star_of_number_of] | 
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 138 | |
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 139 | lemmas star_of_simps [simp] = | 
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 140 | star_of_add star_of_diff star_of_minus | 
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 141 | star_of_mult star_of_divide star_of_inverse | 
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 142 | star_of_div star_of_mod | 
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 143 | star_of_power star_of_abs | 
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 144 | star_of_zero star_of_one star_of_number_of | 
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 145 | star_of_less star_of_le star_of_eq | 
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 146 | star_of_0_less star_of_0_le star_of_0_eq | 
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 147 | star_of_less_0 star_of_le_0 star_of_eq_0 | 
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 148 | star_of_1_less star_of_1_le star_of_1_eq | 
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 149 | star_of_less_1 star_of_le_1 star_of_eq_1 | 
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 150 | star_of_number_less star_of_number_le star_of_number_eq | 
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 151 | star_of_less_number star_of_le_number star_of_eq_number | 
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 152 | |
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 153 | subsection {* Ordering classes *}
 | 
| 17296 | 154 | |
| 155 | instance star :: (order) order | |
| 156 | apply (intro_classes) | |
| 157 | apply (transfer, rule order_refl) | |
| 158 | apply (transfer, erule (1) order_trans) | |
| 159 | apply (transfer, erule (1) order_antisym) | |
| 160 | apply (transfer, rule order_less_le) | |
| 161 | done | |
| 162 | ||
| 163 | instance star :: (linorder) linorder | |
| 164 | by (intro_classes, transfer, rule linorder_linear) | |
| 165 | ||
| 166 | ||
| 17332 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 167 | subsection {* Lattice ordering classes *}
 | 
| 17296 | 168 | |
| 169 | text {*
 | |
| 170 | Some extra trouble is necessary because the class axioms | |
| 171 |   for @{term meet} and @{term join} use quantification over
 | |
| 172 | function spaces. | |
| 173 | *} | |
| 174 | ||
| 175 | lemma ex_star_fun: | |
| 17429 
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
 huffman parents: 
17332diff
changeset | 176 |   "\<exists>f::('a \<Rightarrow> 'b) star. P (\<lambda>x. f \<star> x)
 | 
| 17296 | 177 | \<Longrightarrow> \<exists>f::'a star \<Rightarrow> 'b star. P f" | 
| 178 | by (erule exE, erule exI) | |
| 179 | ||
| 180 | lemma ex_star_fun2: | |
| 17429 
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
 huffman parents: 
17332diff
changeset | 181 |   "\<exists>f::('a \<Rightarrow> 'b \<Rightarrow> 'c) star. P (\<lambda>x y. f \<star> x \<star> y)
 | 
| 17296 | 182 | \<Longrightarrow> \<exists>f::'a star \<Rightarrow> 'b star \<Rightarrow> 'c star. P f" | 
| 183 | by (erule exE, erule exI) | |
| 184 | ||
| 185 | instance star :: (join_semilorder) join_semilorder | |
| 186 | apply (intro_classes) | |
| 187 | apply (rule ex_star_fun2) | |
| 188 | apply (transfer is_join_def) | |
| 189 | apply (rule join_exists) | |
| 190 | done | |
| 191 | ||
| 192 | instance star :: (meet_semilorder) meet_semilorder | |
| 193 | apply (intro_classes) | |
| 194 | apply (rule ex_star_fun2) | |
| 195 | apply (transfer is_meet_def) | |
| 196 | apply (rule meet_exists) | |
| 197 | done | |
| 198 | ||
| 199 | instance star :: (lorder) lorder .. | |
| 200 | ||
| 17429 
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
 huffman parents: 
17332diff
changeset | 201 | lemma star_join_def [transfer_unfold]: "join \<equiv> *f2* join" | 
| 
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
 huffman parents: 
17332diff
changeset | 202 | apply (rule is_join_unique [OF is_join_join, THEN eq_reflection]) | 
| 17296 | 203 | apply (transfer is_join_def, rule is_join_join) | 
| 204 | done | |
| 205 | ||
| 17429 
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
 huffman parents: 
17332diff
changeset | 206 | lemma star_meet_def [transfer_unfold]: "meet \<equiv> *f2* meet" | 
| 
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
 huffman parents: 
17332diff
changeset | 207 | apply (rule is_meet_unique [OF is_meet_meet, THEN eq_reflection]) | 
| 17296 | 208 | apply (transfer is_meet_def, rule is_meet_meet) | 
| 209 | done | |
| 210 | ||
| 17332 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 211 | subsection {* Ordered group classes *}
 | 
| 17296 | 212 | |
| 213 | instance star :: (semigroup_add) semigroup_add | |
| 214 | by (intro_classes, transfer, rule add_assoc) | |
| 215 | ||
| 216 | instance star :: (ab_semigroup_add) ab_semigroup_add | |
| 217 | by (intro_classes, transfer, rule add_commute) | |
| 218 | ||
| 219 | instance star :: (semigroup_mult) semigroup_mult | |
| 220 | by (intro_classes, transfer, rule mult_assoc) | |
| 221 | ||
| 222 | instance star :: (ab_semigroup_mult) ab_semigroup_mult | |
| 223 | by (intro_classes, transfer, rule mult_commute) | |
| 224 | ||
| 225 | instance star :: (comm_monoid_add) comm_monoid_add | |
| 226 | by (intro_classes, transfer, rule comm_monoid_add_class.add_0) | |
| 227 | ||
| 228 | instance star :: (monoid_mult) monoid_mult | |
| 229 | apply (intro_classes) | |
| 230 | apply (transfer, rule mult_1_left) | |
| 231 | apply (transfer, rule mult_1_right) | |
| 232 | done | |
| 233 | ||
| 234 | instance star :: (comm_monoid_mult) comm_monoid_mult | |
| 235 | by (intro_classes, transfer, rule mult_1) | |
| 236 | ||
| 237 | instance star :: (cancel_semigroup_add) cancel_semigroup_add | |
| 238 | apply (intro_classes) | |
| 239 | apply (transfer, erule add_left_imp_eq) | |
| 240 | apply (transfer, erule add_right_imp_eq) | |
| 241 | done | |
| 242 | ||
| 243 | instance star :: (cancel_ab_semigroup_add) cancel_ab_semigroup_add | |
| 244 | by (intro_classes, transfer, rule add_imp_eq) | |
| 245 | ||
| 246 | instance star :: (ab_group_add) ab_group_add | |
| 247 | apply (intro_classes) | |
| 248 | apply (transfer, rule left_minus) | |
| 249 | apply (transfer, rule diff_minus) | |
| 250 | done | |
| 251 | ||
| 252 | instance star :: (pordered_ab_semigroup_add) pordered_ab_semigroup_add | |
| 253 | by (intro_classes, transfer, rule add_left_mono) | |
| 254 | ||
| 255 | instance star :: (pordered_cancel_ab_semigroup_add) pordered_cancel_ab_semigroup_add .. | |
| 256 | ||
| 257 | instance star :: (pordered_ab_semigroup_add_imp_le) pordered_ab_semigroup_add_imp_le | |
| 258 | by (intro_classes, transfer, rule add_le_imp_le_left) | |
| 259 | ||
| 260 | instance star :: (pordered_ab_group_add) pordered_ab_group_add .. | |
| 261 | instance star :: (ordered_cancel_ab_semigroup_add) ordered_cancel_ab_semigroup_add .. | |
| 262 | instance star :: (lordered_ab_group_meet) lordered_ab_group_meet .. | |
| 263 | instance star :: (lordered_ab_group_meet) lordered_ab_group_meet .. | |
| 264 | instance star :: (lordered_ab_group) lordered_ab_group .. | |
| 265 | ||
| 266 | instance star :: (lordered_ab_group_abs) lordered_ab_group_abs | |
| 17332 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 267 | by (intro_classes, transfer, rule abs_lattice) | 
| 17296 | 268 | |
| 17429 
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
 huffman parents: 
17332diff
changeset | 269 | subsection {* Ring and field classes *}
 | 
| 17296 | 270 | |
| 271 | instance star :: (semiring) semiring | |
| 272 | apply (intro_classes) | |
| 273 | apply (transfer, rule left_distrib) | |
| 274 | apply (transfer, rule right_distrib) | |
| 275 | done | |
| 276 | ||
| 277 | instance star :: (semiring_0) semiring_0 .. | |
| 278 | instance star :: (semiring_0_cancel) semiring_0_cancel .. | |
| 279 | ||
| 280 | instance star :: (comm_semiring) comm_semiring | |
| 281 | by (intro_classes, transfer, rule distrib) | |
| 282 | ||
| 283 | instance star :: (comm_semiring_0) comm_semiring_0 .. | |
| 284 | instance star :: (comm_semiring_0_cancel) comm_semiring_0_cancel .. | |
| 285 | ||
| 286 | instance star :: (axclass_0_neq_1) axclass_0_neq_1 | |
| 287 | by (intro_classes, transfer, rule zero_neq_one) | |
| 288 | ||
| 289 | instance star :: (semiring_1) semiring_1 .. | |
| 290 | instance star :: (comm_semiring_1) comm_semiring_1 .. | |
| 291 | ||
| 292 | instance star :: (axclass_no_zero_divisors) axclass_no_zero_divisors | |
| 293 | by (intro_classes, transfer, rule no_zero_divisors) | |
| 294 | ||
| 295 | instance star :: (semiring_1_cancel) semiring_1_cancel .. | |
| 296 | instance star :: (comm_semiring_1_cancel) comm_semiring_1_cancel .. | |
| 297 | instance star :: (ring) ring .. | |
| 298 | instance star :: (comm_ring) comm_ring .. | |
| 299 | instance star :: (ring_1) ring_1 .. | |
| 300 | instance star :: (comm_ring_1) comm_ring_1 .. | |
| 301 | instance star :: (idom) idom .. | |
| 302 | ||
| 20540 | 303 | instance star :: (division_ring) division_ring | 
| 304 | apply (intro_classes) | |
| 305 | apply (transfer, erule left_inverse) | |
| 306 | apply (transfer, erule right_inverse) | |
| 307 | done | |
| 308 | ||
| 17296 | 309 | instance star :: (field) field | 
| 310 | apply (intro_classes) | |
| 311 | apply (transfer, erule left_inverse) | |
| 312 | apply (transfer, rule divide_inverse) | |
| 313 | done | |
| 314 | ||
| 315 | instance star :: (division_by_zero) division_by_zero | |
| 316 | by (intro_classes, transfer, rule inverse_zero) | |
| 317 | ||
| 318 | instance star :: (pordered_semiring) pordered_semiring | |
| 319 | apply (intro_classes) | |
| 320 | apply (transfer, erule (1) mult_left_mono) | |
| 321 | apply (transfer, erule (1) mult_right_mono) | |
| 322 | done | |
| 323 | ||
| 324 | instance star :: (pordered_cancel_semiring) pordered_cancel_semiring .. | |
| 325 | ||
| 326 | instance star :: (ordered_semiring_strict) ordered_semiring_strict | |
| 327 | apply (intro_classes) | |
| 328 | apply (transfer, erule (1) mult_strict_left_mono) | |
| 329 | apply (transfer, erule (1) mult_strict_right_mono) | |
| 330 | done | |
| 331 | ||
| 332 | instance star :: (pordered_comm_semiring) pordered_comm_semiring | |
| 333 | by (intro_classes, transfer, rule pordered_comm_semiring_class.mult_mono) | |
| 334 | ||
| 335 | instance star :: (pordered_cancel_comm_semiring) pordered_cancel_comm_semiring .. | |
| 336 | ||
| 337 | instance star :: (ordered_comm_semiring_strict) ordered_comm_semiring_strict | |
| 338 | by (intro_classes, transfer, rule ordered_comm_semiring_strict_class.mult_strict_mono) | |
| 339 | ||
| 340 | instance star :: (pordered_ring) pordered_ring .. | |
| 341 | instance star :: (lordered_ring) lordered_ring .. | |
| 342 | ||
| 343 | instance star :: (axclass_abs_if) axclass_abs_if | |
| 344 | by (intro_classes, transfer, rule abs_if) | |
| 345 | ||
| 346 | instance star :: (ordered_ring_strict) ordered_ring_strict .. | |
| 347 | instance star :: (pordered_comm_ring) pordered_comm_ring .. | |
| 348 | ||
| 349 | instance star :: (ordered_semidom) ordered_semidom | |
| 350 | by (intro_classes, transfer, rule zero_less_one) | |
| 351 | ||
| 352 | instance star :: (ordered_idom) ordered_idom .. | |
| 353 | instance star :: (ordered_field) ordered_field .. | |
| 354 | ||
| 17332 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 355 | subsection {* Power classes *}
 | 
| 17296 | 356 | |
| 357 | text {*
 | |
| 358 |   Proving the class axiom @{thm [source] power_Suc} for type
 | |
| 359 |   @{typ "'a star"} is a little tricky, because it quantifies
 | |
| 360 |   over values of type @{typ nat}. The transfer principle does
 | |
| 361 | not handle quantification over non-star types in general, | |
| 362 |   but we can work around this by fixing an arbitrary @{typ nat}
 | |
| 363 | value, and then applying the transfer principle. | |
| 364 | *} | |
| 365 | ||
| 366 | instance star :: (recpower) recpower | |
| 367 | proof | |
| 368 | show "\<And>a::'a star. a ^ 0 = 1" | |
| 369 | by transfer (rule power_0) | |
| 370 | next | |
| 371 | fix n show "\<And>a::'a star. a ^ Suc n = a * a ^ n" | |
| 372 | by transfer (rule power_Suc) | |
| 373 | qed | |
| 374 | ||
| 17332 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 375 | subsection {* Number classes *}
 | 
| 17296 | 376 | |
| 17332 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 377 | lemma star_of_nat_def [transfer_unfold]: "of_nat n \<equiv> star_of (of_nat n)" | 
| 17296 | 378 | by (rule eq_reflection, induct_tac n, simp_all) | 
| 379 | ||
| 17332 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 380 | lemma star_of_of_nat [simp]: "star_of (of_nat n) = of_nat n" | 
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 381 | by transfer (rule refl) | 
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 382 | |
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 383 | lemma star_of_int_def [transfer_unfold]: "of_int z \<equiv> star_of (of_int z)" | 
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 384 | by (rule eq_reflection, rule_tac z=z in int_diff_cases, simp) | 
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 385 | |
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 386 | lemma star_of_of_int [simp]: "star_of (of_int z) = of_int z" | 
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17296diff
changeset | 387 | by transfer (rule refl) | 
| 17296 | 388 | |
| 389 | instance star :: (number_ring) number_ring | |
| 390 | by (intro_classes, simp only: star_number_def star_of_int_def number_of_eq) | |
| 391 | ||
| 17429 
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
 huffman parents: 
17332diff
changeset | 392 | subsection {* Finite class *}
 | 
| 
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
 huffman parents: 
17332diff
changeset | 393 | |
| 
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
 huffman parents: 
17332diff
changeset | 394 | lemma starset_finite: "finite A \<Longrightarrow> *s* A = star_of ` A" | 
| 
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
 huffman parents: 
17332diff
changeset | 395 | by (erule finite_induct, simp_all) | 
| 
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
 huffman parents: 
17332diff
changeset | 396 | |
| 
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
 huffman parents: 
17332diff
changeset | 397 | instance star :: (finite) finite | 
| 
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
 huffman parents: 
17332diff
changeset | 398 | apply (intro_classes) | 
| 
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
 huffman parents: 
17332diff
changeset | 399 | apply (subst starset_UNIV [symmetric]) | 
| 
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
 huffman parents: 
17332diff
changeset | 400 | apply (subst starset_finite [OF finite]) | 
| 
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
 huffman parents: 
17332diff
changeset | 401 | apply (rule finite_imageI [OF finite]) | 
| 
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
 huffman parents: 
17332diff
changeset | 402 | done | 
| 
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
 huffman parents: 
17332diff
changeset | 403 | |
| 17296 | 404 | end |