Quot/Examples/AbsRepTest.thy
author Christian Urban <urbanc@in.tum.de>
Tue, 05 Jan 2010 17:12:35 +0100
changeset 811 e038928daa67
parent 810 f73e2f5f17b2
child 813 77506496e6fd
permissions -rw-r--r--
merged

theory AbsRepTest
imports "../QuotMain" "../QuotList" "../QuotOption" "../QuotSum" "../QuotProd" List
begin

ML {* open Quotient_Term *}

ML {*
fun test_funs flag ctxt (rty, qty) =
  (absrep_fun_chk flag ctxt (rty, qty)
   |> Syntax.string_of_term ctxt
   |> writeln;
   equiv_relation_chk ctxt (rty, qty) 
   |> Syntax.string_of_term ctxt
   |> writeln;
   new_equiv_relation_chk ctxt (rty, qty) 
   |> Syntax.string_of_term ctxt
   |> writeln)
*}

definition
  erel1 (infixl "\<approx>1" 50)
where
  "erel1 \<equiv> \<lambda>xs ys. \<forall>e. e \<in> set xs \<longleftrightarrow> e \<in> set ys"

quotient_type 
  'a fset = "'a list" / erel1
  apply(rule equivpI)
  unfolding erel1_def reflp_def symp_def transp_def
  by auto

definition
  erel2 (infixl "\<approx>2" 50)
where
  "erel2 \<equiv> \<lambda>(xs::('a * 'a) list) ys. \<forall>e. e \<in> set xs \<longleftrightarrow> e \<in> set ys"

quotient_type 
  'a foo = "('a * 'a) list" / erel2
  apply(rule equivpI)
  unfolding erel2_def reflp_def symp_def transp_def
  by auto

definition
  erel3 (infixl "\<approx>3" 50)
where
  "erel3 \<equiv> \<lambda>(xs::('a * int) list) ys. \<forall>e. e \<in> set xs \<longleftrightarrow> e \<in> set ys"

quotient_type 
  'a bar = "('a * int) list" / "erel3"
  apply(rule equivpI)
  unfolding erel3_def reflp_def symp_def transp_def
  by auto

fun
  intrel :: "(nat \<times> nat) \<Rightarrow> (nat \<times> nat) \<Rightarrow> bool" (infixl "\<approx>4" 50)
where
  "intrel (x, y) (u, v) = (x + v = u + y)"

quotient_type myint = "nat \<times> nat" / intrel
  by (auto simp add: equivp_def expand_fun_eq)

ML {*
test_funs absF @{context} 
     (@{typ "nat \<times> nat"}, 
      @{typ "myint"})
*}

ML {*
test_funs absF @{context} 
     (@{typ "('a * 'a) list"}, 
      @{typ "'a foo"})
*}

ML {*
test_funs repF @{context} 
     (@{typ "(('a * 'a) list * 'b)"}, 
      @{typ "('a foo * 'b)"})
*}

ML {*
test_funs absF @{context} 
     (@{typ "(('a list) * int) list"}, 
      @{typ "('a fset) bar"})
*}

ML {*
test_funs absF @{context} 
     (@{typ "('a list)"}, 
      @{typ "('a fset)"})
*}

ML {*
test_funs absF @{context} 
     (@{typ "('a list) list"}, 
      @{typ "('a fset) fset"})
*}


ML {*
test_funs absF @{context} 
     (@{typ "((nat * nat) list) list"}, 
      @{typ "((myint) fset) fset"})
*}

ML {*
test_funs absF @{context} 
     (@{typ "(('a * 'a) list) list"}, 
      @{typ "(('a * 'a) fset) fset"})
*}

ML {*
test_funs absF @{context} 
      (@{typ "(nat * nat) list"}, 
       @{typ "myint fset"})
*}

ML {*
test_funs absF @{context} 
     (@{typ "('a list) list \<Rightarrow> 'a list"}, 
      @{typ "('a fset) fset \<Rightarrow> 'a fset"})
*}

lemma
  assumes sr: "symp r"
  and     ss: "symp s"
  shows "(r OO s) x y = (s OO r) y x"
using sr ss
unfolding symp_def
apply (metis pred_comp.intros pred_compE ss symp_def)
done

lemma bla:
  assumes a1: "Quotient (op \<approx>1) abs_fset rep_fset"
  and     a2:  "Quotient r2 abs2 rep2"
  shows  "Quotient ((list_rel r2) OO (op \<approx>1) OO (list_rel r2)) 
               (abs_fset \<circ> (map abs2)) ((map rep2) \<circ> rep_fset)"
using a1
apply -
sorry

lemma bla2:
  assumes a2: "Quotient r1 abs1 rep_fset"
  and         "Quotient r2 abs2 rep2"
  shows  "Quotient ((list_rel r2) OO r1 OO (list_rel r2)) (abs1 \<circ> (map abs2)) ((map rep2) \<circ> rep_fset)"
sorry

thm bla [OF Quotient_fset]
thm bla2[OF Quotient_fset]

thm bla [OF Quotient_fset Quotient_fset]
thm bla2[OF Quotient_fset Quotient_fset]

lemma bla:
  assumes a1: "Quotient r1 abs1 rep1"
  and     a2: "Quotient r2 abs2 rep2"
  shows  "Quotient r2 (abs1 \<circ> abs2) (rep2 \<circ> rep1)"
sorry


  unfolding Quotient_def
apply auto
term rep_fset

lemma
  assumes sr: "equivp r"
  and     ss: "equivp s"
  shows "r OO s = s OO r"
apply(rule ext)
apply(rule ext)
using sr ss
nitpick

apply(auto)
apply(rule pred_compI)

definition
  relation_compose
where
  "relation_compose R1 R2 = \<lambda>x y. \<exists> z. (R1 x z \<and> R2 z y)"



end