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>
diff
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
787
5cf83fa5b36c
made the quotient_type definition more like typedef; now type variables need to be explicitly given
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 10
quotient_type
5cf83fa5b36c
made the quotient_type definition more like typedef; now type variables need to be explicitly given
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 11
my_int = "nat \<times> nat" / intrel
597
+ − 12
apply(unfold equivp_def)
+ − 13
apply(auto simp add: mem_def expand_fun_eq)
+ − 14
done
+ − 15
636
+ − 16
thm quot_equiv
597
+ − 17
636
+ − 18
thm quot_thm
597
+ − 19
+ − 20
thm my_int_equivp
+ − 21
+ − 22
print_theorems
+ − 23
print_quotients
+ − 24
767
37285ec4387d
on suggestion of Tobias renamed "quotient_def" to "quotient_definition"; needs new keyword file
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 25
quotient_definition
705
+ − 26
"ZERO :: my_int"
+ − 27
as
663
0dd10a900cae
Different syntax for definitions that allows overloading and retrieving of definitions by matching whole constants.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 28
"(0::nat, 0::nat)"
597
+ − 29
767
37285ec4387d
on suggestion of Tobias renamed "quotient_def" to "quotient_definition"; needs new keyword file
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 30
quotient_definition
705
+ − 31
"ONE :: my_int"
+ − 32
as
663
0dd10a900cae
Different syntax for definitions that allows overloading and retrieving of definitions by matching whole constants.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 33
"(1::nat, 0::nat)"
597
+ − 34
+ − 35
fun
+ − 36
my_plus :: "(nat \<times> nat) \<Rightarrow> (nat \<times> nat) \<Rightarrow> (nat \<times> nat)"
+ − 37
where
+ − 38
"my_plus (x, y) (u, v) = (x + u, y + v)"
+ − 39
767
37285ec4387d
on suggestion of Tobias renamed "quotient_def" to "quotient_definition"; needs new keyword file
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 40
quotient_definition
705
+ − 41
"PLUS :: my_int \<Rightarrow> my_int \<Rightarrow> my_int"
+ − 42
as
663
0dd10a900cae
Different syntax for definitions that allows overloading and retrieving of definitions by matching whole constants.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 43
"my_plus"
597
+ − 44
+ − 45
fun
+ − 46
my_neg :: "(nat \<times> nat) \<Rightarrow> (nat \<times> nat)"
+ − 47
where
+ − 48
"my_neg (x, y) = (y, x)"
+ − 49
767
37285ec4387d
on suggestion of Tobias renamed "quotient_def" to "quotient_definition"; needs new keyword file
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 50
quotient_definition
705
+ − 51
"NEG :: my_int \<Rightarrow> my_int"
+ − 52
as
663
0dd10a900cae
Different syntax for definitions that allows overloading and retrieving of definitions by matching whole constants.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 53
"my_neg"
597
+ − 54
+ − 55
definition
+ − 56
MINUS :: "my_int \<Rightarrow> my_int \<Rightarrow> my_int"
+ − 57
where
+ − 58
"MINUS z w = PLUS z (NEG w)"
+ − 59
+ − 60
fun
+ − 61
my_mult :: "(nat \<times> nat) \<Rightarrow> (nat \<times> nat) \<Rightarrow> (nat \<times> nat)"
+ − 62
where
+ − 63
"my_mult (x, y) (u, v) = (x*u + y*v, x*v + y*u)"
+ − 64
767
37285ec4387d
on suggestion of Tobias renamed "quotient_def" to "quotient_definition"; needs new keyword file
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 65
quotient_definition
705
+ − 66
"MULT :: my_int \<Rightarrow> my_int \<Rightarrow> my_int"
+ − 67
as
663
0dd10a900cae
Different syntax for definitions that allows overloading and retrieving of definitions by matching whole constants.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 68
"my_mult"
597
+ − 69
+ − 70
+ − 71
(* NOT SURE WETHER THIS DEFINITION IS CORRECT *)
+ − 72
fun
+ − 73
my_le :: "(nat \<times> nat) \<Rightarrow> (nat \<times> nat) \<Rightarrow> bool"
+ − 74
where
+ − 75
"my_le (x, y) (u, v) = (x+v \<le> u+y)"
+ − 76
767
37285ec4387d
on suggestion of Tobias renamed "quotient_def" to "quotient_definition"; needs new keyword file
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 77
quotient_definition
705
+ − 78
"LE :: my_int \<Rightarrow> my_int \<Rightarrow> bool"
+ − 79
as
663
0dd10a900cae
Different syntax for definitions that allows overloading and retrieving of definitions by matching whole constants.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 80
"my_le"
597
+ − 81
+ − 82
term LE
+ − 83
thm LE_def
+ − 84
+ − 85
+ − 86
definition
+ − 87
LESS :: "my_int \<Rightarrow> my_int \<Rightarrow> bool"
+ − 88
where
+ − 89
"LESS z w = (LE z w \<and> z \<noteq> w)"
+ − 90
+ − 91
term LESS
+ − 92
thm LESS_def
+ − 93
+ − 94
definition
+ − 95
ABS :: "my_int \<Rightarrow> my_int"
+ − 96
where
+ − 97
"ABS i = (if (LESS i ZERO) then (NEG i) else i)"
+ − 98
+ − 99
definition
+ − 100
SIGN :: "my_int \<Rightarrow> my_int"
+ − 101
where
+ − 102
"SIGN i = (if i = ZERO then ZERO else if (LESS ZERO i) then ONE else (NEG ONE))"
+ − 103
764
a603aa6c9d01
with "isabelle make keywords" you can create automatically a "quot" keywordfile, provided all Logics are in place
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 104
print_quotconsts
597
+ − 105
+ − 106
lemma plus_sym_pre:
+ − 107
shows "my_plus a b \<approx> my_plus b a"
+ − 108
apply(cases a)
+ − 109
apply(cases b)
+ − 110
apply(auto)
+ − 111
done
+ − 112
636
+ − 113
lemma plus_rsp[quot_respect]:
597
+ − 114
shows "(intrel ===> intrel ===> intrel) my_plus my_plus"
+ − 115
by (simp)
+ − 116
692
c9231c2903bc
Moved 'int_induct' to IntEx to keep IntEx2 being just theory of integers in order.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 117
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>
diff
changeset
+ − 118
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>
diff
changeset
+ − 119
by simp
c9231c2903bc
Moved 'int_induct' to IntEx to keep IntEx2 being just theory of integers in order.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 120
597
+ − 121
lemma test1: "my_plus a b = my_plus a b"
+ − 122
apply(rule refl)
+ − 123
done
+ − 124
+ − 125
lemma "PLUS a b = PLUS a b"
637
+ − 126
apply(lifting_setup test1)
632
+ − 127
apply(regularize)
+ − 128
apply(injection)
+ − 129
apply(cleaning)
597
+ − 130
done
+ − 131
+ − 132
thm lambda_prs
+ − 133
+ − 134
lemma test2: "my_plus a = my_plus a"
+ − 135
apply(rule refl)
+ − 136
done
+ − 137
+ − 138
lemma "PLUS a = PLUS a"
758
+ − 139
apply(lifting_setup test2)
606
+ − 140
apply(rule impI)
597
+ − 141
apply(rule ballI)
+ − 142
apply(rule apply_rsp[OF Quotient_my_int plus_rsp])
+ − 143
apply(simp only: in_respects)
758
+ − 144
apply(injection)
+ − 145
apply(cleaning)
597
+ − 146
done
+ − 147
+ − 148
lemma test3: "my_plus = my_plus"
+ − 149
apply(rule refl)
+ − 150
done
+ − 151
+ − 152
lemma "PLUS = PLUS"
758
+ − 153
apply(lifting_setup test3)
606
+ − 154
apply(rule impI)
597
+ − 155
apply(rule plus_rsp)
636
+ − 156
apply(injection)
+ − 157
apply(cleaning)
597
+ − 158
done
+ − 159
+ − 160
+ − 161
lemma "PLUS a b = PLUS b a"
758
+ − 162
apply(lifting plus_sym_pre)
597
+ − 163
done
+ − 164
+ − 165
lemma plus_assoc_pre:
+ − 166
shows "my_plus (my_plus i j) k \<approx> my_plus i (my_plus j k)"
+ − 167
apply (cases i)
+ − 168
apply (cases j)
+ − 169
apply (cases k)
+ − 170
apply (simp)
+ − 171
done
+ − 172
+ − 173
lemma plus_assoc: "PLUS (PLUS x xa) xb = PLUS x (PLUS xa xb)"
758
+ − 174
apply(lifting plus_assoc_pre)
597
+ − 175
done
+ − 176
692
c9231c2903bc
Moved 'int_induct' to IntEx to keep IntEx2 being just theory of integers in order.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 177
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>
diff
changeset
+ − 178
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>
diff
changeset
+ − 179
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>
diff
changeset
+ − 180
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>
diff
changeset
+ − 181
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>
diff
changeset
+ − 182
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>
diff
changeset
+ − 183
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>
diff
changeset
+ − 184
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>
diff
changeset
+ − 185
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>
diff
changeset
+ − 186
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>
diff
changeset
+ − 187
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>
diff
changeset
+ − 188
done
c9231c2903bc
Moved 'int_induct' to IntEx to keep IntEx2 being just theory of integers in order.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 189
c9231c2903bc
Moved 'int_induct' to IntEx to keep IntEx2 being just theory of integers in order.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 190
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>
diff
changeset
+ − 191
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>
diff
changeset
+ − 192
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>
diff
changeset
+ − 193
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>
diff
changeset
+ − 194
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>
diff
changeset
+ − 195
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>
diff
changeset
+ − 196
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>
diff
changeset
+ − 197
597
+ − 198
lemma ho_tst: "foldl my_plus x [] = x"
+ − 199
apply simp
+ − 200
done
+ − 201
+ − 202
lemma "foldl PLUS x [] = x"
637
+ − 203
apply(lifting ho_tst)
597
+ − 204
done
+ − 205
+ − 206
lemma ho_tst2: "foldl my_plus x (h # t) \<approx> my_plus h (foldl my_plus x t)"
+ − 207
sorry
+ − 208
+ − 209
lemma "foldl PLUS x (h # t) = PLUS h (foldl PLUS x t)"
758
+ − 210
apply(lifting ho_tst2)
597
+ − 211
done
+ − 212
+ − 213
lemma ho_tst3: "foldl f (s::nat \<times> nat) ([]::(nat \<times> nat) list) = s"
+ − 214
by simp
+ − 215
+ − 216
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>
diff
changeset
+ − 217
apply(lifting ho_tst3)
597
+ − 218
done
+ − 219
+ − 220
lemma lam_tst: "(\<lambda>x. (x, x)) y = (y, (y :: nat \<times> nat))"
+ − 221
by simp
+ − 222
622
+ − 223
(* Don't know how to keep the goal non-contracted... *)
597
+ − 224
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>
diff
changeset
+ − 225
apply(lifting lam_tst)
597
+ − 226
done
+ − 227
+ − 228
lemma lam_tst2: "(\<lambda>(y :: nat \<times> nat). y) = (\<lambda>(x :: nat \<times> nat). x)"
+ − 229
by simp
+ − 230
605
120e479ed367
first attempt to deal with Babs in regularise and cleaning (not yet working)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 231
lemma
120e479ed367
first attempt to deal with Babs in regularise and cleaning (not yet working)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 232
shows "equivp (op \<approx>)"
622
+ − 233
and "equivp ((op \<approx>) ===> (op \<approx>))"
+ − 234
(* 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>
diff
changeset
+ − 235
oops
120e479ed367
first attempt to deal with Babs in regularise and cleaning (not yet working)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 236
606
+ − 237
lemma lam_tst3a: "(\<lambda>(y :: nat \<times> nat). y) = (\<lambda>(x :: nat \<times> nat). x)"
597
+ − 238
by auto
+ − 239
620
+ − 240
lemma id_rsp:
+ − 241
shows "(R ===> R) id id"
+ − 242
by simp
+ − 243
+ − 244
lemma lam_tst3a_reg: "(op \<approx> ===> op \<approx>) (Babs (Respects op \<approx>) (\<lambda>y. y)) (Babs (Respects op \<approx>) (\<lambda>x. x))"
+ − 245
apply (rule babs_rsp[OF Quotient_my_int])
+ − 246
apply (simp add: id_rsp)
+ − 247
done
+ − 248
606
+ − 249
lemma "(\<lambda>(y :: my_int). y) = (\<lambda>(x :: my_int). x)"
758
+ − 250
apply(lifting lam_tst3a)
621
c10a46fa0de9
Added a 'rep_abs' in inj_repabs_trm of babs; and proved two lam examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 251
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>
diff
changeset
+ − 252
apply(rule lam_tst3a_reg)
768
e9e205b904e2
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 253
apply(injection)
e9e205b904e2
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 254
sorry
e9e205b904e2
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 255
(* PROBLEM
e9e205b904e2
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 256
done*)
606
+ − 257
+ − 258
lemma lam_tst3b: "(\<lambda>(y :: nat \<times> nat \<Rightarrow> nat \<times> nat). y) = (\<lambda>(x :: nat \<times> nat \<Rightarrow> nat \<times> nat). x)"
+ − 259
by auto
+ − 260
+ − 261
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>
diff
changeset
+ − 262
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>
diff
changeset
+ − 263
apply(rule impI)
648
830b58c2fa94
decoupled QuotProd from QuotMain and also started new cleaning strategy
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 264
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>
diff
changeset
+ − 265
apply(simp add: id_rsp)
768
e9e205b904e2
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 266
apply(injection)
e9e205b904e2
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 267
sorry
e9e205b904e2
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 268
(* PROBLEM
e9e205b904e2
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 269
done*)
622
+ − 270
625
5d6a2b5fb222
Another lambda example theorem proved. Seems it starts working properly.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 271
term map
5d6a2b5fb222
Another lambda example theorem proved. Seems it starts working properly.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 272
5d6a2b5fb222
Another lambda example theorem proved. Seems it starts working properly.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 273
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>
diff
changeset
+ − 274
apply (induct l)
5d6a2b5fb222
Another lambda example theorem proved. Seems it starts working properly.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 275
apply simp
5d6a2b5fb222
Another lambda example theorem proved. Seems it starts working properly.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 276
apply (case_tac a)
5d6a2b5fb222
Another lambda example theorem proved. Seems it starts working properly.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 277
apply simp
5d6a2b5fb222
Another lambda example theorem proved. Seems it starts working properly.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 278
done
5d6a2b5fb222
Another lambda example theorem proved. Seems it starts working properly.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 279
5d6a2b5fb222
Another lambda example theorem proved. Seems it starts working properly.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 280
lemma "map (\<lambda>x. PLUS x ZERO) l = l"
637
+ − 281
apply(lifting lam_tst4)
656
c86a47d4966e
Temporarily repeated fun_map_tac 4 times. Cleaning for all examples work.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 282
done
601
+ − 283
617
ca37f4b6457c
An example of working cleaning for lambda lifting. Still not sure why Babs helps.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 284
end