Quot/Examples/FSet.thy
author Christian Urban <urbanc@in.tum.de>
Tue, 22 Dec 2009 21:31:44 +0100
changeset 776 d1064fa29424
parent 774 b4ffb8826105
child 779 3b21b24a5fb6
permissions -rw-r--r--
renamed get_fun to absrep_fun; introduced explicit checked versions of the term functions
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
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
     2
imports "../QuotMain" "../QuotList" 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 
df053507edba renamed "quotient" command to "quotient_type"; needs new keyword file to be installed
Christian Urban <urbanc@in.tum.de>
parents: 758
diff changeset
    26
  fset = "'a list" / "list_eq"
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
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
    29
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
    30
fun
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
    31
  intrel :: "(nat \<times> nat) \<Rightarrow> (nat \<times> nat) \<Rightarrow> bool"
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
    32
where
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
    33
  "intrel (x, y) (u, v) = (x + v = u + y)"
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
    34
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
    35
quotient_type int = "nat \<times> nat" / intrel
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
    36
  apply(unfold equivp_def)
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
    37
  apply(auto simp add: mem_def expand_fun_eq)
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
    38
  done
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
    39
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
    40
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
    41
ML {*
774
b4ffb8826105 moved get_fun into quotient_term; this simplifies the overall including structure of the package
Christian Urban <urbanc@in.tum.de>
parents: 768
diff changeset
    42
open Quotient_Term;
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
    43
*}
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
    44
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
    45
ML {*
776
d1064fa29424 renamed get_fun to absrep_fun; introduced explicit checked versions of the term functions
Christian Urban <urbanc@in.tum.de>
parents: 774
diff changeset
    46
absrep_fun absF @{context} (@{typ "'a list"}, 
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
    47
                         @{typ "'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
    48
|> Syntax.check_term @{context}
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
    49
|> Syntax.string_of_term @{context}
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
    50
|> writeln
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
    51
*}
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
    52
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
    53
term "map id"
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
    54
term "abs_fset o (map id)"
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
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
ML {*
776
d1064fa29424 renamed get_fun to absrep_fun; introduced explicit checked versions of the term functions
Christian Urban <urbanc@in.tum.de>
parents: 774
diff changeset
    57
absrep_fun absF @{context} (@{typ "(nat * nat) list"}, 
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
    58
                         @{typ "int 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
    59
|> Syntax.check_term @{context}
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
|> Syntax.string_of_term @{context}
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
|> writeln
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
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
    64
term "map abs_int"
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
    65
term "abs_fset o map abs_int"
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
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
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
    68
ML {*
776
d1064fa29424 renamed get_fun to absrep_fun; introduced explicit checked versions of the term functions
Christian Urban <urbanc@in.tum.de>
parents: 774
diff changeset
    69
absrep_fun absF @{context} (@{typ "('a list) list"}, 
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
    70
                         @{typ "('a fset) 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
    71
|> Syntax.check_term @{context}
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
    72
|> Syntax.string_of_term @{context}
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
    73
|> writeln
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
    74
*}
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
    75
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
    76
ML {*
776
d1064fa29424 renamed get_fun to absrep_fun; introduced explicit checked versions of the term functions
Christian Urban <urbanc@in.tum.de>
parents: 774
diff changeset
    77
absrep_fun absF @{context} (@{typ "('a list) list \<Rightarrow> 'a"}, 
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
    78
                         @{typ "('a fset) fset \<Rightarrow> 'a"})
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
    79
|> Syntax.check_term @{context}
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
    80
|> Syntax.string_of_term @{context}
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
    81
|> writeln
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
    82
*}
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
    83
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
    84
quotient_definition
705
f51c6069cd17 New syntax for definitions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 697
diff changeset
    85
   "EMPTY :: 'a fset"
f51c6069cd17 New syntax for definitions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 697
diff changeset
    86
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
    87
  "[]::'a list"
163
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    88
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
    89
quotient_definition
705
f51c6069cd17 New syntax for definitions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 697
diff changeset
    90
   "INSERT :: 'a \<Rightarrow> 'a fset \<Rightarrow> 'a fset"
f51c6069cd17 New syntax for definitions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 697
diff changeset
    91
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
    92
  "op #"
163
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    93
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
    94
quotient_definition
705
f51c6069cd17 New syntax for definitions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 697
diff changeset
    95
   "FUNION :: 'a fset \<Rightarrow> 'a fset \<Rightarrow> 'a fset"
f51c6069cd17 New syntax for definitions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 697
diff changeset
    96
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
    97
  "op @"
163
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    98
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    99
fun
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   100
  card1 :: "'a list \<Rightarrow> nat"
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   101
where
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   102
  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
   103
| 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
   104
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
   105
quotient_definition
705
f51c6069cd17 New syntax for definitions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 697
diff changeset
   106
   "CARD :: 'a fset \<Rightarrow> nat"
f51c6069cd17 New syntax for definitions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 697
diff changeset
   107
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
   108
  "card1"
163
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   109
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
   110
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
   111
  "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
   112
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
   113
  "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
   114
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
   115
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
   116
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
   117
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
   118
text {*
163
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   119
 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
   120
 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
   121
 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
   122
*}
163
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   123
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   124
lemma card1_0:
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   125
  fixes a :: "'a list"
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   126
  shows "(card1 a = 0) = (a = [])"
214
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   127
  by (induct a) auto
163
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   128
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   129
lemma not_mem_card1:
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   130
  fixes x :: "'a"
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   131
  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
   132
  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
   133
  by auto
163
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   134
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   135
lemma mem_cons:
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   136
  fixes x :: "'a"
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   137
  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
   138
  assumes a : "x mem xs"
163
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   139
  shows "x # xs \<approx> xs"
214
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   140
  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
   141
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   142
lemma card1_suc:
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   143
  fixes xs :: "'a list"
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   144
  fixes n :: "nat"
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   145
  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
   146
  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
   147
  using c
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   148
apply(induct xs)
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   149
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
   150
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
   151
done
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   152
294
a092c0b13d83 fold_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 292
diff changeset
   153
definition
a092c0b13d83 fold_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 292
diff changeset
   154
  rsp_fold
a092c0b13d83 fold_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 292
diff changeset
   155
where
a092c0b13d83 fold_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 292
diff changeset
   156
  "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
   157
163
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   158
primrec
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   159
  fold1
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   160
where
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   161
  "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
   162
| "fold1 f g z (a # A) =
294
a092c0b13d83 fold_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 292
diff changeset
   163
     (if rsp_fold f
163
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   164
     then (
683
0d9e8aa1bc7a removed memb and used standard mem (member from List.thy)
Christian Urban <urbanc@in.tum.de>
parents: 681
diff changeset
   165
       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
   166
     ) else z)"
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   167
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   168
lemma fs1_strong_cases:
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   169
  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
   170
  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
   171
  apply (induct X)
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   172
  apply (simp)
685
b12f0321dfb0 moved the interpretation code into Unused.thy
Christian Urban <urbanc@in.tum.de>
parents: 683
diff changeset
   173
  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
   174
  done
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   175
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
   176
quotient_definition
705
f51c6069cd17 New syntax for definitions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 697
diff changeset
   177
   "IN :: 'a \<Rightarrow> 'a fset \<Rightarrow> bool"
f51c6069cd17 New syntax for definitions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 697
diff changeset
   178
as
683
0d9e8aa1bc7a removed memb and used standard mem (member from List.thy)
Christian Urban <urbanc@in.tum.de>
parents: 681
diff changeset
   179
  "op mem"
163
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   180
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
   181
quotient_definition
705
f51c6069cd17 New syntax for definitions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 697
diff changeset
   182
   "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
   183
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
   184
  "fold1"
194
03c03e88efa9 Simplifying Int and Working on map
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 190
diff changeset
   185
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
   186
quotient_definition
705
f51c6069cd17 New syntax for definitions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 697
diff changeset
   187
  "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
   188
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
   189
  "map"
194
03c03e88efa9 Simplifying Int and Working on map
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 190
diff changeset
   190
683
0d9e8aa1bc7a removed memb and used standard mem (member from List.thy)
Christian Urban <urbanc@in.tum.de>
parents: 681
diff changeset
   191
lemma mem_rsp:
163
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   192
  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
   193
  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
   194
  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
   195
  using a by induct auto
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   196
636
520a4084d064 changed names of attributes
Christian Urban <urbanc@in.tum.de>
parents: 631
diff changeset
   197
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
   198
  "(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
   199
  by (simp add: mem_rsp)
164
4f00ca4f5ef4 Stronger tactic, simpler proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 163
diff changeset
   200
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
   201
lemma card1_rsp:
163
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   202
  fixes a b :: "'a list"
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   203
  assumes e: "a \<approx> b"
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   204
  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
   205
  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
   206
636
520a4084d064 changed names of attributes
Christian Urban <urbanc@in.tum.de>
parents: 631
diff changeset
   207
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
   208
  "(op \<approx> ===> op =) card1 card1"
214
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   209
  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
   210
680
d003f9e00c29 removed quot_respect attribute of a non-standard lemma
Christian Urban <urbanc@in.tum.de>
parents: 664
diff changeset
   211
lemma cons_rsp:
163
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   212
  fixes z
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   213
  assumes a: "xs \<approx> ys"
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   214
  shows "(z # xs) \<approx> (z # ys)"
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   215
  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
   216
636
520a4084d064 changed names of attributes
Christian Urban <urbanc@in.tum.de>
parents: 631
diff changeset
   217
lemma ho_cons_rsp[quot_respect]:
228
268a727b0f10 disambiguate ===> syntax
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 226
diff changeset
   218
  "(op = ===> op \<approx> ===> op \<approx>) op # op #"
214
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   219
  by (simp add: cons_rsp)
164
4f00ca4f5ef4 Stronger tactic, simpler proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 163
diff changeset
   220
681
3c6419a5a7fc simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 680
diff changeset
   221
lemma append_rsp_aux1:
3c6419a5a7fc simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 680
diff changeset
   222
  assumes a : "l2 \<approx> r2 "
3c6419a5a7fc simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 680
diff changeset
   223
  shows "(h @ l2) \<approx> (h @ r2)"
3c6419a5a7fc simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 680
diff changeset
   224
using a
3c6419a5a7fc simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 680
diff changeset
   225
apply(induct h)
3c6419a5a7fc simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 680
diff changeset
   226
apply(auto intro: list_eq.intros(5))
3c6419a5a7fc simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 680
diff changeset
   227
done
163
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   228
681
3c6419a5a7fc simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 680
diff changeset
   229
lemma append_rsp_aux2:
3c6419a5a7fc simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 680
diff changeset
   230
  assumes a : "l1 \<approx> r1" "l2 \<approx> r2 "
3c6419a5a7fc simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 680
diff changeset
   231
  shows "(l1 @ l2) \<approx> (r1 @ r2)"
3c6419a5a7fc simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 680
diff changeset
   232
using a
3c6419a5a7fc simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 680
diff changeset
   233
apply(induct arbitrary: l2 r2)
3c6419a5a7fc simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 680
diff changeset
   234
apply(simp_all)
3c6419a5a7fc simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 680
diff changeset
   235
apply(blast intro: list_eq.intros append_rsp_aux1)+
3c6419a5a7fc simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 680
diff changeset
   236
done
214
a66f81c264aa Proof of append_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 213
diff changeset
   237
681
3c6419a5a7fc simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 680
diff changeset
   238
lemma append_rsp[quot_respect]:
228
268a727b0f10 disambiguate ===> syntax
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 226
diff changeset
   239
  "(op \<approx> ===> op \<approx> ===> op \<approx>) op @ op @"
681
3c6419a5a7fc simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 680
diff changeset
   240
  by (auto simp add: append_rsp_aux2)
175
f7602653dddd Preparing infrastructire for LAMBDA_PRS
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 173
diff changeset
   241
451
586e3dc4afdb Added 'TRY' to refl in clean_tac to get as far as possible. Removed unnecessary [quot_rsp] in FSet. Added necessary [quot_rsp] and one lifted thm in LamEx.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 450
diff changeset
   242
lemma map_rsp:
194
03c03e88efa9 Simplifying Int and Working on map
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 190
diff changeset
   243
  assumes a: "a \<approx> b"
03c03e88efa9 Simplifying Int and Working on map
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 190
diff changeset
   244
  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
   245
  using a
03c03e88efa9 Simplifying Int and Working on map
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 190
diff changeset
   246
  apply (induct)
03c03e88efa9 Simplifying Int and Working on map
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 190
diff changeset
   247
  apply(auto intro: list_eq.intros)
03c03e88efa9 Simplifying Int and Working on map
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 190
diff changeset
   248
  done
03c03e88efa9 Simplifying Int and Working on map
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 190
diff changeset
   249
636
520a4084d064 changed names of attributes
Christian Urban <urbanc@in.tum.de>
parents: 631
diff changeset
   250
lemma ho_map_rsp[quot_respect]:
294
a092c0b13d83 fold_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 292
diff changeset
   251
  "(op = ===> op \<approx> ===> op \<approx>) map map"
a092c0b13d83 fold_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 292
diff changeset
   252
  by (simp add: map_rsp)
194
03c03e88efa9 Simplifying Int and Working on map
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 190
diff changeset
   253
294
a092c0b13d83 fold_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 292
diff changeset
   254
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
   255
  "(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
   256
 by simp (rule list_eq_refl)
194
03c03e88efa9 Simplifying Int and Working on map
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 190
diff changeset
   257
636
520a4084d064 changed names of attributes
Christian Urban <urbanc@in.tum.de>
parents: 631
diff changeset
   258
lemma ho_fold_rsp[quot_respect]:
294
a092c0b13d83 fold_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 292
diff changeset
   259
  "(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
   260
  apply (auto)
294
a092c0b13d83 fold_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 292
diff changeset
   261
  apply (case_tac "rsp_fold x")
a092c0b13d83 fold_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 292
diff changeset
   262
  prefer 2
a092c0b13d83 fold_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 292
diff changeset
   263
  apply (erule_tac list_eq.induct)
a092c0b13d83 fold_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 292
diff changeset
   264
  apply (simp_all)
a092c0b13d83 fold_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 292
diff changeset
   265
  apply (erule_tac list_eq.induct)
a092c0b13d83 fold_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 292
diff changeset
   266
  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
   267
  apply (auto simp add: mem_rsp rsp_fold_def)
294
a092c0b13d83 fold_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 292
diff changeset
   268
done
241
60acf3d3a4a0 Finding applications and duplicates filtered out in abstractions
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 239
diff changeset
   269
636
520a4084d064 changed names of attributes
Christian Urban <urbanc@in.tum.de>
parents: 631
diff changeset
   270
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
   271
  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
   272
by (auto intro: list_eq.intros)
f178958d3d81 not yet quite functional treatment of constants
Christian Urban <urbanc@in.tum.de>
parents: 536
diff changeset
   273
364
4c455d58ac99 Fixes to the tactic after quotient_tac changed.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 356
diff changeset
   274
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
   275
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
   276
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
   277
364
4c455d58ac99 Fixes to the tactic after quotient_tac changed.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 356
diff changeset
   278
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
   279
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
   280
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
   281
364
4c455d58ac99 Fixes to the tactic after quotient_tac changed.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 356
diff changeset
   282
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
   283
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
   284
done
4c455d58ac99 Fixes to the tactic after quotient_tac changed.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 356
diff changeset
   285
367
d444389fe3f9 The non-working procedure_tac.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 364
diff changeset
   286
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
   287
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
   288
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
   289
367
d444389fe3f9 The non-working procedure_tac.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 364
diff changeset
   290
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
   291
apply (lifting card1_suc)
364
4c455d58ac99 Fixes to the tactic after quotient_tac changed.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 356
diff changeset
   292
done
4c455d58ac99 Fixes to the tactic after quotient_tac changed.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 356
diff changeset
   293
4c455d58ac99 Fixes to the tactic after quotient_tac changed.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 356
diff changeset
   294
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
   295
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
   296
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
   297
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
   298
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
   299
  (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
   300
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
   301
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
   302
368
c5c49d240cde Conversion
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 367
diff changeset
   303
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
   304
apply (lifting map_append)
368
c5c49d240cde Conversion
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 367
diff changeset
   305
done
c5c49d240cde Conversion
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 367
diff changeset
   306
367
d444389fe3f9 The non-working procedure_tac.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 364
diff changeset
   307
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
   308
apply (lifting append_assoc)
367
d444389fe3f9 The non-working procedure_tac.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 364
diff changeset
   309
done
d444389fe3f9 The non-working procedure_tac.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 364
diff changeset
   310
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
   311
376
e99c0334d8bf lambda_prs and cleaning the existing examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 375
diff changeset
   312
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
   313
apply(lifting list.induct)
414
4dad34ca50db Minor cleaning
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 401
diff changeset
   314
done
390
1dd6a21cdd1c test with monos
Christian Urban <urbanc@in.tum.de>
parents: 387
diff changeset
   315
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
   316
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
   317
  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
   318
  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
   319
  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
   320
  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
   321
  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
   322
  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
   323
  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
   324
  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
   325
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
   326
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
   327
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
   328
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
   329
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
   330
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
   331
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
   332
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
   333
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
   334
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
   335
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
   336
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
   337
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
   338
quotient_type 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
   339
  by (rule equivp_list_eq)
483
74348dc2f8bb Lifting to 2 different types :)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 482
diff changeset
   340
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
   341
quotient_definition
705
f51c6069cd17 New syntax for definitions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 697
diff changeset
   342
   "EMPTY2 :: 'a fset2"
f51c6069cd17 New syntax for definitions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 697
diff changeset
   343
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
   344
  "[]::'a list"
483
74348dc2f8bb Lifting to 2 different types :)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 482
diff changeset
   345
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
   346
quotient_definition
705
f51c6069cd17 New syntax for definitions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 697
diff changeset
   347
   "INSERT2 :: 'a \<Rightarrow> 'a fset2 \<Rightarrow> 'a fset2"
f51c6069cd17 New syntax for definitions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 697
diff changeset
   348
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
   349
  "op #"
483
74348dc2f8bb Lifting to 2 different types :)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 482
diff changeset
   350
74348dc2f8bb Lifting to 2 different types :)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 482
diff changeset
   351
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
   352
apply (lifting list_induct_part)
483
74348dc2f8bb Lifting to 2 different types :)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 482
diff changeset
   353
done
74348dc2f8bb Lifting to 2 different types :)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 482
diff changeset
   354
74348dc2f8bb Lifting to 2 different types :)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 482
diff changeset
   355
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
   356
apply (lifting list_induct_part)
483
74348dc2f8bb Lifting to 2 different types :)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 482
diff changeset
   357
done
74348dc2f8bb Lifting to 2 different types :)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 482
diff changeset
   358
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
   359
quotient_definition
705
f51c6069cd17 New syntax for definitions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 697
diff changeset
   360
  "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
   361
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
   362
  "list_rec"
273
b82e765ca464 Lifting 'fold1.simps(2)' and some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 270
diff changeset
   363
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
   364
quotient_definition
705
f51c6069cd17 New syntax for definitions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 697
diff changeset
   365
  "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
   366
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
   367
  "list_case"
292
bd76f0398aa9 More functionality for lifting list.cases and list.recs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 291
diff changeset
   368
296
eab108c8d4b7 Minor changes
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 294
diff changeset
   369
(* 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
   370
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
   371
  "(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
   372
  apply (auto)
296
eab108c8d4b7 Minor changes
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 294
diff changeset
   373
  apply (erule_tac list_eq.induct)
eab108c8d4b7 Minor changes
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 294
diff changeset
   374
  apply (simp_all)
292
bd76f0398aa9 More functionality for lifting list.cases and list.recs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 291
diff changeset
   375
  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
   376
636
520a4084d064 changed names of attributes
Christian Urban <urbanc@in.tum.de>
parents: 631
diff changeset
   377
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
   378
  "(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
   379
  apply (auto)
292
bd76f0398aa9 More functionality for lifting list.cases and list.recs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 291
diff changeset
   380
  sorry
bd76f0398aa9 More functionality for lifting list.cases and list.recs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 291
diff changeset
   381
376
e99c0334d8bf lambda_prs and cleaning the existing examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 375
diff changeset
   382
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
   383
apply (lifting list.recs(2))
376
e99c0334d8bf lambda_prs and cleaning the existing examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 375
diff changeset
   384
done
e99c0334d8bf lambda_prs and cleaning the existing examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 375
diff changeset
   385
e99c0334d8bf lambda_prs and cleaning the existing examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 375
diff changeset
   386
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
   387
apply (lifting list.cases(2))
376
e99c0334d8bf lambda_prs and cleaning the existing examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 375
diff changeset
   388
done
348
b1f83c7a8674 More theorems lifted in the goal-directed way.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 338
diff changeset
   389
609
6ce4f274b0fa 3 lambda examples in FSet. In the last one regularize_term fails.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 600
diff changeset
   390
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
   391
sorry
654
02fd9de9d45e tuned the examples and flagged the problematic cleaning lemmas in FSet
Christian Urban <urbanc@in.tum.de>
parents: 653
diff changeset
   392
609
6ce4f274b0fa 3 lambda examples in FSet. In the last one regularize_term fails.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 600
diff changeset
   393
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
   394
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
   395
done
6ce4f274b0fa 3 lambda examples in FSet. In the last one regularize_term fails.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 600
diff changeset
   396
658
d616a0912245 improved fun_map_conv
Christian Urban <urbanc@in.tum.de>
parents: 656
diff changeset
   397
609
6ce4f274b0fa 3 lambda examples in FSet. In the last one regularize_term fails.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 600
diff changeset
   398
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
   399
sorry
6ce4f274b0fa 3 lambda examples in FSet. In the last one regularize_term fails.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 600
diff changeset
   400
6ce4f274b0fa 3 lambda examples in FSet. In the last one regularize_term fails.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 600
diff changeset
   401
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
   402
apply(lifting ttt2)
639
820c64273ce0 cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 634
diff changeset
   403
apply(regularize)
634
54573efed527 Manual regularization of a goal in FSet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 631
diff changeset
   404
apply(rule impI)
54573efed527 Manual regularization of a goal in FSet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 631
diff changeset
   405
apply(simp)
54573efed527 Manual regularization of a goal in FSet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 631
diff changeset
   406
apply(rule allI)
54573efed527 Manual regularization of a goal in FSet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 631
diff changeset
   407
apply(rule list_eq_refl)
54573efed527 Manual regularization of a goal in FSet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 631
diff changeset
   408
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
   409
695
2eba169533b5 Found the problem with ttt3.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 685
diff changeset
   410
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
   411
sorry
6ce4f274b0fa 3 lambda examples in FSet. In the last one regularize_term fails.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 600
diff changeset
   412
695
2eba169533b5 Found the problem with ttt3.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 685
diff changeset
   413
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
   414
apply(lifting ttt3)
2eba169533b5 Found the problem with ttt3.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 685
diff changeset
   415
apply(regularize)
2eba169533b5 Found the problem with ttt3.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 685
diff changeset
   416
apply(auto simp add: cons_rsp)
2eba169533b5 Found the problem with ttt3.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 685
diff changeset
   417
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
   418
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
   419
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
   420
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
   421
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
   422
apply(lifting hard)
658
d616a0912245 improved fun_map_conv
Christian Urban <urbanc@in.tum.de>
parents: 656
diff changeset
   423
apply(regularize)
697
57944c1ef728 Regularized the hard lemma.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 695
diff changeset
   424
apply(rule fun_rel_id_asm)
57944c1ef728 Regularized the hard lemma.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 695
diff changeset
   425
apply(subst babs_simp)
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents: 705
diff changeset
   426
apply(tactic {* Quotient_Tacs.quotient_tac @{context} 1 *})
697
57944c1ef728 Regularized the hard lemma.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 695
diff changeset
   427
apply(rule fun_rel_id_asm)
57944c1ef728 Regularized the hard lemma.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 695
diff changeset
   428
apply(rule impI)
57944c1ef728 Regularized the hard lemma.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 695
diff changeset
   429
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
   430
apply(drule fun_cong)
57944c1ef728 Regularized the hard lemma.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 695
diff changeset
   431
apply(drule fun_cong)
57944c1ef728 Regularized the hard lemma.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 695
diff changeset
   432
apply(assumption)
57944c1ef728 Regularized the hard lemma.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 695
diff changeset
   433
done
57944c1ef728 Regularized the hard lemma.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 695
diff changeset
   434
57944c1ef728 Regularized the hard lemma.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 695
diff changeset
   435
650
bbaa07eea396 manually cleaned the hard lemma.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 646
diff changeset
   436
163
3da18bf6886c Split Finite Set example into separate file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   437
end