Quot/QuotMain.thy
author Cezary Kaliszyk <kaliszyk@in.tum.de>
Wed, 10 Feb 2010 17:22:18 +0100
changeset 1124 4a4c714ff795
parent 1122 d1eaed018e5d
child 1127 243a5ceaa088
permissions -rw-r--r--
lowercase locale
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
919
c46b6abad24b cleaned some theorems
Christian Urban <urbanc@in.tum.de>
parents: 911
diff changeset
     1
(*  Title:      QuotMain.thy
759
119f7d6a3556 minor cleaning
Christian Urban <urbanc@in.tum.de>
parents: 758
diff changeset
     2
    Author:     Cezary Kaliszyk and Christian Urban
119f7d6a3556 minor cleaning
Christian Urban <urbanc@in.tum.de>
parents: 758
diff changeset
     3
*)
119f7d6a3556 minor cleaning
Christian Urban <urbanc@in.tum.de>
parents: 758
diff changeset
     4
597
8a1c8dc72b5c directory re-arrangement
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     5
theory QuotMain
920
dae99175f584 renamed QuotScript to QuotBase
Christian Urban <urbanc@in.tum.de>
parents: 919
diff changeset
     6
imports QuotBase
1122
d1eaed018e5d Changes from Makarius's code review + some noticed fixes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1116
diff changeset
     7
uses
d1eaed018e5d Changes from Makarius's code review + some noticed fixes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1116
diff changeset
     8
  ("quotient_info.ML")
d1eaed018e5d Changes from Makarius's code review + some noticed fixes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1116
diff changeset
     9
  ("quotient_typ.ML")
d1eaed018e5d Changes from Makarius's code review + some noticed fixes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1116
diff changeset
    10
  ("quotient_def.ML")
d1eaed018e5d Changes from Makarius's code review + some noticed fixes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1116
diff changeset
    11
  ("quotient_term.ML")
d1eaed018e5d Changes from Makarius's code review + some noticed fixes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1116
diff changeset
    12
  ("quotient_tacs.ML")
597
8a1c8dc72b5c directory re-arrangement
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    13
begin
8a1c8dc72b5c directory re-arrangement
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    14
1124
4a4c714ff795 lowercase locale
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1122
diff changeset
    15
locale quot_type =
597
8a1c8dc72b5c directory re-arrangement
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    16
  fixes R :: "'a \<Rightarrow> 'a \<Rightarrow> bool"
8a1c8dc72b5c directory re-arrangement
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    17
  and   Abs :: "('a \<Rightarrow> bool) \<Rightarrow> 'b"
8a1c8dc72b5c directory re-arrangement
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    18
  and   Rep :: "'b \<Rightarrow> ('a \<Rightarrow> bool)"
8a1c8dc72b5c directory re-arrangement
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    19
  assumes equivp: "equivp R"
8a1c8dc72b5c directory re-arrangement
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    20
  and     rep_prop: "\<And>y. \<exists>x. Rep y = R x"
8a1c8dc72b5c directory re-arrangement
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    21
  and     rep_inverse: "\<And>x. Abs (Rep x) = x"
8a1c8dc72b5c directory re-arrangement
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    22
  and     abs_inverse: "\<And>x. (Rep (Abs (R x))) = (R x)"
8a1c8dc72b5c directory re-arrangement
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    23
  and     rep_inject: "\<And>x y. (Rep x = Rep y) = (x = y)"
8a1c8dc72b5c directory re-arrangement
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    24
begin
8a1c8dc72b5c directory re-arrangement
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    25
8a1c8dc72b5c directory re-arrangement
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    26
definition
719
a9e55e1ef64c Renaming
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 715
diff changeset
    27
  abs::"'a \<Rightarrow> 'b"
597
8a1c8dc72b5c directory re-arrangement
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    28
where
719
a9e55e1ef64c Renaming
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 715
diff changeset
    29
  "abs x \<equiv> Abs (R x)"
597
8a1c8dc72b5c directory re-arrangement
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    30
8a1c8dc72b5c directory re-arrangement
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    31
definition
719
a9e55e1ef64c Renaming
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 715
diff changeset
    32
  rep::"'b \<Rightarrow> 'a"
597
8a1c8dc72b5c directory re-arrangement
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    33
where
719
a9e55e1ef64c Renaming
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 715
diff changeset
    34
  "rep a = Eps (Rep a)"
597
8a1c8dc72b5c directory re-arrangement
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    35
1122
d1eaed018e5d Changes from Makarius's code review + some noticed fixes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1116
diff changeset
    36
lemma homeier_lem9:
597
8a1c8dc72b5c directory re-arrangement
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    37
  shows "R (Eps (R x)) = R x"
8a1c8dc72b5c directory re-arrangement
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    38
proof -
8a1c8dc72b5c directory re-arrangement
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    39
  have a: "R x x" using equivp by (simp add: equivp_reflp_symp_transp reflp_def)
8a1c8dc72b5c directory re-arrangement
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    40
  then have "R x (Eps (R x))" by (rule someI)
8a1c8dc72b5c directory re-arrangement
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    41
  then show "R (Eps (R x)) = R x"
8a1c8dc72b5c directory re-arrangement
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    42
    using equivp unfolding equivp_def by simp
8a1c8dc72b5c directory re-arrangement
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    43
qed
8a1c8dc72b5c directory re-arrangement
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    44
1122
d1eaed018e5d Changes from Makarius's code review + some noticed fixes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1116
diff changeset
    45
theorem homeier_thm10:
930
68c1f378a70a All eq_reflections apart from the one of 'id_apply' can be removed.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 925
diff changeset
    46
  shows "abs (rep a) = a"
719
a9e55e1ef64c Renaming
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 715
diff changeset
    47
  unfolding abs_def rep_def
597
8a1c8dc72b5c directory re-arrangement
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    48
proof -
8a1c8dc72b5c directory re-arrangement
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    49
  from rep_prop
8a1c8dc72b5c directory re-arrangement
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    50
  obtain x where eq: "Rep a = R x" by auto
8a1c8dc72b5c directory re-arrangement
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    51
  have "Abs (R (Eps (Rep a))) = Abs (R (Eps (R x)))" using eq by simp
1122
d1eaed018e5d Changes from Makarius's code review + some noticed fixes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1116
diff changeset
    52
  also have "\<dots> = Abs (R x)" using homeier_lem9 by simp
597
8a1c8dc72b5c directory re-arrangement
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    53
  also have "\<dots> = Abs (Rep a)" using eq by simp
8a1c8dc72b5c directory re-arrangement
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    54
  also have "\<dots> = a" using rep_inverse by simp
8a1c8dc72b5c directory re-arrangement
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    55
  finally
8a1c8dc72b5c directory re-arrangement
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    56
  show "Abs (R (Eps (Rep a))) = a" by simp
8a1c8dc72b5c directory re-arrangement
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    57
qed
8a1c8dc72b5c directory re-arrangement
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    58
1122
d1eaed018e5d Changes from Makarius's code review + some noticed fixes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1116
diff changeset
    59
lemma homeier_lem7:
919
c46b6abad24b cleaned some theorems
Christian Urban <urbanc@in.tum.de>
parents: 911
diff changeset
    60
  shows "(R x = R y) = (Abs (R x) = Abs (R y))" (is "?LHS = ?RHS")
c46b6abad24b cleaned some theorems
Christian Urban <urbanc@in.tum.de>
parents: 911
diff changeset
    61
proof -
c46b6abad24b cleaned some theorems
Christian Urban <urbanc@in.tum.de>
parents: 911
diff changeset
    62
  have "?RHS = (Rep (Abs (R x)) = Rep (Abs (R y)))" by (simp add: rep_inject)
c46b6abad24b cleaned some theorems
Christian Urban <urbanc@in.tum.de>
parents: 911
diff changeset
    63
  also have "\<dots> = ?LHS" by (simp add: abs_inverse)
c46b6abad24b cleaned some theorems
Christian Urban <urbanc@in.tum.de>
parents: 911
diff changeset
    64
  finally show "?LHS = ?RHS" by simp
c46b6abad24b cleaned some theorems
Christian Urban <urbanc@in.tum.de>
parents: 911
diff changeset
    65
qed
597
8a1c8dc72b5c directory re-arrangement
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    66
1122
d1eaed018e5d Changes from Makarius's code review + some noticed fixes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1116
diff changeset
    67
theorem homeier_thm11:
719
a9e55e1ef64c Renaming
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 715
diff changeset
    68
  shows "R r r' = (abs r = abs r')"
1116
3acc7d25627a Some cleaning of proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1113
diff changeset
    69
  unfolding abs_def
1122
d1eaed018e5d Changes from Makarius's code review + some noticed fixes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1116
diff changeset
    70
  by (simp only: equivp[simplified equivp_def] homeier_lem7)
597
8a1c8dc72b5c directory re-arrangement
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    71
920
dae99175f584 renamed QuotScript to QuotBase
Christian Urban <urbanc@in.tum.de>
parents: 919
diff changeset
    72
lemma rep_refl:
dae99175f584 renamed QuotScript to QuotBase
Christian Urban <urbanc@in.tum.de>
parents: 919
diff changeset
    73
  shows "R (rep a) (rep a)"
1116
3acc7d25627a Some cleaning of proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1113
diff changeset
    74
  unfolding rep_def
3acc7d25627a Some cleaning of proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1113
diff changeset
    75
  by (simp add: equivp[simplified equivp_def])
920
dae99175f584 renamed QuotScript to QuotBase
Christian Urban <urbanc@in.tum.de>
parents: 919
diff changeset
    76
dae99175f584 renamed QuotScript to QuotBase
Christian Urban <urbanc@in.tum.de>
parents: 919
diff changeset
    77
719
a9e55e1ef64c Renaming
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 715
diff changeset
    78
lemma rep_abs_rsp:
a9e55e1ef64c Renaming
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 715
diff changeset
    79
  shows "R f (rep (abs g)) = R f g"
a9e55e1ef64c Renaming
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 715
diff changeset
    80
  and   "R (rep (abs g)) f = R g f"
1122
d1eaed018e5d Changes from Makarius's code review + some noticed fixes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1116
diff changeset
    81
  by (simp_all add: homeier_thm10 homeier_thm11)
597
8a1c8dc72b5c directory re-arrangement
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    82
8a1c8dc72b5c directory re-arrangement
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    83
lemma Quotient:
773
d6acae26d027 tuned comments; renamed QUOT_TRUE to Quot_True; atomize_eqv seems to not be neccessary (has it been added to Isabelle)...it is now comented out and everything still works
Christian Urban <urbanc@in.tum.de>
parents: 768
diff changeset
    84
  shows "Quotient R abs rep"
1116
3acc7d25627a Some cleaning of proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1113
diff changeset
    85
  unfolding Quotient_def
1122
d1eaed018e5d Changes from Makarius's code review + some noticed fixes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1116
diff changeset
    86
  apply(simp add: homeier_thm10)
1116
3acc7d25627a Some cleaning of proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1113
diff changeset
    87
  apply(simp add: rep_refl)
1122
d1eaed018e5d Changes from Makarius's code review + some noticed fixes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1116
diff changeset
    88
  apply(subst homeier_thm11[symmetric])
1116
3acc7d25627a Some cleaning of proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1113
diff changeset
    89
  apply(simp add: equivp[simplified equivp_def])
3acc7d25627a Some cleaning of proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1113
diff changeset
    90
  done
597
8a1c8dc72b5c directory re-arrangement
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    91
8a1c8dc72b5c directory re-arrangement
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    92
end
8a1c8dc72b5c directory re-arrangement
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    93
773
d6acae26d027 tuned comments; renamed QUOT_TRUE to Quot_True; atomize_eqv seems to not be neccessary (has it been added to Isabelle)...it is now comented out and everything still works
Christian Urban <urbanc@in.tum.de>
parents: 768
diff changeset
    94
section {* ML setup *}
d6acae26d027 tuned comments; renamed QUOT_TRUE to Quot_True; atomize_eqv seems to not be neccessary (has it been added to Isabelle)...it is now comented out and everything still works
Christian Urban <urbanc@in.tum.de>
parents: 768
diff changeset
    95
919
c46b6abad24b cleaned some theorems
Christian Urban <urbanc@in.tum.de>
parents: 911
diff changeset
    96
text {* Auxiliary data for the quotient package *}
869
ce5f78f0eac5 Finished organising an efficient datastructure for qconst_info.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 854
diff changeset
    97
597
8a1c8dc72b5c directory re-arrangement
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    98
use "quotient_info.ML"
8a1c8dc72b5c directory re-arrangement
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    99
8a1c8dc72b5c directory re-arrangement
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   100
declare [[map "fun" = (fun_map, fun_rel)]]
615
386a6b1a5203 the lift_tac produces a warning message if one of the three automatic proofs fails
Christian Urban <urbanc@in.tum.de>
parents: 614
diff changeset
   101
909
3e7a6ec549d1 Using Bexeq_rsp, and manually lifted lemma with Ex1.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 907
diff changeset
   102
lemmas [quot_thm] = fun_quotient
911
95ee248b3832 Automatic injection of Bexeq
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 909
diff changeset
   103
lemmas [quot_respect] = quot_rel_rsp
648
830b58c2fa94 decoupled QuotProd from QuotMain and also started new cleaning strategy
Christian Urban <urbanc@in.tum.de>
parents: 643
diff changeset
   104
lemmas [quot_equiv] = identity_equivp
597
8a1c8dc72b5c directory re-arrangement
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   105
919
c46b6abad24b cleaned some theorems
Christian Urban <urbanc@in.tum.de>
parents: 911
diff changeset
   106
c46b6abad24b cleaned some theorems
Christian Urban <urbanc@in.tum.de>
parents: 911
diff changeset
   107
text {* Lemmas about simplifying id's. *}
773
d6acae26d027 tuned comments; renamed QUOT_TRUE to Quot_True; atomize_eqv seems to not be neccessary (has it been added to Isabelle)...it is now comented out and everything still works
Christian Urban <urbanc@in.tum.de>
parents: 768
diff changeset
   108
lemmas [id_simps] =
925
8d51795ef54d ids *cannot* be object equalities
Christian Urban <urbanc@in.tum.de>
parents: 920
diff changeset
   109
  id_def[symmetric]
930
68c1f378a70a All eq_reflections apart from the one of 'id_apply' can be removed.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 925
diff changeset
   110
  fun_map_id
939
ce774af6b964 Changed the lambda_prs_simple_conv to use id_apply, now last eq_reflection can be removed from id_simps.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 930
diff changeset
   111
  id_apply
930
68c1f378a70a All eq_reflections apart from the one of 'id_apply' can be removed.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 925
diff changeset
   112
  id_o
68c1f378a70a All eq_reflections apart from the one of 'id_apply' can be removed.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 925
diff changeset
   113
  o_id
825
970e86082cd7 Modifictaions for new_relation.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 789
diff changeset
   114
  eq_comp_r
774
b4ffb8826105 moved get_fun into quotient_term; this simplifies the overall including structure of the package
Christian Urban <urbanc@in.tum.de>
parents: 773
diff changeset
   115
919
c46b6abad24b cleaned some theorems
Christian Urban <urbanc@in.tum.de>
parents: 911
diff changeset
   116
text {* Translation functions for the lifting process. *}
774
b4ffb8826105 moved get_fun into quotient_term; this simplifies the overall including structure of the package
Christian Urban <urbanc@in.tum.de>
parents: 773
diff changeset
   117
use "quotient_term.ML" 
b4ffb8826105 moved get_fun into quotient_term; this simplifies the overall including structure of the package
Christian Urban <urbanc@in.tum.de>
parents: 773
diff changeset
   118
b4ffb8826105 moved get_fun into quotient_term; this simplifies the overall including structure of the package
Christian Urban <urbanc@in.tum.de>
parents: 773
diff changeset
   119
919
c46b6abad24b cleaned some theorems
Christian Urban <urbanc@in.tum.de>
parents: 911
diff changeset
   120
text {* Definitions of the quotient types. *}
725
0d98a4c7f8dc renamed quotient.ML to quotient_typ.ML
Christian Urban <urbanc@in.tum.de>
parents: 720
diff changeset
   121
use "quotient_typ.ML"
597
8a1c8dc72b5c directory re-arrangement
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   122
773
d6acae26d027 tuned comments; renamed QUOT_TRUE to Quot_True; atomize_eqv seems to not be neccessary (has it been added to Isabelle)...it is now comented out and everything still works
Christian Urban <urbanc@in.tum.de>
parents: 768
diff changeset
   123
919
c46b6abad24b cleaned some theorems
Christian Urban <urbanc@in.tum.de>
parents: 911
diff changeset
   124
text {* Definitions for quotient constants. *}
597
8a1c8dc72b5c directory re-arrangement
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   125
use "quotient_def.ML"
8a1c8dc72b5c directory re-arrangement
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   126
759
119f7d6a3556 minor cleaning
Christian Urban <urbanc@in.tum.de>
parents: 758
diff changeset
   127
1122
d1eaed018e5d Changes from Makarius's code review + some noticed fixes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1116
diff changeset
   128
text {*
d1eaed018e5d Changes from Makarius's code review + some noticed fixes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1116
diff changeset
   129
  An auxiliary constant for recording some information
d1eaed018e5d Changes from Makarius's code review + some noticed fixes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1116
diff changeset
   130
  about the lifted theorem in a tactic.
919
c46b6abad24b cleaned some theorems
Christian Urban <urbanc@in.tum.de>
parents: 911
diff changeset
   131
*}
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents: 752
diff changeset
   132
definition
773
d6acae26d027 tuned comments; renamed QUOT_TRUE to Quot_True; atomize_eqv seems to not be neccessary (has it been added to Isabelle)...it is now comented out and everything still works
Christian Urban <urbanc@in.tum.de>
parents: 768
diff changeset
   133
  "Quot_True x \<equiv> True"
605
120e479ed367 first attempt to deal with Babs in regularise and cleaning (not yet working)
Christian Urban <urbanc@in.tum.de>
parents: 602
diff changeset
   134
773
d6acae26d027 tuned comments; renamed QUOT_TRUE to Quot_True; atomize_eqv seems to not be neccessary (has it been added to Isabelle)...it is now comented out and everything still works
Christian Urban <urbanc@in.tum.de>
parents: 768
diff changeset
   135
lemma 
d6acae26d027 tuned comments; renamed QUOT_TRUE to Quot_True; atomize_eqv seems to not be neccessary (has it been added to Isabelle)...it is now comented out and everything still works
Christian Urban <urbanc@in.tum.de>
parents: 768
diff changeset
   136
  shows QT_all: "Quot_True (All P) \<Longrightarrow> Quot_True P"
d6acae26d027 tuned comments; renamed QUOT_TRUE to Quot_True; atomize_eqv seems to not be neccessary (has it been added to Isabelle)...it is now comented out and everything still works
Christian Urban <urbanc@in.tum.de>
parents: 768
diff changeset
   137
  and   QT_ex:  "Quot_True (Ex P) \<Longrightarrow> Quot_True P"
907
e576a97e9a3e The missing rule.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 896
diff changeset
   138
  and   QT_ex1: "Quot_True (Ex1 P) \<Longrightarrow> Quot_True P"
1113
9f6c606d5b59 more minor space and bracket modifications.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1077
diff changeset
   139
  and   QT_lam: "Quot_True (\<lambda>x. P x) \<Longrightarrow> (\<And>x. Quot_True (P x))"
773
d6acae26d027 tuned comments; renamed QUOT_TRUE to Quot_True; atomize_eqv seems to not be neccessary (has it been added to Isabelle)...it is now comented out and everything still works
Christian Urban <urbanc@in.tum.de>
parents: 768
diff changeset
   140
  and   QT_ext: "(\<And>x. Quot_True (a x) \<Longrightarrow> f x = g x) \<Longrightarrow> (Quot_True a \<Longrightarrow> f = g)"
1116
3acc7d25627a Some cleaning of proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1113
diff changeset
   141
  by (simp_all add: Quot_True_def ext)
751
670131bcba4a some commenting
Christian Urban <urbanc@in.tum.de>
parents: 750
diff changeset
   142
773
d6acae26d027 tuned comments; renamed QUOT_TRUE to Quot_True; atomize_eqv seems to not be neccessary (has it been added to Isabelle)...it is now comented out and everything still works
Christian Urban <urbanc@in.tum.de>
parents: 768
diff changeset
   143
lemma QT_imp: "Quot_True a \<equiv> Quot_True b"
1116
3acc7d25627a Some cleaning of proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1113
diff changeset
   144
  by (simp add: Quot_True_def)
751
670131bcba4a some commenting
Christian Urban <urbanc@in.tum.de>
parents: 750
diff changeset
   145
1122
d1eaed018e5d Changes from Makarius's code review + some noticed fixes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1116
diff changeset
   146
919
c46b6abad24b cleaned some theorems
Christian Urban <urbanc@in.tum.de>
parents: 911
diff changeset
   147
text {* Tactics for proving the lifted theorems *}
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents: 752
diff changeset
   148
use "quotient_tacs.ML"
751
670131bcba4a some commenting
Christian Urban <urbanc@in.tum.de>
parents: 750
diff changeset
   149
689
Christian Urban <urbanc@in.tum.de>
parents: 688
diff changeset
   150
section {* Methods / Interface *}
Christian Urban <urbanc@in.tum.de>
parents: 688
diff changeset
   151
1122
d1eaed018e5d Changes from Makarius's code review + some noticed fixes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1116
diff changeset
   152
(* TODO inline *)
632
d23416464f62 added methods for the lifting_tac and the other tacs
Christian Urban <urbanc@in.tum.de>
parents: 624
diff changeset
   153
ML {*
896
4e0b89d886ac liftin and lifing_tac can now lift several "and"-separated goals at once; the raw-theorems have to be given in the order of goals
Christian Urban <urbanc@in.tum.de>
parents: 870
diff changeset
   154
fun mk_method1 tac thms ctxt =
4e0b89d886ac liftin and lifing_tac can now lift several "and"-separated goals at once; the raw-theorems have to be given in the order of goals
Christian Urban <urbanc@in.tum.de>
parents: 870
diff changeset
   155
  SIMPLE_METHOD (HEADGOAL (tac ctxt thms)) 
632
d23416464f62 added methods for the lifting_tac and the other tacs
Christian Urban <urbanc@in.tum.de>
parents: 624
diff changeset
   156
d23416464f62 added methods for the lifting_tac and the other tacs
Christian Urban <urbanc@in.tum.de>
parents: 624
diff changeset
   157
fun mk_method2 tac ctxt =
d23416464f62 added methods for the lifting_tac and the other tacs
Christian Urban <urbanc@in.tum.de>
parents: 624
diff changeset
   158
  SIMPLE_METHOD (HEADGOAL (tac ctxt)) 
d23416464f62 added methods for the lifting_tac and the other tacs
Christian Urban <urbanc@in.tum.de>
parents: 624
diff changeset
   159
*}
d23416464f62 added methods for the lifting_tac and the other tacs
Christian Urban <urbanc@in.tum.de>
parents: 624
diff changeset
   160
d23416464f62 added methods for the lifting_tac and the other tacs
Christian Urban <urbanc@in.tum.de>
parents: 624
diff changeset
   161
method_setup lifting =
896
4e0b89d886ac liftin and lifing_tac can now lift several "and"-separated goals at once; the raw-theorems have to be given in the order of goals
Christian Urban <urbanc@in.tum.de>
parents: 870
diff changeset
   162
  {* Attrib.thms >> (mk_method1 Quotient_Tacs.lift_tac) *}
1122
d1eaed018e5d Changes from Makarius's code review + some noticed fixes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1116
diff changeset
   163
  {* lifts theorems to quotient types *}
632
d23416464f62 added methods for the lifting_tac and the other tacs
Christian Urban <urbanc@in.tum.de>
parents: 624
diff changeset
   164
d23416464f62 added methods for the lifting_tac and the other tacs
Christian Urban <urbanc@in.tum.de>
parents: 624
diff changeset
   165
method_setup lifting_setup =
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents: 752
diff changeset
   166
  {* Attrib.thm >> (mk_method1 Quotient_Tacs.procedure_tac) *}
1122
d1eaed018e5d Changes from Makarius's code review + some noticed fixes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1116
diff changeset
   167
  {* sets up the three goals for the quotient lifting procedure *}
632
d23416464f62 added methods for the lifting_tac and the other tacs
Christian Urban <urbanc@in.tum.de>
parents: 624
diff changeset
   168
d23416464f62 added methods for the lifting_tac and the other tacs
Christian Urban <urbanc@in.tum.de>
parents: 624
diff changeset
   169
method_setup regularize =
1113
9f6c606d5b59 more minor space and bracket modifications.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1077
diff changeset
   170
  {* Scan.succeed (mk_method2 Quotient_Tacs.regularize_tac) *}
1122
d1eaed018e5d Changes from Makarius's code review + some noticed fixes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1116
diff changeset
   171
  {* proves the regularization goals from the quotient lifting procedure *}
632
d23416464f62 added methods for the lifting_tac and the other tacs
Christian Urban <urbanc@in.tum.de>
parents: 624
diff changeset
   172
d23416464f62 added methods for the lifting_tac and the other tacs
Christian Urban <urbanc@in.tum.de>
parents: 624
diff changeset
   173
method_setup injection =
773
d6acae26d027 tuned comments; renamed QUOT_TRUE to Quot_True; atomize_eqv seems to not be neccessary (has it been added to Isabelle)...it is now comented out and everything still works
Christian Urban <urbanc@in.tum.de>
parents: 768
diff changeset
   174
  {* Scan.succeed (mk_method2 Quotient_Tacs.all_injection_tac) *}
1122
d1eaed018e5d Changes from Makarius's code review + some noticed fixes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1116
diff changeset
   175
  {* proves the rep/abs injection goals from the quotient lifting procedure *}
632
d23416464f62 added methods for the lifting_tac and the other tacs
Christian Urban <urbanc@in.tum.de>
parents: 624
diff changeset
   176
d23416464f62 added methods for the lifting_tac and the other tacs
Christian Urban <urbanc@in.tum.de>
parents: 624
diff changeset
   177
method_setup cleaning =
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents: 752
diff changeset
   178
  {* Scan.succeed (mk_method2 Quotient_Tacs.clean_tac) *}
1122
d1eaed018e5d Changes from Makarius's code review + some noticed fixes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1116
diff changeset
   179
  {* proves the cleaning goals from the quotient lifting procedure *}
632
d23416464f62 added methods for the lifting_tac and the other tacs
Christian Urban <urbanc@in.tum.de>
parents: 624
diff changeset
   180
1068
62e54830590f A proper version of the attribute
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 939
diff changeset
   181
attribute_setup quot_lifted =
1077
44461d5615eb some tuning
Christian Urban <urbanc@in.tum.de>
parents: 1068
diff changeset
   182
  {* Scan.succeed Quotient_Tacs.lifted_attrib *}
1122
d1eaed018e5d Changes from Makarius's code review + some noticed fixes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1116
diff changeset
   183
  {* lifts theorems to quotient types *}
1068
62e54830590f A proper version of the attribute
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 939
diff changeset
   184
597
8a1c8dc72b5c directory re-arrangement
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   185
end
8a1c8dc72b5c directory re-arrangement
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   186