Quot/Examples/FSet.thy
author Christian Urban <urbanc@in.tum.de>
Sun, 24 Jan 2010 23:41:27 +0100
changeset 918 7be9b054f672
parent 787 5cf83fa5b36c
child 931 0879d144aaa3
permissions -rw-r--r--
test with splits
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
918
7be9b054f672 test with splits
Christian Urban <urbanc@in.tum.de>
parents: 787
diff changeset
     2
imports "../QuotMain" "../QuotList" "../QuotProd" List
163
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     3
begin
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     4
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     5
inductive
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     6
  list_eq (infix "\<approx>" 50)
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     7
where
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     8
  "a#b#xs \<approx> b#a#xs"
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     9
| "[] \<approx> []"
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    10
| "xs \<approx> ys \<Longrightarrow> ys \<approx> xs"
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    11
| "a#a#xs \<approx> a#xs"
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    12
| "xs \<approx> ys \<Longrightarrow> a#xs \<approx> a#ys"
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    13
| "\<lbrakk>xs1 \<approx> xs2; xs2 \<approx> xs3\<rbrakk> \<Longrightarrow> xs1 \<approx> xs3"
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    14
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    15
lemma list_eq_refl:
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    16
  shows "xs \<approx> xs"
451
586e3dc4afdb Added 'TRY' to refl in clean_tac to get as far as possible. Removed unnecessary [quot_rsp] in FSet. Added necessary [quot_rsp] and one lifted thm in LamEx.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 450
diff changeset
    17
  by (induct xs) (auto intro: list_eq.intros)
163
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    18
529
6348c2a57ec2 More name changes
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 526
diff changeset
    19
lemma equivp_list_eq:
6348c2a57ec2 More name changes
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 526
diff changeset
    20
  shows "equivp list_eq"
6348c2a57ec2 More name changes
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 526
diff changeset
    21
  unfolding equivp_reflp_symp_transp reflp_def symp_def transp_def
163
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    22
  apply(auto intro: list_eq.intros list_eq_refl)
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    23
  done
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    24
766
df053507edba renamed "quotient" command to "quotient_type"; needs new keyword file to be installed
Christian Urban <urbanc@in.tum.de>
parents: 758
diff changeset
    25
quotient_type 
787
5cf83fa5b36c made the quotient_type definition more like typedef; now type variables need to be explicitly given
Christian Urban <urbanc@in.tum.de>
parents: 779
diff changeset
    26
  'a fset = "'a list" / "list_eq"
766
df053507edba renamed "quotient" command to "quotient_type"; needs new keyword file to be installed
Christian Urban <urbanc@in.tum.de>
parents: 758
diff changeset
    27
  by (rule equivp_list_eq)
163
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    28
767
37285ec4387d on suggestion of Tobias renamed "quotient_def" to "quotient_definition"; needs new keyword file
Christian Urban <urbanc@in.tum.de>
parents: 766
diff changeset
    29
quotient_definition
705
f51c6069cd17 New syntax for definitions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 697
diff changeset
    30
   "EMPTY :: 'a fset"
f51c6069cd17 New syntax for definitions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 697
diff changeset
    31
as
663
0dd10a900cae Different syntax for definitions that allows overloading and retrieving of definitions by matching whole constants.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 656
diff changeset
    32
  "[]::'a list"
163
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    33
767
37285ec4387d on suggestion of Tobias renamed "quotient_def" to "quotient_definition"; needs new keyword file
Christian Urban <urbanc@in.tum.de>
parents: 766
diff changeset
    34
quotient_definition
705
f51c6069cd17 New syntax for definitions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 697
diff changeset
    35
   "INSERT :: 'a \<Rightarrow> 'a fset \<Rightarrow> 'a fset"
f51c6069cd17 New syntax for definitions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 697
diff changeset
    36
as
663
0dd10a900cae Different syntax for definitions that allows overloading and retrieving of definitions by matching whole constants.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 656
diff changeset
    37
  "op #"
163
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    38
767
37285ec4387d on suggestion of Tobias renamed "quotient_def" to "quotient_definition"; needs new keyword file
Christian Urban <urbanc@in.tum.de>
parents: 766
diff changeset
    39
quotient_definition
705
f51c6069cd17 New syntax for definitions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 697
diff changeset
    40
   "FUNION :: 'a fset \<Rightarrow> 'a fset \<Rightarrow> 'a fset"
f51c6069cd17 New syntax for definitions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 697
diff changeset
    41
as
663
0dd10a900cae Different syntax for definitions that allows overloading and retrieving of definitions by matching whole constants.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 656
diff changeset
    42
  "op @"
163
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    43
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    44
fun
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    45
  card1 :: "'a list \<Rightarrow> nat"
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    46
where
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    47
  card1_nil: "(card1 []) = 0"
683
0d9e8aa1bc7a removed memb and used standard mem (member from List.thy)
Christian Urban <urbanc@in.tum.de>
parents: 681
diff changeset
    48
| card1_cons: "(card1 (x # xs)) = (if (x mem xs) then (card1 xs) else (Suc (card1 xs)))"
163
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    49
767
37285ec4387d on suggestion of Tobias renamed "quotient_def" to "quotient_definition"; needs new keyword file
Christian Urban <urbanc@in.tum.de>
parents: 766
diff changeset
    50
quotient_definition
705
f51c6069cd17 New syntax for definitions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 697
diff changeset
    51
   "CARD :: 'a fset \<Rightarrow> nat"
f51c6069cd17 New syntax for definitions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 697
diff changeset
    52
as
663
0dd10a900cae Different syntax for definitions that allows overloading and retrieving of definitions by matching whole constants.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 656
diff changeset
    53
  "card1"
163
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    54
768
e9e205b904e2 get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
    55
quotient_definition
e9e205b904e2 get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
    56
  "fconcat :: ('a fset) fset \<Rightarrow> 'a fset"
e9e205b904e2 get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
    57
as
e9e205b904e2 get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
    58
  "concat"
e9e205b904e2 get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
    59
e9e205b904e2 get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
    60
term concat
e9e205b904e2 get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
    61
term fconcat
e9e205b904e2 get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
    62
e9e205b904e2 get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
    63
text {*
163
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    64
 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
    65
 respects the relation. With it such a definition would be impossible:
768
e9e205b904e2 get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
    66
 make_const_def CARD @{term "length"} NoSyn @{typ "'a list"} @{typ "'a fset"} #> snd
e9e205b904e2 get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
    67
*}
163
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
lemma card1_0:
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    70
  fixes a :: "'a list"
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    71
  shows "(card1 a = 0) = (a = [])"
214
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
    72
  by (induct a) auto
163
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    73
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    74
lemma not_mem_card1:
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    75
  fixes x :: "'a"
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    76
  fixes xs :: "'a list"
683
0d9e8aa1bc7a removed memb and used standard mem (member from List.thy)
Christian Urban <urbanc@in.tum.de>
parents: 681
diff changeset
    77
  shows "(~(x mem xs)) = (card1 (x # xs) = Suc (card1 xs))"
309
20fa8dd8fb93 Lifting towards goal and manually finished the proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 305
diff changeset
    78
  by auto
163
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    79
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    80
lemma mem_cons:
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    81
  fixes x :: "'a"
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    82
  fixes xs :: "'a list"
683
0d9e8aa1bc7a removed memb and used standard mem (member from List.thy)
Christian Urban <urbanc@in.tum.de>
parents: 681
diff changeset
    83
  assumes a : "x mem xs"
163
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    84
  shows "x # xs \<approx> xs"
214
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
    85
  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
    86
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    87
lemma card1_suc:
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    88
  fixes xs :: "'a list"
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    89
  fixes n :: "nat"
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    90
  assumes c: "card1 xs = Suc n"
683
0d9e8aa1bc7a removed memb and used standard mem (member from List.thy)
Christian Urban <urbanc@in.tum.de>
parents: 681
diff changeset
    91
  shows "\<exists>a ys. ~(a mem ys) \<and> xs \<approx> (a # ys)"
163
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    92
  using c
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    93
apply(induct xs)
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    94
apply (metis Suc_neq_Zero card1_0)
685
b12f0321dfb0 moved the interpretation code into Unused.thy
Christian Urban <urbanc@in.tum.de>
parents: 683
diff changeset
    95
apply (metis FSet.card1_cons list_eq.intros(6) list_eq_refl mem_cons)
163
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    96
done
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    97
294
a092c0b13d83 fold_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 292
diff changeset
    98
definition
a092c0b13d83 fold_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 292
diff changeset
    99
  rsp_fold
a092c0b13d83 fold_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 292
diff changeset
   100
where
a092c0b13d83 fold_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 292
diff changeset
   101
  "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
   102
163
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   103
primrec
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   104
  fold1
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   105
where
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   106
  "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
   107
| "fold1 f g z (a # A) =
294
a092c0b13d83 fold_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 292
diff changeset
   108
     (if rsp_fold f
163
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   109
     then (
683
0d9e8aa1bc7a removed memb and used standard mem (member from List.thy)
Christian Urban <urbanc@in.tum.de>
parents: 681
diff changeset
   110
       if (a mem A) then (fold1 f g z A) else (f (g a) (fold1 f g z A))
163
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   111
     ) else z)"
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   112
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   113
lemma fs1_strong_cases:
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   114
  fixes X :: "'a list"
683
0d9e8aa1bc7a removed memb and used standard mem (member from List.thy)
Christian Urban <urbanc@in.tum.de>
parents: 681
diff changeset
   115
  shows "(X = []) \<or> (\<exists>a. \<exists> Y. (~(a mem Y) \<and> (X \<approx> a # Y)))"
163
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   116
  apply (induct X)
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   117
  apply (simp)
685
b12f0321dfb0 moved the interpretation code into Unused.thy
Christian Urban <urbanc@in.tum.de>
parents: 683
diff changeset
   118
  apply (metis List.member.simps(1) list_eq.intros(6) list_eq_refl mem_cons)
163
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   119
  done
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   120
767
37285ec4387d on suggestion of Tobias renamed "quotient_def" to "quotient_definition"; needs new keyword file
Christian Urban <urbanc@in.tum.de>
parents: 766
diff changeset
   121
quotient_definition
705
f51c6069cd17 New syntax for definitions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 697
diff changeset
   122
   "IN :: 'a \<Rightarrow> 'a fset \<Rightarrow> bool"
f51c6069cd17 New syntax for definitions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 697
diff changeset
   123
as
683
0d9e8aa1bc7a removed memb and used standard mem (member from List.thy)
Christian Urban <urbanc@in.tum.de>
parents: 681
diff changeset
   124
  "op mem"
163
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   125
767
37285ec4387d on suggestion of Tobias renamed "quotient_def" to "quotient_definition"; needs new keyword file
Christian Urban <urbanc@in.tum.de>
parents: 766
diff changeset
   126
quotient_definition
705
f51c6069cd17 New syntax for definitions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 697
diff changeset
   127
   "FOLD :: ('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> ('b \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'b fset \<Rightarrow> 'a"
f51c6069cd17 New syntax for definitions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 697
diff changeset
   128
as
663
0dd10a900cae Different syntax for definitions that allows overloading and retrieving of definitions by matching whole constants.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 656
diff changeset
   129
  "fold1"
194
03c03e88efa9 Simplifying Int and Working on map
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 190
diff changeset
   130
767
37285ec4387d on suggestion of Tobias renamed "quotient_def" to "quotient_definition"; needs new keyword file
Christian Urban <urbanc@in.tum.de>
parents: 766
diff changeset
   131
quotient_definition
705
f51c6069cd17 New syntax for definitions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 697
diff changeset
   132
  "fmap :: ('a \<Rightarrow> 'b) \<Rightarrow> 'a fset \<Rightarrow> 'b fset"
f51c6069cd17 New syntax for definitions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 697
diff changeset
   133
as
663
0dd10a900cae Different syntax for definitions that allows overloading and retrieving of definitions by matching whole constants.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 656
diff changeset
   134
  "map"
194
03c03e88efa9 Simplifying Int and Working on map
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 190
diff changeset
   135
683
0d9e8aa1bc7a removed memb and used standard mem (member from List.thy)
Christian Urban <urbanc@in.tum.de>
parents: 681
diff changeset
   136
lemma mem_rsp:
163
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   137
  fixes z
450
2dc708ddb93a introduced a global list of respectfulness lemmas; the attribute is [quot_rsp]
Christian Urban <urbanc@in.tum.de>
parents: 448
diff changeset
   138
  assumes a: "x \<approx> y"
683
0d9e8aa1bc7a removed memb and used standard mem (member from List.thy)
Christian Urban <urbanc@in.tum.de>
parents: 681
diff changeset
   139
  shows "(z mem x) = (z mem y)"
163
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   140
  using a by induct auto
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   141
636
520a4084d064 changed names of attributes
Christian Urban <urbanc@in.tum.de>
parents: 631
diff changeset
   142
lemma ho_memb_rsp[quot_respect]:
683
0d9e8aa1bc7a removed memb and used standard mem (member from List.thy)
Christian Urban <urbanc@in.tum.de>
parents: 681
diff changeset
   143
  "(op = ===> (op \<approx> ===> op =)) (op mem) (op mem)"
0d9e8aa1bc7a removed memb and used standard mem (member from List.thy)
Christian Urban <urbanc@in.tum.de>
parents: 681
diff changeset
   144
  by (simp add: mem_rsp)
164
4f00ca4f5ef4 Stronger tactic, simpler proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 163
diff changeset
   145
451
586e3dc4afdb Added 'TRY' to refl in clean_tac to get as far as possible. Removed unnecessary [quot_rsp] in FSet. Added necessary [quot_rsp] and one lifted thm in LamEx.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 450
diff changeset
   146
lemma card1_rsp:
163
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   147
  fixes a b :: "'a list"
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   148
  assumes e: "a \<approx> b"
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   149
  shows "card1 a = card1 b"
683
0d9e8aa1bc7a removed memb and used standard mem (member from List.thy)
Christian Urban <urbanc@in.tum.de>
parents: 681
diff changeset
   150
  using e by induct (simp_all add: mem_rsp)
163
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   151
636
520a4084d064 changed names of attributes
Christian Urban <urbanc@in.tum.de>
parents: 631
diff changeset
   152
lemma ho_card1_rsp[quot_respect]: 
450
2dc708ddb93a introduced a global list of respectfulness lemmas; the attribute is [quot_rsp]
Christian Urban <urbanc@in.tum.de>
parents: 448
diff changeset
   153
  "(op \<approx> ===> op =) card1 card1"
214
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   154
  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
   155
680
d003f9e00c29 removed quot_respect attribute of a non-standard lemma
Christian Urban <urbanc@in.tum.de>
parents: 664
diff changeset
   156
lemma cons_rsp:
163
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   157
  fixes z
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   158
  assumes a: "xs \<approx> ys"
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   159
  shows "(z # xs) \<approx> (z # ys)"
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   160
  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
   161
636
520a4084d064 changed names of attributes
Christian Urban <urbanc@in.tum.de>
parents: 631
diff changeset
   162
lemma ho_cons_rsp[quot_respect]:
228
268a727b0f10 disambiguate ===> syntax
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 226
diff changeset
   163
  "(op = ===> op \<approx> ===> op \<approx>) op # op #"
214
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   164
  by (simp add: cons_rsp)
164
4f00ca4f5ef4 Stronger tactic, simpler proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 163
diff changeset
   165
681
3c6419a5a7fc simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 680
diff changeset
   166
lemma append_rsp_aux1:
3c6419a5a7fc simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 680
diff changeset
   167
  assumes a : "l2 \<approx> r2 "
3c6419a5a7fc simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 680
diff changeset
   168
  shows "(h @ l2) \<approx> (h @ r2)"
3c6419a5a7fc simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 680
diff changeset
   169
using a
3c6419a5a7fc simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 680
diff changeset
   170
apply(induct h)
3c6419a5a7fc simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 680
diff changeset
   171
apply(auto intro: list_eq.intros(5))
3c6419a5a7fc simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 680
diff changeset
   172
done
163
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   173
681
3c6419a5a7fc simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 680
diff changeset
   174
lemma append_rsp_aux2:
3c6419a5a7fc simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 680
diff changeset
   175
  assumes a : "l1 \<approx> r1" "l2 \<approx> r2 "
3c6419a5a7fc simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 680
diff changeset
   176
  shows "(l1 @ l2) \<approx> (r1 @ r2)"
3c6419a5a7fc simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 680
diff changeset
   177
using a
3c6419a5a7fc simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 680
diff changeset
   178
apply(induct arbitrary: l2 r2)
3c6419a5a7fc simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 680
diff changeset
   179
apply(simp_all)
3c6419a5a7fc simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 680
diff changeset
   180
apply(blast intro: list_eq.intros append_rsp_aux1)+
3c6419a5a7fc simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 680
diff changeset
   181
done
214
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   182
681
3c6419a5a7fc simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 680
diff changeset
   183
lemma append_rsp[quot_respect]:
228
268a727b0f10 disambiguate ===> syntax
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 226
diff changeset
   184
  "(op \<approx> ===> op \<approx> ===> op \<approx>) op @ op @"
681
3c6419a5a7fc simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 680
diff changeset
   185
  by (auto simp add: append_rsp_aux2)
175
f7602653dddd Preparing infrastructire for LAMBDA_PRS
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 173
diff changeset
   186
451
586e3dc4afdb Added 'TRY' to refl in clean_tac to get as far as possible. Removed unnecessary [quot_rsp] in FSet. Added necessary [quot_rsp] and one lifted thm in LamEx.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 450
diff changeset
   187
lemma map_rsp:
194
03c03e88efa9 Simplifying Int and Working on map
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 190
diff changeset
   188
  assumes a: "a \<approx> b"
03c03e88efa9 Simplifying Int and Working on map
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 190
diff changeset
   189
  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
   190
  using a
03c03e88efa9 Simplifying Int and Working on map
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 190
diff changeset
   191
  apply (induct)
03c03e88efa9 Simplifying Int and Working on map
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 190
diff changeset
   192
  apply(auto intro: list_eq.intros)
03c03e88efa9 Simplifying Int and Working on map
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 190
diff changeset
   193
  done
03c03e88efa9 Simplifying Int and Working on map
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 190
diff changeset
   194
636
520a4084d064 changed names of attributes
Christian Urban <urbanc@in.tum.de>
parents: 631
diff changeset
   195
lemma ho_map_rsp[quot_respect]:
294
a092c0b13d83 fold_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 292
diff changeset
   196
  "(op = ===> op \<approx> ===> op \<approx>) map map"
a092c0b13d83 fold_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 292
diff changeset
   197
  by (simp add: map_rsp)
194
03c03e88efa9 Simplifying Int and Working on map
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 190
diff changeset
   198
294
a092c0b13d83 fold_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 292
diff changeset
   199
lemma map_append:
450
2dc708ddb93a introduced a global list of respectfulness lemmas; the attribute is [quot_rsp]
Christian Urban <urbanc@in.tum.de>
parents: 448
diff changeset
   200
  "(map f (a @ b)) \<approx> (map f a) @ (map f b)"
215
89a2ff3f82c7 More finshed proofs and cleaning
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 214
diff changeset
   201
 by simp (rule list_eq_refl)
194
03c03e88efa9 Simplifying Int and Working on map
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 190
diff changeset
   202
636
520a4084d064 changed names of attributes
Christian Urban <urbanc@in.tum.de>
parents: 631
diff changeset
   203
lemma ho_fold_rsp[quot_respect]:
294
a092c0b13d83 fold_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 292
diff changeset
   204
  "(op = ===> op = ===> op = ===> op \<approx> ===> op =) fold1 fold1"
536
44fa9df44e6f More code cleaning and name changes
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 529
diff changeset
   205
  apply (auto)
294
a092c0b13d83 fold_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 292
diff changeset
   206
  apply (case_tac "rsp_fold x")
a092c0b13d83 fold_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 292
diff changeset
   207
  prefer 2
a092c0b13d83 fold_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 292
diff changeset
   208
  apply (erule_tac list_eq.induct)
a092c0b13d83 fold_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 292
diff changeset
   209
  apply (simp_all)
a092c0b13d83 fold_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 292
diff changeset
   210
  apply (erule_tac list_eq.induct)
a092c0b13d83 fold_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 292
diff changeset
   211
  apply (simp_all)
683
0d9e8aa1bc7a removed memb and used standard mem (member from List.thy)
Christian Urban <urbanc@in.tum.de>
parents: 681
diff changeset
   212
  apply (auto simp add: mem_rsp rsp_fold_def)
294
a092c0b13d83 fold_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 292
diff changeset
   213
done
241
60acf3d3a4a0 Finding applications and duplicates filtered out in abstractions
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 239
diff changeset
   214
636
520a4084d064 changed names of attributes
Christian Urban <urbanc@in.tum.de>
parents: 631
diff changeset
   215
lemma list_equiv_rsp[quot_respect]:
549
f178958d3d81 not yet quite functional treatment of constants
Christian Urban <urbanc@in.tum.de>
parents: 536
diff changeset
   216
  shows "(op \<approx> ===> op \<approx> ===> op =) op \<approx> op \<approx>"
f178958d3d81 not yet quite functional treatment of constants
Christian Urban <urbanc@in.tum.de>
parents: 536
diff changeset
   217
by (auto intro: list_eq.intros)
f178958d3d81 not yet quite functional treatment of constants
Christian Urban <urbanc@in.tum.de>
parents: 536
diff changeset
   218
364
4c455d58ac99 Fixes to the tactic after quotient_tac changed.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 356
diff changeset
   219
lemma "IN x EMPTY = False"
683
0d9e8aa1bc7a removed memb and used standard mem (member from List.thy)
Christian Urban <urbanc@in.tum.de>
parents: 681
diff changeset
   220
apply(lifting member.simps(1))
455
9cb45d022524 tried to improve the inj_repabs_trm function but left the new part commented out
Christian Urban <urbanc@in.tum.de>
parents: 452
diff changeset
   221
done
353
9a0e8ab42ee8 fixed the error by a temporary fix (the data of the eqivalence relation should be only its name)
Christian Urban <urbanc@in.tum.de>
parents: 350
diff changeset
   222
364
4c455d58ac99 Fixes to the tactic after quotient_tac changed.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 356
diff changeset
   223
lemma "IN x (INSERT y xa) = (x = y \<or> IN x xa)"
768
e9e205b904e2 get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   224
apply (lifting member.simps(2))
e9e205b904e2 get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   225
done
356
51aafebf4d06 Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 353
diff changeset
   226
364
4c455d58ac99 Fixes to the tactic after quotient_tac changed.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 356
diff changeset
   227
lemma "INSERT a (INSERT a x) = INSERT a x"
654
02fd9de9d45e tuned the examples and flagged the problematic cleaning lemmas in FSet
Christian Urban <urbanc@in.tum.de>
parents: 653
diff changeset
   228
apply (lifting list_eq.intros(4))
364
4c455d58ac99 Fixes to the tactic after quotient_tac changed.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 356
diff changeset
   229
done
4c455d58ac99 Fixes to the tactic after quotient_tac changed.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 356
diff changeset
   230
367
d444389fe3f9 The non-working procedure_tac.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 364
diff changeset
   231
lemma "x = xa \<Longrightarrow> INSERT a x = INSERT a xa"
654
02fd9de9d45e tuned the examples and flagged the problematic cleaning lemmas in FSet
Christian Urban <urbanc@in.tum.de>
parents: 653
diff changeset
   232
apply (lifting list_eq.intros(5))
364
4c455d58ac99 Fixes to the tactic after quotient_tac changed.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 356
diff changeset
   233
done
353
9a0e8ab42ee8 fixed the error by a temporary fix (the data of the eqivalence relation should be only its name)
Christian Urban <urbanc@in.tum.de>
parents: 350
diff changeset
   234
367
d444389fe3f9 The non-working procedure_tac.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 364
diff changeset
   235
lemma "CARD x = Suc n \<Longrightarrow> (\<exists>a b. \<not> IN a b & x = INSERT a b)"
654
02fd9de9d45e tuned the examples and flagged the problematic cleaning lemmas in FSet
Christian Urban <urbanc@in.tum.de>
parents: 653
diff changeset
   236
apply (lifting card1_suc)
364
4c455d58ac99 Fixes to the tactic after quotient_tac changed.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 356
diff changeset
   237
done
4c455d58ac99 Fixes to the tactic after quotient_tac changed.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 356
diff changeset
   238
4c455d58ac99 Fixes to the tactic after quotient_tac changed.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 356
diff changeset
   239
lemma "(\<not> IN x xa) = (CARD (INSERT x xa) = Suc (CARD xa))"
654
02fd9de9d45e tuned the examples and flagged the problematic cleaning lemmas in FSet
Christian Urban <urbanc@in.tum.de>
parents: 653
diff changeset
   240
apply (lifting not_mem_card1)
364
4c455d58ac99 Fixes to the tactic after quotient_tac changed.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 356
diff changeset
   241
done
356
51aafebf4d06 Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 353
diff changeset
   242
442
7beed9b75ea2 renamed LAMBDA_RES_TAC and WEAK_LAMBDA_RES_TAC to lower case names
Christian Urban <urbanc@in.tum.de>
parents: 435
diff changeset
   243
lemma "FOLD f g (z::'b) (INSERT a x) =
364
4c455d58ac99 Fixes to the tactic after quotient_tac changed.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 356
diff changeset
   244
  (if rsp_fold f then if IN a x then FOLD f g z x else f (g a) (FOLD f g z x) else z)"
654
02fd9de9d45e tuned the examples and flagged the problematic cleaning lemmas in FSet
Christian Urban <urbanc@in.tum.de>
parents: 653
diff changeset
   245
apply(lifting fold1.simps(2))
364
4c455d58ac99 Fixes to the tactic after quotient_tac changed.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 356
diff changeset
   246
done
356
51aafebf4d06 Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 353
diff changeset
   247
368
c5c49d240cde Conversion
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 367
diff changeset
   248
lemma "fmap f (FUNION (x::'b fset) (xa::'b fset)) = FUNION (fmap f x) (fmap f xa)"
654
02fd9de9d45e tuned the examples and flagged the problematic cleaning lemmas in FSet
Christian Urban <urbanc@in.tum.de>
parents: 653
diff changeset
   249
apply (lifting map_append)
368
c5c49d240cde Conversion
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 367
diff changeset
   250
done
c5c49d240cde Conversion
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 367
diff changeset
   251
367
d444389fe3f9 The non-working procedure_tac.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 364
diff changeset
   252
lemma "FUNION (FUNION x xa) xb = FUNION x (FUNION xa xb)"
654
02fd9de9d45e tuned the examples and flagged the problematic cleaning lemmas in FSet
Christian Urban <urbanc@in.tum.de>
parents: 653
diff changeset
   253
apply (lifting append_assoc)
367
d444389fe3f9 The non-working procedure_tac.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 364
diff changeset
   254
done
d444389fe3f9 The non-working procedure_tac.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 364
diff changeset
   255
477
6c88b42da228 A bit of progress; but the object-logic vs meta-logic distinction is troublesome.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 475
diff changeset
   256
376
e99c0334d8bf lambda_prs and cleaning the existing examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 375
diff changeset
   257
lemma "\<lbrakk>P EMPTY; \<And>a x. P x \<Longrightarrow> P (INSERT a x)\<rbrakk> \<Longrightarrow> P l"
654
02fd9de9d45e tuned the examples and flagged the problematic cleaning lemmas in FSet
Christian Urban <urbanc@in.tum.de>
parents: 653
diff changeset
   258
apply(lifting list.induct)
414
4dad34ca50db Minor cleaning
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 401
diff changeset
   259
done
390
1dd6a21cdd1c test with monos
Christian Urban <urbanc@in.tum.de>
parents: 387
diff changeset
   260
482
767baada01dc New APPLY_RSP which finally does automatic partial lifting :). Doesn't support same relation yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 481
diff changeset
   261
lemma list_induct_part:
767baada01dc New APPLY_RSP which finally does automatic partial lifting :). Doesn't support same relation yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 481
diff changeset
   262
  assumes a: "P (x :: 'a list) ([] :: 'c list)"
767baada01dc New APPLY_RSP which finally does automatic partial lifting :). Doesn't support same relation yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 481
diff changeset
   263
  assumes b: "\<And>e t. P x t \<Longrightarrow> P x (e # t)"
767baada01dc New APPLY_RSP which finally does automatic partial lifting :). Doesn't support same relation yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 481
diff changeset
   264
  shows "P x l"
767baada01dc New APPLY_RSP which finally does automatic partial lifting :). Doesn't support same relation yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 481
diff changeset
   265
  apply (rule_tac P="P x" in list.induct)
767baada01dc New APPLY_RSP which finally does automatic partial lifting :). Doesn't support same relation yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 481
diff changeset
   266
  apply (rule a)
767baada01dc New APPLY_RSP which finally does automatic partial lifting :). Doesn't support same relation yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 481
diff changeset
   267
  apply (rule b)
767baada01dc New APPLY_RSP which finally does automatic partial lifting :). Doesn't support same relation yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 481
diff changeset
   268
  apply (assumption)
767baada01dc New APPLY_RSP which finally does automatic partial lifting :). Doesn't support same relation yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 481
diff changeset
   269
  done
767baada01dc New APPLY_RSP which finally does automatic partial lifting :). Doesn't support same relation yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 481
diff changeset
   270
767baada01dc New APPLY_RSP which finally does automatic partial lifting :). Doesn't support same relation yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 481
diff changeset
   271
lemma "P (x :: 'a list) (EMPTY :: 'c fset) \<Longrightarrow> (\<And>e t. P x t \<Longrightarrow> P x (INSERT e t)) \<Longrightarrow> P x l"
654
02fd9de9d45e tuned the examples and flagged the problematic cleaning lemmas in FSet
Christian Urban <urbanc@in.tum.de>
parents: 653
diff changeset
   272
apply (lifting list_induct_part)
482
767baada01dc New APPLY_RSP which finally does automatic partial lifting :). Doesn't support same relation yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 481
diff changeset
   273
done
767baada01dc New APPLY_RSP which finally does automatic partial lifting :). Doesn't support same relation yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 481
diff changeset
   274
767baada01dc New APPLY_RSP which finally does automatic partial lifting :). Doesn't support same relation yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 481
diff changeset
   275
lemma "P (x :: 'a fset) (EMPTY :: 'c fset) \<Longrightarrow> (\<And>e t. P x t \<Longrightarrow> P x (INSERT e t)) \<Longrightarrow> P x l"
654
02fd9de9d45e tuned the examples and flagged the problematic cleaning lemmas in FSet
Christian Urban <urbanc@in.tum.de>
parents: 653
diff changeset
   276
apply (lifting list_induct_part)
482
767baada01dc New APPLY_RSP which finally does automatic partial lifting :). Doesn't support same relation yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 481
diff changeset
   277
done
767baada01dc New APPLY_RSP which finally does automatic partial lifting :). Doesn't support same relation yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 481
diff changeset
   278
767baada01dc New APPLY_RSP which finally does automatic partial lifting :). Doesn't support same relation yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 481
diff changeset
   279
lemma "P (x :: 'a fset) ([] :: 'c list) \<Longrightarrow> (\<And>e t. P x t \<Longrightarrow> P x (e # t)) \<Longrightarrow> P x l"
654
02fd9de9d45e tuned the examples and flagged the problematic cleaning lemmas in FSet
Christian Urban <urbanc@in.tum.de>
parents: 653
diff changeset
   280
apply (lifting list_induct_part)
482
767baada01dc New APPLY_RSP which finally does automatic partial lifting :). Doesn't support same relation yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 481
diff changeset
   281
done
767baada01dc New APPLY_RSP which finally does automatic partial lifting :). Doesn't support same relation yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 481
diff changeset
   282
787
5cf83fa5b36c made the quotient_type definition more like typedef; now type variables need to be explicitly given
Christian Urban <urbanc@in.tum.de>
parents: 779
diff changeset
   283
quotient_type 'a fset2 = "'a list" / "list_eq"
654
02fd9de9d45e tuned the examples and flagged the problematic cleaning lemmas in FSet
Christian Urban <urbanc@in.tum.de>
parents: 653
diff changeset
   284
  by (rule equivp_list_eq)
483
74348dc2f8bb Lifting to 2 different types :)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 482
diff changeset
   285
767
37285ec4387d on suggestion of Tobias renamed "quotient_def" to "quotient_definition"; needs new keyword file
Christian Urban <urbanc@in.tum.de>
parents: 766
diff changeset
   286
quotient_definition
705
f51c6069cd17 New syntax for definitions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 697
diff changeset
   287
   "EMPTY2 :: 'a fset2"
f51c6069cd17 New syntax for definitions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 697
diff changeset
   288
as
663
0dd10a900cae Different syntax for definitions that allows overloading and retrieving of definitions by matching whole constants.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 656
diff changeset
   289
  "[]::'a list"
483
74348dc2f8bb Lifting to 2 different types :)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 482
diff changeset
   290
767
37285ec4387d on suggestion of Tobias renamed "quotient_def" to "quotient_definition"; needs new keyword file
Christian Urban <urbanc@in.tum.de>
parents: 766
diff changeset
   291
quotient_definition
705
f51c6069cd17 New syntax for definitions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 697
diff changeset
   292
   "INSERT2 :: 'a \<Rightarrow> 'a fset2 \<Rightarrow> 'a fset2"
f51c6069cd17 New syntax for definitions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 697
diff changeset
   293
as
663
0dd10a900cae Different syntax for definitions that allows overloading and retrieving of definitions by matching whole constants.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 656
diff changeset
   294
  "op #"
483
74348dc2f8bb Lifting to 2 different types :)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 482
diff changeset
   295
74348dc2f8bb Lifting to 2 different types :)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 482
diff changeset
   296
lemma "P (x :: 'a fset2) (EMPTY :: 'c fset) \<Longrightarrow> (\<And>e t. P x t \<Longrightarrow> P x (INSERT e t)) \<Longrightarrow> P x l"
654
02fd9de9d45e tuned the examples and flagged the problematic cleaning lemmas in FSet
Christian Urban <urbanc@in.tum.de>
parents: 653
diff changeset
   297
apply (lifting list_induct_part)
483
74348dc2f8bb Lifting to 2 different types :)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 482
diff changeset
   298
done
74348dc2f8bb Lifting to 2 different types :)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 482
diff changeset
   299
74348dc2f8bb Lifting to 2 different types :)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 482
diff changeset
   300
lemma "P (x :: 'a fset) (EMPTY2 :: 'c fset2) \<Longrightarrow> (\<And>e t. P x t \<Longrightarrow> P x (INSERT2 e t)) \<Longrightarrow> P x l"
654
02fd9de9d45e tuned the examples and flagged the problematic cleaning lemmas in FSet
Christian Urban <urbanc@in.tum.de>
parents: 653
diff changeset
   301
apply (lifting list_induct_part)
483
74348dc2f8bb Lifting to 2 different types :)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 482
diff changeset
   302
done
74348dc2f8bb Lifting to 2 different types :)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 482
diff changeset
   303
767
37285ec4387d on suggestion of Tobias renamed "quotient_def" to "quotient_definition"; needs new keyword file
Christian Urban <urbanc@in.tum.de>
parents: 766
diff changeset
   304
quotient_definition
705
f51c6069cd17 New syntax for definitions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 697
diff changeset
   305
  "fset_rec :: 'a \<Rightarrow> ('b \<Rightarrow> 'b fset \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> 'b fset \<Rightarrow> 'a"
f51c6069cd17 New syntax for definitions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 697
diff changeset
   306
as
663
0dd10a900cae Different syntax for definitions that allows overloading and retrieving of definitions by matching whole constants.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 656
diff changeset
   307
  "list_rec"
273
b82e765ca464 Lifting 'fold1.simps(2)' and some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 270
diff changeset
   308
767
37285ec4387d on suggestion of Tobias renamed "quotient_def" to "quotient_definition"; needs new keyword file
Christian Urban <urbanc@in.tum.de>
parents: 766
diff changeset
   309
quotient_definition
705
f51c6069cd17 New syntax for definitions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 697
diff changeset
   310
  "fset_case :: 'a \<Rightarrow> ('b \<Rightarrow> 'b fset \<Rightarrow> 'a) \<Rightarrow> 'b fset \<Rightarrow> 'a"
f51c6069cd17 New syntax for definitions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 697
diff changeset
   311
as
663
0dd10a900cae Different syntax for definitions that allows overloading and retrieving of definitions by matching whole constants.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 656
diff changeset
   312
  "list_case"
292
bd76f0398aa9 More functionality for lifting list.cases and list.recs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 291
diff changeset
   313
296
eab108c8d4b7 Minor changes
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 294
diff changeset
   314
(* Probably not true without additional assumptions about the function *)
636
520a4084d064 changed names of attributes
Christian Urban <urbanc@in.tum.de>
parents: 631
diff changeset
   315
lemma list_rec_rsp[quot_respect]:
292
bd76f0398aa9 More functionality for lifting list.cases and list.recs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 291
diff changeset
   316
  "(op = ===> (op = ===> op \<approx> ===> op =) ===> op \<approx> ===> op =) list_rec list_rec"
536
44fa9df44e6f More code cleaning and name changes
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 529
diff changeset
   317
  apply (auto)
296
eab108c8d4b7 Minor changes
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 294
diff changeset
   318
  apply (erule_tac list_eq.induct)
eab108c8d4b7 Minor changes
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 294
diff changeset
   319
  apply (simp_all)
292
bd76f0398aa9 More functionality for lifting list.cases and list.recs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 291
diff changeset
   320
  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
   321
636
520a4084d064 changed names of attributes
Christian Urban <urbanc@in.tum.de>
parents: 631
diff changeset
   322
lemma list_case_rsp[quot_respect]:
292
bd76f0398aa9 More functionality for lifting list.cases and list.recs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 291
diff changeset
   323
  "(op = ===> (op = ===> op \<approx> ===> op =) ===> op \<approx> ===> op =) list_case list_case"
536
44fa9df44e6f More code cleaning and name changes
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 529
diff changeset
   324
  apply (auto)
292
bd76f0398aa9 More functionality for lifting list.cases and list.recs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 291
diff changeset
   325
  sorry
bd76f0398aa9 More functionality for lifting list.cases and list.recs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 291
diff changeset
   326
376
e99c0334d8bf lambda_prs and cleaning the existing examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 375
diff changeset
   327
lemma "fset_rec (f1::'t) x (INSERT a xa) = x a xa (fset_rec f1 x xa)"
654
02fd9de9d45e tuned the examples and flagged the problematic cleaning lemmas in FSet
Christian Urban <urbanc@in.tum.de>
parents: 653
diff changeset
   328
apply (lifting list.recs(2))
376
e99c0334d8bf lambda_prs and cleaning the existing examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 375
diff changeset
   329
done
e99c0334d8bf lambda_prs and cleaning the existing examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 375
diff changeset
   330
e99c0334d8bf lambda_prs and cleaning the existing examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 375
diff changeset
   331
lemma "fset_case (f1::'t) f2 (INSERT a xa) = f2 a xa"
654
02fd9de9d45e tuned the examples and flagged the problematic cleaning lemmas in FSet
Christian Urban <urbanc@in.tum.de>
parents: 653
diff changeset
   332
apply (lifting list.cases(2))
376
e99c0334d8bf lambda_prs and cleaning the existing examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 375
diff changeset
   333
done
348
b1f83c7a8674 More theorems lifted in the goal-directed way.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 338
diff changeset
   334
609
6ce4f274b0fa 3 lambda examples in FSet. In the last one regularize_term fails.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 600
diff changeset
   335
lemma ttt: "((op @) x ((op #) e [])) = (((op #) e x))"
6ce4f274b0fa 3 lambda examples in FSet. In the last one regularize_term fails.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 600
diff changeset
   336
sorry
654
02fd9de9d45e tuned the examples and flagged the problematic cleaning lemmas in FSet
Christian Urban <urbanc@in.tum.de>
parents: 653
diff changeset
   337
609
6ce4f274b0fa 3 lambda examples in FSet. In the last one regularize_term fails.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 600
diff changeset
   338
lemma "(FUNION x (INSERT e EMPTY)) = ((INSERT e x))"
654
02fd9de9d45e tuned the examples and flagged the problematic cleaning lemmas in FSet
Christian Urban <urbanc@in.tum.de>
parents: 653
diff changeset
   339
apply (lifting ttt)
609
6ce4f274b0fa 3 lambda examples in FSet. In the last one regularize_term fails.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 600
diff changeset
   340
done
6ce4f274b0fa 3 lambda examples in FSet. In the last one regularize_term fails.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 600
diff changeset
   341
658
d616a0912245 improved fun_map_conv
Christian Urban <urbanc@in.tum.de>
parents: 656
diff changeset
   342
609
6ce4f274b0fa 3 lambda examples in FSet. In the last one regularize_term fails.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 600
diff changeset
   343
lemma ttt2: "(\<lambda>e. ((op @) x ((op #) e []))) = (\<lambda>e. ((op #) e x))"
6ce4f274b0fa 3 lambda examples in FSet. In the last one regularize_term fails.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 600
diff changeset
   344
sorry
6ce4f274b0fa 3 lambda examples in FSet. In the last one regularize_term fails.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 600
diff changeset
   345
6ce4f274b0fa 3 lambda examples in FSet. In the last one regularize_term fails.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 600
diff changeset
   346
lemma "(\<lambda>e. (FUNION x (INSERT e EMPTY))) = (\<lambda>e. (INSERT e x))"
641
b98d64dc98d9 added preserve rules to the cleaning_tac
Christian Urban <urbanc@in.tum.de>
parents: 638
diff changeset
   347
apply(lifting ttt2)
639
820c64273ce0 cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 634
diff changeset
   348
apply(regularize)
634
54573efed527 Manual regularization of a goal in FSet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 631
diff changeset
   349
apply(rule impI)
54573efed527 Manual regularization of a goal in FSet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 631
diff changeset
   350
apply(simp)
54573efed527 Manual regularization of a goal in FSet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 631
diff changeset
   351
apply(rule allI)
54573efed527 Manual regularization of a goal in FSet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 631
diff changeset
   352
apply(rule list_eq_refl)
54573efed527 Manual regularization of a goal in FSet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 631
diff changeset
   353
done
609
6ce4f274b0fa 3 lambda examples in FSet. In the last one regularize_term fails.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 600
diff changeset
   354
695
2eba169533b5 Found the problem with ttt3.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 685
diff changeset
   355
lemma ttt3: "(\<lambda>x. ((op @) x (e # []))) = (op #) e"
609
6ce4f274b0fa 3 lambda examples in FSet. In the last one regularize_term fails.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 600
diff changeset
   356
sorry
6ce4f274b0fa 3 lambda examples in FSet. In the last one regularize_term fails.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 600
diff changeset
   357
695
2eba169533b5 Found the problem with ttt3.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 685
diff changeset
   358
lemma "(\<lambda>x. (FUNION x (INSERT e EMPTY))) = INSERT e"
2eba169533b5 Found the problem with ttt3.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 685
diff changeset
   359
apply(lifting ttt3)
2eba169533b5 Found the problem with ttt3.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 685
diff changeset
   360
apply(regularize)
2eba169533b5 Found the problem with ttt3.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 685
diff changeset
   361
apply(auto simp add: cons_rsp)
2eba169533b5 Found the problem with ttt3.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 685
diff changeset
   362
done
646
10d04ee52101 An example which is hard to lift because of the interplay between lambda_prs and unfolding.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 642
diff changeset
   363
lemma hard: "(\<lambda>P. \<lambda>Q. P (Q (x::'a list))) = (\<lambda>P. \<lambda>Q. Q (P (x::'a list)))"
10d04ee52101 An example which is hard to lift because of the interplay between lambda_prs and unfolding.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 642
diff changeset
   364
sorry
10d04ee52101 An example which is hard to lift because of the interplay between lambda_prs and unfolding.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 642
diff changeset
   365
10d04ee52101 An example which is hard to lift because of the interplay between lambda_prs and unfolding.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 642
diff changeset
   366
lemma hard_lift: "(\<lambda>P. \<lambda>Q. P (Q (x::'a fset))) = (\<lambda>P. \<lambda>Q. Q (P (x::'a fset)))"
656
c86a47d4966e Temporarily repeated fun_map_tac 4 times. Cleaning for all examples work.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 654
diff changeset
   367
apply(lifting hard)
658
d616a0912245 improved fun_map_conv
Christian Urban <urbanc@in.tum.de>
parents: 656
diff changeset
   368
apply(regularize)
697
57944c1ef728 Regularized the hard lemma.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 695
diff changeset
   369
apply(rule fun_rel_id_asm)
57944c1ef728 Regularized the hard lemma.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 695
diff changeset
   370
apply(subst babs_simp)
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents: 705
diff changeset
   371
apply(tactic {* Quotient_Tacs.quotient_tac @{context} 1 *})
697
57944c1ef728 Regularized the hard lemma.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 695
diff changeset
   372
apply(rule fun_rel_id_asm)
57944c1ef728 Regularized the hard lemma.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 695
diff changeset
   373
apply(rule impI)
57944c1ef728 Regularized the hard lemma.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 695
diff changeset
   374
apply(rule mp[OF eq_imp_rel[OF fset_equivp]])
57944c1ef728 Regularized the hard lemma.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 695
diff changeset
   375
apply(drule fun_cong)
57944c1ef728 Regularized the hard lemma.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 695
diff changeset
   376
apply(drule fun_cong)
57944c1ef728 Regularized the hard lemma.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 695
diff changeset
   377
apply(assumption)
57944c1ef728 Regularized the hard lemma.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 695
diff changeset
   378
done
57944c1ef728 Regularized the hard lemma.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 695
diff changeset
   379
918
7be9b054f672 test with splits
Christian Urban <urbanc@in.tum.de>
parents: 787
diff changeset
   380
lemma [quot_respect]:
7be9b054f672 test with splits
Christian Urban <urbanc@in.tum.de>
parents: 787
diff changeset
   381
  "((op \<approx> ===> op \<approx> ===> op =) ===> prod_rel op \<approx> op \<approx> ===> op =) split split"
7be9b054f672 test with splits
Christian Urban <urbanc@in.tum.de>
parents: 787
diff changeset
   382
apply(simp)
7be9b054f672 test with splits
Christian Urban <urbanc@in.tum.de>
parents: 787
diff changeset
   383
done
7be9b054f672 test with splits
Christian Urban <urbanc@in.tum.de>
parents: 787
diff changeset
   384
7be9b054f672 test with splits
Christian Urban <urbanc@in.tum.de>
parents: 787
diff changeset
   385
thm quot_preserve
7be9b054f672 test with splits
Christian Urban <urbanc@in.tum.de>
parents: 787
diff changeset
   386
term split
7be9b054f672 test with splits
Christian Urban <urbanc@in.tum.de>
parents: 787
diff changeset
   387
7be9b054f672 test with splits
Christian Urban <urbanc@in.tum.de>
parents: 787
diff changeset
   388
7be9b054f672 test with splits
Christian Urban <urbanc@in.tum.de>
parents: 787
diff changeset
   389
lemma [quot_preserve]:
7be9b054f672 test with splits
Christian Urban <urbanc@in.tum.de>
parents: 787
diff changeset
   390
  shows "(((abs_fset ---> abs_fset ---> id) ---> prod_fun rep_fset rep_fset ---> id) split) = split"
7be9b054f672 test with splits
Christian Urban <urbanc@in.tum.de>
parents: 787
diff changeset
   391
apply(simp add: expand_fun_eq)
7be9b054f672 test with splits
Christian Urban <urbanc@in.tum.de>
parents: 787
diff changeset
   392
apply(simp add: Quotient_abs_rep[OF Quotient_fset])
7be9b054f672 test with splits
Christian Urban <urbanc@in.tum.de>
parents: 787
diff changeset
   393
done
7be9b054f672 test with splits
Christian Urban <urbanc@in.tum.de>
parents: 787
diff changeset
   394
7be9b054f672 test with splits
Christian Urban <urbanc@in.tum.de>
parents: 787
diff changeset
   395
7be9b054f672 test with splits
Christian Urban <urbanc@in.tum.de>
parents: 787
diff changeset
   396
lemma test: "All (\<lambda>(x::'a list, y). x = y)"
7be9b054f672 test with splits
Christian Urban <urbanc@in.tum.de>
parents: 787
diff changeset
   397
sorry
7be9b054f672 test with splits
Christian Urban <urbanc@in.tum.de>
parents: 787
diff changeset
   398
7be9b054f672 test with splits
Christian Urban <urbanc@in.tum.de>
parents: 787
diff changeset
   399
lemma  "All (\<lambda>(x::'a fset, y). x = y)"
7be9b054f672 test with splits
Christian Urban <urbanc@in.tum.de>
parents: 787
diff changeset
   400
apply(lifting test)
7be9b054f672 test with splits
Christian Urban <urbanc@in.tum.de>
parents: 787
diff changeset
   401
apply(cleaning)
7be9b054f672 test with splits
Christian Urban <urbanc@in.tum.de>
parents: 787
diff changeset
   402
thm all_prs
7be9b054f672 test with splits
Christian Urban <urbanc@in.tum.de>
parents: 787
diff changeset
   403
apply(rule all_prs)
7be9b054f672 test with splits
Christian Urban <urbanc@in.tum.de>
parents: 787
diff changeset
   404
apply(tactic {* Quotient_Tacs.quotient_tac @{context} 1*})
7be9b054f672 test with splits
Christian Urban <urbanc@in.tum.de>
parents: 787
diff changeset
   405
done
7be9b054f672 test with splits
Christian Urban <urbanc@in.tum.de>
parents: 787
diff changeset
   406
7be9b054f672 test with splits
Christian Urban <urbanc@in.tum.de>
parents: 787
diff changeset
   407
lemma test2: "Ex (\<lambda>(x::'a list, y). x = y)"
7be9b054f672 test with splits
Christian Urban <urbanc@in.tum.de>
parents: 787
diff changeset
   408
sorry
7be9b054f672 test with splits
Christian Urban <urbanc@in.tum.de>
parents: 787
diff changeset
   409
7be9b054f672 test with splits
Christian Urban <urbanc@in.tum.de>
parents: 787
diff changeset
   410
lemma  "Ex (\<lambda>(x::'a fset, y). x = y)"
7be9b054f672 test with splits
Christian Urban <urbanc@in.tum.de>
parents: 787
diff changeset
   411
apply(lifting test2)
7be9b054f672 test with splits
Christian Urban <urbanc@in.tum.de>
parents: 787
diff changeset
   412
apply(cleaning)
7be9b054f672 test with splits
Christian Urban <urbanc@in.tum.de>
parents: 787
diff changeset
   413
apply(rule ex_prs)
7be9b054f672 test with splits
Christian Urban <urbanc@in.tum.de>
parents: 787
diff changeset
   414
apply(tactic {* Quotient_Tacs.quotient_tac @{context} 1*})
7be9b054f672 test with splits
Christian Urban <urbanc@in.tum.de>
parents: 787
diff changeset
   415
done
7be9b054f672 test with splits
Christian Urban <urbanc@in.tum.de>
parents: 787
diff changeset
   416
7be9b054f672 test with splits
Christian Urban <urbanc@in.tum.de>
parents: 787
diff changeset
   417
ML {* @{term "Ex (\<lambda>(x::'a fset, y). x = y)"} *}
7be9b054f672 test with splits
Christian Urban <urbanc@in.tum.de>
parents: 787
diff changeset
   418
7be9b054f672 test with splits
Christian Urban <urbanc@in.tum.de>
parents: 787
diff changeset
   419
697
57944c1ef728 Regularized the hard lemma.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 695
diff changeset
   420
650
bbaa07eea396 manually cleaned the hard lemma.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 646
diff changeset
   421
163
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   422
end