FSet.thy
author Christian Urban <urbanc@in.tum.de>
Mon, 23 Nov 2009 13:24:12 +0100
changeset 334 5a7024be9083
parent 333 7851e2a74f85
child 335 87eae6a2e5ff
permissions -rw-r--r--
code review with Cezary
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 *}
172
da38ce2711a6 More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 171
diff changeset
   307
314
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 309
diff changeset
   308
thm m2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 309
diff changeset
   309
317
d3c7f6d19c7f Still don't know how to do the proof automatically.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 314
diff changeset
   310
thm neq_Nil_conv
d3c7f6d19c7f Still don't know how to do the proof automatically.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 314
diff changeset
   311
term REP_fset
d3c7f6d19c7f Still don't know how to do the proof automatically.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 314
diff changeset
   312
term "op --->"
d3c7f6d19c7f Still don't know how to do the proof automatically.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 314
diff changeset
   313
thm INSERT_def
d3c7f6d19c7f Still don't know how to do the proof automatically.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 314
diff changeset
   314
ML {* val defs_sym = flat (map (add_lower_defs @{context}) @{thms INSERT_def}) *}
d3c7f6d19c7f Still don't know how to do the proof automatically.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 314
diff changeset
   315
(*ML {* lift_thm_fset @{context} @{thm neq_Nil_conv} *}*)
226
2a28e7ef3048 cleaned FSet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 225
diff changeset
   316
ML {* lift_thm_fset @{context} @{thm m1} *}
2a28e7ef3048 cleaned FSet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 225
diff changeset
   317
ML {* lift_thm_fset @{context} @{thm m2} *}
2a28e7ef3048 cleaned FSet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 225
diff changeset
   318
ML {* lift_thm_fset @{context} @{thm list_eq.intros(4)} *}
2a28e7ef3048 cleaned FSet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 225
diff changeset
   319
ML {* lift_thm_fset @{context} @{thm list_eq.intros(5)} *}
2a28e7ef3048 cleaned FSet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 225
diff changeset
   320
ML {* lift_thm_fset @{context} @{thm card1_suc} *}
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
   321
ML {* lift_thm_fset @{context} @{thm map_append} *}
251
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 248
diff changeset
   322
ML {* lift_thm_fset @{context} @{thm append_assoc} *}
267
3764566c1151 Automatic FORALL_PRS. 'list.induct' lifts automatically. Faster ALLEX_RSP
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 260
diff changeset
   323
ML {* lift_thm_fset @{context} @{thm list.induct} *}
273
b82e765ca464 Lifting 'fold1.simps(2)' and some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 270
diff changeset
   324
ML {* lift_thm_fset @{context} @{thm fold1.simps(2)} *}
309
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   325
ML {* lift_thm_fset @{context} @{thm not_mem_card1} *}
171
13aab4c59096 More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 168
diff changeset
   326
273
b82e765ca464 Lifting 'fold1.simps(2)' and some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 270
diff changeset
   327
quotient_def
276
783d6c940e45 Experiments in Int
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 274
diff changeset
   328
  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
   329
where
b82e765ca464 Lifting 'fold1.simps(2)' and some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 270
diff changeset
   330
  "fset_rec \<equiv> list_rec"
b82e765ca464 Lifting 'fold1.simps(2)' and some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 270
diff changeset
   331
292
bd76f0398aa9 More functionality for lifting list.cases and list.recs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 291
diff changeset
   332
quotient_def
bd76f0398aa9 More functionality for lifting list.cases and list.recs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 291
diff changeset
   333
  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
   334
where
bd76f0398aa9 More functionality for lifting list.cases and list.recs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 291
diff changeset
   335
  "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
   336
296
eab108c8d4b7 Minor changes
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 294
diff changeset
   337
(* 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
   338
lemma list_rec_rsp:
bd76f0398aa9 More functionality for lifting list.cases and list.recs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 291
diff changeset
   339
  "(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
   340
  apply (auto simp add: FUN_REL_EQ)
296
eab108c8d4b7 Minor changes
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 294
diff changeset
   341
  apply (erule_tac list_eq.induct)
eab108c8d4b7 Minor changes
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 294
diff changeset
   342
  apply (simp_all)
292
bd76f0398aa9 More functionality for lifting list.cases and list.recs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 291
diff changeset
   343
  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
   344
292
bd76f0398aa9 More functionality for lifting list.cases and list.recs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 291
diff changeset
   345
lemma list_case_rsp:
bd76f0398aa9 More functionality for lifting list.cases and list.recs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 291
diff changeset
   346
  "(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
   347
  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
   348
  sorry
bd76f0398aa9 More functionality for lifting list.cases and list.recs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 291
diff changeset
   349
bd76f0398aa9 More functionality for lifting list.cases and list.recs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 291
diff changeset
   350
bd76f0398aa9 More functionality for lifting list.cases and list.recs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 291
diff changeset
   351
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
   352
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
   353
bd76f0398aa9 More functionality for lifting list.cases and list.recs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 291
diff changeset
   354
ML {* fun lift_thm_fset lthy t = lift_thm lthy qty "fset" rsp_thms defs t *}
bd76f0398aa9 More functionality for lifting list.cases and list.recs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 291
diff changeset
   355
300
c6a9b4e4d548 Fixes for the other get_fun implementation.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 296
diff changeset
   356
292
bd76f0398aa9 More functionality for lifting list.cases and list.recs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 291
diff changeset
   357
ML {* map (lift_thm_fset @{context}) @{thms list.recs} *}
bd76f0398aa9 More functionality for lifting list.cases and list.recs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 291
diff changeset
   358
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
   359
304
e741c735b867 Atomizing a "goal" theorems.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 300
diff changeset
   360
lemma list_induct_part:
e741c735b867 Atomizing a "goal" theorems.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 300
diff changeset
   361
  assumes a: "P (x :: 'a list) ([] :: 'a list)"
e741c735b867 Atomizing a "goal" theorems.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 300
diff changeset
   362
  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
   363
  shows "P x l"
e741c735b867 Atomizing a "goal" theorems.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 300
diff changeset
   364
  apply (rule_tac P="P x" in list.induct)
e741c735b867 Atomizing a "goal" theorems.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 300
diff changeset
   365
  apply (rule a)
e741c735b867 Atomizing a "goal" theorems.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 300
diff changeset
   366
  apply (rule b)
e741c735b867 Atomizing a "goal" theorems.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 300
diff changeset
   367
  apply (assumption)
e741c735b867 Atomizing a "goal" theorems.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 300
diff changeset
   368
  done
273
b82e765ca464 Lifting 'fold1.simps(2)' and some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 270
diff changeset
   369
292
bd76f0398aa9 More functionality for lifting list.cases and list.recs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 291
diff changeset
   370
bd76f0398aa9 More functionality for lifting list.cases and list.recs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 291
diff changeset
   371
(* Construction site starts here *)
bd76f0398aa9 More functionality for lifting list.cases and list.recs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 291
diff changeset
   372
bd76f0398aa9 More functionality for lifting list.cases and list.recs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 291
diff changeset
   373
273
b82e765ca464 Lifting 'fold1.simps(2)' and some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 270
diff changeset
   374
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
   375
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
   376
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
   377
296
eab108c8d4b7 Minor changes
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 294
diff changeset
   378
thm list.recs(2)
304
e741c735b867 Atomizing a "goal" theorems.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 300
diff changeset
   379
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
   380
5a7024be9083 code review with Cezary
Christian Urban <urbanc@in.tum.de>
parents: 333
diff changeset
   381
285
8ebdef196fd5 Infrastructure for polymorphic types
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 276
diff changeset
   382
(* prove {* build_regularize_goal t_a rty rel @{context}  *}
8ebdef196fd5 Infrastructure for polymorphic types
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 276
diff changeset
   383
 ML_prf {*  fun tac ctxt = FIRST' [
251
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 248
diff changeset
   384
      rtac rel_refl,
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 248
diff changeset
   385
      atac,
285
8ebdef196fd5 Infrastructure for polymorphic types
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 276
diff changeset
   386
      rtac @{thm universal_twice},
8ebdef196fd5 Infrastructure for polymorphic types
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 276
diff changeset
   387
      (rtac @{thm impI} THEN' atac),
8ebdef196fd5 Infrastructure for polymorphic types
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 276
diff changeset
   388
      rtac @{thm implication_twice},
334
5a7024be9083 code review with Cezary
Christian Urban <urbanc@in.tum.de>
parents: 333
diff changeset
   389
      //comented out  rtac @{thm equality_twice}, //
285
8ebdef196fd5 Infrastructure for polymorphic types
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 276
diff changeset
   390
      EqSubst.eqsubst_tac ctxt [0]
239
02b14a21761a Cleaning of the interface to lift.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 233
diff changeset
   391
        [(@{thm equiv_res_forall} OF [rel_eqv]),
285
8ebdef196fd5 Infrastructure for polymorphic types
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 276
diff changeset
   392
         (@{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
   393
      (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
   394
      (rtac @{thm RIGHT_RES_FORALL_REGULAR})
285
8ebdef196fd5 Infrastructure for polymorphic types
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 276
diff changeset
   395
     ]; *}
252
e30997c88050 Regularize for equalities and a better tactic. "alpha.cases" now lifts.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 251
diff changeset
   396
  apply (atomize(full))
285
8ebdef196fd5 Infrastructure for polymorphic types
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 276
diff changeset
   397
  apply (tactic {* REPEAT_ALL_NEW (tac @{context}) 1 *})
334
5a7024be9083 code review with Cezary
Christian Urban <urbanc@in.tum.de>
parents: 333
diff changeset
   398
  done  *)
305
d7b60303adb8 Removed 'Toplevel.program' for polyml 5.3
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 304
diff changeset
   399
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
   400
ML {*
285
8ebdef196fd5 Infrastructure for polymorphic types
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 276
diff changeset
   401
  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
   402
  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
   403
*}
300
c6a9b4e4d548 Fixes for the other get_fun implementation.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 296
diff changeset
   404
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
   405
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
   406
apply(atomize(full))
194
03c03e88efa9 Simplifying Int and Working on map
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 190
diff changeset
   407
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
   408
done
305
d7b60303adb8 Removed 'Toplevel.program' for polyml 5.3
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 304
diff changeset
   409
ML {*
d7b60303adb8 Removed 'Toplevel.program' for polyml 5.3
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 304
diff changeset
   410
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
   411
*}
241
60acf3d3a4a0 Finding applications and duplicates filtered out in abstractions
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 239
diff changeset
   412
292
bd76f0398aa9 More functionality for lifting list.cases and list.recs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 291
diff changeset
   413
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
   414
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
   415
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
   416
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
   417
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
   418
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
   419
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
   420
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
   421
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
   422
ML app_prs_thms
296
eab108c8d4b7 Minor changes
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 294
diff changeset
   423
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
   424
ML lam_prs_thms
e741c735b867 Atomizing a "goal" theorems.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 300
diff changeset
   425
ML {* val t_id = simp_ids @{context} t_l *}
e741c735b867 Atomizing a "goal" theorems.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 300
diff changeset
   426
thm INSERT_def
292
bd76f0398aa9 More functionality for lifting list.cases and list.recs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 291
diff changeset
   427
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
   428
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
   429
ML allthms
e741c735b867 Atomizing a "goal" theorems.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 300
diff changeset
   430
thm FORALL_PRS
309
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   431
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
   432
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
   433
ML {* ObjectLogic.rulify t_s *}
e741c735b867 Atomizing a "goal" theorems.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 300
diff changeset
   434
e741c735b867 Atomizing a "goal" theorems.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 300
diff changeset
   435
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"} *}
e741c735b867 Atomizing a "goal" theorems.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 300
diff changeset
   436
ML {* val vars = map Free (Term.add_frees gl []) *}
e741c735b867 Atomizing a "goal" theorems.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 300
diff changeset
   437
ML {* fun lambda_all (var as Free(_, T)) trm = (Term.all T) $ lambda var trm *}
e741c735b867 Atomizing a "goal" theorems.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 300
diff changeset
   438
ML {* val gla = fold lambda_all vars gl *}
332
87f5fbebd6d5 Fixes for recent changes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 317
diff changeset
   439
ML {* val glf = Type.legacy_freeze gla *}
304
e741c735b867 Atomizing a "goal" theorems.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 300
diff changeset
   440
ML {* val glac = (snd o Thm.dest_equals o cprop_of) (ObjectLogic.atomize (cterm_of @{theory} glf)) *}
309
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   441
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   442
ML {*
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   443
fun apply_subt2 f trm trm2 =
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   444
  case (trm, trm2) of
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   445
    (Abs (x, T, t), Abs (x2, T2, t2)) =>
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   446
       let
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   447
         val (x', t') = Term.dest_abs (x, T, t);
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   448
         val (x2', t2') = Term.dest_abs (x2, T2, t2)
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   449
         val (s1, s2) = f t' t2';
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   450
       in
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   451
         (Term.absfree (x', T, s1),
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   452
          Term.absfree (x2', T2, s2))
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   453
       end
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   454
  | _ => f trm trm2
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   455
*}
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   456
317
d3c7f6d19c7f Still don't know how to do the proof automatically.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 314
diff changeset
   457
(*ML_prf {*
d3c7f6d19c7f Still don't know how to do the proof automatically.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 314
diff changeset
   458
val concl = #concl (fst (Subgoal.focus @{context} 1 (#goal (Isar.goal ()))))
d3c7f6d19c7f Still don't know how to do the proof automatically.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 314
diff changeset
   459
val pat = Drule.strip_imp_concl (cprop_of @{thm APPLY_RSP2})
d3c7f6d19c7f Still don't know how to do the proof automatically.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 314
diff changeset
   460
val insts = Thm.first_order_match (pat, concl)
d3c7f6d19c7f Still don't know how to do the proof automatically.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 314
diff changeset
   461
val t = Drule.instantiate insts @{thm APPLY_RSP2}
d3c7f6d19c7f Still don't know how to do the proof automatically.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 314
diff changeset
   462
*}*)
d3c7f6d19c7f Still don't know how to do the proof automatically.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 314
diff changeset
   463
309
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   464
ML {*
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   465
fun tyRel2 lthy ty gty =
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   466
  if ty = gty then HOLogic.eq_const ty else
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   467
317
d3c7f6d19c7f Still don't know how to do the proof automatically.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 314
diff changeset
   468
  case quotdata_lookup lthy (fst (dest_Type gty)) of
309
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   469
    SOME quotdata =>
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   470
      if Sign.typ_instance (ProofContext.theory_of lthy) (ty, #rtyp quotdata) then
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   471
        case #rel quotdata of
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   472
          Const(s, _) => Const(s, dummyT)
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   473
        | _ => error "tyRel2 fail: relation is not a constant"
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   474
      else error "tyRel2 fail: a non-lifted type lifted to a lifted type"
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   475
  | NONE => (case (ty, gty) of
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   476
      (Type (s, tys), Type (s2, tys2)) =>
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   477
      if s = s2 andalso length tys = length tys2 then
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   478
        let
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   479
          val tys_rel = map (fn ty => ty --> ty --> @{typ bool}) tys;
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   480
          val ty_out = ty --> ty --> @{typ bool};
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   481
          val tys_out = tys_rel ---> ty_out;
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   482
        in
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   483
        (case (maps_lookup (ProofContext.theory_of lthy) s) of
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   484
          SOME (info) => list_comb (Const (#relfun info, tys_out),
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   485
                              map2 (tyRel2 lthy) tys tys2)
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   486
        | NONE  => HOLogic.eq_const ty
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   487
        )
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   488
        end
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   489
      else error "tyRel2 fail: different type structures"
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   490
    | _ => HOLogic.eq_const ty)
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   491
*}
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   492
332
87f5fbebd6d5 Fixes for recent changes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 317
diff changeset
   493
ML mk_babs
87f5fbebd6d5 Fixes for recent changes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 317
diff changeset
   494
87f5fbebd6d5 Fixes for recent changes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 317
diff changeset
   495
ML {*
87f5fbebd6d5 Fixes for recent changes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 317
diff changeset
   496
fun mk_babs ty ty' = Const (@{const_name "Babs"}, [ty' --> @{typ bool}, ty] ---> ty)
87f5fbebd6d5 Fixes for recent changes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 317
diff changeset
   497
fun mk_ball ty = Const (@{const_name "Ball"}, [ty, ty] ---> @{typ bool})
87f5fbebd6d5 Fixes for recent changes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 317
diff changeset
   498
fun mk_bex ty = Const (@{const_name "Bex"}, [ty, ty] ---> @{typ bool})
87f5fbebd6d5 Fixes for recent changes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 317
diff changeset
   499
fun mk_resp ty = Const (@{const_name Respects}, [[ty, ty] ---> @{typ bool}, ty] ---> @{typ bool})
87f5fbebd6d5 Fixes for recent changes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 317
diff changeset
   500
*}
87f5fbebd6d5 Fixes for recent changes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 317
diff changeset
   501
87f5fbebd6d5 Fixes for recent changes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 317
diff changeset
   502
309
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   503
ML {*
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   504
fun my_reg2 lthy trm gtrm =
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   505
  case (trm, gtrm) of
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   506
    (Abs (x, T, t), Abs (x2, T2, t2)) =>
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   507
       if not (T = T2) then
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   508
         let
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   509
           val rrel = tyRel2 lthy T T2;
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   510
           val (s1, s2) = apply_subt2 (my_reg2 lthy) trm gtrm
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   511
         in
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   512
           (((mk_babs (fastype_of trm) T) $ (mk_resp T $ rrel) $ s1),
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   513
           ((mk_babs (fastype_of gtrm) T2) $ (mk_resp T2 $ (HOLogic.eq_const dummyT)) $ s2))
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   514
         end
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   515
       else
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   516
         let
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   517
           val (s1, s2) = apply_subt2 (my_reg2 lthy) t t2
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   518
         in
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   519
           (Abs(x, T, s1), Abs(x2, T2, s2))
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   520
         end
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   521
  | (Const (@{const_name "All"}, ty) $ (t as Abs (_, T, _)),
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   522
     Const (@{const_name "All"}, ty') $ (t2 as Abs (_, T2, _))) =>
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   523
       if not (T = T2) then
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   524
         let
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   525
            val ty1 = domain_type ty;
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   526
            val ty2 = domain_type ty1;
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   527
            val ty'1 = domain_type ty';
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   528
            val ty'2 = domain_type ty'1;
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   529
            val rrel = tyRel2 lthy T T2;
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   530
            val (s1, s2) = apply_subt2 (my_reg2 lthy) t t2;
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   531
         in
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   532
           (((mk_ball ty1) $ (mk_resp ty2 $ rrel) $ s1),
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   533
           ((mk_ball ty'1) $ (mk_resp ty'2 $ (HOLogic.eq_const dummyT)) $ s2))
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   534
         end
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   535
       else
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   536
         let
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   537
           val (s1, s2) = apply_subt2 (my_reg2 lthy) t t2
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   538
         in
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   539
           ((Const (@{const_name "All"}, ty) $ s1),
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   540
           (Const (@{const_name "All"}, ty') $ s2))
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   541
         end
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   542
  | (Const (@{const_name "Ex"}, ty) $ (t as Abs (_, T, _)),
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   543
     Const (@{const_name "Ex"}, ty') $ (t2 as Abs (_, T2, _))) =>
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   544
       if not (T = T2) then
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   545
         let
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   546
            val ty1 = domain_type ty
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   547
            val ty2 = domain_type ty1
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   548
            val ty'1 = domain_type ty'
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   549
            val ty'2 = domain_type ty'1
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   550
            val rrel = tyRel2 lthy T T2
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   551
            val (s1, s2) = apply_subt2 (my_reg2 lthy) t t2
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   552
         in
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   553
           (((mk_bex ty1) $ (mk_resp ty2 $ rrel) $ s1),
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   554
           ((mk_bex ty'1) $ (mk_resp ty'2 $ (HOLogic.eq_const dummyT)) $ s2))
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   555
         end
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   556
       else
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   557
         let
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   558
           val (s1, s2) = apply_subt2 (my_reg2 lthy) t t2
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   559
         in
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   560
           ((Const (@{const_name "Ex"}, ty) $ s1),
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   561
           (Const (@{const_name "Ex"}, ty') $ s2))
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   562
         end
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   563
  | (Const (@{const_name "op ="}, T) $ t, (Const (@{const_name "op ="}, T2) $ t2)) =>
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   564
      let
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   565
        val (s1, s2) = apply_subt2 (my_reg2 lthy) t t2
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   566
        val rhs = Const (@{const_name "op ="}, T2) $ s2
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   567
      in
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   568
        if not (T = T2) then
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   569
          ((tyRel2 lthy T T2) $ s1, rhs)
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   570
        else
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   571
          (Const (@{const_name "op ="}, T) $ s1, rhs)
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   572
      end
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   573
  | (t $ t', t2 $ t2') =>
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   574
      let
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   575
        val (s1, s2) = apply_subt2 (my_reg2 lthy) t t2
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   576
        val (s1', s2') = apply_subt2 (my_reg2 lthy) t' t2'
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   577
      in
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   578
        (s1 $ s1', s2 $ s2')
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   579
      end
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   580
  | (Const c1, Const c2) => (Const c1, Const c2) (* c2 may be lifted *)
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   581
  | (Bound i, Bound j) => (* Bounds are replaced, so should never happen? *)
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   582
      if i = j then (Bound i, Bound j) else error "my_reg2: different Bounds"
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   583
  | (Free (n, T), Free(n2, T2)) => if n = n2 then (Free (n, T), Free (n2, T2))
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   584
      else error ("my_ref2: different variables: " ^ n ^ ", " ^ n2)
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   585
  | _ => error "my_reg2: terms don't agree"
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   586
*}
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   587
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   588
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   589
ML {* prop_of t_a *}
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   590
ML {* term_of glac *}
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   591
ML {* val (tta, ttb) = (my_reg2 @{context} (prop_of t_a) (term_of glac)) *}
333
7851e2a74f85 The other branch does not seem to work...
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 332
diff changeset
   592
7851e2a74f85 The other branch does not seem to work...
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 332
diff changeset
   593
(* Does not work. *)
334
5a7024be9083 code review with Cezary
Christian Urban <urbanc@in.tum.de>
parents: 333
diff changeset
   594
ML {* 
5a7024be9083 code review with Cezary
Christian Urban <urbanc@in.tum.de>
parents: 333
diff changeset
   595
  prop_of t_a 
5a7024be9083 code review with Cezary
Christian Urban <urbanc@in.tum.de>
parents: 333
diff changeset
   596
  |> Syntax.string_of_term @{context}
5a7024be9083 code review with Cezary
Christian Urban <urbanc@in.tum.de>
parents: 333
diff changeset
   597
  |> writeln
5a7024be9083 code review with Cezary
Christian Urban <urbanc@in.tum.de>
parents: 333
diff changeset
   598
*}
5a7024be9083 code review with Cezary
Christian Urban <urbanc@in.tum.de>
parents: 333
diff changeset
   599
5a7024be9083 code review with Cezary
Christian Urban <urbanc@in.tum.de>
parents: 333
diff changeset
   600
ML {* 
5a7024be9083 code review with Cezary
Christian Urban <urbanc@in.tum.de>
parents: 333
diff changeset
   601
  (term_of glac) 
5a7024be9083 code review with Cezary
Christian Urban <urbanc@in.tum.de>
parents: 333
diff changeset
   602
  |> Syntax.string_of_term @{context}
5a7024be9083 code review with Cezary
Christian Urban <urbanc@in.tum.de>
parents: 333
diff changeset
   603
  |> writeln
5a7024be9083 code review with Cezary
Christian Urban <urbanc@in.tum.de>
parents: 333
diff changeset
   604
*}
5a7024be9083 code review with Cezary
Christian Urban <urbanc@in.tum.de>
parents: 333
diff changeset
   605
5a7024be9083 code review with Cezary
Christian Urban <urbanc@in.tum.de>
parents: 333
diff changeset
   606
ML {* val ttar = REGULARIZE_trm @{context} (prop_of t_a) (term_of glac) *} 
333
7851e2a74f85 The other branch does not seem to work...
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 332
diff changeset
   607
309
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   608
ML {* val tt = Syntax.check_term @{context} tta *}
334
5a7024be9083 code review with Cezary
Christian Urban <urbanc@in.tum.de>
parents: 333
diff changeset
   609
ML {* val ttr = Syntax.check_term @{context} ttar *}
5a7024be9083 code review with Cezary
Christian Urban <urbanc@in.tum.de>
parents: 333
diff changeset
   610
5a7024be9083 code review with Cezary
Christian Urban <urbanc@in.tum.de>
parents: 333
diff changeset
   611
309
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   612
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   613
prove t_r: {* Logic.mk_implies
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   614
       ((prop_of t_a), tt) *}
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   615
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
   616
      rtac rel_refl,
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   617
      atac,
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   618
      rtac @{thm universal_twice},
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   619
      (rtac @{thm impI} THEN' atac),
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   620
      rtac @{thm implication_twice},
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   621
      (*rtac @{thm equality_twice},*)
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   622
      EqSubst.eqsubst_tac ctxt [0]
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   623
        [(@{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
   624
         (@{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
   625
      (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
   626
      (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
   627
     ]; *}
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   628
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   629
  apply (atomize(full))
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   630
  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
   631
  done
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   632
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   633
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
   634
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   635
ML {* val ttg = Syntax.check_term @{context} ttb *}
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   636
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   637
ML {*
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   638
fun is_lifted_const h gh = is_Const h andalso is_Const gh andalso not (h = gh)
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   639
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   640
fun mkrepabs lthy ty gty t =
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   641
  let
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   642
    val qenv = distinct (op=) (diff (gty, ty) [])
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   643
(*  val _ = sanity_chk qenv lthy *)
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   644
    val ty = fastype_of t
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   645
    val abs = get_fun absF qenv lthy gty $ t
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   646
    val rep = get_fun repF qenv lthy gty $ abs
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   647
  in
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   648
    Syntax.check_term lthy rep
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   649
  end
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   650
*}
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   651
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   652
ML {*
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   653
  cterm_of @{theory} (mkrepabs @{context} @{typ "'a list \<Rightarrow> bool"} @{typ "'a fset \<Rightarrow> bool"} @{term "f :: ('a list \<Rightarrow> bool)"})
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   654
*}
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   655
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   656
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   657
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   658
ML {*
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   659
fun build_repabs_term lthy trm gtrm =
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   660
  case (trm, gtrm) of
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   661
    (Abs (a as (_, T, _)), Abs (a2 as (_, T2, _))) =>
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   662
      let
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   663
        val (vs, t) = Term.dest_abs a;
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   664
        val (_,  g) = Term.dest_abs a2;
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   665
        val v = Free(vs, T);
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   666
        val t' = lambda v (build_repabs_term lthy t g);
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   667
        val ty = fastype_of trm;
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   668
        val gty = fastype_of gtrm;
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   669
      in
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   670
        if (ty = gty) then t' else
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   671
        mkrepabs lthy ty gty (
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   672
          if (T = T2) then t' else
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   673
          lambda v (t' $ (mkrepabs lthy T T2 v))
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   674
        )
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   675
      end
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   676
  | _ =>
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   677
    case (Term.strip_comb trm, Term.strip_comb gtrm) of
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   678
      ((Const(@{const_name Respects}, _), _), _) => trm
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   679
    | ((h, tms), (gh, gtms)) =>
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   680
      let
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   681
        val ty = fastype_of trm;
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   682
        val gty = fastype_of gtrm;
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   683
        val tms' = map2 (build_repabs_term lthy) tms gtms
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   684
        val t' = list_comb(h, tms')
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   685
      in
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   686
        if ty = gty then t' else
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   687
        if is_lifted_const h gh then mkrepabs lthy ty gty t' else
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   688
        if (Term.is_Free h) andalso (length tms > 0) then mkrepabs lthy ty gty t' else t'
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   689
      end
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   690
*}
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   691
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   692
ML {* val ttt = build_repabs_term @{context} tt ttg *}
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   693
ML {* val si = simp_ids_trm (cterm_of @{theory} ttt) *}
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   694
prove t_t: {* Logic.mk_equals ((Thm.prop_of t_r), term_of si) *}
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   695
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
   696
apply(atomize(full))
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   697
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
   698
apply (rule FUN_QUOTIENT)
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   699
apply (rule FUN_QUOTIENT)
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   700
apply (rule IDENTITY_QUOTIENT)
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   701
apply (rule FUN_QUOTIENT)
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   702
apply (rule QUOTIENT_fset)
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   703
apply (rule IDENTITY_QUOTIENT)
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   704
apply (rule IDENTITY_QUOTIENT)
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   705
apply (rule IDENTITY_QUOTIENT)
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   706
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
   707
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
   708
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
   709
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
   710
apply (rule IDENTITY_QUOTIENT)
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   711
apply (rule IDENTITY_QUOTIENT)
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   712
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
   713
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
   714
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
   715
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
   716
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
   717
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
   718
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
   719
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
   720
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
   721
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
   722
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
   723
apply (rule IDENTITY_QUOTIENT)
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   724
apply (rule FUN_QUOTIENT)
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   725
apply (rule QUOTIENT_fset)
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   726
apply (rule IDENTITY_QUOTIENT)
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   727
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
   728
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
   729
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
   730
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
   731
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
   732
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
   733
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
   734
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
   735
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
   736
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
   737
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
   738
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
   739
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
   740
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
   741
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
   742
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
   743
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
   744
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
   745
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
   746
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
   747
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
   748
apply assumption
d3c7f6d19c7f Still don't know how to do the proof automatically.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 314
diff changeset
   749
apply (rule refl)
309
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   750
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
   751
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
   752
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
   753
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
   754
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
   755
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
   756
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
   757
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
   758
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
   759
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
   760
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
   761
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
   762
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
   763
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
   764
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
   765
done
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   766
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   767
thm t_t
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
   768
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
   769
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
   770
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
   771
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
   772
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
   773
178
945786a68ec6 Finally lifted induction, with some manually added simplification lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 176
diff changeset
   774
163
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   775
ML {*
226
2a28e7ef3048 cleaned FSet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 225
diff changeset
   776
  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
   777
    let
226
2a28e7ef3048 cleaned FSet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 225
diff changeset
   778
      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
   779
      val (_, lthy2) = note (name, lifted_thm) lthy;
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   780
    in
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   781
      lthy2
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   782
    end;
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   783
*}
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   784
226
2a28e7ef3048 cleaned FSet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 225
diff changeset
   785
local_setup {*
2a28e7ef3048 cleaned FSet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 225
diff changeset
   786
  lift_thm_fset_note @{binding "m1l"} @{thm m1} #>
2a28e7ef3048 cleaned FSet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 225
diff changeset
   787
  lift_thm_fset_note @{binding "m2l"} @{thm m2} #>
2a28e7ef3048 cleaned FSet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 225
diff changeset
   788
  lift_thm_fset_note @{binding "leqi4l"} @{thm list_eq.intros(4)} #>
2a28e7ef3048 cleaned FSet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 225
diff changeset
   789
  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
   790
*}
226
2a28e7ef3048 cleaned FSet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 225
diff changeset
   791
thm m1l
2a28e7ef3048 cleaned FSet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 225
diff changeset
   792
thm m2l
2a28e7ef3048 cleaned FSet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 225
diff changeset
   793
thm leqi4l
2a28e7ef3048 cleaned FSet
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 225
diff changeset
   794
thm leqi5l
163
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   795
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   796
end