Quot/Quotients.thy
changeset 698 ed44eaaea63e
parent 697 57944c1ef728
child 701 1f4dfcd9351b
equal deleted inserted replaced
697:57944c1ef728 698:ed44eaaea63e
     1 theory Quotients
       
     2 imports Main
       
     3 begin
       
     4 
       
     5 (* Other quotients that have not been proved yet *)
       
     6 
       
     7 fun
       
     8   option_rel
       
     9 where
       
    10   "option_rel R None None = True"
       
    11 | "option_rel R (Some x) None = False"
       
    12 | "option_rel R None (Some x) = False"
       
    13 | "option_rel R (Some x) (Some y) = R x y"
       
    14 
       
    15 fun
       
    16   option_map
       
    17 where
       
    18   "option_map f None = None"
       
    19 | "option_map f (Some x) = Some (f x)"
       
    20 
       
    21 fun
       
    22   prod_rel
       
    23 where
       
    24   "prod_rel R1 R2 (a1,a2) (b1,b2) = (R1 a1 b1 \<and> R2 a2 b2)"
       
    25 
       
    26 fun
       
    27   prod_map
       
    28 where
       
    29   "prod_map f1 f2 (a,b) = (f1 a, f2 b)"
       
    30 
       
    31 fun
       
    32   sum_rel
       
    33 where
       
    34   "sum_rel R1 R2 (Inl a1) (Inl b1) = R1 a1 b1"
       
    35 | "sum_rel R1 R2 (Inl a1) (Inr b2) = False"
       
    36 | "sum_rel R1 R2 (Inr a2) (Inl b1) = False"
       
    37 | "sum_rel R1 R2 (Inr a2) (Inr b2) = R2 a2 b2"
       
    38 
       
    39 fun
       
    40   sum_map
       
    41 where
       
    42   "sum_map f1 f2 (Inl a) = Inl (f1 a)"
       
    43 | "sum_map f1 f2 (Inr a) = Inr (f2 a)"
       
    44 
       
    45 fun
       
    46   noption_map::"('a \<Rightarrow> 'b) \<Rightarrow> ('a noption) \<Rightarrow> ('b noption)"
       
    47 where
       
    48   "noption_map f (nSome x) = nSome (f x)"
       
    49 | "noption_map f nNone = nNone"
       
    50 
       
    51 fun
       
    52   noption_rel
       
    53 where
       
    54   "noption_rel r (nSome x) (nSome y) = r x y"
       
    55 | "noption_rel r _ _ = False"
       
    56 
       
    57 declare [[map noption = (noption_map, noption_rel)]]
       
    58 
       
    59 lemma "noption_map id = id"
       
    60 sorry
       
    61 
       
    62 lemma noption_Quotient:
       
    63   assumes q: "Quotient R Abs Rep"
       
    64   shows "Quotient (noption_rel R) (noption_map Abs) (noption_map Rep)"
       
    65   apply (unfold Quotient_def)
       
    66   apply (auto)
       
    67   using q
       
    68   apply (unfold Quotient_def)
       
    69   apply (case_tac "a :: 'b noption")
       
    70   apply (simp)
       
    71   apply (simp)
       
    72   apply (case_tac "a :: 'b noption")
       
    73   apply (simp only: option_map.simps)
       
    74   apply (subst option_rel.simps)
       
    75   (* Simp starts hanging so don't know how to continue *)
       
    76   sorry