src/HOL/Ctr_Sugar.thy
author haftmann
Sat, 25 Jan 2014 23:50:49 +0100
changeset 55147 bce3dbc11f95
parent 54701 4ed7454aebde
child 55160 2d69438b1b0c
permissions -rw-r--r--
prefer explicit code symbol type over ad-hoc name mangling
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
54397
f4b4fa25ce56 tuned headers
blanchet
parents: 54396
diff changeset
     1
(*  Title:      HOL/Ctr_Sugar.thy
49486
64cc57c0d0fe generate "expand" property
blanchet
parents: 49312
diff changeset
     2
    Author:     Jasmin Blanchette, TU Muenchen
54397
f4b4fa25ce56 tuned headers
blanchet
parents: 54396
diff changeset
     3
    Copyright   2012, 2013
49075
ed769978dc8d rearrange dependencies
blanchet
parents:
diff changeset
     4
51797
182454c06a80 tuned ML and thy file names
blanchet
parents: 51781
diff changeset
     5
Wrapping existing freely generated type's constructors.
49075
ed769978dc8d rearrange dependencies
blanchet
parents:
diff changeset
     6
*)
ed769978dc8d rearrange dependencies
blanchet
parents:
diff changeset
     7
51797
182454c06a80 tuned ML and thy file names
blanchet
parents: 51781
diff changeset
     8
header {* Wrapping Existing Freely Generated Type's Constructors *}
49075
ed769978dc8d rearrange dependencies
blanchet
parents:
diff changeset
     9
54006
9fe1bd54d437 renamed theory file
blanchet
parents: 52969
diff changeset
    10
theory Ctr_Sugar
54398
100c0eaf63d5 moved 'Ctr_Sugar' further up the theory hierarchy, so that 'Datatype' can use it
blanchet
parents: 54397
diff changeset
    11
imports HOL
49075
ed769978dc8d rearrange dependencies
blanchet
parents:
diff changeset
    12
keywords
54398
100c0eaf63d5 moved 'Ctr_Sugar' further up the theory hierarchy, so that 'Datatype' can use it
blanchet
parents: 54397
diff changeset
    13
  "print_case_translations" :: diag and
54626
8a5e82425e55 added 'no_code' option
blanchet
parents: 54615
diff changeset
    14
  "wrap_free_constructors" :: thy_goal
49075
ed769978dc8d rearrange dependencies
blanchet
parents:
diff changeset
    15
begin
ed769978dc8d rearrange dependencies
blanchet
parents:
diff changeset
    16
54398
100c0eaf63d5 moved 'Ctr_Sugar' further up the theory hierarchy, so that 'Datatype' can use it
blanchet
parents: 54397
diff changeset
    17
consts
100c0eaf63d5 moved 'Ctr_Sugar' further up the theory hierarchy, so that 'Datatype' can use it
blanchet
parents: 54397
diff changeset
    18
  case_guard :: "bool \<Rightarrow> 'a \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'b"
100c0eaf63d5 moved 'Ctr_Sugar' further up the theory hierarchy, so that 'Datatype' can use it
blanchet
parents: 54397
diff changeset
    19
  case_nil :: "'a \<Rightarrow> 'b"
100c0eaf63d5 moved 'Ctr_Sugar' further up the theory hierarchy, so that 'Datatype' can use it
blanchet
parents: 54397
diff changeset
    20
  case_cons :: "('a \<Rightarrow> 'b) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b"
100c0eaf63d5 moved 'Ctr_Sugar' further up the theory hierarchy, so that 'Datatype' can use it
blanchet
parents: 54397
diff changeset
    21
  case_elem :: "'a \<Rightarrow> 'b \<Rightarrow> 'a \<Rightarrow> 'b"
100c0eaf63d5 moved 'Ctr_Sugar' further up the theory hierarchy, so that 'Datatype' can use it
blanchet
parents: 54397
diff changeset
    22
  case_abs :: "('c \<Rightarrow> 'b) \<Rightarrow> 'b"
100c0eaf63d5 moved 'Ctr_Sugar' further up the theory hierarchy, so that 'Datatype' can use it
blanchet
parents: 54397
diff changeset
    23
declare [[coercion_args case_guard - + -]]
100c0eaf63d5 moved 'Ctr_Sugar' further up the theory hierarchy, so that 'Datatype' can use it
blanchet
parents: 54397
diff changeset
    24
declare [[coercion_args case_cons - -]]
100c0eaf63d5 moved 'Ctr_Sugar' further up the theory hierarchy, so that 'Datatype' can use it
blanchet
parents: 54397
diff changeset
    25
declare [[coercion_args case_abs -]]
100c0eaf63d5 moved 'Ctr_Sugar' further up the theory hierarchy, so that 'Datatype' can use it
blanchet
parents: 54397
diff changeset
    26
declare [[coercion_args case_elem - +]]
100c0eaf63d5 moved 'Ctr_Sugar' further up the theory hierarchy, so that 'Datatype' can use it
blanchet
parents: 54397
diff changeset
    27
54701
4ed7454aebde tuning -- moved ML files to subdirectory
blanchet
parents: 54626
diff changeset
    28
ML_file "Tools/Ctr_Sugar/case_translation.ML"
54398
100c0eaf63d5 moved 'Ctr_Sugar' further up the theory hierarchy, so that 'Datatype' can use it
blanchet
parents: 54397
diff changeset
    29
setup Case_Translation.setup
100c0eaf63d5 moved 'Ctr_Sugar' further up the theory hierarchy, so that 'Datatype' can use it
blanchet
parents: 54397
diff changeset
    30
49312
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    31
lemma iffI_np: "\<lbrakk>x \<Longrightarrow> \<not> y; \<not> x \<Longrightarrow> y\<rbrakk> \<Longrightarrow> \<not> x \<longleftrightarrow> y"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    32
by (erule iffI) (erule contrapos_pn)
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    33
49486
64cc57c0d0fe generate "expand" property
blanchet
parents: 49312
diff changeset
    34
lemma iff_contradict:
64cc57c0d0fe generate "expand" property
blanchet
parents: 49312
diff changeset
    35
"\<not> P \<Longrightarrow> P \<longleftrightarrow> Q \<Longrightarrow> Q \<Longrightarrow> R"
64cc57c0d0fe generate "expand" property
blanchet
parents: 49312
diff changeset
    36
"\<not> Q \<Longrightarrow> P \<longleftrightarrow> Q \<Longrightarrow> P \<Longrightarrow> R"
64cc57c0d0fe generate "expand" property
blanchet
parents: 49312
diff changeset
    37
by blast+
64cc57c0d0fe generate "expand" property
blanchet
parents: 49312
diff changeset
    38
54701
4ed7454aebde tuning -- moved ML files to subdirectory
blanchet
parents: 54626
diff changeset
    39
ML_file "Tools/Ctr_Sugar/ctr_sugar_util.ML"
4ed7454aebde tuning -- moved ML files to subdirectory
blanchet
parents: 54626
diff changeset
    40
ML_file "Tools/Ctr_Sugar/ctr_sugar_tactics.ML"
4ed7454aebde tuning -- moved ML files to subdirectory
blanchet
parents: 54626
diff changeset
    41
ML_file "Tools/Ctr_Sugar/ctr_sugar_code.ML"
4ed7454aebde tuning -- moved ML files to subdirectory
blanchet
parents: 54626
diff changeset
    42
ML_file "Tools/Ctr_Sugar/ctr_sugar.ML"
49309
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents: 49286
diff changeset
    43
49075
ed769978dc8d rearrange dependencies
blanchet
parents:
diff changeset
    44
end