FSet.thy
author Cezary Kaliszyk <kaliszyk@in.tum.de>
Tue, 24 Nov 2009 08:36:28 +0100
changeset 356 51aafebf4d06
parent 353 9a0e8ab42ee8
child 364 4c455d58ac99
permissions -rw-r--r--
Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
163
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     1
theory FSet
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     2
imports QuotMain
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     3
begin
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     4
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     5
inductive
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     6
  list_eq (infix "\<approx>" 50)
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     7
where
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     8
  "a#b#xs \<approx> b#a#xs"
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     9
| "[] \<approx> []"
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    10
| "xs \<approx> ys \<Longrightarrow> ys \<approx> xs"
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    11
| "a#a#xs \<approx> a#xs"
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    12
| "xs \<approx> ys \<Longrightarrow> a#xs \<approx> a#ys"
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    13
| "\<lbrakk>xs1 \<approx> xs2; xs2 \<approx> xs3\<rbrakk> \<Longrightarrow> xs1 \<approx> xs3"
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    14
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    15
lemma list_eq_refl:
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    16
  shows "xs \<approx> xs"
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    17
  apply (induct xs)
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    18
   apply (auto intro: list_eq.intros)
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    19
  done
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    20
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    21
lemma equiv_list_eq:
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    22
  shows "EQUIV list_eq"
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    23
  unfolding EQUIV_REFL_SYM_TRANS REFL_def SYM_def TRANS_def
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    24
  apply(auto intro: list_eq.intros list_eq_refl)
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    25
  done
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    26
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    27
quotient fset = "'a list" / "list_eq"
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    28
  apply(rule equiv_list_eq)
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    29
  done
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    30
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    31
print_theorems
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    32
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    33
typ "'a fset"
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    34
thm "Rep_fset"
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    35
thm "ABS_fset_def"
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    36
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: 267
diff changeset
    37
quotient_def 
231
c643938b846a updated some definitions; had to give sometimes different names; somewhere I introduced a bug, since not everything is working anymore (needs fixing!)
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
    38
  EMPTY :: "'a fset"
c643938b846a updated some definitions; had to give sometimes different names; somewhere I introduced a bug, since not everything is working anymore (needs fixing!)
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
    39
where
c643938b846a updated some definitions; had to give sometimes different names; somewhere I introduced a bug, since not everything is working anymore (needs fixing!)
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
    40
  "EMPTY \<equiv> ([]::'a list)"
163
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    41
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    42
term Nil
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    43
term EMPTY
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    44
thm EMPTY_def
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    45
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: 267
diff changeset
    46
quotient_def 
254
77ff9624cfd6 fixed the problem with types in map
Christian Urban <urbanc@in.tum.de>
parents: 252
diff changeset
    47
  INSERT :: "'a \<Rightarrow> 'a fset \<Rightarrow> 'a fset"
231
c643938b846a updated some definitions; had to give sometimes different names; somewhere I introduced a bug, since not everything is working anymore (needs fixing!)
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
    48
where
c643938b846a updated some definitions; had to give sometimes different names; somewhere I introduced a bug, since not everything is working anymore (needs fixing!)
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
    49
  "INSERT \<equiv> op #"
163
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    50
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    51
term Cons
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    52
term INSERT
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    53
thm INSERT_def
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    54
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: 267
diff changeset
    55
quotient_def 
231
c643938b846a updated some definitions; had to give sometimes different names; somewhere I introduced a bug, since not everything is working anymore (needs fixing!)
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
    56
  FUNION :: "'a fset \<Rightarrow> 'a fset \<Rightarrow> 'a fset"
c643938b846a updated some definitions; had to give sometimes different names; somewhere I introduced a bug, since not everything is working anymore (needs fixing!)
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
    57
where
c643938b846a updated some definitions; had to give sometimes different names; somewhere I introduced a bug, since not everything is working anymore (needs fixing!)
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
    58
  "FUNION \<equiv> (op @)"
163
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    59
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    60
term append
231
c643938b846a updated some definitions; had to give sometimes different names; somewhere I introduced a bug, since not everything is working anymore (needs fixing!)
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
    61
term FUNION
c643938b846a updated some definitions; had to give sometimes different names; somewhere I introduced a bug, since not everything is working anymore (needs fixing!)
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
    62
thm FUNION_def
163
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    63
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    64
thm QUOTIENT_fset
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    65
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    66
thm QUOT_TYPE_I_fset.thm11
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    67
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    68
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    69
fun
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    70
  membship :: "'a \<Rightarrow> 'a list \<Rightarrow> bool" (infix "memb" 100)
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    71
where
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    72
  m1: "(x memb []) = False"
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    73
| m2: "(x memb (y#xs)) = ((x=y) \<or> (x memb xs))"
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    74
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    75
fun
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    76
  card1 :: "'a list \<Rightarrow> nat"
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    77
where
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    78
  card1_nil: "(card1 []) = 0"
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    79
| card1_cons: "(card1 (x # xs)) = (if (x memb xs) then (card1 xs) else (Suc (card1 xs)))"
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    80
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: 267
diff changeset
    81
quotient_def 
231
c643938b846a updated some definitions; had to give sometimes different names; somewhere I introduced a bug, since not everything is working anymore (needs fixing!)
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
    82
  CARD :: "'a fset \<Rightarrow> nat"
c643938b846a updated some definitions; had to give sometimes different names; somewhere I introduced a bug, since not everything is working anymore (needs fixing!)
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
    83
where
c643938b846a updated some definitions; had to give sometimes different names; somewhere I introduced a bug, since not everything is working anymore (needs fixing!)
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
    84
  "CARD \<equiv> card1"
163
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    85
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    86
term card1
231
c643938b846a updated some definitions; had to give sometimes different names; somewhere I introduced a bug, since not everything is working anymore (needs fixing!)
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
    87
term CARD
c643938b846a updated some definitions; had to give sometimes different names; somewhere I introduced a bug, since not everything is working anymore (needs fixing!)
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
    88
thm CARD_def
163
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    89
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    90
(* text {*
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    91
 Maybe make_const_def should require a theorem that says that the particular lifted function
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    92
 respects the relation. With it such a definition would be impossible:
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    93
 make_const_def @{binding CARD} @{term "length"} NoSyn @{typ "'a list"} @{typ "'a fset"} #> snd
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    94
*}*)
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    95
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    96
lemma card1_0:
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    97
  fixes a :: "'a list"
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    98
  shows "(card1 a = 0) = (a = [])"
214
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
    99
  by (induct a) auto
163
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   100
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   101
lemma not_mem_card1:
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   102
  fixes x :: "'a"
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   103
  fixes xs :: "'a list"
309
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   104
  shows "(~(x memb xs)) = (card1 (x # xs) = Suc (card1 xs))"
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   105
  by auto
163
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   106
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   107
lemma mem_cons:
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   108
  fixes x :: "'a"
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   109
  fixes xs :: "'a list"
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   110
  assumes a : "x memb xs"
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   111
  shows "x # xs \<approx> xs"
214
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   112
  using a by (induct xs) (auto intro: list_eq.intros )
163
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   113
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   114
lemma card1_suc:
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   115
  fixes xs :: "'a list"
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   116
  fixes n :: "nat"
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   117
  assumes c: "card1 xs = Suc n"
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   118
  shows "\<exists>a ys. ~(a memb ys) \<and> xs \<approx> (a # ys)"
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   119
  using c
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   120
apply(induct xs)
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   121
apply (metis Suc_neq_Zero card1_0)
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   122
apply (metis QUOT_TYPE_I_fset.R_trans card1_cons list_eq_refl mem_cons)
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   123
done
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   124
294
a092c0b13d83 fold_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 292
diff changeset
   125
definition
a092c0b13d83 fold_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 292
diff changeset
   126
  rsp_fold
a092c0b13d83 fold_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 292
diff changeset
   127
where
a092c0b13d83 fold_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 292
diff changeset
   128
  "rsp_fold f = ((!u v. (f u v = f v u)) \<and> (!u v w. ((f u (f v w) = f (f u v) w))))"
a092c0b13d83 fold_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 292
diff changeset
   129
163
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   130
primrec
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   131
  fold1
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   132
where
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   133
  "fold1 f (g :: 'a \<Rightarrow> 'b) (z :: 'b) [] = z"
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   134
| "fold1 f g z (a # A) =
294
a092c0b13d83 fold_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 292
diff changeset
   135
     (if rsp_fold f
163
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   136
     then (
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   137
       if (a memb A) then (fold1 f g z A) else (f (g a) (fold1 f g z A))
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   138
     ) else z)"
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   139
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   140
(* fold1_def is not usable, but: *)
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   141
thm fold1.simps
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   142
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   143
lemma fs1_strong_cases:
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   144
  fixes X :: "'a list"
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   145
  shows "(X = []) \<or> (\<exists>a. \<exists> Y. (~(a memb Y) \<and> (X \<approx> a # Y)))"
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   146
  apply (induct X)
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   147
  apply (simp)
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   148
  apply (metis QUOT_TYPE_I_fset.thm11 list_eq_refl mem_cons m1)
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   149
  done
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   150
296
eab108c8d4b7 Minor changes
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 294
diff changeset
   151
quotient_def
231
c643938b846a updated some definitions; had to give sometimes different names; somewhere I introduced a bug, since not everything is working anymore (needs fixing!)
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
   152
  IN :: "'a \<Rightarrow> 'a fset \<Rightarrow> bool"
c643938b846a updated some definitions; had to give sometimes different names; somewhere I introduced a bug, since not everything is working anymore (needs fixing!)
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
   153
where
c643938b846a updated some definitions; had to give sometimes different names; somewhere I introduced a bug, since not everything is working anymore (needs fixing!)
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
   154
  "IN \<equiv> membship"
163
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   155
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   156
term membship
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   157
term IN
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   158
thm IN_def
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   159
274
df225aa45770 simplified the quotient_def code
Christian Urban <urbanc@in.tum.de>
parents: 273
diff changeset
   160
term fold1
df225aa45770 simplified the quotient_def code
Christian Urban <urbanc@in.tum.de>
parents: 273
diff changeset
   161
quotient_def 
df225aa45770 simplified the quotient_def code
Christian Urban <urbanc@in.tum.de>
parents: 273
diff changeset
   162
  FOLD :: "('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> ('b \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'b fset \<Rightarrow> 'a"
231
c643938b846a updated some definitions; had to give sometimes different names; somewhere I introduced a bug, since not everything is working anymore (needs fixing!)
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
   163
where
c643938b846a updated some definitions; had to give sometimes different names; somewhere I introduced a bug, since not everything is working anymore (needs fixing!)
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
   164
  "FOLD \<equiv> fold1"
194
03c03e88efa9 Simplifying Int and Working on map
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 190
diff changeset
   165
03c03e88efa9 Simplifying Int and Working on map
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 190
diff changeset
   166
term fold1
03c03e88efa9 Simplifying Int and Working on map
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 190
diff changeset
   167
term fold
03c03e88efa9 Simplifying Int and Working on map
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 190
diff changeset
   168
thm fold_def
03c03e88efa9 Simplifying Int and Working on map
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 190
diff changeset
   169
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: 267
diff changeset
   170
quotient_def 
254
77ff9624cfd6 fixed the problem with types in map
Christian Urban <urbanc@in.tum.de>
parents: 252
diff changeset
   171
  fmap::"('a \<Rightarrow> 'b) \<Rightarrow> 'a fset \<Rightarrow> 'b fset"
225
9b8e039ae960 Some cleaning
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 221
diff changeset
   172
where
254
77ff9624cfd6 fixed the problem with types in map
Christian Urban <urbanc@in.tum.de>
parents: 252
diff changeset
   173
  "fmap \<equiv> map"
194
03c03e88efa9 Simplifying Int and Working on map
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 190
diff changeset
   174
03c03e88efa9 Simplifying Int and Working on map
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 190
diff changeset
   175
term map
03c03e88efa9 Simplifying Int and Working on map
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 190
diff changeset
   176
term fmap
03c03e88efa9 Simplifying Int and Working on map
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 190
diff changeset
   177
thm fmap_def
03c03e88efa9 Simplifying Int and Working on map
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 190
diff changeset
   178
290
a0be84b0c707 removed typing information from get_fun in quotient_def; *potentially* dangerous
Christian Urban <urbanc@in.tum.de>
parents: 285
diff changeset
   179
ML {* prop_of @{thm fmap_def} *}
a0be84b0c707 removed typing information from get_fun in quotient_def; *potentially* dangerous
Christian Urban <urbanc@in.tum.de>
parents: 285
diff changeset
   180
274
df225aa45770 simplified the quotient_def code
Christian Urban <urbanc@in.tum.de>
parents: 273
diff changeset
   181
ML {* val defs = @{thms EMPTY_def IN_def FUNION_def CARD_def INSERT_def fmap_def FOLD_def} *}
163
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   182
164
4f00ca4f5ef4 Stronger tactic, simpler proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 163
diff changeset
   183
lemma memb_rsp:
163
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   184
  fixes z
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   185
  assumes a: "list_eq x y"
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   186
  shows "(z memb x) = (z memb y)"
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   187
  using a by induct auto
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   188
164
4f00ca4f5ef4 Stronger tactic, simpler proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 163
diff changeset
   189
lemma ho_memb_rsp:
4f00ca4f5ef4 Stronger tactic, simpler proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 163
diff changeset
   190
  "(op = ===> (op \<approx> ===> op =)) (op memb) (op memb)"
214
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   191
  by (simp add: memb_rsp)
164
4f00ca4f5ef4 Stronger tactic, simpler proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 163
diff changeset
   192
163
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   193
lemma card1_rsp:
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   194
  fixes a b :: "'a list"
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   195
  assumes e: "a \<approx> b"
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   196
  shows "card1 a = card1 b"
214
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   197
  using e by induct (simp_all add:memb_rsp)
163
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   198
228
268a727b0f10 disambiguate ===> syntax
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 226
diff changeset
   199
lemma ho_card1_rsp: "(op \<approx> ===> op =) card1 card1"
214
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   200
  by (simp add: card1_rsp)
171
13aab4c59096 More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 168
diff changeset
   201
164
4f00ca4f5ef4 Stronger tactic, simpler proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 163
diff changeset
   202
lemma cons_rsp:
163
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   203
  fixes z
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   204
  assumes a: "xs \<approx> ys"
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   205
  shows "(z # xs) \<approx> (z # ys)"
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   206
  using a by (rule list_eq.intros(5))
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   207
164
4f00ca4f5ef4 Stronger tactic, simpler proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 163
diff changeset
   208
lemma ho_cons_rsp:
228
268a727b0f10 disambiguate ===> syntax
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 226
diff changeset
   209
  "(op = ===> op \<approx> ===> op \<approx>) op # op #"
214
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   210
  by (simp add: cons_rsp)
164
4f00ca4f5ef4 Stronger tactic, simpler proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 163
diff changeset
   211
175
f7602653dddd Preparing infrastructire for LAMBDA_PRS
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 173
diff changeset
   212
lemma append_rsp_fst:
163
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   213
  assumes a : "list_eq l1 l2"
214
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   214
  shows "(l1 @ s) \<approx> (l2 @ s)"
163
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   215
  using a
214
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   216
  by (induct) (auto intro: list_eq.intros list_eq_refl)
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   217
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   218
lemma append_end:
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   219
  shows "(e # l) \<approx> (l @ [e])"
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   220
  apply (induct l)
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   221
  apply (auto intro: list_eq.intros list_eq_refl)
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   222
  done
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   223
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   224
lemma rev_rsp:
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   225
  shows "a \<approx> rev a"
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   226
  apply (induct a)
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   227
  apply simp
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   228
  apply (rule list_eq_refl)
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   229
  apply simp_all
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   230
  apply (rule list_eq.intros(6))
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   231
  prefer 2
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   232
  apply (rule append_rsp_fst)
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   233
  apply assumption
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   234
  apply (rule append_end)
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   235
  done
163
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   236
214
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   237
lemma append_sym_rsp:
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   238
  shows "(a @ b) \<approx> (b @ a)"
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   239
  apply (rule list_eq.intros(6))
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   240
  apply (rule append_rsp_fst)
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   241
  apply (rule rev_rsp)
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   242
  apply (rule list_eq.intros(6))
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   243
  apply (rule rev_rsp)
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   244
  apply (simp)
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   245
  apply (rule append_rsp_fst)
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   246
  apply (rule list_eq.intros(3))
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   247
  apply (rule rev_rsp)
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   248
  done
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   249
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   250
lemma append_rsp:
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   251
  assumes a : "list_eq l1 r1"
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   252
  assumes b : "list_eq l2 r2 "
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   253
  shows "(l1 @ l2) \<approx> (r1 @ r2)"
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   254
  apply (rule list_eq.intros(6))
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   255
  apply (rule append_rsp_fst)
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   256
  using a apply (assumption)
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   257
  apply (rule list_eq.intros(6))
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   258
  apply (rule append_sym_rsp)
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   259
  apply (rule list_eq.intros(6))
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   260
  apply (rule append_rsp_fst)
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   261
  using b apply (assumption)
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   262
  apply (rule append_sym_rsp)
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   263
  done
175
f7602653dddd Preparing infrastructire for LAMBDA_PRS
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 173
diff changeset
   264
194
03c03e88efa9 Simplifying Int and Working on map
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 190
diff changeset
   265
lemma ho_append_rsp:
228
268a727b0f10 disambiguate ===> syntax
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 226
diff changeset
   266
  "(op \<approx> ===> op \<approx> ===> op \<approx>) op @ op @"
214
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   267
  by (simp add: append_rsp)
175
f7602653dddd Preparing infrastructire for LAMBDA_PRS
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 173
diff changeset
   268
194
03c03e88efa9 Simplifying Int and Working on map
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 190
diff changeset
   269
lemma map_rsp:
03c03e88efa9 Simplifying Int and Working on map
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 190
diff changeset
   270
  assumes a: "a \<approx> b"
03c03e88efa9 Simplifying Int and Working on map
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 190
diff changeset
   271
  shows "map f a \<approx> map f b"
03c03e88efa9 Simplifying Int and Working on map
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 190
diff changeset
   272
  using a
03c03e88efa9 Simplifying Int and Working on map
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 190
diff changeset
   273
  apply (induct)
03c03e88efa9 Simplifying Int and Working on map
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 190
diff changeset
   274
  apply(auto intro: list_eq.intros)
03c03e88efa9 Simplifying Int and Working on map
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 190
diff changeset
   275
  done
03c03e88efa9 Simplifying Int and Working on map
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 190
diff changeset
   276
03c03e88efa9 Simplifying Int and Working on map
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 190
diff changeset
   277
lemma ho_map_rsp:
294
a092c0b13d83 fold_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 292
diff changeset
   278
  "(op = ===> op \<approx> ===> op \<approx>) map map"
a092c0b13d83 fold_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 292
diff changeset
   279
  by (simp add: map_rsp)
194
03c03e88efa9 Simplifying Int and Working on map
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 190
diff changeset
   280
294
a092c0b13d83 fold_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 292
diff changeset
   281
lemma map_append:
258
93ea455b29f1 Map does not fully work yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 257
diff changeset
   282
  "(map f (a @ b)) \<approx>
93ea455b29f1 Map does not fully work yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 257
diff changeset
   283
  (map f a) @ (map f b)"
215
89a2ff3f82c7 More finshed proofs and cleaning
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 214
diff changeset
   284
 by simp (rule list_eq_refl)
194
03c03e88efa9 Simplifying Int and Working on map
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 190
diff changeset
   285
273
b82e765ca464 Lifting 'fold1.simps(2)' and some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 270
diff changeset
   286
lemma ho_fold_rsp:
294
a092c0b13d83 fold_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 292
diff changeset
   287
  "(op = ===> op = ===> op = ===> op \<approx> ===> op =) fold1 fold1"
292
bd76f0398aa9 More functionality for lifting list.cases and list.recs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 291
diff changeset
   288
  apply (auto simp add: FUN_REL_EQ)
294
a092c0b13d83 fold_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 292
diff changeset
   289
  apply (case_tac "rsp_fold x")
a092c0b13d83 fold_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 292
diff changeset
   290
  prefer 2
a092c0b13d83 fold_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 292
diff changeset
   291
  apply (erule_tac list_eq.induct)
a092c0b13d83 fold_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 292
diff changeset
   292
  apply (simp_all)
a092c0b13d83 fold_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 292
diff changeset
   293
  apply (erule_tac list_eq.induct)
a092c0b13d83 fold_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 292
diff changeset
   294
  apply (simp_all)
a092c0b13d83 fold_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 292
diff changeset
   295
  apply (auto simp add: memb_rsp rsp_fold_def)
a092c0b13d83 fold_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 292
diff changeset
   296
done
241
60acf3d3a4a0 Finding applications and duplicates filtered out in abstractions
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 239
diff changeset
   297
254
77ff9624cfd6 fixed the problem with types in map
Christian Urban <urbanc@in.tum.de>
parents: 252
diff changeset
   298
print_quotients
77ff9624cfd6 fixed the problem with types in map
Christian Urban <urbanc@in.tum.de>
parents: 252
diff changeset
   299
77ff9624cfd6 fixed the problem with types in map
Christian Urban <urbanc@in.tum.de>
parents: 252
diff changeset
   300
226
2a28e7ef3048 cleaned FSet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 225
diff changeset
   301
ML {* val qty = @{typ "'a fset"} *}
2a28e7ef3048 cleaned FSet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 225
diff changeset
   302
ML {* val rsp_thms =
273
b82e765ca464 Lifting 'fold1.simps(2)' and some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 270
diff changeset
   303
  @{thms ho_memb_rsp ho_cons_rsp ho_card1_rsp ho_map_rsp ho_append_rsp ho_fold_rsp}
226
2a28e7ef3048 cleaned FSet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 225
diff changeset
   304
  @ @{thms ho_all_prs ho_ex_prs} *}
206
1e227c9ee915 Fixed APPLY_RSP vs Cong in the InjRepAbs tactic.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 202
diff changeset
   305
239
02b14a21761a Cleaning of the interface to lift.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 233
diff changeset
   306
ML {* fun lift_thm_fset lthy t = lift_thm lthy qty "fset" rsp_thms defs t *}
348
b1f83c7a8674 More theorems lifted in the goal-directed way.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 338
diff changeset
   307
ML {* fun lift_thm_g_fset lthy t g = lift_thm_goal lthy qty "fset" rsp_thms defs t g *}
314
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 309
diff changeset
   308
226
2a28e7ef3048 cleaned FSet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 225
diff changeset
   309
ML {* lift_thm_fset @{context} @{thm m1} *}
348
b1f83c7a8674 More theorems lifted in the goal-directed way.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 338
diff changeset
   310
ML {* lift_thm_g_fset @{context} @{thm m1} @{term "IN x EMPTY = False"} *}
353
9a0e8ab42ee8 fixed the error by a temporary fix (the data of the eqivalence relation should be only its name)
Christian Urban <urbanc@in.tum.de>
parents: 350
diff changeset
   311
356
51aafebf4d06 Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 353
diff changeset
   312
ML {* lift_thm_fset @{context} @{thm m2} *}
348
b1f83c7a8674 More theorems lifted in the goal-directed way.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 338
diff changeset
   313
ML {* lift_thm_g_fset @{context} @{thm m2} @{term "IN x (INSERT y xa) = (x = y \<or> IN x xa)"} *}
356
51aafebf4d06 Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 353
diff changeset
   314
226
2a28e7ef3048 cleaned FSet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 225
diff changeset
   315
ML {* lift_thm_fset @{context} @{thm list_eq.intros(4)} *}
348
b1f83c7a8674 More theorems lifted in the goal-directed way.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 338
diff changeset
   316
ML {* lift_thm_g_fset @{context} @{thm list_eq.intros(4)} @{term "INSERT a (INSERT a x) = INSERT a x"} *}
353
9a0e8ab42ee8 fixed the error by a temporary fix (the data of the eqivalence relation should be only its name)
Christian Urban <urbanc@in.tum.de>
parents: 350
diff changeset
   317
226
2a28e7ef3048 cleaned FSet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 225
diff changeset
   318
ML {* lift_thm_fset @{context} @{thm list_eq.intros(5)} *}
348
b1f83c7a8674 More theorems lifted in the goal-directed way.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 338
diff changeset
   319
ML {* lift_thm_g_fset @{context} @{thm list_eq.intros(5)} @{term "x = xa \<Longrightarrow> INSERT a x = INSERT a xa"} *}
356
51aafebf4d06 Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 353
diff changeset
   320
226
2a28e7ef3048 cleaned FSet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 225
diff changeset
   321
ML {* lift_thm_fset @{context} @{thm card1_suc} *}
356
51aafebf4d06 Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 353
diff changeset
   322
ML {* lift_thm_g_fset @{context} @{thm card1_suc} @{term "CARD x = Suc n \<Longrightarrow> \<exists>a b. \<not> IN a b \<and> x = INSERT a b"} *}
51aafebf4d06 Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 353
diff changeset
   323
348
b1f83c7a8674 More theorems lifted in the goal-directed way.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 338
diff changeset
   324
ML {* lift_thm_fset @{context} @{thm not_mem_card1} *}
b1f83c7a8674 More theorems lifted in the goal-directed way.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 338
diff changeset
   325
ML {* lift_thm_g_fset @{context} @{thm not_mem_card1} @{term "(\<not> IN x xa) = (CARD (INSERT x xa) = Suc (CARD xa))"} *}
356
51aafebf4d06 Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 353
diff changeset
   326
348
b1f83c7a8674 More theorems lifted in the goal-directed way.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 338
diff changeset
   327
ML {* lift_thm_fset @{context} @{thm fold1.simps(2)} *}
356
51aafebf4d06 Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 353
diff changeset
   328
348
b1f83c7a8674 More theorems lifted in the goal-directed way.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 338
diff changeset
   329
(* Doesn't work with 'a, 'b, but works with 'b, 'a *)
b1f83c7a8674 More theorems lifted in the goal-directed way.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 338
diff changeset
   330
ML {* lift_thm_g_fset @{context} @{thm fold1.simps(2)} @{term "FOLD f g (z::'b) (INSERT a x) =
b1f83c7a8674 More theorems lifted in the goal-directed way.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 338
diff changeset
   331
    (if rsp_fold f then if IN a x then FOLD f g z x else f (g a) (FOLD f g z x) else z)"} *}
b1f83c7a8674 More theorems lifted in the goal-directed way.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 338
diff changeset
   332
b1f83c7a8674 More theorems lifted in the goal-directed way.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 338
diff changeset
   333
ML {* lift_thm_fset @{context} @{thm append_assoc} *}
356
51aafebf4d06 Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 353
diff changeset
   334
ML {* lift_thm_g_fset @{context} @{thm append_assoc} @{term "FUNION (FUNION x xa) xb = FUNION x (FUNION xa xb)"} *}
353
9a0e8ab42ee8 fixed the error by a temporary fix (the data of the eqivalence relation should be only its name)
Christian Urban <urbanc@in.tum.de>
parents: 350
diff changeset
   335
356
51aafebf4d06 Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 353
diff changeset
   336
ML {* lift_thm_fset @{context} @{thm map_append} *}
51aafebf4d06 Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 353
diff changeset
   337
ML {* lift_thm_g_fset @{context} @{thm map_append} @{term "fmap f (FUNION (x::'b fset) (xa::'b fset)) = FUNION (fmap f x) (fmap f xa)"} *}
349
f507f088de73 domain_type in regularizing equality
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 348
diff changeset
   338
ML {* lift_thm_fset @{context} @{thm list.induct} *}
f507f088de73 domain_type in regularizing equality
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 348
diff changeset
   339
ML {* lift_thm_g_fset @{context} @{thm list.induct} @{term "\<lbrakk>P EMPTY; \<And>a x. P x \<Longrightarrow> P (INSERT a x)\<rbrakk> \<Longrightarrow> P l"} *}
348
b1f83c7a8674 More theorems lifted in the goal-directed way.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 338
diff changeset
   340
b1f83c7a8674 More theorems lifted in the goal-directed way.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 338
diff changeset
   341
(*ML {* lift_thm_fset @{context} @{thm neq_Nil_conv} *}*)
171
13aab4c59096 More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 168
diff changeset
   342
273
b82e765ca464 Lifting 'fold1.simps(2)' and some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 270
diff changeset
   343
quotient_def
276
783d6c940e45 Experiments in Int
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 274
diff changeset
   344
  fset_rec::"'a \<Rightarrow> ('b \<Rightarrow> 'b fset \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> 'b fset \<Rightarrow> 'a"
273
b82e765ca464 Lifting 'fold1.simps(2)' and some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 270
diff changeset
   345
where
b82e765ca464 Lifting 'fold1.simps(2)' and some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 270
diff changeset
   346
  "fset_rec \<equiv> list_rec"
b82e765ca464 Lifting 'fold1.simps(2)' and some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 270
diff changeset
   347
292
bd76f0398aa9 More functionality for lifting list.cases and list.recs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 291
diff changeset
   348
quotient_def
bd76f0398aa9 More functionality for lifting list.cases and list.recs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 291
diff changeset
   349
  fset_case::"'a \<Rightarrow> ('b \<Rightarrow> 'b fset \<Rightarrow> 'a) \<Rightarrow> 'b fset \<Rightarrow> 'a"
bd76f0398aa9 More functionality for lifting list.cases and list.recs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 291
diff changeset
   350
where
bd76f0398aa9 More functionality for lifting list.cases and list.recs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 291
diff changeset
   351
  "fset_case \<equiv> list_case"
bd76f0398aa9 More functionality for lifting list.cases and list.recs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 291
diff changeset
   352
296
eab108c8d4b7 Minor changes
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 294
diff changeset
   353
(* Probably not true without additional assumptions about the function *)
292
bd76f0398aa9 More functionality for lifting list.cases and list.recs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 291
diff changeset
   354
lemma list_rec_rsp:
bd76f0398aa9 More functionality for lifting list.cases and list.recs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 291
diff changeset
   355
  "(op = ===> (op = ===> op \<approx> ===> op =) ===> op \<approx> ===> op =) list_rec list_rec"
bd76f0398aa9 More functionality for lifting list.cases and list.recs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 291
diff changeset
   356
  apply (auto simp add: FUN_REL_EQ)
296
eab108c8d4b7 Minor changes
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 294
diff changeset
   357
  apply (erule_tac list_eq.induct)
eab108c8d4b7 Minor changes
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 294
diff changeset
   358
  apply (simp_all)
292
bd76f0398aa9 More functionality for lifting list.cases and list.recs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 291
diff changeset
   359
  sorry
289
7e8617f20b59 Remaining fixes for polymorphic types. map_append now lifts properly with 'a list and 'b list.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 285
diff changeset
   360
292
bd76f0398aa9 More functionality for lifting list.cases and list.recs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 291
diff changeset
   361
lemma list_case_rsp:
bd76f0398aa9 More functionality for lifting list.cases and list.recs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 291
diff changeset
   362
  "(op = ===> (op = ===> op \<approx> ===> op =) ===> op \<approx> ===> op =) list_case list_case"
bd76f0398aa9 More functionality for lifting list.cases and list.recs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 291
diff changeset
   363
  apply (auto simp add: FUN_REL_EQ)
bd76f0398aa9 More functionality for lifting list.cases and list.recs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 291
diff changeset
   364
  sorry
bd76f0398aa9 More functionality for lifting list.cases and list.recs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 291
diff changeset
   365
bd76f0398aa9 More functionality for lifting list.cases and list.recs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 291
diff changeset
   366
bd76f0398aa9 More functionality for lifting list.cases and list.recs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 291
diff changeset
   367
ML {* val rsp_thms = @{thms list_rec_rsp list_case_rsp} @ rsp_thms *}
bd76f0398aa9 More functionality for lifting list.cases and list.recs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 291
diff changeset
   368
ML {* val defs = @{thms fset_rec_def fset_case_def} @ defs *}
bd76f0398aa9 More functionality for lifting list.cases and list.recs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 291
diff changeset
   369
bd76f0398aa9 More functionality for lifting list.cases and list.recs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 291
diff changeset
   370
ML {* fun lift_thm_fset lthy t = lift_thm lthy qty "fset" rsp_thms defs t *}
356
51aafebf4d06 Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 353
diff changeset
   371
ML {* fun lift_thm_g_fset lthy t g = lift_thm_goal lthy qty "fset" rsp_thms defs t g *}
51aafebf4d06 Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 353
diff changeset
   372
51aafebf4d06 Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 353
diff changeset
   373
thm list.recs(2)
51aafebf4d06 Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 353
diff changeset
   374
ML {* lift_thm_fset @{context} @{thm list.recs(2)} *}
51aafebf4d06 Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 353
diff changeset
   375
ML {* lift_thm_g_fset @{context} @{thm list.recs(2)} @{term "fset_rec (f1::'t) x (INSERT a xa) = x a xa (fset_rec f1 x xa)"} *}
51aafebf4d06 Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 353
diff changeset
   376
51aafebf4d06 Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 353
diff changeset
   377
ML {* val (rty, rel, rel_refl, rel_eqv) = lookup_quot_data @{context} qty *}
51aafebf4d06 Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 353
diff changeset
   378
ML {* val (trans2, reps_same, absrep, quot) = lookup_quot_thms @{context} "fset"; *}
51aafebf4d06 Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 353
diff changeset
   379
ML {* val consts = lookup_quot_consts defs *}
51aafebf4d06 Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 353
diff changeset
   380
51aafebf4d06 Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 353
diff changeset
   381
ML {* val gl = @{term "fset_rec (f1::'t) x (INSERT a xa) = x a xa (fset_rec f1 x xa)"} *}
51aafebf4d06 Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 353
diff changeset
   382
ML {* val t_a = atomize_thm @{thm list.recs(2)} *}
51aafebf4d06 Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 353
diff changeset
   383
ML {* val qtrm = atomize_goal @{theory} gl *}
51aafebf4d06 Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 353
diff changeset
   384
ML {* val rg = cterm_of @{theory}(Syntax.check_term @{context} (REGULARIZE_trm @{context} (prop_of t_a) qtrm)) *}
51aafebf4d06 Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 353
diff changeset
   385
ML {* val rg2 = cterm_of @{theory}(my_reg @{context} rel rty (prop_of t_a)) *}
51aafebf4d06 Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 353
diff changeset
   386
51aafebf4d06 Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 353
diff changeset
   387
ML {* val t_r = regularize_goal @{context} t_a rel_eqv rel_refl qtrm *}
51aafebf4d06 Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 353
diff changeset
   388
ML {* val t_r = regularize t_a rty rel rel_eqv rel_refl @{context} *}
51aafebf4d06 Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 353
diff changeset
   389
51aafebf4d06 Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 353
diff changeset
   390
ML {* val rg = cterm_of @{theory}(Syntax.check_term @{context} (inj_REPABS @{context} ((prop_of t_r), qtrm))) *}
51aafebf4d06 Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 353
diff changeset
   391
ML {* val rg2 = cterm_of @{theory} (build_repabs_term @{context} t_r consts rty qty) *}
51aafebf4d06 Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 353
diff changeset
   392
ML {* val t_t = repabs_goal @{context} t_r rty quot rel_refl trans2 rsp_thms qtrm *}
51aafebf4d06 Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 353
diff changeset
   393
ML {* val t_t = repabs @{context} t_r consts rty qty quot rel_refl trans2 rsp_thms *}
51aafebf4d06 Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 353
diff changeset
   394
ML {*
51aafebf4d06 Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 353
diff changeset
   395
  val lthy = @{context}
51aafebf4d06 Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 353
diff changeset
   396
  val (alls, exs) = findallex lthy rty qty (prop_of t_a);
51aafebf4d06 Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 353
diff changeset
   397
  val allthms = map (make_allex_prs_thm lthy quot @{thm FORALL_PRS}) alls
51aafebf4d06 Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 353
diff changeset
   398
  val exthms = map (make_allex_prs_thm lthy quot @{thm EXISTS_PRS}) exs
51aafebf4d06 Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 353
diff changeset
   399
  val t_a = MetaSimplifier.rewrite_rule (allthms @ exthms) t_t
51aafebf4d06 Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 353
diff changeset
   400
  val abs = findabs rty (prop_of t_a);
51aafebf4d06 Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 353
diff changeset
   401
  val aps = findaps rty (prop_of t_a);
51aafebf4d06 Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 353
diff changeset
   402
  val app_prs_thms = map (applic_prs lthy rty qty absrep) aps;
51aafebf4d06 Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 353
diff changeset
   403
  val lam_prs_thms = map (make_simp_prs_thm lthy quot @{thm LAMBDA_PRS}) abs;
51aafebf4d06 Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 353
diff changeset
   404
  val t_l = repeat_eqsubst_thm lthy (lam_prs_thms @ app_prs_thms) t_a;
51aafebf4d06 Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 353
diff changeset
   405
  val defs_sym = flat (map (add_lower_defs lthy) defs);
51aafebf4d06 Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 353
diff changeset
   406
  val defs_sym_eq = map (fn x => eq_reflection OF [x]) defs_sym;
51aafebf4d06 Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 353
diff changeset
   407
  val t_id = simp_ids lthy t_l;
51aafebf4d06 Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 353
diff changeset
   408
  val t_d0 = MetaSimplifier.rewrite_rule defs_sym_eq t_id;
51aafebf4d06 Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 353
diff changeset
   409
  val t_d = repeat_eqsubst_thm lthy defs_sym t_d0;
51aafebf4d06 Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 353
diff changeset
   410
  val t_r = MetaSimplifier.rewrite_rule [reps_same] t_d;
51aafebf4d06 Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 353
diff changeset
   411
  val t_rv = ObjectLogic.rulify t_r
51aafebf4d06 Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 353
diff changeset
   412
51aafebf4d06 Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 353
diff changeset
   413
*}
292
bd76f0398aa9 More functionality for lifting list.cases and list.recs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 291
diff changeset
   414
300
c6a9b4e4d548 Fixes for the other get_fun implementation.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 296
diff changeset
   415
356
51aafebf4d06 Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 353
diff changeset
   416
51aafebf4d06 Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 353
diff changeset
   417
51aafebf4d06 Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 353
diff changeset
   418
51aafebf4d06 Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 353
diff changeset
   419
292
bd76f0398aa9 More functionality for lifting list.cases and list.recs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 291
diff changeset
   420
ML {* map (lift_thm_fset @{context}) @{thms list.cases} *}
248
6ed87b3d358c Finally merged the code of the versions of regularize and tested examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 244
diff changeset
   421
356
51aafebf4d06 Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 353
diff changeset
   422
51aafebf4d06 Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 353
diff changeset
   423
51aafebf4d06 Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 353
diff changeset
   424
348
b1f83c7a8674 More theorems lifted in the goal-directed way.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 338
diff changeset
   425
ML {* atomize_thm @{thm m1} *}
b1f83c7a8674 More theorems lifted in the goal-directed way.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 338
diff changeset
   426
ML {* cterm_of @{theory} (atomize_goal @{theory} @{term "IN x EMPTY = False"}) *}
b1f83c7a8674 More theorems lifted in the goal-directed way.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 338
diff changeset
   427
ML {* lift_thm_fset @{context} @{thm m1} *}
b1f83c7a8674 More theorems lifted in the goal-directed way.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 338
diff changeset
   428
(* ML {* lift_thm_g_fset @{context} @{thm m1} @{term "IN x EMPTY = False"}) *} *)
b1f83c7a8674 More theorems lifted in the goal-directed way.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 338
diff changeset
   429
b1f83c7a8674 More theorems lifted in the goal-directed way.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 338
diff changeset
   430
304
e741c735b867 Atomizing a "goal" theorems.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 300
diff changeset
   431
lemma list_induct_part:
e741c735b867 Atomizing a "goal" theorems.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 300
diff changeset
   432
  assumes a: "P (x :: 'a list) ([] :: 'a list)"
e741c735b867 Atomizing a "goal" theorems.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 300
diff changeset
   433
  assumes b: "\<And>e t. P x t \<Longrightarrow> P x (e # t)"
e741c735b867 Atomizing a "goal" theorems.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 300
diff changeset
   434
  shows "P x l"
e741c735b867 Atomizing a "goal" theorems.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 300
diff changeset
   435
  apply (rule_tac P="P x" in list.induct)
e741c735b867 Atomizing a "goal" theorems.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 300
diff changeset
   436
  apply (rule a)
e741c735b867 Atomizing a "goal" theorems.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 300
diff changeset
   437
  apply (rule b)
e741c735b867 Atomizing a "goal" theorems.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 300
diff changeset
   438
  apply (assumption)
e741c735b867 Atomizing a "goal" theorems.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 300
diff changeset
   439
  done
273
b82e765ca464 Lifting 'fold1.simps(2)' and some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 270
diff changeset
   440
292
bd76f0398aa9 More functionality for lifting list.cases and list.recs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 291
diff changeset
   441
bd76f0398aa9 More functionality for lifting list.cases and list.recs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 291
diff changeset
   442
(* Construction site starts here *)
bd76f0398aa9 More functionality for lifting list.cases and list.recs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 291
diff changeset
   443
bd76f0398aa9 More functionality for lifting list.cases and list.recs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 291
diff changeset
   444
273
b82e765ca464 Lifting 'fold1.simps(2)' and some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 270
diff changeset
   445
ML {* val consts = lookup_quot_consts defs *}
b82e765ca464 Lifting 'fold1.simps(2)' and some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 270
diff changeset
   446
ML {* val (rty, rel, rel_refl, rel_eqv) = lookup_quot_data @{context} qty *}
b82e765ca464 Lifting 'fold1.simps(2)' and some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 270
diff changeset
   447
ML {* val (trans2, reps_same, absrep, quot) = lookup_quot_thms @{context} "fset" *}
b82e765ca464 Lifting 'fold1.simps(2)' and some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 270
diff changeset
   448
296
eab108c8d4b7 Minor changes
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 294
diff changeset
   449
thm list.recs(2)
304
e741c735b867 Atomizing a "goal" theorems.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 300
diff changeset
   450
ML {* val t_a = atomize_thm @{thm list_induct_part} *}
334
5a7024be9083 code review with Cezary
Christian Urban <urbanc@in.tum.de>
parents: 333
diff changeset
   451
5a7024be9083 code review with Cezary
Christian Urban <urbanc@in.tum.de>
parents: 333
diff changeset
   452
285
8ebdef196fd5 Infrastructure for polymorphic types
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 276
diff changeset
   453
(* prove {* build_regularize_goal t_a rty rel @{context}  *}
8ebdef196fd5 Infrastructure for polymorphic types
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 276
diff changeset
   454
 ML_prf {*  fun tac ctxt = FIRST' [
251
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 248
diff changeset
   455
      rtac rel_refl,
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 248
diff changeset
   456
      atac,
285
8ebdef196fd5 Infrastructure for polymorphic types
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 276
diff changeset
   457
      rtac @{thm universal_twice},
8ebdef196fd5 Infrastructure for polymorphic types
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 276
diff changeset
   458
      (rtac @{thm impI} THEN' atac),
8ebdef196fd5 Infrastructure for polymorphic types
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 276
diff changeset
   459
      rtac @{thm implication_twice},
334
5a7024be9083 code review with Cezary
Christian Urban <urbanc@in.tum.de>
parents: 333
diff changeset
   460
      //comented out  rtac @{thm equality_twice}, //
285
8ebdef196fd5 Infrastructure for polymorphic types
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 276
diff changeset
   461
      EqSubst.eqsubst_tac ctxt [0]
239
02b14a21761a Cleaning of the interface to lift.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 233
diff changeset
   462
        [(@{thm equiv_res_forall} OF [rel_eqv]),
285
8ebdef196fd5 Infrastructure for polymorphic types
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 276
diff changeset
   463
         (@{thm equiv_res_exists} OF [rel_eqv])],
252
e30997c88050 Regularize for equalities and a better tactic. "alpha.cases" now lifts.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 251
diff changeset
   464
      (rtac @{thm impI} THEN' (asm_full_simp_tac (Simplifier.context ctxt HOL_ss)) THEN' rtac rel_refl),
e30997c88050 Regularize for equalities and a better tactic. "alpha.cases" now lifts.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 251
diff changeset
   465
      (rtac @{thm RIGHT_RES_FORALL_REGULAR})
285
8ebdef196fd5 Infrastructure for polymorphic types
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 276
diff changeset
   466
     ]; *}
252
e30997c88050 Regularize for equalities and a better tactic. "alpha.cases" now lifts.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 251
diff changeset
   467
  apply (atomize(full))
285
8ebdef196fd5 Infrastructure for polymorphic types
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 276
diff changeset
   468
  apply (tactic {* REPEAT_ALL_NEW (tac @{context}) 1 *})
334
5a7024be9083 code review with Cezary
Christian Urban <urbanc@in.tum.de>
parents: 333
diff changeset
   469
  done  *)
305
d7b60303adb8 Removed 'Toplevel.program' for polyml 5.3
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 304
diff changeset
   470
ML {* val t_r = regularize t_a rty rel rel_eqv rel_refl @{context} *}
273
b82e765ca464 Lifting 'fold1.simps(2)' and some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 270
diff changeset
   471
ML {*
285
8ebdef196fd5 Infrastructure for polymorphic types
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 276
diff changeset
   472
  val rt = build_repabs_term @{context} t_r consts rty qty
8ebdef196fd5 Infrastructure for polymorphic types
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 276
diff changeset
   473
  val rg = Logic.mk_equals ((Thm.prop_of t_r), rt);
172
da38ce2711a6 More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 171
diff changeset
   474
*}
300
c6a9b4e4d548 Fixes for the other get_fun implementation.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 296
diff changeset
   475
prove {* Syntax.check_term @{context} rg *}
309
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   476
ML_prf {* fun r_mk_comb_tac_fset lthy = r_mk_comb_tac lthy rty quot rel_refl trans2 rsp_thms *}
172
da38ce2711a6 More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 171
diff changeset
   477
apply(atomize(full))
194
03c03e88efa9 Simplifying Int and Working on map
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 190
diff changeset
   478
apply (tactic {* REPEAT_ALL_NEW (r_mk_comb_tac_fset @{context}) 1 *})
273
b82e765ca464 Lifting 'fold1.simps(2)' and some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 270
diff changeset
   479
done
305
d7b60303adb8 Removed 'Toplevel.program' for polyml 5.3
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 304
diff changeset
   480
ML {*
d7b60303adb8 Removed 'Toplevel.program' for polyml 5.3
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 304
diff changeset
   481
val t_t = repabs @{context} t_r consts rty qty quot rel_refl trans2 rsp_thms
172
da38ce2711a6 More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 171
diff changeset
   482
*}
241
60acf3d3a4a0 Finding applications and duplicates filtered out in abstractions
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 239
diff changeset
   483
292
bd76f0398aa9 More functionality for lifting list.cases and list.recs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 291
diff changeset
   484
ML {* val abs = findabs rty (prop_of (t_a)) *}
bd76f0398aa9 More functionality for lifting list.cases and list.recs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 291
diff changeset
   485
ML {* val aps = findaps rty (prop_of (t_a)) *}
285
8ebdef196fd5 Infrastructure for polymorphic types
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 276
diff changeset
   486
ML {* val lam_prs_thms = map (make_simp_prs_thm @{context} quot @{thm LAMBDA_PRS}) abs *}
8ebdef196fd5 Infrastructure for polymorphic types
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 276
diff changeset
   487
ML {* val app_prs_thms = map (applic_prs @{context} rty qty absrep) aps *}
289
7e8617f20b59 Remaining fixes for polymorphic types. map_append now lifts properly with 'a list and 'b list.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 285
diff changeset
   488
ML {* val lam_prs_thms = map Thm.varifyT lam_prs_thms *}
292
bd76f0398aa9 More functionality for lifting list.cases and list.recs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 291
diff changeset
   489
ML {* t_t *}
289
7e8617f20b59 Remaining fixes for polymorphic types. map_append now lifts properly with 'a list and 'b list.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 285
diff changeset
   490
ML {* val (alls, exs) = findallex @{context} rty qty (prop_of t_a); *}
285
8ebdef196fd5 Infrastructure for polymorphic types
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 276
diff changeset
   491
ML {* val allthms = map (make_allex_prs_thm @{context} quot @{thm FORALL_PRS}) alls *}
304
e741c735b867 Atomizing a "goal" theorems.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 300
diff changeset
   492
ML {* val t_l0 = repeat_eqsubst_thm @{context} (app_prs_thms) t_t *}
e741c735b867 Atomizing a "goal" theorems.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 300
diff changeset
   493
ML app_prs_thms
296
eab108c8d4b7 Minor changes
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 294
diff changeset
   494
ML {* val t_l = repeat_eqsubst_thm @{context} (lam_prs_thms) t_l0 *}
304
e741c735b867 Atomizing a "goal" theorems.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 300
diff changeset
   495
ML lam_prs_thms
e741c735b867 Atomizing a "goal" theorems.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 300
diff changeset
   496
ML {* val t_id = simp_ids @{context} t_l *}
e741c735b867 Atomizing a "goal" theorems.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 300
diff changeset
   497
thm INSERT_def
292
bd76f0398aa9 More functionality for lifting list.cases and list.recs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 291
diff changeset
   498
ML {* val defs_sym = flat (map (add_lower_defs @{context}) defs) *}
bd76f0398aa9 More functionality for lifting list.cases and list.recs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 291
diff changeset
   499
ML {* val t_d = repeat_eqsubst_thm @{context} defs_sym t_id *}
304
e741c735b867 Atomizing a "goal" theorems.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 300
diff changeset
   500
ML allthms
e741c735b867 Atomizing a "goal" theorems.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 300
diff changeset
   501
thm FORALL_PRS
309
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   502
ML {* val t_al = MetaSimplifier.rewrite_rule (allthms) t_d *}
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   503
ML {* val t_s = MetaSimplifier.rewrite_rule @{thms QUOT_TYPE_I_fset.REPS_same} t_al *}
304
e741c735b867 Atomizing a "goal" theorems.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 300
diff changeset
   504
ML {* ObjectLogic.rulify t_s *}
e741c735b867 Atomizing a "goal" theorems.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 300
diff changeset
   505
e741c735b867 Atomizing a "goal" theorems.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 300
diff changeset
   506
ML {* val gl = @{term "P (x :: 'a list) (EMPTY :: 'a fset) \<Longrightarrow> (\<And>e t. P x t \<Longrightarrow> P x (INSERT e t)) \<Longrightarrow> P x l"} *}
338
62b188959c8a Move atomize_goal to QuotMain
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 337
diff changeset
   507
ML {* val gla = atomize_goal @{theory} gl *}
309
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   508
338
62b188959c8a Move atomize_goal to QuotMain
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 337
diff changeset
   509
prove t_r: {* mk_REGULARIZE_goal @{context} (prop_of t_a) gla *}
332
87f5fbebd6d5 Fixes for recent changes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 317
diff changeset
   510
309
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   511
ML_prf {*  fun tac ctxt = FIRST' [
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   512
      rtac rel_refl,
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   513
      atac,
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   514
      rtac @{thm universal_twice},
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   515
      (rtac @{thm impI} THEN' atac),
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   516
      rtac @{thm implication_twice},
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   517
      (*rtac @{thm equality_twice},*)
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   518
      EqSubst.eqsubst_tac ctxt [0]
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   519
        [(@{thm equiv_res_forall} OF [rel_eqv]),
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   520
         (@{thm equiv_res_exists} OF [rel_eqv])],
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   521
      (rtac @{thm impI} THEN' (asm_full_simp_tac (Simplifier.context ctxt HOL_ss)) THEN' rtac rel_refl),
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   522
      (rtac @{thm RIGHT_RES_FORALL_REGULAR})
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   523
     ]; *}
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   524
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   525
  apply (atomize(full))
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   526
  apply (tactic {* REPEAT_ALL_NEW (tac @{context}) 1 *})
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   527
  done
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   528
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   529
ML {* val t_r = @{thm t_r} OF [t_a] *}
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   530
338
62b188959c8a Move atomize_goal to QuotMain
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 337
diff changeset
   531
ML {* val ttt = mk_inj_REPABS_goal @{context} (prop_of t_r, gla) *}
309
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   532
ML {* val si = simp_ids_trm (cterm_of @{theory} ttt) *}
337
553bef083318 Removed second implementation of Regularize/Inject from FSet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 335
diff changeset
   533
prove t_t: {* term_of si *}
309
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   534
ML_prf {* fun r_mk_comb_tac_fset lthy = r_mk_comb_tac lthy rty quot rel_refl trans2 rsp_thms *}
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   535
apply(atomize(full))
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   536
apply (tactic {* (APPLY_RSP_TAC rty @{context}) 1 *})
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   537
apply (rule FUN_QUOTIENT)
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   538
apply (rule FUN_QUOTIENT)
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   539
apply (rule IDENTITY_QUOTIENT)
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   540
apply (rule FUN_QUOTIENT)
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   541
apply (rule QUOTIENT_fset)
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   542
apply (rule IDENTITY_QUOTIENT)
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   543
apply (rule IDENTITY_QUOTIENT)
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   544
apply (rule IDENTITY_QUOTIENT)
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   545
apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   546
apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   547
apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   548
apply (tactic {* (APPLY_RSP_TAC rty @{context}) 1 *})
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   549
apply (rule IDENTITY_QUOTIENT)
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   550
apply (rule IDENTITY_QUOTIENT)
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   551
apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   552
apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   553
apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   554
apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   555
apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   556
apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   557
apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   558
apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   559
apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   560
apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   561
apply (tactic {* (APPLY_RSP_TAC rty @{context}) 1 *})
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   562
apply (rule IDENTITY_QUOTIENT)
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   563
apply (rule FUN_QUOTIENT)
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   564
apply (rule QUOTIENT_fset)
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   565
apply (rule IDENTITY_QUOTIENT)
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   566
apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   567
apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   568
apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   569
apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   570
apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   571
apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   572
apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   573
apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   574
apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   575
apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   576
apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   577
apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   578
apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   579
apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   580
apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   581
apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   582
apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   583
apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
317
d3c7f6d19c7f Still don't know how to do the proof automatically.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 314
diff changeset
   584
apply (tactic {* instantiate_tac @{thm APPLY_RSP2} @{context} 1 *})
d3c7f6d19c7f Still don't know how to do the proof automatically.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 314
diff changeset
   585
apply (tactic {* instantiate_tac @{thm APPLY_RSP2} @{context} 1 *})
d3c7f6d19c7f Still don't know how to do the proof automatically.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 314
diff changeset
   586
apply (tactic {* (instantiate_tac @{thm REP_ABS_RSP(1)} @{context} THEN' (RANGE [quotient_tac quot])) 1 *})
d3c7f6d19c7f Still don't know how to do the proof automatically.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 314
diff changeset
   587
apply assumption
d3c7f6d19c7f Still don't know how to do the proof automatically.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 314
diff changeset
   588
apply (rule refl)
309
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   589
apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   590
apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
317
d3c7f6d19c7f Still don't know how to do the proof automatically.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 314
diff changeset
   591
apply (tactic {* instantiate_tac @{thm APPLY_RSP2} @{context} 1 *})
d3c7f6d19c7f Still don't know how to do the proof automatically.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 314
diff changeset
   592
apply (tactic {* instantiate_tac @{thm APPLY_RSP2} @{context} 1 *})
d3c7f6d19c7f Still don't know how to do the proof automatically.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 314
diff changeset
   593
apply (tactic {* (instantiate_tac @{thm REP_ABS_RSP(1)} @{context} THEN' (RANGE [quotient_tac quot])) 1 *})
309
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   594
apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   595
apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   596
apply (tactic {* REPEAT_ALL_NEW (r_mk_comb_tac_fset @{context}) 1 *})
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   597
apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
317
d3c7f6d19c7f Still don't know how to do the proof automatically.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 314
diff changeset
   598
apply (tactic {* instantiate_tac @{thm APPLY_RSP2} @{context} 1 *})
d3c7f6d19c7f Still don't know how to do the proof automatically.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 314
diff changeset
   599
apply (tactic {* (instantiate_tac @{thm REP_ABS_RSP(1)} @{context} THEN' (RANGE [quotient_tac quot])) 1 *})
309
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   600
apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   601
apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   602
apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   603
apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   604
done
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   605
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   606
thm t_t
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   607
ML {* val t_t = @{thm Pure.equal_elim_rule1} OF [@{thm t_t}, t_r] *}
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   608
ML {* val t_l = repeat_eqsubst_thm @{context} (lam_prs_thms) t_t *}
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   609
ML {* val t_d = repeat_eqsubst_thm @{context} defs_sym t_l *}
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   610
ML {* val t_al = MetaSimplifier.rewrite_rule (allthms) t_d *}
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   611
ML {* val t_s = MetaSimplifier.rewrite_rule @{thms QUOT_TYPE_I_fset.REPS_same} t_al *}
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   612
178
945786a68ec6 Finally lifted induction, with some manually added simplification lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 176
diff changeset
   613
163
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   614
ML {*
226
2a28e7ef3048 cleaned FSet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 225
diff changeset
   615
  fun lift_thm_fset_note name thm lthy =
163
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   616
    let
226
2a28e7ef3048 cleaned FSet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 225
diff changeset
   617
      val lifted_thm = lift_thm_fset lthy thm;
163
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   618
      val (_, lthy2) = note (name, lifted_thm) lthy;
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   619
    in
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   620
      lthy2
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   621
    end;
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   622
*}
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   623
226
2a28e7ef3048 cleaned FSet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 225
diff changeset
   624
local_setup {*
2a28e7ef3048 cleaned FSet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 225
diff changeset
   625
  lift_thm_fset_note @{binding "m1l"} @{thm m1} #>
2a28e7ef3048 cleaned FSet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 225
diff changeset
   626
  lift_thm_fset_note @{binding "m2l"} @{thm m2} #>
2a28e7ef3048 cleaned FSet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 225
diff changeset
   627
  lift_thm_fset_note @{binding "leqi4l"} @{thm list_eq.intros(4)} #>
2a28e7ef3048 cleaned FSet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 225
diff changeset
   628
  lift_thm_fset_note @{binding "leqi5l"} @{thm list_eq.intros(5)}
163
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   629
*}
226
2a28e7ef3048 cleaned FSet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 225
diff changeset
   630
thm m1l
2a28e7ef3048 cleaned FSet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 225
diff changeset
   631
thm m2l
2a28e7ef3048 cleaned FSet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 225
diff changeset
   632
thm leqi4l
2a28e7ef3048 cleaned FSet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 225
diff changeset
   633
thm leqi5l
163
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   634
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   635
end