Quot/Examples/FSet3.thy
author Christian Urban <urbanc@in.tum.de>
Sun, 20 Dec 2009 00:53:35 +0100
changeset 767 37285ec4387d
parent 766 df053507edba
child 768 e9e205b904e2
permissions -rw-r--r--
on suggestion of Tobias renamed "quotient_def" to "quotient_definition"; needs new keyword file
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
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     2
imports "../QuotMain" List
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
722
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
     5
fun
714
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     6
  list_eq :: "'a list \<Rightarrow> 'a list \<Rightarrow> bool" (infix "\<approx>" 50)
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     7
where
722
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
     8
  "list_eq xs ys = (\<forall>e. (e \<in> set xs) = (e \<in> set ys))"
714
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     9
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    10
lemma list_eq_equivp:
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    11
  shows "equivp list_eq"
722
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
    12
unfolding equivp_reflp_symp_transp reflp_def symp_def transp_def
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
    13
by auto
714
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    14
766
df053507edba renamed "quotient" command to "quotient_type"; needs new keyword file to be installed
Christian Urban <urbanc@in.tum.de>
parents: 743
diff changeset
    15
quotient_type fset = "'a list" / "list_eq"
714
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    16
  by (rule list_eq_equivp)
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    17
722
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
    18
lemma not_nil_equiv_cons: 
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
    19
  "\<not>[] \<approx> a # A" 
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
    20
by auto
714
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    21
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    22
lemma nil_rsp[quot_respect]:
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    23
  shows "[] \<approx> []"
722
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
    24
  by simp
714
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    25
722
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
    26
lemma cons_rsp[quot_respect]: 
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
    27
  shows "(op = ===> op \<approx> ===> op \<approx>) op # op #"
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
    28
  by simp
714
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    29
722
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
    30
(*
714
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    31
lemma mem_rsp[quot_respect]:
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    32
  "(op = ===> op \<approx> ===> op =) (op mem) (op mem)"
722
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
    33
*)
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
    34
714
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    35
722
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
    36
lemma no_mem_nil: 
729
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 728
diff changeset
    37
  "(\<forall>a. a \<notin> set A) = (A = [])"
722
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
    38
by (induct A) (auto)
714
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    39
722
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
    40
lemma none_mem_nil: 
729
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 728
diff changeset
    41
  "(\<forall>a. a \<notin> set A) = (A \<approx> [])"
722
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
    42
by simp
714
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    43
722
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
    44
lemma mem_cons: 
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
    45
  "a \<in> set A \<Longrightarrow> a # A \<approx> A"
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
    46
by auto
714
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    47
722
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
    48
lemma cons_left_comm: 
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
    49
  "x #y # A \<approx> y # x # A"
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
    50
by auto
714
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    51
722
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
    52
lemma cons_left_idem: 
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
    53
  "x # x # A \<approx> x # A"
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
    54
by auto
714
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    55
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    56
lemma finite_set_raw_strong_cases:
727
2cfe6f3d6352 Proof of finite_set_storng_cases_raw.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 726
diff changeset
    57
  "(X = []) \<or> (\<exists>a Y. ((a \<notin> set Y) \<and> (X \<approx> a # Y)))"
714
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    58
  apply (induct X)
727
2cfe6f3d6352 Proof of finite_set_storng_cases_raw.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 726
diff changeset
    59
  apply (simp)
2cfe6f3d6352 Proof of finite_set_storng_cases_raw.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 726
diff changeset
    60
  apply (rule disjI2)
2cfe6f3d6352 Proof of finite_set_storng_cases_raw.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 726
diff changeset
    61
  apply (erule disjE)
2cfe6f3d6352 Proof of finite_set_storng_cases_raw.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 726
diff changeset
    62
  apply (rule_tac x="a" in exI)
2cfe6f3d6352 Proof of finite_set_storng_cases_raw.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 726
diff changeset
    63
  apply (rule_tac x="[]" in exI)
2cfe6f3d6352 Proof of finite_set_storng_cases_raw.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 726
diff changeset
    64
  apply (simp)
2cfe6f3d6352 Proof of finite_set_storng_cases_raw.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 726
diff changeset
    65
  apply (erule exE)+
2cfe6f3d6352 Proof of finite_set_storng_cases_raw.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 726
diff changeset
    66
  apply (case_tac "a = aa")
2cfe6f3d6352 Proof of finite_set_storng_cases_raw.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 726
diff changeset
    67
  apply (rule_tac x="a" in exI)
2cfe6f3d6352 Proof of finite_set_storng_cases_raw.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 726
diff changeset
    68
  apply (rule_tac x="Y" in exI)
2cfe6f3d6352 Proof of finite_set_storng_cases_raw.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 726
diff changeset
    69
  apply (simp)
2cfe6f3d6352 Proof of finite_set_storng_cases_raw.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 726
diff changeset
    70
  apply (rule_tac x="aa" in exI)
2cfe6f3d6352 Proof of finite_set_storng_cases_raw.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 726
diff changeset
    71
  apply (rule_tac x="a # Y" in exI)
722
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
    72
  apply (auto)
727
2cfe6f3d6352 Proof of finite_set_storng_cases_raw.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 726
diff changeset
    73
  done
714
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    74
722
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
    75
fun
714
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    76
  delete_raw :: "'a list \<Rightarrow> 'a \<Rightarrow> 'a list"
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    77
where
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    78
  "delete_raw [] x = []"
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    79
| "delete_raw (a # A) x = (if (a = x) then delete_raw A x else a # (delete_raw A x))"
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    80
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    81
lemma mem_delete_raw:
728
0015159fee96 Some proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 727
diff changeset
    82
  "x \<in> set (delete_raw A a) = (x \<in> set A \<and> \<not>(x = a))"
726
1a777307f57f A bracket was missing; with it proved the 'definitely false' lemma.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 724
diff changeset
    83
  by (induct A arbitrary: x a) (auto)
714
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    84
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    85
lemma mem_delete_raw_ident:
722
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
    86
  "\<not>(a \<in> set (delete_raw A a))"
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
    87
by (induct A) (auto)
714
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    88
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    89
lemma not_mem_delete_raw_ident:
722
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
    90
  "b \<notin> set A \<Longrightarrow> (delete_raw A b = A)"
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
    91
by (induct A) (auto)
714
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    92
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    93
lemma delete_raw_RSP:
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    94
  "A \<approx> B \<Longrightarrow> delete_raw A a \<approx> delete_raw B a"
722
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
    95
apply(induct A arbitrary: B a)
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
    96
apply(auto)
714
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    97
sorry
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    98
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    99
lemma cons_delete_raw:
728
0015159fee96 Some proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 727
diff changeset
   100
  "a # (delete_raw A a) \<approx> (if a \<in> set A then A else (a # A))"
714
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   101
sorry
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   102
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   103
lemma mem_cons_delete_raw:
726
1a777307f57f A bracket was missing; with it proved the 'definitely false' lemma.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 724
diff changeset
   104
    "a \<in> set A \<Longrightarrow> a # (delete_raw A a) \<approx> A"
714
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   105
sorry
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   106
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   107
lemma finite_set_raw_delete_raw_cases:
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   108
    "X = [] \<or> (\<exists>a. a mem X \<and> X \<approx> a # delete_raw X a)"
726
1a777307f57f A bracket was missing; with it proved the 'definitely false' lemma.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 724
diff changeset
   109
  by (induct X) (auto)
714
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   110
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   111
fun
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   112
  card_raw :: "'a list \<Rightarrow> nat"
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   113
where
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   114
  card_raw_nil: "card_raw [] = 0"
728
0015159fee96 Some proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 727
diff changeset
   115
| card_raw_cons: "card_raw (x # xs) = (if x \<in> set xs then card_raw xs else Suc (card_raw xs))"
714
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   116
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   117
lemma not_mem_card_raw:
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   118
  fixes x :: "'a"
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   119
  fixes xs :: "'a list"
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   120
  shows "(\<not>(x mem xs)) = (card_raw (x # xs) = Suc (card_raw xs))"
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   121
  sorry
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   122
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   123
lemma card_raw_suc:
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   124
  assumes c: "card_raw xs = Suc n"
728
0015159fee96 Some proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 727
diff changeset
   125
  shows "\<exists>a ys. (a \<notin> set ys) \<and> xs \<approx> (a # ys)"
0015159fee96 Some proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 727
diff changeset
   126
  using c apply(induct xs)
0015159fee96 Some proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 727
diff changeset
   127
  apply(simp)
0015159fee96 Some proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 727
diff changeset
   128
  sorry
722
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
   129
728
0015159fee96 Some proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 727
diff changeset
   130
lemma mem_card_raw_gt_0:
0015159fee96 Some proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 727
diff changeset
   131
  "a \<in> set A \<Longrightarrow> 0 < card_raw A"
0015159fee96 Some proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 727
diff changeset
   132
  by (induct A) (auto)
714
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   133
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   134
lemma card_raw_cons_gt_0:
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   135
  "0 < card_raw (a # A)"
728
0015159fee96 Some proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 727
diff changeset
   136
  by (induct A) (auto)
714
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   137
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   138
lemma card_raw_delete_raw:
728
0015159fee96 Some proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 727
diff changeset
   139
  "card_raw (delete_raw A a) = (if a \<in> set A then card_raw A - 1 else card_raw A)"
714
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   140
sorry
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   141
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   142
lemma card_raw_rsp_aux:
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   143
  assumes e: "a \<approx> b"
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   144
  shows "card_raw a = card_raw b"
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   145
  using e sorry
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   146
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   147
lemma card_raw_rsp[quot_respect]:
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   148
  "(op \<approx> ===> op =) card_raw card_raw"
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   149
  by (simp add: card_raw_rsp_aux)
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   150
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   151
lemma card_raw_0:
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   152
  "(card_raw A = 0) = (A = [])"
728
0015159fee96 Some proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 727
diff changeset
   153
  by (induct A) (auto)
714
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   154
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   155
lemma list2set_thm:
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   156
  shows "set [] = {}"
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   157
  and "set (h # t) = insert h (set t)"
728
0015159fee96 Some proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 727
diff changeset
   158
  by (auto)
714
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   159
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   160
lemma list2set_RSP:
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   161
  "A \<approx> B \<Longrightarrow> set A = set B"
728
0015159fee96 Some proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 727
diff changeset
   162
  by auto
714
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   163
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   164
definition
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   165
  rsp_fold
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   166
where
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   167
  "rsp_fold f = (\<forall>u v w. (f u (f v w) = f v (f u w)))"
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   168
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   169
primrec
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   170
  fold_raw :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'b \<Rightarrow> 'a list \<Rightarrow> 'b"
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   171
where
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   172
  "fold_raw f z [] = z"
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   173
| "fold_raw f z (a # A) =
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   174
     (if (rsp_fold f) then
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   175
       if a mem A then fold_raw f z A
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   176
       else f a (fold_raw f z A)
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   177
     else z)"
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   178
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   179
lemma mem_lcommuting_fold_raw:
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   180
  "rsp_fold f \<Longrightarrow> h mem B \<Longrightarrow> fold_raw f z B = f h (fold_raw f z (delete_raw B h))"
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   181
sorry
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   182
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   183
lemma fold_rsp[quot_respect]:
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   184
  "(op = ===> op = ===> op \<approx> ===> op =) fold_raw fold_raw"
722
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
   185
apply(auto)
714
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   186
sorry
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   187
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   188
lemma append_rsp[quot_respect]:
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   189
  "(op \<approx> ===> op \<approx> ===> op \<approx>) op @ op @"
722
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
   190
by auto
714
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   191
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   192
primrec
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   193
  inter_raw
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   194
where
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   195
  "inter_raw [] B = []"
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   196
| "inter_raw (a # A) B = (if a mem B then a # inter_raw A B else inter_raw A B)"
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   197
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   198
lemma mem_inter_raw:
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   199
  "x mem (inter_raw A B) = x mem A \<and> x mem B"
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   200
sorry
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   201
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   202
lemma inter_raw_RSP:
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   203
  "A1 \<approx> A2 \<and> B1 \<approx> B2 \<Longrightarrow> (inter_raw A1 B1) \<approx> (inter_raw A2 B2)"
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   204
sorry
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   205
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   206
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   207
(* LIFTING DEFS *)
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   208
722
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
   209
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
   210
section {* Constants on the Quotient Type *} 
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
   211
767
37285ec4387d on suggestion of Tobias renamed "quotient_def" to "quotient_definition"; needs new keyword file
Christian Urban <urbanc@in.tum.de>
parents: 766
diff changeset
   212
quotient_definition
722
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
   213
  "fempty :: 'a fset" 
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
   214
  as "[]::'a list"
714
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   215
767
37285ec4387d on suggestion of Tobias renamed "quotient_def" to "quotient_definition"; needs new keyword file
Christian Urban <urbanc@in.tum.de>
parents: 766
diff changeset
   216
quotient_definition
722
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
   217
  "finsert :: 'a \<Rightarrow> 'a fset \<Rightarrow> 'a fset" 
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
   218
  as "op #"
714
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   219
767
37285ec4387d on suggestion of Tobias renamed "quotient_def" to "quotient_definition"; needs new keyword file
Christian Urban <urbanc@in.tum.de>
parents: 766
diff changeset
   220
quotient_definition
722
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
   221
  "fin :: 'a \<Rightarrow> 'a fset \<Rightarrow> bool" ("_ \<in>f _" [50, 51] 50)
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
   222
  as "\<lambda>x X. x \<in> set X"
714
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   223
722
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
   224
abbreviation
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
   225
  fnotin :: "'a \<Rightarrow> 'a fset \<Rightarrow> bool" ("_ \<notin>f _" [50, 51] 50)
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
   226
where
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
   227
  "a \<notin>f S \<equiv> \<not>(a \<in>f S)"
714
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   228
767
37285ec4387d on suggestion of Tobias renamed "quotient_def" to "quotient_definition"; needs new keyword file
Christian Urban <urbanc@in.tum.de>
parents: 766
diff changeset
   229
quotient_definition
722
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
   230
  "fcard :: 'a fset \<Rightarrow> nat" 
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
   231
  as "card_raw"
714
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   232
767
37285ec4387d on suggestion of Tobias renamed "quotient_def" to "quotient_definition"; needs new keyword file
Christian Urban <urbanc@in.tum.de>
parents: 766
diff changeset
   233
quotient_definition
722
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
   234
  "fdelete :: 'a fset \<Rightarrow> 'a \<Rightarrow> 'a fset" 
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
   235
  as "delete_raw"
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
   236
767
37285ec4387d on suggestion of Tobias renamed "quotient_def" to "quotient_definition"; needs new keyword file
Christian Urban <urbanc@in.tum.de>
parents: 766
diff changeset
   237
quotient_definition
729
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 728
diff changeset
   238
  "funion :: 'a fset \<Rightarrow> 'a fset \<Rightarrow> 'a fset" ("_ \<union>f _" [50, 51] 50)
722
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
   239
  as "op @"
714
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   240
767
37285ec4387d on suggestion of Tobias renamed "quotient_def" to "quotient_definition"; needs new keyword file
Christian Urban <urbanc@in.tum.de>
parents: 766
diff changeset
   241
quotient_definition
722
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
   242
  "finter :: 'a fset \<Rightarrow> 'a fset \<Rightarrow> 'a fset" ("_ \<inter>f _" [70, 71] 70)
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
   243
  as "inter_raw"
714
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   244
767
37285ec4387d on suggestion of Tobias renamed "quotient_def" to "quotient_definition"; needs new keyword file
Christian Urban <urbanc@in.tum.de>
parents: 766
diff changeset
   245
quotient_definition
722
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
   246
  "ffold :: ('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'b \<Rightarrow> 'a fset \<Rightarrow> 'b" 
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
   247
  as "fold_raw"
714
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   248
767
37285ec4387d on suggestion of Tobias renamed "quotient_def" to "quotient_definition"; needs new keyword file
Christian Urban <urbanc@in.tum.de>
parents: 766
diff changeset
   249
quotient_definition
722
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
   250
  "fset_to_set :: 'a fset \<Rightarrow> 'a set" 
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
   251
  as "set"
714
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   252
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   253
722
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
   254
section {* Lifted Theorems *}
714
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   255
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   256
thm list.cases (* ??? *)
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   257
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   258
thm cons_left_comm
722
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
   259
lemma "finsert a (finsert b S) = finsert b (finsert a S)"
714
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   260
by (lifting cons_left_comm)
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   261
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   262
thm cons_left_idem
722
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
   263
lemma "finsert a (finsert a S) = finsert a S"
714
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   264
by (lifting cons_left_idem)
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   265
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   266
(* thm MEM:
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   267
  MEM x [] = F
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   268
  MEM x (h::t) = (x=h) \/ MEM x t *)
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   269
thm none_mem_nil
729
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 728
diff changeset
   270
(*lemma "(\<forall>a. a \<notin>f A) = (A = fempty)"*)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 728
diff changeset
   271
714
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   272
thm mem_cons
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   273
thm finite_set_raw_strong_cases
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   274
thm card_raw.simps
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   275
thm not_mem_card_raw
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   276
thm card_raw_suc
722
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
   277
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
   278
lemma "fcard X = Suc n \<Longrightarrow> (\<exists>a S. a \<notin>f S & X = finsert a S)"
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
   279
(*by (lifting card_raw_suc)*)
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
   280
sorry
714
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   281
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   282
thm card_raw_cons_gt_0
743
4b3822d1ed24 Replies to questions from the weekend: Uncommenting the renamed theorem commented out in 734.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 734
diff changeset
   283
thm mem_card_raw_gt_0
714
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   284
thm not_nil_equiv_cons
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   285
thm delete_raw.simps
722
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
   286
(*thm mem_delete_raw*)
714
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   287
thm card_raw_delete_raw
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   288
thm cons_delete_raw
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   289
thm mem_cons_delete_raw
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   290
thm finite_set_raw_delete_raw_cases
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   291
thm append.simps
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   292
(* MEM_APPEND: MEM e (APPEND l1 l2) = MEM e l1 \/ MEM e l2 *)
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   293
thm inter_raw.simps
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   294
thm mem_inter_raw
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   295
thm fold_raw.simps
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   296
thm list2set_thm
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   297
thm list_eq_def
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   298
thm list.induct
722
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
   299
lemma "\<lbrakk>P fempty; \<And>a x. P x \<Longrightarrow> P (finsert a x)\<rbrakk> \<Longrightarrow> P l"
714
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   300
by (lifting list.induct)
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   301
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   302
(* We also have map and some properties of it in FSet *)
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   303
(* and the following which still lifts ok *)
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   304
lemma "funion (funion x xa) xb = funion x (funion xa xb)"
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   305
by (lifting append_assoc)
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   306
767
37285ec4387d on suggestion of Tobias renamed "quotient_def" to "quotient_definition"; needs new keyword file
Christian Urban <urbanc@in.tum.de>
parents: 766
diff changeset
   307
quotient_definition
716
1e08743b6997 FSet3 minor fixes + cases
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 714
diff changeset
   308
  "fset_case :: 'a \<Rightarrow> ('b \<Rightarrow> 'b fset \<Rightarrow> 'a) \<Rightarrow> 'b fset \<Rightarrow> 'a"
1e08743b6997 FSet3 minor fixes + cases
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 714
diff changeset
   309
as
1e08743b6997 FSet3 minor fixes + cases
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 714
diff changeset
   310
  "list_case"
1e08743b6997 FSet3 minor fixes + cases
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 714
diff changeset
   311
1e08743b6997 FSet3 minor fixes + cases
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 714
diff changeset
   312
(* NOT SURE IF TRUE *)
1e08743b6997 FSet3 minor fixes + cases
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 714
diff changeset
   313
lemma list_case_rsp[quot_respect]:
1e08743b6997 FSet3 minor fixes + cases
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 714
diff changeset
   314
  "(op = ===> (op = ===> op \<approx> ===> op =) ===> op \<approx> ===> op =) list_case list_case"
1e08743b6997 FSet3 minor fixes + cases
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 714
diff changeset
   315
  apply (auto)
1e08743b6997 FSet3 minor fixes + cases
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 714
diff changeset
   316
  sorry
1e08743b6997 FSet3 minor fixes + cases
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 714
diff changeset
   317
722
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
   318
lemma "fset_case (f1::'t) f2 (finsert a xa) = f2 a xa"
716
1e08743b6997 FSet3 minor fixes + cases
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 714
diff changeset
   319
apply (lifting list.cases(2))
1e08743b6997 FSet3 minor fixes + cases
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 714
diff changeset
   320
done
1e08743b6997 FSet3 minor fixes + cases
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 714
diff changeset
   321
1e08743b6997 FSet3 minor fixes + cases
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 714
diff changeset
   322
714
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   323
end