Attic/Quot/Examples/FSet3.thy
author Cezary Kaliszyk <kaliszyk@in.tum.de>
Wed, 21 Apr 2010 10:24:39 +0200
changeset 1916 c8b31085cb5b
parent 1914 c50601bc617e
child 1917 efbc22a6f1a4
permissions -rw-r--r--
FSet3 cleaning part2
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
1914
c50601bc617e Remove the part already in FSet and leave the experiments
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1873
diff changeset
     2
imports "../../../Nominal/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
     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
1914
c50601bc617e Remove the part already in FSet and leave the experiments
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1873
diff changeset
     5
notation
c50601bc617e Remove the part already in FSet and leave the experiments
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1873
diff changeset
     6
  list_eq (infix "\<approx>" 50)
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
     7
e9e205b904e2 get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
     8
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
     9
  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
    10
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
    11
793
09dff5ef8f74 as expected problems occure when lifting concat lemmas
Christian Urban <urbanc@in.tum.de>
parents: 787
diff changeset
    12
(* 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
    13
(* 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
    14
(* 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
    15
(* 'a_1                                               *)
1914
c50601bc617e Remove the part already in FSet and leave the experiments
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1873
diff changeset
    16
830
89d772dae4d4 New_relations, all works again including concat examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 827
diff changeset
    17
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
    18
  shows "concat [] \<approx> []"
1914
c50601bc617e Remove the part already in FSet and leave the experiments
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1873
diff changeset
    19
by (simp)
793
09dff5ef8f74 as expected problems occure when lifting concat lemmas
Christian Urban <urbanc@in.tum.de>
parents: 787
diff changeset
    20
830
89d772dae4d4 New_relations, all works again including concat examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 827
diff changeset
    21
lemma concat2:
89d772dae4d4 New_relations, all works again including concat examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 827
diff changeset
    22
  shows "concat (x # xs) \<approx> x @ concat xs"
1850
05b2dd2b0e8a Prove insert_rsp2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1260
diff changeset
    23
by (simp)
722
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
    24
1873
a08eaea622d1 minor FSet3 edits.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1850
diff changeset
    25
lemma concat_rsp:
a08eaea622d1 minor FSet3 edits.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1850
diff changeset
    26
  "\<lbrakk>list_rel op \<approx> x x'; x' \<approx> y'; list_rel op \<approx> y' y\<rbrakk> \<Longrightarrow> concat x \<approx> concat y"
a08eaea622d1 minor FSet3 edits.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1850
diff changeset
    27
  apply (induct x y arbitrary: x' y' rule: list_induct2')
a08eaea622d1 minor FSet3 edits.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1850
diff changeset
    28
  apply simp
a08eaea622d1 minor FSet3 edits.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1850
diff changeset
    29
  defer defer
a08eaea622d1 minor FSet3 edits.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1850
diff changeset
    30
  apply (simp only: concat.simps)
a08eaea622d1 minor FSet3 edits.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1850
diff changeset
    31
  sorry
a08eaea622d1 minor FSet3 edits.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1850
diff changeset
    32
a08eaea622d1 minor FSet3 edits.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1850
diff changeset
    33
lemma [quot_respect]:
851
e1473b4b2886 Change OO to OOO in FSet3.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 833
diff changeset
    34
  shows "(list_rel op \<approx> OOO op \<approx> ===> op \<approx>) concat concat"
1873
a08eaea622d1 minor FSet3 edits.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1850
diff changeset
    35
  apply (simp only: fun_rel_def)
a08eaea622d1 minor FSet3 edits.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1850
diff changeset
    36
  apply clarify
a08eaea622d1 minor FSet3 edits.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1850
diff changeset
    37
  apply (rule concat_rsp)
a08eaea622d1 minor FSet3 edits.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1850
diff changeset
    38
  apply assumption+
a08eaea622d1 minor FSet3 edits.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1850
diff changeset
    39
  done
814
cd3fa86be45f Sledgehammer bug.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 798
diff changeset
    40
851
e1473b4b2886 Change OO to OOO in FSet3.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 833
diff changeset
    41
lemma nil_rsp2[quot_respect]: "(list_rel op \<approx> OOO op \<approx>) [] []"
1914
c50601bc617e Remove the part already in FSet and leave the experiments
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1873
diff changeset
    42
  by (metis nil_rsp list_rel.simps(1) pred_compI)
814
cd3fa86be45f Sledgehammer bug.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 798
diff changeset
    43
824
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
    44
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
    45
  apply (rule eq_reflection)
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
    46
  apply auto
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
    47
  done
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
    48
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
    49
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
    50
  unfolding list_eq.simps
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
    51
  apply(simp only: set_map set_in_eq)
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
    52
  done
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
    53
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
    54
lemma quotient_compose_list_pre:
851
e1473b4b2886 Change OO to OOO in FSet3.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 833
diff changeset
    55
  "(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
    56
  ((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
    57
  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
    58
  apply rule
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
    59
  apply rule
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
    60
  apply rule
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
    61
  apply (rule list_rel_refl)
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
    62
  apply (metis equivp_def fset_equivp)
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
    63
  apply rule
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
    64
  apply (rule equivp_reflp[OF fset_equivp])
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
    65
  apply (rule list_rel_refl)
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
    66
  apply (metis equivp_def fset_equivp)
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
    67
  apply(rule)
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
    68
  apply rule
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
    69
  apply (rule list_rel_refl)
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
    70
  apply (metis equivp_def fset_equivp)
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
    71
  apply rule
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
    72
  apply (rule equivp_reflp[OF fset_equivp])
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
    73
  apply (rule list_rel_refl)
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
    74
  apply (metis equivp_def fset_equivp)
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
    75
  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
    76
  apply (metis Quotient_rel[OF Quotient_fset])
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
    77
  apply (auto simp only:)[1]
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
    78
  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
    79
  prefer 2
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
    80
  apply (metis Quotient_rel[OF list_quotient[OF Quotient_fset]])
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
    81
  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
    82
  prefer 2
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
    83
  apply (metis Quotient_rel[OF list_quotient[OF Quotient_fset]])
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
    84
  apply (simp only: map_rel_cong)
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
    85
  apply rule
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
    86
  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
    87
  apply (tactic {* Quotient_Tacs.quotient_tac @{context} 1 *})
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
    88
  apply (rule list_rel_refl)
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
    89
  apply (metis equivp_def fset_equivp)
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
    90
  apply rule
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
    91
  prefer 2
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
    92
  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
    93
  apply (tactic {* Quotient_Tacs.quotient_tac @{context} 1 *})
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
    94
  apply (rule list_rel_refl)
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
    95
  apply (metis equivp_def fset_equivp)
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
    96
  apply (erule conjE)+
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
    97
  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
    98
  prefer 2
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
    99
  apply (metis Quotient_def Quotient_fset equivp_reflp fset_equivp)
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
   100
  apply (rule map_rel_cong)
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
   101
  apply (assumption)
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
   102
  done
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
   103
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
   104
lemma quotient_compose_list[quot_thm]:
851
e1473b4b2886 Change OO to OOO in FSet3.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 833
diff changeset
   105
  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
   106
    (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
   107
  unfolding Quotient_def comp_def
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
   108
  apply (rule)+
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
   109
  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
   110
  apply (rule)
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
   111
  apply (rule)
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
   112
  apply (rule)
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
   113
  apply (rule list_rel_refl)
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
   114
  apply (metis equivp_def fset_equivp)
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
   115
  apply (rule)
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
   116
  apply (rule equivp_reflp[OF fset_equivp])
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
   117
  apply (rule list_rel_refl)
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
   118
  apply (metis equivp_def fset_equivp)
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
   119
  apply rule
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
   120
  apply rule
1850
05b2dd2b0e8a Prove insert_rsp2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1260
diff changeset
   121
  apply (rule quotient_compose_list_pre)
824
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
   122
  done
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
   123
793
09dff5ef8f74 as expected problems occure when lifting concat lemmas
Christian Urban <urbanc@in.tum.de>
parents: 787
diff changeset
   124
lemma fconcat_empty:
09dff5ef8f74 as expected problems occure when lifting concat lemmas
Christian Urban <urbanc@in.tum.de>
parents: 787
diff changeset
   125
  shows "fconcat {||} = {||}"
1850
05b2dd2b0e8a Prove insert_rsp2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1260
diff changeset
   126
  apply(lifting concat1)
05b2dd2b0e8a Prove insert_rsp2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1260
diff changeset
   127
  apply(cleaning)
1914
c50601bc617e Remove the part already in FSet and leave the experiments
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1873
diff changeset
   128
  apply(simp add: comp_def bot_fset_def)
1850
05b2dd2b0e8a Prove insert_rsp2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1260
diff changeset
   129
  done
793
09dff5ef8f74 as expected problems occure when lifting concat lemmas
Christian Urban <urbanc@in.tum.de>
parents: 787
diff changeset
   130
830
89d772dae4d4 New_relations, all works again including concat examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 827
diff changeset
   131
lemma insert_rsp2[quot_respect]:
851
e1473b4b2886 Change OO to OOO in FSet3.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 833
diff changeset
   132
  "(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
   133
  apply auto
05b2dd2b0e8a Prove insert_rsp2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1260
diff changeset
   134
  apply (simp add: set_in_eq)
05b2dd2b0e8a Prove insert_rsp2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1260
diff changeset
   135
  apply (rule_tac b="x # b" in pred_compI)
05b2dd2b0e8a Prove insert_rsp2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1260
diff changeset
   136
  apply auto
05b2dd2b0e8a Prove insert_rsp2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1260
diff changeset
   137
  apply (rule_tac b="x # ba" in pred_compI)
1914
c50601bc617e Remove the part already in FSet and leave the experiments
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1873
diff changeset
   138
  apply auto
1850
05b2dd2b0e8a Prove insert_rsp2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1260
diff changeset
   139
  done
830
89d772dae4d4 New_relations, all works again including concat examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 827
diff changeset
   140
89d772dae4d4 New_relations, all works again including concat examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 827
diff changeset
   141
lemma append_rsp[quot_respect]:
89d772dae4d4 New_relations, all works again including concat examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 827
diff changeset
   142
  "(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
   143
  by (auto)
89d772dae4d4 New_relations, all works again including concat examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 827
diff changeset
   144
793
09dff5ef8f74 as expected problems occure when lifting concat lemmas
Christian Urban <urbanc@in.tum.de>
parents: 787
diff changeset
   145
lemma fconcat_insert:
09dff5ef8f74 as expected problems occure when lifting concat lemmas
Christian Urban <urbanc@in.tum.de>
parents: 787
diff changeset
   146
  shows "fconcat (finsert x S) = x |\<union>| fconcat S"
1850
05b2dd2b0e8a Prove insert_rsp2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1260
diff changeset
   147
  apply(lifting concat2)
05b2dd2b0e8a Prove insert_rsp2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1260
diff changeset
   148
  apply(cleaning)
05b2dd2b0e8a Prove insert_rsp2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1260
diff changeset
   149
  apply (simp add: finsert_def fconcat_def comp_def)
05b2dd2b0e8a Prove insert_rsp2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1260
diff changeset
   150
  apply cleaning
05b2dd2b0e8a Prove insert_rsp2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1260
diff changeset
   151
  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
   152
e9e205b904e2 get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas 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
(* 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
   154
e9e205b904e2 get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas 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
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
   156
e9e205b904e2 get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas 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
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
   158
e9e205b904e2 get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas 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
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
   160
"_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
   161
"_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
   162
"_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
   163
"_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
   164
"_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
   165
"_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
   166
e9e205b904e2 get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas 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
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
   168
"_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
   169
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
   170
"_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
   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
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
   173
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
   174
  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
   175
  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
   176
  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
   177
  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
   178
  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
   179
  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
   180
e9e205b904e2 get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas 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
  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
   182
    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
   183
      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
   184
      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
   185
      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
   186
      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
   187
                                        $ 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
   188
      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
   189
      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
   190
                 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
   191
    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
   192
e9e205b904e2 get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas 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
  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
   194
        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
   195
            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
   196
        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
   197
           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
   198
           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
   199
        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
   200
    | 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
   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
  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
   203
        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
   204
          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
   205
                      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
   206
                    | 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
   207
        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
   208
          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
   209
        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
   210
e9e205b904e2 get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas 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
    | 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
   212
         (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
   213
            (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
   214
          | (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
   215
       
e9e205b904e2 get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas 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
    | 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
   217
        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
   218
          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
   219
        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
   220
          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
   221
        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
   222
e9e205b904e2 get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas 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
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
   224
*}
e9e205b904e2 get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas 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
1260
9df6144e281b moved Quot package to Attic (still compiles there with "isabelle make")
Christian Urban <urbanc@in.tum.de>
parents: 1144
diff changeset
   226
9df6144e281b moved Quot package to Attic (still compiles there with "isabelle make")
Christian Urban <urbanc@in.tum.de>
parents: 1144
diff changeset
   227
(* 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
   228
(* examles *)
1260
9df6144e281b moved Quot package to Attic (still compiles there with "isabelle make")
Christian Urban <urbanc@in.tum.de>
parents: 1144
diff changeset
   229
(*
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
   230
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
   231
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
   232
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
   233
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
   234
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
   235
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
   236
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
   237
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
   238
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
   239
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
   240
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
   241
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
   242
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
   243
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
   244
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
   245
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
   246
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
   247
*)
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
   248
e9e205b904e2 get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   249
(* 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
   250
714
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   251
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   252
end