Attic/Quot/Examples/FSet3.thy
author Cezary Kaliszyk <kaliszyk@in.tum.de>
Thu, 15 Apr 2010 11:42:28 +0200
changeset 1850 05b2dd2b0e8a
parent 1260 9df6144e281b
child 1873 a08eaea622d1
permissions -rw-r--r--
Prove insert_rsp2
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
714
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     1
theory FSet3
1128
17ca92ab4660 Main renaming + fixes for new Isabelle in IntEx2.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 856
diff changeset
     2
imports "../Quotient" "../Quotient_List" List
714
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     3
begin
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     4
798
a422a51bb0eb some small changes
cu@localhost
parents: 797
diff changeset
     5
ML {*
a422a51bb0eb some small changes
cu@localhost
parents: 797
diff changeset
     6
structure QuotientRules = Named_Thms
a422a51bb0eb some small changes
cu@localhost
parents: 797
diff changeset
     7
  (val name = "quot_thm"
a422a51bb0eb some small changes
cu@localhost
parents: 797
diff changeset
     8
   val description = "Quotient theorems.")
a422a51bb0eb some small changes
cu@localhost
parents: 797
diff changeset
     9
*}
a422a51bb0eb some small changes
cu@localhost
parents: 797
diff changeset
    10
a422a51bb0eb some small changes
cu@localhost
parents: 797
diff changeset
    11
ML {*
a422a51bb0eb some small changes
cu@localhost
parents: 797
diff changeset
    12
open QuotientRules
a422a51bb0eb some small changes
cu@localhost
parents: 797
diff changeset
    13
*}
a422a51bb0eb some small changes
cu@localhost
parents: 797
diff changeset
    14
722
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
    15
fun
714
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    16
  list_eq :: "'a list \<Rightarrow> 'a list \<Rightarrow> bool" (infix "\<approx>" 50)
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    17
where
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
    18
  "list_eq xs ys = (\<forall>x. x \<in> set xs \<longleftrightarrow> x \<in> set ys)"
714
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    19
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    20
lemma list_eq_equivp:
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    21
  shows "equivp list_eq"
722
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
    22
unfolding equivp_reflp_symp_transp reflp_def symp_def transp_def
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
    23
by auto
714
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    24
784
da75568e7f12 renamed some fields in the info records
Christian Urban <urbanc@in.tum.de>
parents: 768
diff changeset
    25
(* FIXME-TODO: because of beta-reduction, one cannot give the *)
da75568e7f12 renamed some fields in the info records
Christian Urban <urbanc@in.tum.de>
parents: 768
diff changeset
    26
(* FIXME-TODO: relation as a term or abbreviation             *) 
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
    27
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: 784
diff changeset
    28
  'a fset = "'a list" / "list_eq"
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
by (rule list_eq_equivp)
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
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
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
section {* empty fset, finsert and membership *} 
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
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
quotient_definition
1144
538daee762e6 Added a binding to the parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1141
diff changeset
    35
  fempty  ("{||}")
538daee762e6 Added a binding to the parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1141
diff changeset
    36
where
538daee762e6 Added a binding to the parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1141
diff changeset
    37
  "fempty :: 'a fset"
1141
3c8ad149a4d3 Undid the read_terms change; now compiles.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1139
diff changeset
    38
is "[]::'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
    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
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
    41
  "finsert :: 'a \<Rightarrow> 'a fset \<Rightarrow> 'a fset" 
1141
3c8ad149a4d3 Undid the read_terms change; now compiles.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1139
diff changeset
    42
is "op #"
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
syntax
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
  "@Finset"     :: "args => 'a fset"  ("{|(_)|}")
714
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    46
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
translations
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
  "{|x, xs|}" == "CONST finsert x {|xs|}"
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
  "{|x|}"     == "CONST finsert x {||}"
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
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
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
    52
  memb :: "'a \<Rightarrow> 'a list \<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
    53
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
    54
  "memb x xs \<equiv> x \<in> set xs"
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
quotient_definition
1144
538daee762e6 Added a binding to the parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1141
diff changeset
    57
  fin ("_ |\<in>| _" [50, 51] 50)
538daee762e6 Added a binding to the parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1141
diff changeset
    58
where
538daee762e6 Added a binding to the parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1141
diff changeset
    59
  "fin :: 'a \<Rightarrow> 'a fset \<Rightarrow> bool"
1141
3c8ad149a4d3 Undid the read_terms change; now compiles.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1139
diff changeset
    60
is "memb"
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
    61
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
abbreviation
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
  fnotin :: "'a \<Rightarrow> 'a fset \<Rightarrow> bool" ("_ |\<notin>| _" [50, 51] 50)
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
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
    65
  "a |\<notin>| S \<equiv> \<not>(a |\<in>| S)"
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
lemma memb_rsp[quot_respect]: 
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
  shows "(op = ===> op \<approx> ===> op =) memb memb"
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
    69
by (auto simp add: memb_def)
714
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    70
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    71
lemma nil_rsp[quot_respect]:
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    72
  shows "[] \<approx> []"
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
    73
by simp
714
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    74
1850
05b2dd2b0e8a Prove insert_rsp2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1260
diff changeset
    75
lemma cons_rsp:
05b2dd2b0e8a Prove insert_rsp2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1260
diff changeset
    76
  "xa \<approx> ya \<Longrightarrow> y # xa \<approx> y # ya"
05b2dd2b0e8a Prove insert_rsp2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1260
diff changeset
    77
  by simp
05b2dd2b0e8a Prove insert_rsp2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1260
diff changeset
    78
05b2dd2b0e8a Prove insert_rsp2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1260
diff changeset
    79
lemma [quot_respect]:
722
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
    80
  shows "(op = ===> op \<approx> ===> op \<approx>) op # op #"
1850
05b2dd2b0e8a Prove insert_rsp2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1260
diff changeset
    81
  by simp
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
    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
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
    84
section {* Augmenting a set -- @{const finsert} *}
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
    85
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
    86
text {* raw section *}
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
    87
830
89d772dae4d4 New_relations, all works again including concat examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 827
diff changeset
    88
lemma nil_not_cons:
89d772dae4d4 New_relations, all works again including concat examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 827
diff changeset
    89
  shows "\<not>[] \<approx> x # xs"
89d772dae4d4 New_relations, all works again including concat examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 827
diff changeset
    90
  by auto
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
    91
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
    92
lemma memb_cons_iff:
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
    93
  shows "memb x (y # xs) = (x = y \<or> memb x xs)"
830
89d772dae4d4 New_relations, all works again including concat examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 827
diff changeset
    94
  by (induct xs) (auto simp add: memb_def)
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
    95
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
    96
lemma memb_consI1:
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
    97
  shows "memb x (x # xs)"
830
89d772dae4d4 New_relations, all works again including concat examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 827
diff changeset
    98
  by (simp add: memb_def)
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
    99
830
89d772dae4d4 New_relations, all works again including concat examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 827
diff changeset
   100
lemma memb_consI2:
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
   101
  shows "memb x xs \<Longrightarrow> memb x (y # xs)"
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
   102
  by (simp add: memb_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
   103
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
   104
lemma memb_absorb:
830
89d772dae4d4 New_relations, all works again including concat examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 827
diff changeset
   105
  shows "memb x xs \<Longrightarrow> x # xs \<approx> xs"
89d772dae4d4 New_relations, all works again including concat examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 827
diff changeset
   106
  by (induct xs) (auto simp add: memb_def id_simps)
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
   107
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
   108
text {* lifted section *}
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
   109
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
lemma fin_finsert_iff[simp]:
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
  "x |\<in>| finsert y S = (x = y \<or> x |\<in>| S)"
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
by (lifting memb_cons_iff) 
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
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
lemma
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
  shows finsertI1: "x |\<in>| finsert x S"
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
  and   finsertI2: "x |\<in>| S \<Longrightarrow> x |\<in>| finsert y S"
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
  by (lifting memb_consI1, lifting memb_consI2)
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
830
89d772dae4d4 New_relations, all works again including concat examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 827
diff changeset
   119
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
   120
lemma finsert_absorb [simp]: 
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
  shows "x |\<in>| S \<Longrightarrow> finsert x S = S"
830
89d772dae4d4 New_relations, all works again including concat examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 827
diff changeset
   122
  by (lifting memb_absorb)
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
   123
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
   124
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
   125
section {* Singletons *}
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
   126
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
   127
text {* raw section *}
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
   128
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
   129
lemma singleton_list_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
   130
  shows "[x] \<approx> [y] \<longleftrightarrow> x = y"
830
89d772dae4d4 New_relations, all works again including concat examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 827
diff changeset
   131
  by (simp add: id_simps) auto
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
   132
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
   133
text {* lifted section *}
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
   134
830
89d772dae4d4 New_relations, all works again including concat examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 827
diff changeset
   135
lemma fempty_not_finsert[simp]:
89d772dae4d4 New_relations, all works again including concat examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 827
diff changeset
   136
  shows "{||} \<noteq> finsert x S"
89d772dae4d4 New_relations, all works again including concat examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 827
diff changeset
   137
  by (lifting nil_not_cons)
89d772dae4d4 New_relations, all works again including concat examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 827
diff changeset
   138
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
   139
lemma fsingleton_eq[simp]:
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
   140
  shows "{|x|} = {|y|} \<longleftrightarrow> x = y"
830
89d772dae4d4 New_relations, all works again including concat examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 827
diff changeset
   141
  by (lifting singleton_list_eq)
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
   142
793
09dff5ef8f74 as expected problems occure when lifting concat lemmas
Christian Urban <urbanc@in.tum.de>
parents: 787
diff changeset
   143
section {* Union *}
09dff5ef8f74 as expected problems occure when lifting concat lemmas
Christian Urban <urbanc@in.tum.de>
parents: 787
diff changeset
   144
09dff5ef8f74 as expected problems occure when lifting concat lemmas
Christian Urban <urbanc@in.tum.de>
parents: 787
diff changeset
   145
quotient_definition
1144
538daee762e6 Added a binding to the parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1141
diff changeset
   146
  funion  (infixl "|\<union>|" 65)
538daee762e6 Added a binding to the parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1141
diff changeset
   147
where
538daee762e6 Added a binding to the parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1141
diff changeset
   148
  "funion :: 'a fset \<Rightarrow> 'a fset \<Rightarrow> 'a fset"
1139
c4001cda9da3 renamed 'as' to 'is' everywhere.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1128
diff changeset
   149
is
793
09dff5ef8f74 as expected problems occure when lifting concat lemmas
Christian Urban <urbanc@in.tum.de>
parents: 787
diff changeset
   150
  "op @"
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
   151
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
   152
section {* Cardinality of finite sets *}
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
   153
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
   154
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
   155
  fcard_raw :: "'a list \<Rightarrow> nat"
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
   156
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
   157
  fcard_raw_nil:  "fcard_raw [] = 0"
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
   158
| fcard_raw_cons: "fcard_raw (x # xs) = (if memb x xs then fcard_raw xs else Suc (fcard_raw xs))"
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
   159
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
   160
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
   161
  "fcard :: 'a fset \<Rightarrow> nat" 
1141
3c8ad149a4d3 Undid the read_terms change; now compiles.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1139
diff changeset
   162
is "fcard_raw"
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
   163
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
   164
text {* raw section *}
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
   165
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
   166
lemma fcard_raw_ge_0:
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
   167
  assumes a: "x \<in> set xs"
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
   168
  shows "0 < fcard_raw xs"
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
   169
using 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
   170
by (induct xs) (auto simp add: memb_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
   171
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
   172
lemma fcard_raw_delete_one:
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
   173
  "fcard_raw ([x \<leftarrow> xs. x \<noteq> y]) = (if memb y xs then fcard_raw xs - 1 else fcard_raw xs)"
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
   174
by (induct xs) (auto dest: fcard_raw_ge_0 simp add: memb_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
   175
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
   176
lemma fcard_raw_rsp_aux:
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
   177
  assumes a: "a \<approx> b"
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
   178
  shows "fcard_raw a = fcard_raw b"
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
   179
using 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
   180
apply(induct a arbitrary: b)
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
   181
apply(auto simp add: memb_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
   182
apply(metis)
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
   183
apply(drule_tac x="[x \<leftarrow> b. x \<noteq> a1]" in meta_spec)
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
   184
apply(simp add: fcard_raw_delete_one)
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
   185
apply(metis Suc_pred' fcard_raw_ge_0 fcard_raw_delete_one memb_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
   186
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
   187
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
   188
lemma fcard_raw_rsp[quot_respect]:
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
   189
  "(op \<approx> ===> op =) fcard_raw fcard_raw"
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
   190
  by (simp add: fcard_raw_rsp_aux)
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
   191
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
   192
text {* lifted section *}
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
   193
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
   194
lemma fcard_fempty [simp]:
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
   195
  shows "fcard {||} = 0"
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
   196
by (lifting fcard_raw_nil)
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
   197
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
   198
lemma fcard_finsert_if [simp]:
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
   199
  shows "fcard (finsert x S) = (if x |\<in>| S then fcard S else Suc (fcard S))"
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
   200
by (lifting fcard_raw_cons)
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
   201
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
   202
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
   203
section {* Induction and Cases rules for finite sets *}
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
   204
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
   205
lemma fset_exhaust[case_names fempty finsert, cases type: 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
   206
  shows "\<lbrakk>S = {||} \<Longrightarrow> P; \<And>x S'. S = finsert x S' \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P"
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
   207
by (lifting list.exhaust)
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
   208
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
   209
lemma fset_induct[case_names fempty finsert]:
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
   210
  shows "\<lbrakk>P {||}; \<And>x S. P S \<Longrightarrow> P (finsert x S)\<rbrakk> \<Longrightarrow> P S"
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
   211
by (lifting list.induct)
714
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   212
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
   213
lemma fset_induct2[case_names fempty finsert, induct type: 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
   214
  assumes prem1: "P {||}" 
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
   215
  and     prem2: "\<And>x S. \<lbrakk>x |\<notin>| S; P S\<rbrakk> \<Longrightarrow> P (finsert x S)"
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
   216
  shows "P S"
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
   217
proof(induct S rule: fset_induct)
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
   218
  case fempty
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
   219
  show "P {||}" by (rule prem1)
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
   220
next
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
   221
  case (finsert x S)
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
   222
  have asm: "P S" by fact
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
   223
  show "P (finsert x S)"
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
  proof(cases "x |\<in>| S")
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
    case True
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
   226
    have "x |\<in>| S" by fact
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
   227
    then show "P (finsert x S)" using asm by simp
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
   228
  next
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
   229
    case False
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
   230
    have "x |\<notin>| S" by fact
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
   231
    then show "P (finsert x S)" using prem2 asm by simp
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
   232
  qed
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
   233
qed
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
   234
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
   235
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
   236
section {* fmap and fset comprehension *}
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
   237
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
   238
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
   239
  "fmap :: ('a \<Rightarrow> 'b) \<Rightarrow> 'a fset \<Rightarrow> 'b fset"
1139
c4001cda9da3 renamed 'as' to 'is' everywhere.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1128
diff changeset
   240
is
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
   241
 "map"
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
   242
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
   243
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
   244
  "fconcat :: ('a fset) fset => 'a fset"
1139
c4001cda9da3 renamed 'as' to 'is' everywhere.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1128
diff changeset
   245
is
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
   246
 "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
   247
814
cd3fa86be45f Sledgehammer bug.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 798
diff changeset
   248
(*lemma fconcat_rsp[quot_respect]:
793
09dff5ef8f74 as expected problems occure when lifting concat lemmas
Christian Urban <urbanc@in.tum.de>
parents: 787
diff changeset
   249
  shows "((list_rel op \<approx>) ===> op \<approx>) concat concat"
09dff5ef8f74 as expected problems occure when lifting concat lemmas
Christian Urban <urbanc@in.tum.de>
parents: 787
diff changeset
   250
apply(auto)
09dff5ef8f74 as expected problems occure when lifting concat lemmas
Christian Urban <urbanc@in.tum.de>
parents: 787
diff changeset
   251
sorry
814
cd3fa86be45f Sledgehammer bug.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 798
diff changeset
   252
*)
793
09dff5ef8f74 as expected problems occure when lifting concat lemmas
Christian Urban <urbanc@in.tum.de>
parents: 787
diff changeset
   253
09dff5ef8f74 as expected problems occure when lifting concat lemmas
Christian Urban <urbanc@in.tum.de>
parents: 787
diff changeset
   254
(* PROBLEM: these lemmas needs to be restated, since  *)
09dff5ef8f74 as expected problems occure when lifting concat lemmas
Christian Urban <urbanc@in.tum.de>
parents: 787
diff changeset
   255
(* concat.simps(1) and concat.simps(2) contain the    *)
09dff5ef8f74 as expected problems occure when lifting concat lemmas
Christian Urban <urbanc@in.tum.de>
parents: 787
diff changeset
   256
(* type variables ?'a1.0 (which are turned into frees *)
09dff5ef8f74 as expected problems occure when lifting concat lemmas
Christian Urban <urbanc@in.tum.de>
parents: 787
diff changeset
   257
(* 'a_1                                               *)
830
89d772dae4d4 New_relations, all works again including concat examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 827
diff changeset
   258
lemma concat1:
797
35436401f00d added a functor that allows checking what is added to the theorem lists
Christian Urban <urbanc@in.tum.de>
parents: 796
diff changeset
   259
  shows "concat [] \<approx> []"
830
89d772dae4d4 New_relations, all works again including concat examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 827
diff changeset
   260
by (simp add: id_simps)
793
09dff5ef8f74 as expected problems occure when lifting concat lemmas
Christian Urban <urbanc@in.tum.de>
parents: 787
diff changeset
   261
830
89d772dae4d4 New_relations, all works again including concat examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 827
diff changeset
   262
lemma concat2:
89d772dae4d4 New_relations, all works again including concat examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 827
diff changeset
   263
  shows "concat (x # xs) \<approx> x @ concat xs"
1850
05b2dd2b0e8a Prove insert_rsp2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1260
diff changeset
   264
by (simp)
722
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
   265
814
cd3fa86be45f Sledgehammer bug.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 798
diff changeset
   266
lemma concat_rsp[quot_respect]:
851
e1473b4b2886 Change OO to OOO in FSet3.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 833
diff changeset
   267
  shows "(list_rel op \<approx> OOO op \<approx> ===> op \<approx>) concat concat"
1850
05b2dd2b0e8a Prove insert_rsp2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1260
diff changeset
   268
  sorry
814
cd3fa86be45f Sledgehammer bug.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 798
diff changeset
   269
851
e1473b4b2886 Change OO to OOO in FSet3.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 833
diff changeset
   270
lemma nil_rsp2[quot_respect]: "(list_rel op \<approx> OOO op \<approx>) [] []"
830
89d772dae4d4 New_relations, all works again including concat examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 827
diff changeset
   271
  apply (metis FSet3.nil_rsp list_rel.simps(1) pred_comp.intros)
89d772dae4d4 New_relations, all works again including concat examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 827
diff changeset
   272
  done
814
cd3fa86be45f Sledgehammer bug.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 798
diff changeset
   273
824
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
   274
lemma set_in_eq: "(\<forall>e. ((e \<in> A) \<longleftrightarrow> (e \<in> B))) \<equiv> A = B"
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
   275
  apply (rule eq_reflection)
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
   276
  apply auto
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
   277
  done
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
   278
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
   279
lemma map_rel_cong: "b \<approx> ba \<Longrightarrow> map f b \<approx> map f ba"
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
   280
  unfolding list_eq.simps
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
   281
  apply(simp only: set_map set_in_eq)
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
   282
  done
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
   283
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
   284
lemma quotient_compose_list_pre:
851
e1473b4b2886 Change OO to OOO in FSet3.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 833
diff changeset
   285
  "(list_rel op \<approx> OOO op \<approx>) r s =
e1473b4b2886 Change OO to OOO in FSet3.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 833
diff changeset
   286
  ((list_rel op \<approx> OOO op \<approx>) r r \<and> (list_rel op \<approx> OOO op \<approx>) s s \<and>
824
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
   287
  abs_fset (map abs_fset r) = abs_fset (map abs_fset s))"
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
   288
  apply rule
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
   289
  apply rule
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
   290
  apply rule
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
   291
  apply (rule list_rel_refl)
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
   292
  apply (metis equivp_def fset_equivp)
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
   293
  apply rule
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
   294
  apply (rule equivp_reflp[OF fset_equivp])
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
   295
  apply (rule list_rel_refl)
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
   296
  apply (metis equivp_def fset_equivp)
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
   297
  apply(rule)
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
   298
  apply rule
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
   299
  apply (rule list_rel_refl)
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
   300
  apply (metis equivp_def fset_equivp)
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
   301
  apply rule
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
   302
  apply (rule equivp_reflp[OF fset_equivp])
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
   303
  apply (rule list_rel_refl)
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
   304
  apply (metis equivp_def fset_equivp)
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
   305
  apply (subgoal_tac "map abs_fset r \<approx> map abs_fset s")
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
   306
  apply (metis Quotient_rel[OF Quotient_fset])
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
   307
  apply (auto simp only:)[1]
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
   308
  apply (subgoal_tac "map abs_fset r = map abs_fset b")
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
   309
  prefer 2
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
   310
  apply (metis Quotient_rel[OF list_quotient[OF Quotient_fset]])
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
   311
  apply (subgoal_tac "map abs_fset s = map abs_fset ba")
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
   312
  prefer 2
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
   313
  apply (metis Quotient_rel[OF list_quotient[OF Quotient_fset]])
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
   314
  apply (simp only: map_rel_cong)
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
   315
  apply rule
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
   316
  apply (rule rep_abs_rsp[of "list_rel op \<approx>" "map abs_fset"])
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
   317
  apply (tactic {* Quotient_Tacs.quotient_tac @{context} 1 *})
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
   318
  apply (rule list_rel_refl)
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
   319
  apply (metis equivp_def fset_equivp)
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
   320
  apply rule
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
   321
  prefer 2
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
   322
  apply (rule rep_abs_rsp_left[of "list_rel op \<approx>" "map abs_fset"])
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
   323
  apply (tactic {* Quotient_Tacs.quotient_tac @{context} 1 *})
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
   324
  apply (rule list_rel_refl)
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
   325
  apply (metis equivp_def fset_equivp)
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
   326
  apply (erule conjE)+
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
   327
  apply (subgoal_tac "map abs_fset r \<approx> map abs_fset s")
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
   328
  prefer 2
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
   329
  apply (metis Quotient_def Quotient_fset equivp_reflp fset_equivp)
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
   330
  apply (rule map_rel_cong)
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
   331
  apply (assumption)
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
   332
  done
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
   333
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
   334
lemma quotient_compose_list[quot_thm]:
851
e1473b4b2886 Change OO to OOO in FSet3.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 833
diff changeset
   335
  shows  "Quotient ((list_rel op \<approx>) OOO (op \<approx>))
e1473b4b2886 Change OO to OOO in FSet3.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 833
diff changeset
   336
    (abs_fset \<circ> (map abs_fset)) ((map rep_fset) \<circ> rep_fset)"
824
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
   337
  unfolding Quotient_def comp_def
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
   338
  apply (rule)+
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
   339
  apply (simp add: abs_o_rep[OF Quotient_fset] id_simps Quotient_abs_rep[OF Quotient_fset])
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
   340
  apply (rule)
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
   341
  apply (rule)
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
   342
  apply (rule)
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
   343
  apply (rule list_rel_refl)
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
   344
  apply (metis equivp_def fset_equivp)
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
   345
  apply (rule)
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
   346
  apply (rule equivp_reflp[OF fset_equivp])
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
   347
  apply (rule list_rel_refl)
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
   348
  apply (metis equivp_def fset_equivp)
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
   349
  apply rule
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
   350
  apply rule
1850
05b2dd2b0e8a Prove insert_rsp2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1260
diff changeset
   351
  apply (rule quotient_compose_list_pre)
824
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
   352
  done
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
   353
793
09dff5ef8f74 as expected problems occure when lifting concat lemmas
Christian Urban <urbanc@in.tum.de>
parents: 787
diff changeset
   354
lemma fconcat_empty:
09dff5ef8f74 as expected problems occure when lifting concat lemmas
Christian Urban <urbanc@in.tum.de>
parents: 787
diff changeset
   355
  shows "fconcat {||} = {||}"
1850
05b2dd2b0e8a Prove insert_rsp2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1260
diff changeset
   356
  apply(lifting concat1)
05b2dd2b0e8a Prove insert_rsp2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1260
diff changeset
   357
  apply(cleaning)
05b2dd2b0e8a Prove insert_rsp2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1260
diff changeset
   358
  apply(simp add: comp_def)
05b2dd2b0e8a Prove insert_rsp2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1260
diff changeset
   359
  apply(fold fempty_def[simplified id_simps])
05b2dd2b0e8a Prove insert_rsp2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1260
diff changeset
   360
  apply(rule refl)
05b2dd2b0e8a Prove insert_rsp2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1260
diff changeset
   361
  done
793
09dff5ef8f74 as expected problems occure when lifting concat lemmas
Christian Urban <urbanc@in.tum.de>
parents: 787
diff changeset
   362
830
89d772dae4d4 New_relations, all works again including concat examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 827
diff changeset
   363
(* Should be true *)
89d772dae4d4 New_relations, all works again including concat examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 827
diff changeset
   364
89d772dae4d4 New_relations, all works again including concat examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 827
diff changeset
   365
lemma insert_rsp2[quot_respect]:
851
e1473b4b2886 Change OO to OOO in FSet3.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 833
diff changeset
   366
  "(op \<approx> ===> list_rel op \<approx> OOO op \<approx> ===> list_rel op \<approx> OOO op \<approx>) op # op #"
1850
05b2dd2b0e8a Prove insert_rsp2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1260
diff changeset
   367
  apply auto
05b2dd2b0e8a Prove insert_rsp2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1260
diff changeset
   368
  apply (simp add: set_in_eq)
05b2dd2b0e8a Prove insert_rsp2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1260
diff changeset
   369
  apply (rule_tac b="x # b" in pred_compI)
05b2dd2b0e8a Prove insert_rsp2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1260
diff changeset
   370
  apply auto
05b2dd2b0e8a Prove insert_rsp2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1260
diff changeset
   371
  apply (rule_tac b="x # ba" in pred_compI)
05b2dd2b0e8a Prove insert_rsp2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1260
diff changeset
   372
  apply (rule cons_rsp)
05b2dd2b0e8a Prove insert_rsp2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1260
diff changeset
   373
  apply simp
05b2dd2b0e8a Prove insert_rsp2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1260
diff changeset
   374
  apply (auto)[1]
05b2dd2b0e8a Prove insert_rsp2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1260
diff changeset
   375
  done
830
89d772dae4d4 New_relations, all works again including concat examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 827
diff changeset
   376
89d772dae4d4 New_relations, all works again including concat examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 827
diff changeset
   377
lemma append_rsp[quot_respect]:
89d772dae4d4 New_relations, all works again including concat examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 827
diff changeset
   378
  "(op \<approx> ===> op \<approx> ===> op \<approx>) op @ op @"
89d772dae4d4 New_relations, all works again including concat examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 827
diff changeset
   379
  by (auto)
89d772dae4d4 New_relations, all works again including concat examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 827
diff changeset
   380
793
09dff5ef8f74 as expected problems occure when lifting concat lemmas
Christian Urban <urbanc@in.tum.de>
parents: 787
diff changeset
   381
lemma fconcat_insert:
09dff5ef8f74 as expected problems occure when lifting concat lemmas
Christian Urban <urbanc@in.tum.de>
parents: 787
diff changeset
   382
  shows "fconcat (finsert x S) = x |\<union>| fconcat S"
1850
05b2dd2b0e8a Prove insert_rsp2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1260
diff changeset
   383
  apply(lifting concat2)
05b2dd2b0e8a Prove insert_rsp2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1260
diff changeset
   384
  apply(cleaning)
05b2dd2b0e8a Prove insert_rsp2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1260
diff changeset
   385
  apply (simp add: finsert_def fconcat_def comp_def)
05b2dd2b0e8a Prove insert_rsp2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1260
diff changeset
   386
  apply cleaning
05b2dd2b0e8a Prove insert_rsp2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1260
diff changeset
   387
  done
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
   388
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
   389
text {* raw section *}
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
   390
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
   391
lemma map_rsp_aux:
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
   392
  assumes a: "a \<approx> b"
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
   393
  shows "map f a \<approx> map f b"
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
   394
  using 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
   395
apply(induct a arbitrary: b)
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
   396
apply(auto)
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
   397
apply(metis rev_image_eqI)
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
   398
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
   399
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
   400
lemma map_rsp[quot_respect]:
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
   401
  shows "(op = ===> op \<approx> ===> op \<approx>) map map"
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
   402
by (auto simp add: map_rsp_aux)
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
   403
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
   404
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
   405
text {* lifted section *}
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
   406
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
   407
(* TBD *)
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
   408
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
   409
text {* syntax for fset comprehensions (adapted from lists) *}
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
   410
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
   411
nonterminals fsc_qual fsc_quals
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
   412
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
   413
syntax
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
   414
"_fsetcompr" :: "'a \<Rightarrow> fsc_qual \<Rightarrow> fsc_quals \<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
   415
"_fsc_gen" :: "'a \<Rightarrow> 'a fset \<Rightarrow> fsc_qual" ("_ <- _")
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
   416
"_fsc_test" :: "bool \<Rightarrow> fsc_qual" ("_")
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
   417
"_fsc_end" :: "fsc_quals" ("|}")
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
   418
"_fsc_quals" :: "fsc_qual \<Rightarrow> fsc_quals \<Rightarrow> fsc_quals" (", __")
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
   419
"_fsc_abs" :: "'a => 'b fset => 'b 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
   420
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
   421
syntax (xsymbols)
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
   422
"_fsc_gen" :: "'a \<Rightarrow> 'a fset \<Rightarrow> fsc_qual" ("_ \<leftarrow> _")
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
   423
syntax (HTML output)
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
   424
"_fsc_gen" :: "'a \<Rightarrow> 'a fset \<Rightarrow> fsc_qual" ("_ \<leftarrow> _")
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
   425
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
   426
parse_translation (advanced) {*
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
   427
let
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
   428
  val femptyC = Syntax.const @{const_name fempty};
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
   429
  val finsertC = Syntax.const @{const_name finsert};
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
   430
  val fmapC = Syntax.const @{const_name fmap};
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
   431
  val fconcatC = Syntax.const @{const_name 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
   432
  val IfC = Syntax.const @{const_name If};
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
   433
  fun fsingl x = finsertC $ x $ femptyC;
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
   434
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
   435
  fun pat_tr ctxt p e opti = (* %x. case x of p => e | _ => [] *)
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
   436
    let
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
   437
      val x = Free (Name.variant (fold Term.add_free_names [p, e] []) "x", dummyT);
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
   438
      val e = if opti then fsingl e else e;
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
   439
      val case1 = Syntax.const "_case1" $ p $ e;
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
   440
      val case2 = Syntax.const "_case1" $ Syntax.const Term.dummy_patternN
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
   441
                                        $ femptyC;
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
   442
      val cs = Syntax.const "_case2" $ case1 $ case2
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
   443
      val ft = Datatype_Case.case_tr false Datatype.info_of_constr
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
   444
                 ctxt [x, cs]
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
   445
    in lambda x ft end;
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
   446
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
   447
  fun abs_tr ctxt (p as Free(s,T)) e opti =
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
   448
        let val thy = ProofContext.theory_of ctxt;
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
   449
            val s' = Sign.intern_const thy s
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
   450
        in if Sign.declared_const thy s'
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
   451
           then (pat_tr ctxt p e opti, false)
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
   452
           else (lambda p e, true)
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
   453
        end
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
   454
    | abs_tr ctxt p e opti = (pat_tr ctxt p e opti, false);
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
   455
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
   456
  fun fsc_tr ctxt [e, Const("_fsc_test",_) $ b, qs] =
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
   457
        let 
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
   458
          val res = case qs of 
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
   459
                      Const("_fsc_end",_) => fsingl e
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
   460
                    | Const("_fsc_quals",_)$ q $ qs => fsc_tr ctxt [e, q, qs];
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
   461
        in 
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
   462
          IfC $ b $ res $ femptyC 
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
   463
        end
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
   464
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
   465
    | fsc_tr ctxt [e, Const("_fsc_gen",_) $ p $ es, Const("_fsc_end",_)] =
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
   466
         (case abs_tr ctxt p e true of
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
   467
            (f,true) => fmapC $ f $ es
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
   468
          | (f, false) => fconcatC $ (fmapC $ f $ es))
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
   469
       
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
   470
    | fsc_tr ctxt [e, Const("_fsc_gen",_) $ p $ es, Const("_fsc_quals",_) $ q $ qs] =
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
   471
        let
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
   472
          val e' = fsc_tr ctxt [e, q, qs];
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
   473
        in 
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
   474
          fconcatC $ (fmapC $ (fst (abs_tr ctxt p e' false)) $ es) 
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
   475
        end
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
   476
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
   477
in [("_fsetcompr", fsc_tr)] end
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
   478
*}
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
   479
1260
9df6144e281b moved Quot package to Attic (still compiles there with "isabelle make")
Christian Urban <urbanc@in.tum.de>
parents: 1144
diff changeset
   480
9df6144e281b moved Quot package to Attic (still compiles there with "isabelle make")
Christian Urban <urbanc@in.tum.de>
parents: 1144
diff changeset
   481
(* NEEDS FIXING *)
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
   482
(* examles *)
1260
9df6144e281b moved Quot package to Attic (still compiles there with "isabelle make")
Christian Urban <urbanc@in.tum.de>
parents: 1144
diff changeset
   483
(*
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
   484
term "{|(x,y,z). b|}"
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
   485
term "{|x. x \<leftarrow> xs|}"
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
   486
term "{|(x,y,z). x\<leftarrow>xs|}"
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
   487
term "{|e x y. x\<leftarrow>xs, y\<leftarrow>ys|}"
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
   488
term "{|(x,y,z). x<a, x>b|}"
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
   489
term "{|(x,y,z). x\<leftarrow>xs, x>b|}"
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
   490
term "{|(x,y,z). x<a, x\<leftarrow>xs|}"
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
   491
term "{|(x,y). Cons True x \<leftarrow> xs|}"
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
   492
term "{|(x,y,z). Cons x [] \<leftarrow> xs|}"
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
   493
term "{|(x,y,z). x<a, x>b, x=d|}"
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
   494
term "{|(x,y,z). x<a, x>b, y\<leftarrow>ys|}"
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
   495
term "{|(x,y,z). x<a, x\<leftarrow>xs,y>b|}"
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
   496
term "{|(x,y,z). x<a, x\<leftarrow>xs, y\<leftarrow>ys|}"
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
   497
term "{|(x,y,z). x\<leftarrow>xs, x>b, y<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
   498
term "{|(x,y,z). x\<leftarrow>xs, x>b, y\<leftarrow>ys|}"
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
   499
term "{|(x,y,z). x\<leftarrow>xs, y\<leftarrow>ys,y>x|}"
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
   500
term "{|(x,y,z). x\<leftarrow>xs, y\<leftarrow>ys,z\<leftarrow>zs|}"
1260
9df6144e281b moved Quot package to Attic (still compiles there with "isabelle make")
Christian Urban <urbanc@in.tum.de>
parents: 1144
diff changeset
   501
*)
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
   502
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
   503
(* BELOW CONSTRUCTION SITE *)
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
   504
714
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   505
722
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
   506
lemma no_mem_nil: 
729
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 728
diff changeset
   507
  "(\<forall>a. a \<notin> set A) = (A = [])"
722
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
   508
by (induct A) (auto)
714
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   509
722
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
   510
lemma none_mem_nil: 
729
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 728
diff changeset
   511
  "(\<forall>a. a \<notin> set A) = (A \<approx> [])"
722
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
   512
by simp
714
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   513
722
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
   514
lemma mem_cons: 
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
   515
  "a \<in> set A \<Longrightarrow> a # A \<approx> A"
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
   516
by auto
714
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   517
830
89d772dae4d4 New_relations, all works again including concat examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 827
diff changeset
   518
lemma cons_left_comm:
89d772dae4d4 New_relations, all works again including concat examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 827
diff changeset
   519
  "x # y # A \<approx> y # x # A"
89d772dae4d4 New_relations, all works again including concat examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 827
diff changeset
   520
by (auto simp add: id_simps)
714
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   521
830
89d772dae4d4 New_relations, all works again including concat examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 827
diff changeset
   522
lemma cons_left_idem:
722
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
   523
  "x # x # A \<approx> x # A"
830
89d772dae4d4 New_relations, all works again including concat examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 827
diff changeset
   524
by (auto simp add: id_simps)
714
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   525
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   526
lemma finite_set_raw_strong_cases:
727
2cfe6f3d6352 Proof of finite_set_storng_cases_raw.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 726
diff changeset
   527
  "(X = []) \<or> (\<exists>a Y. ((a \<notin> set Y) \<and> (X \<approx> a # Y)))"
714
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   528
  apply (induct X)
727
2cfe6f3d6352 Proof of finite_set_storng_cases_raw.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 726
diff changeset
   529
  apply (simp)
2cfe6f3d6352 Proof of finite_set_storng_cases_raw.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 726
diff changeset
   530
  apply (rule disjI2)
2cfe6f3d6352 Proof of finite_set_storng_cases_raw.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 726
diff changeset
   531
  apply (erule disjE)
2cfe6f3d6352 Proof of finite_set_storng_cases_raw.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 726
diff changeset
   532
  apply (rule_tac x="a" in exI)
2cfe6f3d6352 Proof of finite_set_storng_cases_raw.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 726
diff changeset
   533
  apply (rule_tac x="[]" in exI)
2cfe6f3d6352 Proof of finite_set_storng_cases_raw.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 726
diff changeset
   534
  apply (simp)
2cfe6f3d6352 Proof of finite_set_storng_cases_raw.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 726
diff changeset
   535
  apply (erule exE)+
2cfe6f3d6352 Proof of finite_set_storng_cases_raw.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 726
diff changeset
   536
  apply (case_tac "a = aa")
2cfe6f3d6352 Proof of finite_set_storng_cases_raw.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 726
diff changeset
   537
  apply (rule_tac x="a" in exI)
2cfe6f3d6352 Proof of finite_set_storng_cases_raw.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 726
diff changeset
   538
  apply (rule_tac x="Y" in exI)
2cfe6f3d6352 Proof of finite_set_storng_cases_raw.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 726
diff changeset
   539
  apply (simp)
2cfe6f3d6352 Proof of finite_set_storng_cases_raw.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 726
diff changeset
   540
  apply (rule_tac x="aa" in exI)
2cfe6f3d6352 Proof of finite_set_storng_cases_raw.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 726
diff changeset
   541
  apply (rule_tac x="a # Y" in exI)
722
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
   542
  apply (auto)
727
2cfe6f3d6352 Proof of finite_set_storng_cases_raw.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 726
diff changeset
   543
  done
714
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   544
722
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
   545
fun
714
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   546
  delete_raw :: "'a list \<Rightarrow> 'a \<Rightarrow> 'a list"
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   547
where
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   548
  "delete_raw [] x = []"
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   549
| "delete_raw (a # A) x = (if (a = x) then delete_raw A x else a # (delete_raw A x))"
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   550
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   551
lemma mem_delete_raw:
728
0015159fee96 Some proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 727
diff changeset
   552
  "x \<in> set (delete_raw A a) = (x \<in> set A \<and> \<not>(x = a))"
726
1a777307f57f A bracket was missing; with it proved the 'definitely false' lemma.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 724
diff changeset
   553
  by (induct A arbitrary: x a) (auto)
714
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   554
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   555
lemma mem_delete_raw_ident:
722
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
   556
  "\<not>(a \<in> set (delete_raw A a))"
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
   557
by (induct A) (auto)
714
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   558
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   559
lemma not_mem_delete_raw_ident:
722
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
   560
  "b \<notin> set A \<Longrightarrow> (delete_raw A b = A)"
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
   561
by (induct A) (auto)
714
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   562
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   563
lemma delete_raw_RSP:
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   564
  "A \<approx> B \<Longrightarrow> delete_raw A a \<approx> delete_raw B a"
722
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
   565
apply(induct A arbitrary: B a)
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
   566
apply(auto)
714
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   567
sorry
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   568
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   569
lemma cons_delete_raw:
728
0015159fee96 Some proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 727
diff changeset
   570
  "a # (delete_raw A a) \<approx> (if a \<in> set A then A else (a # A))"
714
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   571
sorry
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   572
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   573
lemma mem_cons_delete_raw:
726
1a777307f57f A bracket was missing; with it proved the 'definitely false' lemma.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 724
diff changeset
   574
    "a \<in> set A \<Longrightarrow> a # (delete_raw A a) \<approx> A"
714
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   575
sorry
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   576
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   577
lemma finite_set_raw_delete_raw_cases:
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   578
    "X = [] \<or> (\<exists>a. a mem X \<and> X \<approx> a # delete_raw X a)"
726
1a777307f57f A bracket was missing; with it proved the 'definitely false' lemma.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 724
diff changeset
   579
  by (induct X) (auto)
714
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   580
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   581
722
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
   582
714
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   583
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   584
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   585
lemma list2set_thm:
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   586
  shows "set [] = {}"
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   587
  and "set (h # t) = insert h (set t)"
728
0015159fee96 Some proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 727
diff changeset
   588
  by (auto)
714
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   589
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   590
lemma list2set_RSP:
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   591
  "A \<approx> B \<Longrightarrow> set A = set B"
728
0015159fee96 Some proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 727
diff changeset
   592
  by auto
714
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   593
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   594
definition
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   595
  rsp_fold
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   596
where
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   597
  "rsp_fold f = (\<forall>u v w. (f u (f v w) = f v (f u w)))"
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   598
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   599
primrec
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   600
  fold_raw :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'b \<Rightarrow> 'a list \<Rightarrow> 'b"
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   601
where
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   602
  "fold_raw f z [] = z"
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   603
| "fold_raw f z (a # A) =
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   604
     (if (rsp_fold f) then
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   605
       if a mem A then fold_raw f z A
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   606
       else f a (fold_raw f z A)
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   607
     else z)"
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   608
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   609
lemma mem_lcommuting_fold_raw:
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   610
  "rsp_fold f \<Longrightarrow> h mem B \<Longrightarrow> fold_raw f z B = f h (fold_raw f z (delete_raw B h))"
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   611
sorry
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   612
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   613
lemma fold_rsp[quot_respect]:
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   614
  "(op = ===> op = ===> op \<approx> ===> op =) fold_raw fold_raw"
722
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
   615
apply(auto)
714
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   616
sorry
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   617
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   618
primrec
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   619
  inter_raw
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   620
where
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   621
  "inter_raw [] B = []"
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   622
| "inter_raw (a # A) B = (if a mem B then a # inter_raw A B else inter_raw A B)"
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   623
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   624
lemma mem_inter_raw:
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   625
  "x mem (inter_raw A B) = x mem A \<and> x mem B"
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   626
sorry
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   627
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   628
lemma inter_raw_RSP:
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   629
  "A1 \<approx> A2 \<and> B1 \<approx> B2 \<Longrightarrow> (inter_raw A1 B1) \<approx> (inter_raw A2 B2)"
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   630
sorry
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   631
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   632
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   633
(* LIFTING DEFS *)
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   634
722
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
   635
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
   636
section {* Constants on the Quotient Type *} 
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
   637
714
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   638
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
   639
quotient_definition
722
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
   640
  "fdelete :: 'a fset \<Rightarrow> 'a \<Rightarrow> 'a fset" 
1139
c4001cda9da3 renamed 'as' to 'is' everywhere.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1128
diff changeset
   641
  is "delete_raw"
722
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
   642
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
   643
quotient_definition
1144
538daee762e6 Added a binding to the parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1141
diff changeset
   644
  finter ("_ \<inter>f _" [70, 71] 70)
538daee762e6 Added a binding to the parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1141
diff changeset
   645
where
538daee762e6 Added a binding to the parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1141
diff changeset
   646
  "finter :: 'a fset \<Rightarrow> 'a fset \<Rightarrow> 'a fset"
1139
c4001cda9da3 renamed 'as' to 'is' everywhere.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1128
diff changeset
   647
  is "inter_raw"
714
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   648
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
   649
quotient_definition
722
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
   650
  "ffold :: ('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'b \<Rightarrow> 'a fset \<Rightarrow> 'b" 
1139
c4001cda9da3 renamed 'as' to 'is' everywhere.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1128
diff changeset
   651
  is "fold_raw"
714
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   652
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
   653
quotient_definition
722
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
   654
  "fset_to_set :: 'a fset \<Rightarrow> 'a set" 
1139
c4001cda9da3 renamed 'as' to 'is' everywhere.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1128
diff changeset
   655
  is "set"
714
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   656
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   657
722
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
   658
section {* Lifted Theorems *}
714
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   659
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   660
thm list.cases (* ??? *)
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   661
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   662
thm cons_left_comm
722
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
   663
lemma "finsert a (finsert b S) = finsert b (finsert a S)"
714
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   664
by (lifting cons_left_comm)
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   665
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   666
thm cons_left_idem
722
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
   667
lemma "finsert a (finsert a S) = finsert a S"
714
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   668
by (lifting cons_left_idem)
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   669
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   670
(* thm MEM:
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   671
  MEM x [] = F
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   672
  MEM x (h::t) = (x=h) \/ MEM x t *)
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   673
thm none_mem_nil
729
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 728
diff changeset
   674
(*lemma "(\<forall>a. a \<notin>f A) = (A = fempty)"*)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 728
diff changeset
   675
714
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   676
thm mem_cons
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   677
thm finite_set_raw_strong_cases
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
   678
(*thm card_raw.simps*)
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
   679
(*thm not_mem_card_raw*)
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
   680
(*thm card_raw_suc*)
722
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
   681
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
   682
lemma "fcard X = Suc n \<Longrightarrow> (\<exists>a S. a \<notin>f S & X = finsert a S)"
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
   683
(*by (lifting card_raw_suc)*)
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
   684
sorry
714
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   685
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
   686
(*thm card_raw_cons_gt_0
743
4b3822d1ed24 Replies to questions from the weekend: Uncommenting the renamed theorem commented out in 734.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 734
diff changeset
   687
thm mem_card_raw_gt_0
714
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   688
thm not_nil_equiv_cons
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
   689
*)
714
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   690
thm delete_raw.simps
722
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
   691
(*thm mem_delete_raw*)
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
   692
(*thm card_raw_delete_raw*)
714
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   693
thm cons_delete_raw
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   694
thm mem_cons_delete_raw
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   695
thm finite_set_raw_delete_raw_cases
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   696
thm append.simps
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   697
(* MEM_APPEND: MEM e (APPEND l1 l2) = MEM e l1 \/ MEM e l2 *)
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   698
thm inter_raw.simps
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   699
thm mem_inter_raw
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   700
thm fold_raw.simps
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   701
thm list2set_thm
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   702
thm list_eq_def
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   703
thm list.induct
722
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
   704
lemma "\<lbrakk>P fempty; \<And>a x. P x \<Longrightarrow> P (finsert a x)\<rbrakk> \<Longrightarrow> P l"
714
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   705
by (lifting list.induct)
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   706
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   707
(* We also have map and some properties of it in FSet *)
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   708
(* and the following which still lifts ok *)
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   709
lemma "funion (funion x xa) xb = funion x (funion xa xb)"
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   710
by (lifting append_assoc)
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   711
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
   712
quotient_definition
716
1e08743b6997 FSet3 minor fixes + cases
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 714
diff changeset
   713
  "fset_case :: 'a \<Rightarrow> ('b \<Rightarrow> 'b fset \<Rightarrow> 'a) \<Rightarrow> 'b fset \<Rightarrow> 'a"
1139
c4001cda9da3 renamed 'as' to 'is' everywhere.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1128
diff changeset
   714
is
716
1e08743b6997 FSet3 minor fixes + cases
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 714
diff changeset
   715
  "list_case"
1e08743b6997 FSet3 minor fixes + cases
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 714
diff changeset
   716
1e08743b6997 FSet3 minor fixes + cases
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 714
diff changeset
   717
(* NOT SURE IF TRUE *)
1e08743b6997 FSet3 minor fixes + cases
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 714
diff changeset
   718
lemma list_case_rsp[quot_respect]:
1e08743b6997 FSet3 minor fixes + cases
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 714
diff changeset
   719
  "(op = ===> (op = ===> op \<approx> ===> op =) ===> op \<approx> ===> op =) list_case list_case"
1e08743b6997 FSet3 minor fixes + cases
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 714
diff changeset
   720
  apply (auto)
1e08743b6997 FSet3 minor fixes + cases
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 714
diff changeset
   721
  sorry
1e08743b6997 FSet3 minor fixes + cases
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 714
diff changeset
   722
722
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
   723
lemma "fset_case (f1::'t) f2 (finsert a xa) = f2 a xa"
716
1e08743b6997 FSet3 minor fixes + cases
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 714
diff changeset
   724
apply (lifting list.cases(2))
1e08743b6997 FSet3 minor fixes + cases
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 714
diff changeset
   725
done
1e08743b6997 FSet3 minor fixes + cases
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 714
diff changeset
   726
714
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   727
end