| author | Cezary Kaliszyk <kaliszyk@in.tum.de> | 
| Fri, 11 Dec 2009 11:08:58 +0100 | |
| changeset 705 | f51c6069cd17 | 
| parent 692 | c9231c2903bc | 
| child 758 | 3104d62e7a16 | 
| permissions | -rw-r--r-- | 
| 597 | 1 | theory IntEx | 
| 663 
0dd10a900cae
Different syntax for definitions that allows overloading and retrieving of definitions by matching whole constants.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
656diff
changeset | 2 | imports "../QuotProd" "../QuotList" | 
| 597 | 3 | begin | 
| 4 | ||
| 5 | fun | |
| 6 | intrel :: "(nat \<times> nat) \<Rightarrow> (nat \<times> nat) \<Rightarrow> bool" (infix "\<approx>" 50) | |
| 7 | where | |
| 8 | "intrel (x, y) (u, v) = (x + v = u + y)" | |
| 9 | ||
| 10 | quotient my_int = "nat \<times> nat" / intrel | |
| 11 | apply(unfold equivp_def) | |
| 12 | apply(auto simp add: mem_def expand_fun_eq) | |
| 13 | done | |
| 14 | ||
| 636 
520a4084d064
changed names of attributes
 Christian Urban <urbanc@in.tum.de> parents: 
632diff
changeset | 15 | thm quot_equiv | 
| 597 | 16 | |
| 636 
520a4084d064
changed names of attributes
 Christian Urban <urbanc@in.tum.de> parents: 
632diff
changeset | 17 | thm quot_thm | 
| 597 | 18 | |
| 19 | thm my_int_equivp | |
| 20 | ||
| 21 | print_theorems | |
| 22 | print_quotients | |
| 23 | ||
| 663 
0dd10a900cae
Different syntax for definitions that allows overloading and retrieving of definitions by matching whole constants.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
656diff
changeset | 24 | quotient_def | 
| 705 
f51c6069cd17
New syntax for definitions.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
692diff
changeset | 25 | "ZERO :: my_int" | 
| 
f51c6069cd17
New syntax for definitions.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
692diff
changeset | 26 | as | 
| 663 
0dd10a900cae
Different syntax for definitions that allows overloading and retrieving of definitions by matching whole constants.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
656diff
changeset | 27 | "(0::nat, 0::nat)" | 
| 597 | 28 | |
| 663 
0dd10a900cae
Different syntax for definitions that allows overloading and retrieving of definitions by matching whole constants.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
656diff
changeset | 29 | quotient_def | 
| 705 
f51c6069cd17
New syntax for definitions.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
692diff
changeset | 30 | "ONE :: my_int" | 
| 
f51c6069cd17
New syntax for definitions.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
692diff
changeset | 31 | as | 
| 663 
0dd10a900cae
Different syntax for definitions that allows overloading and retrieving of definitions by matching whole constants.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
656diff
changeset | 32 | "(1::nat, 0::nat)" | 
| 597 | 33 | |
| 34 | fun | |
| 35 | my_plus :: "(nat \<times> nat) \<Rightarrow> (nat \<times> nat) \<Rightarrow> (nat \<times> nat)" | |
| 36 | where | |
| 37 | "my_plus (x, y) (u, v) = (x + u, y + v)" | |
| 38 | ||
| 663 
0dd10a900cae
Different syntax for definitions that allows overloading and retrieving of definitions by matching whole constants.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
656diff
changeset | 39 | quotient_def | 
| 705 
f51c6069cd17
New syntax for definitions.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
692diff
changeset | 40 | "PLUS :: my_int \<Rightarrow> my_int \<Rightarrow> my_int" | 
| 
f51c6069cd17
New syntax for definitions.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
692diff
changeset | 41 | as | 
| 663 
0dd10a900cae
Different syntax for definitions that allows overloading and retrieving of definitions by matching whole constants.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
656diff
changeset | 42 | "my_plus" | 
| 597 | 43 | |
| 44 | fun | |
| 45 | my_neg :: "(nat \<times> nat) \<Rightarrow> (nat \<times> nat)" | |
| 46 | where | |
| 47 | "my_neg (x, y) = (y, x)" | |
| 48 | ||
| 663 
0dd10a900cae
Different syntax for definitions that allows overloading and retrieving of definitions by matching whole constants.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
656diff
changeset | 49 | quotient_def | 
| 705 
f51c6069cd17
New syntax for definitions.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
692diff
changeset | 50 | "NEG :: my_int \<Rightarrow> my_int" | 
| 
f51c6069cd17
New syntax for definitions.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
692diff
changeset | 51 | as | 
| 663 
0dd10a900cae
Different syntax for definitions that allows overloading and retrieving of definitions by matching whole constants.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
656diff
changeset | 52 | "my_neg" | 
| 597 | 53 | |
| 54 | definition | |
| 55 | MINUS :: "my_int \<Rightarrow> my_int \<Rightarrow> my_int" | |
| 56 | where | |
| 57 | "MINUS z w = PLUS z (NEG w)" | |
| 58 | ||
| 59 | fun | |
| 60 | my_mult :: "(nat \<times> nat) \<Rightarrow> (nat \<times> nat) \<Rightarrow> (nat \<times> nat)" | |
| 61 | where | |
| 62 | "my_mult (x, y) (u, v) = (x*u + y*v, x*v + y*u)" | |
| 63 | ||
| 663 
0dd10a900cae
Different syntax for definitions that allows overloading and retrieving of definitions by matching whole constants.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
656diff
changeset | 64 | quotient_def | 
| 705 
f51c6069cd17
New syntax for definitions.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
692diff
changeset | 65 | "MULT :: my_int \<Rightarrow> my_int \<Rightarrow> my_int" | 
| 
f51c6069cd17
New syntax for definitions.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
692diff
changeset | 66 | as | 
| 663 
0dd10a900cae
Different syntax for definitions that allows overloading and retrieving of definitions by matching whole constants.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
656diff
changeset | 67 | "my_mult" | 
| 597 | 68 | |
| 69 | ||
| 70 | (* NOT SURE WETHER THIS DEFINITION IS CORRECT *) | |
| 71 | fun | |
| 72 | my_le :: "(nat \<times> nat) \<Rightarrow> (nat \<times> nat) \<Rightarrow> bool" | |
| 73 | where | |
| 74 | "my_le (x, y) (u, v) = (x+v \<le> u+y)" | |
| 75 | ||
| 663 
0dd10a900cae
Different syntax for definitions that allows overloading and retrieving of definitions by matching whole constants.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
656diff
changeset | 76 | quotient_def | 
| 705 
f51c6069cd17
New syntax for definitions.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
692diff
changeset | 77 | "LE :: my_int \<Rightarrow> my_int \<Rightarrow> bool" | 
| 
f51c6069cd17
New syntax for definitions.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
692diff
changeset | 78 | as | 
| 663 
0dd10a900cae
Different syntax for definitions that allows overloading and retrieving of definitions by matching whole constants.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
656diff
changeset | 79 | "my_le" | 
| 597 | 80 | |
| 81 | term LE | |
| 82 | thm LE_def | |
| 83 | ||
| 84 | ||
| 85 | definition | |
| 86 | LESS :: "my_int \<Rightarrow> my_int \<Rightarrow> bool" | |
| 87 | where | |
| 88 | "LESS z w = (LE z w \<and> z \<noteq> w)" | |
| 89 | ||
| 90 | term LESS | |
| 91 | thm LESS_def | |
| 92 | ||
| 93 | definition | |
| 94 | ABS :: "my_int \<Rightarrow> my_int" | |
| 95 | where | |
| 96 | "ABS i = (if (LESS i ZERO) then (NEG i) else i)" | |
| 97 | ||
| 98 | definition | |
| 99 | SIGN :: "my_int \<Rightarrow> my_int" | |
| 100 | where | |
| 101 | "SIGN i = (if i = ZERO then ZERO else if (LESS ZERO i) then ONE else (NEG ONE))" | |
| 102 | ||
| 103 | ML {* print_qconstinfo @{context} *}
 | |
| 104 | ||
| 105 | lemma plus_sym_pre: | |
| 106 | shows "my_plus a b \<approx> my_plus b a" | |
| 107 | apply(cases a) | |
| 108 | apply(cases b) | |
| 109 | apply(auto) | |
| 110 | done | |
| 111 | ||
| 636 
520a4084d064
changed names of attributes
 Christian Urban <urbanc@in.tum.de> parents: 
632diff
changeset | 112 | lemma plus_rsp[quot_respect]: | 
| 597 | 113 | shows "(intrel ===> intrel ===> intrel) my_plus my_plus" | 
| 114 | by (simp) | |
| 115 | ||
| 692 
c9231c2903bc
Moved 'int_induct' to IntEx to keep IntEx2 being just theory of integers in order.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
663diff
changeset | 116 | lemma neg_rsp[quot_respect]: | 
| 
c9231c2903bc
Moved 'int_induct' to IntEx to keep IntEx2 being just theory of integers in order.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
663diff
changeset | 117 | shows "(op \<approx> ===> op \<approx>) my_neg my_neg" | 
| 
c9231c2903bc
Moved 'int_induct' to IntEx to keep IntEx2 being just theory of integers in order.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
663diff
changeset | 118 | by simp | 
| 
c9231c2903bc
Moved 'int_induct' to IntEx to keep IntEx2 being just theory of integers in order.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
663diff
changeset | 119 | |
| 597 | 120 | lemma test1: "my_plus a b = my_plus a b" | 
| 121 | apply(rule refl) | |
| 122 | done | |
| 123 | ||
| 124 | lemma "PLUS a b = PLUS a b" | |
| 637 
b029f242d85d
chnaged syntax to "lifting theorem"
 Christian Urban <urbanc@in.tum.de> parents: 
636diff
changeset | 125 | apply(lifting_setup test1) | 
| 632 
d23416464f62
added methods for the lifting_tac and the other tacs
 Christian Urban <urbanc@in.tum.de> parents: 
625diff
changeset | 126 | apply(regularize) | 
| 
d23416464f62
added methods for the lifting_tac and the other tacs
 Christian Urban <urbanc@in.tum.de> parents: 
625diff
changeset | 127 | apply(injection) | 
| 
d23416464f62
added methods for the lifting_tac and the other tacs
 Christian Urban <urbanc@in.tum.de> parents: 
625diff
changeset | 128 | apply(cleaning) | 
| 597 | 129 | done | 
| 130 | ||
| 131 | thm lambda_prs | |
| 132 | ||
| 133 | lemma test2: "my_plus a = my_plus a" | |
| 134 | apply(rule refl) | |
| 135 | done | |
| 136 | ||
| 137 | lemma "PLUS a = PLUS a" | |
| 138 | apply(tactic {* procedure_tac @{context} @{thm test2} 1 *})
 | |
| 606 
38aa6b67a80b
clarified the function examples
 Christian Urban <urbanc@in.tum.de> parents: 
605diff
changeset | 139 | apply(rule impI) | 
| 597 | 140 | apply(rule ballI) | 
| 141 | apply(rule apply_rsp[OF Quotient_my_int plus_rsp]) | |
| 142 | apply(simp only: in_respects) | |
| 610 
2bee5ca44ef5
removed "global" data and lookup functions; had to move a tactic out from the inj_repabs_match tactic since apply_rsp interferes with a trans2 rule for ===>
 Christian Urban <urbanc@in.tum.de> parents: 
606diff
changeset | 143 | apply(tactic {* all_inj_repabs_tac @{context} 1*})
 | 
| 597 | 144 | apply(tactic {* clean_tac @{context} 1 *})
 | 
| 145 | done | |
| 146 | ||
| 147 | lemma test3: "my_plus = my_plus" | |
| 148 | apply(rule refl) | |
| 149 | done | |
| 150 | ||
| 151 | lemma "PLUS = PLUS" | |
| 152 | apply(tactic {* procedure_tac @{context} @{thm test3} 1 *})
 | |
| 606 
38aa6b67a80b
clarified the function examples
 Christian Urban <urbanc@in.tum.de> parents: 
605diff
changeset | 153 | apply(rule impI) | 
| 597 | 154 | apply(rule plus_rsp) | 
| 636 
520a4084d064
changed names of attributes
 Christian Urban <urbanc@in.tum.de> parents: 
632diff
changeset | 155 | apply(injection) | 
| 
520a4084d064
changed names of attributes
 Christian Urban <urbanc@in.tum.de> parents: 
632diff
changeset | 156 | apply(cleaning) | 
| 597 | 157 | done | 
| 158 | ||
| 159 | ||
| 160 | lemma "PLUS a b = PLUS b a" | |
| 161 | apply(tactic {* procedure_tac @{context} @{thm plus_sym_pre} 1 *})
 | |
| 162 | apply(tactic {* regularize_tac @{context} 1 *})
 | |
| 610 
2bee5ca44ef5
removed "global" data and lookup functions; had to move a tactic out from the inj_repabs_match tactic since apply_rsp interferes with a trans2 rule for ===>
 Christian Urban <urbanc@in.tum.de> parents: 
606diff
changeset | 163 | apply(tactic {* all_inj_repabs_tac @{context} 1*})
 | 
| 597 | 164 | apply(tactic {* clean_tac @{context} 1 *})
 | 
| 165 | done | |
| 166 | ||
| 167 | lemma plus_assoc_pre: | |
| 168 | shows "my_plus (my_plus i j) k \<approx> my_plus i (my_plus j k)" | |
| 169 | apply (cases i) | |
| 170 | apply (cases j) | |
| 171 | apply (cases k) | |
| 172 | apply (simp) | |
| 173 | done | |
| 174 | ||
| 175 | lemma plus_assoc: "PLUS (PLUS x xa) xb = PLUS x (PLUS xa xb)" | |
| 176 | apply(tactic {* procedure_tac @{context} @{thm plus_assoc_pre} 1 *})
 | |
| 177 | apply(tactic {* regularize_tac @{context} 1 *})
 | |
| 610 
2bee5ca44ef5
removed "global" data and lookup functions; had to move a tactic out from the inj_repabs_match tactic since apply_rsp interferes with a trans2 rule for ===>
 Christian Urban <urbanc@in.tum.de> parents: 
606diff
changeset | 178 | apply(tactic {* all_inj_repabs_tac @{context} 1*})
 | 
| 597 | 179 | apply(tactic {* clean_tac @{context} 1 *})
 | 
| 180 | done | |
| 181 | ||
| 692 
c9231c2903bc
Moved 'int_induct' to IntEx to keep IntEx2 being just theory of integers in order.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
663diff
changeset | 182 | lemma int_induct_raw: | 
| 
c9231c2903bc
Moved 'int_induct' to IntEx to keep IntEx2 being just theory of integers in order.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
663diff
changeset | 183 | assumes a: "P (0::nat, 0)" | 
| 
c9231c2903bc
Moved 'int_induct' to IntEx to keep IntEx2 being just theory of integers in order.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
663diff
changeset | 184 | and b: "\<And>i. P i \<Longrightarrow> P (my_plus i (1,0))" | 
| 
c9231c2903bc
Moved 'int_induct' to IntEx to keep IntEx2 being just theory of integers in order.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
663diff
changeset | 185 | and c: "\<And>i. P i \<Longrightarrow> P (my_plus i (my_neg (1,0)))" | 
| 
c9231c2903bc
Moved 'int_induct' to IntEx to keep IntEx2 being just theory of integers in order.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
663diff
changeset | 186 | shows "P x" | 
| 
c9231c2903bc
Moved 'int_induct' to IntEx to keep IntEx2 being just theory of integers in order.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
663diff
changeset | 187 | apply(case_tac x) apply(simp) | 
| 
c9231c2903bc
Moved 'int_induct' to IntEx to keep IntEx2 being just theory of integers in order.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
663diff
changeset | 188 | apply(rule_tac x="b" in spec) | 
| 
c9231c2903bc
Moved 'int_induct' to IntEx to keep IntEx2 being just theory of integers in order.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
663diff
changeset | 189 | apply(rule_tac Nat.induct) | 
| 
c9231c2903bc
Moved 'int_induct' to IntEx to keep IntEx2 being just theory of integers in order.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
663diff
changeset | 190 | apply(rule allI) | 
| 
c9231c2903bc
Moved 'int_induct' to IntEx to keep IntEx2 being just theory of integers in order.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
663diff
changeset | 191 | apply(rule_tac Nat.induct) | 
| 
c9231c2903bc
Moved 'int_induct' to IntEx to keep IntEx2 being just theory of integers in order.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
663diff
changeset | 192 | using a b c apply(auto) | 
| 
c9231c2903bc
Moved 'int_induct' to IntEx to keep IntEx2 being just theory of integers in order.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
663diff
changeset | 193 | done | 
| 
c9231c2903bc
Moved 'int_induct' to IntEx to keep IntEx2 being just theory of integers in order.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
663diff
changeset | 194 | |
| 
c9231c2903bc
Moved 'int_induct' to IntEx to keep IntEx2 being just theory of integers in order.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
663diff
changeset | 195 | lemma int_induct: | 
| 
c9231c2903bc
Moved 'int_induct' to IntEx to keep IntEx2 being just theory of integers in order.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
663diff
changeset | 196 | assumes a: "P ZERO" | 
| 
c9231c2903bc
Moved 'int_induct' to IntEx to keep IntEx2 being just theory of integers in order.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
663diff
changeset | 197 | and b: "\<And>i. P i \<Longrightarrow> P (PLUS i ONE)" | 
| 
c9231c2903bc
Moved 'int_induct' to IntEx to keep IntEx2 being just theory of integers in order.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
663diff
changeset | 198 | and c: "\<And>i. P i \<Longrightarrow> P (PLUS i (NEG ONE))" | 
| 
c9231c2903bc
Moved 'int_induct' to IntEx to keep IntEx2 being just theory of integers in order.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
663diff
changeset | 199 | shows "P x" | 
| 
c9231c2903bc
Moved 'int_induct' to IntEx to keep IntEx2 being just theory of integers in order.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
663diff
changeset | 200 | using a b c | 
| 
c9231c2903bc
Moved 'int_induct' to IntEx to keep IntEx2 being just theory of integers in order.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
663diff
changeset | 201 | by (lifting int_induct_raw) | 
| 
c9231c2903bc
Moved 'int_induct' to IntEx to keep IntEx2 being just theory of integers in order.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
663diff
changeset | 202 | |
| 597 | 203 | lemma ho_tst: "foldl my_plus x [] = x" | 
| 204 | apply simp | |
| 205 | done | |
| 206 | ||
| 207 | lemma "foldl PLUS x [] = x" | |
| 637 
b029f242d85d
chnaged syntax to "lifting theorem"
 Christian Urban <urbanc@in.tum.de> parents: 
636diff
changeset | 208 | apply(lifting ho_tst) | 
| 597 | 209 | done | 
| 210 | ||
| 211 | lemma ho_tst2: "foldl my_plus x (h # t) \<approx> my_plus h (foldl my_plus x t)" | |
| 212 | sorry | |
| 213 | ||
| 214 | lemma "foldl PLUS x (h # t) = PLUS h (foldl PLUS x t)" | |
| 648 
830b58c2fa94
decoupled QuotProd from QuotMain and also started new cleaning strategy
 Christian Urban <urbanc@in.tum.de> parents: 
645diff
changeset | 215 | apply(lifting_setup ho_tst2) | 
| 
830b58c2fa94
decoupled QuotProd from QuotMain and also started new cleaning strategy
 Christian Urban <urbanc@in.tum.de> parents: 
645diff
changeset | 216 | apply(regularize) | 
| 
830b58c2fa94
decoupled QuotProd from QuotMain and also started new cleaning strategy
 Christian Urban <urbanc@in.tum.de> parents: 
645diff
changeset | 217 | apply(injection) | 
| 
830b58c2fa94
decoupled QuotProd from QuotMain and also started new cleaning strategy
 Christian Urban <urbanc@in.tum.de> parents: 
645diff
changeset | 218 | apply(cleaning) | 
| 597 | 219 | done | 
| 220 | ||
| 221 | lemma ho_tst3: "foldl f (s::nat \<times> nat) ([]::(nat \<times> nat) list) = s" | |
| 222 | by simp | |
| 223 | ||
| 224 | lemma "foldl f (x::my_int) ([]::my_int list) = x" | |
| 648 
830b58c2fa94
decoupled QuotProd from QuotMain and also started new cleaning strategy
 Christian Urban <urbanc@in.tum.de> parents: 
645diff
changeset | 225 | apply(lifting ho_tst3) | 
| 597 | 226 | done | 
| 227 | ||
| 228 | lemma lam_tst: "(\<lambda>x. (x, x)) y = (y, (y :: nat \<times> nat))" | |
| 229 | by simp | |
| 230 | ||
| 622 
df7a2f76daae
Nitpick found a counterexample for one lemma.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
621diff
changeset | 231 | (* Don't know how to keep the goal non-contracted... *) | 
| 597 | 232 | lemma "(\<lambda>x. (x, x)) (y::my_int) = (y, y)" | 
| 648 
830b58c2fa94
decoupled QuotProd from QuotMain and also started new cleaning strategy
 Christian Urban <urbanc@in.tum.de> parents: 
645diff
changeset | 233 | apply(lifting lam_tst) | 
| 597 | 234 | done | 
| 235 | ||
| 236 | lemma lam_tst2: "(\<lambda>(y :: nat \<times> nat). y) = (\<lambda>(x :: nat \<times> nat). x)" | |
| 237 | by simp | |
| 238 | ||
| 605 
120e479ed367
first attempt to deal with Babs in regularise and cleaning (not yet working)
 Christian Urban <urbanc@in.tum.de> parents: 
603diff
changeset | 239 | (* test about lifting identity equations *) | 
| 597 | 240 | |
| 605 
120e479ed367
first attempt to deal with Babs in regularise and cleaning (not yet working)
 Christian Urban <urbanc@in.tum.de> parents: 
603diff
changeset | 241 | ML {*
 | 
| 
120e479ed367
first attempt to deal with Babs in regularise and cleaning (not yet working)
 Christian Urban <urbanc@in.tum.de> parents: 
603diff
changeset | 242 | (* helper code from QuotMain *) | 
| 
120e479ed367
first attempt to deal with Babs in regularise and cleaning (not yet working)
 Christian Urban <urbanc@in.tum.de> parents: 
603diff
changeset | 243 | val pat_ball = @{term "Ball (Respects (R1 ===> R2)) P"}
 | 
| 
120e479ed367
first attempt to deal with Babs in regularise and cleaning (not yet working)
 Christian Urban <urbanc@in.tum.de> parents: 
603diff
changeset | 244 | val pat_bex  = @{term "Bex (Respects (R1 ===> R2)) P"}
 | 
| 
120e479ed367
first attempt to deal with Babs in regularise and cleaning (not yet working)
 Christian Urban <urbanc@in.tum.de> parents: 
603diff
changeset | 245 | val simproc = Simplifier.simproc_i @{theory} "" [pat_ball, pat_bex] (K (ball_bex_range_simproc))
 | 
| 
120e479ed367
first attempt to deal with Babs in regularise and cleaning (not yet working)
 Christian Urban <urbanc@in.tum.de> parents: 
603diff
changeset | 246 | val simpset = (mk_minimal_ss @{context}) 
 | 
| 
120e479ed367
first attempt to deal with Babs in regularise and cleaning (not yet working)
 Christian Urban <urbanc@in.tum.de> parents: 
603diff
changeset | 247 |                        addsimps @{thms ball_reg_eqv bex_reg_eqv babs_reg_eqv}
 | 
| 
120e479ed367
first attempt to deal with Babs in regularise and cleaning (not yet working)
 Christian Urban <urbanc@in.tum.de> parents: 
603diff
changeset | 248 | addsimprocs [simproc] addSolver equiv_solver | 
| 
120e479ed367
first attempt to deal with Babs in regularise and cleaning (not yet working)
 Christian Urban <urbanc@in.tum.de> parents: 
603diff
changeset | 249 | *} | 
| 
120e479ed367
first attempt to deal with Babs in regularise and cleaning (not yet working)
 Christian Urban <urbanc@in.tum.de> parents: 
603diff
changeset | 250 | |
| 
120e479ed367
first attempt to deal with Babs in regularise and cleaning (not yet working)
 Christian Urban <urbanc@in.tum.de> parents: 
603diff
changeset | 251 | (* What regularising does *) | 
| 
120e479ed367
first attempt to deal with Babs in regularise and cleaning (not yet working)
 Christian Urban <urbanc@in.tum.de> parents: 
603diff
changeset | 252 | (*========================*) | 
| 
120e479ed367
first attempt to deal with Babs in regularise and cleaning (not yet working)
 Christian Urban <urbanc@in.tum.de> parents: 
603diff
changeset | 253 | |
| 
120e479ed367
first attempt to deal with Babs in regularise and cleaning (not yet working)
 Christian Urban <urbanc@in.tum.de> parents: 
603diff
changeset | 254 | (* 0. preliminary simplification step *) | 
| 606 
38aa6b67a80b
clarified the function examples
 Christian Urban <urbanc@in.tum.de> parents: 
605diff
changeset | 255 | thm ball_reg_eqv bex_reg_eqv (* of no use: babs_reg_eqv *) | 
| 605 
120e479ed367
first attempt to deal with Babs in regularise and cleaning (not yet working)
 Christian Urban <urbanc@in.tum.de> parents: 
603diff
changeset | 256 | ball_reg_eqv_range bex_reg_eqv_range | 
| 
120e479ed367
first attempt to deal with Babs in regularise and cleaning (not yet working)
 Christian Urban <urbanc@in.tum.de> parents: 
603diff
changeset | 257 | (* 1. first two rtacs *) | 
| 
120e479ed367
first attempt to deal with Babs in regularise and cleaning (not yet working)
 Christian Urban <urbanc@in.tum.de> parents: 
603diff
changeset | 258 | thm ball_reg_right bex_reg_left | 
| 
120e479ed367
first attempt to deal with Babs in regularise and cleaning (not yet working)
 Christian Urban <urbanc@in.tum.de> parents: 
603diff
changeset | 259 | (* 2. monos *) | 
| 
120e479ed367
first attempt to deal with Babs in regularise and cleaning (not yet working)
 Christian Urban <urbanc@in.tum.de> parents: 
603diff
changeset | 260 | (* 3. commutation rules *) | 
| 
120e479ed367
first attempt to deal with Babs in regularise and cleaning (not yet working)
 Christian Urban <urbanc@in.tum.de> parents: 
603diff
changeset | 261 | thm ball_all_comm bex_ex_comm | 
| 
120e479ed367
first attempt to deal with Babs in regularise and cleaning (not yet working)
 Christian Urban <urbanc@in.tum.de> parents: 
603diff
changeset | 262 | (* 4. then rel-equality *) | 
| 
120e479ed367
first attempt to deal with Babs in regularise and cleaning (not yet working)
 Christian Urban <urbanc@in.tum.de> parents: 
603diff
changeset | 263 | thm eq_imp_rel | 
| 
120e479ed367
first attempt to deal with Babs in regularise and cleaning (not yet working)
 Christian Urban <urbanc@in.tum.de> parents: 
603diff
changeset | 264 | (* 5. then simplification like 0 *) | 
| 
120e479ed367
first attempt to deal with Babs in regularise and cleaning (not yet working)
 Christian Urban <urbanc@in.tum.de> parents: 
603diff
changeset | 265 | (* finally jump to 1 *) | 
| 
120e479ed367
first attempt to deal with Babs in regularise and cleaning (not yet working)
 Christian Urban <urbanc@in.tum.de> parents: 
603diff
changeset | 266 | |
| 
120e479ed367
first attempt to deal with Babs in regularise and cleaning (not yet working)
 Christian Urban <urbanc@in.tum.de> parents: 
603diff
changeset | 267 | |
| 
120e479ed367
first attempt to deal with Babs in regularise and cleaning (not yet working)
 Christian Urban <urbanc@in.tum.de> parents: 
603diff
changeset | 268 | (* What cleaning does *) | 
| 
120e479ed367
first attempt to deal with Babs in regularise and cleaning (not yet working)
 Christian Urban <urbanc@in.tum.de> parents: 
603diff
changeset | 269 | (*====================*) | 
| 
120e479ed367
first attempt to deal with Babs in regularise and cleaning (not yet working)
 Christian Urban <urbanc@in.tum.de> parents: 
603diff
changeset | 270 | |
| 
120e479ed367
first attempt to deal with Babs in regularise and cleaning (not yet working)
 Christian Urban <urbanc@in.tum.de> parents: 
603diff
changeset | 271 | (* 1. conversion *) | 
| 
120e479ed367
first attempt to deal with Babs in regularise and cleaning (not yet working)
 Christian Urban <urbanc@in.tum.de> parents: 
603diff
changeset | 272 | thm lambda_prs | 
| 
120e479ed367
first attempt to deal with Babs in regularise and cleaning (not yet working)
 Christian Urban <urbanc@in.tum.de> parents: 
603diff
changeset | 273 | (* 2. simplification with *) | 
| 
120e479ed367
first attempt to deal with Babs in regularise and cleaning (not yet working)
 Christian Urban <urbanc@in.tum.de> parents: 
603diff
changeset | 274 | thm all_prs ex_prs | 
| 
120e479ed367
first attempt to deal with Babs in regularise and cleaning (not yet working)
 Christian Urban <urbanc@in.tum.de> parents: 
603diff
changeset | 275 | (* 3. Simplification with *) | 
| 
120e479ed367
first attempt to deal with Babs in regularise and cleaning (not yet working)
 Christian Urban <urbanc@in.tum.de> parents: 
603diff
changeset | 276 | thm fun_map.simps id_simps Quotient_abs_rep Quotient_rel_rep | 
| 
120e479ed367
first attempt to deal with Babs in regularise and cleaning (not yet working)
 Christian Urban <urbanc@in.tum.de> parents: 
603diff
changeset | 277 | (* 4. Test for refl *) | 
| 
120e479ed367
first attempt to deal with Babs in regularise and cleaning (not yet working)
 Christian Urban <urbanc@in.tum.de> parents: 
603diff
changeset | 278 | |
| 
120e479ed367
first attempt to deal with Babs in regularise and cleaning (not yet working)
 Christian Urban <urbanc@in.tum.de> parents: 
603diff
changeset | 279 | lemma | 
| 
120e479ed367
first attempt to deal with Babs in regularise and cleaning (not yet working)
 Christian Urban <urbanc@in.tum.de> parents: 
603diff
changeset | 280 | shows "equivp (op \<approx>)" | 
| 622 
df7a2f76daae
Nitpick found a counterexample for one lemma.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
621diff
changeset | 281 | and "equivp ((op \<approx>) ===> (op \<approx>))" | 
| 
df7a2f76daae
Nitpick found a counterexample for one lemma.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
621diff
changeset | 282 | (* Nitpick finds a counterexample! *) | 
| 605 
120e479ed367
first attempt to deal with Babs in regularise and cleaning (not yet working)
 Christian Urban <urbanc@in.tum.de> parents: 
603diff
changeset | 283 | oops | 
| 
120e479ed367
first attempt to deal with Babs in regularise and cleaning (not yet working)
 Christian Urban <urbanc@in.tum.de> parents: 
603diff
changeset | 284 | |
| 606 
38aa6b67a80b
clarified the function examples
 Christian Urban <urbanc@in.tum.de> parents: 
605diff
changeset | 285 | lemma lam_tst3a: "(\<lambda>(y :: nat \<times> nat). y) = (\<lambda>(x :: nat \<times> nat). x)" | 
| 597 | 286 | by auto | 
| 287 | ||
| 620 | 288 | lemma id_rsp: | 
| 289 | shows "(R ===> R) id id" | |
| 290 | by simp | |
| 291 | ||
| 292 | lemma lam_tst3a_reg: "(op \<approx> ===> op \<approx>) (Babs (Respects op \<approx>) (\<lambda>y. y)) (Babs (Respects op \<approx>) (\<lambda>x. x))" | |
| 293 | apply (rule babs_rsp[OF Quotient_my_int]) | |
| 294 | apply (simp add: id_rsp) | |
| 295 | done | |
| 296 | ||
| 606 
38aa6b67a80b
clarified the function examples
 Christian Urban <urbanc@in.tum.de> parents: 
605diff
changeset | 297 | lemma "(\<lambda>(y :: my_int). y) = (\<lambda>(x :: my_int). x)" | 
| 
38aa6b67a80b
clarified the function examples
 Christian Urban <urbanc@in.tum.de> parents: 
605diff
changeset | 298 | apply(tactic {* procedure_tac @{context} @{thm lam_tst3a} 1 *})
 | 
| 621 
c10a46fa0de9
Added a 'rep_abs' in inj_repabs_trm of babs; and proved two lam examples.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
620diff
changeset | 299 | apply(rule impI) | 
| 
c10a46fa0de9
Added a 'rep_abs' in inj_repabs_trm of babs; and proved two lam examples.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
620diff
changeset | 300 | apply(rule lam_tst3a_reg) | 
| 610 
2bee5ca44ef5
removed "global" data and lookup functions; had to move a tactic out from the inj_repabs_match tactic since apply_rsp interferes with a trans2 rule for ===>
 Christian Urban <urbanc@in.tum.de> parents: 
606diff
changeset | 301 | apply(tactic {* all_inj_repabs_tac @{context} 1*})
 | 
| 602 
e56eeb9fedb3
make_inst for lambda_prs where the second quotient is not identity.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
600diff
changeset | 302 | apply(tactic {* clean_tac  @{context} 1 *})
 | 
| 621 
c10a46fa0de9
Added a 'rep_abs' in inj_repabs_trm of babs; and proved two lam examples.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
620diff
changeset | 303 | done | 
| 606 
38aa6b67a80b
clarified the function examples
 Christian Urban <urbanc@in.tum.de> parents: 
605diff
changeset | 304 | |
| 
38aa6b67a80b
clarified the function examples
 Christian Urban <urbanc@in.tum.de> parents: 
605diff
changeset | 305 | lemma lam_tst3b: "(\<lambda>(y :: nat \<times> nat \<Rightarrow> nat \<times> nat). y) = (\<lambda>(x :: nat \<times> nat \<Rightarrow> nat \<times> nat). x)" | 
| 
38aa6b67a80b
clarified the function examples
 Christian Urban <urbanc@in.tum.de> parents: 
605diff
changeset | 306 | by auto | 
| 
38aa6b67a80b
clarified the function examples
 Christian Urban <urbanc@in.tum.de> parents: 
605diff
changeset | 307 | |
| 
38aa6b67a80b
clarified the function examples
 Christian Urban <urbanc@in.tum.de> parents: 
605diff
changeset | 308 | lemma "(\<lambda>(y :: my_int => my_int). y) = (\<lambda>(x :: my_int => my_int). x)" | 
| 648 
830b58c2fa94
decoupled QuotProd from QuotMain and also started new cleaning strategy
 Christian Urban <urbanc@in.tum.de> parents: 
645diff
changeset | 309 | apply(lifting lam_tst3b) | 
| 621 
c10a46fa0de9
Added a 'rep_abs' in inj_repabs_trm of babs; and proved two lam examples.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
620diff
changeset | 310 | apply(rule impI) | 
| 648 
830b58c2fa94
decoupled QuotProd from QuotMain and also started new cleaning strategy
 Christian Urban <urbanc@in.tum.de> parents: 
645diff
changeset | 311 | apply(rule babs_rsp[OF fun_quotient[OF Quotient_my_int Quotient_my_int]]) | 
| 
830b58c2fa94
decoupled QuotProd from QuotMain and also started new cleaning strategy
 Christian Urban <urbanc@in.tum.de> parents: 
645diff
changeset | 312 | apply(simp add: id_rsp) | 
| 625 
5d6a2b5fb222
Another lambda example theorem proved. Seems it starts working properly.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
622diff
changeset | 313 | done | 
| 622 
df7a2f76daae
Nitpick found a counterexample for one lemma.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
621diff
changeset | 314 | |
| 625 
5d6a2b5fb222
Another lambda example theorem proved. Seems it starts working properly.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
622diff
changeset | 315 | term map | 
| 
5d6a2b5fb222
Another lambda example theorem proved. Seems it starts working properly.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
622diff
changeset | 316 | |
| 
5d6a2b5fb222
Another lambda example theorem proved. Seems it starts working properly.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
622diff
changeset | 317 | lemma lam_tst4: "map (\<lambda>x. my_plus x (0,0)) l = l" | 
| 
5d6a2b5fb222
Another lambda example theorem proved. Seems it starts working properly.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
622diff
changeset | 318 | apply (induct l) | 
| 
5d6a2b5fb222
Another lambda example theorem proved. Seems it starts working properly.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
622diff
changeset | 319 | apply simp | 
| 
5d6a2b5fb222
Another lambda example theorem proved. Seems it starts working properly.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
622diff
changeset | 320 | apply (case_tac a) | 
| 
5d6a2b5fb222
Another lambda example theorem proved. Seems it starts working properly.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
622diff
changeset | 321 | apply simp | 
| 
5d6a2b5fb222
Another lambda example theorem proved. Seems it starts working properly.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
622diff
changeset | 322 | done | 
| 
5d6a2b5fb222
Another lambda example theorem proved. Seems it starts working properly.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
622diff
changeset | 323 | |
| 
5d6a2b5fb222
Another lambda example theorem proved. Seems it starts working properly.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
622diff
changeset | 324 | lemma "map (\<lambda>x. PLUS x ZERO) l = l" | 
| 637 
b029f242d85d
chnaged syntax to "lifting theorem"
 Christian Urban <urbanc@in.tum.de> parents: 
636diff
changeset | 325 | apply(lifting lam_tst4) | 
| 656 
c86a47d4966e
Temporarily repeated fun_map_tac 4 times. Cleaning for all examples work.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
652diff
changeset | 326 | done | 
| 601 
81f40b8bde7b
added "end" to each example theory
 Christian Urban <urbanc@in.tum.de> parents: 
600diff
changeset | 327 | |
| 617 
ca37f4b6457c
An example of working cleaning for lambda lifting. Still not sure why Babs helps.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
606diff
changeset | 328 | end |