| author | Cezary Kaliszyk <kaliszyk@in.tum.de> | 
| Thu, 29 Oct 2009 13:29:03 +0100 | |
| changeset 236 | 23f9fead8bd6 | 
| parent 228 | 268a727b0f10 | 
| child 239 | 02b14a21761a | 
| permissions | -rw-r--r-- | 
| 181 
3e53081ad53a
added another example file about integers (see HOL/Int.thy)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1 | theory IntEx | 
| 
3e53081ad53a
added another example file about integers (see HOL/Int.thy)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 2 | imports QuotMain | 
| 
3e53081ad53a
added another example file about integers (see HOL/Int.thy)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 3 | begin | 
| 
3e53081ad53a
added another example file about integers (see HOL/Int.thy)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 4 | |
| 
3e53081ad53a
added another example file about integers (see HOL/Int.thy)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 5 | fun | 
| 206 
1e227c9ee915
Fixed APPLY_RSP vs Cong in the InjRepAbs tactic.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
200diff
changeset | 6 | intrel :: "(nat \<times> nat) \<Rightarrow> (nat \<times> nat) \<Rightarrow> bool" | 
| 181 
3e53081ad53a
added another example file about integers (see HOL/Int.thy)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 7 | where | 
| 
3e53081ad53a
added another example file about integers (see HOL/Int.thy)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 8 | "intrel (x, y) (u, v) = (x + v = u + y)" | 
| 
3e53081ad53a
added another example file about integers (see HOL/Int.thy)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 9 | |
| 
3e53081ad53a
added another example file about integers (see HOL/Int.thy)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 10 | quotient my_int = "nat \<times> nat" / intrel | 
| 
3e53081ad53a
added another example file about integers (see HOL/Int.thy)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 11 | apply(unfold EQUIV_def) | 
| 
3e53081ad53a
added another example file about integers (see HOL/Int.thy)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 12 | apply(auto simp add: mem_def expand_fun_eq) | 
| 
3e53081ad53a
added another example file about integers (see HOL/Int.thy)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 13 | done | 
| 
3e53081ad53a
added another example file about integers (see HOL/Int.thy)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 14 | |
| 218 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
210diff
changeset | 15 | quotient_def (for my_int) | 
| 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
210diff
changeset | 16 | ZERO::"my_int" | 
| 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
210diff
changeset | 17 | where | 
| 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
210diff
changeset | 18 | "ZERO \<equiv> (0::nat, 0::nat)" | 
| 181 
3e53081ad53a
added another example file about integers (see HOL/Int.thy)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 19 | |
| 
3e53081ad53a
added another example file about integers (see HOL/Int.thy)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 20 | thm ZERO_def | 
| 
3e53081ad53a
added another example file about integers (see HOL/Int.thy)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 21 | |
| 218 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
210diff
changeset | 22 | quotient_def (for my_int) | 
| 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
210diff
changeset | 23 | ONE::"my_int" | 
| 193 | 24 | where | 
| 218 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
210diff
changeset | 25 | "ONE \<equiv> (1::nat, 0::nat)" | 
| 181 
3e53081ad53a
added another example file about integers (see HOL/Int.thy)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 26 | |
| 
3e53081ad53a
added another example file about integers (see HOL/Int.thy)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 27 | term ONE | 
| 
3e53081ad53a
added another example file about integers (see HOL/Int.thy)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 28 | thm ONE_def | 
| 
3e53081ad53a
added another example file about integers (see HOL/Int.thy)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 29 | |
| 
3e53081ad53a
added another example file about integers (see HOL/Int.thy)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 30 | fun | 
| 
3e53081ad53a
added another example file about integers (see HOL/Int.thy)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 31 | my_plus :: "(nat \<times> nat) \<Rightarrow> (nat \<times> nat) \<Rightarrow> (nat \<times> nat)" | 
| 
3e53081ad53a
added another example file about integers (see HOL/Int.thy)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 32 | where | 
| 
3e53081ad53a
added another example file about integers (see HOL/Int.thy)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 33 | "my_plus (x, y) (u, v) = (x + u, y + v)" | 
| 
3e53081ad53a
added another example file about integers (see HOL/Int.thy)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 34 | |
| 218 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
210diff
changeset | 35 | quotient_def (for my_int) | 
| 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
210diff
changeset | 36 | PLUS::"my_int \<Rightarrow> my_int \<Rightarrow> my_int" | 
| 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
210diff
changeset | 37 | where | 
| 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
210diff
changeset | 38 | "PLUS \<equiv> my_plus" | 
| 181 
3e53081ad53a
added another example file about integers (see HOL/Int.thy)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 39 | |
| 
3e53081ad53a
added another example file about integers (see HOL/Int.thy)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 40 | term PLUS | 
| 
3e53081ad53a
added another example file about integers (see HOL/Int.thy)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 41 | thm PLUS_def | 
| 
3e53081ad53a
added another example file about integers (see HOL/Int.thy)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 42 | |
| 
3e53081ad53a
added another example file about integers (see HOL/Int.thy)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 43 | fun | 
| 
3e53081ad53a
added another example file about integers (see HOL/Int.thy)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 44 | my_neg :: "(nat \<times> nat) \<Rightarrow> (nat \<times> nat)" | 
| 
3e53081ad53a
added another example file about integers (see HOL/Int.thy)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 45 | where | 
| 
3e53081ad53a
added another example file about integers (see HOL/Int.thy)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 46 | "my_neg (x, y) = (y, x)" | 
| 
3e53081ad53a
added another example file about integers (see HOL/Int.thy)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 47 | |
| 220 | 48 | quotient_def (for my_int) | 
| 49 | NEG::"my_int \<Rightarrow> my_int" | |
| 50 | where | |
| 51 | "NEG \<equiv> my_neg" | |
| 181 
3e53081ad53a
added another example file about integers (see HOL/Int.thy)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 52 | |
| 
3e53081ad53a
added another example file about integers (see HOL/Int.thy)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 53 | term NEG | 
| 
3e53081ad53a
added another example file about integers (see HOL/Int.thy)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 54 | thm NEG_def | 
| 
3e53081ad53a
added another example file about integers (see HOL/Int.thy)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 55 | |
| 
3e53081ad53a
added another example file about integers (see HOL/Int.thy)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 56 | definition | 
| 
3e53081ad53a
added another example file about integers (see HOL/Int.thy)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 57 | MINUS :: "my_int \<Rightarrow> my_int \<Rightarrow> my_int" | 
| 
3e53081ad53a
added another example file about integers (see HOL/Int.thy)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 58 | where | 
| 
3e53081ad53a
added another example file about integers (see HOL/Int.thy)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 59 | "MINUS z w = PLUS z (NEG w)" | 
| 
3e53081ad53a
added another example file about integers (see HOL/Int.thy)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 60 | |
| 
3e53081ad53a
added another example file about integers (see HOL/Int.thy)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 61 | fun | 
| 
3e53081ad53a
added another example file about integers (see HOL/Int.thy)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 62 | my_mult :: "(nat \<times> nat) \<Rightarrow> (nat \<times> nat) \<Rightarrow> (nat \<times> nat)" | 
| 
3e53081ad53a
added another example file about integers (see HOL/Int.thy)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 63 | where | 
| 
3e53081ad53a
added another example file about integers (see HOL/Int.thy)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 64 | "my_mult (x, y) (u, v) = (x*u + y*v, x*v + y*u)" | 
| 
3e53081ad53a
added another example file about integers (see HOL/Int.thy)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 65 | |
| 220 | 66 | quotient_def (for my_int) | 
| 67 | MULT::"my_int \<Rightarrow> my_int \<Rightarrow> my_int" | |
| 68 | where | |
| 69 | "MULT \<equiv> my_mult" | |
| 181 
3e53081ad53a
added another example file about integers (see HOL/Int.thy)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 70 | |
| 
3e53081ad53a
added another example file about integers (see HOL/Int.thy)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 71 | term MULT | 
| 
3e53081ad53a
added another example file about integers (see HOL/Int.thy)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 72 | thm MULT_def | 
| 
3e53081ad53a
added another example file about integers (see HOL/Int.thy)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 73 | |
| 
3e53081ad53a
added another example file about integers (see HOL/Int.thy)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 74 | (* NOT SURE WETHER THIS DEFINITION IS CORRECT *) | 
| 
3e53081ad53a
added another example file about integers (see HOL/Int.thy)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 75 | fun | 
| 
3e53081ad53a
added another example file about integers (see HOL/Int.thy)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 76 | my_le :: "(nat \<times> nat) \<Rightarrow> (nat \<times> nat) \<Rightarrow> bool" | 
| 
3e53081ad53a
added another example file about integers (see HOL/Int.thy)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 77 | where | 
| 
3e53081ad53a
added another example file about integers (see HOL/Int.thy)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 78 | "my_le (x, y) (u, v) = (x+v \<le> u+y)" | 
| 
3e53081ad53a
added another example file about integers (see HOL/Int.thy)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 79 | |
| 220 | 80 | quotient_def (for my_int) | 
| 81 | LE :: "my_int \<Rightarrow> my_int \<Rightarrow> bool" | |
| 82 | where | |
| 83 | "LE \<equiv> my_le" | |
| 181 
3e53081ad53a
added another example file about integers (see HOL/Int.thy)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 84 | |
| 
3e53081ad53a
added another example file about integers (see HOL/Int.thy)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 85 | term LE | 
| 
3e53081ad53a
added another example file about integers (see HOL/Int.thy)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 86 | thm LE_def | 
| 
3e53081ad53a
added another example file about integers (see HOL/Int.thy)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 87 | |
| 
3e53081ad53a
added another example file about integers (see HOL/Int.thy)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 88 | definition | 
| 
3e53081ad53a
added another example file about integers (see HOL/Int.thy)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 89 | LESS :: "my_int \<Rightarrow> my_int \<Rightarrow> bool" | 
| 
3e53081ad53a
added another example file about integers (see HOL/Int.thy)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 90 | where | 
| 
3e53081ad53a
added another example file about integers (see HOL/Int.thy)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 91 | "LESS z w = (LE z w \<and> z \<noteq> w)" | 
| 
3e53081ad53a
added another example file about integers (see HOL/Int.thy)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 92 | |
| 
3e53081ad53a
added another example file about integers (see HOL/Int.thy)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 93 | term LESS | 
| 
3e53081ad53a
added another example file about integers (see HOL/Int.thy)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 94 | thm LESS_def | 
| 
3e53081ad53a
added another example file about integers (see HOL/Int.thy)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 95 | |
| 
3e53081ad53a
added another example file about integers (see HOL/Int.thy)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 96 | definition | 
| 
3e53081ad53a
added another example file about integers (see HOL/Int.thy)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 97 | ABS :: "my_int \<Rightarrow> my_int" | 
| 
3e53081ad53a
added another example file about integers (see HOL/Int.thy)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 98 | where | 
| 
3e53081ad53a
added another example file about integers (see HOL/Int.thy)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 99 | "ABS i = (if (LESS i ZERO) then (NEG i) else i)" | 
| 
3e53081ad53a
added another example file about integers (see HOL/Int.thy)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 100 | |
| 
3e53081ad53a
added another example file about integers (see HOL/Int.thy)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 101 | definition | 
| 
3e53081ad53a
added another example file about integers (see HOL/Int.thy)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 102 | SIGN :: "my_int \<Rightarrow> my_int" | 
| 
3e53081ad53a
added another example file about integers (see HOL/Int.thy)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 103 | where | 
| 
3e53081ad53a
added another example file about integers (see HOL/Int.thy)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 104 | "SIGN i = (if i = ZERO then ZERO else if (LESS ZERO i) then ONE else (NEG ONE))" | 
| 
3e53081ad53a
added another example file about integers (see HOL/Int.thy)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 105 | |
| 199 | 106 | lemma plus_sym_pre: | 
| 107 | shows "intrel (my_plus a b) (my_plus b a)" | |
| 108 | apply(cases a) | |
| 109 | apply(cases b) | |
| 110 | apply(auto) | |
| 111 | done | |
| 112 | ||
| 206 
1e227c9ee915
Fixed APPLY_RSP vs Cong in the InjRepAbs tactic.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
200diff
changeset | 113 | lemma ho_plus_rsp: | 
| 228 
268a727b0f10
disambiguate ===> syntax
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
224diff
changeset | 114 | "(intrel ===> intrel ===> intrel) my_plus my_plus" | 
| 192 
a296bf1a3b09
Simplifying code in int
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
191diff
changeset | 115 | by (simp) | 
| 
a296bf1a3b09
Simplifying code in int
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
191diff
changeset | 116 | |
| 
a296bf1a3b09
Simplifying code in int
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
191diff
changeset | 117 | ML {* val consts = [@{const_name "my_plus"}] *}
 | 
| 
a296bf1a3b09
Simplifying code in int
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
191diff
changeset | 118 | ML {* val rty = @{typ "nat \<times> nat"} *}
 | 
| 
a296bf1a3b09
Simplifying code in int
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
191diff
changeset | 119 | ML {* val qty = @{typ "my_int"} *}
 | 
| 
a296bf1a3b09
Simplifying code in int
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
191diff
changeset | 120 | ML {* val rel = @{term "intrel"} *}
 | 
| 224 
f9a25fe22037
Cleaning the unnecessary theorems in 'IntEx'.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
222diff
changeset | 121 | ML {* val rel_eqv = (#equiv_thm o hd) (quotdata_lookup @{context}) *}
 | 
| 
f9a25fe22037
Cleaning the unnecessary theorems in 'IntEx'.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
222diff
changeset | 122 | ML {* val rel_refl = @{thm EQUIV_REFL} OF [rel_eqv] *}
 | 
| 192 
a296bf1a3b09
Simplifying code in int
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
191diff
changeset | 123 | ML {* val quot = @{thm QUOTIENT_my_int} *}
 | 
| 
a296bf1a3b09
Simplifying code in int
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
191diff
changeset | 124 | ML {* val rsp_thms = @{thms ho_plus_rsp} @ @{thms ho_all_prs ho_ex_prs} *}
 | 
| 
a296bf1a3b09
Simplifying code in int
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
191diff
changeset | 125 | ML {* val trans2 = @{thm QUOT_TYPE_I_my_int.R_trans2} *}
 | 
| 197 
c0f2db9a243b
Further reordering in Int code.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
196diff
changeset | 126 | ML {* val reps_same = @{thm QUOT_TYPE_I_my_int.REPS_same} *}
 | 
| 
c0f2db9a243b
Further reordering in Int code.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
196diff
changeset | 127 | ML {* val t_defs = @{thms PLUS_def} *}
 | 
| 
c0f2db9a243b
Further reordering in Int code.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
196diff
changeset | 128 | |
| 198 
ff4425e000db
Completely cleaned Int.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
197diff
changeset | 129 | ML {*
 | 
| 
ff4425e000db
Completely cleaned Int.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
197diff
changeset | 130 | fun lift_thm_my_int lthy t = | 
| 
ff4425e000db
Completely cleaned Int.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
197diff
changeset | 131 | lift_thm lthy consts rty qty rel rel_eqv rel_refl quot rsp_thms trans2 reps_same t_defs t | 
| 
ff4425e000db
Completely cleaned Int.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
197diff
changeset | 132 | *} | 
| 
ff4425e000db
Completely cleaned Int.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
197diff
changeset | 133 | |
| 
ff4425e000db
Completely cleaned Int.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
197diff
changeset | 134 | ML {* lift_thm_my_int @{context} @{thm plus_sym_pre} *}
 | 
| 
ff4425e000db
Completely cleaned Int.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
197diff
changeset | 135 | |
| 
ff4425e000db
Completely cleaned Int.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
197diff
changeset | 136 | lemma plus_assoc_pre: | 
| 
ff4425e000db
Completely cleaned Int.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
197diff
changeset | 137 | shows "intrel (my_plus (my_plus i j) k) (my_plus i (my_plus j k))" | 
| 
ff4425e000db
Completely cleaned Int.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
197diff
changeset | 138 | apply (cases i) | 
| 
ff4425e000db
Completely cleaned Int.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
197diff
changeset | 139 | apply (cases j) | 
| 
ff4425e000db
Completely cleaned Int.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
197diff
changeset | 140 | apply (cases k) | 
| 224 
f9a25fe22037
Cleaning the unnecessary theorems in 'IntEx'.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
222diff
changeset | 141 | apply (simp) | 
| 198 
ff4425e000db
Completely cleaned Int.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
197diff
changeset | 142 | done | 
| 
ff4425e000db
Completely cleaned Int.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
197diff
changeset | 143 | |
| 224 
f9a25fe22037
Cleaning the unnecessary theorems in 'IntEx'.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
222diff
changeset | 144 | ML {* lift_thm_my_int @{context} @{thm plus_assoc_pre} *}
 | 
| 
f9a25fe22037
Cleaning the unnecessary theorems in 'IntEx'.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
222diff
changeset | 145 | |
| 
f9a25fe22037
Cleaning the unnecessary theorems in 'IntEx'.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
222diff
changeset | 146 | |
| 
f9a25fe22037
Cleaning the unnecessary theorems in 'IntEx'.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
222diff
changeset | 147 | |
| 
f9a25fe22037
Cleaning the unnecessary theorems in 'IntEx'.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
222diff
changeset | 148 | |
| 
f9a25fe22037
Cleaning the unnecessary theorems in 'IntEx'.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
222diff
changeset | 149 | |
| 197 
c0f2db9a243b
Further reordering in Int code.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
196diff
changeset | 150 | |
| 191 
b97f3f5fbc18
Symmetry of integer addition
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
184diff
changeset | 151 | |
| 224 
f9a25fe22037
Cleaning the unnecessary theorems in 'IntEx'.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
222diff
changeset | 152 | |
| 
f9a25fe22037
Cleaning the unnecessary theorems in 'IntEx'.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
222diff
changeset | 153 | text {* Below is the construction site code used if things do not work *}
 | 
| 
f9a25fe22037
Cleaning the unnecessary theorems in 'IntEx'.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
222diff
changeset | 154 | |
| 
f9a25fe22037
Cleaning the unnecessary theorems in 'IntEx'.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
222diff
changeset | 155 | |
| 
f9a25fe22037
Cleaning the unnecessary theorems in 'IntEx'.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
222diff
changeset | 156 | ML {* val abs = findabs rty (prop_of (atomize_thm @{thm plus_sym_pre})) *}
 | 
| 
f9a25fe22037
Cleaning the unnecessary theorems in 'IntEx'.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
222diff
changeset | 157 | ML {* val simp_lam_prs_thms = map (make_simp_lam_prs_thm @{context} quot) abs *}
 | 
| 
f9a25fe22037
Cleaning the unnecessary theorems in 'IntEx'.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
222diff
changeset | 158 | ML {* val t_defs_sym = add_lower_defs @{context} t_defs *}
 | 
| 
f9a25fe22037
Cleaning the unnecessary theorems in 'IntEx'.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
222diff
changeset | 159 | |
| 198 
ff4425e000db
Completely cleaned Int.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
197diff
changeset | 160 | |
| 221 | 161 | ML {* val t_a = atomize_thm @{thm plus_sym_pre} *}
 | 
| 192 
a296bf1a3b09
Simplifying code in int
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
191diff
changeset | 162 | ML {* val t_r = regularize t_a rty rel rel_eqv @{context} *}
 | 
| 224 
f9a25fe22037
Cleaning the unnecessary theorems in 'IntEx'.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
222diff
changeset | 163 | ML {* val t_t = repabs @{context} t_r consts rty qty quot rel_refl trans2 rsp_thms *}
 | 
| 221 | 164 | ML {* val t_l = repeat_eqsubst_thm @{context} simp_lam_prs_thms t_t *}
 | 
| 196 | 165 | ML {* val t_a = simp_allex_prs @{context} quot t_l *}
 | 
| 221 | 166 | ML {* val t_d = repeat_eqsubst_thm @{context} t_defs_sym t_a *}
 | 
| 197 
c0f2db9a243b
Further reordering in Int code.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
196diff
changeset | 167 | ML {* val t_r = MetaSimplifier.rewrite_rule [reps_same] t_d *}
 | 
| 191 
b97f3f5fbc18
Symmetry of integer addition
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
184diff
changeset | 168 | ML {* ObjectLogic.rulify t_r *}
 | 
| 181 
3e53081ad53a
added another example file about integers (see HOL/Int.thy)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 169 | |
| 
3e53081ad53a
added another example file about integers (see HOL/Int.thy)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 170 | |
| 
3e53081ad53a
added another example file about integers (see HOL/Int.thy)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 171 |