Attic/Quot/Examples/FSet3.thy
author Cezary Kaliszyk <kaliszyk@in.tum.de>
Wed, 21 Apr 2010 20:01:18 +0200
changeset 1929 f4e241829b80
parent 1927 6c5e3ac737d9
child 1935 266abc3ee228
permissions -rw-r--r--
minor
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"
1927
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
    10
  by (lifting list.exhaust)
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
    11
1917
efbc22a6f1a4 Working lifting of concat with inline proofs of second level preservation.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1916
diff changeset
    12
lemma list_rel_find_element:
efbc22a6f1a4 Working lifting of concat with inline proofs of second level preservation.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1916
diff changeset
    13
  assumes a: "x \<in> set a"
efbc22a6f1a4 Working lifting of concat with inline proofs of second level preservation.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1916
diff changeset
    14
  and b: "list_rel R a b"
efbc22a6f1a4 Working lifting of concat with inline proofs of second level preservation.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1916
diff changeset
    15
  shows "\<exists>y. (y \<in> set b \<and> R x y)"
efbc22a6f1a4 Working lifting of concat with inline proofs of second level preservation.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1916
diff changeset
    16
proof -
efbc22a6f1a4 Working lifting of concat with inline proofs of second level preservation.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1916
diff changeset
    17
  have "length a = length b" using b by (rule list_rel_len)
efbc22a6f1a4 Working lifting of concat with inline proofs of second level preservation.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1916
diff changeset
    18
  then show ?thesis using a b proof (induct a b rule: list_induct2)
efbc22a6f1a4 Working lifting of concat with inline proofs of second level preservation.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1916
diff changeset
    19
    case Nil
efbc22a6f1a4 Working lifting of concat with inline proofs of second level preservation.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1916
diff changeset
    20
    show ?case using Nil.prems by simp
efbc22a6f1a4 Working lifting of concat with inline proofs of second level preservation.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1916
diff changeset
    21
  next
efbc22a6f1a4 Working lifting of concat with inline proofs of second level preservation.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1916
diff changeset
    22
    case (Cons x xs y ys)
efbc22a6f1a4 Working lifting of concat with inline proofs of second level preservation.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1916
diff changeset
    23
    show ?case using Cons by auto
efbc22a6f1a4 Working lifting of concat with inline proofs of second level preservation.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1916
diff changeset
    24
  qed
efbc22a6f1a4 Working lifting of concat with inline proofs of second level preservation.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1916
diff changeset
    25
qed
793
09dff5ef8f74 as expected problems occure when lifting concat lemmas
Christian Urban <urbanc@in.tum.de>
parents: 787
diff changeset
    26
1917
efbc22a6f1a4 Working lifting of concat with inline proofs of second level preservation.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1916
diff changeset
    27
lemma concat_rsp_pre:
1921
e41b7046be2c More about concat
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1917
diff changeset
    28
  assumes a: "list_rel op \<approx> x x'"
e41b7046be2c More about concat
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1917
diff changeset
    29
  and     b: "x' \<approx> y'"
e41b7046be2c More about concat
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1917
diff changeset
    30
  and     c: "list_rel op \<approx> y' y"
e41b7046be2c More about concat
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1917
diff changeset
    31
  and     d: "\<exists>x\<in>set x. xa \<in> set x"
e41b7046be2c More about concat
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1917
diff changeset
    32
  shows "\<exists>x\<in>set y. xa \<in> set x"
e41b7046be2c More about concat
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1917
diff changeset
    33
proof -
e41b7046be2c More about concat
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1917
diff changeset
    34
  obtain xb where e: "xb \<in> set x" and f: "xa \<in> set xb" using d by auto
e41b7046be2c More about concat
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1917
diff changeset
    35
  have "\<exists>y. y \<in> set x' \<and> xb \<approx> y" by (rule list_rel_find_element[OF e a])
e41b7046be2c More about concat
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1917
diff changeset
    36
  then obtain ya where h: "ya \<in> set x'" and i: "xb \<approx> ya" by auto
e41b7046be2c More about concat
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1917
diff changeset
    37
  have j: "ya \<in> set y'" using b h by simp
e41b7046be2c More about concat
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1917
diff changeset
    38
  have "\<exists>yb. yb \<in> set y \<and> ya \<approx> yb" by (rule list_rel_find_element[OF j c])
e41b7046be2c More about concat
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1917
diff changeset
    39
  then show ?thesis using f i by auto
e41b7046be2c More about concat
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1917
diff changeset
    40
qed
e41b7046be2c More about concat
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1917
diff changeset
    41
e41b7046be2c More about concat
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1917
diff changeset
    42
lemma fun_relI [intro]:
e41b7046be2c More about concat
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1917
diff changeset
    43
  assumes "\<And>a b. P a b \<Longrightarrow> Q (x a) (y b)"
e41b7046be2c More about concat
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1917
diff changeset
    44
  shows "(P ===> Q) x y"
e41b7046be2c More about concat
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1917
diff changeset
    45
  using assms by (simp add: fun_rel_def)
1873
a08eaea622d1 minor FSet3 edits.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1850
diff changeset
    46
a08eaea622d1 minor FSet3 edits.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1850
diff changeset
    47
lemma [quot_respect]:
851
e1473b4b2886 Change OO to OOO in FSet3.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 833
diff changeset
    48
  shows "(list_rel op \<approx> OOO op \<approx> ===> op \<approx>) concat concat"
1927
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
    49
proof (rule fun_relI, elim pred_compE)
1921
e41b7046be2c More about concat
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1917
diff changeset
    50
  fix a b ba bb
e41b7046be2c More about concat
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1917
diff changeset
    51
  assume a: "list_rel op \<approx> a ba"
e41b7046be2c More about concat
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1917
diff changeset
    52
  assume b: "ba \<approx> bb"
e41b7046be2c More about concat
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1917
diff changeset
    53
  assume c: "list_rel op \<approx> bb b"
e41b7046be2c More about concat
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1917
diff changeset
    54
  have "\<forall>x. (\<exists>xa\<in>set a. x \<in> set xa) = (\<exists>xa\<in>set b. x \<in> set xa)" proof
e41b7046be2c More about concat
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1917
diff changeset
    55
    fix x
e41b7046be2c More about concat
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1917
diff changeset
    56
    show "(\<exists>xa\<in>set a. x \<in> set xa) = (\<exists>xa\<in>set b. x \<in> set xa)" proof
e41b7046be2c More about concat
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1917
diff changeset
    57
      assume d: "\<exists>xa\<in>set a. x \<in> set xa"
e41b7046be2c More about concat
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1917
diff changeset
    58
      show "\<exists>xa\<in>set b. x \<in> set xa" by (rule concat_rsp_pre[OF a b c d])
e41b7046be2c More about concat
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1917
diff changeset
    59
    next
e41b7046be2c More about concat
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1917
diff changeset
    60
      assume e: "\<exists>xa\<in>set b. x \<in> set xa"
e41b7046be2c More about concat
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1917
diff changeset
    61
      have a': "list_rel op \<approx> ba a" by (rule list_rel_symp[OF list_eq_equivp, OF a])
e41b7046be2c More about concat
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1917
diff changeset
    62
      have b': "bb \<approx> ba" by (rule equivp_symp[OF list_eq_equivp, OF b])
e41b7046be2c More about concat
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1917
diff changeset
    63
      have c': "list_rel op \<approx> b bb" by (rule list_rel_symp[OF list_eq_equivp, OF c])
e41b7046be2c More about concat
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1917
diff changeset
    64
      show "\<exists>xa\<in>set a. x \<in> set xa" by (rule concat_rsp_pre[OF c' b' a' e])
e41b7046be2c More about concat
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1917
diff changeset
    65
    qed
e41b7046be2c More about concat
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1917
diff changeset
    66
  qed
e41b7046be2c More about concat
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1917
diff changeset
    67
  then show "concat a \<approx> concat b" by simp
e41b7046be2c More about concat
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1917
diff changeset
    68
qed
814
cd3fa86be45f Sledgehammer bug.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 798
diff changeset
    69
851
e1473b4b2886 Change OO to OOO in FSet3.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 833
diff changeset
    70
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
    71
  by (metis nil_rsp list_rel.simps(1) pred_compI)
814
cd3fa86be45f Sledgehammer bug.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 798
diff changeset
    72
824
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
    73
lemma set_in_eq: "(\<forall>e. ((e \<in> A) \<longleftrightarrow> (e \<in> B))) \<equiv> A = B"
1917
efbc22a6f1a4 Working lifting of concat with inline proofs of second level preservation.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1916
diff changeset
    74
  by (rule eq_reflection) auto
824
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
    75
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
    76
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
    77
  unfolding list_eq.simps
1917
efbc22a6f1a4 Working lifting of concat with inline proofs of second level preservation.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1916
diff changeset
    78
  by (simp only: set_map set_in_eq)
824
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
    79
1927
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
    80
lemma compose_list_refl:
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
    81
  shows "(list_rel op \<approx> OOO op \<approx>) r r"
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
    82
proof
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
    83
  show c: "list_rel op \<approx> r r" by (rule list_rel_refl) (metis equivp_def fset_equivp)
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
    84
  have d: "r \<approx> r" by (rule equivp_reflp[OF fset_equivp])
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
    85
  show b: "(op \<approx> OO list_rel op \<approx>) r r" by (rule pred_compI) (rule d, rule c)
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
    86
qed
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
    87
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
    88
lemma list_rel_refl:
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
    89
  shows "(list_rel op \<approx>) r r"
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
    90
  by (rule list_rel_refl)(metis equivp_def fset_equivp)
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
    91
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
    92
lemma Quotient_fset_list:
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
    93
  shows "Quotient (list_rel op \<approx>) (map abs_fset) (map rep_fset)"
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
    94
  by (fact list_quotient[OF Quotient_fset])
824
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
    95
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
    96
lemma quotient_compose_list[quot_thm]:
851
e1473b4b2886 Change OO to OOO in FSet3.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 833
diff changeset
    97
  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
    98
    (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
    99
  unfolding Quotient_def comp_def
1927
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   100
proof (intro conjI allI)
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   101
  fix a r s
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   102
  show "abs_fset (map abs_fset (map rep_fset (rep_fset a))) = a"
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   103
    by (simp add: abs_o_rep[OF Quotient_fset] Quotient_abs_rep[OF Quotient_fset] map_id)
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   104
  have b: "list_rel op \<approx> (map rep_fset (rep_fset a)) (map rep_fset (rep_fset a))"
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   105
    by (rule list_rel_refl)
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   106
  have c: "(op \<approx> OO list_rel op \<approx>) (map rep_fset (rep_fset a)) (map rep_fset (rep_fset a))"
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   107
    by (rule, rule equivp_reflp[OF fset_equivp]) (rule b)
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   108
  show "(list_rel op \<approx> OOO op \<approx>) (map rep_fset (rep_fset a)) (map rep_fset (rep_fset a))"
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   109
    by (rule, rule list_rel_refl) (rule c)
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   110
  show "(list_rel op \<approx> OOO op \<approx>) r s = ((list_rel op \<approx> OOO op \<approx>) r r \<and>
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   111
        (list_rel op \<approx> OOO op \<approx>) s s \<and> abs_fset (map abs_fset r) = abs_fset (map abs_fset s))"
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   112
  proof (intro iffI conjI)
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   113
    show "(list_rel op \<approx> OOO op \<approx>) r r" by (rule compose_list_refl)
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   114
    show "(list_rel op \<approx> OOO op \<approx>) s s" by (rule compose_list_refl)
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   115
  next
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   116
    assume a: "(list_rel op \<approx> OOO op \<approx>) r s"
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   117
    then have b: "map abs_fset r \<approx> map abs_fset s" proof (elim pred_compE)
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   118
      fix b ba
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   119
      assume c: "list_rel op \<approx> r b"
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   120
      assume d: "b \<approx> ba"
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   121
      assume e: "list_rel op \<approx> ba s"
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   122
      have f: "map abs_fset r = map abs_fset b"
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   123
        by (metis Quotient_rel[OF Quotient_fset_list] c)
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   124
      have g: "map abs_fset s = map abs_fset ba"
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   125
        by (metis Quotient_rel[OF Quotient_fset_list] e)
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   126
      show "map abs_fset r \<approx> map abs_fset s" using d f g map_rel_cong by simp
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   127
    qed
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   128
    then show "abs_fset (map abs_fset r) = abs_fset (map abs_fset s)"
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   129
      by (metis Quotient_rel[OF Quotient_fset])
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   130
  next
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   131
    assume a: "(list_rel op \<approx> OOO op \<approx>) r r \<and> (list_rel op \<approx> OOO op \<approx>) s s
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   132
      \<and> abs_fset (map abs_fset r) = abs_fset (map abs_fset s)"
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   133
    then have s: "(list_rel op \<approx> OOO op \<approx>) s s" by simp
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   134
    have d: "map abs_fset r \<approx> map abs_fset s"
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   135
      by (subst Quotient_rel[OF Quotient_fset]) (simp add: a)
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   136
    have b: "map rep_fset (map abs_fset r) \<approx> map rep_fset (map abs_fset s)"
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   137
      by (rule map_rel_cong[OF d])
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   138
    have y: "list_rel op \<approx> (map rep_fset (map abs_fset s)) s"
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   139
      by (fact rep_abs_rsp_left[OF Quotient_fset_list, OF list_rel_refl[of s]])
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   140
    have c: "(op \<approx> OO list_rel op \<approx>) (map rep_fset (map abs_fset r)) s"
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   141
      by (rule pred_compI) (rule b, rule y)
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   142
    have z: "list_rel op \<approx> r (map rep_fset (map abs_fset r))"
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   143
      by (fact rep_abs_rsp[OF Quotient_fset_list, OF list_rel_refl[of r]])
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   144
    then show "(list_rel op \<approx> OOO op \<approx>) r s"
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   145
      using a c pred_compI by simp
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   146
  qed
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   147
qed
824
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 815
diff changeset
   148
1921
e41b7046be2c More about concat
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1917
diff changeset
   149
lemma nil_prs2[quot_preserve]: "(abs_fset \<circ> map f) [] = abs_fset []"
e41b7046be2c More about concat
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1917
diff changeset
   150
  by simp
e41b7046be2c More about concat
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1917
diff changeset
   151
793
09dff5ef8f74 as expected problems occure when lifting concat lemmas
Christian Urban <urbanc@in.tum.de>
parents: 787
diff changeset
   152
lemma fconcat_empty:
09dff5ef8f74 as expected problems occure when lifting concat lemmas
Christian Urban <urbanc@in.tum.de>
parents: 787
diff changeset
   153
  shows "fconcat {||} = {||}"
1921
e41b7046be2c More about concat
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1917
diff changeset
   154
  by (lifting concat.simps(1))
793
09dff5ef8f74 as expected problems occure when lifting concat lemmas
Christian Urban <urbanc@in.tum.de>
parents: 787
diff changeset
   155
830
89d772dae4d4 New_relations, all works again including concat examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 827
diff changeset
   156
lemma insert_rsp2[quot_respect]:
851
e1473b4b2886 Change OO to OOO in FSet3.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 833
diff changeset
   157
  "(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
   158
  apply auto
05b2dd2b0e8a Prove insert_rsp2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1260
diff changeset
   159
  apply (simp add: set_in_eq)
05b2dd2b0e8a Prove insert_rsp2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1260
diff changeset
   160
  apply (rule_tac b="x # b" in pred_compI)
05b2dd2b0e8a Prove insert_rsp2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1260
diff changeset
   161
  apply auto
05b2dd2b0e8a Prove insert_rsp2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1260
diff changeset
   162
  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
   163
  apply auto
1850
05b2dd2b0e8a Prove insert_rsp2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1260
diff changeset
   164
  done
830
89d772dae4d4 New_relations, all works again including concat examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 827
diff changeset
   165
89d772dae4d4 New_relations, all works again including concat examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 827
diff changeset
   166
lemma append_rsp[quot_respect]:
89d772dae4d4 New_relations, all works again including concat examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 827
diff changeset
   167
  "(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
   168
  by (auto)
89d772dae4d4 New_relations, all works again including concat examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 827
diff changeset
   169
1921
e41b7046be2c More about concat
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1917
diff changeset
   170
lemma insert_prs2[quot_preserve]:
e41b7046be2c More about concat
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1917
diff changeset
   171
  "(rep_fset ---> (map rep_fset \<circ> rep_fset) ---> (abs_fset \<circ> map abs_fset)) op # = finsert"
e41b7046be2c More about concat
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1917
diff changeset
   172
  by (simp add: expand_fun_eq Quotient_abs_rep[OF Quotient_fset]
e41b7046be2c More about concat
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1917
diff changeset
   173
      abs_o_rep[OF Quotient_fset] map_id finsert_def)
e41b7046be2c More about concat
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1917
diff changeset
   174
793
09dff5ef8f74 as expected problems occure when lifting concat lemmas
Christian Urban <urbanc@in.tum.de>
parents: 787
diff changeset
   175
lemma fconcat_insert:
09dff5ef8f74 as expected problems occure when lifting concat lemmas
Christian Urban <urbanc@in.tum.de>
parents: 787
diff changeset
   176
  shows "fconcat (finsert x S) = x |\<union>| fconcat S"
1921
e41b7046be2c More about concat
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1917
diff changeset
   177
  by (lifting concat.simps(2))
e41b7046be2c More about concat
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1917
diff changeset
   178
e41b7046be2c More about concat
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1917
diff changeset
   179
lemma append_prs2[quot_preserve]:
e41b7046be2c More about concat
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1917
diff changeset
   180
  "((map rep_fset \<circ> rep_fset) ---> (map rep_fset \<circ> rep_fset) ---> abs_fset \<circ> map abs_fset) op @ = funion"
e41b7046be2c More about concat
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1917
diff changeset
   181
  by (simp add: expand_fun_eq Quotient_abs_rep[OF Quotient_fset]
e41b7046be2c More about concat
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1917
diff changeset
   182
      abs_o_rep[OF Quotient_fset] map_id sup_fset_def)
e41b7046be2c More about concat
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1917
diff changeset
   183
1927
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   184
lemma list_rel_app_l:
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   185
  assumes a: "reflp R"
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   186
  and b: "list_rel R l r"
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   187
  shows "list_rel R (z @ l) (z @ r)"
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   188
  by (induct z) (simp_all add: b, metis a reflp_def)
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   189
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   190
lemma append_rsp2_pre0:
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   191
  assumes a:"list_rel op \<approx> x x'"
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   192
  shows "list_rel op \<approx> (x @ z) (x' @ z)"
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   193
  using a apply (induct x x' rule: list_induct2')
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   194
  apply simp_all
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   195
  apply (rule list_rel_refl)
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   196
  done
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   197
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   198
lemma append_rsp2_pre1:
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   199
  assumes a:"list_rel op \<approx> x x'"
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   200
  shows "list_rel op \<approx> (z @ x) (z @ x')"
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   201
  using a apply (induct x x' arbitrary: z rule: list_induct2')
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   202
  apply (rule list_rel_refl)
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   203
  apply (simp_all del: list_eq.simps)
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   204
  apply (rule list_rel_app_l)
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   205
  apply (simp_all add: reflp_def)
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   206
  done
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   207
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   208
lemma append_rsp2_pre:
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   209
  assumes a:"list_rel op \<approx> x x'"
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   210
  and     b: "list_rel op \<approx> z z'"
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   211
  shows "list_rel op \<approx> (x @ z) (x' @ z')"
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   212
  apply (rule list_rel_transp[OF fset_equivp])
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   213
  apply (rule append_rsp2_pre0)
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   214
  apply (rule a)
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   215
  using b apply (induct z z' rule: list_induct2')
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   216
  apply (simp_all only: append_Nil2)
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   217
  apply (rule list_rel_refl)
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   218
  apply simp_all
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   219
  apply (rule append_rsp2_pre1)
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   220
  apply simp
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   221
  done
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   222
1921
e41b7046be2c More about concat
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1917
diff changeset
   223
lemma append_rsp2[quot_respect]:
e41b7046be2c More about concat
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1917
diff changeset
   224
  "(list_rel op \<approx> OOO op \<approx> ===> list_rel op \<approx> OOO op \<approx> ===> list_rel op \<approx> OOO op \<approx>) op @ op @"
1927
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   225
proof (intro fun_relI, elim pred_compE)
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   226
  fix x y z w x' z' y' w' :: "'a list list"
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   227
  assume a:"list_rel op \<approx> x x'"
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   228
  and b:    "x' \<approx> y'"
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   229
  and c:    "list_rel op \<approx> y' y"
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   230
  assume aa: "list_rel op \<approx> z z'"
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   231
  and bb:   "z' \<approx> w'"
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   232
  and cc:   "list_rel op \<approx> w' w"
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   233
  have a': "list_rel op \<approx> (x @ z) (x' @ z')" using a aa append_rsp2_pre by auto
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   234
  have b': "x' @ z' \<approx> y' @ w'" using b bb by simp
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   235
  have c': "list_rel op \<approx> (y' @ w') (y @ w)" using c cc append_rsp2_pre by auto
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   236
  have d': "(op \<approx> OO list_rel op \<approx>) (x' @ z') (y @ w)"
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   237
    by (rule pred_compI) (rule b', rule c')
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   238
  show "(list_rel op \<approx> OOO op \<approx>) (x @ z) (y @ w)"
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   239
    by (rule pred_compI) (rule a', rule d')
6c5e3ac737d9 append_rsp2 + isarification
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1921
diff changeset
   240
qed
1921
e41b7046be2c More about concat
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1917
diff changeset
   241
e41b7046be2c More about concat
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1917
diff changeset
   242
lemma "fconcat (xs |\<union>| ys) = fconcat xs |\<union>| fconcat ys"
e41b7046be2c More about concat
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1917
diff changeset
   243
  by (lifting concat_append)
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
   244
e9e205b904e2 get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas 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
(* 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
   246
e9e205b904e2 get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas 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
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
   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
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
   250
e9e205b904e2 get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   251
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
   252
"_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
   253
"_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
   254
"_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
   255
"_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
   256
"_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
   257
"_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
   258
e9e205b904e2 get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   259
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
   260
"_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
   261
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
   262
"_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
   263
e9e205b904e2 get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   264
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
   265
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
   266
  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
   267
  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
   268
  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
   269
  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
   270
  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
   271
  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
   272
e9e205b904e2 get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   273
  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
   274
    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
   275
      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
   276
      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
   277
      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
   278
      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
   279
                                        $ 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
   280
      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
   281
      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
   282
                 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
   283
    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
   284
e9e205b904e2 get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   285
  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
   286
        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
   287
            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
   288
        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
   289
           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
   290
           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
   291
        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
   292
    | 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
   293
e9e205b904e2 get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   294
  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
   295
        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
   296
          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
   297
                      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
   298
                    | 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
   299
        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
   300
          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
   301
        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
   302
e9e205b904e2 get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   303
    | 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
   304
         (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
   305
            (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
   306
          | (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
   307
       
e9e205b904e2 get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   308
    | 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
   309
        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
   310
          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
   311
        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
   312
          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
   313
        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
   314
e9e205b904e2 get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   315
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
   316
*}
e9e205b904e2 get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   317
1260
9df6144e281b moved Quot package to Attic (still compiles there with "isabelle make")
Christian Urban <urbanc@in.tum.de>
parents: 1144
diff changeset
   318
9df6144e281b moved Quot package to Attic (still compiles there with "isabelle make")
Christian Urban <urbanc@in.tum.de>
parents: 1144
diff changeset
   319
(* 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
   320
(* examles *)
1260
9df6144e281b moved Quot package to Attic (still compiles there with "isabelle make")
Christian Urban <urbanc@in.tum.de>
parents: 1144
diff changeset
   321
(*
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
   322
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
   323
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
   324
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
   325
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
   326
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
   327
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
   328
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
   329
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
   330
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
   331
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
   332
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
   333
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
   334
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
   335
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
   336
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
   337
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
   338
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
   339
*)
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
   340
e9e205b904e2 get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   341
(* 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
   342
714
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   343
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   344
end