FSet.thy
author Christian Urban <urbanc@in.tum.de>
Fri, 04 Dec 2009 15:19:39 +0100
changeset 531 3feed4dbfa45
parent 526 7ba2fc25c6a3
child 529 6348c2a57ec2
permissions -rw-r--r--
merge
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"
451
586e3dc4afdb Added 'TRY' to refl in clean_tac to get as far as possible. Removed unnecessary [quot_rsp] in FSet. Added necessary [quot_rsp] and one lifted thm in LamEx.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 450
diff changeset
    17
  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
    18
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    19
lemma equiv_list_eq:
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    20
  shows "EQUIV list_eq"
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    21
  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
    22
  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
    23
  done
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    24
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    25
quotient fset = "'a list" / "list_eq"
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    26
  apply(rule equiv_list_eq)
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    27
  done
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    28
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    29
print_theorems
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
typ "'a fset"
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    32
thm "Rep_fset"
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    33
thm "ABS_fset_def"
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    34
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
    35
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
    36
  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
    37
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
    38
  "EMPTY \<equiv> ([]::'a list)"
163
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    39
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    40
term Nil
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    41
term EMPTY
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    42
thm EMPTY_def
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    43
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
    44
quotient_def 
254
77ff9624cfd6 fixed the problem with types in map
Christian Urban <urbanc@in.tum.de>
parents: 252
diff changeset
    45
  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
    46
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
    47
  "INSERT \<equiv> op #"
163
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    48
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    49
term Cons
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    50
term INSERT
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    51
thm INSERT_def
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    52
268
4d58c02289ca simplified the quotient_def code; type of the defined constant must now be given; for-part eliminated
Christian Urban <urbanc@in.tum.de>
parents: 267
diff changeset
    53
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
    54
  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
    55
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
    56
  "FUNION \<equiv> (op @)"
163
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    57
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    58
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
    59
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
    60
thm FUNION_def
163
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    61
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    62
thm QUOTIENT_fset
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 QUOT_TYPE_I_fset.thm11
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
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    67
fun
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    68
  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
    69
where
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    70
  m1: "(x memb []) = False"
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    71
| 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
    72
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    73
fun
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    74
  card1 :: "'a list \<Rightarrow> nat"
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    75
where
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    76
  card1_nil: "(card1 []) = 0"
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    77
| 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
    78
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
    79
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
    80
  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
    81
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
    82
  "CARD \<equiv> card1"
163
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    83
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    84
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
    85
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
    86
thm CARD_def
163
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    87
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    88
(* text {*
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    89
 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
    90
 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
    91
 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
    92
*}*)
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    93
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    94
lemma card1_0:
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    95
  fixes a :: "'a list"
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    96
  shows "(card1 a = 0) = (a = [])"
214
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
    97
  by (induct a) auto
163
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    98
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    99
lemma not_mem_card1:
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   100
  fixes x :: "'a"
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   101
  fixes xs :: "'a list"
309
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   102
  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
   103
  by auto
163
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   104
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   105
lemma mem_cons:
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   106
  fixes x :: "'a"
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   107
  fixes xs :: "'a list"
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   108
  assumes a : "x memb xs"
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   109
  shows "x # xs \<approx> xs"
214
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   110
  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
   111
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   112
lemma card1_suc:
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   113
  fixes xs :: "'a list"
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   114
  fixes n :: "nat"
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   115
  assumes c: "card1 xs = Suc n"
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   116
  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
   117
  using c
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   118
apply(induct xs)
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   119
apply (metis Suc_neq_Zero card1_0)
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   120
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
   121
done
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   122
294
a092c0b13d83 fold_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 292
diff changeset
   123
definition
a092c0b13d83 fold_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 292
diff changeset
   124
  rsp_fold
a092c0b13d83 fold_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 292
diff changeset
   125
where
a092c0b13d83 fold_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 292
diff changeset
   126
  "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
   127
163
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   128
primrec
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   129
  fold1
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   130
where
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   131
  "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
   132
| "fold1 f g z (a # A) =
294
a092c0b13d83 fold_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 292
diff changeset
   133
     (if rsp_fold f
163
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   134
     then (
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   135
       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
   136
     ) else z)"
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   137
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   138
lemma fs1_strong_cases:
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   139
  fixes X :: "'a list"
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   140
  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
   141
  apply (induct X)
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   142
  apply (simp)
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   143
  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
   144
  done
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   145
296
eab108c8d4b7 Minor changes
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 294
diff changeset
   146
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
   147
  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
   148
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
   149
  "IN \<equiv> membship"
163
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   150
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   151
term membship
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   152
term IN
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   153
thm IN_def
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   154
274
df225aa45770 simplified the quotient_def code
Christian Urban <urbanc@in.tum.de>
parents: 273
diff changeset
   155
term fold1
df225aa45770 simplified the quotient_def code
Christian Urban <urbanc@in.tum.de>
parents: 273
diff changeset
   156
quotient_def 
df225aa45770 simplified the quotient_def code
Christian Urban <urbanc@in.tum.de>
parents: 273
diff changeset
   157
  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
   158
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
   159
  "FOLD \<equiv> fold1"
194
03c03e88efa9 Simplifying Int and Working on map
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 190
diff changeset
   160
03c03e88efa9 Simplifying Int and Working on map
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 190
diff changeset
   161
term fold1
03c03e88efa9 Simplifying Int and Working on map
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 190
diff changeset
   162
term fold
03c03e88efa9 Simplifying Int and Working on map
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 190
diff changeset
   163
thm fold_def
03c03e88efa9 Simplifying Int and Working on map
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 190
diff changeset
   164
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
   165
quotient_def 
254
77ff9624cfd6 fixed the problem with types in map
Christian Urban <urbanc@in.tum.de>
parents: 252
diff changeset
   166
  fmap::"('a \<Rightarrow> 'b) \<Rightarrow> 'a fset \<Rightarrow> 'b fset"
225
9b8e039ae960 Some cleaning
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 221
diff changeset
   167
where
254
77ff9624cfd6 fixed the problem with types in map
Christian Urban <urbanc@in.tum.de>
parents: 252
diff changeset
   168
  "fmap \<equiv> map"
194
03c03e88efa9 Simplifying Int and Working on map
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 190
diff changeset
   169
03c03e88efa9 Simplifying Int and Working on map
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 190
diff changeset
   170
term map
03c03e88efa9 Simplifying Int and Working on map
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 190
diff changeset
   171
term fmap
03c03e88efa9 Simplifying Int and Working on map
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 190
diff changeset
   172
thm fmap_def
03c03e88efa9 Simplifying Int and Working on map
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 190
diff changeset
   173
451
586e3dc4afdb Added 'TRY' to refl in clean_tac to get as far as possible. Removed unnecessary [quot_rsp] in FSet. Added necessary [quot_rsp] and one lifted thm in LamEx.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 450
diff changeset
   174
lemma memb_rsp:
163
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   175
  fixes z
450
2dc708ddb93a introduced a global list of respectfulness lemmas; the attribute is [quot_rsp]
Christian Urban <urbanc@in.tum.de>
parents: 448
diff changeset
   176
  assumes a: "x \<approx> y"
163
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   177
  shows "(z memb x) = (z memb y)"
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   178
  using a by induct auto
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   179
506
91c374abde06 removed quot argument...not all examples work anymore
Christian Urban <urbanc@in.tum.de>
parents: 498
diff changeset
   180
lemma ho_memb_rsp[quotient_rsp]:
164
4f00ca4f5ef4 Stronger tactic, simpler proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 163
diff changeset
   181
  "(op = ===> (op \<approx> ===> op =)) (op memb) (op memb)"
214
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   182
  by (simp add: memb_rsp)
164
4f00ca4f5ef4 Stronger tactic, simpler proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 163
diff changeset
   183
451
586e3dc4afdb Added 'TRY' to refl in clean_tac to get as far as possible. Removed unnecessary [quot_rsp] in FSet. Added necessary [quot_rsp] and one lifted thm in LamEx.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 450
diff changeset
   184
lemma card1_rsp:
163
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   185
  fixes a b :: "'a list"
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   186
  assumes e: "a \<approx> b"
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   187
  shows "card1 a = card1 b"
214
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   188
  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
   189
506
91c374abde06 removed quot argument...not all examples work anymore
Christian Urban <urbanc@in.tum.de>
parents: 498
diff changeset
   190
lemma ho_card1_rsp[quotient_rsp]: 
450
2dc708ddb93a introduced a global list of respectfulness lemmas; the attribute is [quot_rsp]
Christian Urban <urbanc@in.tum.de>
parents: 448
diff changeset
   191
  "(op \<approx> ===> op =) card1 card1"
214
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   192
  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
   193
506
91c374abde06 removed quot argument...not all examples work anymore
Christian Urban <urbanc@in.tum.de>
parents: 498
diff changeset
   194
lemma cons_rsp[quotient_rsp]:
163
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   195
  fixes z
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   196
  assumes a: "xs \<approx> ys"
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   197
  shows "(z # xs) \<approx> (z # ys)"
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   198
  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
   199
506
91c374abde06 removed quot argument...not all examples work anymore
Christian Urban <urbanc@in.tum.de>
parents: 498
diff changeset
   200
lemma ho_cons_rsp[quotient_rsp]:
228
268a727b0f10 disambiguate ===> syntax
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 226
diff changeset
   201
  "(op = ===> op \<approx> ===> op \<approx>) op # op #"
214
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   202
  by (simp add: cons_rsp)
164
4f00ca4f5ef4 Stronger tactic, simpler proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 163
diff changeset
   203
175
f7602653dddd Preparing infrastructire for LAMBDA_PRS
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 173
diff changeset
   204
lemma append_rsp_fst:
450
2dc708ddb93a introduced a global list of respectfulness lemmas; the attribute is [quot_rsp]
Christian Urban <urbanc@in.tum.de>
parents: 448
diff changeset
   205
  assumes a : "l1 \<approx> l2"
214
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   206
  shows "(l1 @ s) \<approx> (l2 @ s)"
163
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   207
  using a
214
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   208
  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
   209
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   210
lemma append_end:
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   211
  shows "(e # l) \<approx> (l @ [e])"
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   212
  apply (induct l)
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   213
  apply (auto intro: list_eq.intros list_eq_refl)
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   214
  done
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   215
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   216
lemma rev_rsp:
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   217
  shows "a \<approx> rev a"
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   218
  apply (induct a)
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   219
  apply simp
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   220
  apply (rule list_eq_refl)
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   221
  apply simp_all
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   222
  apply (rule list_eq.intros(6))
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   223
  prefer 2
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   224
  apply (rule append_rsp_fst)
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   225
  apply assumption
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   226
  apply (rule append_end)
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   227
  done
163
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   228
214
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   229
lemma append_sym_rsp:
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   230
  shows "(a @ b) \<approx> (b @ a)"
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   231
  apply (rule list_eq.intros(6))
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 (rule rev_rsp)
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   234
  apply (rule list_eq.intros(6))
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   235
  apply (rule rev_rsp)
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   236
  apply (simp)
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   237
  apply (rule append_rsp_fst)
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   238
  apply (rule list_eq.intros(3))
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   239
  apply (rule rev_rsp)
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   240
  done
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   241
451
586e3dc4afdb Added 'TRY' to refl in clean_tac to get as far as possible. Removed unnecessary [quot_rsp] in FSet. Added necessary [quot_rsp] and one lifted thm in LamEx.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 450
diff changeset
   242
lemma append_rsp:
450
2dc708ddb93a introduced a global list of respectfulness lemmas; the attribute is [quot_rsp]
Christian Urban <urbanc@in.tum.de>
parents: 448
diff changeset
   243
  assumes a : "l1 \<approx> r1"
2dc708ddb93a introduced a global list of respectfulness lemmas; the attribute is [quot_rsp]
Christian Urban <urbanc@in.tum.de>
parents: 448
diff changeset
   244
  assumes b : "l2 \<approx> r2 "
214
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   245
  shows "(l1 @ l2) \<approx> (r1 @ r2)"
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   246
  apply (rule list_eq.intros(6))
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   247
  apply (rule append_rsp_fst)
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   248
  using a apply (assumption)
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   249
  apply (rule list_eq.intros(6))
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   250
  apply (rule append_sym_rsp)
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   251
  apply (rule list_eq.intros(6))
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   252
  apply (rule append_rsp_fst)
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   253
  using b apply (assumption)
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   254
  apply (rule append_sym_rsp)
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   255
  done
175
f7602653dddd Preparing infrastructire for LAMBDA_PRS
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 173
diff changeset
   256
506
91c374abde06 removed quot argument...not all examples work anymore
Christian Urban <urbanc@in.tum.de>
parents: 498
diff changeset
   257
lemma ho_append_rsp[quotient_rsp]:
228
268a727b0f10 disambiguate ===> syntax
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 226
diff changeset
   258
  "(op \<approx> ===> op \<approx> ===> op \<approx>) op @ op @"
214
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   259
  by (simp add: append_rsp)
175
f7602653dddd Preparing infrastructire for LAMBDA_PRS
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 173
diff changeset
   260
451
586e3dc4afdb Added 'TRY' to refl in clean_tac to get as far as possible. Removed unnecessary [quot_rsp] in FSet. Added necessary [quot_rsp] and one lifted thm in LamEx.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 450
diff changeset
   261
lemma map_rsp:
194
03c03e88efa9 Simplifying Int and Working on map
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 190
diff changeset
   262
  assumes a: "a \<approx> b"
03c03e88efa9 Simplifying Int and Working on map
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 190
diff changeset
   263
  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
   264
  using a
03c03e88efa9 Simplifying Int and Working on map
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 190
diff changeset
   265
  apply (induct)
03c03e88efa9 Simplifying Int and Working on map
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 190
diff changeset
   266
  apply(auto intro: list_eq.intros)
03c03e88efa9 Simplifying Int and Working on map
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 190
diff changeset
   267
  done
03c03e88efa9 Simplifying Int and Working on map
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 190
diff changeset
   268
506
91c374abde06 removed quot argument...not all examples work anymore
Christian Urban <urbanc@in.tum.de>
parents: 498
diff changeset
   269
lemma ho_map_rsp[quotient_rsp]:
294
a092c0b13d83 fold_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 292
diff changeset
   270
  "(op = ===> op \<approx> ===> op \<approx>) map map"
a092c0b13d83 fold_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 292
diff changeset
   271
  by (simp add: map_rsp)
194
03c03e88efa9 Simplifying Int and Working on map
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 190
diff changeset
   272
294
a092c0b13d83 fold_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 292
diff changeset
   273
lemma map_append:
450
2dc708ddb93a introduced a global list of respectfulness lemmas; the attribute is [quot_rsp]
Christian Urban <urbanc@in.tum.de>
parents: 448
diff changeset
   274
  "(map f (a @ b)) \<approx> (map f a) @ (map f b)"
215
89a2ff3f82c7 More finshed proofs and cleaning
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 214
diff changeset
   275
 by simp (rule list_eq_refl)
194
03c03e88efa9 Simplifying Int and Working on map
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 190
diff changeset
   276
506
91c374abde06 removed quot argument...not all examples work anymore
Christian Urban <urbanc@in.tum.de>
parents: 498
diff changeset
   277
lemma ho_fold_rsp[quotient_rsp]:
294
a092c0b13d83 fold_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 292
diff changeset
   278
  "(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
   279
  apply (auto simp add: FUN_REL_EQ)
294
a092c0b13d83 fold_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 292
diff changeset
   280
  apply (case_tac "rsp_fold x")
a092c0b13d83 fold_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 292
diff changeset
   281
  prefer 2
a092c0b13d83 fold_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 292
diff changeset
   282
  apply (erule_tac list_eq.induct)
a092c0b13d83 fold_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 292
diff changeset
   283
  apply (simp_all)
a092c0b13d83 fold_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 292
diff changeset
   284
  apply (erule_tac list_eq.induct)
a092c0b13d83 fold_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 292
diff changeset
   285
  apply (simp_all)
a092c0b13d83 fold_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 292
diff changeset
   286
  apply (auto simp add: memb_rsp rsp_fold_def)
a092c0b13d83 fold_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 292
diff changeset
   287
done
241
60acf3d3a4a0 Finding applications and duplicates filtered out in abstractions
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 239
diff changeset
   288
254
77ff9624cfd6 fixed the problem with types in map
Christian Urban <urbanc@in.tum.de>
parents: 252
diff changeset
   289
print_quotients
77ff9624cfd6 fixed the problem with types in map
Christian Urban <urbanc@in.tum.de>
parents: 252
diff changeset
   290
226
2a28e7ef3048 cleaned FSet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 225
diff changeset
   291
ML {* val qty = @{typ "'a fset"} *}
2a28e7ef3048 cleaned FSet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 225
diff changeset
   292
ML {* val rsp_thms =
458
44a70e69ef92 Code cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 455
diff changeset
   293
  @{thms ho_memb_rsp ho_cons_rsp ho_card1_rsp ho_map_rsp ho_append_rsp ho_fold_rsp} *}
206
1e227c9ee915 Fixed APPLY_RSP vs Cong in the InjRepAbs tactic.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 202
diff changeset
   294
364
4c455d58ac99 Fixes to the tactic after quotient_tac changed.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 356
diff changeset
   295
ML {* val (rty, rel, rel_refl, rel_eqv) = lookup_quot_data @{context} qty *}
4c455d58ac99 Fixes to the tactic after quotient_tac changed.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 356
diff changeset
   296
ML {* val (trans2, reps_same, absrep, quot) = lookup_quot_thms @{context} "fset"; *}
506
91c374abde06 removed quot argument...not all examples work anymore
Christian Urban <urbanc@in.tum.de>
parents: 498
diff changeset
   297
ML {* fun lift_tac_fset lthy t = lift_tac lthy t [rel_eqv] *}
314
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 309
diff changeset
   298
364
4c455d58ac99 Fixes to the tactic after quotient_tac changed.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 356
diff changeset
   299
lemma "IN x EMPTY = False"
455
9cb45d022524 tried to improve the inj_repabs_trm function but left the new part commented out
Christian Urban <urbanc@in.tum.de>
parents: 452
diff changeset
   300
apply(tactic {* procedure_tac @{context} @{thm m1} 1 *})
477
6c88b42da228 A bit of progress; but the object-logic vs meta-logic distinction is troublesome.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 475
diff changeset
   301
apply(tactic {* regularize_tac @{context} [rel_eqv] 1 *})
525
3f657c4fbefa Removed previous inj_repabs_tac
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 516
diff changeset
   302
apply(tactic {* all_inj_repabs_tac @{context} [rel_refl] [trans2] 1 *})
506
91c374abde06 removed quot argument...not all examples work anymore
Christian Urban <urbanc@in.tum.de>
parents: 498
diff changeset
   303
apply(tactic {* clean_tac @{context} 1*})
455
9cb45d022524 tried to improve the inj_repabs_trm function but left the new part commented out
Christian Urban <urbanc@in.tum.de>
parents: 452
diff changeset
   304
done
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
   305
364
4c455d58ac99 Fixes to the tactic after quotient_tac changed.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 356
diff changeset
   306
lemma "IN x (INSERT y xa) = (x = y \<or> IN x xa)"
4c455d58ac99 Fixes to the tactic after quotient_tac changed.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 356
diff changeset
   307
by (tactic {* lift_tac_fset @{context} @{thm m2} 1 *})
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
   308
364
4c455d58ac99 Fixes to the tactic after quotient_tac changed.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 356
diff changeset
   309
lemma "INSERT a (INSERT a x) = INSERT a x"
4c455d58ac99 Fixes to the tactic after quotient_tac changed.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 356
diff changeset
   310
apply (tactic {* lift_tac_fset @{context} @{thm list_eq.intros(4)} 1 *})
4c455d58ac99 Fixes to the tactic after quotient_tac changed.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 356
diff changeset
   311
done
4c455d58ac99 Fixes to the tactic after quotient_tac changed.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 356
diff changeset
   312
367
d444389fe3f9 The non-working procedure_tac.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 364
diff changeset
   313
lemma "x = xa \<Longrightarrow> INSERT a x = INSERT a xa"
364
4c455d58ac99 Fixes to the tactic after quotient_tac changed.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 356
diff changeset
   314
apply (tactic {* lift_tac_fset @{context} @{thm list_eq.intros(5)} 1 *})
4c455d58ac99 Fixes to the tactic after quotient_tac changed.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 356
diff changeset
   315
done
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
   316
367
d444389fe3f9 The non-working procedure_tac.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 364
diff changeset
   317
lemma "CARD x = Suc n \<Longrightarrow> (\<exists>a b. \<not> IN a b & x = INSERT a b)"
364
4c455d58ac99 Fixes to the tactic after quotient_tac changed.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 356
diff changeset
   318
apply (tactic {* lift_tac_fset @{context} @{thm card1_suc} 1 *})
4c455d58ac99 Fixes to the tactic after quotient_tac changed.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 356
diff changeset
   319
done
4c455d58ac99 Fixes to the tactic after quotient_tac changed.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 356
diff changeset
   320
4c455d58ac99 Fixes to the tactic after quotient_tac changed.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 356
diff changeset
   321
lemma "(\<not> IN x xa) = (CARD (INSERT x xa) = Suc (CARD xa))"
4c455d58ac99 Fixes to the tactic after quotient_tac changed.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 356
diff changeset
   322
apply (tactic {* lift_tac_fset @{context} @{thm not_mem_card1} 1 *})
4c455d58ac99 Fixes to the tactic after quotient_tac changed.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 356
diff changeset
   323
done
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
   324
442
7beed9b75ea2 renamed LAMBDA_RES_TAC and WEAK_LAMBDA_RES_TAC to lower case names
Christian Urban <urbanc@in.tum.de>
parents: 435
diff changeset
   325
lemma "FOLD f g (z::'b) (INSERT a x) =
364
4c455d58ac99 Fixes to the tactic after quotient_tac changed.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 356
diff changeset
   326
  (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)"
442
7beed9b75ea2 renamed LAMBDA_RES_TAC and WEAK_LAMBDA_RES_TAC to lower case names
Christian Urban <urbanc@in.tum.de>
parents: 435
diff changeset
   327
apply(tactic {* lift_tac_fset @{context} @{thm fold1.simps(2)} 1 *})
364
4c455d58ac99 Fixes to the tactic after quotient_tac changed.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 356
diff changeset
   328
done
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
   329
525
3f657c4fbefa Removed previous inj_repabs_tac
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 516
diff changeset
   330
ML {* fun inj_repabs_tac_fset lthy = inj_repabs_tac lthy [rel_refl] [trans2] *}
482
767baada01dc New APPLY_RSP which finally does automatic partial lifting :). Doesn't support same relation yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 481
diff changeset
   331
368
c5c49d240cde Conversion
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 367
diff changeset
   332
lemma "fmap f (FUNION (x::'b fset) (xa::'b fset)) = FUNION (fmap f x) (fmap f xa)"
c5c49d240cde Conversion
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 367
diff changeset
   333
apply (tactic {* lift_tac_fset @{context} @{thm map_append} 1 *})
c5c49d240cde Conversion
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 367
diff changeset
   334
done
c5c49d240cde Conversion
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 367
diff changeset
   335
367
d444389fe3f9 The non-working procedure_tac.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 364
diff changeset
   336
lemma "FUNION (FUNION x xa) xb = FUNION x (FUNION xa xb)"
d444389fe3f9 The non-working procedure_tac.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 364
diff changeset
   337
apply (tactic {* lift_tac_fset @{context} @{thm append_assoc} 1 *})
d444389fe3f9 The non-working procedure_tac.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 364
diff changeset
   338
done
d444389fe3f9 The non-working procedure_tac.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 364
diff changeset
   339
477
6c88b42da228 A bit of progress; but the object-logic vs meta-logic distinction is troublesome.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 475
diff changeset
   340
376
e99c0334d8bf lambda_prs and cleaning the existing examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 375
diff changeset
   341
lemma "\<lbrakk>P EMPTY; \<And>a x. P x \<Longrightarrow> P (INSERT a x)\<rbrakk> \<Longrightarrow> P l"
462
0911d3aabf47 clean_tac rewrites the definitions the other way
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 459
diff changeset
   342
apply (tactic {* (ObjectLogic.full_atomize_tac THEN' gen_frees_tac @{context}) 1 *})
392
98ccde1c184c Fixed FSet after merge.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 391
diff changeset
   343
apply(tactic {* procedure_tac @{context} @{thm list.induct} 1 *})
432
9c33c0809733 Finished and tested the new regularize
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 423
diff changeset
   344
apply(tactic {* regularize_tac @{context} [rel_eqv] 1 *})
398
fafcc54e531d some diagnostic code for r_mk_comb
Christian Urban <urbanc@in.tum.de>
parents: 397
diff changeset
   345
prefer 2
506
91c374abde06 removed quot argument...not all examples work anymore
Christian Urban <urbanc@in.tum.de>
parents: 498
diff changeset
   346
apply(tactic {* clean_tac @{context} 1 *})
481
7f97c52021c9 Fixed unlam for non-abstractions and updated list_induct_part proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 480
diff changeset
   347
apply(tactic {* inj_repabs_tac_fset @{context} 1*}) (* 3 *) (* Ball-Ball *)
7f97c52021c9 Fixed unlam for non-abstractions and updated list_induct_part proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 480
diff changeset
   348
apply(tactic {* inj_repabs_tac_fset @{context} 1*}) (* 9 *) (* Rep-Abs-elim - can be complex Rep-Abs *)
7f97c52021c9 Fixed unlam for non-abstractions and updated list_induct_part proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 480
diff changeset
   349
apply(tactic {* inj_repabs_tac_fset @{context} 1*}) (* 2 *) (* lam-lam-elim for R = (===>) *)
7f97c52021c9 Fixed unlam for non-abstractions and updated list_induct_part proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 480
diff changeset
   350
apply(tactic {* inj_repabs_tac_fset @{context} 1*}) (* 3 *) (* Ball-Ball *)
7f97c52021c9 Fixed unlam for non-abstractions and updated list_induct_part proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 480
diff changeset
   351
apply(tactic {* inj_repabs_tac_fset @{context} 1*}) (* 9 *) (* Rep-Abs-elim - can be complex Rep-Abs *)
7f97c52021c9 Fixed unlam for non-abstractions and updated list_induct_part proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 480
diff changeset
   352
apply(tactic {* inj_repabs_tac_fset @{context} 1*}) (* 2 *) (* lam-lam-elim for R = (===>) *)
7f97c52021c9 Fixed unlam for non-abstractions and updated list_induct_part proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 480
diff changeset
   353
apply(tactic {* inj_repabs_tac_fset @{context} 1*}) (* B *) (* Cong *)
7f97c52021c9 Fixed unlam for non-abstractions and updated list_induct_part proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 480
diff changeset
   354
apply(tactic {* inj_repabs_tac_fset @{context} 1*}) (* B *) (* Cong *)
7f97c52021c9 Fixed unlam for non-abstractions and updated list_induct_part proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 480
diff changeset
   355
apply(tactic {* inj_repabs_tac_fset @{context} 1*}) (* 8 *) (* = reflexivity arising from cong *)
7f97c52021c9 Fixed unlam for non-abstractions and updated list_induct_part proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 480
diff changeset
   356
apply(tactic {* inj_repabs_tac_fset @{context} 1*}) (* A *) (* application if type needs lifting *)
7f97c52021c9 Fixed unlam for non-abstractions and updated list_induct_part proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 480
diff changeset
   357
apply(tactic {* inj_repabs_tac_fset @{context} 1*}) (* 9 *) (* Rep-Abs-elim - can be complex Rep-Abs *)
7f97c52021c9 Fixed unlam for non-abstractions and updated list_induct_part proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 480
diff changeset
   358
apply(tactic {* inj_repabs_tac_fset @{context} 1*}) (* E *) (* R x y assumptions *)
7f97c52021c9 Fixed unlam for non-abstractions and updated list_induct_part proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 480
diff changeset
   359
apply(tactic {* inj_repabs_tac_fset @{context} 1*}) (* 9 *) (* Rep-Abs-elim - can be complex Rep-Abs *)
7f97c52021c9 Fixed unlam for non-abstractions and updated list_induct_part proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 480
diff changeset
   360
apply(tactic {* inj_repabs_tac_fset @{context} 1*}) (* D *) (* reflexivity of basic relations *)
7f97c52021c9 Fixed unlam for non-abstractions and updated list_induct_part proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 480
diff changeset
   361
apply(tactic {* inj_repabs_tac_fset @{context} 1*}) (* B *) (* Cong *)
7f97c52021c9 Fixed unlam for non-abstractions and updated list_induct_part proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 480
diff changeset
   362
apply(tactic {* inj_repabs_tac_fset @{context} 1*}) (* B *) (* Cong *)
7f97c52021c9 Fixed unlam for non-abstractions and updated list_induct_part proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 480
diff changeset
   363
apply(tactic {* inj_repabs_tac_fset @{context} 1*}) (* 8 *) (* = reflexivity arising from cong *)
7f97c52021c9 Fixed unlam for non-abstractions and updated list_induct_part proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 480
diff changeset
   364
apply(tactic {* inj_repabs_tac_fset @{context} 1*}) (* B *) (* Cong *)
7f97c52021c9 Fixed unlam for non-abstractions and updated list_induct_part proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 480
diff changeset
   365
apply(tactic {* inj_repabs_tac_fset @{context} 1*}) (* 8 *) (* = reflexivity arising from cong *)
7f97c52021c9 Fixed unlam for non-abstractions and updated list_induct_part proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 480
diff changeset
   366
apply(tactic {* inj_repabs_tac_fset @{context} 1*}) (* C *) (* = and extensionality *)
7f97c52021c9 Fixed unlam for non-abstractions and updated list_induct_part proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 480
diff changeset
   367
apply(tactic {* inj_repabs_tac_fset @{context} 1*}) (* 3 *) (* Ball-Ball *)
7f97c52021c9 Fixed unlam for non-abstractions and updated list_induct_part proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 480
diff changeset
   368
apply(tactic {* inj_repabs_tac_fset @{context} 1*}) (* 9 *) (* Rep-Abs-elim - can be complex Rep-Abs *)
7f97c52021c9 Fixed unlam for non-abstractions and updated list_induct_part proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 480
diff changeset
   369
apply(tactic {* inj_repabs_tac_fset @{context} 1*}) (* 2 *) (* lam-lam-elim for R = (===>) *) 
7f97c52021c9 Fixed unlam for non-abstractions and updated list_induct_part proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 480
diff changeset
   370
apply(tactic {* inj_repabs_tac_fset @{context} 1*}) (* B *) (* Cong *)
7f97c52021c9 Fixed unlam for non-abstractions and updated list_induct_part proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 480
diff changeset
   371
apply(tactic {* inj_repabs_tac_fset @{context} 1*}) (* B *) (* Cong *)
7f97c52021c9 Fixed unlam for non-abstractions and updated list_induct_part proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 480
diff changeset
   372
apply(tactic {* inj_repabs_tac_fset @{context} 1*}) (* 8 *) (* = reflexivity arising from cong *)
7f97c52021c9 Fixed unlam for non-abstractions and updated list_induct_part proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 480
diff changeset
   373
apply(tactic {* inj_repabs_tac_fset @{context} 1*}) (* A *) (* application if type needs lifting *)
7f97c52021c9 Fixed unlam for non-abstractions and updated list_induct_part proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 480
diff changeset
   374
apply(tactic {* inj_repabs_tac_fset @{context} 1*}) (* 9 *) (* Rep-Abs-elim - can be complex Rep-Abs *)
7f97c52021c9 Fixed unlam for non-abstractions and updated list_induct_part proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 480
diff changeset
   375
apply(tactic {* inj_repabs_tac_fset @{context} 1*}) (* E *) (* R x y assumptions *)
7f97c52021c9 Fixed unlam for non-abstractions and updated list_induct_part proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 480
diff changeset
   376
apply(tactic {* inj_repabs_tac_fset @{context} 1*}) (* 9 *) (* Rep-Abs-elim - can be complex Rep-Abs *)
7f97c52021c9 Fixed unlam for non-abstractions and updated list_induct_part proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 480
diff changeset
   377
apply(tactic {* inj_repabs_tac_fset @{context} 1*}) (* E *) (* R x y assumptions *)
7f97c52021c9 Fixed unlam for non-abstractions and updated list_induct_part proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 480
diff changeset
   378
apply(tactic {* inj_repabs_tac_fset @{context} 1*}) (* A *) (* application if type needs lifting *)
7f97c52021c9 Fixed unlam for non-abstractions and updated list_induct_part proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 480
diff changeset
   379
apply(tactic {* inj_repabs_tac_fset @{context} 1*}) (* 9 *) (* Rep-Abs-elim - can be complex Rep-Abs *)
7f97c52021c9 Fixed unlam for non-abstractions and updated list_induct_part proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 480
diff changeset
   380
apply(tactic {* inj_repabs_tac_fset @{context} 1*}) (* E *) (* R x y assumptions *)
7f97c52021c9 Fixed unlam for non-abstractions and updated list_induct_part proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 480
diff changeset
   381
apply(tactic {* inj_repabs_tac_fset @{context} 1*}) (* 9 *) (* Rep-Abs-elim - can be complex Rep-Abs *)
7f97c52021c9 Fixed unlam for non-abstractions and updated list_induct_part proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 480
diff changeset
   382
apply(tactic {* inj_repabs_tac_fset @{context} 1*}) (* A *) (* application if type needs lifting *)
7f97c52021c9 Fixed unlam for non-abstractions and updated list_induct_part proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 480
diff changeset
   383
apply(tactic {* inj_repabs_tac_fset @{context} 1*}) (* A *) (* application if type needs lifting *)
7f97c52021c9 Fixed unlam for non-abstractions and updated list_induct_part proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 480
diff changeset
   384
apply(tactic {* inj_repabs_tac_fset @{context} 1*}) (* 7 *) (* respectfulness *)
7f97c52021c9 Fixed unlam for non-abstractions and updated list_induct_part proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 480
diff changeset
   385
apply(tactic {* inj_repabs_tac_fset @{context} 1*}) (* 8 *) (* = reflexivity arising from cong *)
7f97c52021c9 Fixed unlam for non-abstractions and updated list_induct_part proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 480
diff changeset
   386
apply(tactic {* inj_repabs_tac_fset @{context} 1*}) (* 9 *) (* Rep-Abs-elim - can be complex Rep-Abs *)
7f97c52021c9 Fixed unlam for non-abstractions and updated list_induct_part proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 480
diff changeset
   387
apply(tactic {* inj_repabs_tac_fset @{context} 1*}) (* E *) (* R x y assumptions *)
7f97c52021c9 Fixed unlam for non-abstractions and updated list_induct_part proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 480
diff changeset
   388
apply(tactic {* inj_repabs_tac_fset @{context} 1*}) (* A *) (* application if type needs lifting *)
7f97c52021c9 Fixed unlam for non-abstractions and updated list_induct_part proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 480
diff changeset
   389
apply(tactic {* inj_repabs_tac_fset @{context} 1*}) (* 9 *) (* Rep-Abs-elim - can be complex Rep-Abs *)
516
bed81795848c Using APPLY_RSP1; again a little bit faster.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 509
diff changeset
   390
apply(tactic {* inj_repabs_tac_fset @{context} 1*}) (* E *) (* R x y assumptions *)
bed81795848c Using APPLY_RSP1; again a little bit faster.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 509
diff changeset
   391
apply(tactic {* inj_repabs_tac_fset @{context} 1*}) (* 9 *) (* Rep-Abs-elim - can be complex Rep-Abs *)
bed81795848c Using APPLY_RSP1; again a little bit faster.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 509
diff changeset
   392
apply(tactic {* inj_repabs_tac_fset @{context} 1*}) (* E *) (* R x y assumptions *)
414
4dad34ca50db Minor cleaning
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 401
diff changeset
   393
done
390
1dd6a21cdd1c test with monos
Christian Urban <urbanc@in.tum.de>
parents: 387
diff changeset
   394
482
767baada01dc New APPLY_RSP which finally does automatic partial lifting :). Doesn't support same relation yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 481
diff changeset
   395
lemma list_induct_part:
767baada01dc New APPLY_RSP which finally does automatic partial lifting :). Doesn't support same relation yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 481
diff changeset
   396
  assumes a: "P (x :: 'a list) ([] :: 'c list)"
767baada01dc New APPLY_RSP which finally does automatic partial lifting :). Doesn't support same relation yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 481
diff changeset
   397
  assumes b: "\<And>e t. P x t \<Longrightarrow> P x (e # t)"
767baada01dc New APPLY_RSP which finally does automatic partial lifting :). Doesn't support same relation yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 481
diff changeset
   398
  shows "P x l"
767baada01dc New APPLY_RSP which finally does automatic partial lifting :). Doesn't support same relation yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 481
diff changeset
   399
  apply (rule_tac P="P x" in list.induct)
767baada01dc New APPLY_RSP which finally does automatic partial lifting :). Doesn't support same relation yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 481
diff changeset
   400
  apply (rule a)
767baada01dc New APPLY_RSP which finally does automatic partial lifting :). Doesn't support same relation yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 481
diff changeset
   401
  apply (rule b)
767baada01dc New APPLY_RSP which finally does automatic partial lifting :). Doesn't support same relation yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 481
diff changeset
   402
  apply (assumption)
767baada01dc New APPLY_RSP which finally does automatic partial lifting :). Doesn't support same relation yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 481
diff changeset
   403
  done
767baada01dc New APPLY_RSP which finally does automatic partial lifting :). Doesn't support same relation yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 481
diff changeset
   404
506
91c374abde06 removed quot argument...not all examples work anymore
Christian Urban <urbanc@in.tum.de>
parents: 498
diff changeset
   405
ML {* quot *}
91c374abde06 removed quot argument...not all examples work anymore
Christian Urban <urbanc@in.tum.de>
parents: 498
diff changeset
   406
thm quotient_thm
91c374abde06 removed quot argument...not all examples work anymore
Christian Urban <urbanc@in.tum.de>
parents: 498
diff changeset
   407
482
767baada01dc New APPLY_RSP which finally does automatic partial lifting :). Doesn't support same relation yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 481
diff changeset
   408
lemma "P (x :: 'a list) (EMPTY :: 'c fset) \<Longrightarrow> (\<And>e t. P x t \<Longrightarrow> P x (INSERT e t)) \<Longrightarrow> P x l"
767baada01dc New APPLY_RSP which finally does automatic partial lifting :). Doesn't support same relation yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 481
diff changeset
   409
apply (tactic {* lift_tac_fset @{context} @{thm list_induct_part} 1 *})
767baada01dc New APPLY_RSP which finally does automatic partial lifting :). Doesn't support same relation yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 481
diff changeset
   410
done
767baada01dc New APPLY_RSP which finally does automatic partial lifting :). Doesn't support same relation yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 481
diff changeset
   411
767baada01dc New APPLY_RSP which finally does automatic partial lifting :). Doesn't support same relation yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 481
diff changeset
   412
lemma "P (x :: 'a fset) (EMPTY :: 'c fset) \<Longrightarrow> (\<And>e t. P x t \<Longrightarrow> P x (INSERT e t)) \<Longrightarrow> P x l"
767baada01dc New APPLY_RSP which finally does automatic partial lifting :). Doesn't support same relation yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 481
diff changeset
   413
apply (tactic {* lift_tac_fset @{context} @{thm list_induct_part} 1 *})
767baada01dc New APPLY_RSP which finally does automatic partial lifting :). Doesn't support same relation yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 481
diff changeset
   414
done
767baada01dc New APPLY_RSP which finally does automatic partial lifting :). Doesn't support same relation yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 481
diff changeset
   415
767baada01dc New APPLY_RSP which finally does automatic partial lifting :). Doesn't support same relation yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 481
diff changeset
   416
lemma "P (x :: 'a fset) ([] :: 'c list) \<Longrightarrow> (\<And>e t. P x t \<Longrightarrow> P x (e # t)) \<Longrightarrow> P x l"
767baada01dc New APPLY_RSP which finally does automatic partial lifting :). Doesn't support same relation yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 481
diff changeset
   417
apply (tactic {* lift_tac_fset @{context} @{thm list_induct_part} 1 *})
767baada01dc New APPLY_RSP which finally does automatic partial lifting :). Doesn't support same relation yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 481
diff changeset
   418
done
767baada01dc New APPLY_RSP which finally does automatic partial lifting :). Doesn't support same relation yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 481
diff changeset
   419
483
74348dc2f8bb Lifting to 2 different types :)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 482
diff changeset
   420
quotient fset2 = "'a list" / "list_eq"
74348dc2f8bb Lifting to 2 different types :)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 482
diff changeset
   421
  apply(rule equiv_list_eq)
74348dc2f8bb Lifting to 2 different types :)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 482
diff changeset
   422
  done
74348dc2f8bb Lifting to 2 different types :)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 482
diff changeset
   423
74348dc2f8bb Lifting to 2 different types :)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 482
diff changeset
   424
quotient_def
74348dc2f8bb Lifting to 2 different types :)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 482
diff changeset
   425
  EMPTY2 :: "'a fset2"
74348dc2f8bb Lifting to 2 different types :)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 482
diff changeset
   426
where
74348dc2f8bb Lifting to 2 different types :)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 482
diff changeset
   427
  "EMPTY2 \<equiv> ([]::'a list)"
74348dc2f8bb Lifting to 2 different types :)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 482
diff changeset
   428
74348dc2f8bb Lifting to 2 different types :)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 482
diff changeset
   429
quotient_def
74348dc2f8bb Lifting to 2 different types :)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 482
diff changeset
   430
  INSERT2 :: "'a \<Rightarrow> 'a fset2 \<Rightarrow> 'a fset2"
74348dc2f8bb Lifting to 2 different types :)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 482
diff changeset
   431
where
74348dc2f8bb Lifting to 2 different types :)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 482
diff changeset
   432
  "INSERT2 \<equiv> op #"
74348dc2f8bb Lifting to 2 different types :)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 482
diff changeset
   433
74348dc2f8bb Lifting to 2 different types :)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 482
diff changeset
   434
ML {* val quot = @{thms QUOTIENT_fset QUOTIENT_fset2} *}
507
f7569f994195 the error occurs in APPLY_RSP_TAC where the wrong quotient (for LIST_REL) is applied
Christian Urban <urbanc@in.tum.de>
parents: 506
diff changeset
   435
ML {* fun inj_repabs_tac_fset lthy = inj_repabs_tac lthy [rel_refl] [trans2] *}
f7569f994195 the error occurs in APPLY_RSP_TAC where the wrong quotient (for LIST_REL) is applied
Christian Urban <urbanc@in.tum.de>
parents: 506
diff changeset
   436
ML {* fun lift_tac_fset lthy t = lift_tac lthy t [rel_eqv] *}
483
74348dc2f8bb Lifting to 2 different types :)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 482
diff changeset
   437
74348dc2f8bb Lifting to 2 different types :)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 482
diff changeset
   438
lemma "P (x :: 'a fset2) (EMPTY :: 'c fset) \<Longrightarrow> (\<And>e t. P x t \<Longrightarrow> P x (INSERT e t)) \<Longrightarrow> P x l"
74348dc2f8bb Lifting to 2 different types :)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 482
diff changeset
   439
apply (tactic {* lift_tac_fset @{context} @{thm list_induct_part} 1 *})
74348dc2f8bb Lifting to 2 different types :)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 482
diff changeset
   440
done
74348dc2f8bb Lifting to 2 different types :)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 482
diff changeset
   441
74348dc2f8bb Lifting to 2 different types :)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 482
diff changeset
   442
lemma "P (x :: 'a fset) (EMPTY2 :: 'c fset2) \<Longrightarrow> (\<And>e t. P x t \<Longrightarrow> P x (INSERT2 e t)) \<Longrightarrow> P x l"
74348dc2f8bb Lifting to 2 different types :)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 482
diff changeset
   443
apply (tactic {* lift_tac_fset @{context} @{thm list_induct_part} 1 *})
74348dc2f8bb Lifting to 2 different types :)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 482
diff changeset
   444
done
74348dc2f8bb Lifting to 2 different types :)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 482
diff changeset
   445
273
b82e765ca464 Lifting 'fold1.simps(2)' and some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 270
diff changeset
   446
quotient_def
276
783d6c940e45 Experiments in Int
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 274
diff changeset
   447
  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
   448
where
b82e765ca464 Lifting 'fold1.simps(2)' and some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 270
diff changeset
   449
  "fset_rec \<equiv> list_rec"
b82e765ca464 Lifting 'fold1.simps(2)' and some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 270
diff changeset
   450
292
bd76f0398aa9 More functionality for lifting list.cases and list.recs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 291
diff changeset
   451
quotient_def
bd76f0398aa9 More functionality for lifting list.cases and list.recs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 291
diff changeset
   452
  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
   453
where
bd76f0398aa9 More functionality for lifting list.cases and list.recs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 291
diff changeset
   454
  "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
   455
296
eab108c8d4b7 Minor changes
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 294
diff changeset
   456
(* Probably not true without additional assumptions about the function *)
506
91c374abde06 removed quot argument...not all examples work anymore
Christian Urban <urbanc@in.tum.de>
parents: 498
diff changeset
   457
lemma list_rec_rsp[quotient_rsp]:
292
bd76f0398aa9 More functionality for lifting list.cases and list.recs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 291
diff changeset
   458
  "(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
   459
  apply (auto simp add: FUN_REL_EQ)
296
eab108c8d4b7 Minor changes
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 294
diff changeset
   460
  apply (erule_tac list_eq.induct)
eab108c8d4b7 Minor changes
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 294
diff changeset
   461
  apply (simp_all)
292
bd76f0398aa9 More functionality for lifting list.cases and list.recs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 291
diff changeset
   462
  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
   463
506
91c374abde06 removed quot argument...not all examples work anymore
Christian Urban <urbanc@in.tum.de>
parents: 498
diff changeset
   464
lemma list_case_rsp[quotient_rsp]:
292
bd76f0398aa9 More functionality for lifting list.cases and list.recs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 291
diff changeset
   465
  "(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
   466
  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
   467
  sorry
bd76f0398aa9 More functionality for lifting list.cases and list.recs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 291
diff changeset
   468
bd76f0398aa9 More functionality for lifting list.cases and list.recs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 291
diff changeset
   469
ML {* val rsp_thms = @{thms list_rec_rsp list_case_rsp} @ rsp_thms *}
508
fac6069d8e80 Made your version work again with LIST_REL_EQ.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 506
diff changeset
   470
ML {* fun lift_tac_fset lthy t = lift_tac lthy t [rel_eqv] *}
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
   471
376
e99c0334d8bf lambda_prs and cleaning the existing examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 375
diff changeset
   472
lemma "fset_rec (f1::'t) x (INSERT a xa) = x a xa (fset_rec f1 x xa)"
e99c0334d8bf lambda_prs and cleaning the existing examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 375
diff changeset
   473
apply (tactic {* lift_tac_fset @{context} @{thm list.recs(2)} 1 *})
e99c0334d8bf lambda_prs and cleaning the existing examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 375
diff changeset
   474
done
e99c0334d8bf lambda_prs and cleaning the existing examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 375
diff changeset
   475
e99c0334d8bf lambda_prs and cleaning the existing examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 375
diff changeset
   476
lemma "fset_case (f1::'t) f2 (INSERT a xa) = f2 a xa"
e99c0334d8bf lambda_prs and cleaning the existing examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 375
diff changeset
   477
apply (tactic {* lift_tac_fset @{context} @{thm list.cases(2)} 1 *})
e99c0334d8bf lambda_prs and cleaning the existing examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 375
diff changeset
   478
done
348
b1f83c7a8674 More theorems lifted in the goal-directed way.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 338
diff changeset
   479
496
8f1bf5266ebc Added the definition to quotient constant data.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 495
diff changeset
   480
163
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   481
end