merged
authorChristian Urban <urbanc@in.tum.de>
Fri, 11 Dec 2009 19:22:30 +0100
changeset 724 d705d7ae2410
parent 723 93dce7c71929 (current diff)
parent 721 19032e71fb1c (diff)
child 725 0d98a4c7f8dc
merged
Quot/Examples/FSet3.thy
--- a/Quot/Examples/IntEx2.thy	Fri Dec 11 19:19:50 2009 +0100
+++ b/Quot/Examples/IntEx2.thy	Fri Dec 11 19:22:30 2009 +0100
@@ -194,8 +194,8 @@
    (simp)
 
 lemma add:
-     "(ABS_int (x,y)) + (ABS_int (u,v)) =
-      (ABS_int (x + u, y + v))"
+     "(abs_int (x,y)) + (abs_int (u,v)) =
+      (abs_int (x + u, y + v))"
 apply(simp add: plus_int_def)
 apply(fold plus_raw.simps)
 apply(rule Quotient_rel_abs[OF Quotient_int])
--- a/Quot/QuotMain.thy	Fri Dec 11 19:19:50 2009 +0100
+++ b/Quot/QuotMain.thy	Fri Dec 11 19:22:30 2009 +0100
@@ -17,14 +17,14 @@
 begin
 
 definition
-  ABS::"'a \<Rightarrow> 'b"
+  abs::"'a \<Rightarrow> 'b"
 where
-  "ABS x \<equiv> Abs (R x)"
+  "abs x \<equiv> Abs (R x)"
 
 definition
-  REP::"'b \<Rightarrow> 'a"
+  rep::"'b \<Rightarrow> 'a"
 where
-  "REP a = Eps (Rep a)"
+  "rep a = Eps (Rep a)"
 
 lemma lem9:
   shows "R (Eps (R x)) = R x"
@@ -36,9 +36,9 @@
 qed
 
 theorem thm10:
-  shows "ABS (REP a) \<equiv> a"
+  shows "abs (rep a) \<equiv> a"
   apply  (rule eq_reflection)
-  unfolding ABS_def REP_def
+  unfolding abs_def rep_def
 proof -
   from rep_prop
   obtain x where eq: "Rep a = R x" by auto
@@ -50,9 +50,9 @@
   show "Abs (R (Eps (Rep a))) = a" by simp
 qed
 
-lemma REP_refl:
-  shows "R (REP a) (REP a)"
-unfolding REP_def
+lemma rep_refl:
+  shows "R (rep a) (rep a)"
+unfolding rep_def
 by (simp add: equivp[simplified equivp_def])
 
 lemma lem7:
@@ -64,21 +64,21 @@
 done
 
 theorem thm11:
-  shows "R r r' = (ABS r = ABS r')"
-unfolding ABS_def
+  shows "R r r' = (abs r = abs r')"
+unfolding abs_def
 by (simp only: equivp[simplified equivp_def] lem7)
 
 
-lemma REP_ABS_rsp:
-  shows "R f (REP (ABS g)) = R f g"
-  and   "R (REP (ABS g)) f = R g f"
+lemma rep_abs_rsp:
+  shows "R f (rep (abs g)) = R f g"
+  and   "R (rep (abs g)) f = R g f"
 by (simp_all add: thm10 thm11)
 
 lemma Quotient:
-  "Quotient R ABS REP"
+  "Quotient R abs rep"
 apply(unfold Quotient_def)
 apply(simp add: thm10)
-apply(simp add: REP_refl)
+apply(simp add: rep_refl)
 apply(subst thm11[symmetric])
 apply(simp add: equivp[simplified equivp_def])
 done
--- a/Quot/QuotScript.thy	Fri Dec 11 19:19:50 2009 +0100
+++ b/Quot/QuotScript.thy	Fri Dec 11 19:22:30 2009 +0100
@@ -29,12 +29,12 @@
 
 lemma equivp_transp:
   shows "equivp E \<Longrightarrow> (\<And>x y z. E x y \<Longrightarrow> E y z \<Longrightarrow> E x z)"
-by (metis equivp_reflp_symp_transp transp_def)
+  by (metis equivp_reflp_symp_transp transp_def)
 
 lemma equivpI:
   assumes "reflp R" "symp R" "transp R"
   shows "equivp R"
-using assms by (simp add: equivp_reflp_symp_transp)
+  using assms by (simp add: equivp_reflp_symp_transp)
 
 definition
   "part_equivp E \<equiv> (\<exists>x. E x x) \<and> (\<forall>x y. E x y = (E x x \<and> E y y \<and> (E x = E y)))"
@@ -81,6 +81,12 @@
   using a unfolding Quotient_def
   by blast
 
+lemma Quotient_rel_abs:
+  assumes a: "Quotient E Abs Rep"
+  shows "E r s \<Longrightarrow> Abs r = Abs s"
+  using a unfolding Quotient_def
+  by blast
+
 lemma identity_equivp:
   shows "equivp (op =)"
   unfolding equivp_def
@@ -129,7 +135,7 @@
 
 lemma fun_rel_eq:
   "(op =) ===> (op =) \<equiv> (op =)"
-by (rule eq_reflection) (simp add: expand_fun_eq)
+  by (rule eq_reflection) (simp add: expand_fun_eq)
 
 lemma fun_quotient:
   assumes q1: "Quotient R1 abs1 rep1"
@@ -203,6 +209,12 @@
   shows "R x1 (Rep (Abs x2))"
   using q a by (metis Quotient_rel[OF q] Quotient_abs_rep[OF q] Quotient_rep_reflp[OF q])
 
+lemma rep_abs_rsp_left:
+  assumes q: "Quotient R Abs Rep"
+  and     a: "R x1 x2"
+  shows "R (Rep (Abs x1)) x2"
+using q a by (metis Quotient_rel[OF q] Quotient_abs_rep[OF q] Quotient_rep_reflp[OF q])
+
 (* In the following theorem R1 can be instantiated with anything,
    but we know some of the types of the Rep and Abs functions;
    so by solving Quotient assumptions we can get a unique R1 that
@@ -309,11 +321,11 @@
 
 lemma ball_all_comm:
   "(\<And>y. (\<forall>x\<in>P. A x y) \<longrightarrow> (\<forall>x. B x y)) \<Longrightarrow> ((\<forall>x\<in>P. \<forall>y. A x y) \<longrightarrow> (\<forall>x. \<forall>y. B x y))"
-by auto
+  by auto
 
 lemma bex_ex_comm:
   "((\<exists>y. \<exists>x. A x y) \<longrightarrow> (\<exists>y. \<exists>x\<in>P. B x y)) \<Longrightarrow> ((\<exists>x. \<exists>y. A x y) \<longrightarrow> (\<exists>x\<in>P. \<exists>y. B x y))"
-by auto
+  by auto
 
 (* Bounded abstraction *)
 definition
@@ -358,20 +370,20 @@
 lemma babs_simp:
   assumes q: "Quotient R1 Abs Rep"
   shows "((R1 ===> R2) (Babs (Respects R1) f) (Babs (Respects R1) g)) = ((R1 ===> R2) f g)"
-apply(rule iffI)
-apply(simp_all only: babs_rsp[OF q])
-apply(auto simp add: Babs_def)
-apply (subgoal_tac "x \<in> Respects R1 \<and> y \<in> Respects R1")
-apply(metis Babs_def)
-apply (simp add: in_respects)
-using Quotient_rel[OF q]
-by metis
+  apply(rule iffI)
+  apply(simp_all only: babs_rsp[OF q])
+  apply(auto simp add: Babs_def)
+  apply (subgoal_tac "x \<in> Respects R1 \<and> y \<in> Respects R1")
+  apply(metis Babs_def)
+  apply (simp add: in_respects)
+  using Quotient_rel[OF q]
+  by metis
 
-
-(* needed for regularising? *)
+(* If a user proves that a particular functional relation is an equivalence
+   this may be useful in regularising *)
 lemma babs_reg_eqv:
   shows "equivp R \<Longrightarrow> Babs (Respects R) P = P"
-by (simp add: expand_fun_eq Babs_def in_respects equivp_reflp)
+  by (simp add: expand_fun_eq Babs_def in_respects equivp_reflp)
 
 (* 2 lemmas needed for cleaning of quantifiers *)
 lemma all_prs:
@@ -389,12 +401,12 @@
 lemma fun_rel_id:
   assumes a: "\<And>x y. R1 x y \<Longrightarrow> R2 (f x) (g y)"
   shows "(R1 ===> R2) f g"
-using a by simp
+  using a by simp
 
 lemma fun_rel_id_asm:
   assumes a: "\<And>x y. R1 x y \<Longrightarrow> (A \<longrightarrow> R2 (f x) (g y))"
   shows "A \<longrightarrow> (R1 ===> R2) f g"
-using a by auto
+  using a by auto
 
 lemma quot_rel_rsp:
   assumes a: "Quotient R Abs Rep"
@@ -404,7 +416,53 @@
   apply(assumption)+
   done
 
+lemma o_prs:
+  assumes q1: "Quotient R1 Abs1 Rep1"
+  and     q2: "Quotient R2 Abs2 Rep2"
+  and     q3: "Quotient R3 Abs3 Rep3"
+  shows "(Rep1 ---> Abs3) (((Abs2 ---> Rep3) f) o ((Abs1 ---> Rep2) g)) = f o g"
+  using Quotient_abs_rep[OF q1] Quotient_abs_rep[OF q2] Quotient_abs_rep[OF q3]
+  unfolding o_def expand_fun_eq by simp
 
+lemma o_rsp:
+  assumes q1: "Quotient R1 Abs1 Rep1"
+  and     q2: "Quotient R2 Abs2 Rep2"
+  and     q3: "Quotient R3 Abs3 Rep3"
+  and     a1: "(R2 ===> R3) f1 f2"
+  and     a2: "(R1 ===> R2) g1 g2"
+  shows "(R1 ===> R3) (f1 o g1) (f2 o g2)"
+  using a1 a2 unfolding o_def expand_fun_eq
+  by (auto)
+
+lemma cond_prs:
+  assumes a: "Quotient R absf repf"
+  shows "absf (if a then repf b else repf c) = (if a then b else c)"
+  using a unfolding Quotient_def by auto
+
+lemma if_prs:
+  assumes q: "Quotient R Abs Rep"
+  shows "Abs (If a (Rep b) (Rep c)) = If a b c"
+using Quotient_abs_rep[OF q] by auto
+
+(* q not used *)
+lemma if_rsp:
+  assumes q: "Quotient R Abs Rep"
+  and     a: "a1 = a2" "R b1 b2" "R c1 c2"
+  shows "R (If a1 b1 c1) (If a2 b2 c2)"
+using a by auto
+
+lemma let_prs:
+  assumes q1: "Quotient R1 Abs1 Rep1"
+  and     q2: "Quotient R2 Abs2 Rep2"
+  shows "Abs2 (Let (Rep1 x) ((Abs1 ---> Rep2) f)) = Let x f"
+  using Quotient_abs_rep[OF q1] Quotient_abs_rep[OF q2] by auto
+
+lemma let_rsp:
+  assumes q1: "Quotient R1 Abs1 Rep1"
+  and     a1: "(R1 ===> R2) f g"
+  and     a2: "R1 x y"
+  shows "R2 ((Let x f)::'c) ((Let y g)::'c)"
+  using apply_rsp[OF q1 a1] a2 by auto
 
 
 
@@ -413,104 +471,62 @@
 (* REST OF THE FILE IS UNUSED (until now) *)
 (******************************************)
 
-(* Always safe to apply, but not too helpful *)
-lemma app_prs2:
-  assumes q1: "Quotient R1 abs1 rep1"
-  and     q2: "Quotient R2 abs2 rep2"
-  shows  "((abs1 ---> rep2) ((rep1 ---> abs2) f) (rep1 x)) = rep2 (((rep1 ---> abs2) f) x)"
-unfolding expand_fun_eq
-using Quotient_abs_rep[OF q1] Quotient_abs_rep[OF q2]
-by simp
-
-lemma Quotient_rel_abs:
-  assumes a: "Quotient E Abs Rep"
-  shows "E r s \<Longrightarrow> Abs r = Abs s"
-using a unfolding Quotient_def
-by blast
-
-lemma Quotient_rel_abs_eq:
-  assumes a: "Quotient E Abs Rep"
-  shows "E r r \<Longrightarrow> E s s \<Longrightarrow> E r s = (Abs r = Abs s)"
-using a unfolding Quotient_def
-by blast
-
 lemma in_fun:
   shows "x \<in> ((f ---> g) s) = g (f x \<in> s)"
-by (simp add: mem_def)
+  by (simp add: mem_def)
 
-lemma RESPECTS_THM:
+lemma respects_thm:
   shows "Respects (R1 ===> R2) f = (\<forall>x y. R1 x y \<longrightarrow> R2 (f x) (f y))"
-unfolding Respects_def
-by (simp add: expand_fun_eq) 
+  unfolding Respects_def
+  by (simp add: expand_fun_eq)
 
-lemma RESPECTS_REP_ABS:
+lemma respects_rep_abs:
   assumes a: "Quotient R1 Abs1 Rep1"
   and     b: "Respects (R1 ===> R2) f"
   and     c: "R1 x x"
   shows "R2 (f (Rep1 (Abs1 x))) (f x)"
-using a b[simplified RESPECTS_THM] c unfolding Quotient_def
-by blast
+  using a b[simplified respects_thm] c unfolding Quotient_def
+  by blast
 
-lemma RESPECTS_MP:
+lemma respects_mp:
   assumes a: "Respects (R1 ===> R2) f"
   and     b: "R1 x y"
   shows "R2 (f x) (f y)"
-using a b unfolding Respects_def
-by simp
+  using a b unfolding Respects_def
+  by simp
 
-lemma RESPECTS_o:
+lemma respects_o:
   assumes a: "Respects (R2 ===> R3) f"
   and     b: "Respects (R1 ===> R2) g"
   shows "Respects (R1 ===> R3) (f o g)"
-using a b unfolding Respects_def
-by simp
+  using a b unfolding Respects_def
+  by simp
 
-lemma fun_rel_EQ_REL:
+lemma fun_rel_eq_rel:
   assumes q1: "Quotient R1 Abs1 Rep1"
   and     q2: "Quotient R2 Abs2 Rep2"
-  shows "(R1 ===> R2) f g = ((Respects (R1 ===> R2) f) \<and> (Respects (R1 ===> R2) g) 
+  shows "(R1 ===> R2) f g = ((Respects (R1 ===> R2) f) \<and> (Respects (R1 ===> R2) g)
                              \<and> ((Rep1 ---> Abs2) f = (Rep1 ---> Abs2) g))"
-using fun_quotient[OF q1 q2] unfolding Respects_def Quotient_def expand_fun_eq
-by blast
-
-(* Not used since in the end we just unfold fun_map *)
-lemma APP_PRS:
-  assumes q1: "Quotient R1 abs1 rep1"
-  and     q2: "Quotient R2 abs2 rep2"
-  shows "abs2 ((abs1 ---> rep2) f (rep1 x)) = f x"
-unfolding expand_fun_eq
-using Quotient_abs_rep[OF q1] Quotient_abs_rep[OF q2]
-by simp
+  using fun_quotient[OF q1 q2] unfolding Respects_def Quotient_def expand_fun_eq
+  by blast
 
-(* Ask Peter: assumption q1 and q2 not used and lemma is the 'identity' *)
-lemma LAMBDA_RSP:
-  assumes q1: "Quotient R1 Abs1 Rep1"
-  and     q2: "Quotient R2 Abs2 Rep2"
-  and     a: "(R1 ===> R2) f1 f2"
-  shows "(R1 ===> R2) (\<lambda>x. f1 x) (\<lambda>y. f2 y)"
-by (rule a)
+lemma let_babs:
+  "v \<in> r \<Longrightarrow> Let v (Babs r lam) = Let v lam"
+  by (simp add: Babs_def)
 
-(* ASK Peter about next four lemmas in quotientScript
-lemma ABSTRACT_PRS:
-  assumes q1: "Quotient R1 Abs1 Rep1"
-  and     q2: "Quotient R2 Abs2 Rep2"
-  shows "f = (Rep1 ---> Abs2) ???"
-*)
-
-
-lemma fun_rel_EQUALS:
+lemma fun_rel_equals:
   assumes q1: "Quotient R1 Abs1 Rep1"
   and     q2: "Quotient R2 Abs2 Rep2"
   and     r1: "Respects (R1 ===> R2) f"
   and     r2: "Respects (R1 ===> R2) g" 
   shows "((Rep1 ---> Abs2) f = (Rep1 ---> Abs2) g) = (\<forall>x y. R1 x y \<longrightarrow> R2 (f x) (g y))"
-apply(rule_tac iffI)
-using fun_quotient[OF q1 q2] r1 r2 unfolding Quotient_def Respects_def
-apply(metis apply_rsp')
-using r1 unfolding Respects_def expand_fun_eq
-apply(simp (no_asm_use))
-apply(metis Quotient_rel[OF q2] Quotient_rel_rep[OF q1])
-done
+  apply(rule_tac iffI)
+  using fun_quotient[OF q1 q2] r1 r2 unfolding Quotient_def Respects_def
+  apply(metis apply_rsp')
+  using r1 unfolding Respects_def expand_fun_eq
+  apply(simp (no_asm_use))
+  apply(metis Quotient_rel[OF q2] Quotient_rel_rep[OF q1])
+  done
 
 (* ask Peter: fun_rel_IMP used twice *) 
 lemma fun_rel_IMP2:
@@ -520,100 +536,31 @@
   and     r2: "Respects (R1 ===> R2) g" 
   and     a:  "(Rep1 ---> Abs2) f = (Rep1 ---> Abs2) g"
   shows "R1 x y \<Longrightarrow> R2 (f x) (g y)"
-using q1 q2 r1 r2 a
-by (simp add: fun_rel_EQUALS)
+  using q1 q2 r1 r2 a
+  by (simp add: fun_rel_equals)
 
-lemma LAMBDA_REP_ABS_RSP:
+lemma lambda_rep_abs_rsp:
   assumes r1: "\<And>r r'. R1 r r' \<Longrightarrow>R1 r (Rep1 (Abs1 r'))"
   and     r2: "\<And>r r'. R2 r r' \<Longrightarrow>R2 r (Rep2 (Abs2 r'))"
   shows "(R1 ===> R2) f1 f2 \<Longrightarrow> (R1 ===> R2) f1 ((Abs1 ---> Rep2) ((Rep1 ---> Abs2) f2))"
-using r1 r2 by auto
-
-(* Not used *)
-lemma rep_abs_rsp_left:
-  assumes q: "Quotient R Abs Rep"
-  and     a: "R x1 x2"
-  shows "R (Rep (Abs x1)) x2"
-using q a by (metis Quotient_rel[OF q] Quotient_abs_rep[OF q] Quotient_rep_reflp[OF q])
-
-
-
-(* bool theory: COND, LET *)
-lemma IF_PRS:
-  assumes q: "Quotient R Abs Rep"
-  shows "If a b c = Abs (If a (Rep b) (Rep c))"
-using Quotient_abs_rep[OF q] by auto
-
-(* ask peter: no use of q *)
-lemma IF_RSP:
-  assumes q: "Quotient R Abs Rep"
-  and     a: "a1 = a2" "R b1 b2" "R c1 c2"
-  shows "R (If a1 b1 c1) (If a2 b2 c2)"
-using a by auto
-
-lemma LET_PRS:
-  assumes q1: "Quotient R1 Abs1 Rep1"
-  and     q2: "Quotient R2 Abs2 Rep2"
-  shows "Let x f = Abs2 (Let (Rep1 x) ((Abs1 ---> Rep2) f))"
-using Quotient_abs_rep[OF q1] Quotient_abs_rep[OF q2] by auto
-
-lemma LET_RSP:
-  assumes q1: "Quotient R1 Abs1 Rep1"
-  and     a1: "(R1 ===> R2) f g"
-  and     a2: "R1 x y"
-  shows "R2 ((Let x f)::'c) ((Let y g)::'c)"
-using apply_rsp[OF q1 a1] a2
-by auto
-
-
+  using r1 r2 by auto
 
 (* ask peter what are literal_case *)
 (* literal_case_PRS *)
 (* literal_case_RSP *)
-
-
-
-
-
-(* combinators: I, K, o, C, W *)
+(* Cez: !f x. literal_case f x = f x *)
 
 (* We use id_simps which includes id_apply; so these 2 theorems can be removed *)
-
-lemma I_PRS:
+lemma id_prs:
   assumes q: "Quotient R Abs Rep"
-  shows "id e = Abs (id (Rep e))"
-using Quotient_abs_rep[OF q] by auto
+  shows "Abs (id (Rep e)) = id e"
+  using Quotient_abs_rep[OF q] by auto
 
-lemma I_RSP:
+lemma id_rsp:
   assumes q: "Quotient R Abs Rep"
   and     a: "R e1 e2"
   shows "R (id e1) (id e2)"
-using a by auto
-
-lemma o_PRS:
-  assumes q1: "Quotient R1 Abs1 Rep1"
-  and     q2: "Quotient R2 Abs2 Rep2"
-  and     q3: "Quotient R3 Abs3 Rep3"
-  shows "f o g = (Rep1 ---> Abs3) (((Abs2 ---> Rep3) f) o ((Abs1 ---> Rep2) g))"
-using Quotient_abs_rep[OF q1] Quotient_abs_rep[OF q2] Quotient_abs_rep[OF q3]
-unfolding o_def expand_fun_eq
-by simp
-
-lemma o_RSP:
-  assumes q1: "Quotient R1 Abs1 Rep1"
-  and     q2: "Quotient R2 Abs2 Rep2"
-  and     q3: "Quotient R3 Abs3 Rep3"
-  and     a1: "(R2 ===> R3) f1 f2"
-  and     a2: "(R1 ===> R2) g1 g2"
-  shows "(R1 ===> R3) (f1 o g1) (f2 o g2)"
-using a1 a2 unfolding o_def expand_fun_eq
-by (auto)
-
-lemma COND_PRS:
-  assumes a: "Quotient R absf repf"
-  shows "(if a then b else c) = absf (if a then repf b else repf c)"
-  using a unfolding Quotient_def by auto
-
+  using a by auto
 
 end
 
--- a/Quot/ROOT.ML	Fri Dec 11 19:19:50 2009 +0100
+++ b/Quot/ROOT.ML	Fri Dec 11 19:22:30 2009 +0100
@@ -4,8 +4,10 @@
    ["QuotMain",
     "Examples/FSet",
     "Examples/FSet2",
+    "Examples/FSet3",
     "Examples/IntEx",
     "Examples/IntEx2",
     "Examples/LFex",
     "Examples/LamEx",
-    "Examples/LarryDatatype"];
+    "Examples/LarryDatatype",
+    "Examples/LarryInt"];
--- a/Quot/quotient.ML	Fri Dec 11 19:19:50 2009 +0100
+++ b/Quot/quotient.ML	Fri Dec 11 19:22:30 2009 +0100
@@ -132,12 +132,12 @@
   val rep = Const (rep_name, abs_ty --> rep_ty)
 
   (* ABS and REP definitions *)
-  val ABS_const = Const (@{const_name "QUOT_TYPE.ABS"}, dummyT )
-  val REP_const = Const (@{const_name "QUOT_TYPE.REP"}, dummyT )
+  val ABS_const = Const (@{const_name "QUOT_TYPE.abs"}, dummyT )
+  val REP_const = Const (@{const_name "QUOT_TYPE.rep"}, dummyT )
   val ABS_trm = Syntax.check_term lthy1 (ABS_const $ rel $ abs)
   val REP_trm = Syntax.check_term lthy1 (REP_const $ rep)
-  val ABS_name = Binding.prefix_name "ABS_" qty_name
-  val REP_name = Binding.prefix_name "REP_" qty_name
+  val ABS_name = Binding.prefix_name "abs_" qty_name
+  val REP_name = Binding.prefix_name "rep_" qty_name
   val (((ABS, ABS_def), (REP, REP_def)), lthy2) =
          lthy1 |> define (ABS_name, NoSyn, ABS_trm)
                ||>> define (REP_name, NoSyn, REP_trm)
--- a/Quot/quotient_def.ML	Fri Dec 11 19:19:50 2009 +0100
+++ b/Quot/quotient_def.ML	Fri Dec 11 19:22:30 2009 +0100
@@ -31,18 +31,18 @@
 fun get_fun_aux lthy s fs =
   case (maps_lookup (ProofContext.theory_of lthy) s) of
     SOME info => list_comb (Const (#mapfun info, dummyT), fs)
-  | NONE      => raise 
+  | NONE      => raise
         (LIFT_MATCH (space_implode " " ["get_fun_aux: no map for type", quote s, "."]))
 
 fun get_const flag lthy _ qty =
 (* FIXME: check here that _ and qty are related *)
-let 
+let
   val thy = ProofContext.theory_of lthy
   val qty_name = Long_Name.base_name (fst (dest_Type qty))
 in
   case flag of
-    absF => Const (Sign.full_bname thy ("ABS_" ^ qty_name), dummyT)
-  | repF => Const (Sign.full_bname thy ("REP_" ^ qty_name), dummyT)
+    absF => Const (Sign.full_bname thy ("abs_" ^ qty_name), dummyT)
+  | repF => Const (Sign.full_bname thy ("rep_" ^ qty_name), dummyT)
 end
 
 
@@ -53,17 +53,17 @@
 
 fun get_fun flag lthy (rty, qty) =
   if rty = qty then mk_identity qty else
-  case (rty, qty) of 
+  case (rty, qty) of
     (Type ("fun", [ty1, ty2]), Type ("fun", [ty1', ty2'])) =>
      let
        val fs_ty1 = get_fun (negF flag) lthy (ty1, ty1')
        val fs_ty2 = get_fun flag lthy (ty2, ty2')
-     in  
+     in
        get_fun_aux lthy "fun" [fs_ty1, fs_ty2]
-     end 
+     end
   | (Type (s, []), Type (s', [])) =>
      if s = s'
-     then mk_identity qty 
+     then mk_identity qty
      else get_const flag lthy rty qty
   | (Type (s, tys), Type (s', tys')) =>
      if s = s'
@@ -71,7 +71,7 @@
      else get_const flag lthy rty qty
   | (TFree x, TFree x') =>
      if x = x'
-     then mk_identity qty 
+     then mk_identity qty
      else raise (LIFT_MATCH "get_fun")
   | (TVar _, TVar _) => raise (LIFT_MATCH "get_fun")
   | _ => raise (LIFT_MATCH "get_fun")