IntEx.thy
author Cezary Kaliszyk <kaliszyk@in.tum.de>
Mon, 07 Dec 2009 04:41:42 +0100
changeset 592 66f39908df95
parent 591 01a0da807f50
parent 588 2c95d0436a2b
child 593 18eac4596ef1
permissions -rw-r--r--
merge
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
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
281
46e6d06efe3f Experiments with lifting partially applied constants.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 276
diff changeset
     6
  intrel :: "(nat \<times> nat) \<Rightarrow> (nat \<times> nat) \<Rightarrow> bool" (infix "\<approx>" 50)
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
529
6348c2a57ec2 More name changes
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 506
diff changeset
    11
  apply(unfold equivp_def)
181
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
582
a082e2d138ab added a theorem list for equivalence theorems
Christian Urban <urbanc@in.tum.de>
parents: 572
diff changeset
    15
thm quotient_equiv
a082e2d138ab added a theorem list for equivalence theorems
Christian Urban <urbanc@in.tum.de>
parents: 572
diff changeset
    16
506
91c374abde06 removed quot argument...not all examples work anymore
Christian Urban <urbanc@in.tum.de>
parents: 505
diff changeset
    17
thm quotient_thm
91c374abde06 removed quot argument...not all examples work anymore
Christian Urban <urbanc@in.tum.de>
parents: 505
diff changeset
    18
529
6348c2a57ec2 More name changes
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 506
diff changeset
    19
thm my_int_equivp
321
f46dc0ca08c3 simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
parents: 320
diff changeset
    20
239
02b14a21761a Cleaning of the interface to lift.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 228
diff changeset
    21
print_theorems
274
df225aa45770 simplified the quotient_def code
Christian Urban <urbanc@in.tum.de>
parents: 268
diff changeset
    22
print_quotients
239
02b14a21761a Cleaning of the interface to lift.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 228
diff changeset
    23
268
4d58c02289ca simplified the quotient_def code; type of the defined constant must now be given; for-part eliminated
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
    24
quotient_def 
218
df05cd030d2f added infrastructure for defining lifted constants
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
    25
  ZERO::"my_int"
df05cd030d2f added infrastructure for defining lifted constants
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
    26
where
df05cd030d2f added infrastructure for defining lifted constants
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
    27
  "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
    28
318
746b17e1d6d8 fixed the storage of qconst definitions
Christian Urban <urbanc@in.tum.de>
parents: 312
diff changeset
    29
ML {* print_qconstinfo @{context} *}
746b17e1d6d8 fixed the storage of qconst definitions
Christian Urban <urbanc@in.tum.de>
parents: 312
diff changeset
    30
268
4d58c02289ca simplified the quotient_def code; type of the defined constant must now be given; for-part eliminated
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
    31
term ZERO
4d58c02289ca simplified the quotient_def code; type of the defined constant must now be given; for-part eliminated
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
    32
thm ZERO_def
181
3e53081ad53a added another example file about integers (see HOL/Int.thy)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    33
290
a0be84b0c707 removed typing information from get_fun in quotient_def; *potentially* dangerous
Christian Urban <urbanc@in.tum.de>
parents: 286
diff changeset
    34
ML {* prop_of @{thm ZERO_def} *}
a0be84b0c707 removed typing information from get_fun in quotient_def; *potentially* dangerous
Christian Urban <urbanc@in.tum.de>
parents: 286
diff changeset
    35
318
746b17e1d6d8 fixed the storage of qconst definitions
Christian Urban <urbanc@in.tum.de>
parents: 312
diff changeset
    36
ML {* separate *}
746b17e1d6d8 fixed the storage of qconst definitions
Christian Urban <urbanc@in.tum.de>
parents: 312
diff changeset
    37
268
4d58c02289ca simplified the quotient_def code; type of the defined constant must now be given; for-part eliminated
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
    38
quotient_def 
218
df05cd030d2f added infrastructure for defining lifted constants
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
    39
  ONE::"my_int"
193
Christian Urban <urbanc@in.tum.de>
parents: 192
diff changeset
    40
where
218
df05cd030d2f added infrastructure for defining lifted constants
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
    41
  "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
    42
318
746b17e1d6d8 fixed the storage of qconst definitions
Christian Urban <urbanc@in.tum.de>
parents: 312
diff changeset
    43
ML {* print_qconstinfo @{context} *}
746b17e1d6d8 fixed the storage of qconst definitions
Christian Urban <urbanc@in.tum.de>
parents: 312
diff changeset
    44
181
3e53081ad53a added another example file about integers (see HOL/Int.thy)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    45
term ONE
3e53081ad53a added another example file about integers (see HOL/Int.thy)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    46
thm ONE_def
3e53081ad53a added another example file about integers (see HOL/Int.thy)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    47
3e53081ad53a added another example file about integers (see HOL/Int.thy)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    48
fun
3e53081ad53a added another example file about integers (see HOL/Int.thy)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    49
  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
    50
where
3e53081ad53a added another example file about integers (see HOL/Int.thy)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    51
  "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
    52
268
4d58c02289ca simplified the quotient_def code; type of the defined constant must now be given; for-part eliminated
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
    53
quotient_def 
218
df05cd030d2f added infrastructure for defining lifted constants
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
    54
  PLUS::"my_int \<Rightarrow> my_int \<Rightarrow> my_int"
df05cd030d2f added infrastructure for defining lifted constants
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
    55
where
275
34ad627ac5d5 fixed definition of PLUS
Christian Urban <urbanc@in.tum.de>
parents: 274
diff changeset
    56
  "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
    57
318
746b17e1d6d8 fixed the storage of qconst definitions
Christian Urban <urbanc@in.tum.de>
parents: 312
diff changeset
    58
term my_plus
181
3e53081ad53a added another example file about integers (see HOL/Int.thy)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    59
term PLUS
3e53081ad53a added another example file about integers (see HOL/Int.thy)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    60
thm PLUS_def
3e53081ad53a added another example file about integers (see HOL/Int.thy)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    61
3e53081ad53a added another example file about integers (see HOL/Int.thy)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    62
fun
3e53081ad53a added another example file about integers (see HOL/Int.thy)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    63
  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
    64
where
3e53081ad53a added another example file about integers (see HOL/Int.thy)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    65
  "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
    66
268
4d58c02289ca simplified the quotient_def code; type of the defined constant must now be given; for-part eliminated
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
    67
quotient_def 
220
af951c8fb80a updated all definitions
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    68
  NEG::"my_int \<Rightarrow> my_int"
af951c8fb80a updated all definitions
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    69
where
af951c8fb80a updated all definitions
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    70
  "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
    71
3e53081ad53a added another example file about integers (see HOL/Int.thy)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    72
term NEG
3e53081ad53a added another example file about integers (see HOL/Int.thy)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    73
thm NEG_def
3e53081ad53a added another example file about integers (see HOL/Int.thy)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    74
3e53081ad53a added another example file about integers (see HOL/Int.thy)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    75
definition
3e53081ad53a added another example file about integers (see HOL/Int.thy)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    76
  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
    77
where
3e53081ad53a added another example file about integers (see HOL/Int.thy)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    78
  "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
    79
3e53081ad53a added another example file about integers (see HOL/Int.thy)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    80
fun
3e53081ad53a added another example file about integers (see HOL/Int.thy)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    81
  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
    82
where
3e53081ad53a added another example file about integers (see HOL/Int.thy)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    83
  "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
    84
268
4d58c02289ca simplified the quotient_def code; type of the defined constant must now be given; for-part eliminated
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
    85
quotient_def 
220
af951c8fb80a updated all definitions
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    86
  MULT::"my_int \<Rightarrow> my_int \<Rightarrow> my_int"
af951c8fb80a updated all definitions
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    87
where
af951c8fb80a updated all definitions
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    88
  "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
    89
3e53081ad53a added another example file about integers (see HOL/Int.thy)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    90
term MULT
3e53081ad53a added another example file about integers (see HOL/Int.thy)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    91
thm MULT_def
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
(* 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
    94
fun
3e53081ad53a added another example file about integers (see HOL/Int.thy)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    95
  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
    96
where
3e53081ad53a added another example file about integers (see HOL/Int.thy)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    97
  "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
    98
268
4d58c02289ca simplified the quotient_def code; type of the defined constant must now be given; for-part eliminated
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
    99
quotient_def 
220
af951c8fb80a updated all definitions
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
   100
  LE :: "my_int \<Rightarrow> my_int \<Rightarrow> bool"
af951c8fb80a updated all definitions
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
   101
where
af951c8fb80a updated all definitions
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
   102
  "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
   103
3e53081ad53a added another example file about integers (see HOL/Int.thy)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   104
term LE
3e53081ad53a added another example file about integers (see HOL/Int.thy)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   105
thm LE_def
3e53081ad53a added another example file about integers (see HOL/Int.thy)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   106
321
f46dc0ca08c3 simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
parents: 320
diff changeset
   107
181
3e53081ad53a added another example file about integers (see HOL/Int.thy)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   108
definition
3e53081ad53a added another example file about integers (see HOL/Int.thy)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   109
  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
   110
where
3e53081ad53a added another example file about integers (see HOL/Int.thy)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   111
  "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
   112
3e53081ad53a added another example file about integers (see HOL/Int.thy)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   113
term LESS
3e53081ad53a added another example file about integers (see HOL/Int.thy)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   114
thm LESS_def
3e53081ad53a added another example file about integers (see HOL/Int.thy)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   115
3e53081ad53a added another example file about integers (see HOL/Int.thy)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   116
definition
3e53081ad53a added another example file about integers (see HOL/Int.thy)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   117
  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
   118
where
3e53081ad53a added another example file about integers (see HOL/Int.thy)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   119
  "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
   120
3e53081ad53a added another example file about integers (see HOL/Int.thy)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   121
definition
3e53081ad53a added another example file about integers (see HOL/Int.thy)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   122
  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
   123
where
3e53081ad53a added another example file about integers (see HOL/Int.thy)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   124
 "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
   125
310
fec6301a1989 added a container for quotient constants (does not work yet though)
Christian Urban <urbanc@in.tum.de>
parents: 306
diff changeset
   126
ML {* print_qconstinfo @{context} *}
fec6301a1989 added a container for quotient constants (does not work yet though)
Christian Urban <urbanc@in.tum.de>
parents: 306
diff changeset
   127
199
Christian Urban <urbanc@in.tum.de>
parents: 198
diff changeset
   128
lemma plus_sym_pre:
281
46e6d06efe3f Experiments with lifting partially applied constants.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 276
diff changeset
   129
  shows "my_plus a b \<approx> my_plus b a"
199
Christian Urban <urbanc@in.tum.de>
parents: 198
diff changeset
   130
  apply(cases a)
Christian Urban <urbanc@in.tum.de>
parents: 198
diff changeset
   131
  apply(cases b)
Christian Urban <urbanc@in.tum.de>
parents: 198
diff changeset
   132
  apply(auto)
Christian Urban <urbanc@in.tum.de>
parents: 198
diff changeset
   133
  done
Christian Urban <urbanc@in.tum.de>
parents: 198
diff changeset
   134
561
41500cd131dc Added nil_rsp and cons_rsp to quotient_rsp; simplified IntEx.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 559
diff changeset
   135
lemma plus_rsp[quotient_rsp]:
450
2dc708ddb93a introduced a global list of respectfulness lemmas; the attribute is [quot_rsp]
Christian Urban <urbanc@in.tum.de>
parents: 446
diff changeset
   136
  shows "(intrel ===> intrel ===> intrel) my_plus my_plus"
2dc708ddb93a introduced a global list of respectfulness lemmas; the attribute is [quot_rsp]
Christian Urban <urbanc@in.tum.de>
parents: 446
diff changeset
   137
by (simp)
192
a296bf1a3b09 Simplifying code in int
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 191
diff changeset
   138
a296bf1a3b09 Simplifying code in int
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 191
diff changeset
   139
ML {* val qty = @{typ "my_int"} *}
377
edd71fd83a2d cleaning in MyInt
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 374
diff changeset
   140
ML {* val (rty, rel, rel_refl, rel_eqv) = lookup_quot_data @{context} qty *}
edd71fd83a2d cleaning in MyInt
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 374
diff changeset
   141
ML {* val (trans2, reps_same, absrep, quot) = lookup_quot_thms @{context} "my_int"; *}
197
c0f2db9a243b Further reordering in Int code.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 196
diff changeset
   142
582
a082e2d138ab added a theorem list for equivalence theorems
Christian Urban <urbanc@in.tum.de>
parents: 572
diff changeset
   143
ML {* fun lift_tac_intex lthy t = lift_tac lthy t *}
423
2f0ad33f0241 annotated a proof with all steps and simplified LAMBDA_RES_TAC
Christian Urban <urbanc@in.tum.de>
parents: 419
diff changeset
   144
506
91c374abde06 removed quot argument...not all examples work anymore
Christian Urban <urbanc@in.tum.de>
parents: 505
diff changeset
   145
ML {* fun inj_repabs_tac_intex lthy = inj_repabs_tac lthy [rel_refl] [trans2] *}
91c374abde06 removed quot argument...not all examples work anymore
Christian Urban <urbanc@in.tum.de>
parents: 505
diff changeset
   146
ML {* fun all_inj_repabs_tac_intex lthy = all_inj_repabs_tac lthy [rel_refl] [trans2] *}
423
2f0ad33f0241 annotated a proof with all steps and simplified LAMBDA_RES_TAC
Christian Urban <urbanc@in.tum.de>
parents: 419
diff changeset
   147
552
d9151fa76f84 added three examples to IntEx for testing ideas - regularisation and injection seem to be not quite right yet
Christian Urban <urbanc@in.tum.de>
parents: 549
diff changeset
   148
lemma test1: "my_plus a b = my_plus a b"
d9151fa76f84 added three examples to IntEx for testing ideas - regularisation and injection seem to be not quite right yet
Christian Urban <urbanc@in.tum.de>
parents: 549
diff changeset
   149
apply(rule refl)
d9151fa76f84 added three examples to IntEx for testing ideas - regularisation and injection seem to be not quite right yet
Christian Urban <urbanc@in.tum.de>
parents: 549
diff changeset
   150
done
d9151fa76f84 added three examples to IntEx for testing ideas - regularisation and injection seem to be not quite right yet
Christian Urban <urbanc@in.tum.de>
parents: 549
diff changeset
   151
d9151fa76f84 added three examples to IntEx for testing ideas - regularisation and injection seem to be not quite right yet
Christian Urban <urbanc@in.tum.de>
parents: 549
diff changeset
   152
lemma "PLUS a b = PLUS a b"
d9151fa76f84 added three examples to IntEx for testing ideas - regularisation and injection seem to be not quite right yet
Christian Urban <urbanc@in.tum.de>
parents: 549
diff changeset
   153
apply(tactic {* procedure_tac @{context} @{thm test1} 1 *})
582
a082e2d138ab added a theorem list for equivalence theorems
Christian Urban <urbanc@in.tum.de>
parents: 572
diff changeset
   154
apply(tactic {* regularize_tac @{context} 1 *})
553
09cd71fac4ec made some slight simplifications to the examples
Christian Urban <urbanc@in.tum.de>
parents: 552
diff changeset
   155
apply(tactic {* all_inj_repabs_tac_intex @{context} 1*})
09cd71fac4ec made some slight simplifications to the examples
Christian Urban <urbanc@in.tum.de>
parents: 552
diff changeset
   156
apply(tactic {* clean_tac @{context} 1 *}) 
09cd71fac4ec made some slight simplifications to the examples
Christian Urban <urbanc@in.tum.de>
parents: 552
diff changeset
   157
done
552
d9151fa76f84 added three examples to IntEx for testing ideas - regularisation and injection seem to be not quite right yet
Christian Urban <urbanc@in.tum.de>
parents: 549
diff changeset
   158
588
2c95d0436a2b removed usage of HOL_basic_ss by using a slighly extended version of empty_ss
Christian Urban <urbanc@in.tum.de>
parents: 586
diff changeset
   159
thm lambda_prs
2c95d0436a2b removed usage of HOL_basic_ss by using a slighly extended version of empty_ss
Christian Urban <urbanc@in.tum.de>
parents: 586
diff changeset
   160
552
d9151fa76f84 added three examples to IntEx for testing ideas - regularisation and injection seem to be not quite right yet
Christian Urban <urbanc@in.tum.de>
parents: 549
diff changeset
   161
lemma test2: "my_plus a = my_plus a"
d9151fa76f84 added three examples to IntEx for testing ideas - regularisation and injection seem to be not quite right yet
Christian Urban <urbanc@in.tum.de>
parents: 549
diff changeset
   162
apply(rule refl)
d9151fa76f84 added three examples to IntEx for testing ideas - regularisation and injection seem to be not quite right yet
Christian Urban <urbanc@in.tum.de>
parents: 549
diff changeset
   163
done
d9151fa76f84 added three examples to IntEx for testing ideas - regularisation and injection seem to be not quite right yet
Christian Urban <urbanc@in.tum.de>
parents: 549
diff changeset
   164
d9151fa76f84 added three examples to IntEx for testing ideas - regularisation and injection seem to be not quite right yet
Christian Urban <urbanc@in.tum.de>
parents: 549
diff changeset
   165
lemma "PLUS a = PLUS a"
d9151fa76f84 added three examples to IntEx for testing ideas - regularisation and injection seem to be not quite right yet
Christian Urban <urbanc@in.tum.de>
parents: 549
diff changeset
   166
apply(tactic {* procedure_tac @{context} @{thm test2} 1 *})
554
8395fc6a6945 Solutions to IntEx tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 553
diff changeset
   167
apply(rule ballI)
561
41500cd131dc Added nil_rsp and cons_rsp to quotient_rsp; simplified IntEx.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 559
diff changeset
   168
apply(rule apply_rsp[OF Quotient_my_int plus_rsp])
554
8395fc6a6945 Solutions to IntEx tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 553
diff changeset
   169
apply(simp only: in_respects)
555
b460565dbb58 Handling of respects in the fast inj_repabs_tac; includes respects with quotient assumptions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 554
diff changeset
   170
apply(tactic {* all_inj_repabs_tac_intex @{context} 1*})
554
8395fc6a6945 Solutions to IntEx tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 553
diff changeset
   171
apply(tactic {* clean_tac @{context} 1 *})
8395fc6a6945 Solutions to IntEx tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 553
diff changeset
   172
done
552
d9151fa76f84 added three examples to IntEx for testing ideas - regularisation and injection seem to be not quite right yet
Christian Urban <urbanc@in.tum.de>
parents: 549
diff changeset
   173
d9151fa76f84 added three examples to IntEx for testing ideas - regularisation and injection seem to be not quite right yet
Christian Urban <urbanc@in.tum.de>
parents: 549
diff changeset
   174
lemma test3: "my_plus = my_plus"
d9151fa76f84 added three examples to IntEx for testing ideas - regularisation and injection seem to be not quite right yet
Christian Urban <urbanc@in.tum.de>
parents: 549
diff changeset
   175
apply(rule refl)
d9151fa76f84 added three examples to IntEx for testing ideas - regularisation and injection seem to be not quite right yet
Christian Urban <urbanc@in.tum.de>
parents: 549
diff changeset
   176
done
d9151fa76f84 added three examples to IntEx for testing ideas - regularisation and injection seem to be not quite right yet
Christian Urban <urbanc@in.tum.de>
parents: 549
diff changeset
   177
d9151fa76f84 added three examples to IntEx for testing ideas - regularisation and injection seem to be not quite right yet
Christian Urban <urbanc@in.tum.de>
parents: 549
diff changeset
   178
lemma "PLUS = PLUS"
d9151fa76f84 added three examples to IntEx for testing ideas - regularisation and injection seem to be not quite right yet
Christian Urban <urbanc@in.tum.de>
parents: 549
diff changeset
   179
apply(tactic {* procedure_tac @{context} @{thm test3} 1 *})
561
41500cd131dc Added nil_rsp and cons_rsp to quotient_rsp; simplified IntEx.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 559
diff changeset
   180
apply(rule plus_rsp)
555
b460565dbb58 Handling of respects in the fast inj_repabs_tac; includes respects with quotient assumptions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 554
diff changeset
   181
apply(tactic {* all_inj_repabs_tac_intex @{context} 1*})
554
8395fc6a6945 Solutions to IntEx tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 553
diff changeset
   182
apply(tactic {* clean_tac @{context} 1 *})
8395fc6a6945 Solutions to IntEx tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 553
diff changeset
   183
done
552
d9151fa76f84 added three examples to IntEx for testing ideas - regularisation and injection seem to be not quite right yet
Christian Urban <urbanc@in.tum.de>
parents: 549
diff changeset
   184
423
2f0ad33f0241 annotated a proof with all steps and simplified LAMBDA_RES_TAC
Christian Urban <urbanc@in.tum.de>
parents: 419
diff changeset
   185
377
edd71fd83a2d cleaning in MyInt
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 374
diff changeset
   186
lemma "PLUS a b = PLUS b a"
423
2f0ad33f0241 annotated a proof with all steps and simplified LAMBDA_RES_TAC
Christian Urban <urbanc@in.tum.de>
parents: 419
diff changeset
   187
apply(tactic {* procedure_tac @{context} @{thm plus_sym_pre} 1 *})
582
a082e2d138ab added a theorem list for equivalence theorems
Christian Urban <urbanc@in.tum.de>
parents: 572
diff changeset
   188
apply(tactic {* regularize_tac @{context} 1 *})
559
Christian Urban <urbanc@in.tum.de>
parents: 558 555
diff changeset
   189
apply(tactic {* all_inj_repabs_tac_intex @{context} 1*})
549
f178958d3d81 not yet quite functional treatment of constants
Christian Urban <urbanc@in.tum.de>
parents: 540
diff changeset
   190
apply(tactic {* clean_tac @{context} 1 *})
423
2f0ad33f0241 annotated a proof with all steps and simplified LAMBDA_RES_TAC
Christian Urban <urbanc@in.tum.de>
parents: 419
diff changeset
   191
done
342
eb15be678ac4 lift_thm with a goal.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 340
diff changeset
   192
198
ff4425e000db Completely cleaned Int.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 197
diff changeset
   193
lemma plus_assoc_pre:
281
46e6d06efe3f Experiments with lifting partially applied constants.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 276
diff changeset
   194
  shows "my_plus (my_plus i j) k \<approx> my_plus i (my_plus j k)"
198
ff4425e000db Completely cleaned Int.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 197
diff changeset
   195
  apply (cases i)
ff4425e000db Completely cleaned Int.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 197
diff changeset
   196
  apply (cases j)
ff4425e000db Completely cleaned Int.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 197
diff changeset
   197
  apply (cases k)
224
f9a25fe22037 Cleaning the unnecessary theorems in 'IntEx'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 222
diff changeset
   198
  apply (simp)
198
ff4425e000db Completely cleaned Int.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 197
diff changeset
   199
  done
ff4425e000db Completely cleaned Int.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 197
diff changeset
   200
377
edd71fd83a2d cleaning in MyInt
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 374
diff changeset
   201
lemma plus_assoc: "PLUS (PLUS x xa) xb = PLUS x (PLUS xa xb)"
423
2f0ad33f0241 annotated a proof with all steps and simplified LAMBDA_RES_TAC
Christian Urban <urbanc@in.tum.de>
parents: 419
diff changeset
   202
apply(tactic {* procedure_tac @{context} @{thm plus_assoc_pre} 1 *})
582
a082e2d138ab added a theorem list for equivalence theorems
Christian Urban <urbanc@in.tum.de>
parents: 572
diff changeset
   203
apply(tactic {* regularize_tac @{context} 1 *})
445
f1c0a66284d3 renamed r_mk_comb_tac to inj_repabs_tac
Christian Urban <urbanc@in.tum.de>
parents: 433
diff changeset
   204
apply(tactic {* all_inj_repabs_tac_intex @{context} 1*})
506
91c374abde06 removed quot argument...not all examples work anymore
Christian Urban <urbanc@in.tum.de>
parents: 505
diff changeset
   205
apply(tactic {* clean_tac @{context} 1 *})
423
2f0ad33f0241 annotated a proof with all steps and simplified LAMBDA_RES_TAC
Christian Urban <urbanc@in.tum.de>
parents: 419
diff changeset
   206
done
354
2eb6d527dfe4 addded a tactic, which sets up the three goals of the `algorithm'
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
   207
379
57bde65f6eb2 Removed unused things from QuotMain.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 377
diff changeset
   208
lemma ho_tst: "foldl my_plus x [] = x"
57bde65f6eb2 Removed unused things from QuotMain.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 377
diff changeset
   209
apply simp
57bde65f6eb2 Removed unused things from QuotMain.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 377
diff changeset
   210
done
57bde65f6eb2 Removed unused things from QuotMain.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 377
diff changeset
   211
57bde65f6eb2 Removed unused things from QuotMain.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 377
diff changeset
   212
lemma "foldl PLUS x [] = x"
445
f1c0a66284d3 renamed r_mk_comb_tac to inj_repabs_tac
Christian Urban <urbanc@in.tum.de>
parents: 433
diff changeset
   213
apply(tactic {* procedure_tac @{context} @{thm ho_tst} 1 *})
582
a082e2d138ab added a theorem list for equivalence theorems
Christian Urban <urbanc@in.tum.de>
parents: 572
diff changeset
   214
apply(tactic {* regularize_tac @{context} 1 *})
445
f1c0a66284d3 renamed r_mk_comb_tac to inj_repabs_tac
Christian Urban <urbanc@in.tum.de>
parents: 433
diff changeset
   215
apply(tactic {* all_inj_repabs_tac_intex @{context} 1*})
506
91c374abde06 removed quot argument...not all examples work anymore
Christian Urban <urbanc@in.tum.de>
parents: 505
diff changeset
   216
apply(tactic {* clean_tac @{context} 1 *})
561
41500cd131dc Added nil_rsp and cons_rsp to quotient_rsp; simplified IntEx.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 559
diff changeset
   217
apply(simp only: foldl_prs[OF Quotient_my_int Quotient_my_int] nil_prs[OF Quotient_my_int])
556
287ea842a7d4 simpler version of clean_tac
Christian Urban <urbanc@in.tum.de>
parents: 553
diff changeset
   218
done
359
64c3c83e0ed4 New cleaning tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 358
diff changeset
   219
485
4efb104514f7 More experiments with higher order quotients and theorems with non-lifted constants.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 466
diff changeset
   220
lemma ho_tst2: "foldl my_plus x (h # t) \<approx> my_plus h (foldl my_plus x t)"
4efb104514f7 More experiments with higher order quotients and theorems with non-lifted constants.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 466
diff changeset
   221
sorry
327
22c8bf90cadb tried to prove the repabs_inj lemma, but failed for the moment
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
   222
485
4efb104514f7 More experiments with higher order quotients and theorems with non-lifted constants.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 466
diff changeset
   223
lemma "foldl PLUS x (h # t) = PLUS h (foldl PLUS x t)"
4efb104514f7 More experiments with higher order quotients and theorems with non-lifted constants.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 466
diff changeset
   224
apply(tactic {* procedure_tac @{context} @{thm ho_tst2} 1 *})
582
a082e2d138ab added a theorem list for equivalence theorems
Christian Urban <urbanc@in.tum.de>
parents: 572
diff changeset
   225
apply(tactic {* regularize_tac @{context} 1 *})
485
4efb104514f7 More experiments with higher order quotients and theorems with non-lifted constants.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 466
diff changeset
   226
apply(tactic {* all_inj_repabs_tac_intex @{context} 1*})
506
91c374abde06 removed quot argument...not all examples work anymore
Christian Urban <urbanc@in.tum.de>
parents: 505
diff changeset
   227
apply(tactic {* clean_tac @{context} 1 *})
561
41500cd131dc Added nil_rsp and cons_rsp to quotient_rsp; simplified IntEx.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 559
diff changeset
   228
apply(simp only: foldl_prs[OF Quotient_my_int Quotient_my_int] cons_prs[OF Quotient_my_int])
485
4efb104514f7 More experiments with higher order quotients and theorems with non-lifted constants.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 466
diff changeset
   229
done
4efb104514f7 More experiments with higher order quotients and theorems with non-lifted constants.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 466
diff changeset
   230
591
01a0da807f50 3 new example thms in MyInt; reveal problems with handling of lambdas; regularize fails with "Loose Bound".
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 582
diff changeset
   231
lemma ho_tst3: "foldl f (s::nat \<times> nat) ([]::(nat \<times> nat) list) = s"
01a0da807f50 3 new example thms in MyInt; reveal problems with handling of lambdas; regularize fails with "Loose Bound".
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 582
diff changeset
   232
by simp
01a0da807f50 3 new example thms in MyInt; reveal problems with handling of lambdas; regularize fails with "Loose Bound".
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 582
diff changeset
   233
01a0da807f50 3 new example thms in MyInt; reveal problems with handling of lambdas; regularize fails with "Loose Bound".
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 582
diff changeset
   234
lemma "foldl f (x::my_int) ([]::my_int list) = x"
01a0da807f50 3 new example thms in MyInt; reveal problems with handling of lambdas; regularize fails with "Loose Bound".
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 582
diff changeset
   235
apply(tactic {* procedure_tac @{context} @{thm ho_tst3} 1 *})
01a0da807f50 3 new example thms in MyInt; reveal problems with handling of lambdas; regularize fails with "Loose Bound".
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 582
diff changeset
   236
apply(tactic {* regularize_tac @{context} 1 *})
01a0da807f50 3 new example thms in MyInt; reveal problems with handling of lambdas; regularize fails with "Loose Bound".
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 582
diff changeset
   237
apply(tactic {* all_inj_repabs_tac_intex @{context} 1*})
01a0da807f50 3 new example thms in MyInt; reveal problems with handling of lambdas; regularize fails with "Loose Bound".
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 582
diff changeset
   238
(* TODO: does not work when this is added *)
01a0da807f50 3 new example thms in MyInt; reveal problems with handling of lambdas; regularize fails with "Loose Bound".
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 582
diff changeset
   239
(* apply(tactic {* lambda_prs_tac @{context} 1 *})*)
01a0da807f50 3 new example thms in MyInt; reveal problems with handling of lambdas; regularize fails with "Loose Bound".
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 582
diff changeset
   240
apply(tactic {* clean_tac @{context} 1 *})
01a0da807f50 3 new example thms in MyInt; reveal problems with handling of lambdas; regularize fails with "Loose Bound".
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 582
diff changeset
   241
apply(simp only: foldl_prs[OF Quotient_my_int Quotient_my_int] nil_prs[OF Quotient_my_int])
01a0da807f50 3 new example thms in MyInt; reveal problems with handling of lambdas; regularize fails with "Loose Bound".
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 582
diff changeset
   242
done
01a0da807f50 3 new example thms in MyInt; reveal problems with handling of lambdas; regularize fails with "Loose Bound".
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 582
diff changeset
   243
01a0da807f50 3 new example thms in MyInt; reveal problems with handling of lambdas; regularize fails with "Loose Bound".
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 582
diff changeset
   244
lemma lam_tst: "(\<lambda>x. (x, x)) y = (y, (y :: nat \<times> nat))"
01a0da807f50 3 new example thms in MyInt; reveal problems with handling of lambdas; regularize fails with "Loose Bound".
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 582
diff changeset
   245
by simp
01a0da807f50 3 new example thms in MyInt; reveal problems with handling of lambdas; regularize fails with "Loose Bound".
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 582
diff changeset
   246
01a0da807f50 3 new example thms in MyInt; reveal problems with handling of lambdas; regularize fails with "Loose Bound".
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 582
diff changeset
   247
lemma "(\<lambda>x. (x, x)) (y::my_int) = (y, y)"
01a0da807f50 3 new example thms in MyInt; reveal problems with handling of lambdas; regularize fails with "Loose Bound".
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 582
diff changeset
   248
apply(tactic {* procedure_tac @{context} @{thm lam_tst} 1 *})
01a0da807f50 3 new example thms in MyInt; reveal problems with handling of lambdas; regularize fails with "Loose Bound".
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 582
diff changeset
   249
defer
01a0da807f50 3 new example thms in MyInt; reveal problems with handling of lambdas; regularize fails with "Loose Bound".
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 582
diff changeset
   250
apply(tactic {* inj_repabs_tac_intex @{context} 1*})
01a0da807f50 3 new example thms in MyInt; reveal problems with handling of lambdas; regularize fails with "Loose Bound".
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 582
diff changeset
   251
apply(tactic {* inj_repabs_tac_intex @{context} 1*})
01a0da807f50 3 new example thms in MyInt; reveal problems with handling of lambdas; regularize fails with "Loose Bound".
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 582
diff changeset
   252
apply(tactic {* inj_repabs_tac_intex @{context} 1*})
01a0da807f50 3 new example thms in MyInt; reveal problems with handling of lambdas; regularize fails with "Loose Bound".
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 582
diff changeset
   253
apply(simp only: prod_rel.simps)
01a0da807f50 3 new example thms in MyInt; reveal problems with handling of lambdas; regularize fails with "Loose Bound".
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 582
diff changeset
   254
defer
01a0da807f50 3 new example thms in MyInt; reveal problems with handling of lambdas; regularize fails with "Loose Bound".
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 582
diff changeset
   255
apply(tactic {* clean_tac @{context} 1 *})
01a0da807f50 3 new example thms in MyInt; reveal problems with handling of lambdas; regularize fails with "Loose Bound".
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 582
diff changeset
   256
apply(simp add: prod_rel.simps)
01a0da807f50 3 new example thms in MyInt; reveal problems with handling of lambdas; regularize fails with "Loose Bound".
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 582
diff changeset
   257
apply(tactic {* clean_tac @{context} 1 *})
01a0da807f50 3 new example thms in MyInt; reveal problems with handling of lambdas; regularize fails with "Loose Bound".
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 582
diff changeset
   258
apply(simp)
01a0da807f50 3 new example thms in MyInt; reveal problems with handling of lambdas; regularize fails with "Loose Bound".
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 582
diff changeset
   259
(*apply(tactic {* regularize_tac @{context} 1 *})
01a0da807f50 3 new example thms in MyInt; reveal problems with handling of lambdas; regularize fails with "Loose Bound".
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 582
diff changeset
   260
apply(tactic {* inj_repabs_tac_intex @{context} 1*})*)
01a0da807f50 3 new example thms in MyInt; reveal problems with handling of lambdas; regularize fails with "Loose Bound".
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 582
diff changeset
   261
sorry
01a0da807f50 3 new example thms in MyInt; reveal problems with handling of lambdas; regularize fails with "Loose Bound".
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 582
diff changeset
   262
01a0da807f50 3 new example thms in MyInt; reveal problems with handling of lambdas; regularize fails with "Loose Bound".
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 582
diff changeset
   263
lemma lam_tst2: "(\<lambda>(y :: nat \<times> nat). y) = (\<lambda>(x :: nat \<times> nat). x)"
01a0da807f50 3 new example thms in MyInt; reveal problems with handling of lambdas; regularize fails with "Loose Bound".
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 582
diff changeset
   264
by simp
01a0da807f50 3 new example thms in MyInt; reveal problems with handling of lambdas; regularize fails with "Loose Bound".
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 582
diff changeset
   265
01a0da807f50 3 new example thms in MyInt; reveal problems with handling of lambdas; regularize fails with "Loose Bound".
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 582
diff changeset
   266
lemma "(\<lambda>(y :: my_int). y) = (\<lambda>(x :: my_int). x)"
01a0da807f50 3 new example thms in MyInt; reveal problems with handling of lambdas; regularize fails with "Loose Bound".
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 582
diff changeset
   267
apply(tactic {* ObjectLogic.full_atomize_tac 1 *})
01a0da807f50 3 new example thms in MyInt; reveal problems with handling of lambdas; regularize fails with "Loose Bound".
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 582
diff changeset
   268
apply(tactic {* gen_frees_tac @{context} 1 *})
01a0da807f50 3 new example thms in MyInt; reveal problems with handling of lambdas; regularize fails with "Loose Bound".
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 582
diff changeset
   269
ML_prf {* procedure_inst @{context} (prop_of @{thm lam_tst2}) @{prop "(\<lambda>(y :: my_int). y) = (\<lambda>(x :: my_int). x)"} *}
01a0da807f50 3 new example thms in MyInt; reveal problems with handling of lambdas; regularize fails with "Loose Bound".
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 582
diff changeset
   270
01a0da807f50 3 new example thms in MyInt; reveal problems with handling of lambdas; regularize fails with "Loose Bound".
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 582
diff changeset
   271
apply(tactic {* procedure_tac @{context} @{thm lam_tst2} 1 *})
01a0da807f50 3 new example thms in MyInt; reveal problems with handling of lambdas; regularize fails with "Loose Bound".
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 582
diff changeset
   272
apply(tactic {* regularize_tac @{context} 1 *})
01a0da807f50 3 new example thms in MyInt; reveal problems with handling of lambdas; regularize fails with "Loose Bound".
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 582
diff changeset
   273
apply(tactic {* all_inj_repabs_tac_intex @{context} 1*})
01a0da807f50 3 new example thms in MyInt; reveal problems with handling of lambdas; regularize fails with "Loose Bound".
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 582
diff changeset
   274
apply(tactic {* clean_tac @{context} 1 *})
01a0da807f50 3 new example thms in MyInt; reveal problems with handling of lambdas; regularize fails with "Loose Bound".
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 582
diff changeset
   275
done
01a0da807f50 3 new example thms in MyInt; reveal problems with handling of lambdas; regularize fails with "Loose Bound".
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 582
diff changeset
   276