597
+ − 1
theory IntEx
622
+ − 2
imports "../QuotList" Nitpick
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
+ − 15
thm quot_equiv
597
+ − 16
636
+ − 17
thm quot_thm
597
+ − 18
+ − 19
thm my_int_equivp
+ − 20
+ − 21
print_theorems
+ − 22
print_quotients
+ − 23
+ − 24
quotient_def
+ − 25
ZERO::"my_int"
+ − 26
where
+ − 27
"ZERO \<equiv> (0::nat, 0::nat)"
+ − 28
+ − 29
ML {* print_qconstinfo @{context} *}
+ − 30
+ − 31
term ZERO
+ − 32
thm ZERO_def
+ − 33
+ − 34
ML {* prop_of @{thm ZERO_def} *}
+ − 35
+ − 36
ML {* separate *}
+ − 37
+ − 38
quotient_def
+ − 39
ONE::"my_int"
+ − 40
where
+ − 41
"ONE \<equiv> (1::nat, 0::nat)"
+ − 42
+ − 43
ML {* print_qconstinfo @{context} *}
+ − 44
+ − 45
term ONE
+ − 46
thm ONE_def
+ − 47
+ − 48
fun
+ − 49
my_plus :: "(nat \<times> nat) \<Rightarrow> (nat \<times> nat) \<Rightarrow> (nat \<times> nat)"
+ − 50
where
+ − 51
"my_plus (x, y) (u, v) = (x + u, y + v)"
+ − 52
+ − 53
quotient_def
+ − 54
PLUS::"my_int \<Rightarrow> my_int \<Rightarrow> my_int"
+ − 55
where
+ − 56
"PLUS \<equiv> my_plus"
+ − 57
+ − 58
term my_plus
+ − 59
term PLUS
+ − 60
thm PLUS_def
+ − 61
+ − 62
fun
+ − 63
my_neg :: "(nat \<times> nat) \<Rightarrow> (nat \<times> nat)"
+ − 64
where
+ − 65
"my_neg (x, y) = (y, x)"
+ − 66
+ − 67
quotient_def
+ − 68
NEG::"my_int \<Rightarrow> my_int"
+ − 69
where
+ − 70
"NEG \<equiv> my_neg"
+ − 71
+ − 72
term NEG
+ − 73
thm NEG_def
+ − 74
+ − 75
definition
+ − 76
MINUS :: "my_int \<Rightarrow> my_int \<Rightarrow> my_int"
+ − 77
where
+ − 78
"MINUS z w = PLUS z (NEG w)"
+ − 79
+ − 80
fun
+ − 81
my_mult :: "(nat \<times> nat) \<Rightarrow> (nat \<times> nat) \<Rightarrow> (nat \<times> nat)"
+ − 82
where
+ − 83
"my_mult (x, y) (u, v) = (x*u + y*v, x*v + y*u)"
+ − 84
+ − 85
quotient_def
+ − 86
MULT::"my_int \<Rightarrow> my_int \<Rightarrow> my_int"
+ − 87
where
+ − 88
"MULT \<equiv> my_mult"
+ − 89
+ − 90
term MULT
+ − 91
thm MULT_def
+ − 92
+ − 93
(* NOT SURE WETHER THIS DEFINITION IS CORRECT *)
+ − 94
fun
+ − 95
my_le :: "(nat \<times> nat) \<Rightarrow> (nat \<times> nat) \<Rightarrow> bool"
+ − 96
where
+ − 97
"my_le (x, y) (u, v) = (x+v \<le> u+y)"
+ − 98
+ − 99
quotient_def
+ − 100
LE :: "my_int \<Rightarrow> my_int \<Rightarrow> bool"
+ − 101
where
+ − 102
"LE \<equiv> my_le"
+ − 103
+ − 104
term LE
+ − 105
thm LE_def
+ − 106
+ − 107
+ − 108
definition
+ − 109
LESS :: "my_int \<Rightarrow> my_int \<Rightarrow> bool"
+ − 110
where
+ − 111
"LESS z w = (LE z w \<and> z \<noteq> w)"
+ − 112
+ − 113
term LESS
+ − 114
thm LESS_def
+ − 115
+ − 116
definition
+ − 117
ABS :: "my_int \<Rightarrow> my_int"
+ − 118
where
+ − 119
"ABS i = (if (LESS i ZERO) then (NEG i) else i)"
+ − 120
+ − 121
definition
+ − 122
SIGN :: "my_int \<Rightarrow> my_int"
+ − 123
where
+ − 124
"SIGN i = (if i = ZERO then ZERO else if (LESS ZERO i) then ONE else (NEG ONE))"
+ − 125
+ − 126
ML {* print_qconstinfo @{context} *}
+ − 127
+ − 128
lemma plus_sym_pre:
+ − 129
shows "my_plus a b \<approx> my_plus b a"
+ − 130
apply(cases a)
+ − 131
apply(cases b)
+ − 132
apply(auto)
+ − 133
done
+ − 134
636
+ − 135
lemma plus_rsp[quot_respect]:
597
+ − 136
shows "(intrel ===> intrel ===> intrel) my_plus my_plus"
+ − 137
by (simp)
+ − 138
+ − 139
lemma test1: "my_plus a b = my_plus a b"
+ − 140
apply(rule refl)
+ − 141
done
+ − 142
+ − 143
lemma "PLUS a b = PLUS a b"
637
+ − 144
apply(lifting_setup test1)
632
+ − 145
apply(regularize)
+ − 146
apply(injection)
+ − 147
apply(cleaning)
597
+ − 148
done
+ − 149
+ − 150
thm lambda_prs
+ − 151
+ − 152
lemma test2: "my_plus a = my_plus a"
+ − 153
apply(rule refl)
+ − 154
done
+ − 155
+ − 156
lemma "PLUS a = PLUS a"
+ − 157
apply(tactic {* procedure_tac @{context} @{thm test2} 1 *})
606
+ − 158
apply(rule impI)
597
+ − 159
apply(rule ballI)
+ − 160
apply(rule apply_rsp[OF Quotient_my_int plus_rsp])
+ − 161
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>
diff
changeset
+ − 162
apply(tactic {* all_inj_repabs_tac @{context} 1*})
597
+ − 163
apply(tactic {* clean_tac @{context} 1 *})
+ − 164
done
+ − 165
+ − 166
lemma test3: "my_plus = my_plus"
+ − 167
apply(rule refl)
+ − 168
done
+ − 169
+ − 170
lemma "PLUS = PLUS"
+ − 171
apply(tactic {* procedure_tac @{context} @{thm test3} 1 *})
606
+ − 172
apply(rule impI)
597
+ − 173
apply(rule plus_rsp)
636
+ − 174
apply(injection)
+ − 175
apply(cleaning)
597
+ − 176
done
+ − 177
+ − 178
+ − 179
lemma "PLUS a b = PLUS b a"
+ − 180
apply(tactic {* procedure_tac @{context} @{thm plus_sym_pre} 1 *})
+ − 181
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>
diff
changeset
+ − 182
apply(tactic {* all_inj_repabs_tac @{context} 1*})
597
+ − 183
apply(tactic {* clean_tac @{context} 1 *})
+ − 184
done
+ − 185
+ − 186
lemma plus_assoc_pre:
+ − 187
shows "my_plus (my_plus i j) k \<approx> my_plus i (my_plus j k)"
+ − 188
apply (cases i)
+ − 189
apply (cases j)
+ − 190
apply (cases k)
+ − 191
apply (simp)
+ − 192
done
+ − 193
+ − 194
lemma plus_assoc: "PLUS (PLUS x xa) xb = PLUS x (PLUS xa xb)"
+ − 195
apply(tactic {* procedure_tac @{context} @{thm plus_assoc_pre} 1 *})
+ − 196
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>
diff
changeset
+ − 197
apply(tactic {* all_inj_repabs_tac @{context} 1*})
597
+ − 198
apply(tactic {* clean_tac @{context} 1 *})
+ − 199
done
+ − 200
+ − 201
lemma ho_tst: "foldl my_plus x [] = x"
+ − 202
apply simp
+ − 203
done
+ − 204
+ − 205
lemma "foldl PLUS x [] = x"
637
+ − 206
apply(lifting ho_tst)
597
+ − 207
apply(tactic {* clean_tac @{context} 1 *})
+ − 208
apply(simp only: foldl_prs[OF Quotient_my_int Quotient_my_int] nil_prs[OF Quotient_my_int])
+ − 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)"
+ − 215
apply(tactic {* procedure_tac @{context} @{thm ho_tst2} 1 *})
+ − 216
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>
diff
changeset
+ − 217
apply(tactic {* all_inj_repabs_tac @{context} 1*})
597
+ − 218
apply(tactic {* clean_tac @{context} 1 *})
+ − 219
apply(simp only: foldl_prs[OF Quotient_my_int Quotient_my_int] cons_prs[OF Quotient_my_int])
+ − 220
done
+ − 221
+ − 222
lemma ho_tst3: "foldl f (s::nat \<times> nat) ([]::(nat \<times> nat) list) = s"
+ − 223
by simp
+ − 224
+ − 225
lemma "foldl f (x::my_int) ([]::my_int list) = x"
+ − 226
apply(tactic {* procedure_tac @{context} @{thm ho_tst3} 1 *})
+ − 227
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>
diff
changeset
+ − 228
apply(tactic {* all_inj_repabs_tac @{context} 1*})
597
+ − 229
(* TODO: does not work when this is added *)
+ − 230
(* apply(tactic {* lambda_prs_tac @{context} 1 *})*)
+ − 231
apply(tactic {* clean_tac @{context} 1 *})
+ − 232
apply(simp only: foldl_prs[OF Quotient_my_int Quotient_my_int] nil_prs[OF Quotient_my_int])
+ − 233
done
+ − 234
+ − 235
lemma lam_tst: "(\<lambda>x. (x, x)) y = (y, (y :: nat \<times> nat))"
+ − 236
by simp
+ − 237
622
+ − 238
(* Don't know how to keep the goal non-contracted... *)
597
+ − 239
lemma "(\<lambda>x. (x, x)) (y::my_int) = (y, y)"
+ − 240
apply(tactic {* procedure_tac @{context} @{thm lam_tst} 1 *})
+ − 241
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>
diff
changeset
+ − 242
apply(tactic {* all_inj_repabs_tac @{context} 1*})
597
+ − 243
apply(tactic {* clean_tac @{context} 1 *})
+ − 244
apply(simp add: pair_prs)
+ − 245
done
+ − 246
+ − 247
lemma lam_tst2: "(\<lambda>(y :: nat \<times> nat). y) = (\<lambda>(x :: nat \<times> nat). x)"
+ − 248
by simp
+ − 249
605
120e479ed367
first attempt to deal with Babs in regularise and cleaning (not yet working)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 250
(* test about lifting identity equations *)
597
+ − 251
605
120e479ed367
first attempt to deal with Babs in regularise and cleaning (not yet working)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 252
ML {*
120e479ed367
first attempt to deal with Babs in regularise and cleaning (not yet working)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 253
(* helper code from QuotMain *)
120e479ed367
first attempt to deal with Babs in regularise and cleaning (not yet working)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 254
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>
diff
changeset
+ − 255
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>
diff
changeset
+ − 256
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>
diff
changeset
+ − 257
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>
diff
changeset
+ − 258
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>
diff
changeset
+ − 259
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>
diff
changeset
+ − 260
*}
120e479ed367
first attempt to deal with Babs in regularise and cleaning (not yet working)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 261
120e479ed367
first attempt to deal with Babs in regularise and cleaning (not yet working)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 262
(* What regularising does *)
120e479ed367
first attempt to deal with Babs in regularise and cleaning (not yet working)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 263
(*========================*)
120e479ed367
first attempt to deal with Babs in regularise and cleaning (not yet working)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 264
120e479ed367
first attempt to deal with Babs in regularise and cleaning (not yet working)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 265
(* 0. preliminary simplification step *)
606
+ − 266
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>
diff
changeset
+ − 267
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>
diff
changeset
+ − 268
(* 1. first two rtacs *)
120e479ed367
first attempt to deal with Babs in regularise and cleaning (not yet working)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 269
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>
diff
changeset
+ − 270
(* 2. monos *)
120e479ed367
first attempt to deal with Babs in regularise and cleaning (not yet working)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 271
(* 3. commutation rules *)
120e479ed367
first attempt to deal with Babs in regularise and cleaning (not yet working)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 272
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>
diff
changeset
+ − 273
(* 4. then rel-equality *)
120e479ed367
first attempt to deal with Babs in regularise and cleaning (not yet working)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 274
thm eq_imp_rel
120e479ed367
first attempt to deal with Babs in regularise and cleaning (not yet working)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 275
(* 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>
diff
changeset
+ − 276
(* finally jump to 1 *)
120e479ed367
first attempt to deal with Babs in regularise and cleaning (not yet working)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 277
120e479ed367
first attempt to deal with Babs in regularise and cleaning (not yet working)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 278
120e479ed367
first attempt to deal with Babs in regularise and cleaning (not yet working)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 279
(* What cleaning does *)
120e479ed367
first attempt to deal with Babs in regularise and cleaning (not yet working)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 280
(*====================*)
120e479ed367
first attempt to deal with Babs in regularise and cleaning (not yet working)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 281
120e479ed367
first attempt to deal with Babs in regularise and cleaning (not yet working)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 282
(* 1. conversion *)
120e479ed367
first attempt to deal with Babs in regularise and cleaning (not yet working)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 283
thm lambda_prs
120e479ed367
first attempt to deal with Babs in regularise and cleaning (not yet working)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 284
(* 2. simplification with *)
120e479ed367
first attempt to deal with Babs in regularise and cleaning (not yet working)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 285
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>
diff
changeset
+ − 286
(* 3. Simplification with *)
120e479ed367
first attempt to deal with Babs in regularise and cleaning (not yet working)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 287
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>
diff
changeset
+ − 288
(* 4. Test for refl *)
120e479ed367
first attempt to deal with Babs in regularise and cleaning (not yet working)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 289
120e479ed367
first attempt to deal with Babs in regularise and cleaning (not yet working)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 290
lemma
120e479ed367
first attempt to deal with Babs in regularise and cleaning (not yet working)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 291
shows "equivp (op \<approx>)"
622
+ − 292
and "equivp ((op \<approx>) ===> (op \<approx>))"
+ − 293
(* 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
+ − 294
oops
120e479ed367
first attempt to deal with Babs in regularise and cleaning (not yet working)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 295
606
+ − 296
lemma lam_tst3a: "(\<lambda>(y :: nat \<times> nat). y) = (\<lambda>(x :: nat \<times> nat). x)"
597
+ − 297
by auto
+ − 298
620
+ − 299
lemma id_rsp:
+ − 300
shows "(R ===> R) id id"
+ − 301
by simp
+ − 302
+ − 303
lemma lam_tst3a_reg: "(op \<approx> ===> op \<approx>) (Babs (Respects op \<approx>) (\<lambda>y. y)) (Babs (Respects op \<approx>) (\<lambda>x. x))"
+ − 304
apply (rule babs_rsp[OF Quotient_my_int])
+ − 305
apply (simp add: id_rsp)
+ − 306
done
+ − 307
606
+ − 308
lemma "(\<lambda>(y :: my_int). y) = (\<lambda>(x :: my_int). x)"
+ − 309
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>
diff
changeset
+ − 310
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
+ − 311
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>
diff
changeset
+ − 312
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>
diff
changeset
+ − 313
apply(tactic {* clean_tac @{context} 1 *})
625
5d6a2b5fb222
Another lambda example theorem proved. Seems it starts working properly.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 314
apply(simp only: babs_prs[OF Quotient_my_int Quotient_my_int])
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
+ − 315
done
606
+ − 316
+ − 317
lemma lam_tst3b: "(\<lambda>(y :: nat \<times> nat \<Rightarrow> nat \<times> nat). y) = (\<lambda>(x :: nat \<times> nat \<Rightarrow> nat \<times> nat). x)"
+ − 318
by auto
+ − 319
+ − 320
lemma "(\<lambda>(y :: my_int => my_int). y) = (\<lambda>(x :: my_int => my_int). x)"
+ − 321
apply(tactic {* procedure_tac @{context} @{thm lam_tst3b} 1 *})
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
+ − 322
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
+ − 323
apply (rule babs_rsp[OF fun_quotient[OF Quotient_my_int Quotient_my_int]])
c10a46fa0de9
Added a 'rep_abs' in inj_repabs_trm of babs; and proved two lam examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 324
apply (simp add: id_rsp)
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>
diff
changeset
+ − 325
apply(tactic {* all_inj_repabs_tac @{context} 1*})
606
+ − 326
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>
diff
changeset
+ − 327
apply(subst babs_prs)
c10a46fa0de9
Added a 'rep_abs' in inj_repabs_trm of babs; and proved two lam examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 328
apply(tactic {* quotient_tac @{context} 1 *})
c10a46fa0de9
Added a 'rep_abs' in inj_repabs_trm of babs; and proved two lam examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 329
apply(tactic {* quotient_tac @{context} 1 *})
c10a46fa0de9
Added a 'rep_abs' in inj_repabs_trm of babs; and proved two lam examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 330
apply(subst babs_prs)
c10a46fa0de9
Added a 'rep_abs' in inj_repabs_trm of babs; and proved two lam examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 331
apply(tactic {* quotient_tac @{context} 1 *})
c10a46fa0de9
Added a 'rep_abs' in inj_repabs_trm of babs; and proved two lam examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 332
apply(tactic {* quotient_tac @{context} 1 *})
c10a46fa0de9
Added a 'rep_abs' in inj_repabs_trm of babs; and proved two lam examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 333
apply(rule refl)
625
5d6a2b5fb222
Another lambda example theorem proved. Seems it starts working properly.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 334
done
622
+ − 335
625
5d6a2b5fb222
Another lambda example theorem proved. Seems it starts working properly.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 336
term map
5d6a2b5fb222
Another lambda example theorem proved. Seems it starts working properly.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 337
5d6a2b5fb222
Another lambda example theorem proved. Seems it starts working properly.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 338
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
+ − 339
apply (induct l)
5d6a2b5fb222
Another lambda example theorem proved. Seems it starts working properly.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 340
apply simp
5d6a2b5fb222
Another lambda example theorem proved. Seems it starts working properly.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 341
apply (case_tac a)
5d6a2b5fb222
Another lambda example theorem proved. Seems it starts working properly.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 342
apply simp
5d6a2b5fb222
Another lambda example theorem proved. Seems it starts working properly.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 343
done
5d6a2b5fb222
Another lambda example theorem proved. Seems it starts working properly.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 344
5d6a2b5fb222
Another lambda example theorem proved. Seems it starts working properly.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 345
lemma "map (\<lambda>x. PLUS x ZERO) l = l"
637
+ − 346
apply(lifting lam_tst4)
636
+ − 347
apply(cleaning)
625
5d6a2b5fb222
Another lambda example theorem proved. Seems it starts working properly.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 348
apply(simp only: babs_prs[OF Quotient_my_int Quotient_my_int])
636
+ − 349
apply(simp only: map_prs[OF Quotient_my_int Quotient_my_int, symmetric])
625
5d6a2b5fb222
Another lambda example theorem proved. Seems it starts working properly.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 350
done
601
+ − 351
617
ca37f4b6457c
An example of working cleaning for lambda lifting. Still not sure why Babs helps.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 352
end