Quot/Examples/FSet3.thy
author Cezary Kaliszyk <kaliszyk@in.tum.de>
Sat, 12 Dec 2009 04:25:47 +0100
changeset 726 1a777307f57f
parent 724 d705d7ae2410
child 727 2cfe6f3d6352
permissions -rw-r--r--
A bracket was missing; with it proved the 'definitely false' lemma.
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
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    15
quotient fset = "'a list" / "list_eq"
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: 
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
    37
  "(\<forall>a. \<not>(a \<in> set A)) = (A = [])"
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: 
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
    41
  "(\<forall>a. \<not>(a \<in> set A)) = (A \<approx> [])"
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:
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    57
  "(X = []) \<or> (\<exists>a. \<exists> Y. (~(a mem Y) \<and> (X \<approx> a # Y)))"
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)
722
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
    59
  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
    60
  sorry
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    61
722
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
    62
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
    63
  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
    64
where
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    65
  "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
    66
| "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
    67
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    68
lemma mem_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
    69
  "x mem (delete_raw A a) = (x mem A \<and> \<not>(x = a))"
1a777307f57f A bracket was missing; with it proved the 'definitely false' lemma.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 724
diff changeset
    70
  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
    71
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    72
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
    73
  "\<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
    74
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
    75
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    76
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
    77
  "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
    78
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
    79
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    80
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
    81
  "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
    82
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
    83
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
    84
sorry
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    85
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    86
lemma 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
    87
  "a # (delete_raw A a) \<approx> (if a mem A then A else (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
    88
sorry
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    89
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    90
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
    91
    "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
    92
sorry
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    93
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    94
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
    95
    "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
    96
  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
    97
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    98
fun
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    99
  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
   100
where
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   101
  card_raw_nil: "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
   102
| card_raw_cons: "card_raw (x # xs) = (if x mem xs then card_raw xs else 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
   103
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   104
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
   105
  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
   106
  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
   107
  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
   108
  sorry
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   109
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   110
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
   111
  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
   112
  fixes n :: "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
  assumes c: "card_raw xs = Suc n"
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   114
  shows "\<exists>a ys. \<not>(a mem ys) \<and> xs \<approx> (a # ys)"
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   115
  using c
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   116
apply(induct xs)
722
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
   117
(*apply(metis 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
   118
apply(metis mem_delete_raw)
722
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
   119
done*)
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
   120
sorry
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
   121
714
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 mem_card_raw_not_0:
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   124
  "a mem A \<Longrightarrow> \<not>(card_raw A = 0)"
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   125
sorry
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   126
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   127
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
   128
  "0 < card_raw (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
   129
sorry
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   130
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   131
lemma 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
   132
  "card_raw (delete_raw A a) = (if a mem A then card_raw A - 1 else card_raw A)"
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   133
sorry
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   134
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   135
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
   136
  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
   137
  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
   138
  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
   139
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   140
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
   141
  "(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
   142
  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
   143
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   144
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
   145
  "(card_raw A = 0) = (A = [])"
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   146
sorry
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   147
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   148
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
   149
  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
   150
  and "set (h # t) = insert h (set t)"
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   151
sorry
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   152
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   153
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
   154
  "A \<approx> B \<Longrightarrow> set A = set B"
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   155
sorry
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   156
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   157
definition
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   158
  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
   159
where
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   160
  "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
   161
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   162
primrec
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   163
  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
   164
where
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   165
  "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
   166
| "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
   167
     (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
   168
       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
   169
       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
   170
     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
   171
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   172
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
   173
  "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
   174
sorry
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   175
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   176
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
   177
  "(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
   178
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
   179
sorry
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   180
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   181
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
   182
  "(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
   183
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
   184
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   185
primrec
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   186
  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
   187
where
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   188
  "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
   189
| "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
   190
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   191
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
   192
  "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
   193
sorry
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   194
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   195
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
   196
  "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
   197
sorry
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   198
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   199
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   200
(* 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
   201
722
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
   202
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
   203
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
   204
714
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   205
quotient_def
722
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
   206
  "fempty :: 'a fset" 
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
   207
  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
   208
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   209
quotient_def
722
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
   210
  "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
   211
  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
   212
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   213
quotient_def
722
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
   214
  "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
   215
  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
   216
722
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
   217
abbreviation
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
   218
  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
   219
where
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
   220
  "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
   221
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   222
quotient_def
722
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
   223
  "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
   224
  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
   225
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   226
quotient_def
722
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
   227
  "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
   228
  as "delete_raw"
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
   229
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
   230
quotient_def
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
   231
  "funion :: 'a fset \<Rightarrow> 'a fset \<Rightarrow> 'a fset" ("_ \<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
   232
  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
   233
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   234
quotient_def
722
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
   235
  "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
   236
  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
   237
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   238
quotient_def
722
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
   239
  "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
   240
  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
   241
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   242
quotient_def
722
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
   243
  "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
   244
  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
   245
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   246
722
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
   247
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
   248
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   249
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
   250
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   251
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
   252
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
   253
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
   254
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   255
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
   256
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
   257
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
   258
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   259
(* 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
   260
  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
   261
  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
   262
thm none_mem_nil
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   263
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
   264
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
   265
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
   266
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
   267
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
   268
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
   269
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
   270
(*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
   271
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
   272
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 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
   274
thm mem_card_raw_not_0
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_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
   276
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
   277
(*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
   278
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
   279
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
   280
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
   281
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
   282
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
   283
(* 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
   284
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
   285
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
   286
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
   287
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
   288
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
   289
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
   290
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
   291
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
   292
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   293
(* 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
   294
(* 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
   295
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
   296
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
   297
716
1e08743b6997 FSet3 minor fixes + cases
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 714
diff changeset
   298
quotient_def
1e08743b6997 FSet3 minor fixes + cases
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 714
diff changeset
   299
  "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
   300
as
1e08743b6997 FSet3 minor fixes + cases
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 714
diff changeset
   301
  "list_case"
1e08743b6997 FSet3 minor fixes + cases
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 714
diff changeset
   302
1e08743b6997 FSet3 minor fixes + cases
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 714
diff changeset
   303
(* NOT SURE IF TRUE *)
1e08743b6997 FSet3 minor fixes + cases
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 714
diff changeset
   304
lemma list_case_rsp[quot_respect]:
1e08743b6997 FSet3 minor fixes + cases
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 714
diff changeset
   305
  "(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
   306
  apply (auto)
1e08743b6997 FSet3 minor fixes + cases
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 714
diff changeset
   307
  sorry
1e08743b6997 FSet3 minor fixes + cases
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 714
diff changeset
   308
722
d5fce1ead432 started to have a look at it; redefined the relation
Christian Urban <urbanc@in.tum.de>
parents: 716
diff changeset
   309
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
   310
apply (lifting list.cases(2))
1e08743b6997 FSet3 minor fixes + cases
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 714
diff changeset
   311
done
1e08743b6997 FSet3 minor fixes + cases
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 714
diff changeset
   312
1e08743b6997 FSet3 minor fixes + cases
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 714
diff changeset
   313
714
37f7dc85b61b Added FSet3 with a formalisation of finite sets based on Michael's one.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   314
end