[m-rev.] For review: (re-send with the diff inlined) adding region-based memory management (RBMM) analysis

Quan Phan Quan.Phan at cs.kuleuven.be
Tue May 15 17:34:07 AEST 2007


Estimated hours taken: 50 (for integrating into MMC 0.13.x).

The analysis divides the heap used by a program into regions and reasons
about their lifetime, based on those data it then detects the points
where instructions for creating and destroying regions can be added.
Currently it just gathers the information and generates region-annotated
Mercury code that can be used for simulation purpose. The analysis works
for single-module programs only.

compiler/smm_data.m
compiler/smm_utils.m
    Data and procedures that are used in RBMM as well as CTGC 

compiler/rbmm_utils.m

compiler/rpta_info.m
compiler/rpt_graph.m
compiler/rpt_alpha_mapping.m
    Define data types used in region points-to analysis.

compiler/rpta_fixpoint_table.m
    Fixpoint computation for region points-to analysis.
compiler/rpta_run.m
    Implementation of the region points-to analysis.

compiler/lra_data.m
    Define data types used in live region analysis.

compiler/lra_lva.m
compiler/lra_rules.m
    Implementation of the live region analysis

compiler/region_transform.m
    Introducing region instructions.

compiler/pp_region_annotated_code.m
    Generate region annotated code for simulating purpose.
The above are new files.

compiler/fixpoint_table.m
    re-added but this file is not used before, to perform fixpoint
    computation.
compiler/mercury_compile.m
    Add calls to the region analysis.
compiler/options.m
    add option --region-analysis.

Quan.

cvs diff: Diffing .
Index: fixpoint_table.m
===================================================================
RCS file: fixpoint_table.m
diff -N fixpoint_table.m
--- /dev/null	1 Jan 1970 00:00:00 -0000
+++ fixpoint_table.m	15 May 2007 04:05:09 -0000
@@ -0,0 +1,217 @@
+%-----------------------------------------------------------------------------%
+% vim: ft=mercury ts=4 sw=4 et
+%-----------------------------------------------------------------------------%
+% Copyright (C) 2000-2001 The University of Melbourne.
+% This file may only be copied under the terms of the GNU General
+% Public License - see the file COPYING in the Mercury distribution.
+%-----------------------------------------------------------------------------%
+% File: fixpoint_table.m
+% Main author: nancy.
+%
+% This modules defines a generic table to be used for fixpoint
computations. 
+% For each key (usually pred_proc_id), it will map a given abstract
+% substitution. Here the notion of abstract substitution is abstracted 
+% away as a typeclass which required only two operations: equal and init.
+%
+%-----------------------------------------------------------------------------%
+
+:- module fixpoint_table.
+:- interface.
+
+:- import_module list.
+
+:- type fixpoint_table(K, E). 
+
+	% Initialise the table.
+	% The first parameter is a list of keys which will be allowed in 
+	% the table. 
+	% fp_init(Initializer, Keys, Table).
+    %
+:- pred fp_init(pred(K, E), list(K), fixpoint_table(K, E)).
+:- mode fp_init(pred(in, out) is det, in, out) is det.
+
+	% Inform the table that a new run has begun.
+    %
+:- pred fp_new_run(fixpoint_table(K, E)::in, fixpoint_table(K, E)::out) 
+        is det.
+
+	% Which run of the fix point are we up to?
+    %
+:- func fp_which_run(fixpoint_table(K, E)) = int.
+
+	% Check whether a fixpoint has been reached.
+    %
+:- pred fp_stable(fixpoint_table(K, E)::in) is semidet.
+
+	% Check whether the entries are recursive.
+    %
+:- pred fp_is_recursive(fixpoint_table(K,E)::in) is semidet.
+
+	% Add a new element (E) associated with key (K) to the table.
+	%   - if an element is already recorded with that key, 
+	%      * and if both elements are equal, then a fixpoint is obtained
+	%        as far as this key is concerned
+	%      * if the elements are different, fixpoint is not reached yet, 
+	%	 and the new information is simply kept
+	%   - if the element was not yet present in the table, add it
+	%     to the table (which does not change the stability of the
+	%     table) 
+	% fp_add( EqualityTest, Key, Element, TableIn, TableOut).
+    %
+:- pred fp_add(pred(E, E), K, E, fixpoint_table(K, E),
fixpoint_table(K, E)).
+:- mode fp_add(pred(in, in) is semidet, in, in, in, out) is det.
+
+	% Retrieve an element (E) associated with key (K) from the table.
+	% This operation will change the state of the table if the
+	% element _is_ present in the table. This means we're facing
+	% a recursive calltree. If the key is not an element of the
+	% allowed keys, then the procedure will fail.
+    %
+:- pred fp_get(K::in, E::out, fixpoint_table(K, E)::in, 
+        fixpoint_table(K, E)::out) is semidet.
+
+	% Retrieve an element (E) associated with key (K) from the table. 
+	% The operation reports an error when the element is not present. 
+    %
+:- pred fp_get_final(K::in, E::out, fixpoint_table(K,E)::in) is det.
+
+	% Same as fp_get_final, but the predicate fails instead
+	% of aborting when the element is not present. 
+    %
+:- pred fp_get_final_semidet(K::in, E::out, fixpoint_table(K,E)::in) 
+        is semidet.
+
+%-----------------------------------------------------------------------------%
+
+:- implementation. 
+
+:- import_module bool, map, list, int.
+:- import_module libs.
+:- import_module libs.compiler_util.
+
+:- import_module string.
+
+:- type fixpoint_table(K, E)
+	--->	ft(  
+		     keys	:: list(K),	% list of allowed keys
+		     run	:: int,		% number of runs
+		     recursive	:: bool,	% is recursive or not
+		     mapping 	:: map(K, fp_entry(E))
+		).
+
+:- type fp_entry(E) 
+	--->	entry(
+			bool, 	% stability: yes = stable, no = unstable
+			E). 
+		   
+:- func fp_entry_init(E) = fp_entry(E).
+:- func fp_entry_stability(fp_entry(E)) = bool. 
+:- func fp_entry_elem(fp_entry(E)) = E. 
+:- func fp_entry_init(bool, E) = fp_entry(E). 
+fp_entry_init(Elem) = entry(no, Elem).  
+fp_entry_init(Bool, Elem) = entry(Bool, Elem). 
+fp_entry_stability(entry(S, _)) = S. 
+fp_entry_elem(entry(_, Elem)) = Elem. 
+
+fp_init(Init, Ks, ft(Ks, Run, IsRecursive, Map)) :- 
+	Run = 0,
+	IsRecursive = no,
+	map.init(Map0),
+	list.foldl(
+		(pred(K::in, M0::in, M::out) is det :- 
+			Init(K, E),
+			map.det_insert(M0, K, fp_entry_init(E), M)
+		),
+		Ks, 
+		Map0, 
+		Map
+	).
+
+fp_new_run(T0, T0 ^ run := T0 ^ run + 1).
+fp_which_run(T0) = T0 ^ run.
+
+fp_is_recursive(T) :- T ^ recursive = yes.
+
+fp_stable(T) :- 
+	(
+		T ^ recursive = yes
+	->
+		map.foldl(
+			pred(_K::in, Entry::in, S0::in, S::out) is det :- 
+			(
+				( S0 = no -> 
+					S = no
+				;
+					S = fp_entry_stability(Entry)
+				)
+			),
+			T ^ mapping,
+			yes, 
+			yes)
+	;
+		true
+	). 
+	
+fp_add(Equal, Index, Elem, Tin, Tout) :- 
+	Map = Tin ^ mapping, 
+	( 
+		map.search(Map, Index, Entry),
+		TabledElem = fp_entry_elem(Entry)
+	->
+		(
+			Equal(TabledElem, Elem)
+		->
+			S = yes
+		;
+			S = no 
+		),
+        % whether or not the tabled element is equal to
+        % the new element, the final tabled element will 
+        % always be set to the new one. This is handy
+        % for performing the following trick: equality
+        % can be checked on some partial piece of the 
+        % elements (for deciding stability), but some other
+        % part might have changed, a change that should 
+        % become visible in the table too. 
+        % (in fact this is necessary for the reuse-fixpoint
+        % table where not only the reuses are kept (the
+        % abstract substitution), but also the goal that
+        % might have changed. 
+		FinalTabledElem = fp_entry_init(S, Elem),
+		map.det_update(Map, Index, FinalTabledElem, MapOut)
+	;
+		% should not occur!
+		map.det_insert(Map, Index, fp_entry_init(Elem), MapOut)
+	),
+	Tout = (Tin ^ mapping := MapOut).
+
+fp_get(Index, Elem, Tin, Tout) :-
+	List = Tin ^ keys, 
+	list.member(Index, List), % can fail
+	Mapin = Tin ^ mapping,
+	(	
+		map.search(Mapin, Index, Entry), 
+		TabledElem = fp_entry_elem(Entry)
+	->
+		Elem = TabledElem,
+		Mapout = Mapin
+	;
+		unexpected(thisfile, "fp_get: key not in map")
+	),
+	Tout = (Tin ^ mapping := Mapout) ^ recursive := yes.
+
+fp_get_final(Index, Elem, T) :- 
+	(
+		fp_get_final_semidet(Index, TabledElem, T)
+	->
+		Elem = TabledElem
+	; 
+		unexpected(thisfile, "fp_get_final: final element not found")
+	).
+
+fp_get_final_semidet(Index, Elem, T):- 
+	map.search(T^mapping, Index, Entry),
+	Elem = fp_entry_elem(Entry). 
+
+:- func thisfile = string.
+thisfile = "fixpoint_table.m".
Index: lra_data.m
===================================================================
RCS file: lra_data.m
diff -N lra_data.m
--- /dev/null	1 Jan 1970 00:00:00 -0000
+++ lra_data.m	15 May 2007 04:10:51 -0000
@@ -0,0 +1,107 @@
+%-----------------------------------------------------------------------------%
+% vim: ft=mercury ts=4 sw=4 et
+%-----------------------------------------------------------------------------%
+% Copyright (C) 2005-2007 The University of Melbourne.
+% This file may only be copied under the terms of the GNU General
+% Public License - see the file COPYING in the Mercury distribution.
+%-----------------------------------------------------------------------------%
+
+% file lra_data.m.
+% main author: Quan Phan
+%
+% Defines the data structures used in the live region analysis.
+
+:- module lra_data.
+
+:- interface.
+
+:- import_module hlds.
+:- import_module hlds.hlds_pred.
+:- import_module hlds.hlds_goal.
+
+:- import_module parse_tree.
+:- import_module parse_tree.prog_data.
+
+:- import_module rpt_graph.
+:- import_module smm_data.
+:- import_module list.
+:- import_module assoc_list.
+:- import_module set.
+:- import_module map.
+
+:- type execution_path == assoc_list(smm_data.program_point, hlds_goal).
+:- type ep_table == map(pred_proc_id, list(execution_path)).
+
+%----------------------------------------------------------------------------%
+%
+% part for program variables
+%
+
+    % to represent a set of program variables
+    %
+:- type variable_set == set(prog_var).
+
+    % to represent the relation between a program point and a set of 
+    % variables. E.g., a map of this type can be the sets of live
+    % variables before program points in a procedure.
+    %
+:- type pp_varset_table == map(smm_data.program_point, variable_set).
+
+    % to represent the relation between a procedure and sets of variables
+    % associated with its program points.
+    %
+:- type proc_pp_varset_table == map(pred_proc_id, pp_varset_table).
+
+%----------------------------------------------------------------------------%
+%
+% part for region/node
+%
+
+    % represent a set of regions, e.g., deadR, bornR, ..., live regions
before
+    % and after a program point.
+    %
+:- type region_set == set(rptg_node).
+
+:- pred region_set_equal(region_set::in, region_set::in) is semidet.
+
+    % to represent the relation between a procedure and its interesting
sets 
+    % of regions, such as deadR, bornR, constantR
+    %
+:- type proc_region_set_table == map(pred_proc_id, region_set).
+
+:- pred proc_region_set_table_equal(proc_region_set_table::in,
+    proc_region_set_table::in) is semidet.
+
+    % to represent the relation between a program point and its region set
+    % e.g., live regions before the pp, live regions after the pp
+    %
+:- type pp_region_set_table == map(smm_data.program_point, region_set).
+
+    % represent the relation between a procedure and sets of regions
+    % associated with its program points
+    %
+:- type proc_pp_region_set_table == map(pred_proc_id, pp_region_set_table).
+
+:- implementation.
+
+region_set_equal(RegionSet1, RegionSet2) :-
+	set.equal(RegionSet1, RegionSet2).
+
+proc_region_set_table_equal(ProcRegionSetTable1, ProcRegionSetTable2)
:-    
+    map.count(ProcRegionSetTable1, C1),
+    map.count(ProcRegionSetTable2, C2),
+    C1 = C2,
+
+    map.keys(ProcRegionSetTable1, PredProcIds1),
+    prst_equal_2(PredProcIds1, ProcRegionSetTable1, ProcRegionSetTable2).
+
+:- pred prst_equal_2(list(pred_proc_id)::in, proc_region_set_table::in,
+    proc_region_set_table::in) is semidet.
+
+prst_equal_2([], _, _).
+prst_equal_2([Pid|Pids], PRST1, PRST2) :-
+    map.search(PRST2, Pid, RS2),
+
+    map.lookup(PRST1, Pid, RS1),
+    set.equal(RS1, RS2),
+    prst_equal_2(Pids, PRST1, PRST2).
Index: lra_lva.m
===================================================================
RCS file: lra_lva.m
diff -N lra_lva.m
--- /dev/null	1 Jan 1970 00:00:00 -0000
+++ lra_lva.m	15 May 2007 04:13:52 -0000
@@ -0,0 +1,1002 @@
+%-----------------------------------------------------------------------------%
+% vim: ft=mercury ts=4 sw=4 et
+%-----------------------------------------------------------------------------%
+% Copyright (C) 2005-2007 The University of Melbourne.
+% This file may only be copied under the terms of the GNU General
+% Public License - see the file COPYING in the Mercury distribution.
+%-----------------------------------------------------------------------------%
+%
+% File lra_lva.m.
+% Main author: Quan Phan.
+%
+% This module implements the live region analysis as described in the 
+% technical report. It is composed of:
+% 1. collecting all execution paths (EP) of procedures,
+% 2. based on those EPs collecting live variables at each program point,
+% 3. based on those EPs with live variables collecting live regions at 
+% each program point.
+% 4. updating the initial bornR and deadR, then computing constantR for
each 
+% procedure.
+
+:- module lra_lva.
+
+:- interface.
+
+:- import_module hlds.
+:- import_module hlds.hlds_module.
+
+:- import_module io. 
+
+:- import_module rpta_info.
+:- import_module lra_data.
+
+    % this collects EPs for each procedure.
+    %
+:- pred execution_path_analysis(module_info::in, ep_table::out, io::di, 
+    io::uo) is det.
+
+	% this collects live variable sets.
+    %
+:- pred live_variable_analysis(module_info::in, ep_table::in, 
+    proc_pp_varset_table::out, proc_pp_varset_table::out, 
+    proc_pp_varset_table::out, io::di, io::uo) is det.
+
+	% this collects live region sets.
+    %
+:- pred live_region_analysis(module_info::in, rpta_info_table::in, 
+    proc_pp_varset_table::in, proc_pp_varset_table::in, 
+    proc_pp_varset_table::in, proc_pp_region_set_table::out, 
+    proc_pp_region_set_table::out, proc_pp_region_set_table::out, 
+    proc_region_set_table::out, proc_region_set_table::out, 
+    proc_region_set_table::out, proc_region_set_table::out, 
+    proc_region_set_table::out, io::di, io::uo) is det.
+
+    % this predicate reasons about lifetime of regions across procedure
+    % boundary. It will update the initial deadR and bornR sets and compute
+    % constantR set.
+    %
+:- pred compute_interproc_region_lifetime(module_info::in, 
+    rpta_info_table::in, ep_table::in, proc_pp_region_set_table::in,
+    proc_pp_region_set_table::in, proc_region_set_table::in,
+    proc_region_set_table::in, proc_region_set_table::out,
+    proc_region_set_table::in, proc_region_set_table::out,
+    proc_region_set_table::in, proc_region_set_table::out,
+    io::di, io::uo) is det.
+
+    % this predicate removes regions of primitive types from the input
data 
+    % structures. 
+    % The reason for this is the assumption that primitive values are not 
+    % boxed (i.e., not store in regions).
+    %
+:- pred ignore_primitive_regions(module_info::in, rpta_info_table::in,
+    proc_region_set_table::in, proc_region_set_table::out,
+    proc_region_set_table::in, proc_region_set_table::out,
+    proc_region_set_table::in, proc_region_set_table::out,
+    proc_region_set_table::in, proc_region_set_table::out,
+    proc_pp_region_set_table::in, proc_pp_region_set_table::out,
+    proc_pp_region_set_table::in, proc_pp_region_set_table::out,
+    proc_pp_region_set_table::in, proc_pp_region_set_table::out) is det.
+%-----------------------------------------------------------------------------%
+
+:- implementation. 
+
+:- import_module hlds.hlds_pred. 
+:- import_module hlds.hlds_goal.
+:- import_module hlds.arg_info.
+
+:- import_module parse_tree.
+:- import_module parse_tree.prog_data.
+:- import_module parse_tree.mercury_to_mercury.
+:- import_module check_hlds.
+:- import_module check_hlds.goal_path.
+:- import_module check_hlds.type_util. 
+:- import_module libs.
+:- import_module libs.compiler_util.
+:- import_module smm_data.
+:- import_module smm_utils.
+:- import_module rpt_graph.
+:- import_module lra_rules.
+
+:- import_module bool.
+:- import_module varset.
+:- import_module map.
+:- import_module list.
+:- import_module assoc_list.
+:- import_module set.
+:- import_module string.
+:- import_module term.
+:- import_module pair.
+:- import_module svmap.
+:- import_module svset.
+
+%-----------------------------------------------------------------------------%
+%
+% part for execution path analysis
+%
+
+execution_path_analysis(HLDS, EPTable, !IO) :-
+	module_info_predids(PredIds, HLDS, _),
+	map.init(EPTable0),
+	list.foldl2(execution_path_analysis_pred(HLDS), PredIds, EPTable0, 
+        EPTable, !IO).
+
+:- pred execution_path_analysis_pred(module_info::in, pred_id::in, 
+    ep_table::in, ep_table::out, io::di, io::uo) is det.
+    
+execution_path_analysis_pred(HLDS, PredId, !EPTable, !IO) :-
+	module_info_pred_info(HLDS, PredId, PredInfo),
+	pred_info_non_imported_procids(PredInfo) = ProcIds,
+	list.foldl2(execution_path_analysis_proc(HLDS, PredId), ProcIds, 
+        !EPTable, !IO).
+
+:- pred execution_path_analysis_proc(module_info::in, pred_id::in, 
+    proc_id::in, ep_table::in, ep_table::out, io::di, io::uo) is det.
+
+execution_path_analysis_proc(HLDS, PredId, ProcId, EPTable0, EPTable,
!IO) :-
+	PredProcId = proc(PredId, ProcId),
+	( if
+		smm_utils.some_are_special_preds([PredProcId], HLDS)
+	  then
+		EPTable = EPTable0
+	  else
+		module_info_pred_proc_info(HLDS, PredProcId, _PredInfo, ProcInfo),
+		execution_paths(ProcInfo, HLDS, EPs),
+		
+		map.set(EPTable0, PredProcId, EPs, EPTable)
+	).
+
+:- pred execution_paths(proc_info::in, module_info::in, 
+    list(execution_path)::out) is det.
+
+execution_paths(ProcInfo, HLDS, EPs) :-
+    % to fill the goals with program point information
+	fill_goal_path_slots(HLDS, ProcInfo, ProcInfo1),
+
+	proc_info_get_goal(ProcInfo1, Goal),
+	execution_paths_covered_goal(ProcInfo, Goal, [[]], EPs).
+	
+	% construct execution paths that cover (are up to and include) this goal.
+    %
+:- pred execution_paths_covered_goal(proc_info::in, hlds_goal::in, 
+    list(execution_path)::in, list(execution_path)::out) is det.
+
+execution_paths_covered_goal(ProcInfo, Goal, EPsUntilGoal,
EPsCoveredGoal) :- 
+	Goal = hlds_goal(Expr, Info),
+	(
+		goal_is_atomic(Expr)
+	->
+		(
+			(
+				Expr = unify(_, _, _, _, _) 
+			;	
+				Expr = plain_call(_, _, _, _, _, _) 
+			; 
+				Expr = conj(_ConjType, [])
+			;
+				Expr = disj([])
+			)
+		->
+            % retrieve the program point of this goal
+			program_point_init(Info, PP),
+			append_to_each_ep(EPsUntilGoal, [[pair(PP, Goal)]], EPsCoveredGoal)
+		;
+            % XXX: other kinds of atomic calls (generic_call, 
+            % foreign_proc), TEMPORARILY ignored their corresponding pps.
+			append_to_each_ep(EPsUntilGoal, [[]], EPsCoveredGoal)
+		)
+	;
+		execution_paths_covered_compound_goal(ProcInfo, Goal, 
+            EPsUntilGoal, EPsCoveredGoal)
+	).
+
+	% extend current eps to cover this compound goal.
+    %
+:- pred execution_paths_covered_compound_goal(proc_info::in,
hlds_goal::in, 
+    list(execution_path)::in, list(execution_path)::out) is det.
+
+execution_paths_covered_compound_goal(ProcInfo, CompoundGoal,
EPsBeforeGoal, 
+    EPsCoveredGoal) :- 
+	CompoundGoal = hlds_goal(Expr, _Info),
+	(
+		Expr = conj(_ConjType, Goals)
+	->
+		execution_paths_covered_conj(ProcInfo, Goals, 
+            EPsBeforeGoal, EPsCoveredGoal)
+	;
+		Expr = switch(_A, _B, Cases)
+	->
+		execution_paths_covered_cases(ProcInfo, CompoundGoal, Cases, 
+            EPsBeforeGoal, EPsCoveredGoal)
+	;
+		Expr = disj(Goals)
+	->
+		execution_paths_covered_disj(ProcInfo, Goals, 
+            EPsBeforeGoal, EPsCoveredGoal)
+	;
+		Expr = negation(Goal)
+	->
+		execution_paths_covered_goal(ProcInfo, Goal, 
+            EPsBeforeGoal, EPsCoveredGoal)
+	;
+		Expr = scope(_ScopeReason, Goal)
+	->
+		execution_paths_covered_goal(ProcInfo, Goal, 
+            EPsBeforeGoal, EPsCoveredGoal)
+	;
+		Expr = if_then_else(_V, Cond, Then, Else)
+	->		
+		execution_paths_covered_goal(ProcInfo, Cond, 
+            EPsBeforeGoal, EPsCoveredCond),
+		execution_paths_covered_goal(ProcInfo, Then, 
+            EPsCoveredCond, EPsCoveredThen),
+		execution_paths_covered_goal(ProcInfo, Else, 
+            EPsBeforeGoal, EPsCoveredElse),
+			% In the operational semantics of Mercury (MMC), 
+            % if g1 then g2 else g3 is equivalent to 
+            % ((g1,g2);(g3))
+			%
+		list.append(EPsCoveredThen, EPsCoveredElse, EPsCoveredGoal)
+	;
+		unexpected(thisfile, "collect_execution_path_in_compound_goal: " 
+            ++ "found atomic call")
+	). 
+
+	% the idea is to extend all the current execution paths with all the
goals 
+    % in this conjunction.
+    %
+:- pred execution_paths_covered_conj(proc_info::in, list(hlds_goal)::in, 
+    list(execution_path)::in, list(execution_path)::out) is det.
+
+execution_paths_covered_conj(_ProcInfo, [], EPsBeforeGoal, EPsBeforeGoal).
+execution_paths_covered_conj(ProcInfo, [Goal|Others], 
+    EPsBeforeGoal, EPsCoveredGoal) :-
+    % extend all current EPs to this Goal
+	execution_paths_covered_goal(ProcInfo, Goal, EPsBeforeGoal, 
+        EPsCoveredThisGoal),
+	
+    % extend them with to the other goals
+	execution_paths_covered_conj(ProcInfo, Others, EPsCoveredThisGoal, 
+        EPsCoveredGoal).
+
+	% the idea is to extend all the current ep with each branch of this 
+    % disjunction. 
+	% E.g. if before this disj we have 2 eps, this disj has 3 branches, 
+    % then after this extension we will have 6 eps: 2 extended to the 1st 
+    % branch, 2 to the 2nd, 2 to the 3rd.
+    %
+:- pred execution_paths_covered_disj(proc_info::in, list(hlds_goal)::in, 
+    list(execution_path)::in, list(execution_path)::out) is det.
+
+execution_paths_covered_disj(_ProcInfo, [], _EPsBeforeGoal, []).
+execution_paths_covered_disj(ProcInfo, [Goal|Others], EPsBeforeGoal, 
+    EPsCoveredGoal) :- 
+    % extend all current EPs to this goal
+	execution_paths_covered_goal(ProcInfo, Goal, EPsBeforeGoal, 
+        EPsCoveredThisGoal),
+	
+    % extend all current EPs to each of the other goals (branches)
+	execution_paths_covered_disj(ProcInfo, Others, EPsBeforeGoal, 
+        EPsCoveredOtherGoals),
+	
+    % they are all EPs that cover this disjunction
+	list.append(EPsCoveredThisGoal, EPsCoveredOtherGoals, EPsCoveredGoal).
+	
+    % the next pp, if any, will be a first-shared-point.
+	% This is for the situation where MMC detects a disj is actually a 
+    % switch (semidet or det disj). 
+	% One noticeable issue here is that MMC removes the unification of the 
+    % switched-on variable with a constant and/or a functor of arity 0
(f/0). 
+    % Those removals cause eps to have fewer pps than we want. The
solution 
+    % here is to add an invented program point for the removed
unification. 
+    % The corresponding goal of the pp is the switch goal itself so
that we 
+    % can get information about the switched-on variable in live variable 
+    % analysis.
+	% Other than the above point, the algorithm is the same as a normal disj.
+    %
+:- pred execution_paths_covered_cases(proc_info::in, hlds_goal::in, 
+    list(case)::in, list(execution_path)::in, list(execution_path)::out) 
+    is det.
+
+execution_paths_covered_cases(_ProcInfo, _Switch, [], _EPsBeforeGoal, []).
+execution_paths_covered_cases(ProcInfo, Switch, [Case|Others], 
+    EPsBeforeGoal, EPsCoveredGoal) :-
+    % deal with the removed unification, if any, and we add a dummy
program 
+    % point for the removed one because the MMC removes it
+	Case = case(ConsId, Goal),
+	Switch = hlds_goal(_SwitchExpr, Info),
+	smm_data.program_point_init(Info, PP),
+
+	(
+		ConsId = cons(_SymName, Arity)
+	->
+		( if
+			Arity = 0
+		  then
+			append_to_each_ep(EPsBeforeGoal, [[pair(PP, Switch)]], 
+                EPsBeforeGoal1)
+		  else
+			EPsBeforeGoal1 = EPsBeforeGoal
+		)
+	; 
+		ConsId = int_const(_Int)
+	->
+		% need to add a dummy pp
+		append_to_each_ep(EPsBeforeGoal, [[pair(PP, Switch)]], EPsBeforeGoal1)
+	;
+		ConsId = string_const(_String)
+	->
+		% need to add a pp
+		append_to_each_ep(EPsBeforeGoal, [[pair(PP, Switch)]], EPsBeforeGoal1)
+	;
+		ConsId = float_const(_Float)
+	->
+		% need to add a pp
+		append_to_each_ep(EPsBeforeGoal, [[pair(PP, Switch)]], EPsBeforeGoal1)
+	;
+		unexpected(thisfile, "execution_paths_covered_cases: new cons_id "
+            ++ "encountered")
+	),
+    % as disj
+	execution_paths_covered_goal(ProcInfo, Goal, 
+        EPsBeforeGoal1, EPsCoveredThisGoal),
+	execution_paths_covered_cases(ProcInfo, Switch, Others, 
+        EPsBeforeGoal, EPsCoveredOtherGoals),
+	list.append(EPsCoveredThisGoal, EPsCoveredOtherGoals, EPsCoveredGoal).
+
+	% extend each execution path in the first list with each in the 
+    % second list, all the extended execution paths are put in the
third list
+    %
+:- pred append_to_each_ep(list(execution_path)::in,
list(execution_path)::in, 
+    list(execution_path)::out) is det.
+
+append_to_each_ep([], _ToBeAppended, []).
+append_to_each_ep([EP|Others], ToBeAppended, ExtendedEPList) :-
+    % collect all the extensions of this EP
+	extend_ep(ToBeAppended, EP, ExtendedEPs),
+	
+    % collect the extensions of the others
+	append_to_each_ep(Others, ToBeAppended, OtherExtendedEPs),
+	
+    % put them all in the result
+	list.append(ExtendedEPs, OtherExtendedEPs, ExtendedEPList).
+
+	% extend the execution path with each in the list,
+	% all the extensions are returned.
+    %
+:- pred extend_ep(list(execution_path)::in, execution_path::in, 
+    list(execution_path)::out) is det.
+
+extend_ep([], _EP, []).
+extend_ep([Extension|Others], EP, ExtendedEPs) :-
+	append(EP, Extension, ExtendedEP),
+	extend_ep(Others, EP, OtherExtendedEPs),
+	append([ExtendedEP], OtherExtendedEPs, ExtendedEPs).
+
+%----------------------------------------------------------------------------%
+%
+% part for live variable analysis
+%
+% This analysis computes, for each procedure, the sets of live variables 
+% before and after each of its program points.
+% Currently, it also calculates set of void variables (i.e., ones whose
names
+% start with "_") after each program point. Those variables are considered 
+% dead at that point.
+%
+
+live_variable_analysis(HLDS, EPTable, LVBeforeTable, LVAfterTable, 
+    VoidVarTable, !IO) :-
+	module_info_predids(PredIds, HLDS, _HLDS),
+	map.init(LVBeforeTable0),
+	map.init(LVAfterTable0),
+	map.init(VoidVarTable0),
+	list.foldl4(live_variable_analysis_pred(HLDS, EPTable), PredIds, 
+		LVBeforeTable0, LVBeforeTable, LVAfterTable0, LVAfterTable, 
+        VoidVarTable0, VoidVarTable, !IO).
+
+:- pred live_variable_analysis_pred(module_info::in, ep_table::in, 
+    pred_id::in, proc_pp_varset_table::in, proc_pp_varset_table::out,
+    proc_pp_varset_table::in, proc_pp_varset_table::out,
+	proc_pp_varset_table::in, proc_pp_varset_table::out,
+	io::di, io::uo) is det.
+
+live_variable_analysis_pred(HLDS, EPTable, PredId, !LVBeforeTable, 
+    !LVAfterTable, !VoidVarTable, !IO) :-
+	module_info_pred_info(HLDS, PredId, PredInfo),
+	pred_info_non_imported_procids(PredInfo) = ProcIds,
+	list.foldl4(live_variable_analysis_proc(HLDS, EPTable, PredId), 
+		ProcIds, !LVBeforeTable, !LVAfterTable, !VoidVarTable, !IO).
+
+:- pred live_variable_analysis_proc(module_info::in, ep_table::in, 
+    pred_id::in, proc_id::in, proc_pp_varset_table::in, 
+    proc_pp_varset_table::out, proc_pp_varset_table::in, 
+    proc_pp_varset_table::out, proc_pp_varset_table::in, 
+    proc_pp_varset_table::out, io::di, io::uo) is det.
+
+live_variable_analysis_proc(HLDS, EPTable, PredId, ProcId, !LVBeforeTable, 
+    !LVAfterTable, !VoidVarTable, !IO) :-	
+	PredProcId = proc(PredId, ProcId),
+	( if 
+		smm_utils.some_are_special_preds([PredProcId], HLDS)
+	  then
+        true 
+	  else
+		module_info_pred_proc_info(HLDS, PredProcId, _PredInfo, ProcInfo),
+		find_input_output_params(HLDS, ProcInfo, Inputs, Outputs),
+		map.lookup(EPTable, PredProcId, EPs),
+		live_variable_analysis_eps(EPs, Inputs, Outputs, HLDS, ProcInfo,
+			map.init, ProcLVBefore, map.init, ProcLVAfter, map.init, 
+                ProcVoidVar, !IO),
+		
+		map.set(!.LVBeforeTable, PredProcId, ProcLVBefore, !:LVBeforeTable),
+		map.set(!.LVAfterTable, PredProcId, ProcLVAfter, !:LVAfterTable),
+		map.set(!.VoidVarTable, PredProcId, ProcVoidVar, !:VoidVarTable)
+	).
+
+:- pred live_variable_analysis_eps(list(execution_path)::in, 
+    list(prog_var)::in, list(prog_var)::in, module_info::in,
proc_info::in, 
+    pp_varset_table::in, pp_varset_table::out, pp_varset_table::in, 
+    pp_varset_table::out, pp_varset_table::in, pp_varset_table::out, 
+    io::di, io::uo) is det.
+
+    % Live variable analysis is backward, so we reverse the execution path
+    % before starting
+    %
+live_variable_analysis_eps([], _Inputs, _Outputs, _HLDS, _ProcInfo, 
+    !ProcLVBefore, !ProcLVAfter, !ProcVoidVar, !IO).
+live_variable_analysis_eps([EP0|OtherEPs], Inputs, Outputs, HLDS,
ProcInfo, 
+    !ProcLVBefore, !ProcLVAfter, !ProcVoidVar, !IO) :-
+	list.reverse(EP0, EP1),
+	( if
+		list.length(EP1) = 1
+	  then
+        % i.e. there's only 1 pp in this EP, so the last pp is also 
+        % the first pp
+	  	live_variable_analysis_singleton_ep(EP1, Inputs, Outputs, HLDS, 
+            ProcInfo, !ProcLVBefore, !ProcLVAfter, !ProcVoidVar, !IO)
+	  else
+        % start with the last pp
+		live_variable_analysis_ep(EP1, Inputs, Outputs, HLDS, ProcInfo, 
+			yes, set.init, !ProcLVBefore, !ProcLVAfter, !ProcVoidVar, !IO)
+	),
+	live_variable_analysis_eps(OtherEPs, Inputs, Outputs, HLDS, ProcInfo, 
+		!ProcLVBefore, !ProcLVAfter, !ProcVoidVar, !IO).
+
+:- pred live_variable_analysis_ep(execution_path::in, list(prog_var)::in, 
+    list(prog_var)::in, module_info::in, proc_info::in, bool::in, 
+    set(prog_var)::in, pp_varset_table::in, pp_varset_table::out, 
+    pp_varset_table::in, pp_varset_table::out, pp_varset_table::in, 
+    pp_varset_table::out, io::di, io::uo) is det.
+
+live_variable_analysis_ep([], _Inputs, _Outputs, _HLDS, _ProcInfo,_Last, 
+	_LVBeforeNext, !ProcLVBefore, !ProcLVAfter, !ProcVoidVar, !IO).
+	% process the last program point in an execution path
+    %
+live_variable_analysis_ep([(PP - Goal)|OtherPPGoals], Inputs, Outputs,
HLDS, 
+    ProcInfo, yes, _LVBeforeNext, !ProcLVBefore, ProcLVAfter0, 
+    ProcLVAfter, !ProcVoidVar, !IO) :-
+    % live variable after of the last pp is ALWAYS set of output variables
+	( if
+		map.search(ProcLVAfter0, PP, LastAfter)
+	  then
+		ProcLVAfter1 = ProcLVAfter0,
+		LVAfter = LastAfter
+	  else
+		LVAfter = set.list_to_set(Outputs),
+		map.set(ProcLVAfter0, PP, LVAfter, ProcLVAfter1)
+	),
+
+    % live variable before this last pp
+	compute_useds_produceds(HLDS, Goal, UsedSet, ProducedSet),
+	set.union(set.difference(LVAfter, ProducedSet), UsedSet,
LVBeforeInThisEP),
+	record_live_vars_at_pp(PP, LVBeforeInThisEP, !ProcLVBefore),
+
+    % collect void variables after this pp
+	collect_void_vars(PP, ProducedSet, ProcInfo, !ProcVoidVar),
+
+    % continue analysing this ep
+	live_variable_analysis_ep(OtherPPGoals, Inputs, Outputs, HLDS, ProcInfo, 
+        no, LVBeforeInThisEP, !ProcLVBefore, ProcLVAfter1, 
+        ProcLVAfter, !ProcVoidVar, !IO).
+
+	% process a middle program point
+    %
+live_variable_analysis_ep([(PP - Goal), Another|OtherPPGoals], Inputs, 
+    Outputs, HLDS, ProcInfo, no, LVBeforeNext, !ProcLVBefore,
!ProcLVAfter, 
+    !ProcVoidVar, !IO) :-
+    % this pp's lv after is the union of this pp's lv after in the 
+    % other eps 	
+	record_live_vars_at_pp(PP, LVBeforeNext, !ProcLVAfter),
+
+    % lv before this pp
+	compute_useds_produceds(HLDS, Goal, UsedSet, ProducedSet),
+	set.union(set.difference(LVBeforeNext, ProducedSet), UsedSet, 
+        LVBeforeInThisEP),
+	record_live_vars_at_pp(PP, LVBeforeInThisEP, !ProcLVBefore),
+
+    % collect void variables after this pp
+	collect_void_vars(PP, ProducedSet, ProcInfo, !ProcVoidVar),
+
+	live_variable_analysis_ep([Another|OtherPPGoals], Inputs, Outputs, HLDS, 
+        ProcInfo, no, LVBeforeInThisEP, !ProcLVBefore, !ProcLVAfter, 
+        !ProcVoidVar, !IO).
+	
+	% the first program point
+    %
+live_variable_analysis_ep([PP - Goal], Inputs, _Outputs, HLDS,
ProcInfo, no, 
+    LVBeforeNext, ProcLVBefore0, ProcLVBefore, !ProcLVAfter, !ProcVoidVar, 
+    !IO) :-
+    % lv before the first pp is ALWAYS inputs
+	( if
+		map.search(ProcLVBefore0, PP, _FirstBefore)
+	  then
+		ProcLVBefore = ProcLVBefore0
+	  else
+		LVBefore = set.list_to_set(Inputs),
+		map.set(ProcLVBefore0, PP, LVBefore, ProcLVBefore)
+	),
+
+    % update lv after this first pp
+	record_live_vars_at_pp(PP, LVBeforeNext, !ProcLVAfter),
+
+    % collect void vars after this pp
+	compute_useds_produceds(HLDS, Goal, _UsedSet, ProducedSet),
+	collect_void_vars(PP, ProducedSet, ProcInfo, !ProcVoidVar).
+
+:- pred live_variable_analysis_singleton_ep(execution_path::in, 
+    list(prog_var)::in, list(prog_var)::in, module_info::in, proc_info::in,
+    pp_varset_table::in, pp_varset_table::out, pp_varset_table::in, 
+    pp_varset_table::out, pp_varset_table::in, pp_varset_table::out, 
+    io::di, io::uo) is det.
+live_variable_analysis_singleton_ep([PP - Goal|_Others], Inputs, Outputs, 
+    HLDS, ProcInfo, !ProcLVBefore, !ProcLVAfter, !ProcVoidVar, !IO) :-
+	LVBefore = set.list_to_set(Inputs),
+	map.set(!.ProcLVBefore, PP, LVBefore, !:ProcLVBefore),
+	LVAfter = set.list_to_set(Outputs),
+	map.set(!.ProcLVAfter, PP, LVAfter, !:ProcLVAfter),
+        
+    % collect void vars after this pp
+	compute_useds_produceds(HLDS, Goal, _UsedSet, ProducedSet),
+	collect_void_vars(PP, ProducedSet, ProcInfo, !ProcVoidVar).
+	
+live_variable_analysis_singleton_ep([], _Inputs, _Outputs, _HLDS,
_ProcInfo,
+			!ProcLVBefore, !ProcLVAfter, !ProcVoidVar, !IO) :-
+	unexpected(thisfile, "live_variable_analysis_singleton_ep: impossible").
+		
+    % A live variable is live at a program point if it is live in one of 
+    % the execution paths that cover it.
+    % Therefore for a program point we need to union the existing live 
+    % variable set with the newly found.
+    %
+:- pred record_live_vars_at_pp(program_point::in, variable_set::in, 
+    pp_varset_table::in, pp_varset_table::out) is det.
+record_live_vars_at_pp(PP, LV, ProcLV0, ProcLV) :- 
+	( if
+		map.search(ProcLV0, PP, ExistingLV)
+	  then
+		map.set(ProcLV0, PP, set.union(ExistingLV, LV), ProcLV)
+	  else
+		map.set(ProcLV0, PP, LV, ProcLV)
+	).
+
+:- pred compute_useds_produceds(module_info::in, hlds_goal::in, 
+    variable_set::out, variable_set::out) is det.
+compute_useds_produceds(HLDS, Goal, UsedSet, ProducedSet) :-
+	( if
+        % a removed switch
+        Goal = hlds_goal(switch(Var, _B, _Cases), _SwitchInfo) 	
+	  then
+		Useds = [Var],
+		Produceds = []
+	  else
+		Goal = hlds_goal(Expr, _Info),
+		(
+			Expr = plain_call(CalleePredId, CalleeProcId, Args, 
+                _BuiltIn, _Context, _Name)
+		->
+			actual_inputs_outputs_proc(Args, 
+				proc(CalleePredId, CalleeProcId), HLDS, Useds, Produceds)
+		;
+			Expr = unify(_, _, _, Unification, _)		
+		->
+			actual_inputs_outputs_unification(Unification, Useds, 
+                Produceds)
+		;
+			(Expr = conj(_, []) ; Expr = disj([]))
+		->
+			Useds = [],
+			Produceds = []
+		;
+			unexpected(thisfile, "compute_useds_produceds: the expression "
+                ++ "must be either call, unify, true, or fail")
+		)
+	),
+	set.list_to_set(Useds, UsedSet),
+	set.list_to_set(Produceds, ProducedSet).
+
+	% find Useds and Produceds of unification.
+    %
+:- pred actual_inputs_outputs_unification(unification::in, 
+    list(prog_var)::out, list(prog_var)::out) is det.
+actual_inputs_outputs_unification(construct(LVar, _, Args, _, _, _, _), 
+    Args, [LVar]).
+actual_inputs_outputs_unification(deconstruct(LVar, _, Args, _, _, _), 
+    [LVar], Args).
+actual_inputs_outputs_unification(assign(LVar, RVar), [RVar], [LVar]). 
+actual_inputs_outputs_unification(simple_test(LVar, RVar), [LVar,
RVar], []).
+actual_inputs_outputs_unification(complicated_unify(_, _, _), [], []).
+
+	% find Used and Produceds of called procedure.
+    %
+:- pred actual_inputs_outputs_proc(list(prog_var)::in, pred_proc_id::in, 
+    module_info::in, list(prog_var)::out, list(prog_var)::out) is det.
+actual_inputs_outputs_proc(ActualArgs, CalleeId, HLDS, ActualInputs, 
+    ActualOutputs) :-
+	module_info_pred_proc_info(HLDS, CalleeId, _PredInfo, CalleeInfo),
+	find_input_output_params(HLDS, CalleeInfo, Inputs, Outputs),
+
+	proc_info_get_headvars(CalleeInfo, FormalParams),
+	actual_inputs_outputs(FormalParams, ActualArgs, Inputs, Outputs, 
+        [], ActualInputs, [], ActualOutputs).
+
+:- pred actual_inputs_outputs(list(prog_var)::in, list(prog_var)::in, 
+    list(prog_var)::in, list(prog_var)::in, list(prog_var)::in, 
+    list(prog_var)::out, list(prog_var)::in, list(prog_var)::out) is det.
+actual_inputs_outputs([], [], _Inputs, _Outputs, 
+    ActualInputs, ActualInputs, ActualOutputs, ActualOutputs).
+actual_inputs_outputs([], [_|_], _Inputs, _Outputs, 
+    ActualInputs, ActualInputs, ActualOutputs, ActualOutputs) :-
+	unexpected(thisfile, 
+        "actual_inputs_outputs: more actual parameters than formal ones").
+actual_inputs_outputs([_|_], [], _Inputs, _Outputs, 
+    ActualInputs, ActualInputs, ActualOutputs, ActualOutputs) :-
+	unexpected(thisfile, 
+        "actual_inputs_outputs: more formal parameters that actual ones").
+actual_inputs_outputs([FormalParam|OtherFormalParams], 
+    [ActualArg|OtherActualArgs], Inputs, Outputs, 
+    ActualInputs0, ActualInputs, ActualOutputs0, ActualOutputs) :-
+	(
+		list.member(FormalParam, Outputs)
+	->
+        % this formal param is an output, so the corresponding arg is 
+        % an actual output arg
+		ActualOutputs1 = [ActualArg|ActualOutputs0],
+		ActualInputs1 = ActualInputs0
+	;
+		list.member(FormalParam, Inputs)
+	->
+        % this formal param is an input,
+		ActualInputs1 = [ActualArg|ActualInputs0],
+		ActualOutputs1 = ActualOutputs0
+	;
+        % this formal param is neither an output nor an input, so ignore 
+        % the corresponding arg.
+		ActualOutputs1 = ActualOutputs0,
+		ActualInputs1 = ActualInputs0
+	),
+    % continue with the others
+	actual_inputs_outputs(OtherFormalParams, OtherActualArgs, Inputs,
Outputs, 
+        ActualInputs1, ActualInputs, ActualOutputs1, ActualOutputs).
+
+	% collect variables whose names start with _, i.e., void variables.
+	% I am considering those vars dead right after created in the live 
+    % variable and region analyses.
+    %
+:- pred collect_void_vars(program_point::in, variable_set::in,
proc_info::in,
+    pp_varset_table::in, pp_varset_table::out) is det.
+collect_void_vars(PP, ProducedSet, ProcInfo, !ProcVoidVar) :-
+	( if
+		map.search(!.ProcVoidVar, PP, _DeadVars)
+	  then
+        true
+	  else
+		proc_info_get_varset(ProcInfo, Varset),
+		set.fold(void_var(Varset), ProducedSet, set.init, VoidVars),
+		map.set(!.ProcVoidVar, PP, VoidVars, !:ProcVoidVar)
+	).
+
+    % to be used with the fold above, 
+    % if Var is a void variable, add it to VoidVars set
+    %
+:- pred void_var(prog_varset::in, prog_var::in, variable_set::in, 
+    variable_set::out) is det.
+void_var(Varset, Var, !VoidVars) :-
+	mercury_var_to_string(Varset, no, Var) = VarName,
+	string.substring(VarName, 0, 1, FirstChar),
+	( if
+		FirstChar = "_"
+	  then
+		set.insert(!.VoidVars, Var, !:VoidVars)
+	  else
+        true 
+	).
+
+%----------------------------------------------------------------------------%
+%
+% part for live region analysis
+%
+% This analysis "converts" the variable-based result of live variable 
+% analysis into region-based information. 
+% It also computes the inputR, outputR, localR, and the initial bornR and 
+% deadR for a procedure.
+%
+
+live_region_analysis(HLDS, RptaInfoTable, LVBeforeTable, LVAfterTable, 
+    VoidVarTable, LRBeforeTable, LRAfterTable, VoidVarRegionTable,
+	InputRTable, OutputRTable, BornRTable, DeadRTable, LocalRTable, !IO) :-
+	module_info_predids(PredIds, HLDS, _),
+	map.init(LRBeforeTable0),
+	map.init(LRAfterTable0),
+	map.init(VoidVarRegionTable0),
+	map.init(InputRTable0),
+	map.init(OutputRTable0),
+	map.init(BornRTable0),
+	map.init(DeadRTable0),
+	map.init(LocalRTable0),
+	foldl9(live_region_analysis_pred(HLDS, RptaInfoTable, LVBeforeTable, 
+        LVAfterTable, VoidVarTable), 
+		PredIds, 
+		LRBeforeTable0, LRBeforeTable, 
+		LRAfterTable0, LRAfterTable,
+		VoidVarRegionTable0, VoidVarRegionTable,  
+		InputRTable0, InputRTable,
+		OutputRTable0, OutputRTable, 
+		BornRTable0, BornRTable, 
+		DeadRTable0, DeadRTable, 
+		LocalRTable0, LocalRTable, !IO).
+
+:- pred live_region_analysis_pred(module_info::in, rpta_info_table::in, 
+    proc_pp_varset_table::in, proc_pp_varset_table::in, 
+	proc_pp_varset_table::in, pred_id::in, 
+	proc_pp_region_set_table::in, proc_pp_region_set_table::out,
+	proc_pp_region_set_table::in, proc_pp_region_set_table::out,
+	proc_pp_region_set_table::in, proc_pp_region_set_table::out,
+	proc_region_set_table::in, proc_region_set_table::out, 
+	proc_region_set_table::in, proc_region_set_table::out, 
+	proc_region_set_table::in, proc_region_set_table::out, 
+	proc_region_set_table::in, proc_region_set_table::out, 
+	proc_region_set_table::in, proc_region_set_table::out, 
+	io::di, io::uo) is det.
+live_region_analysis_pred(HLDS, RptaInfoTable, LVBeforeTable,
LVAfterTable, 
+    VoidVarTable, PredId, 
+    !LRBeforeTable, !LRAfterTable, !VoidVarRegionTable, !InputRTable,
+	!OutputRTable, !BornRTable, !DeadRTable, !LocalRTable, !IO) :-
+	module_info_pred_info(HLDS, PredId, PredInfo),
+    % this statement eliminates quite a lot compiler-generated predicates 
+    % (it returns an empty list for those predicates)
+	pred_info_non_imported_procids(PredInfo) = ProcIds,
+
+	foldl9(live_region_analysis_proc(HLDS, RptaInfoTable, LVBeforeTable, 
+        LVAfterTable, VoidVarTable, PredId), ProcIds, 
+		!LRBeforeTable, !LRAfterTable, !VoidVarRegionTable, !InputRTable,
+		!OutputRTable, !BornRTable, !DeadRTable, !LocalRTable, !IO).
+	
+:- pred live_region_analysis_proc(module_info::in, rpta_info_table::in, 
+    proc_pp_varset_table::in, proc_pp_varset_table::in, 
+	proc_pp_varset_table::in, pred_id::in, proc_id::in, 
+	proc_pp_region_set_table::in, proc_pp_region_set_table::out,
+	proc_pp_region_set_table::in, proc_pp_region_set_table::out,
+	proc_pp_region_set_table::in, proc_pp_region_set_table::out,
+	proc_region_set_table::in, proc_region_set_table::out, 
+	proc_region_set_table::in, proc_region_set_table::out, 
+	proc_region_set_table::in, proc_region_set_table::out, 
+	proc_region_set_table::in, proc_region_set_table::out, 
+	proc_region_set_table::in, proc_region_set_table::out, 
+	io::di, io::uo) is det.
+
+live_region_analysis_proc(HLDS, RptaInfoTable, LVBeforeTable,
LVAfterTable, 
+    VoidVarTable, PredId, ProcId, 
+	!LRBeforeTable, !LRAfterTable, !VoidVarRegionTable, !InputRTable,
+	!OutputRTable, !BornRTable, !DeadRTable, !LocalRTable, !IO) :-
+	PredProcId = proc(PredId, ProcId),
+	( if
+		smm_utils.some_are_special_preds([PredProcId], HLDS)
+	then
+		% XXX: This action seems to be overkilled, it never goes in this 
+        % branch.
+		% XXX: for now just ignore these special predicates
+		% such as __Unify__ and others or non-defined-in-module ones
+		% for the later ones, they should have been analysed when their 
+        % modules analysed and those tables of them will be integrated. 
+        % But it is not the case at the moment.
+        %
+		true
+	else
+		module_info_pred_proc_info(HLDS, PredProcId, _PredInfo, ProcInfo),
+		
+        % the region points-to graph of this procedure
+		( 
+			RptaInfo = rpta_info_table_search_rpta_info(PredProcId, 
+                RptaInfoTable)
+		-> 
+			RptaInfo = rpta_info(Graph, _ALpha),
+
+            % compute region sets
+			find_input_output_params(HLDS, ProcInfo, Inputs, Outputs),
+			lv_to_lr(set.from_list(Inputs), Graph, HLDS, ProcInfo, InputR),
+			lv_to_lr(set.from_list(Outputs), Graph, HLDS, ProcInfo, OutputR),
+                
+            % set of input regions
+			map.set(!.InputRTable, PredProcId, InputR, !:InputRTable),
+
+            % set of output regions
+			map.set(!.OutputRTable, PredProcId, OutputR, !:OutputRTable),
+            
+            % bornR
+			set.difference(OutputR, InputR, BornR),
+			map.set(!.BornRTable, PredProcId, BornR, !:BornRTable),
+            
+            % deadR
+			set.difference(InputR, OutputR, DeadR),
+			map.set(!.DeadRTable, PredProcId, DeadR, !:DeadRTable),
+
+            % localR
+			rptg_get_nodes(Graph, Nodes),
+			set.difference(set.difference(set.from_list(Nodes), InputR),
+                OutputR, LocalR),
+			map.set(!.LocalRTable, PredProcId, LocalR, !:LocalRTable),
+
+			map.lookup(LVBeforeTable, PredProcId, ProcLVBefore),
+			map.foldl(proc_lv_to_proc_lr(Graph, HLDS, ProcInfo), 
+				ProcLVBefore, map.init, ProcLRBefore),
+			map.set(!.LRBeforeTable, PredProcId, ProcLRBefore, 
+                !:LRBeforeTable),
+ 
+			map.lookup(LVAfterTable, PredProcId, ProcLVAfter),
+			map.foldl(proc_lv_to_proc_lr(Graph, HLDS, ProcInfo), 
+				ProcLVAfter, map.init, ProcLRAfter),
+			map.set(!.LRAfterTable, PredProcId, ProcLRAfter, !:LRAfterTable),
+
+			map.lookup(VoidVarTable, PredProcId, ProcVoidVar),
+			map.foldl(proc_lv_to_proc_lr(Graph, HLDS, ProcInfo), 
+				ProcVoidVar, map.init, ProcVoidVarRegion),
+			map.set(!.VoidVarRegionTable, PredProcId, ProcVoidVarRegion, 
+                !:VoidVarRegionTable)
+		;
+			unexpected(thisfile, "live_region_analysis_proc: rpta_info must"
+                ++ " exist")
+		)
+	).
+
+:- pred proc_lv_to_proc_lr(rpt_graph::in, module_info::in, proc_info::in, 
+	program_point::in, variable_set::in, pp_region_set_table::in, 
+    pp_region_set_table::out) is det.
+proc_lv_to_proc_lr(Graph, HLDS, ProcInfo, PP, LVs, !ProcLRMap) :-
+	lv_to_lr(LVs, Graph, HLDS, ProcInfo, LRs),
+	map.set(!.ProcLRMap, PP, LRs, !:ProcLRMap).
+
+:- pred foldl9(pred(L, A, A, B, B, C, C, D, D, E, E, F, F, G, G, H, H,
I, I), 
+    list(L),
+	A, A, B, B, C, C, D, D, E, E, F, F, G, G, H, H, I, I).
+:- mode foldl9(pred(in, in, out, in, out, in, out, in, out, in, out,
in, out, 
+    in, out, in, out, di, uo) is det,
+	in, in, out, in, out, in, out, in, out, in, out, in, out, in, out, in, 
+    out, di, uo) is det.
+foldl9(_, [], !A, !B, !C, !D, !E, !F, !G, !H, !I).
+foldl9(P, [H|T], !A, !B, !C, !D, !E, !F, !G, !H, !I) :-
+	call(P, H, !A, !B, !C, !D, !E, !F, !G, !H, !I),
+	foldl9(P, T, !A, !B, !C, !D, !E, !F, !G, !H, !I).
+
+:- pred lv_to_lr(variable_set::in, rpt_graph::in, module_info::in, 
+    proc_info::in, region_set::out) is det.
+lv_to_lr(LVSet, Graph, HLDS, ProcInfo, LRSet) :-
+	(
+		set.empty(LVSet)
+	->
+		set.init(LRSet)
+	;
+        % collect reachable regions at this pp
+		set.init(LRSet0),
+		set.fold(reach_from_a_variable(Graph, HLDS, ProcInfo), LVSet, 
+            LRSet0, LRSet)
+	).
+
+:- pred find_input_output_params(module_info::in, proc_info::in,
+    list(prog_var)::out, list(prog_var)::out) is det.
+find_input_output_params(HLDS, CalleeProcInfo, Inputs, Outputs) :-
+    proc_info_get_headvars(CalleeProcInfo, ArgVars),
+    proc_info_get_vartypes(CalleeProcInfo, VarTypes),
+    list.map(map.lookup(VarTypes), ArgVars, ArgTypes),
+    proc_info_get_argmodes(CalleeProcInfo, ArgModes),
+    arg_info.compute_in_and_out_vars(HLDS, ArgVars, ArgModes, ArgTypes,
+        Inputs, Outputs).
+
+%----------------------------------------------------------------------------%
+%
+% part for computing across procedure region lifetime
+%
+
+compute_interproc_region_lifetime(HLDS, RptaInfoTable, EPTable,
+    LRBeforeTable, LRAfterTable, InputRTable, OutputRTable, 
+    ConstantRTable, !BornRTable, !DeadRTable, !IO) :-
+    
+    apply_live_region_born_removal_rules(HLDS, RptaInfoTable, EPTable, 
+        LRBeforeTable, LRAfterTable, !BornRTable, !IO),
+
+	apply_live_region_dead_removal_rules(HLDS, RptaInfoTable, EPTable, 
+        LRBeforeTable, LRAfterTable, !DeadRTable, !IO),
+
+	map.foldl(compute_constantR(InputRTable, OutputRTable, !.BornRTable), 
+		!.DeadRTable, map.init, ConstantRTable).
+
+:- pred compute_constantR(proc_region_set_table::in, 
+    proc_region_set_table::in, proc_region_set_table::in, 
+	pred_proc_id::in, region_set::in, proc_region_set_table::in, 
+    proc_region_set_table::out) is det.
+compute_constantR(InputRTable, OutputRTable, BornRTable, PredProcId,
DeadR, 
+    !ConstantRTable) :-
+	map.lookup(InputRTable, PredProcId, InputR),
+	map.lookup(OutputRTable, PredProcId, OutputR),
+	map.lookup(BornRTable, PredProcId, BornR),
+	set.union(InputR, OutputR, IOR),
+	set.difference(IOR, BornR, IOR1),
+	set.difference(IOR1, DeadR, ConstantR),
+	svmap.set(PredProcId, ConstantR, !ConstantRTable).
+
+ignore_primitive_regions(HLDS, RptaInfoTable, !BornRTable, !DeadRTable,
+    !ConstantRTable, !LocalRTable, !LRBeforeTable, !LRAfterTable, 
+    !VoidVarRegionTable) :-
+	map.foldl(eliminate_primitive_regions(HLDS, RptaInfoTable), 
+		!.BornRTable, !BornRTable),
+	map.foldl(eliminate_primitive_regions(HLDS, RptaInfoTable), 
+		!.DeadRTable, !DeadRTable),
+	map.foldl(eliminate_primitive_regions(HLDS, RptaInfoTable), 
+		!.ConstantRTable, !ConstantRTable),
+	map.foldl(eliminate_primitive_regions(HLDS, RptaInfoTable), 
+		!.LocalRTable, !LocalRTable),
+
+	map.foldl(eliminate_primitive_regions_2(HLDS, RptaInfoTable), 
+		!.LRBeforeTable, !LRBeforeTable),
+	map.foldl(eliminate_primitive_regions_2(HLDS, RptaInfoTable), 
+		!.LRAfterTable, !LRAfterTable),
+	map.foldl(eliminate_primitive_regions_2(HLDS, RptaInfoTable), 
+		!.VoidVarRegionTable, !VoidVarRegionTable).
+
+:- pred eliminate_primitive_regions(module_info::in, rpta_info_table::in, 
+    pred_proc_id::in, region_set::in, proc_region_set_table::in, 
+    proc_region_set_table::out) is det.
+eliminate_primitive_regions(HLDS, RptaInfoTable, PredProcId, RegionSet0, 
+    !RegionSetTable) :-
+	map.lookup(RptaInfoTable, PredProcId, RptaInfo),
+	RptaInfo = rpta_info(Graph, _Alpha),
+	set.fold(non_primitive_regions(HLDS, Graph), RegionSet0, set.init, 
+        RegionSet),
+	svmap.det_update(PredProcId, RegionSet, !RegionSetTable).
+
+:- pred non_primitive_regions(module_info::in, rpt_graph::in,
rptg_node::in, 
+	region_set::in, region_set::out) is det.
+non_primitive_regions(HLDS, Graph, Region, !RegionSet) :-
+	rptg_node_contents(Graph, Region, Content),
+	rptg_node_content_get_node_type(Content, NodeType),
+	( 
+		type_util.type_is_atomic(HLDS, NodeType)
+	  ->
+		true
+	  ;
+		type_util.is_dummy_argument_type(HLDS, NodeType)
+	  ->
+		true
+	  ;
+		svset.insert(Region, !RegionSet)
+	).
+
+:- pred eliminate_primitive_regions_2(module_info::in,
rpta_info_table::in, 
+    pred_proc_id::in, pp_region_set_table::in,
proc_pp_region_set_table::in, 
+    proc_pp_region_set_table::out) is det.
+eliminate_primitive_regions_2(HLDS, RptaInfoTable, PredProcId, LRProc0, 
+    !LRTable) :-
+	map.lookup(RptaInfoTable, PredProcId, RptaInfo),
+	RptaInfo = rpta_info(Graph, _Alpha),
+	map.foldl(non_primitive_regions_at_pp(HLDS, Graph), LRProc0, map.init, 
+        LRProc),
+	svmap.det_update(PredProcId, LRProc, !LRTable).
+
+:- pred non_primitive_regions_at_pp(module_info::in, rpt_graph::in, 
+	smm_data.program_point::in, region_set::in, pp_region_set_table::in, 
+    pp_region_set_table::out) is det.
+non_primitive_regions_at_pp(HLDS, Graph, PP, RegionSet0, !LRProc) :-
+	set.fold(non_primitive_regions(HLDS, Graph), RegionSet0, set.init, 
+        RegionSet),
+	svmap.set(PP, RegionSet, !LRProc).
+
+:- func thisfile = string.
+thisfile = "lra_lva.m".
Index: lra_rules.m
===================================================================
RCS file: lra_rules.m
diff -N lra_rules.m
--- /dev/null	1 Jan 1970 00:00:00 -0000
+++ lra_rules.m	15 May 2007 04:15:13 -0000
@@ -0,0 +1,490 @@
+%-----------------------------------------------------------------------------%
+% vim: ft=mercury ts=4 sw=4 et
+%-----------------------------------------------------------------------------%
+% Copyright (C) 2005-2007 The University of Melbourne.
+% This file may only be copied under the terms of the GNU General
+% Public License - see the file COPYING in the Mercury distribution.
+%-----------------------------------------------------------------------------%
+%
+% File: lra_rules.m.
+% Main author: Quan Phan
+%
+% This module implements the process of updating bornR and deadR sets of 
+% each procedure, which basically means to detect lifetime of regions
across
+% procedure boundary. bornR is the set of regions that the procedure will 
+% create, deadR is the set of regions that the procedure can safely
removed.
+% 
+% The updating will only happen at call sites and when analysing a
procedure
+% we will not update its deadR and bornR sets but update those sets of the
+% procedures that it invokes.  This process requires a fixpoint
computation 
+% and we will follow the strongly connected component graph in the top-down
+% manner.
+
+:- module lra_rules. 
+
+:- interface.
+
+:- import_module hlds.
+:- import_module hlds.hlds_module.
+
+:- import_module io.
+
+:- import_module rpta_info.
+:- import_module lra_data.
+
+	% Apply the live region analysis rules to update bornR and deadR 
+	% sets of each procedure.
+    %
+:- pred apply_live_region_dead_removal_rules(module_info::in, 
+	rpta_info_table::in, ep_table::in, proc_pp_region_set_table::in, 
+    proc_pp_region_set_table::in, proc_region_set_table::in, 
+    proc_region_set_table::out, io::di, io::uo) is det.
+
+:- pred apply_live_region_born_removal_rules(module_info::in, 
+    rpta_info_table::in, ep_table::in, proc_pp_region_set_table::in, 
+    proc_pp_region_set_table::in, proc_region_set_table::in, 
+    proc_region_set_table::out, io::di, io::uo) is det.
+
+%-----------------------------------------------------------------------------%
+
+:- implementation. 
+
+:- import_module hlds.hlds_pred. 
+:- import_module hlds.hlds_data.
+:- import_module hlds.hlds_goal.
+:- import_module hlds.hlds_llds.
+:- import_module transform_hlds.
+:- import_module transform_hlds.dependency_graph.
+:- import_module check_hlds.
+:- import_module check_hlds.goal_path.
+:- import_module check_hlds.type_util.
+:- import_module libs.
+:- import_module libs.compiler_util.
+:- import_module rpt_graph.
+:- import_module smm_utils.
+
+:- import_module list.
+:- import_module map.
+:- import_module bool.
+:- import_module varset.
+:- import_module string.
+:- import_module term.
+:- import_module set.
+:- import_module int.
+:- import_module std_util.
+:- import_module pair.
+:- import_module solutions.
+:- import_module maybe.
+
+apply_live_region_dead_removal_rules(HLDS, RptaInfoTable, EPTable, 
+	LRBeforeTable, LRAfterTable, !DeadRTable, !IO) :-
+	apply_live_region_rule(dead_removal_rules, HLDS, RptaInfoTable, EPTable, 
+        LRBeforeTable, LRAfterTable, !DeadRTable, !IO).
+
+apply_live_region_born_removal_rules(HLDS, RptaInfoTable, EPTable, 
+	LRBeforeTable, LRAfterTable, !BornRTable, !IO) :-
+	apply_live_region_rule(born_removal_rules, HLDS, RptaInfoTable, EPTable, 
+        LRBeforeTable, LRAfterTable, !BornRTable, !IO).
+
+:- type rule_pred == ( 
+        pred(pred_proc_id, region_set, region_set, proc_region_set_table, 
+            map(rptg_node, rptg_node), region_set) 
+    ).
+:- inst rule_pred == ( pred(in, in, in, in, in, out) is det ).
+
+:- pred apply_live_region_rule(rule_pred::in(rule_pred), module_info::in, 
+    rpta_info_table::in, ep_table::in, proc_pp_region_set_table::in, 
+    proc_pp_region_set_table::in, proc_region_set_table::in, 
+    proc_region_set_table::out, io::di, io::uo) is det.
+
+apply_live_region_rule(Rule, HLDS, RptaInfoTable, EPTable, LRBeforeTable, 
+    LRAfterTable, !ProcRegionSetTable, !IO) :-
+    % strongly connected components needed
+	transform_hlds.dependency_graph.module_info_ensure_dependency_info(
+		HLDS, HLDS1),
+	module_info_get_maybe_dependency_info(HLDS1, MaybeDepInfo),
+	(
+		MaybeDepInfo = yes(DepInfo)
+	->
+		hlds.hlds_module.hlds_dependency_info_get_dependency_ordering(
+			DepInfo, DepOrdering),
+        % perform the analysis and update the rpta_info of 
+        % the procedures
+		run_with_dependencies(Rule, DepOrdering, HLDS1, RptaInfoTable, 
+			EPTable, LRBeforeTable, LRAfterTable, !ProcRegionSetTable, !IO)
+	;
+		unexpected(thisfile, "lra_rules: no dependency info")
+	).
+
+:- pred run_with_dependencies(rule_pred::in(rule_pred),
+    dependency_ordering::in, module_info::in, rpta_info_table::in, 
+    ep_table::in, proc_pp_region_set_table::in, 
+    proc_pp_region_set_table::in, proc_region_set_table::in, 
+    proc_region_set_table::out, io::di, io::uo) is det.
+    
+run_with_dependencies(Rule, Deps, HLDS, RptaInfoTable, EPTable,
LRBeforeTable, 
+    LRAfterTable, !ProcRegionSetTable, !IO) :-
+    % We want to proceed the SCC graph top-down so reverse the list
+    % (the process is foldr2, but it is not yet in list module)
+	list.reverse(Deps, Deps1), 
+	list.foldl2(run_with_dependency(Rule, HLDS, RptaInfoTable, EPTable, 
+		LRBeforeTable, LRAfterTable), 
+		Deps1, !ProcRegionSetTable, !IO).
+
+:- pred run_with_dependency(rule_pred::in(rule_pred), 
+    module_info::in, rpta_info_table::in, ep_table::in,
+    proc_pp_region_set_table::in, proc_pp_region_set_table::in, 
+    list(pred_proc_id)::in, proc_region_set_table::in, 
+    proc_region_set_table::out, io::di, io::uo) is det.
+
+run_with_dependency(Rule, HLDS, RptaInfoTable, EPTable, LRBeforeTable, 
+    LRAfterTable, SCC, !ProcRegionSetTable, 
+    !IO) :- 
+	(
+        % analysis ignores special predicates
+		smm_utils.some_are_special_preds(SCC, HLDS)
+	->
+		true
+	;
+        % for each list of strongly connected components, 
+        % perform a fixpoint computation.
+		run_with_dependency_until_fixpoint(Rule, SCC, HLDS, RptaInfoTable, 
+            EPTable, LRBeforeTable, LRAfterTable, !ProcRegionSetTable, !IO)
+	).
+
+:- pred run_with_dependency_until_fixpoint(rule_pred::in(rule_pred),
+    list(pred_proc_id)::in, module_info::in, rpta_info_table::in, 
+    ep_table::in, proc_pp_region_set_table::in,
proc_pp_region_set_table::in, 
+    proc_region_set_table::in, proc_region_set_table::out, io::di, io::uo) 
+    is det.
+
+run_with_dependency_until_fixpoint(Rule, SCC, HLDS, RptaInfoTable,
EPTable, 
+	LRBeforeTable, LRAfterTable, ProcRegionSetTable0, 
+    ProcRegionSetTable, !IO) :-
+    % this call calculates the region set for each procedure in SCC
+	list.foldl2(apply_rule_pred_proc(Rule, HLDS, RptaInfoTable, EPTable, 
+        LRBeforeTable, LRAfterTable), SCC, ProcRegionSetTable0, 
+        ProcRegionSetTable1, !IO),
+	(
+		proc_region_set_table_equal(ProcRegionSetTable1, ProcRegionSetTable0)
+	->
+		% if all region_set's in the FPTable are intact update the main 
+        % ProcRegionSetTable.
+		ProcRegionSetTable = ProcRegionSetTable1
+	;
+		% some is not fixed, start all over again 
+		run_with_dependency_until_fixpoint(Rule, SCC, HLDS, RptaInfoTable, 
+            EPTable, LRBeforeTable, LRAfterTable, ProcRegionSetTable1, 
+            ProcRegionSetTable, !IO)
+	).
+
+:- pred apply_rule_pred_proc(rule_pred::in(rule_pred),
+    module_info::in, rpta_info_table::in, ep_table::in, 
+    proc_pp_region_set_table::in, proc_pp_region_set_table::in, 
+    pred_proc_id::in, proc_region_set_table::in,
proc_region_set_table::out, 
+    io::di, io::uo) 
+    is det.
+
+apply_rule_pred_proc(Rule, HLDS, RptaInfoTable, EPTable, LRBeforeTable, 
+    LRAfterTable, PredProcId, !ProcRegionSetTable, !IO) :- 
+    % We need to follow each execution path and apply the two rules when 
+    % possible
+	map.lookup(RptaInfoTable, PredProcId, RptaInfo),
+	map.lookup(EPTable, PredProcId, EPs), 
+	map.lookup(LRBeforeTable, PredProcId, ProcLRBefore),
+	map.lookup(LRAfterTable, PredProcId, ProcLRAfter),
+
+    % here we analysing a caller but will update the region sets of its 
+    % callees.
+	apply_live_region_rules_eps(Rule, EPs, EPTable, HLDS, PredProcId, 
+        RptaInfo, RptaInfoTable, ProcLRBefore, ProcLRAfter, 
+        !ProcRegionSetTable, !IO).
+
+:- pred apply_live_region_rules_eps(rule_pred::in(rule_pred),
+    list(execution_path)::in, ep_table::in, module_info::in,
pred_proc_id::in, 
+    rpta_info::in, rpta_info_table::in, pp_region_set_table::in, 
+    pp_region_set_table::in, proc_region_set_table::in, 
+    proc_region_set_table::out, io::di, io::uo) is det.
+
+apply_live_region_rules_eps(_Rule, [], _, _, _, _, _, _ProcLRBefore, 
+    _ProcLRAfter, !ProcRegionSetTable, !IO).
+apply_live_region_rules_eps(Rule, [EP|OtherEPs], EPTable, HLDS,
PredProcId, 
+    RptaInfo, RptaInfoTable, ProcLRBefore, ProcLRAfter,
!ProcRegionSetTable, 
+    !IO) :-
+	apply_live_region_rules_ep(Rule, EP, EPTable, HLDS, PredProcId, RptaInfo, 
+		RptaInfoTable, ProcLRBefore, ProcLRAfter, !ProcRegionSetTable, !IO),
+	apply_live_region_rules_eps(Rule, OtherEPs, EPTable, HLDS, PredProcId, 
+        RptaInfo, RptaInfoTable, ProcLRBefore, ProcLRAfter, 
+        !ProcRegionSetTable, !IO).
+
+	% follow each execution path of a procedure and update deadR and bornR 
+    % sets
+    %
+:- pred apply_live_region_rules_ep(rule_pred::in(rule_pred),
+    execution_path::in, ep_table::in, module_info::in, pred_proc_id::in, 
+    rpta_info::in, rpta_info_table::in, pp_region_set_table::in, 
+    pp_region_set_table::in, proc_region_set_table::in, 
+    proc_region_set_table::out, io::di, io::uo) is det.
+
+apply_live_region_rules_ep(_Rule, [], _, _, _, _, _, _ProcLRBefore, 
+    _ProcLRAfter, !ProcRegionSetTable, !IO).
+apply_live_region_rules_ep(Rule, [PP - Goal|OtherPP_Goals], EPTable, HLDS, 
+    PredProcId, RptaInfo, RptaInfoTable, ProcLRBefore, ProcLRAfter, 
+    !ProcRegionSetTable, !IO) :-
+	Goal = hlds_goal(Expr, _Info),
+    % The updating will only happen at call sites, i.e., when the goal
+    % at the program point is a procedure call.
+	( if
+		Expr = plain_call(PredId_q, ProcId_q, _, _, _, _)
+	  then
+	  	PredProcId_q = proc(PredId_q, ProcId_q),
+		( if 
+			smm_utils.some_are_special_preds([PredProcId_q], HLDS)
+		  then
+			true
+		  else
+			RptaInfo = rpta_info(_Graph_p, AlphaMapping),
+			map.lookup(AlphaMapping, PP, AlphaAtPP),
+			
+            map.lookup(ProcLRBefore, PP, LRBefore), 
+			map.lookup(ProcLRAfter, PP, LRAfter),
+
+            % apply a rule
+			call(Rule, PredProcId_q, LRBefore, LRAfter, 
+                !.ProcRegionSetTable, AlphaAtPP, RegionSet),
+				
+			map.lookup(!.ProcRegionSetTable, PredProcId_q, RegionSet0),
+			( if 
+				set.equal(RegionSet, RegionSet0)
+			  then
+                % i.e., no region is removed, so everything is the same 
+                % as before
+				true
+			  else
+                % some regions are removed, record the new set for q
and ...
+				map.set(!.ProcRegionSetTable, PredProcId_q, RegionSet, 
+                    !:ProcRegionSetTable),
+
+                % ... those removals need to be propagated to the ones 
+                % called by q
+				set.difference(RegionSet0, RegionSet, ToBeRemoved),
+				list.foldl2(remove_this_region_from_callees_of_proc(
+					PredProcId_q, EPTable, HLDS, RptaInfoTable), 
+					set.to_sorted_list(ToBeRemoved), !ProcRegionSetTable, !IO)
+			)	
+		)
+  	  else
+        % ignore other sorts of goal
+		true
+	),
+    apply_live_region_rules_ep(Rule, OtherPP_Goals, EPTable, HLDS,
PredProcId, 
+        RptaInfo, RptaInfoTable, ProcLRBefore, ProcLRAfter, 
+        !ProcRegionSetTable, !IO).
+
+%-----------------------------------------------------------------------------%
+%
+% implementation of the live region analysis rules.
+%
+% the idea of those rules is to ensure that:
+% 1. when it is not safe for a procedure to remove a region, that
region must 
+% not be in the procedure's deadR seti,
+% 2. when a region exists before the procedure is called, that region
must not
+% be in the procedure's bornR set.
+% 
+
+    % rules for eliminating regions from deadR set.
+    %
+:- pred dead_removal_rules(pred_proc_id::in, region_set::in,
region_set::in, 
+    proc_region_set_table::in, map(rptg_node, rptg_node)::in, 
+	region_set::out) is det.
+dead_removal_rules(Q_Id, LRBefore, LRAfter, DeadRTable, AlphaAtPP,
DeadR_q) :-
+    % the current deadR of q
+	map.lookup(DeadRTable, Q_Id, DeadR_q0), 
+
+    % apply dead removal rule L1 for r that is live before and after the 
+    % call to q
+	set.intersect(LRBefore, LRAfter, Rule1_Candidate),
+	set.fold(dead_removal_rule_1(AlphaAtPP), Rule1_Candidate, DeadR_q0, 
+        DeadR_q1),
+	
+    % remove deadR rule L2,
+	targets_with_more_than_one_source(AlphaAtPP, Targets),
+	set.fold(dead_removal_rule_2(AlphaAtPP), Targets, DeadR_q1, DeadR_q).
+	
+:- pred dead_removal_rule_1(map(rptg_node, rptg_node)::in, rptg_node::in, 
+    region_set::in, region_set::out) is det.
+dead_removal_rule_1(AlphaAtPP, Region, DeadR_q0, DeadR_q) :-
+    % find r' such that alpha(r') = Region
+	solutions(map.inverse_search(AlphaAtPP, Region), SourceList),
+	set.list_to_set(SourceList, RPrimes),
+
+    % remove any r' that is in deadR(q)
+	set.difference(DeadR_q0, RPrimes, DeadR_q).
+
+:- pred dead_removal_rule_2(map(rptg_node, rptg_node)::in, rptg_node::in,
+    set(rptg_node)::in, set(rptg_node)::out) is det.
+dead_removal_rule_2(AlphaAtPP, Region, DeadR_q0, DeadR_q) :-
+	solutions(map.inverse_search(AlphaAtPP, Region), SourceList),
+	set.list_to_set(SourceList, RPrimes),	
+	set.difference(DeadR_q0, RPrimes, DeadR_q).
+
+    % rules for eliminating regions from bornR set.
+    %
+:- pred born_removal_rules(pred_proc_id::in, region_set::in,
region_set::in, 
+	proc_region_set_table::in, map(rptg_node, rptg_node)::in, 
+    region_set::out) is det.
+born_removal_rules(Q_Id, LRBefore, _LRAfter, BornRTable, AlphaAtPP,
BornR_q) :-
+    % the current bornR of q
+	map.lookup(BornRTable, Q_Id, BornR_q0),
+
+    % apply born removal rule L3 for r that is live before and after the 
+    % call to q
+	set.fold(born_removal_rule_1(AlphaAtPP), LRBefore, BornR_q0, BornR_q1),
+
+    % remove bornR rule L4,
+	targets_with_more_than_one_source(AlphaAtPP, Targets),
+	set.fold(born_removal_rule_2(AlphaAtPP), Targets, BornR_q1, BornR_q).
+	
+:- pred born_removal_rule_1(map(rptg_node, rptg_node)::in, rptg_node::in,
+    set(rptg_node)::in, set(rptg_node)::out) is det.
+born_removal_rule_1(AlphaAtPP, Region, BornR_q0, BornR_q) :-
+	solutions(map.inverse_search(AlphaAtPP, Region), SourceList),
+	set.list_to_set(SourceList, RPrimes),
+	set.difference(BornR_q0, RPrimes, BornR_q).
+	
+	% alpha(r') = r, alpha(r'') = r, r', r'' in bornR(q) imply remove r', 
+    % r'' from bornR(q).
+	%
+:- pred born_removal_rule_2(map(rptg_node, rptg_node)::in, rptg_node::in,
+    set(rptg_node)::in, set(rptg_node)::out) is det.
+born_removal_rule_2(AlphaAtPP, Region, BornR_q0, BornR_q) :-
+	solutions(map.inverse_search(AlphaAtPP, Region), SourceList),
+	set.list_to_set(SourceList, RPrimes),
+	set.difference(BornR_q0, RPrimes, BornR_q).
+	
+	% find targets of alpha mapping that are mapped to by more than one 
+    % source. 
+    %
+:- pred targets_with_more_than_one_source(map(rptg_node, rptg_node)::in, 
+    region_set::out) is det.
+targets_with_more_than_one_source(AlphaAtPP, Targets) :-
+    map.foldl2(process_one_mapping, AlphaAtPP, set.init, _Processed,
set.init, 
+        Targets).
+
+:- pred process_one_mapping(rptg_node::in, rptg_node::in, region_set::in, 
+    region_set::out, region_set::in, region_set::out) is det.
+process_one_mapping(_Source, Target, Candidates0, Candidates, Targets0, 
+    Targets) :-
+	( if
+		set.contains(Candidates0, Target)
+	  then
+		set.insert(Targets0, Target, Targets),
+		Candidates = Candidates0
+	  else
+		set.insert(Candidates0, Target, Candidates),
+		Targets = Targets0
+	).
+	
+    % This predicate propagates the removal of a region from a deadR or
+    % bornR sets of a procedure to the ones it calls, i.e., also remove
+    % the region from the corresponding sets of them.
+    %
+:- pred remove_this_region_from_callees_of_proc(pred_proc_id::in, 
+    ep_table::in, module_info::in, rpta_info_table::in, rptg_node::in,
+	proc_region_set_table::in, proc_region_set_table::out, io::di, io::uo)
+    is det.
+remove_this_region_from_callees_of_proc(PredProcId, EPTable, HLDS, 
+    RptaInfoTable, Region, !ProcRegionSetTable, !IO) :-
+    % to have the goal at each pp
+	map.lookup(EPTable, PredProcId, EPs), 
+
+    % follow eps of this proc and remove the region from this proc's 
+    % callees.
+	remove_this_from_eps(EPs, PredProcId, Region, EPTable, HLDS, 
+        RptaInfoTable, !ProcRegionSetTable, !IO).
+
+	% follow each execution path of a procedure and update deadR and 
+    % bornR sets
+    %
+:- pred remove_this_from_eps(list(execution_path)::in, pred_proc_id::in, 
+    rptg_node::in, ep_table::in, module_info::in, rpta_info_table::in, 
+	proc_region_set_table::in, proc_region_set_table::out, 
+    io::di, io::uo) is det.
+remove_this_from_eps([], _, _, _, _, _, !ProcRegionSetTable, !IO).
+remove_this_from_eps([EP|OtherEPs], PredProcId, Region, EPTable, HLDS, 
+    RptaInfoTable, !ProcRegionSetTable, !IO) :-
+	remove_this_from_ep(EP, PredProcId, Region, EPTable, HLDS, RptaInfoTable, 
+        !ProcRegionSetTable, !IO),
+	remove_this_from_eps(OtherEPs, PredProcId, Region, EPTable, HLDS, 
+        RptaInfoTable, !ProcRegionSetTable, !IO).
+
+:- pred remove_this_from_ep(execution_path::in, pred_proc_id::in, 
+    rptg_node::in, ep_table::in, module_info::in, rpta_info_table::in, 
+	proc_region_set_table::in, proc_region_set_table::out, 
+    io::di, io::uo) is det.
+remove_this_from_ep([], _, _, _, _, _, !ProcRegionSetTable, !IO).
+remove_this_from_ep([PP - Goal|OtherPP_Goals], PredProcId,
ToBeRemovedRegion, 
+    EPTable, HLDS, RptaInfoTable, !ProcRegionSetTable, !IO) :-
+	Goal = hlds_goal(Expr, _Info),
+	( if
+		Expr = plain_call(PredId_q, ProcId_q, _, _, _, _)
+	  then
+	  	PredProcId_q = proc(PredId_q, ProcId_q),
+		( if 
+			smm_utils.some_are_special_preds([PredProcId_q], HLDS)
+		  then
+			true
+		  else
+            % find the alpha mapping: alpha(_, R) = ToBeRemovedRegion
+			map.lookup(RptaInfoTable, PredProcId, RptaInfo_p),
+			RptaInfo_p = rpta_info(_Graph_p, AlphaMapping),
+			map.lookup(AlphaMapping, PP, AlphaAtPP),
+		        
+			map.foldl(find_alpha_source(ToBeRemovedRegion), AlphaAtPP, 
+                set.init, Rs),
+ 
+            % remove the sources from the RegionSet (either deadR or 
+            % bornR) of this callee
+			map.lookup(!.ProcRegionSetTable, PredProcId_q, RegionSet0),
+			set.difference(RegionSet0, Rs, RegionSet1),
+				
+            % update the table and continue
+			( if
+				set.equal(RegionSet0, RegionSet1)
+			  then
+                % no need to update
+				true
+			  else
+                % some is removed from deadR or bornR of this callee, 
+                % so we update the entry of this called and analyse it
+				map.set(!.ProcRegionSetTable, PredProcId_q, RegionSet1, 
+                    !:ProcRegionSetTable),
+				set.difference(RegionSet0, RegionSet1, RemovedFromQ),
+
+                % call this one mutually recursively
+				list.foldl2(remove_this_region_from_callees_of_proc(
+                    PredProcId_q, EPTable, HLDS, RptaInfoTable), 
+                set.to_sorted_list(RemovedFromQ), !ProcRegionSetTable, !IO)
+			) 
+		)
+  	  else
+			% ignore other sorts of goals
+			%
+		true
+	),
+	remove_this_from_ep(OtherPP_Goals, PredProcId, ToBeRemovedRegion, 
+		EPTable, HLDS, RptaInfoTable, !ProcRegionSetTable, !IO).
+
+:- pred find_alpha_source(rptg_node::in, rptg_node::in, rptg_node::in,
+	set(rptg_node)::in, set(rptg_node)::out) is det.
+find_alpha_source(ToBeRemovedRegion, Source, Target, !Rs) :-
+	( if 
+		ToBeRemovedRegion = Target
+	  then
+		set.insert(!.Rs, Source, !:Rs)
+	  else
+		true
+	).
+
+:- func thisfile = string.
+thisfile = "lra_rules.m".
Index: mercury_compile.m
===================================================================
RCS file:
/home/mercury/mercury1/repository/mercury/compiler/mercury_compile.m,v
retrieving revision 1.432
diff -u -u -r1.432 mercury_compile.m
--- mercury_compile.m	12 Apr 2007 06:40:09 -0000	1.432
+++ mercury_compile.m	15 May 2007 06:00:14 -0000
@@ -101,6 +101,13 @@
 :- import_module transform_hlds.size_prof.
 :- import_module ll_backend.deep_profiling.
 
+    % ======== added by Quan, 03/5/2007 rbmm analysis =========== %
+:- import_module smm_utils.
+:- import_module rpta_run.
+:- import_module lra_lva.
+:- import_module region_transform.
+:- import_module pp_region_annotated_code.
+
     % the LLDS back-end
 :- import_module ll_backend.continuation_info.
 :- import_module ll_backend.dupproc.
@@ -125,9 +132,7 @@
     %
 :- import_module ll_backend.llds_to_x86_64.
 :- import_module ll_backend.llds_to_x86_64_out.
-:- import_module ll_backend.x86_64_instrs.
-:- import_module ll_backend.x86_64_out.
-:- import_module ll_backend.x86_64_regs.
+%:- import_module ll_backend.x86_64_instrs.
 
     % the MLDS back-end
 :- import_module ml_backend.add_trail_ops.         % HLDS -> HLDS
@@ -2647,6 +2652,10 @@
     maybe_experimental_complexity(Verbose, Stats, !HLDS, !IO),
     maybe_dump_hlds(!.HLDS, 230, "complexity", !DumpInfo, !IO),
 
+    % XXX This may be moved to later.
+    maybe_region_analysis(!HLDS, !IO),
+    maybe_dump_hlds(!.HLDS, 240, "region_analysis", !DumpInfo, !IO),
+
     maybe_dump_hlds(!.HLDS, 299, "middle_pass", !DumpInfo, !IO).
 
 %-----------------------------------------------------------------------------%
@@ -4211,6 +4220,46 @@
         maybe_report_stats(Stats, !IO)
     ).
 
+:- pred maybe_region_analysis(module_info::in, module_info::out, io::di, 
+    io::uo) is det.
+
+maybe_region_analysis(!HLDS, !IO) :-
+    module_info_get_globals(!.HLDS, Globals),
+    globals.lookup_bool_option(Globals, region_analysis, Analysis),
+    (
+        Analysis = yes,
+        rpta_run.region_points_to_analysis(RptaInfoTable, !HLDS, !IO),
+        smm_utils.annotate_all_liveness_in_module(!HLDS, !IO),
+        lra_lva.execution_path_analysis(!.HLDS, EPTable, !IO),
+        lra_lva.live_variable_analysis(!.HLDS, EPTable, LVBeforeTable, 
+            LVAfterTable, VoidVarTable, !IO),
+        lra_lva.live_region_analysis(!.HLDS, RptaInfoTable, 
+            LVBeforeTable, LVAfterTable, VoidVarTable,
+            LRBeforeTable0, LRAfterTable0, VoidVarRegionTable0, 
+            InputRTable, OutputRTable, BornRTable0, DeadRTable0,
LocalRTable0,
+            !IO),
+        lra_lva.compute_interproc_region_lifetime(!.HLDS, RptaInfoTable,
+            EPTable, LRBeforeTable0, LRAfterTable0, InputRTable,
OutputRTable,
+            ConstantRTable0, BornRTable0, BornRTable1, DeadRTable0, 
+            DeadRTable1, !IO),
+        lra_lva.ignore_primitive_regions(!.HLDS, RptaInfoTable, 
+            BornRTable1, BornRTable, DeadRTable1, DeadRTable, 
+            ConstantRTable0, ConstantRTable, LocalRTable0, LocalRTable,
+            LRBeforeTable0, LRBeforeTable, LRAfterTable0, LRAfterTable,
+            VoidVarRegionTable0, VoidVarRegionTable),
+        region_transform.transform(!.HLDS, RptaInfoTable, EPTable,
+            LRBeforeTable, LRAfterTable, VoidVarRegionTable, BornRTable, 
+            DeadRTable, LocalRTable, AnnotationTable, !IO),
+        % XXX: this prints out the annotated code. It will be changed with
+        % something that annotates the HLDS. But for now it is the only way
+        % to verify the analysis' output.
+        pp_region_annotated_code.pp_region_annotations(!.HLDS, 
+            AnnotationTable, RptaInfoTable, ConstantRTable, DeadRTable, 
+            BornRTable, !IO)
+    ;
+        Analysis = no
+    ).
+
 :- pred maybe_deep_profiling(bool::in, bool::in,
     module_info::in, module_info::out, io::di, io::uo) is det.
 
Index: options.m
===================================================================
RCS file: /home/mercury/mercury1/repository/mercury/compiler/options.m,v
retrieving revision 1.556
diff -u -u -r1.556 options.m
--- options.m	24 Apr 2007 09:23:33 -0000	1.556
+++ options.m	3 May 2007 05:54:51 -0000
@@ -291,6 +291,7 @@
     ;       record_term_sizes_as_words
     ;       record_term_sizes_as_cells
     ;       experimental_complexity
+    ;       region_analysis
 
     % (c) Miscellaneous
     ;       gc
@@ -1059,6 +1060,7 @@
     record_term_sizes_as_words          -   bool(no),
     record_term_sizes_as_cells          -   bool(no),
     experimental_complexity             -   string(""),
+    region_analysis                     -   bool(no),
     % (c) Miscellaneous optional features
     gc                                  -   string("boehm"),
     parallel                            -   bool(no),
@@ -1814,6 +1816,7 @@
 long_option("record-term-sizes-as-words", record_term_sizes_as_words).
 long_option("record-term-sizes-as-cells", record_term_sizes_as_cells).
 long_option("experimental-complexity",  experimental_complexity).
+long_option("region-analysis",      region_analysis).
 % (c) miscellaneous optional features
 long_option("gc",                   gc).
 long_option("garbage-collection",   gc).
@@ -3617,7 +3620,9 @@
         "\tEnable experimental complexity analysis for the predicates",
         "\tlisted in the given file.",
         "\tThis option is supported for the C back-end, with",
-        "\t--no-highlevel-code."
+        "\t--no-highlevel-code.",
+        "--region-analysis",
+        "\tEnable the analysis for region based memory management"
     ]),
 
     io.write_string("      Miscellaneous optional features\n"),
Index: pp_region_annotated_code.m
===================================================================
RCS file: pp_region_annotated_code.m
diff -N pp_region_annotated_code.m
--- /dev/null	1 Jan 1970 00:00:00 -0000
+++ pp_region_annotated_code.m	15 May 2007 05:56:42 -0000
@@ -0,0 +1,1320 @@
+%_-----------------------------------------------------------------------------%%
vim: ft=mercury ts=4 sw=4 et
+%-----------------------------------------------------------------------------%
+% Copyright (C) 2006, 2007 The University of Melbourne.
+% This file may only be copied under the terms of the GNU General
+% Public License - see the file COPYING in the Mercury distribution.
+%-----------------------------------------------------------------------------%
+%
+% File: pp_region_annotated_code.m.
+% Main author: Quan Phan.
+%
+% Pretty print Mercury code with region annotations that can be
compiled with
+% a slightly changed MMC for the PURPOSE OF SIMULATING the memory behaviour
+% of RBMM.
+% The region annotations consist of: 
+% 1. region instructions: create, remove, reg_alloc, make_ite_frame, 
+% commit_removal, undo_ite_changes. 
+% 2. formal and actual region parameters for procedures.
+% 3. The type of region parameters is int, their modes are either "in" for 
+% dead and constant regions or "out" for born ones.
+%
+% Date: 20/12/2006
+
+%-----------------------------------------------------------------------------%
+
+:- module pp_region_annotated_code.
+
+:- interface.
+
+:- import_module hlds.
+:- import_module hlds.hlds_module.
+
+:- import_module io.
+
+:- import_module rpta_info.
+:- import_module lra_data.
+:- import_module region_transform.
+
+:- pred pp_region_annotations(module_info::in, annotation_table::in, 
+	rpta_info_table::in, proc_region_set_table::in,
proc_region_set_table::in, 
+	proc_region_set_table::in, io::di, io::uo) is det.
+
+:- implementation. 
+
+:- type region == rptg_node.
+
+:- import_module hlds.hlds_data.
+:- import_module hlds.hlds_goal.
+:- import_module hlds.hlds_pred.
+:- import_module hlds.hlds_out.
+
+:- import_module parse_tree.
+:- import_module parse_tree.prog_data.
+:- import_module parse_tree.prog_type.
+:- import_module parse_tree.prog_out.
+:- import_module parse_tree.prog_util.
+:- import_module parse_tree.mercury_to_mercury.
+:- import_module mdbcomp.
+:- import_module mdbcomp.prim_data.
+:- import_module check_hlds.
+:- import_module check_hlds.goal_path.
+:- import_module check_hlds.type_util.
+:- import_module check_hlds.mode_util.
+:- import_module rpt_graph.
+:- import_module rbmm_utils.
+:- import_module smm_utils.
+:- import_module smm_data.
+
+:- import_module bool.
+:- import_module varset.
+:- import_module list.
+:- import_module map.
+:- import_module string.
+:- import_module int.
+:- import_module term.
+:- import_module term_io.
+:- import_module set.
+:- import_module maybe.
+:- import_module std_util.
+:- import_module pair.
+:- import_module libs.
+:- import_module libs.compiler_util.
+
+pp_region_annotations(HLDS, AnnotationTable, RptaInfoTable, ConstantRTable,
+    DeadRTable, BornRTable, !IO) :- 
+
+	pp_header(HLDS, ConstantRTable, DeadRTable, BornRTable, !IO),
+	
+	module_info_predids(PredIds, HLDS, _),
+	list.foldl(pp_region_annotations_pred(HLDS, AnnotationTable, 
+		RptaInfoTable, ConstantRTable, DeadRTable, BornRTable), 
+		PredIds, !IO).
+
+:- pred pp_header(module_info::in, proc_region_set_table::in, 
+    proc_region_set_table::in, proc_region_set_table::in, io::di, io::uo) 
+    is det.
+
+pp_header(HLDS, ConstantRTable, DeadRTable, BornRTable, !IO) :-
+	
+	module_info_get_name(HLDS, Name),
+	io.write_string(":- module ", !IO),
+	prog_out.write_sym_name(Name, !IO),
+	io.write_string(".\n\n", !IO),
+	
+	io.write_string(":- interface.\n\n", !IO),
+	
+	module_info_get_imported_module_specifiers(HLDS, Imports),
+	pp_imports(Imports, !IO),
+	
+    % predicate declarations
+	module_info_predids(PredIds, HLDS, _),
+	list.foldl(pp_pred_declaration(HLDS, ConstantRTable, DeadRTable,
+			BornRTable), PredIds, !IO),
+	io.nl(!IO),
+	io.write_string(":- implementation.\n\n", !IO),
+	io.write_string(":- import_module reg_prof_int.\n\n", !IO).
+	
+:- pred pp_imports(set(module_specifier)::in, io::di, io::uo) is det.
+
+pp_imports(ImportSet, !IO) :-
+	io.write_string(":- import_module ", !IO),
+	io.write_list(set.to_sorted_list(ImportSet), ", ",
+		prog_out.write_sym_name, !IO),
+	io.write_string(".\n", !IO).
+
+	% print out :- pred ....
+	% XXX: this predicate does not work for functions.
+	%
+:- pred pp_pred_declaration(module_info::in, proc_region_set_table::in, 
+    proc_region_set_table::in, proc_region_set_table::in, pred_id::in, 
+    io::di, io::uo) is det.
+pp_pred_declaration(HLDS, ConstantRTable, DeadRTable, BornRTable, PredId, 
+    !IO) :-
+	module_info_pred_info(HLDS, PredId, PredInfo),
+	
+	pred_info_non_imported_procids(PredInfo) = ProcIds,
+	list.foldl(
+		pred(ProcId::in, A0::in, A::out) is det :- 
+		(
+			A = [proc(PredId, ProcId)|A0]
+		), 
+		ProcIds, 
+		[], PredProcIds
+	),
+	( if 
+		(
+			smm_utils.some_are_special_preds(PredProcIds, HLDS);
+			PredProcIds = []
+		)	
+	  then
+	  	true
+	  else	
+		PredName = pred_info_name(PredInfo),
+		PredOrFunc = pred_info_is_pred_or_func(PredInfo),
+		pred_info_get_arg_types(PredInfo, ArgTypes),
+		pred_info_get_exist_quant_tvars(PredInfo, ExistQVars),
+		pred_info_get_typevarset(PredInfo, TVarSet),
+		pred_info_context(PredInfo, Context),
+		pred_info_get_class_context(PredInfo, ClassContext),
+		pred_info_get_purity(PredInfo, Purity),
+		AppendVarNums = no,
+		(
+			PredOrFunc = pf_predicate,
+			PredType = mercury_pred_type_to_string(TVarSet, 
+				ExistQVars,
+				unqualified(PredName),
+				ArgTypes, no, Purity, ClassContext, Context,
+				AppendVarNums)
+		;
+			PredOrFunc = pf_function,
+			pred_args_to_func_args(ArgTypes, FuncArgTypes,
+				FuncRetType),
+			PredType = mercury_func_type_to_string(TVarSet, 
+				ExistQVars,
+				unqualified(PredName), FuncArgTypes,
+				FuncRetType, no, Purity, ClassContext,
+				Context, AppendVarNums)
+		),
+		( if 
+			ArgTypes \= []
+		  then	% has at least 1 regular param, 
+		  		
+				% print their types
+				%
+			PrintedPredType = string.left(PredType, 
+				string.length(PredType) - 3),
+			io.nl(!IO),	
+			io.write_string(PrintedPredType, !IO),
+			
+            % print "int" as type of regions.
+			( if 
+				PredProcIds = [PredProcId|_]
+			  then
+			  	( if
+					PredName = "main"	
+				  then
+				  	true
+				  else
+				  	pp_reg_param_types(PredProcId, ConstantRTable, DeadRTable, 
+                        BornRTable, !IO)
+				)	
+			  else
+			  	true
+			),
+			io.write_string(").\n", !IO)
+		  else 
+            % has no params, so should use no regions either.
+		  	true
+		)
+	),
+	list.foldl(pp_proc_declaration(HLDS, PredId, ConstantRTable, DeadRTable,
+		BornRTable), ProcIds, !IO).
+
+	% print out :- mode ...
+	%
+:- pred pp_proc_declaration(module_info::in, pred_id::in, 
+    proc_region_set_table::in, proc_region_set_table::in, 
+    proc_region_set_table::in, proc_id::in, io::di, io::uo) is det.
+
+pp_proc_declaration(HLDS, PredId, ConstantRTable, DeadRTable, BornRTable, 
+    ProcId, !IO) :-
+	PredProcId = proc(PredId, ProcId),
+	( if
+		smm_utils.some_are_special_preds([PredProcId], HLDS)
+	then
+		true
+	else
+		io.write_string(":- mode ", !IO),
+		module_info_pred_proc_info(HLDS, PredProcId, PredInfo, ProcInfo),
+		PredName = pred_info_name(PredInfo),
+		io.write_string(PredName, !IO),
+		
+		( if
+			PredName = "main"
+		  then
+            % This should not be necessary now because
+            % we can deal with di and uo modes.
+		  	io.write_string("(di, uo)", !IO)
+		  else
+			hlds.hlds_pred.proc_info_get_argmodes(ProcInfo, Modes),
+		  	hlds.hlds_pred.proc_info_get_headvars(ProcInfo, 
+				HeadVars),
+			hlds.hlds_pred.proc_info_get_vartypes(ProcInfo, 
+				VarTypes),
+			modes_to_string(Modes, HLDS, HeadVars, VarTypes, "", S),
+			( if 
+				S \= ""
+			  then	
+				io.write_string("(", !IO),
+				io.write_string(S, !IO),
+				pp_reg_param_modes(PredProcId, ConstantRTable, DeadRTable, 
+					BornRTable, yes, !IO),
+				io.write_string(")", !IO)
+			  else
+				pp_reg_param_modes(PredProcId, ConstantRTable,
+					DeadRTable, BornRTable, no, !IO)
+			)
+		),	
+		
+		io.write_string(" is ", !IO),
+		proc_info_get_inferred_determinism(ProcInfo, Determinism),
+		io.write_string(determinism_to_string(Determinism), !IO),
+		io.write_string(".\n", !IO)
+	).
+
+	% print "int" as the type of regions,
+	% if there is at least 1 region, print ", " to separate with regular
+	% parameters
+	%
+:- pred pp_reg_param_types(pred_proc_id::in, proc_region_set_table::in, 
+    proc_region_set_table::in, proc_region_set_table::in, io::di, io::uo) 
+    is det.
+
+pp_reg_param_types(PredProcId, ConstantRTable, DeadRTable, BornRTable, 
+    !IO) :-
+	map.lookup(ConstantRTable, PredProcId, ConstantR),
+	map.lookup(DeadRTable, PredProcId, DeadR),
+	map.lookup(BornRTable, PredProcId, BornR),
+	set.to_sorted_list(ConstantR, LConstantR),
+	set.to_sorted_list(DeadR, LDeadR),
+	set.to_sorted_list(BornR, LBornR),
+	list.append(LConstantR, LDeadR, L0),
+	list.append(L0, LBornR, L),
+	( if
+		L \= []
+	  then % has at least 1 region
+	  	io.write_string(",", !IO),
+		write_list(L, ", ", print_string("int"), !IO)
+	  else
+	  	true
+	).
+
+    % A bit difficult to find a good name for this predicate. It just
prints 
+    % out the provided string but 
+	%
+:- pred print_string(string::in, rptg_node::in, io::di, io::uo) is det.
+print_string(S, _L, !IO) :- 
+	io.write_string(S, !IO).
+
+	% print in/out, ..., in/out, according to the modes 
+	%
+:- pred modes_to_string(list(mer_mode)::in, module_info::in, 
+	list(prog_var)::in, vartypes::in,
+	string::in, string::out) is det.
+modes_to_string([], _HLDS, _HeadVars, _VarTypes, S, S).
+modes_to_string([M|Ms], HLDS, HeadVars, VarTypes, S0, S) :-
+	( if
+		check_hlds.mode_util.mode_is_input(HLDS, M)
+	  then
+		( if
+			HeadVars = [V|_Vs]
+		  then
+			map.lookup(VarTypes, V, VType),
+			( if
+				type_is_io_state(VType)
+			  then
+				S1 = S0 ++ "di, "
+			  else
+				S1 = S0 ++ "in, "
+			)
+		  else
+			unexpected(thisfile, "mode_to_string: impossible")
+		)
+	  else
+		( if
+			HeadVars = [V|_Vs]
+		  then
+			map.lookup(VarTypes, V, VType),
+			( if
+				type_is_io_state(VType)
+			  then
+				S1 = S0 ++ "uo, "
+			  else
+				S1 = S0 ++ "out, "
+			)
+		  else
+			unexpected(thisfile, "mode_to_string: impossible")
+		)
+	),
+	( if
+		Ms = []
+	  then
+	  	S = string.left(S1, string.length(S1) - 2)
+	  else
+		HeadVars = [_V|Vs],
+		modes_to_string(Ms, HLDS, Vs, VarTypes, S1, S)
+	).
+	
+	% print in, ..., in, out, ..., out
+	%	
+:- pred pp_reg_param_modes(pred_proc_id::in, proc_region_set_table::in, 
+    proc_region_set_table::in, proc_region_set_table::in, bool::in,
io::di, 
+    io::uo) is det.
+
+pp_reg_param_modes(PredProcId, ConstantRTable, DeadRTable, BornRTable, 
+    HasParam, !IO) :-
+	map.lookup(ConstantRTable, PredProcId, ConstantR),
+	map.lookup(DeadRTable, PredProcId, DeadR),
+	map.lookup(BornRTable, PredProcId, BornR),
+	set.to_sorted_list(ConstantR, LConstantR),
+	set.to_sorted_list(DeadR, LDeadR),
+	set.to_sorted_list(BornR, Outs),
+	list.append(LConstantR, LDeadR, Ins),
+	( if
+		HasParam = yes
+	  then
+	 	( if
+			(Ins = [],Outs=[])
+		  then 
+            % the proc has no region params
+			true
+		  else 
+	  		( if
+				Ins \= []
+			  then
+				io.write_string(",", !IO),
+				io.write_list(Ins, ", ", print_string("in"), !IO)
+			  else
+				true
+			),
+			( if
+				Outs \= []
+			  then
+				io.write_string(",", !IO),
+				io.write_list(Outs, ", ", print_string("out"), !IO)
+			  else
+				true
+			)
+		)
+	  else	
+        % the proc has no regular param, so take care everything
+	  	( if
+			(Ins = [],Outs=[])
+		  then  
+            % the proc has no region params either, no print
+			true
+		  else
+		  	io.write_string("(", !IO),
+			io.write_list(Ins, ", ", print_string("in"), !IO),
+			( if
+				Outs \= []
+			  then
+				io.write_string(",", !IO),
+				io.write_list(Outs, ", ", print_string("out"), !IO)
+			  else
+				true
+			),
+			io.write_string(")", !IO)
+       		)
+	).
+
+
+:- pred pp_region_annotations_pred(module_info::in, annotation_table::in, 
+	rpta_info_table::in, proc_region_set_table::in,
proc_region_set_table::in, 
+    proc_region_set_table::in, pred_id::in, io::di, io::uo) is det.
+pp_region_annotations_pred(HLDS, AnnotationTable, RptaInfoTable, 
+    ConstantRTable, DeadRTable, BornRTable, PredId, !IO) :-
+	module_info_pred_info(HLDS, PredId, PredInfo),
+	pred_info_non_imported_procids(PredInfo) = ProcIds,
+	list.foldl(pp_region_annotations_proc(HLDS, PredId, AnnotationTable, 
+		RptaInfoTable, ConstantRTable, DeadRTable, BornRTable), ProcIds, !IO).
+
+	% generate Mercury code for a procedure with region support
+	% - region parameters are at the end of the parameter list
+	% - create(R), remove(R), make_ite_frame, commit_removal, undo_ite_changes
+	%	
+:- pred pp_region_annotations_proc(module_info::in, pred_id::in, 
+    annotation_table::in, rpta_info_table::in, proc_region_set_table::in, 
+    proc_region_set_table::in, proc_region_set_table::in, proc_id::in,
+    io::di, io::uo) is det.
+pp_region_annotations_proc(HLDS, PredId, AnnotationTable, RptaInfoTable, 
+    ConstantRTable, DeadRTable, BornRTable, ProcId, !IO) :-
+	PredProcId = proc(PredId, ProcId),
+	( if
+		smm_utils.some_are_special_preds([PredProcId], HLDS)
+	then
+		true
+	else
+		module_info_pred_proc_info(HLDS, PredProcId, PredInfo, ProcInfo),
+		
+        % Print out all the annotations in this procedure	
+        %
+
+        % First type and mode declarations 
+        %	
+		pred_info_name(PredInfo) = Name,
+		io.write_string(Name, !IO),
+		
+        % print out argument list: normal arguments and region arguments
+		proc_info_get_headvars(ProcInfo, HeadVars),
+		( if
+			HeadVars \= []
+		  then
+			io.write_string("(", !IO),
+
+			% ... non-region parameters
+			pp_proc_headvars(ProcInfo, !IO),
+			
+			% print formal region parameters in the order: constantR, deadR,
+			% bornR
+			%
+			( if 
+				Name = "main"
+	  		  then
+				% 'main' has no region parameters
+				true
+			  else
+				pp_formal_region_parameters(PredProcId, RptaInfoTable,
+					ConstantRTable, DeadRTable, BornRTable, yes, !IO)
+			),
+			io.write_string(")", !IO)
+		  else
+			% this procedure has no non-region parameters (note that it
+			% cannot be 'main', so no need to consider 'main' here.
+			%
+			pp_formal_region_parameters(PredProcId, RptaInfoTable,
+				ConstantRTable, DeadRTable, BornRTable, no, !IO)
+		),
+		
+		io.write_string(" :-", !IO),
+		io.write_string("\n", !IO),
+			
+		( if 
+			Name = "main"
+		  then
+			io.write_string("\tinit,\n", !IO)
+		  else 
+			true
+		),
+		
+		% to fill out program point infomation
+		fill_goal_path_slots(HLDS, ProcInfo, ProcInfo1),
+		proc_info_get_goal(ProcInfo1, Goal),
+		map.lookup(AnnotationTable, PredProcId, AnnotationProc),
+		pp_region_annotations_goal(HLDS, PredInfo, ProcInfo1, PredProcId, 
+			AnnotationProc, RptaInfoTable, ConstantRTable, DeadRTable, 
+            BornRTable, "\t", Goal, !IO),
+		( if 
+			Name = "main"
+		  then
+			io.write_string(",\n\tprint_profile_data", !IO)
+		  else 
+			true
+		),
+		io.write_string(".\n\n", !IO)
+	).
+		
+	% generate Mercury code for a goal. The goal is aligned with Tabs
+	%
+:- pred pp_region_annotations_goal(module_info::in, pred_info::in, 
+    proc_info::in, pred_proc_id::in, annotation_proc::in,
rpta_info_table::in, 
+    proc_region_set_table::in, proc_region_set_table::in, 
+    proc_region_set_table::in, string::in, hlds_goal::in, io::di, io::uo) 
+    is det.
+
+pp_region_annotations_goal(HLDS, PredInfo, ProcInfo, ProcId_p,
AnnotationProc, 
+    RptaInfoTable, ConstantRTable, DeadRTable, BornRTable, Tabs, Goal,
!IO) :-
+	Goal = hlds_goal(Expr, Info),
+	(
+		goal_is_atomic(Expr)
+	->
+        % We want to print out the annotations, if any so find them in 
+        % the Annotation table of the procedure
+		program_point_init(Info, PP),
+		( if 
+			map.search(AnnotationProc, PP, before_after(Before, After))
+		  then
+            % print out annotations before the goal
+			prepare_annotations(Before, [], InstructionsBefore),
+			print_annotation_list(InstructionsBefore, Tabs, HLDS, PredInfo, 
+                ProcInfo, !IO),
+			( if
+				InstructionsBefore = []
+			  then
+				true
+			  else
+				io.write_string(",\n", !IO)
+			),
+
+            % print out the goal itself
+			pp_atomic_goal(HLDS, ProcInfo, Goal, ProcId_p, PP, 
+				RptaInfoTable, ConstantRTable, DeadRTable, 
+				BornRTable, Tabs, !IO),
+				
+            % if the goal is a construction then we may need to add
+            % ralloc
+			add_ralloc(Expr, Tabs, RptaInfoTable, ProcId_p, !IO),
+            % print out annotations after the goal
+			( if
+				After = []
+			  then
+				true
+			  else
+				io.write_string(",\n", !IO),
+				prepare_annotations(After, [], InstructionsAfter),
+				print_annotation_list(InstructionsAfter, 
+							Tabs, HLDS, PredInfo, ProcInfo, !IO)
+			)
+		  else
+		  	pp_atomic_goal(HLDS, ProcInfo, Goal, ProcId_p, PP, 
+				RptaInfoTable, ConstantRTable, DeadRTable, 
+				BornRTable, Tabs, !IO),
+			add_ralloc(Expr, Tabs, RptaInfoTable, ProcId_p, !IO)
+		)
+	;
+		pp_region_annotations_compound_goal(HLDS, PredInfo, ProcInfo, 
+            ProcId_p, AnnotationProc, RptaInfoTable, ConstantRTable,
+			DeadRTable, BornRTable, Tabs, Goal, !IO)
+	).
+
+	% pretty printing a compound goal, aligned with Tabs
+	%
+:- pred pp_region_annotations_compound_goal(module_info::in,
pred_info::in, 
+    proc_info::in, pred_proc_id::in, annotation_proc::in,
rpta_info_table::in, 
+    proc_region_set_table::in, proc_region_set_table::in, 
+    proc_region_set_table::in, string::in, hlds_goal::in, io::di, io::uo) 
+    is det.
+pp_region_annotations_compound_goal(HLDS, PredInfo, ProcInfo, ProcId_p,
+    AnnotationProc, RptaInfoTable, ConstantRTable, DeadRTable, BornRTable, 
+    Tabs, CompoundGoal, !IO) :- 
+	CompoundGoal = hlds_goal(Expr, _Info),
+	(
+		Expr = conj(_, Goals)
+	->
+		pp_region_annotations_conj(Goals, HLDS, PredInfo, ProcInfo, ProcId_p,
+			AnnotationProc, RptaInfoTable, ConstantRTable,
+			DeadRTable, BornRTable, Tabs, !IO)
+	;
+		Expr = switch(_Var, _B, Cases)
+	->
+		io.write_string(Tabs, !IO),
+		io.write_string("(\n", !IO),
+		pp_region_annotations_cases(Cases, CompoundGoal, HLDS, PredInfo,
+			ProcInfo, ProcId_p, AnnotationProc, RptaInfoTable,
+			ConstantRTable, DeadRTable, BornRTable, Tabs, !IO),
+		io.write_string(Tabs, !IO),
+		io.write_string(")", !IO)
+	;
+		Expr = disj(Goals)
+	->
+		io.write_string(Tabs, !IO),
+		io.write_string("(\n", !IO),
+		pp_region_annotations_disj(Goals, HLDS, PredInfo, ProcInfo, ProcId_p,
+			AnnotationProc, RptaInfoTable, ConstantRTable,
+			DeadRTable, BornRTable, Tabs, !IO),
+		io.write_string(Tabs, !IO),
+		io.write_string(")", !IO)
+	;
+		Expr = negation(Goal)
+	->
+		io.write_string(Tabs, !IO),
+		io.write_string("not ", !IO),
+		pp_region_annotations_goal(HLDS, PredInfo, ProcInfo, ProcId_p,
+			AnnotationProc, RptaInfoTable, ConstantRTable,
+			DeadRTable, BornRTable, Tabs, Goal, !IO)
+	;
+		Expr = scope(_Reason,  Goal)
+	->
+		io.write_string(Tabs, !IO),
+		pp_region_annotations_goal(HLDS, PredInfo, ProcInfo, ProcId_p,
+			AnnotationProc, RptaInfoTable, ConstantRTable,
+			DeadRTable, BornRTable, Tabs, Goal, !IO)
+	;
+		Expr = if_then_else(_V, Cond, Then, Else)
+	->
+		( if 
+			Else = hlds_goal(if_then_else(_, _, _, _), _InfoExpr)
+		  then
+            % for the case where if cond1 then (1) else if cond2 then (2)
+            % we print out using the format 
+            %	(
+            %		cond1
+            %	->
+            %		(1)
+            %	;
+            %		cond2
+            %	->	
+            %		(2)
+            %	; 
+            %	...
+		  	io.write_string(Tabs, !IO),
+			io.write_string("(\n", !IO),
+			
+			string.append(Tabs, "\t", NewTabs),
+			
+			print_make_ite_frame(NewTabs, !IO),
+			io.write_string(",\n", !IO),
+				
+            % print out the condition itself
+			pp_region_annotations_goal(HLDS, PredInfo, ProcInfo, ProcId_p,
+				AnnotationProc, RptaInfoTable, ConstantRTable, DeadRTable, 
+                BornRTable, NewTabs, Cond, !IO),
+			io.write_string("\n", !IO),
+		
+			io.write_string(Tabs, !IO),
+			io.write_string("->\n", !IO),
+			
+			print_commit_removal(NewTabs, !IO),
+			io.write_string(",\n", !IO),
+				
+            % the Then goal itself
+			pp_region_annotations_goal(HLDS, PredInfo, ProcInfo, ProcId_p,
+				AnnotationProc, RptaInfoTable, ConstantRTable, DeadRTable, 
+                BornRTable, NewTabs, Then, !IO),
+			io.write_string("\n", !IO),
+		
+			io.write_string(Tabs, !IO),
+			io.write_string(";\n", !IO),
+			
+			print_undo_ite_changes(NewTabs, !IO),
+			io.write_string(",\n", !IO),
+						
+            % print out the Else goal, which is also an if-then-else.
+            % Tabs is input instead of NewTabs because the next ";", "->"
+            % are aligned with Tabs.
+			pp_region_annotations_inner_else(HLDS, PredInfo, ProcInfo, 
+                ProcId_p, AnnotationProc, RptaInfoTable, ConstantRTable,
+				DeadRTable, BornRTable, Tabs, Else, !IO),
+			io.write_string("\n", !IO),
+			io.write_string(Tabs, !IO),
+			io.write_string(")", !IO)
+		  else
+		  		% using the normal if ... then ... else format. 
+				%
+		  	io.write_string(Tabs, !IO),
+			io.write_string("( if\n", !IO),
+			string.append(Tabs, "\t", NewTabs),
+			
+			print_make_ite_frame(NewTabs, !IO),
+			io.write_string(",\n", !IO),
+			
+			pp_region_annotations_goal(HLDS, PredInfo, ProcInfo, ProcId_p,
+				AnnotationProc, RptaInfoTable, ConstantRTable,
+				DeadRTable, BornRTable, NewTabs, Cond, !IO),
+			io.write_string("\n", !IO),
+		
+			io.write_string(Tabs, !IO),
+			io.write_string("  then\n", !IO),
+			
+			print_commit_removal(NewTabs, !IO),
+			io.write_string(",\n", !IO),
+			
+			pp_region_annotations_goal(HLDS, PredInfo, ProcInfo, ProcId_p,
+				AnnotationProc, RptaInfoTable, ConstantRTable,
+				DeadRTable, BornRTable, NewTabs, Then, !IO),
+			io.write_string("\n", !IO),
+		
+			io.write_string(Tabs, !IO),
+			io.write_string("  else\n", !IO),
+			
+			print_undo_ite_changes(NewTabs, !IO),
+			io.write_string(",\n", !IO),
+			
+			pp_region_annotations_goal(HLDS, PredInfo, ProcInfo, ProcId_p,
+				AnnotationProc, RptaInfoTable, ConstantRTable,
+				DeadRTable, BornRTable, NewTabs, Else, !IO),
+			io.write_string("\n", !IO),
+			io.write_string(Tabs, !IO),
+			io.write_string(")", !IO)
+		)
+	;
+		unexpected(thisfile, 
+            "pp_region_annotations_compound_goal: found atomic goal")
+	). 
+	
+	% this predicate is to print out an else branch that is an if-then-else
+	%
+:- pred pp_region_annotations_inner_else(module_info::in, pred_info::in, 
+    proc_info::in, pred_proc_id::in, annotation_proc::in,
rpta_info_table::in, 
+    proc_region_set_table::in, proc_region_set_table::in, 
+    proc_region_set_table::in, string::in, hlds_goal::in, io::di, io::uo) 
+    is det.
+
+pp_region_annotations_inner_else(HLDS, PredInfo, ProcInfo, ProcId_p, 
+    AnnotationProc, RptaInfoTable, ConstantRTable, DeadRTable, BornRTable, 
+    Tabs, InnerElse, !IO) :-
+	( if 
+        InnerElse = hlds_goal(if_then_else(_V, Cond, Then, Else), _Info)
+	  then
+        % the condition part is printed right below the ";"
+		string.append(Tabs, "\t", NewTabs),
+		
+		print_make_ite_frame(NewTabs, !IO),
+		io.write_string(",\n", !IO),
+		
+		pp_region_annotations_goal(HLDS, PredInfo, ProcInfo, ProcId_p,
+			AnnotationProc, RptaInfoTable, ConstantRTable,
+			DeadRTable, BornRTable, NewTabs, Cond, !IO),
+		io.write_string("\n", !IO),
+			
+        % ... the then part follows
+		io.write_string(Tabs, !IO),
+		io.write_string("->\n", !IO),
+		
+		print_commit_removal(NewTabs, !IO),
+		io.write_string(",\n", !IO),
+		
+		pp_region_annotations_goal(HLDS, PredInfo, ProcInfo, ProcId_p,
+			AnnotationProc, RptaInfoTable, ConstantRTable,
+			DeadRTable, BornRTable, NewTabs, Then, !IO),
+		io.write_string("\n", !IO),
+
+        % ... the else part: if this again is an if-then-else, call this
+        % predicate. 
+        % Otherwise call the general predicate for a goal
+		io.write_string(Tabs, !IO),
+		io.write_string(";\n", !IO),
+		
+		print_undo_ite_changes(NewTabs, !IO),
+		io.write_string(",\n", !IO),
+		
+		(if 
+			Else = hlds_goal(if_then_else(_, _, _, _), _InfoExpr)
+		  then
+			pp_region_annotations_inner_else(HLDS, PredInfo, ProcInfo, 
+                ProcId_p, AnnotationProc, RptaInfoTable, ConstantRTable,
+				DeadRTable, BornRTable, Tabs, Else, !IO)
+		  else
+			pp_region_annotations_goal(HLDS, PredInfo, ProcInfo, ProcId_p,
+				AnnotationProc, RptaInfoTable, ConstantRTable,
+				DeadRTable, BornRTable, NewTabs, Else, !IO)
+		)
+	  else
+	  	unexpected(thisfile, "pp_region_annotations_inner_else: impossible")
+	).
+
+:- pred pp_region_annotations_conj(list(hlds_goal)::in, module_info::in, 
+    pred_info::in, proc_info::in, pred_proc_id::in, annotation_proc::in, 
+    rpta_info_table::in, proc_region_set_table::in,
proc_region_set_table::in, 
+    proc_region_set_table::in, string::in, io::di, io::uo) is det.
+
+pp_region_annotations_conj([], _HLDS, _PredInfo, _ProcInfo, _ProcId_p,
+    _AnnotationProc, _RptaInfoTable, _ConstantRTable, _DeadRTable, 
+    _BornRTable, _Tabs, !IO).
+pp_region_annotations_conj([Goal|Others], HLDS, PredInfo, ProcInfo,
ProcId_p,
+    AnnotationProc, RptaInfoTable, ConstantRTable, DeadRTable, BornRTable, 
+    Tabs, !IO) :-
+    % execution paths covered this Goal
+	pp_region_annotations_goal(HLDS, PredInfo, ProcInfo, ProcId_p,
+		AnnotationProc, RptaInfoTable, ConstantRTable, DeadRTable,
+		BornRTable, Tabs, Goal, !IO),
+	( if 
+		Others = []
+	  then
+	  	true
+	  else
+	  	io.write_string(",\n", !IO)
+	),
+    % continue with the other goals in the conj
+	pp_region_annotations_conj(Others, HLDS, PredInfo, ProcInfo, ProcId_p,
+        AnnotationProc, RptaInfoTable, ConstantRTable, DeadRTable,
BornRTable, 
+        Tabs, !IO).
+	
+:- pred pp_region_annotations_disj(list(hlds_goal)::in, module_info::in, 
+    pred_info::in, proc_info::in, pred_proc_id::in, annotation_proc::in, 
+    rpta_info_table::in, proc_region_set_table::in,
proc_region_set_table::in, 
+    proc_region_set_table::in, string::in, io::di, io::uo) is det.
+pp_region_annotations_disj([], _HLDS, _PredInfo, _ProcInfo, _ProcId_p,
+    _AnnotationProc, _RptaInfoTable, _ConstantRTable, _DeadRTable, 
+    _BornRTable, _Tabs, !IO).
+pp_region_annotations_disj([Goal|Others], HLDS, PredInfo, ProcInfo,
ProcId_p,
+    AnnotationProc, RptaInfoTable, ConstantRTable, DeadRTable, BornRTable, 
+    Tabs, !IO) :- 
+    % collect execution paths covered this goal
+	string.append(Tabs, "\t", NewTabs),
+	pp_region_annotations_goal(HLDS, PredInfo, ProcInfo, ProcId_p,
+        AnnotationProc, RptaInfoTable, ConstantRTable, DeadRTable, 
+        BornRTable, NewTabs, Goal, !IO),
+	io.write_string("\n", !IO),
+	
+	( if 
+		Others = []
+	  then
+	  	true
+	  else
+		io.write_string(Tabs, !IO),
+		io.write_string(";\n", !IO)
+	),
+    % collect execution paths covered the other goals (branches)
+	pp_region_annotations_disj(Others, HLDS, PredInfo, ProcInfo, ProcId_p,
+		AnnotationProc, RptaInfoTable, ConstantRTable, DeadRTable, BornRTable, 
+        Tabs, !IO).
+
+:- pred pp_region_annotations_cases(list(case)::in, hlds_goal::in, 
+    module_info::in, pred_info::in, proc_info::in, pred_proc_id::in, 
+    annotation_proc::in, rpta_info_table::in, proc_region_set_table::in, 
+    proc_region_set_table::in, proc_region_set_table::in, string::in, 
+    io::di, io::uo) is det.
+pp_region_annotations_cases([], _Switch,_HLDS, _PredInfo, _ProcInfo, 
+    _ProcId_p, _AnnotationProc, _RptaInfoTable, _ConstantRTable,
_DeadRTable, 
+    _BornRTable, _Tabs, !IO).
+pp_region_annotations_cases([Case|Others], Switch, HLDS, PredInfo,
ProcInfo,
+    ProcId_p, AnnotationProc, RptaInfoTable, ConstantRTable, DeadRTable, 
+    BornRTable, Tabs, !IO) :-
+	string.append(Tabs, "\t", NewTabs),
+	Case = case(ConsId, Goal),
+    % print the removed unification if MMC removed
+    % its annotations if any also printed.
+	( if
+		(
+			(ConsId = cons(_SymName, Arity), Arity = 0)
+		;
+			ConsId = int_const(_Int)
+		;
+			ConsId = string_const(_String)
+		;
+			ConsId = float_const(_Float)
+		)
+	  then
+		( if
+			Switch = hlds_goal(switch(Var, _B, _Cases), Info)
+		  then	
+			program_point_init(Info, PP),
+			( if 
+				map.search(AnnotationProc, PP, before_after(Before, After))
+			then
+				prepare_annotations(Before, [], InstructionsBefore),
+				print_annotation_list(InstructionsBefore, 
+							NewTabs, HLDS, PredInfo, ProcInfo, !IO),
+				( if
+					Before = []
+				then
+					true
+				else
+					io.write_string(",\n", !IO)
+				),				
+				io.write_string(NewTabs, !IO),
+				print_removed_deconstruction(Var, Case, HLDS, ProcInfo, !IO),	
+				( if
+					After = []
+				then
+					true
+				else
+					io.write_string(",\n", !IO),
+					prepare_annotations(After, [], InstructionsAfter),
+					print_annotation_list(InstructionsAfter, 
+								NewTabs, HLDS, PredInfo, ProcInfo, !IO)
+				)
+			else
+				io.write_string(NewTabs, !IO),
+				print_removed_deconstruction(Var, Case, HLDS, ProcInfo, !IO)
+			),
+			io.write_string(",\n", !IO)
+		  else
+		  	unexpected(thisfile, "pp_region_annotation_cases: impossible")
+		)  	
+	  else
+		true
+	),	  
+    % print this branch
+	pp_region_annotations_goal(HLDS, PredInfo, ProcInfo, ProcId_p,
+        AnnotationProc, RptaInfoTable, ConstantRTable, DeadRTable,
BornRTable, 
+        NewTabs, Goal, !IO),
+	io.write_string("\n", !IO),
+	
+	( if 
+		Others = []
+	  then
+		true
+	  else
+		io.write_string(Tabs, !IO),
+		io.write_string(";\n", !IO)
+	),
+    % print other branches
+	pp_region_annotations_cases(Others, Switch, HLDS, PredInfo, ProcInfo,
+        ProcId_p, AnnotationProc, RptaInfoTable, ConstantRTable,
DeadRTable, 
+        BornRTable, Tabs, !IO).
+
+:- pred print_removed_deconstruction(prog_var::in, case::in,
module_info::in, 
+    proc_info::in, io::di, io::uo) is det.
+
+print_removed_deconstruction(Var, Case, HLDS, ProcInfo, !IO) :-
+	proc_info_get_varset(ProcInfo, Varset),
+	mercury_var_to_string(Varset, no, Var) = VarStr,
+	Case = case(ConsId, _Goal),
+	(
+		ConsId = cons(_SymName, Arity)
+	->
+		( if
+			Arity = 0
+		  then	
+			io.write_string(VarStr, !IO),
+			io.write_string(" = ", !IO),
+			hlds_out.write_functor_cons_id(ConsId, [], Varset, HLDS, no, !IO)
+		  else
+			true
+		)
+	; 
+		( ConsId = int_const(_Int);
+		  ConsId = string_const(_String); 
+		  ConsId = float_const(_Float)
+		)
+	->
+		io.write_string(VarStr, !IO),
+		io.write_string(" = ", !IO),
+		hlds_out.write_cons_id(ConsId, !IO)
+	;
+		unexpected(thisfile, "pp_region_annotations_cases: cons_id unexpected")
+	).
+	
+:- pred print_annotation_list(list(string)::in, string::in,
module_info::in, 
+    pred_info::in, proc_info::in, io::di, io::uo) is det.
+print_annotation_list([], _Tabs, _HLDS, _PredInfo, _ProcInfo, !IO).
+print_annotation_list([Instruction|Others], Tabs, HLDS, PredInfo,
ProcInfo, 
+    !IO) :-
+	io.write_string(Tabs, !IO),
+	
+    % build the string to print out
+	pred_info_name(PredInfo) = _PredNameString,
+	(
+		string.sub_string_search(Instruction, "keep", _Index1)
+	->
+		unexpected(thisfile, "There should be no keep instructions")
+	;
+		string.sub_string_search(Instruction, "use", _Index2)
+	->
+        % there is no "use" instruction here.
+		unexpected(thisfile, "There should be no use instructions")
+	;	
+		string.sub_string_search(Instruction, "conditionally_create", _Index3)
+	->
+		unexpected(thisfile, "There should be no conditionally_create "
+            ++ "instructions")
+	;	
+		string.substring(Instruction, 0, 6, Kind),
+		string.substring(Instruction, 6, string.length(Instruction) - 6,
+            RegionName),
+		string.format("%s(%s)", [s(Kind), s(RegionName)], S)
+	),
+	
+	io.write_string(S, !IO),
+	( if
+		Others = []
+	  then
+		true
+	  else
+		io.write_string(",\n", !IO)
+	),
+	print_annotation_list(Others, Tabs, HLDS, PredInfo, ProcInfo, !IO).
+
+	% return a list of region annotations that are ready to be print out.
+    %
+:- pred prepare_annotations(list(string)::in, list(string)::in, 
+    list(string)::out) is det.
+
+prepare_annotations([], Results, Results).
+prepare_annotations([Annotation|Others], Results0, Results) :-
+	(
+		string.sub_string_search(Annotation, "keep", Index1)
+	->
+		Kind = "keep",
+		Index = Index1 + 4
+	;
+		string.sub_string_search(Annotation, "use", Index2)
+	->
+		Kind = "use",
+		Index = Index2 + 3	
+	;
+		string.sub_string_search(Annotation, "conditionally_create", Index3)
+	->
+		Kind = "conditionally_create",
+		Index = Index3 + 21
+	;
+		string.substring(Annotation, 4, 6, Kind),
+		Index = 11
+	),
+	string.substring(Annotation, Index, string.length(Annotation) - Index, 
+        RegionName),
+	Instruction = Kind ++ RegionName,
+	( if
+		list.member(Instruction, Results0)
+	  then
+		prepare_annotations(Others, Results0, Results)
+	  else
+		prepare_annotations(Others, [Instruction|Results0], Results)
+	).
+
+:- pred pp_proc_headvars(proc_info::in, io::di, io::uo) is det.
+
+pp_proc_headvars(ProcInfo, !IO) :-
+	proc_info_get_varset(ProcInfo, VarSet),
+	PrintVar = (pred(VarName::in, IOState0::di, IOState::uo) is det :-
+		term_io.write_variable(VarName, VarSet, IOState0, IOState)
+	),
+	proc_info_get_headvars(ProcInfo, HeadVars),
+	io.write_list(HeadVars, ", ", PrintVar, !IO).
+
+:- pred pp_proc_args(proc_info::in, list(prog_var)::in, io::di, io::uo)
is det.
+pp_proc_args(ProcInfo, Args, !IO) :-
+	proc_info_get_varset(ProcInfo, VarSet),
+	PrintVar = (pred(Var::in, IOState0::di, IOState::uo) is det :-
+		mercury_var_to_string(VarSet, no, Var) = VarName,
+		string.substring(VarName, 0, 1, FirstChar),
+		string.substring(VarName, 2, 1, ThirdChar),
+		( if
+			FirstChar = "_"
+		  then
+			( if
+			  	string.to_int(ThirdChar, _Int)
+			  then 
+				io.write_string("V", IOState0, IOState1)
+			  else
+				io.write_string("", IOState0, IOState1)
+			)
+		  else
+			IOState1 = IOState0
+		),
+		io.write_string(VarName, IOState1, IOState)
+	),
+	io.write_list(Args, ", ", PrintVar, !IO).
+	
+	% print out the atomic goal, aligned with Tabs 
+	% note: the expr is in the proc	
+	%
+:- pred pp_atomic_goal(module_info::in, proc_info::in, hlds_goal::in, 
+    pred_proc_id::in, program_point::in, rpta_info_table::in, 
+    proc_region_set_table::in, proc_region_set_table::in,
+    proc_region_set_table::in, string::in, io::di, io::uo) is det.
+
+pp_atomic_goal(HLDS, ProcInfo, Goal, ProcId_p, PP, RptaInfoTable,
+    ConstantRTable, DeadRTable, BornRTable, Tabs, !IO) :- 
+	io.write_string(Tabs, !IO),
+	Goal = hlds_goal(Expr, _Info),
+	( if
+		Expr = plain_call(PredId, ProcId, Args, _Builtin_State, 
+            _Unify_Context, Name)
+	  then
+		% expr is a procedure call
+		ProcId_q = proc(PredId, ProcId),
+		( if 
+			smm_utils.some_are_special_preds([ProcId_q], HLDS)
+		  then
+		  	proc_info_get_varset(ProcInfo, Varset),
+			hlds_out.write_goal(Goal, HLDS, Varset, no, 0, "", !IO)
+		  else		
+			mdbcomp.prim_data.sym_name_to_string(Name) = NameString,
+			io.write_string(NameString, !IO),
+
+			module_info_pred_proc_info(HLDS, ProcId_q, _PredInfo_q, 
+                ProcInfo_q),
+			proc_info_get_headvars(ProcInfo_q, HeadVars),
+			( if
+                HeadVars \= []
+			  then
+				io.write_string("(", !IO),
+				pp_proc_args(ProcInfo, Args, !IO),
+			
+				% print out actual region parameters
+				pp_actual_region_parameters(ProcId_p, PP, ProcId_q, 
+					RptaInfoTable, ConstantRTable, DeadRTable, 
+					BornRTable, yes, !IO),
+				io.write_string(")", !IO)
+			  else
+				pp_actual_region_parameters(ProcId_p, PP, ProcId_q, 
+					RptaInfoTable, ConstantRTable, DeadRTable, 
+					BornRTable, no, !IO)
+			)
+		)	
+	  else	
+		proc_info_get_varset(ProcInfo, Varset),
+		hlds_out.write_goal(Goal, HLDS, Varset, no, 0, "", !IO)
+	).
+
+	% print out formal region parameters.
+	%
+:- pred pp_formal_region_parameters(pred_proc_id::in, rpta_info_table::in, 
+	proc_region_set_table::in, proc_region_set_table::in, 
+    proc_region_set_table::in, bool::in, io::di, io::uo) is det.
+
+pp_formal_region_parameters(PredProcId, RptaInfoTable, ConstantRTable, 
+    DeadRTable, BornRTable, HasParam, !IO) :-
+	map.lookup(RptaInfoTable, PredProcId, RptaInfo),
+	RptaInfo = rpta_info(Graph, _Alpha),
+	map.lookup(ConstantRTable, PredProcId, ConstantR),
+	map.lookup(DeadRTable, PredProcId, DeadR),
+	map.lookup(BornRTable, PredProcId, BornR),
+
+    % constant regions
+	set.to_sorted_list(ConstantR, LConstantR),
+	set.to_sorted_list(DeadR, LDeadR),
+	set.to_sorted_list(BornR, LBornR),
+	list.append(LConstantR, LDeadR, L0),
+	list.append(L0, LBornR, L),
+	( if
+		HasParam = yes
+	  then
+	  	( if 
+			L \= []
+		  then
+			io.write_string(",", !IO),
+			write_list(L, ",", print_region_name(Graph), !IO)
+		  else
+			true
+		)
+	  else
+	  	( if 
+			L \= []
+		  then
+			io.write_string("(", !IO),
+			write_list(L, ",", print_region_name(Graph), !IO),
+			io.write_string(")", !IO)
+		  else
+			true
+		)		
+	).
+
+:- pred print_region_name(rpt_graph::in, rptg_node::in, io::di, io::uo) 
+    is det.
+print_region_name(Graph, Node, !IO) :-
+	write_string(rptg_lookup_region_name(Graph, Node), !IO).
+	
+	% print out bool parameters corresponding to regions in bornR set.
+	%
+:- pred pp_formal_bool_parameters(pred_proc_id::in, rpta_info_table::in, 
+    proc_region_set_table::in, io::di, io::uo) is det.
+
+pp_formal_bool_parameters(PredProcId, RptaInfoTable, BornRTable, !IO) :-
+	map.lookup(RptaInfoTable, PredProcId, RptaInfo),
+	RptaInfo = rpta_info(Graph, _Alpha),
+	map.lookup(BornRTable, PredProcId, BornR),
+	write_string("<", !IO),
+
+    % born regions
+	set.to_sorted_list(BornR, LBornR),
+	write_list(LBornR, ",", print_bool_name(Graph), !IO),
+	
+	write_string(">", !IO).
+
+:- pred print_bool_name(rpt_graph::in, rptg_node::in, io::di, io::uo)
is det.
+
+print_bool_name(Graph, Node, !IO) :-
+	rptg_lookup_region_name(Graph, Node) = RegionName,
+	( if
+		string.replace(RegionName, "R", "B", BoolName)
+	  then
+	  	write_string(BoolName, !IO)
+	  else
+	  	unexpected(thisfile, "impossible")	
+	).	
+		
+	% print out actual region parameters.
+	%
+:- pred pp_actual_region_parameters(pred_proc_id::in, program_point::in, 
+    pred_proc_id::in, rpta_info_table::in, proc_region_set_table::in, 
+    proc_region_set_table::in, proc_region_set_table::in, bool::in, 
+    io::di, io::uo) is det.
+
+pp_actual_region_parameters(ProcId_p, PP, ProcId_q, RptaInfoTable, 
+    ConstantRTable, DeadRTable, BornRTable, HasParam, !IO) :-
+	map.lookup(RptaInfoTable, ProcId_p, rpta_info(Graph_p, Alpha_p)),
+	map.lookup(Alpha_p, PP, Alpha_PP),
+	
+	map.lookup(ConstantRTable, ProcId_q, ConstantR_q),
+	map.lookup(DeadRTable, ProcId_q, DeadR_q),
+	map.lookup(BornRTable, ProcId_q, BornR_q),
+	
+	set.to_sorted_list(ConstantR_q, LConstantR_q),
+	list.foldl(find_actual_param(Alpha_PP), LConstantR_q, [],
+			LActualConstantR0),
+	list.reverse(LActualConstantR0, LActualConstantR),
+	set.to_sorted_list(DeadR_q, LDeadR_q),
+	list.foldl(find_actual_param(Alpha_PP), LDeadR_q, [], LActualDeadR0),
+	list.reverse(LActualDeadR0, LActualDeadR),
+	set.to_sorted_list(BornR_q, LBornR_q),
+	list.foldl(find_actual_param(Alpha_PP), LBornR_q, [], LActualBornR0),
+	list.reverse(LActualBornR0, LActualBornR),
+	list.append(LActualConstantR, LActualDeadR, L0),
+	list.append(L0, LActualBornR, L),
+
+	( if 
+		HasParam = yes
+	  then
+		( if
+			L \= []
+		  then
+			io.write_string(", ", !IO),
+			io.write_list(L, ",", print_region_name(Graph_p), !IO)
+		  else
+			true
+		)
+	  else
+		( if
+			L \= []
+		  then
+			io.write_string("(", !IO),
+			io.write_list(L, ",", print_region_name(Graph_p), !IO),
+			io.write_string(")", !IO)
+		  else
+			true
+		)
+		
+	).	
+
+	
+:- pred find_actual_param(map(rptg_node, rptg_node)::in, rptg_node::in, 
+    list(rptg_node)::in, list(rptg_node)::out) is det.
+find_actual_param(Alpha_PP, Formal, Actuals0, Actuals) :-
+	map.lookup(Alpha_PP, Formal, Actual),
+	Actuals = [Actual|Actuals0].
+	
+
+:- pred print_make_ite_frame(string::in, io::di, io::uo) is det.
+print_make_ite_frame(Tabs, !IO) :-	
+	io.write_string(Tabs, !IO),
+	io.write_string("make_ite_frame", !IO).
+	
+:- pred print_commit_removal(string::in, io::di, io::uo) is det.
+print_commit_removal(Tabs, !IO) :-
+	io.write_string(Tabs, !IO),
+	io.write_string("commit_removal", !IO).
+
+:- pred print_undo_ite_changes(string::in, io::di, io::uo) is det.
+print_undo_ite_changes(Tabs, !IO) :-
+	io.write_string(Tabs, !IO),
+	io.write_string("undo_ite_changes", !IO).
+
+:- pred print_reg_alloc(string::in, string::in, int::in, io::di, io::uo) 
+    is det.
+print_reg_alloc(Tabs, RegName, Size, !IO) :-
+	io.write_string(Tabs, !IO),
+	S = "ralloc(" ++ RegName ++ ", " ++ string.int_to_string(Size) ++ ")",
+	io.write_string(S, !IO).
+
+	% add ralloc after a construction of a non-primitive term
+	%	
+:- pred add_ralloc(hlds_goal_expr::in, string::in, rpta_info_table::in,
+    pred_proc_id::in, io::di, io::uo) is det.
+add_ralloc(Expr, Tabs, RptaInfoTable, ProcId, !IO) :-
+( if 
+	Expr = unify(_, _, _, Unification, _),
+	Unification = construct(LVar, _ConsId, RVars, _, _, _, _)
+  then
+	list.length(RVars, Size),
+	( if
+		Size > 0
+	  then
+		io.write_string(",\n", !IO),
+		map.lookup(RptaInfoTable, ProcId, RptaInfo),
+		RptaInfo = rpta_info(Graph, _), 
+		get_node_by_variable(Graph, LVar, AllocatedRegion),
+		rptg_lookup_region_name(Graph, AllocatedRegion) = AllocatedRegionName,
+		print_reg_alloc(Tabs, AllocatedRegionName, Size, !IO)
+	  else
+		true
+	)	
+	else	
+		true
+).
+
+:- func pred_prefix = string.
+pred_prefix = ":- pred ".
+
+:- func func_prefix = string.
+func_prefix = ":- func ".
+
+:- func mode_prefix = string.
+mode_prefix = ":- mode ".
+
+:- func thisfile = string.
+thisfile = "pp_region_annotated_code.m".
+
Index: rbmm_utils.m
===================================================================
RCS file: rbmm_utils.m
diff -N rbmm_utils.m
--- /dev/null	1 Jan 1970 00:00:00 -0000
+++ rbmm_utils.m	8 May 2007 00:35:58 -0000
@@ -0,0 +1,47 @@
+%------------------------------------------------------------------------------%
+% vim: ft=mercury ts=4 sw=4 et
+%------------------------------------------------------------------------------%
+% Copyright (C) 2000-2002,2006 The University of Melbourne.
+% This file may only be copied under the terms of the GNU General
+% Public License - see the file COPYING in the Mercury distribution.
+%-----------------------------------------------------------------------------%
+
+% file rbmm_utils.m.
+% main author: Quan Phan.
+%
+% This module implements some utilities that are used in several phases of 
+% rbmm analysis.
+
+:- module rbmm_utils.
+
+:- interface.
+
+:- import_module hlds.
+:- import_module hlds.hlds_module.
+:- import_module hlds.hlds_pred.
+:- import_module parse_tree.
+:- import_module parse_tree.prog_data.
+
+:- import_module list.
+
+:- pred find_input_output_params(module_info::in, proc_info::in,
+	list(prog_var)::out, list(prog_var)::out) is det.
+
+
+:- implementation.
+
+:- import_module hlds.arg_info.
+:- import_module map.
+
+	% Currently, this predicate is somewhat hard-coded.
+	% The idea is that I only care about input and output parameters,
+	% i.e. those of modes "in" and "out"
+	% therefore other parameters are simply ignored.
+    %
+find_input_output_params(HLDS, CalleeProcInfo, Inputs, Outputs) :-
+	proc_info_get_headvars(CalleeProcInfo, ArgVars),
+	proc_info_get_vartypes(CalleeProcInfo, VarTypes),
+	list.map(map.lookup(VarTypes), ArgVars, ArgTypes),
+	proc_info_get_argmodes(CalleeProcInfo, ArgModes),
+	arg_info.compute_in_and_out_vars(HLDS, ArgVars, ArgModes, ArgTypes,
+		Inputs, Outputs).
Index: region_transform.m
===================================================================
RCS file: region_transform.m
diff -N region_transform.m
--- /dev/null	1 Jan 1970 00:00:00 -0000
+++ region_transform.m	14 May 2007 07:18:33 -0000
@@ -0,0 +1,639 @@
+%-----------------------------------------------------------------------------%
vim: ft=mercury ts=4 sw=4 et
+%-----------------------------------------------------------------------------%
+% Copyright (C) 2006-2002,2004 The University of Melbourne.
+% This file may only be copied under the terms of the GNU General
+% Public License - see the file COPYING in the Mercury distribution.
+%-----------------------------------------------------------------------------%
+
+% file region_transform.m.
+% main author: Quan Phan
+%
+% This module implements the process of introducing region instructions to 
+% each program point in a procedure based on its region points-to graph
and 
+% live region information.
+%
+
+:- module region_transform.
+
+:- interface.
+
+:- import_module hlds.
+:- import_module hlds.hlds_module.
+:- import_module hlds.hlds_pred. 
+:- import_module lra_data.
+:- import_module rpta_info.
+:- import_module smm_data.
+
+:- import_module list.
+:- import_module map.
+:- import_module io.
+
+:- type annotation_table == map(pred_proc_id, annotation_proc).
+
+:- type annotation_proc == map(program_point, before_after).
+:- type before_after ---> before_after(list(string), list(string)).
+
+    % Currently the region instructions are strings, which are attached to
+    % either before or after a program point.
+    % 
+:- pred transform(module_info::in, rpta_info_table::in, ep_table::in, 
+	proc_pp_region_set_table::in, proc_pp_region_set_table::in, 
+    proc_pp_region_set_table::in, proc_region_set_table::in, 
+    proc_region_set_table::in, proc_region_set_table::in, 
+    annotation_table::out, io::di, io::uo) is det.
+
+:- implementation. 
+
+:- import_module hlds.hlds_data.
+:- import_module hlds.hlds_goal.
+:- import_module hlds.hlds_llds.
+
+:- import_module check_hlds.
+:- import_module check_hlds.goal_path.
+
+:- import_module rpt_graph.
+:- import_module rbmm_utils.
+:- import_module smm_utils.
+
+:- import_module bool.
+:- import_module varset.
+:- import_module string.
+:- import_module term.
+:- import_module set.
+:- import_module std_util.
+:- import_module require.
+:- import_module svmap.
+:- import_module svset.
+:- import_module pair.
+:- import_module solutions.
+%-----------------------------------------------------------------------------%
+
+transform(HLDS, RptaInfoTable, EPTable, LRBeforeTable, LRAfterTable, 
+    VoidVarRegionTable, BornRTable, DeadRTable, LocalRTable,
AnnotationTable, 
+    !IO) :-
+
+    module_info_predids(PredIds, HLDS, _),
+	map.init(AnnotationTable0),
+	list.foldl2(transform_pred(HLDS, RptaInfoTable, EPTable, LRBeforeTable, 
+        LRAfterTable, VoidVarRegionTable, BornRTable, DeadRTable, 
+        LocalRTable), PredIds, AnnotationTable0, AnnotationTable, !IO).
+
+:- pred transform_pred(module_info::in, rpta_info_table::in, ep_table::in, 
+    proc_pp_region_set_table::in, proc_pp_region_set_table::in, 
+    proc_pp_region_set_table::in, proc_region_set_table::in, 
+    proc_region_set_table::in, proc_region_set_table::in, pred_id::in, 
+    annotation_table::in, annotation_table::out, io::di, io::uo) is det.
+transform_pred(HLDS, RptaInfoTable, EPTable, LRBeforeTable, LRAfterTable, 
+    VoidVarRegionTable, BornRTable, DeadRTable, LocalRTable, PredId, 
+    !AnnotationTable, !IO) :-
+	module_info_pred_info(HLDS, PredId, PredInfo),
+	pred_info_non_imported_procids(PredInfo) = ProcIds,
+	list.foldl2(transform_proc(HLDS, PredId, RptaInfoTable, EPTable, 
+        LRBeforeTable, LRAfterTable, VoidVarRegionTable, BornRTable, 
+        DeadRTable, LocalRTable), ProcIds, !AnnotationTable, !IO ).
+
+:- pred transform_proc(module_info::in, pred_id::in, rpta_info_table::in, 
+    ep_table::in, proc_pp_region_set_table::in,
proc_pp_region_set_table::in, 
+    proc_pp_region_set_table::in, proc_region_set_table::in, 
+    proc_region_set_table::in, proc_region_set_table::in, proc_id::in, 
+    annotation_table::in, annotation_table::out, io::di, io::uo) is det.
+transform_proc(HLDS, PredId, RptaInfoTable, EPTable, LRBeforeTable, 
+    LRAfterTable, VoidVarRegionTable, BornRTable, DeadRTable, LocalRTable, 
+    ProcId, !AnnotationTable, !IO) :-
+	PredProcId = proc(PredId, ProcId),
+	( if
+		smm_utils.some_are_special_preds([PredProcId], HLDS)
+	then
+		true
+	else
+		module_info_pred_proc_info(HLDS, PredProcId, _PredInfo, ProcInfo),
+	
+            % We need to follow each execution path and build up the 
+            % before and after lists for program points in the path,
+            % ignore "out".
+            %
+	
+		map.lookup(RptaInfoTable, PredProcId, RptaInfo),
+		map.lookup(BornRTable, PredProcId, BornR),
+		map.lookup(DeadRTable, PredProcId, DeadR),
+		map.lookup(LocalRTable, PredProcId, LocalR),
+            % to have the goal at each pp
+            %
+		map.lookup(EPTable, PredProcId, EPs), 
+
+		map.lookup(LRBeforeTable, PredProcId, ProcLRBefore),
+		map.lookup(LRAfterTable, PredProcId, ProcLRAfter),
+		map.lookup(VoidVarRegionTable, PredProcId, ProcVoidVarRegion),
+		
+		transform_eps(EPs, RptaInfo, BornR, DeadR, LocalR, RptaInfoTable, 
+			ProcLRBefore, ProcLRAfter, ProcVoidVarRegion, BornRTable, 
+            DeadRTable, HLDS, ProcInfo, map.init, AnnotationProc, !IO),
+		svmap.set(PredProcId, AnnotationProc, !AnnotationTable)
+	).
+
+	% follow each execution path of a procedure and introduce 
+    % region-annotations
+    %
+:- pred transform_eps(list(execution_path)::in, rpta_info::in, 
+    set(rptg_node)::in, set(rptg_node)::in, set(rptg_node)::in, 
+    rpta_info_table::in, pp_region_set_table::in, pp_region_set_table::in, 
+    pp_region_set_table::in, proc_region_set_table::in, 
+    proc_region_set_table::in, module_info::in, proc_info::in, 
+    annotation_proc::in, annotation_proc::out, io::di, io::uo) is det.
+transform_eps([], _, _, _, _, _, _, _, _, _, _, _HLDS, _ProcInfo, 
+    !AnnotationProc, !IO).
+transform_eps([EP|OtherEPs], RptaInfo, BornR, DeadR, LocalR,
RptaInfoTable, 
+    ProcLRBefore, ProcLRAfter, ProcVoidVarRegion, BornRTable, DeadRTable,
+    HLDS, ProcInfo, AnnotationProc0, AnnotationProc, !IO) :-
+	transform_ep(EP, RptaInfo, BornR, DeadR, LocalR, RptaInfoTable, 
+        ProcLRBefore, ProcLRAfter, ProcVoidVarRegion, BornRTable,
DeadRTable,
+        HLDS, ProcInfo, AnnotationProc0, AnnotationProc1, !IO),
+	transform_eps(OtherEPs, RptaInfo, BornR, DeadR, LocalR, RptaInfoTable, 
+        ProcLRBefore, ProcLRAfter, ProcVoidVarRegion, BornRTable,
DeadRTable,
+        HLDS, ProcInfo, AnnotationProc1, AnnotationProc, !IO).
+
+	% build before and after lists for program points in this execution path
+	% we need some tables because we will need to get deadR and bornR region 
+    % sets of callees
+    %
+:- pred transform_ep(execution_path::in, rpta_info::in,
set(rptg_node)::in, 
+    set(rptg_node)::in, set(rptg_node)::in, rpta_info_table::in, 
+	pp_region_set_table::in, pp_region_set_table::in,
pp_region_set_table::in, 
+    proc_region_set_table::in, proc_region_set_table::in, module_info::in, 
+    proc_info::in, annotation_proc::in, annotation_proc::out, io::di, 
+    io::uo) is det.
+transform_ep([], _, _, _, _, _, _, _, _, _, _, _HLDS, _ProcInfo, 
+    !AnnotationProc, !IO).
+transform_ep([PP - Goal, NextPPGoal|OtherPP_Goals], RptaInfo, BornR,
DeadR, 
+    LocalR, RptaInfoTable, ProcLRBefore, ProcLRAfter, ProcVoidVarRegion, 
+    BornRTable, DeadRTable, HLDS, ProcInfo, !AnnotationProc, !IO) :-
+		% this is the first program point or a middle one
+        %
+	map.lookup(ProcLRBefore, PP, LRBefore),
+	map.lookup(ProcLRAfter, PP, LRAfter),
+	map.lookup(ProcVoidVarRegion, PP, VoidVarRegions),
+
+        % XXX new: those regions need to be deleted after this pp, but
only if 
+        % they are in DeadR or LocalR or BornR
+	set.difference(VoidVarRegions, LRAfter, DeadRegions0),
+	set.union(DeadR, LocalR, Dead_Local),
+	set.union(Dead_Local, BornR, Dead_Local_Born),
+	
+	set.intersect(Dead_Local_Born, DeadRegions0, DeadRegions),
+	RptaInfo = rpta_info(Graph_p, _AlphaMapping),
+	set.fold(annotate_after_pp(PP, Graph_p, "Tn", "remove"), DeadRegions, 
+        !AnnotationProc),
+
+        % XXX: to be removed before the next pp ... 
+        %
+	map.lookup(ProcLRBefore, NextPP, LRBeforeNext),
+	set.difference(LRAfter, LRBeforeNext, NolongerNeededRegions0),
+
+		% BUT only if they are in DeadR or LocalR or BornR 
+		% 
+	set.intersect(Dead_Local_Born, NolongerNeededRegions0, 
+        NolongerNeededRegions),
+	
+	NextPPGoal = NextPP - _NextEPGoal,
+	set.fold(annotate_before_pp(NextPP, Graph_p, "T5", "remove"), 
+			NolongerNeededRegions, !AnnotationProc),
+			
+        % XXX: hack???
+        %
+	set.difference(NolongerNeededRegions0, Dead_Local, LiveBeforeNext1),
+	set.difference(LiveBeforeNext1, BornR, LiveBeforeNext),
+	set.union(LiveBeforeNext, LRBeforeNext, LRBeforeNext1),
+	map.set(ProcLRBefore, NextPP, LRBeforeNext1, ProcLRBefore1),
+
+	( if 
+		Goal = hlds_goal(switch(_A, _B, _Cases), _Info)
+	  then
+            % XXX: This is a switch, i.e. unification, only rule 4 applied.
+            %
+		transformation_rule_4_2(LRBefore, LRAfter, PP, 
+			RptaInfo, LocalR, DeadR, BornR, !AnnotationProc)
+	  else 
+		Goal = hlds_goal(Expr, _Info),
+		( if
+			transformation_rule_1(Expr, LRBefore, LRAfter, BornR, LocalR, PP, 
+                RptaInfo, RptaInfoTable, BornRTable, !AnnotationProc),
+			transformation_rule_2(Expr, LRBefore, LRAfter, PP, RptaInfo, 
+                LocalR, BornR, HLDS, ProcInfo, !AnnotationProc),
+			transformation_rule_3(Expr, LRBefore, LRAfter, PP, RptaInfo, 
+                LocalR, DeadR, RptaInfoTable, DeadRTable, !AnnotationProc),
+			transformation_rule_4(Expr, LRBefore, LRAfter, PP, RptaInfo, 
+                LocalR, DeadR, BornR, !AnnotationProc)
+	  	  then
+			true
+	  	  else
+			error("transform_ep -> transformation rule: impossible.")
+		)
+	),
+	transform_ep([NextPPGoal|OtherPP_Goals], RptaInfo, BornR, DeadR, LocalR, 
+        RptaInfoTable, ProcLRBefore1, ProcLRAfter, ProcVoidVarRegion, 
+        BornRTable, DeadRTable, HLDS, ProcInfo, !AnnotationProc, !IO).
+	
+	% this is for the last program point
+transform_ep([PP - Goal], RptaInfo, BornR, DeadR, LocalR, RptaInfoTable, 
+		ProcLRBefore, ProcLRAfter, ProcVoidVarRegion, BornRTable, DeadRTable, 
+		HLDS, ProcInfo, !AnnotationProc, !IO) :-
+		% this is the last pp
+        %
+	map.lookup(ProcLRBefore, PP, LRBefore),
+	map.lookup(ProcLRAfter, PP, LRAfter),
+	map.lookup(ProcVoidVarRegion, PP, VoidVarRegions),
+
+	% XXX new: those regions need to be deleted after this pp
+	% i.e., in DeadR, LocalR, or BornR
+	set.difference(VoidVarRegions, LRAfter, DeadRegions0),
+	set.union(DeadR, LocalR, Dead_Local),
+	set.union(Dead_Local, BornR, Dead_Local_Born),
+	set.intersect(Dead_Local_Born, DeadRegions0, DeadRegions),
+	
+	RptaInfo = rpta_info(Graph_p, _AlphaMapping),
+	set.fold(annotate_after_pp(PP, Graph_p, "Tn", "remove"), DeadRegions, 
+        !AnnotationProc),
+
+	( if 
+		Goal = hlds_goal(switch(_A, _B, _Cases), _Info)
+	  then
+            % XXX: This is a switch, i.e. unification, only rule 4 applied.
+            %
+		transformation_rule_4_2(LRBefore, LRAfter, PP, RptaInfo, LocalR, 
+            DeadR, BornR, !AnnotationProc)
+	  else 
+		Goal = hlds_goal(Expr, _Info),
+		( if
+			transformation_rule_1(Expr, LRBefore, LRAfter, BornR, LocalR, PP, 
+                RptaInfo, RptaInfoTable, BornRTable, !AnnotationProc),
+			transformation_rule_2(Expr, LRBefore, LRAfter, PP, RptaInfo, 
+                LocalR, BornR, HLDS, ProcInfo, !AnnotationProc),
+			transformation_rule_3(Expr, LRBefore, LRAfter, PP, RptaInfo, 
+                LocalR, DeadR, RptaInfoTable, DeadRTable, !AnnotationProc),
+			transformation_rule_4(Expr, LRBefore, LRAfter, PP, RptaInfo, 
+                LocalR, DeadR, BornR, !AnnotationProc)
+	  	  then
+			true
+	  	  else
+			error("transform_ep -> transformation rule: impossible.")
+		)
+	).
+
+	% Rule: if an output region of q is not live, not created by q, 
+	% and p is allowed to create it, it is created before the call to q.
+	%
+	% Algorithm: 
+	% There are two cases: either q is defined in this module or q is
imported.
+	% The former is straightforward
+	% The later: for now, with the assumption that all source code is in only 
+    % one module, imported preds will only be ones from Mercury's
library. We 
+    % do not intent to deal with the library's code now therefore we
have to 
+    % assume here that the caller will always *CREATE* the OUTPUT
REGION for 
+    % those procedures.
+	%
+:- pred transformation_rule_1(hlds_goal_expr::in, 
+	region_set::in, region_set::in, region_set::in, region_set::in, 
+    program_point::in, rpta_info::in, rpta_info_table::in, 
+    proc_region_set_table::in, annotation_proc::in, annotation_proc::out)
+    is semidet.
+transformation_rule_1(plain_call(PredId_q, ProcId_q, _, _, _, _),
LRBefore, 
+    LRAfter, BornR, LocalR, PP, RptaInfo, RptaInfoTable, 
+    BornRTable, !AnnotationProc) :-
+	PredProcId_q = proc(PredId_q, ProcId_q),
+	( if
+            % check if q is defined in this module, if it is not, just
ignore
+            % currently we do not collect BornR for non-defined-in-module 
+            % procedure, so if we cannot find one here then q is an
imported.
+            %
+		map.search(BornRTable, PredProcId_q, _)
+	then
+            % q is defined in this module
+		
+		RptaInfo = rpta_info(Graph, AlphaMapping),
+        map.lookup(AlphaMapping, PP, AlphaAtPP),
+	
+		map.lookup(BornRTable, PredProcId_q, BornR_q),
+		map.lookup(RptaInfoTable, PredProcId_q, RptaInfo_q),
+		RptaInfo_q = rpta_info(Graph_q, _AlphaMapping_q),
+	
+		map.foldl(process_mapping_rule_1(PP, LRBefore, LRAfter, BornR, LocalR, 
+            BornR_q, Graph_q, Graph), AlphaAtPP, !AnnotationProc)
+	else
+            % q is from an imported module. So there is no renaming, no
alpha 
+            % mapping at this pp (actually, there is, but we don't
care), no 
+            % rpt graph of q also. 
+            % BornR of q is empty, so just create whatever becoming live.
+            % provided that they are in BornR or LocalR
+            %
+		set.union(BornR, LocalR, Born_Local),
+		set.difference(LRAfter, LRBefore, ToBeCreated0),
+		set.intersect(ToBeCreated0, Born_Local, ToBeCreated),
+		RptaInfo = rpta_info(Graph, _AlphaMapping),
+		set.fold(process_mapping_rule_1_imported_callee(PP, Graph), 
+            ToBeCreated, !AnnotationProc)
+	).
+
+transformation_rule_1(conj(_, []), _LRBefore, _LRAfter, _BornR,
_LocalR, _PP, 
+    _RptaInfo, _RptaInfoTable, _BornRTable, !AnnotationProc).
+transformation_rule_1(disj([]), _LRBefore, _LRAfter, _BornR, _LocalR, _PP, 
+    _RptaInfo, _RptaInfoTable, _BornRTable, !AnnotationProc).
+transformation_rule_1(unify(_, _, _, _, _), _LRBefore, _LRAfter, _BornR, 
+    _LocalR, _PP, _RptaInfo, _RptaInfoTable, _BornRTable, !AnnotationProc).
+
+:- pred process_mapping_rule_1(program_point::in, 
+	region_set::in, region_set::in, region_set::in, region_set::in, 
+	region_set::in, rpt_graph::in, rpt_graph::in, rptg_node::in,
rptg_node::in,
+	annotation_proc::in, annotation_proc::out) is det.
+process_mapping_rule_1(PP, RenamedLR, LRAfter, BornR, LocalR, BornR_q, 
+    Graph_q, Graph_p, SourceRegion, TargetRegion, !AnnotationProc) :-
+	get_node_by_node(Graph_q, SourceRegion, RPrime),
+	get_node_by_node(Graph_p, TargetRegion, R),
+	set.union(BornR, LocalR, Born_Local),
+	( if
+		(
+			not set.contains(RenamedLR, R),
+			set.contains(LRAfter, R),
+			not set.contains(BornR_q, RPrime),
+			set.contains(Born_Local, R)
+		)
+	  then
+		annotate_before_pp(PP, Graph_p, "T1", "create", R, !AnnotationProc)
+	  else
+		true
+	).
+
+:- pred process_mapping_rule_1_imported_callee(program_point::in, 
+    rpt_graph::in, rptg_node::in, annotation_proc::in,
annotation_proc::out)
+    is det.
+process_mapping_rule_1_imported_callee(PP, Graph_p, Region,
!AnnotationProc) :-
+		% Region here is a real one, just go ahead.
+		% get Before of this pp and add "create r" to it
+        %
+	annotate_before_pp(PP, Graph_p, "T1", "create", Region, !AnnotationProc).
+	
+	% transformation rule 2: if a region is not live before a construction
but 
+    % it is live after and is a local region or in a BornR set, then the 
+    % region is created before the unification.
+	% 
+:- pred transformation_rule_2(hlds_goal_expr::in, 
+	region_set::in, region_set::in, program_point::in, rpta_info::in, 
+    region_set::in, region_set::in, module_info::in, proc_info::in, 
+    annotation_proc::in, annotation_proc::out) is semidet.
+transformation_rule_2(unify(X, _, _, construct(_, _, _, _, _, _, _), _), 
+    LRBefore, LRAfter,  PP, RptaInfo, LocalR, BornR, HLDS, ProcInfo, 
+    !AnnotationProc) :-
+	RptaInfo = rpta_info(Graph, _AlphaMapping),
+		% region of X
+		% need to be regions reachable from X
+        %
+	reach_from_a_variable(Graph, HLDS, ProcInfo, X, set__init, Reach_X),
+
+		% not in LRBefore
+		%
+	set.difference(Reach_X, LRBefore, ToBeCreated0),
+
+		% in LRAfter
+		%
+	set.intersect(ToBeCreated0, LRAfter, ToBeCreated1),
+
+		% in (BornR or LocalR)
+		%
+	set.union(LocalR, BornR, Local_Born),
+	set.intersect(ToBeCreated1, Local_Born, ToBeCreated),
+	
+	set.fold(annotate_before_pp(PP, Graph, "T2", "create"), ToBeCreated, 
+        !AnnotationProc).
+
+transformation_rule_2(unify(_, _, _, deconstruct(_, _, _, _, _, _), _), 
+    _LRBefore, _LRAfter, _PP, _RptaInfo, _LocalR, _BornR, _HLDS,
_ProcInfo, 
+    !AnnotationProc).
+transformation_rule_2(unify(_, _, _, assign(_, _), _), _LRBefore,
_LRAfter, 
+    _PP, _RptaInfo, _LocalR, _BornR,  _HLDS, _ProcInfo, !AnnotationProc).
+transformation_rule_2(unify(_, _, _, simple_test(_, _), _), _LRBefore, 
+    _LRAfter, _PP, _RptaInfo, _LocalR, _BornR,  _HLDS, _ProcInfo, 
+    !AnnotationProc).
+transformation_rule_2(unify(_, _, _, complicated_unify(_, _, _), _), 
+    _LRBefore, _LRAfter, _PP, _RptaInfo, _LocalR, _BornR,  _HLDS,
_ProcInfo, 
+    !AnnotationProc).
+transformation_rule_2(plain_call(_, _, _, _, _, _), _LRBefore, _LRAfter, 
+    _PP, _RptaInfo, _LocalR, _BornR, _HLDS, _ProcInfo, !AnnotationProc).
+transformation_rule_2(conj(_, []), _LRBefore, _LRAfter, _PP, 
+    _RptaInfo, _LocalR, _BornR, _HLDS, _ProcInfo, !AnnotationProc).
+transformation_rule_2(disj([]), _LRBefore, _LRAfter, _PP, 
+    _RptaInfo, _LocalR, _BornR, _HLDS, _ProcInfo, !AnnotationProc).
+
+	% Rule: if a region is live before q but it is not live after q and it is 
+    % not removed by q, the caller will remove it after calling q
+	% Algorithm: 
+	% There are two cases: either q is defined in this module or q is
imported.
+	% The former: is straightforward.
+	% The later: for now, with the assumption that all source code is in only 
+    % one module, imported preds will only be ones from Mercury's
library. We 
+    % do not intent to deal with the library's code now therefore we
have to 
+    % assume here that the caller will always REMOVE the REGION for those 
+    % procedures.
+	% 
+:- pred transformation_rule_3(hlds_goal_expr::in, region_set::in, 
+    region_set::in, program_point::in, rpta_info::in, region_set::in, 
+    region_set::in, rpta_info_table::in, proc_region_set_table::in, 
+    annotation_proc::in, annotation_proc::out) is semidet.
+transformation_rule_3(plain_call(PredId_q, ProcId_q, _, _, _, _),
LRBefore, 
+    LRAfter, PP, RptaInfo, LocalR, DeadR, _RptaInfoTable, DeadRTable, 
+    !AnnotationProc) :-
+	PredProcId_q = proc(PredId_q, ProcId_q),
+	( if
+		map.search(DeadRTable, PredProcId_q, _)
+	  then
+            % alpha mapping at this PP
+            %
+		RptaInfo = rpta_info(Graph, AlphaMapping),
+		map.lookup(AlphaMapping, PP, AlphaAtPP),
+
+		map.lookup(DeadRTable, PredProcId_q, DeadR_q),
+		set.union(LocalR, DeadR, Local_Dead),
+		set.fold(process_mapping_rule_3(PP, LRAfter, Local_Dead, 
+            DeadR_q, AlphaAtPP, Graph), LRBefore, !AnnotationProc)
+	  else
+            % q is from an imported module. So there is no renaming, no
alpha 
+            % mapping at this pp (actually,
+            % there is, but we don't care), no rpt graph of q also. 
+            % deadR of q is empty, so just remove whatever becoming dead.
+            % provided that they are in LocalR or DeadR
+            %
+		set.union(LocalR, DeadR, Local_Dead),
+		set.difference(LRBefore, LRAfter, ToBeRemoved0),
+		set.intersect(ToBeRemoved0, Local_Dead, ToBeRemoved),
+		RptaInfo = rpta_info(Graph, _AlphaMapping),
+		set.fold(process_mapping_rule_3_imported_callee(PP, Graph), 
+            ToBeRemoved, !AnnotationProc)
+	).
+
+transformation_rule_3(unify(_, _, _, _, _), _LRBefore, _LRAfter, _PP, 
+    _RptaInfo, _LocalR, _DeadR, _RptaInfoTable, _DeadRTable,
!AnnotationProc).
+transformation_rule_3(conj(_, []), _LRBefore, _LRAfter, _PP, _RptaInfo, 
+    _LocalR, _DeadR, _RptaInfoTable, _DeadRTable, !AnnotationProc).
+transformation_rule_3(disj([]), _LRBefore, _LRAfter,  _PP, _RptaInfo,
+    _LocalR, _DeadR, _RptaInfoTable, _DeadRTable, !AnnotationProc).
+
+:- pred process_mapping_rule_3(program_point::in, region_set::in, 
+    region_set::in, region_set::in, map(rptg_node, rptg_node)::in, 
+    rpt_graph::in, rptg_node::in, annotation_proc::in,
annotation_proc::out) 
+    is det.
+process_mapping_rule_3(PP, LRAfter, Local_Dead, DeadR_q, AlphaAtPP,
Graph_p, 
+    R, !AnnotationProc) :-
+		% alpha(RPrime) = R, 
+	solutions(map.inverse_search(AlphaAtPP, R), SourceList), 
+	set.list_to_set(SourceList, RPrimes),
+	
+        % RPrime is removed by the callee?
+        %
+	set.intersect(RPrimes, DeadR_q, Removed_By_q),
+ 	( if
+		(
+            not set.contains(LRAfter, R),
+            set.contains(Local_Dead, R),
+            set.empty(Removed_By_q)
+        )
+	  then
+	  	annotate_after_pp(PP, Graph_p, "T3", "remove", R, !AnnotationProc)
+	  else
+		true
+	).
+
+:- pred process_mapping_rule_3_imported_callee(program_point::in, 
+    rpt_graph::in, rptg_node::in, annotation_proc::in,
annotation_proc::out) 
+    is det.
+process_mapping_rule_3_imported_callee(PP, Graph_p, Region,
!AnnotationProc) :-
+		% Region here is a real one, just go ahead
+		% get After of the pp and add "remove r" to it
+        %
+	annotate_after_pp(PP, Graph_p, "T3", "remove", Region, !AnnotationProc).
+
+	% transformation rule 4: if a region is live before a unification but 
+	% it is not live after and it is in the dead region set or is a local
+	% region, then it is removed after the unification.
+    %
+:- pred transformation_rule_4(hlds_goal_expr::in, region_set::in, 
+    region_set::in, program_point::in, rpta_info::in, 
+	region_set::in, region_set::in, region_set::in, annotation_proc::in, 
+    annotation_proc::out) is semidet.
+transformation_rule_4(unify(_, _, _, _, _), LRBefore, LRAfter, PP, 
+    RptaInfo, LocalR, DeadR, BornR, !AnnotationProc) :-
+	set.union(LocalR, DeadR, Local_Dead),
+	set.union(Local_Dead, BornR, Local_Dead_Born),
+	RptaInfo = rpta_info(Graph, _AlphaMapping),
+		% process each live region 
+        %
+	set.fold(process_live_region_rule_4(PP, LRAfter, Local_Dead_Born, 
+        Graph), LRBefore, !AnnotationProc).
+
+transformation_rule_4(plain_call(_, _, _, _, _, _), _LRBefore,
_LRAfter, _PP, 
+    _RptaInfo, _LocalR, _DeadR, _BornR, !AnnotationProc).
+transformation_rule_4(conj(_, []), _LRBefore, _LRAfter, _PP, _RptaInfo, 
+    _LocalR, _DeadR, _BornR, !AnnotationProc).
+transformation_rule_4(disj([]), _LRBefore, _LRAfter, _PP, _RptaInfo,
_LocalR, 
+    _DeadR, _BornR, !AnnotationProc).
+	
+	% Because the goal at this pp is a real unification, not one removed by 
+    % MMC the pp for it exists, and we can *ALWAYS* record the
annotations to 
+    % this pp's After.
+    %
+:- pred process_live_region_rule_4(program_point::in, region_set::in,
+	region_set::in, rpt_graph::in, rptg_node::in, annotation_proc::in, 
+    annotation_proc::out) is det.
+process_live_region_rule_4(PP, LRAfter, Local_Dead_Born, Graph, Region, 
+    !AnnotationProc) :-
+	( if
+		(
+            not set.contains(LRAfter, Region),
+            set.contains(Local_Dead_Born, Region)
+        )
+	  then
+	  	annotate_after_pp(PP, Graph, "T4", "remove", Region, !AnnotationProc)
+	  else
+		true
+	).
+
+	% this is for the case rule 4 applied for a unification in a switch, 
+	% the unification has been removed by MMC.
+	% We will record the annotations in PP's After.
+    %
+:- pred transformation_rule_4_2(region_set::in, region_set::in, 
+	program_point::in, rpta_info::in, region_set::in, region_set::in, 
+    region_set::in, annotation_proc::in, annotation_proc::out) is det.
+transformation_rule_4_2(LRBefore, LRAfter, PP, RptaInfo, LocalR, DeadR,
BornR, 
+    !AnnotationProc) :-
+	set.union(LocalR, DeadR, Local_Dead),
+	set.union(Local_Dead, BornR, Local_Dead_Born),
+	RptaInfo = rpta_info(Graph, _AlphaMapping),
+		% process each live region 
+	set.fold(process_live_region_rule_4(PP, LRAfter, Local_Dead_Born, 
+        Graph), LRBefore, !AnnotationProc).
+
+:- pred process_live_region_rule_4_2(program_point::in, region_set::in,
+	region_set::in, rpt_graph::in, rptg_node::in, annotation_proc::in, 
+    annotation_proc::out) is det.
+process_live_region_rule_4_2(PP, LRAfter, Local_Dead_Born, Graph, Region, 
+    !AnnotationProc) :-
+	( if
+		(
+            not set.contains(LRAfter, Region),
+            set.contains(Local_Dead_Born, Region)
+        )
+	  then
+		annotate_after_pp(PP, Graph, "T4", "remove", Region, !AnnotationProc)
+	  else
+		true
+	).
+
+:- pred annotate_after_pp(program_point::in, rpt_graph::in, string::in, 
+    string::in, rptg_node::in, annotation_proc::in, annotation_proc::out) 
+    is det.
+annotate_after_pp(PP, Graph_p, Rule, Action, Region, !AnnotationProc) :-
+        % Region here is a real one, just go ahead
+        %
+	rptg_node_contents(Graph_p, Region, Content),
+	string.format("%s: %s %s", [s(Rule), s(Action), s(Content^reg_var_name)]) 
+        = RegionAnno,
+        % get After of the pp and add annotation to it
+        %
+	( if
+		map.search(!.AnnotationProc, PP, before_after(Before, After))
+	then
+		( if 
+			list.member(RegionAnno, After)
+		  then
+			true
+		  else
+			svmap.set(PP, before_after(Before, [RegionAnno|After]), 
+                !AnnotationProc)
+		)
+	else
+		svmap.set(PP, before_after([], [RegionAnno]), !AnnotationProc)
+	).
+
+:- pred annotate_before_pp(program_point::in, rpt_graph::in, string::in, 
+    string::in, rptg_node::in, annotation_proc::in, annotation_proc::out) 
+    is det.
+annotate_before_pp(PP, Graph_p, Rule, Action, Region, !AnnotationProc) :-
+        % Region here is a real one, just go ahead
+        %
+	rptg_node_contents(Graph_p, Region, Content),
+	string.format("%s: %s %s", [s(Rule), s(Action), s(Content^reg_var_name)]) 
+        = RegionAnno,
+        % get After of the pp and add annotation to it
+        %
+	( if
+		map.search(!.AnnotationProc, PP, before_after(Before, After))
+	then
+		( if 
+			list.member(RegionAnno, Before)
+		  then
+			true
+		  else
+			svmap.set(PP, before_after([RegionAnno|Before], After), 
+                !AnnotationProc)
+		)
+	else
+		svmap.set(PP, before_after([RegionAnno], []), !AnnotationProc)
+	).
Index: rpt_alpha_mapping.m
===================================================================
RCS file: rpt_alpha_mapping.m
diff -N rpt_alpha_mapping.m
--- /dev/null	1 Jan 1970 00:00:00 -0000
+++ rpt_alpha_mapping.m	10 May 2007 05:26:44 -0000
@@ -0,0 +1,62 @@
+%-----------------------------------------------------------------------------%
+% vim: ft=mercury ts=4 sw=4 et
+%-----------------------------------------------------------------------------%%
Copyright (C) 1994-2007 The University of Melbourne.
+% This file may only be copied under the terms of the GNU Library General
+% Public License - see the file COPYING.LIB in the Mercury distribution.
+%-----------------------------------------------------------------------------%
+%
+% File: rpt_alpha_mapping.m
+% main author: Quan Phan.
+% 
+% This module defines the alpha mapping type.
+
+:- module rpt_alpha_mapping.
+
+:- interface.
+
+:- import_module smm_data.
+:- import_module rpt_graph.
+
+:- import_module io.
+:- import_module map.
+
+:- type rpt_alpha_mapping ==
+    map(
+        program_point,
+        map(rptg_node, rptg_node)
+    ).
+
+:- pred rpt_alpha_mapping_equal(rpt_alpha_mapping::in,
rpt_alpha_mapping::in) 
+    is semidet.
+            
+:- pred print_rpt_alpha_mapping(rpt_alpha_mapping::in, io::di, io::uo)
is det.
+
+:- implementation.
+
+:- import_module list.
+:- import_module string.
+
+rpt_alpha_mapping_equal(AlphaMapping1, AlphaMapping2) :-
+    AlphaMapping1 = AlphaMapping2.
+    
+    % Print alpha mapping
+    %
+print_rpt_alpha_mapping(AlphaMapping, !IO) :-
+    io.write_string("% AlphaMapping.\n", !IO),
+    map.foldl(print_rpt_alpha_mapping_at_pp, AlphaMapping, !IO),
+    io.write_string("% End AlphaMapping.\n", !IO).
+
+:- pred print_rpt_alpha_mapping_at_pp(program_point::in, 
+    map(rptg_node, rptg_node)::in, io::di, io::uo) is det.
+    print_rpt_alpha_mapping_at_pp(PP, AlphaMap, !IO) :-
+    print_program_point(PP, !IO),
+    io.write_string(".\n", !IO),
+    map.foldl(print_rpt_alpha_mapping_at_pp_entry, AlphaMap, !IO).
+
+:- pred print_rpt_alpha_mapping_at_pp_entry(rptg_node::in, 
+    rptg_node::in, io::di, io::uo) is det.
+print_rpt_alpha_mapping_at_pp_entry(FromNode, ToNode, !IO) :-
+    FromNode = rptg_node(F),
+    ToNode = rptg_node(T),
+    io.format("\talpha(node(%d)) = node(%d).\n", [i(F), i(T)], !IO).
+
Index: rpt_graph.m
===================================================================
RCS file: rpt_graph.m
diff -N rpt_graph.m
--- /dev/null	1 Jan 1970 00:00:00 -0000
+++ rpt_graph.m	15 May 2007 05:29:40 -0000
@@ -0,0 +1,1003 @@
+%-----------------------------------------------------------------------------%
+% vim: ft=mercury ts=4 sw=4 et
+%-----------------------------------------------------------------------------%
+% Copyright (C) 2005-2007 The University of Melbourne.
+% This file may only be copied under the terms of the GNU Library General
+% Public License - see the file COPYING.LIB in the Mercury distribution.
+%-----------------------------------------------------------------------------%
+%
+% File: rpt_graph.m.
+% Main author: Quan Phan.
+%
+% This module defines the region points-to graph data structure and the 
+% operations on it. 
+
+:- module rpt_graph.  
+
+:- interface.
+
+:- import_module hlds.
+:- import_module hlds.hlds_module.
+:- import_module hlds.hlds_pred.
+:- import_module parse_tree.
+:- import_module parse_tree.prog_data.
+
+:- import_module list.
+:- import_module set.
+:- import_module map.
+:- import_module int.
+:- import_module string.
+
+    % A region points-to graph rpt_graph(Node, Arc) represents a directed 
+    % graph with information type Node associated with each node, and 
+    % information of % type Arc associated with each arc.
+    %
+:- type rpt_graph
+        --->    rpt_graph(
+                    rptg_node_supply, % the counter for node
+                    rptg_arc_supply,  % the counter for arc
+                    map(rptg_node, rptg_node_content),
+                    map(rptg_arc, rptg_arc_info), 
+                    map(rptg_node, map(rptg_arc, rptg_node))
+                ).
+
+:- type rptg_node_supply    == int.
+:- type rptg_arc_supply     == int.
+
+:- type rptg_node
+        --->    rptg_node(int).
+
+:- type rptg_node_content
+        --->    rptg_node_content(
+                    % set of procedure variables assigned to this node
+                    varset          :: set(prog_var),
+                    % the region variable that names this node
+                    reg_var_name    :: string,
+                    merged_from     :: set(rptg_node),
+                    node_type       :: mer_type
+                ).
+
+:- pred rptg_node_content_get_varset(rptg_node_content::in, 
+    set(prog_var)::out) is det. 
+:- pred rptg_node_content_get_region_name(rptg_node_content::in,
+    string::out) is det.
+:- pred rptg_node_content_get_merged_from(rptg_node_content::in, 
+    set(rptg_node)::out) is det.
+:- pred rptg_node_content_get_node_type(rptg_node_content::in, 
+    mer_type::out) is det.
+
+:- pred rptg_node_content_set_varset(rptg_node_content::in, 
+    set(prog_var)::in, rptg_node_content::out) is det.
+:- pred rptg_node_content_set_region_name(rptg_node_content::in, 
+    string::in, rptg_node_content::out) is det.
+:- pred rptg_node_content_set_merged_from(rptg_node_content::in, 
+    set(rptg_node)::in, rptg_node_content::out) is det.
+:- pred rptg_node_content_set_node_type(rptg_node_content::in, 
+    mer_type::in, rptg_node_content::out) is det.
+
+:- type rptg_arc
+        --->    rptg_arc(int).
+
+:- type rptg_arc_content 
+        --->    rptg_arc_content(
+                    label   :: selector % the label of an edge
+                ).
+
+:- pred rptg_arc_content_get_label(rptg_arc_content::in, selector::out)
is det.
+:- pred rptg_arc_content_set_label(rptg_arc_content::in, selector::in, 
+    rptg_arc_content::out) is det.
+
+:- type rptg_arc_info 
+        --->    rptg_arc_info(
+                    rptg_node,  % the from node
+                    rptg_node,  % the to node
+                    rptg_arc_content    % the label
+                ).
+
+    % rpt_graph_init(Graph) binds Graph to an empty graph
+    % containing no nodes and no arcs. (The graph contains
+    % a counter of the number of nodes allocated in it, so
+    % it is possible for a graph to contain no nodes or arcs
+    % and still fail to unify with the binding of Graph from
+    % rpt_graph_init.)
+    %
+:- pred rpt_graph_init(rpt_graph::out) is det.
+
+:- func rpt_graph_init = rpt_graph.
+
+    % rptg_set_node(NodeInfo, Node, OldGraph, NewGraph) takes
+    % OldGraph and NodeInfo which is the information to be stored
+    % in a new node, and returns a key "Node" which refers to that
+    % node, and the new graph NewGraph containing all of the nodes
+    % and arcs in OldGraph as well as the new node.
+    % It is possible to have two nodes in the graph with the
+    % same information stored in them.
+    %
+    % This operation is O(lgN) for a graph containing N nodes.
+    %
+:- pred rptg_set_node(rptg_node_content::in, rptg_node::out, 
+    rpt_graph::in, rpt_graph::out) is det. 
+
+    % rptg_node_contents(Graph, Node, NodeContent) takes Graph and
+    % Node and returns the information NodeContent stored in Node.
+    %
+    % This operation is O(lgN) for a graph containing N nodes.
+    %
+:- pred rptg_node_contents(rpt_graph::in,
+    rptg_node::in, rptg_node_content::out) is det.
+:- func rptg_node_contents(rpt_graph, rptg_node) = rptg_node_content.
+
+    % rptg_successors(Graph, Node, Nodes) takes a graph Graph and
+    % a node Node and returns the set of nodes Nodes that are reachable
+    % (directly - not transitively) from Node.
+    %
+    % This operation is O(NlgN) for a graph containing N nodes.
+    %
+:- pred rptg_successors(rpt_graph::in, rptg_node::in, set(rptg_node)::out) 
+    is det.
+:- func rptg_successors(rpt_graph, rptg_node) = set(rptg_node).
+
+    % rptg_get_nodes(Graph, Nodes) binds Nodes to the set of nodes in
Graph.
+    %
+:- pred rptg_get_nodes(rpt_graph::in, list(rptg_node)::out) is det.
+:- func rptg_get_nodes(rpt_graph) = list(rptg_node).
+
+    % rptg_set_edge(Start, End, ArcInfo, Arc, OldGraph, NewGraph)
+    % takes a graph OldGraph and adds an arc from Start to End with
+    % the information ArcInfo stored in it, and returns a key for
+    % that arc Arc, and the new graph NewGraph.
+    % If an identical arc already exists then this operation has
+    % no effect.
+    %
+    % This operation is O(lgN+lgM) for a graph with N nodes and M arcs.
+    %
+:- pred rptg_set_edge(rptg_node::in, rptg_node::in, rptg_arc_content::in, 
+    rptg_arc::out, rpt_graph::in, rpt_graph::out) is det.
+
+    % rptg_arc_contents(Graph, Arc, Start, End, ArcInfo) takes a
+    % graph Graph and an arc Arc and returns the start and end nodes
+    % and the content stored in that arc.
+    %
+:- pred rptg_arc_contents(rpt_graph::in, rptg_arc::in, rptg_node::out, 
+    rptg_node::out, rptg_arc_content::out) is det.
+
+    % rptg_path(Graph, Start, End, Path) is true iff there is a path
+    % from the node Start to the node End in Graph that goes through
+    % the sequence of arcs Arcs.
+    % The algorithm will return paths containing at most one cycle.
+    %
+:- pred rptg_path(rpt_graph, rptg_node, rptg_node, list(rptg_arc)). 
+:- mode rptg_path(in, in, in, out) is nondet.
+:- mode rptg_path(in, in, out, out) is nondet.
+
+:- pred reachable_and_having_type(rpt_graph::in, rptg_node::in,
mer_type::in, 
+    rptg_node::out) is semidet.
+
+:- pred rptg_get_node_supply(rpt_graph::in, rptg_node_supply::out) is det.
+:- pred rptg_get_arc_supply(rpt_graph::in, rptg_arc_supply::out) is det.
+:- pred rptg_get_nodemap(rpt_graph::in, 
+    map(rptg_node, rptg_node_content)::out) is det.
+:- pred rptg_get_arcmap(rpt_graph::in, map(rptg_arc, rptg_arc_info)::out) 
+    is det.
+:- pred rptg_get_edgemap(rpt_graph::in, 
+    map(rptg_node, map(rptg_arc, rptg_node))::out) is det.
+
+:- pred rptg_set_node_supply(rptg_node_supply::in, 
+    rpt_graph::in, rpt_graph::out)is det.
+:- pred rptg_set_arc_supply(rptg_arc_supply::in, 
+    rpt_graph::in, rpt_graph::out) is det.
+:- pred rptg_set_nodemap(map(rptg_node, rptg_node_content)::in, 
+    rpt_graph::in, rpt_graph::out) is det.
+:- pred rptg_set_arcmap(map(rptg_arc, rptg_arc_info)::in, 
+    rpt_graph::in, rpt_graph::out) is det.
+:- pred rptg_set_edgemap(map(rptg_node, map(rptg_arc, rptg_node))::in,
+    rpt_graph::in, rpt_graph::out) is det.
+
+    % get a node given the region name (region variable) assigned to it.
+    % There is one and only one node with a given region name. 
+    % Therefore the predicate returns the node as soon as it finds.
+    %
+:- pred get_node_by_region_name(rpt_graph::in, string::in, 
+    rptg_node::out) is det.
+      
+    % Get a node given a set of Mercury variables assigned to it.
+    % There is one and only one node corresponding to a set of variables.
+    % Therefore the predicate returns the node as soon as it finds.
+    %
+:- pred get_node_by_varset(rpt_graph::in, set(prog_var)::in, 
+    rptg_node::out) is det.
+
+    % Get a node given a Mercury variable assigned to it.
+    % There is one and only one node of a variable. 
+    % Therefore the predicate returns the node as soon as it finds.
+    %
+:- pred get_node_by_variable(rpt_graph::in, prog_var::in, 
+    rptg_node::out) is det.
+
+    % Get a node given a node that has been merged into the first one.
+    %
+:- pred get_node_by_node(rpt_graph::in, rptg_node::in, rptg_node::out)
is det.
+ 
+    % compare two graphs
+    %
+:- pred rptg_equal(rpt_graph::in, rpt_graph::in) is semidet.
+
+    % This predicate returns a set of nodes (regions) reachable from a 
+    % variable X
+    %
+:- pred reach_from_a_variable(rpt_graph::in, module_info::in,
proc_info::in, 
+    prog_var::in, set(rptg_node)::in, set(rptg_node)::out) is det.
+
+:- func rptg_lookup_region_name(rpt_graph, rptg_node) = string.
+:- func rptg_lookup_node_type(rpt_graph, rptg_node) = mer_type.
+
+    % Implement the unify operator 
+    % 
+:- pred unify_operator(rptg_node::in, rptg_node::in, 
+    rpt_graph::in, rpt_graph::out) is det.
+
+    % Implement the edge operator 
+    %
+:- pred edge_operator(rptg_node::in, rptg_node::in, rptg_arc_content::in, 
+    rpt_graph::in, rpt_graph::out) is det.
+
+:- implementation.
+
+:- import_module hlds.
+:- import_module hlds.hlds_data.
+:- import_module hlds.hlds_goal.
+:- import_module hlds.hlds_out.
+:- import_module parse_tree.mercury_to_mercury.
+:- import_module libs.
+:- import_module libs.compiler_util.
+:- import_module smm_utils.
+
+:- import_module assoc_list.
+:- import_module bool.
+:- import_module std_util.
+:- import_module term.
+:- import_module varset.
+:- import_module pair.
+:- import_module solutions.
+:- import_module svset.
+
+rpt_graph_init(Graph) :-
+    Graph = rpt_graph(0, 0, Nodes, Arcs, Edges),
+    map.init(Nodes),
+    map.init(Arcs),
+    map.init(Edges).
+
+rpt_graph_init = G :-
+    rpt_graph_init(G).
+
+    % After adding a node with Content into the graph, we need to
update the
+    % Content so that the merged_from set of the node contains itself.
+    % Doing it this way is not completely satisfied because we are adding a
+    % node with the given content but we change the content after all.
+    % But without adding the node first, the node is nonexistant and we 
+    % cannot make add it to the merged_from set. 
+    %
+rptg_set_node(Content0, rptg_node(N), G0, G) :-
+    rptg_get_node_supply(G0, NS0),
+    NS = NS0 + 1,
+    N = NS,
+    rptg_set_node_supply(NS, G0, G1),
+        
+        % make the merged_from set contain this node
+        %
+    set.init(MergedFrom0),
+    set.insert(MergedFrom0, rptg_node(NS), MergedFrom1),
+    rptg_node_content_set_merged_from(Content0, MergedFrom1, Content),
+
+    rptg_get_nodemap(G1, NodeMap0),
+    map.set(NodeMap0, rptg_node(N), Content, NodeMap),
+    rptg_set_nodemap(NodeMap, G1, G2),
+        
+        % set edge
+    rptg_get_edgemap(G2, EdgeMap0),
+    map.set(EdgeMap0, rptg_node(N), map.init, EdgeMap),
+    rptg_set_edgemap(EdgeMap, G2, G).
+
+rptg_node_contents(G, N, Content) :-
+    rptg_get_nodemap(G, NodeMap),
+    map.lookup(NodeMap, N, Content).
+
+rptg_node_contents(G, N) = NI :-
+    rptg_node_contents(G, N, NI).
+
+rptg_get_nodes(G, Ns) :-
+    rptg_get_nodemap(G, NodeMap),
+    map.keys(NodeMap, Ns).
+
+rptg_get_nodes(G) = Ns :-
+    rptg_get_nodes(G,Ns).
+
+rptg_successors(G, N, Ss) :-
+    rptg_get_edgemap(G, EdgeMap),
+    map.lookup(EdgeMap, N, OutEdges),
+    map.values(OutEdges, SsList),
+    set.list_to_set(SsList, Ss).
+
+rptg_successors(G, N) = S :-
+    rptg_successors(G, N, S).
+
+rptg_set_edge(Start, End, Info, Arc, G0, G) :-
+    rptg_get_arc_supply(G0, AS0),
+    AS is AS0 + 1,
+    Arc = rptg_arc(AS),
+    rptg_set_arc_supply(AS, G0, G1),
+
+    rptg_get_arcmap(G1, ArcMap0),
+    map.set(ArcMap0, Arc, rptg_arc_info(Start, End, Info), ArcMap),
+    rptg_set_arcmap(ArcMap, G1, G2),
+    
+    % register into edge set of the Start node
+    rptg_get_edgemap(G2, EdgeMap0),
+    map.lookup(EdgeMap0, Start, OutEdges0),
+    map.set(OutEdges0, Arc, End, OutEdges),
+    map.set(EdgeMap0, Start, OutEdges, EdgeMap),
+    rptg_set_edgemap(EdgeMap, G2, G).
+
+rptg_arc_contents(G, Arc, S, E, Content) :-
+    rptg_get_arcmap(G, ArcMap),
+    map.lookup(ArcMap, Arc, I),
+    I = rptg_arc_info(S, E, Content).
+
+rptg_path(G, S, E, Path) :-
+    rptg_path_2(G, S, E, [], Path).
+
+:- pred rptg_path_2(rpt_graph, rptg_node, rptg_node, list(rptg_node), 
+    list(rptg_arc)).
+:- mode rptg_path_2(in, in, in, in, out) is nondet.
+:- mode rptg_path_2(in, in, out, in, out) is nondet.
+
+rptg_path_2(G, S, E, Nodes0, Path) :-
+    rptg_get_edgemap(G, EdgeMap),
+    map.lookup(EdgeMap, S, OutEdgesOfS),
+    (
+        map.member(OutEdgesOfS, A, E),
+        \+ list.member(E, Nodes0),
+        Path = [A]
+    ;
+        map.member(OutEdgesOfS, A, N),
+        \+ list.member(N, Nodes0),
+        rptg_path_2(G, N, E, [N|Nodes0], Path0),
+        Path = [A|Path0]
+    ).
+
+    % Find a node that is reachable from Start and has type EType.
+    % if not found, fails.
+    %
+reachable_and_having_type(Graph, Start, EType, E) :-
+    rptg_lookup_node_type(Graph, Start) = Type,
+    ( if
+        Type = EType
+      then
+        E = Start
+      else
+        reachable_and_having_type_2(Graph, [Start], [Start], EType, E)
+    ).
+
+    % the algorithm here means that a node may be checked more than once. 
+    % This is not optimal but the algorithm ensures that each node may
become
+    % "Start" node at most once, therefore it will terminate.
+    %
+:- pred reachable_and_having_type_2(rpt_graph::in, list(rptg_node)::in, 
+        list(rptg_node)::in, mer_type::in, rptg_node::out) is semidet.
+
+reachable_and_having_type_2(Graph, [Start|Others], StartedFrom, EType,
E) :-
+    rptg_get_edgemap(Graph, EdgeMap),
+    map.lookup(EdgeMap, Start, OutArcs),
+    map.values(OutArcs, Ends),
+    ( if
+        end_node_with_same_type(Ends, Graph, EType, E1)
+      then 
+        E = E1
+      else
+        % breath-first search
+        list.append(Others, Ends, NewStarts0),
+
+        % only "start" with nodes that we have never started from
+        list.remove_dups([Start|StartedFrom], NewStartedFrom),
+        list.delete_elems(NewStarts0, NewStartedFrom, NewStarts),
+        
+        reachable_and_having_type_2(Graph, NewStarts, NewStartedFrom,
EType, E)
+    ).
+    
+:- pred end_node_with_same_type(list(rptg_node)::in, rpt_graph::in, 
+    mer_type::in, rptg_node::out) is semidet.
+    
+end_node_with_same_type([N|Ns], Graph, Type, End) :-
+    rptg_lookup_node_type(Graph, N) = NType,
+    ( if
+        NType = Type
+      then
+        End = N
+      else
+        end_node_with_same_type(Ns, Graph, Type, End)
+    ).
+
+rptg_get_node_supply(G, NS) :-
+    G = rpt_graph(NS, _AS, _N, _A, _E).
+rptg_get_arc_supply(G, AS) :-
+    G = rpt_graph(_NS, AS, _N, _A, _E).
+rptg_get_nodemap(G, N) :-
+    G = rpt_graph(_NS, _AS, N, _A, _E).
+rptg_get_arcmap(G, A) :-
+    G = rpt_graph(_NS, _AS, _N, A, _E).
+rptg_get_edgemap(G, E) :-
+    G = rpt_graph(_NS, _AS, _N, _A, E).
+rptg_set_node_supply(NS, G0, G) :-
+    G0 = rpt_graph(_, AS, N, A, E),
+    G = rpt_graph(NS, AS, N, A, E).
+
+rptg_set_arc_supply(AS, G0, G) :-
+    G0 = rpt_graph(NS, _, N, A, E),
+    G = rpt_graph(NS, AS, N, A, E).
+rptg_set_nodemap(N, G0, G) :-
+    G0 = rpt_graph(NS, AS, _, A, E),
+    G = rpt_graph(NS, AS, N, A, E).
+rptg_set_arcmap(A, G0, G) :-
+    G0 = rpt_graph(NS, AS, N, _, E),
+    G = rpt_graph(NS, AS, N, A, E).
+rptg_set_edgemap(E, G0, G) :-
+    G0 = rpt_graph(NS, AS, N, A, _),
+    G = rpt_graph(NS, AS, N, A, E).
+    
+rptg_node_content_get_varset(NC, NC^varset).
+rptg_node_content_get_region_name(NC, NC^reg_var_name).
+rptg_node_content_get_merged_from(NC, NC^merged_from).
+rptg_node_content_get_node_type(NC, NC^node_type).
+
+rptg_node_content_set_varset(NC0, VarSet, NC) :-
+    NC0^varset := VarSet = NC.
+rptg_node_content_set_region_name(NC0, Name, NC) :-
+    NC0^reg_var_name := Name = NC.
+rptg_node_content_set_merged_from(NC0, Nodes, NC) :-
+    NC0^merged_from := Nodes = NC.
+rptg_node_content_set_node_type(NC0, NodeType, NC) :-
+    NC0^node_type := NodeType = NC.
+
+rptg_arc_content_get_label(AC, AC^label).
+rptg_arc_content_set_label(AC0, Label, AC) :-
+    AC0^label := Label = AC.
+
+    % find a node in the graph using its region name
+    %
+get_node_by_region_name(Graph, RegName, Node) :-
+    % from all nodes in the graph find a node corresponding to the
region name
+    rptg_get_nodes(Graph, AllNodes),
+    ( if 
+        get_node_by_region_name_from_list(Graph, AllNodes, RegName, Node0)
+      then
+        Node = Node0
+      else
+        unexpected(thisfile, "get_node_by_region_name: No such a node
exists")
+    ).
+
+:- pred get_node_by_region_name_from_list(rpt_graph::in,
list(rptg_node)::in, 
+    string::in, rptg_node::out) is semidet.
+get_node_by_region_name_from_list(Graph, List, RegName, Node) :- 
+    List = [ANode | Rest],
+    rptg_node_contents(Graph, ANode, NodeInfo),
+    ( if 
+        NodeInfo^reg_var_name = RegName
+      then
+        Node = ANode
+      else
+        get_node_by_region_name_from_list(Graph, Rest, RegName, Node)
+    ).
+
+    % find a node in the graph using a set of variables.
+    % Because a variable is assigned to only one node.
+    %
+get_node_by_varset(Graph, Varset, Node) :-
+    rptg_get_nodes(Graph, AllNodes),
+    ( if
+        get_node_by_varset_from_list(Graph, AllNodes, Varset, Node0)
+      then
+        Node = Node0
+      else
+        unexpected(thisfile, "get_node_by_varset: No such a node exists")
+    ).
+
+:- pred get_node_by_varset_from_list(rpt_graph::in, list(rptg_node)::in, 
+    set(prog_var)::in, rptg_node::out) is semidet.
+get_node_by_varset_from_list(Graph, List, Varset, Node) :- 
+    List = [ANode | Rest],
+    rptg_node_contents(Graph, ANode, NodeInfo),
+    ( if 
+          set.subset(Varset, NodeInfo^varset)
+      then
+          Node = ANode
+      else
+          get_node_by_varset_from_list(Graph, Rest, Varset, Node)
+    ).
+
+    % find a node in the graph using a variable assigned to it.
+    % simply call get_node_by_varset.
+    %
+get_node_by_variable(Graph, Var, Node) :- 
+    % make a set(prog_var) containing the variable
+    set.init(Varset0),
+    set.insert(Varset0, Var, Varset),
+    % find the node
+    get_node_by_varset(Graph, Varset, Node).
+
+get_node_by_node(Graph, Node, MergedNode) :-
+    % first the node in the NodeMap
+    rptg_get_nodemap(Graph, NodeMap),
+    ( if 
+        map.search(NodeMap, Node, _NodeContent)
+      then
+        MergedNode = Node
+      else
+        % not directly in the NodeMap, checked if it has been merged
+        rptg_get_nodes(Graph, AllNodes),
+        ( if
+            get_node_by_node_from_list(Graph, AllNodes, Node, MergedNode0)
+          then
+            MergedNode = MergedNode0
+          else
+            unexpected(thisfile, "get_node_by_node: No such a node exists")
+        )
+    ).
+
+:- pred get_node_by_node_from_list(rpt_graph::in, list(rptg_node)::in,
+    rptg_node::in, rptg_node::out) is semidet.
+get_node_by_node_from_list(Graph, List, Node, MergedNode) :- 
+        List = [ANode|Rest],
+        rptg_node_contents(Graph, ANode, NodeInfo),
+        ( if 
+              set.member(Node, NodeInfo^merged_from)
+          then
+              MergedNode = ANode
+          else
+              get_node_by_node_from_list(Graph, Rest, Node, MergedNode)
+        ).
+                
+rptg_lookup_region_name(Graph, Node) = RegionName :-
+    rptg_node_contents(Graph, Node, Content),
+    Content^reg_var_name = RegionName.
+
+rptg_lookup_node_type(Graph, Node) = NodeType :-
+    rptg_node_contents(Graph, Node, Content),
+    Content^node_type = NodeType.
+
+%-----------------------------------------------------------------------------%
+%
+% the two graph-manipulating operators, i.e., unify and edge
+%
+
+unify_operator(Node1, Node2, !Graph) :-
+    % the varset need to be unioned
+    rptg_get_nodemap(!.Graph, NodeMap0), 
+    rptg_node_contents(!.Graph, Node1, NodeContent1),
+    rptg_node_contents(!.Graph, Node2, NodeContent2),
+
+    rptg_node_content_set_varset(NodeContent1, 
+        set.union(NodeContent1^varset, NodeContent2^varset), NewContent0),
+    rptg_node_content_set_merged_from(NewContent0, 
+        set.union(NodeContent1^merged_from, NodeContent2^merged_from), 
+        NewContent1),
+    map.det_update(NodeMap0, Node1, NewContent1, NodeMap1),
+    
+    rptg_set_nodemap(NodeMap1, !Graph),
+    
+    % copy all out-edges of node 2 to node 1
+    transfer_out_edges(Node1, Node2, !Graph), 
+  
+    % copy all in-edges of node 2 to node 1
+    transfer_in_edges(Node1, Node2, !Graph), 
+
+    % remove node 2
+    delete_node(Node2, !Graph).
+	
+    % This predicate receives a graph and returns a new graph in which 
+	% for all the out-edges of node2 in the first graph are copied to 
+	% node1, as out-edges of node1.
+    %
+:- pred transfer_out_edges(rptg_node::in, rptg_node::in, rpt_graph::in, 
+    rpt_graph::out) is det.
+
+transfer_out_edges(Node1, Node2, !Graph) :- 
+    % out-edges from node 2
+    rptg_get_edgemap(!.Graph, EdgeMap),
+    map.lookup(EdgeMap, Node2, OutEdges2Map),
+    map.keys(OutEdges2Map, OutArcs),
+    transfer_out_edges_2(OutArcs, Node1, !Graph).
+
+	% This predicate receives a list of out-edges of node2 and returns a
graph 
+    % with all the edges in the list copied to Node1, but it maintains the 
+    % invariant that "there is only one arc with a specific label from a 
+    % specific node to another specific node".        
+	% The algorithm is as follows.
+	% for each arc in OutArcs2 list:
+	%    if Node1 points to the cor. node
+	%         compare the content of the arc with content of that arc
+	%              if content the same, ignore the arc,
+	%              if content different, copy the arc to Node1
+	%    no, copy the arc to Node1
+    %
+:- pred transfer_out_edges_2(list(rptg_arc)::in, rptg_node::in,
rpt_graph::in, 
+    rpt_graph::out) is det.
+
+transfer_out_edges_2([], _N, !Graph).
+transfer_out_edges_2([Arc|Arcs], Node1, !Graph) :-
+    rptg_arc_contents(!.Graph, Arc, _Node2, CorNode, ArcContent),
+    ( if 
+        % if Node1 points to the cor. node
+        rptg_successors(!.Graph, Node1, SsNode1),
+        set.contains(SsNode1, CorNode)
+      then
+        % need to check if an edge from node1 to the cor. node with the
same 
+        % label already existed
+        rptg_get_edgemap(!.Graph, EdgeMap),
+        map.lookup(EdgeMap, Node1, OutEdges1Map),
+        solutions(map.inverse_search(OutEdges1Map, CorNode), 
+            ArcsFromNode1ToCorNode),
+        ( if 
+            same_label(ArcsFromNode1ToCorNode, !.Graph, Arc)
+          then
+           % existed therefore ignore this Arc and continue with the other 
+           % out-edges of node 2
+            transfer_out_edges_2(Arcs, Node1, !Graph)
+          else 
+               % not existed, copy the Arc as an out-edge of Node1
+               rptg_set_edge(Node1, CorNode, ArcContent, _Arc, !Graph),
+
+               % continue with the other out-edges of node 2
+               transfer_out_edges_2(Arcs, Node1, !Graph)
+        ) 
+      else
+          % No edge from Node1 to CorNode, copy the Arc as an out-edge of 
+          % Node1
+          rptg_set_edge(Node1, CorNode, ArcContent, _Arc, !Graph),
+
+          % continue with the other out-edges of node 2
+          transfer_out_edges_2(Arcs, Node1, !Graph)
+    ).
+
+	% check if any arc in the list has the same label as the label in the 
+    % input Arc
+    %
+:- pred same_label(list(rptg_arc)::in, rpt_graph::in, rptg_arc::in) is
semidet.
+
+same_label([A|As], Graph, Arc) :-
+    rptg_arc_contents(Graph, Arc, _Start, _End, ArcContent), 
+    rptg_arc_contents(Graph, A, _S, _E, AC),
+	( if 
+        ArcContent = AC
+      then 
+        true
+      else
+        same_label(As, Graph, Arc)
+    ).
+
+	% This predicate receives a graph and returns a new graph in which 
+	% for all the in-edges of node2 in the first graph are copied to 
+	% node1, as in-edges of node1.
+    %
+:- pred transfer_in_edges(rptg_node::in, rptg_node::in, rpt_graph::in, 
+    rpt_graph::out) is det.
+
+transfer_in_edges(Node1, Node2, !Graph) :-
+    % in-edges of node 2
+    rptg_get_in_arcs(!.Graph, Node2, InArcs),
+
+    % copy them to node 1
+    transfer_in_edges_2(InArcs, Node1, !Graph).
+
+:- pred rptg_get_in_arcs(rpt_graph::in, rptg_node::in,
list(rptg_arc)::out) 
+    is det.
+
+rptg_get_in_arcs(Graph, Node, Arcs) :-
+    rptg_get_arcmap(Graph, ArcMap),
+    map.foldl(arc_to_node(Node), ArcMap, [], Arcs).
+
+:- pred arc_to_node(rptg_node::in, rptg_arc::in, rptg_arc_info::in, 
+    list(rptg_arc)::in, list(rptg_arc)::out) is det.
+arc_to_node(End, Arc, ArcInfo, !L) :-
+    ArcInfo = rptg_arc_info(_S, E, _C),
+    ( if
+        E = End
+      then
+        !:L = [Arc|!.L]
+      else
+        true
+    ).
+
+:- pred transfer_in_edges_2(list(rptg_arc)::in, rptg_node::in,
rpt_graph::in,  
+    rpt_graph::out) is det.
+transfer_in_edges_2([], _N, !Graph).
+transfer_in_edges_2([Arc|Arcs], Node1, !Graph) :-
+    rptg_arc_contents(!.Graph, Arc, CorNode, _Node2, ArcContent),
+    ( if 
+        % if CorNode points to Node1
+        rptg_successors(!.Graph, CorNode, SsCorNode),
+        set.contains(SsCorNode, Node1)
+      then
+        % need to check if an edge from cor. node to node1 with the same 
+        % label already existed
+        rptg_get_edgemap(!.Graph, EdgeMap),
+        map.lookup(EdgeMap, CorNode, OutEdgesCorNodeMap),
+        solutions(map.inverse_search(OutEdgesCorNodeMap, Node1), 
+            ArcsFromCorNodeToNode1),
+        ( if 
+            same_label(ArcsFromCorNodeToNode1, !.Graph, Arc)
+          then
+            % existed therefore ignore this Arc and continue with the
other 
+            % in-edges of node 2
+            transfer_in_edges_2(Arcs, Node1, !Graph)
+          else 
+            % not existed, copy the Arc as an in-edge of Node1
+            rptg_set_edge(CorNode, Node1, ArcContent, _Arc, !Graph),
+
+            % continue with the other in-edges of node 2
+            transfer_in_edges_2(Arcs, Node1, !Graph)
+        ) 
+          else
+            % No edge from CorNode to Node1, copy the Arc as an in-edge of 
+            % Node1
+            rptg_set_edge(CorNode, Node1, ArcContent, _Arc, !Graph),
+
+            % continue with the other in-edges of node 2
+            transfer_in_edges_2(Arcs, Node1, !Graph)
+).
+
+    % delete a node from the graph
+    %
+:- pred delete_node(rptg_node::in, rpt_graph::in, rpt_graph::out) is det.
+
+delete_node(Node, Graph0, Graph) :-
+    Graph0 = rpt_graph(NS, AS, NodeMap0, ArcMap0, EdgeMap0),
+    
+    % delete the node
+    map.delete(NodeMap0, Node, NodeMap),
+        
+    % we also need to delete all the edges and arcs from the Node
+    delete_all_outedges_and_arcs(Node, EdgeMap0, ArcMap0, EdgeMap1,
ArcMap1),
+    
+    % and also need to delete all edges and arcs to the Node
+    delete_all_inedges_and_arcs(Node, EdgeMap1, ArcMap1, EdgeMap, ArcMap),
+    
+    Graph = rpt_graph(NS, AS, NodeMap, ArcMap, EdgeMap).
+
+    % This predicate deletes all the edges which start from the input 
+    % node Node. 
+    % Note: it works as a helper for delete_node so it does not proceed
+    % on a graph but on the edge map and the arc map.
+    %
+:- pred delete_all_outedges_and_arcs(rptg_node::in,
+    map(rptg_node, map(rptg_arc, rptg_node))::in,
+    map(rptg_arc, rptg_arc_info)::in,
+    map(rptg_node, map(rptg_arc, rptg_node))::out,
+    map(rptg_arc, rptg_arc_info)::out) is det.
+    
+delete_all_outedges_and_arcs(Node, EdgeMap0, ArcMap0, EdgeMap, ArcMap) :-
+    % delete all outcoming edges of the Node from edge map
+    map.delete(EdgeMap0, Node, EdgeMap),
+    
+    % delete the corresponding arcs from arc map
+    map.lookup(EdgeMap0, Node, EdgesFromNodeMap),
+    map.keys(EdgesFromNodeMap, OutArcs),
+    map.delete_list(ArcMap0, OutArcs, ArcMap).
+
+    % This predicate deletes all the incoming edges of the input node
(Node).
+    % We only store outcoming edges therefore to remove incoming ones
of Node
+    % we need to check all the outcoming edges and remove those point
to Node.
+    %
+:- pred delete_all_inedges_and_arcs(rptg_node::in, 
+    map(rptg_node, map(rptg_arc, rptg_node))::in,
+    map(rptg_arc, rptg_arc_info)::in, 
+    map(rptg_node, map(rptg_arc, rptg_node))::out,
+    map(rptg_arc, rptg_arc_info)::out) is det.
+
+delete_all_inedges_and_arcs(Node, EdgeMap0, ArcMap0, EdgeMap, ArcMap) :-
+    map.keys(EdgeMap0, StartNodes),
+    
+        % for each node: find the outcoming edges from it 
+        % and delete ones pointing to Node
+        %
+    delete_all_inedges_and_arcs_2(StartNodes, EdgeMap0, ArcMap0, 
+        Node, EdgeMap, ArcMap).
+
+    % This predicate receives a node (Node) and a list of (start) nodes. 
+    % It deletes all the start nodes' outcoming edges and corresponding
arcs
+    % which point to the node.
+    %
+:- pred delete_all_inedges_and_arcs_2(list(rptg_node)::in, 
+    map(rptg_node, map(rptg_arc, rptg_node))::in,
+    map(rptg_arc, rptg_arc_info)::in, rptg_node::in,
+    map(rptg_node, map(rptg_arc, rptg_node))::out,
+    map(rptg_arc, rptg_arc_info)::out) is det.
+
+delete_all_inedges_and_arcs_2([], EdgeMap0, ArcMap0, _, EdgeMap0, ArcMap0).
+delete_all_inedges_and_arcs_2([N|Ns], EdgeMap0, ArcMap0, Node, EdgeMap, 
+    ArcMap) :-
+    % the corresponding map
+    map.lookup(EdgeMap0, N, EdgesFromNodeMap0), 
+        
+    % find the mapping with end node = Node, there will be only one,
+    % but we do as below because inverse_search is non-det.
+    solutions(map.inverse_search(EdgesFromNodeMap0, Node), ArcList),
+    map.delete_list(EdgesFromNodeMap0, ArcList, EdgesFromNodeMap),
+    map.delete_list(ArcMap0, ArcList, ArcMap1),
+    map.set(EdgeMap0, N, EdgesFromNodeMap, EdgeMap1),
+    delete_all_inedges_and_arcs_2(Ns, EdgeMap1, ArcMap1, Node, 
+        EdgeMap, ArcMap).
+    
+edge_operator(Start, End, Info, !G) :-
+	rptg_set_edge(Start, End, Info, _Arc, !G).
+	
+%-----------------------------------------------------------------------------%
+%-----------------------------------------------------------------------------%
+%
+% equality of region points-to graphs.
+%
+
+	% equal for rpt_graph
+rptg_equal(Graph1, Graph2) :-
+	Graph1 = rpt_graph(NS1, AS1, Nodes1, Arcs1, Edges1),
+	Graph2 = rpt_graph(NS2, AS2, Nodes2, Arcs2, Edges2),
+	NS1 = NS2,
+	AS1 = AS2,
+		% compare these maps
+	simple_map_equal(Nodes1, Nodes2),
+	simple_map_equal(Arcs1, Arcs2),
+	complex_map_equal(Edges1, Edges2).
+
+% The comparisons below may not be necessary, unification can help if it is
+% sure that the elements are added to the maps in the same order.
+% 
+	% The values of the maps are required to be comparable using unification.
+    % i.e., values of type V1 can be compared using unification
+    %
+:- pred simple_map_equal(map(K1, V1)::in, map(K1, V1)::in) is semidet.
+
+simple_map_equal(Map1, Map2) :- 
+    % check if they have the same number of entries?
+	map.count(Map1, C1),
+	map.count(Map2, C2),
+	C1 = C2,
+
+    % if yes, check if all the entries are equal 
+	map.keys(Map1, Ks1),
+	simple_map_equal_2(Ks1, Map1, Map2).
+
+    % With the condition that the two maps have the same number of entries,
+    % verify that all keys in map 1 are also in map 2 and
+    % their corresponding values are equal.
+    %
+:- pred	simple_map_equal_2(list(K1)::in, map(K1, V1)::in, map(K1, V1)::in) 
+    is semidet.
+
+simple_map_equal_2([], _, _).
+simple_map_equal_2([K|Ks], Map1, Map2) :-
+    % K is also in map 2?
+	map.search(Map2, K, V2),
+
+    % Yes, so check whether the values are equal
+	map.lookup(Map1, K, V1),
+	V1 = V2,
+	simple_map_equal_2(Ks, Map1, Map2).
+
+	% The maps need to be of map-in-map structure, namely map(k1, map(k2, v))
+    % and values of type V can be compared by unifying (i.e., in our notion
+    % here map(k2, v) is a "simple" map).
+    %
+:- pred complex_map_equal(map(K1, map(K2, V))::in, map(K1, map(K2,
V))::in) 
+    is semidet.
+
+complex_map_equal(Map1, Map2) :- 
+	map.count(Map1, C1),
+	map.count(Map2, C2),
+	C1 = C2,
+	map.keys(Map1, Ks1),
+	complex_map_equal_2(Ks1, Map1, Map2).
+
+:- pred	complex_map_equal_2(list(K1)::in, map(K1, map(K2, V))::in, 
+    map(K1, map(K2, V))::in) is semidet.
+
+complex_map_equal_2([], _, _).
+complex_map_equal_2([K|Ks], Map1, Map2) :-
+	map.search(Map2, K, V2), 
+    
+    % V2 is "simple" map, so compare it with V1 
+	map.lookup(Map1, K, V1),
+	simple_map_equal(V1, V2),
+	complex_map_equal_2(Ks, Map1, Map2).
+
+reach_from_a_variable(Graph, HLDS, ProcInfo, X, Reach0, Reach) :-
+    % region of X
+	get_node_by_variable(Graph, X, N_X),
+	Node_Selector = pair(N_X, []),
+	
+    % type of X
+	proc_info_get_vartypes(ProcInfo, VarTypes),
+	map.lookup(VarTypes, X, TypeX),
+
+    % find regions reached from X
+	set.init(Reach_X0),
+	reach_from_nodes_reachable_from_a_var([Node_Selector], Graph, HLDS, 
+        TypeX, [], Reach_X0, Reach_X),
+	
+    % accumulate them
+	set.union(Reach0, Reach_X, Reach). 
+	
+	% this predicate finds all regions that are reachable from X
+	% Algorithm: 
+	% 	1. receives a list of to-be-processed nodes, 
+	%	2. each node needs to be added to the reach_from_x set and then 
+	%	3. the target of each of its out-arcs is examined to see if 
+	% 		we should add it to the to-be-processed list or not.
+	%	4. do until the to-be-processed is empty.
+	%
+:- pred reach_from_nodes_reachable_from_a_var(
+    assoc_list(rptg_node, selector)::in, rpt_graph::in, module_info::in, 
+    mer_type::in, list(rptg_node)::in, set(rptg_node)::in, 
+    set(rptg_node)::out) is det.
+reach_from_nodes_reachable_from_a_var([], _Graph, _HLDS, _TypeX,
_Processed, 
+    !Reach_X).
+reach_from_nodes_reachable_from_a_var([Node_Selector|Others0], Graph,
HLDS, 
+    TypeX, Processed0, !Reach_X) :-
+	Node_Selector = Node - Selector,
+
+    % add the Node to reach_from_x set
+	svset.insert(Node, !Reach_X),
+	
+    % add the Node to processed list (Node is not yet in Processed0
because 
+    % if it is in there it will not be in the to-be-processed list.
+    %
+	Processed1 = [Node|Processed0],
+	
+		% take out-arcs of the Node and update the to-be-processed list 
+		%
+	rptg_get_edgemap(Graph, EdgeMap),
+	map.lookup(EdgeMap, Node, OutEdges), 
+	map.keys(OutEdges, OutArcs),
+	list.foldl(
+		update_to_be_processed_list(Selector, HLDS, TypeX, Graph, Processed1), 
+		OutArcs, Others0, Others1),
+	
+		% go on with other nodes
+		%
+	reach_from_nodes_reachable_from_a_var(Others1, Graph, HLDS, TypeX, 
+        Processed1, !Reach_X).
+
+	% Algorithm:
+	%	A target is added to the to-be-processed list only if its 
+	%		arc's selector is valid.
+	% 	The to-be-processed list is added at the end, so it is a 
+    %   breadth-first process.
+	%
+:- pred update_to_be_processed_list(selector::in, module_info::in, 
+    mer_type::in, rpt_graph::in, list(rptg_node)::in, rptg_arc::in, 
+	assoc_list(rptg_node, selector)::in, 
+    assoc_list(rptg_node, selector)::out) is det.
+update_to_be_processed_list(Selector, HLDS, TypeX, Graph, Processed,
OutArc, 
+    List0, List) :-
+	rptg_arc_contents(Graph, OutArc, _Start, End, ArcContent),
+	ArcSelector = ArcContent^label,
+	list.append(Selector, ArcSelector, NewSelector),
+	( if
+		check_type_of_node(HLDS, TypeX, NewSelector)
+	  then
+	  	% the arc's selector is a valid one.
+		( if
+			list.member(End, Processed)
+		  then
+		  	List = List0
+		  else
+            % here is a non-processed node and can be reached from X by a 
+            % valid selector, so it needs to be processed
+            %
+		  	list.append(List0, [pair(End, NewSelector)], List)
+		)
+	  else
+	  	% not a valid one.
+		List = List0
+	).
+
+:- func thisfile = string.
+thisfile = "rpt_graph.m".
Index: rpta_fixpoint_table.m
===================================================================
RCS file: rpta_fixpoint_table.m
diff -N rpta_fixpoint_table.m
--- /dev/null	1 Jan 1970 00:00:00 -0000
+++ rpta_fixpoint_table.m	15 May 2007 05:30:09 -0000
@@ -0,0 +1,137 @@
+%-----------------------------------------------------------------------------%
+% vim: ft=mercury ts=4 sw=4 et
+%-----------------------------------------------------------------------------%
+% Copyright (C) 2005-2007 The University of Melbourne.
+% This file may only be copied under the terms of the GNU General
+% Public License - see the file COPYING in the Mercury distribution.
+%-----------------------------------------------------------------------------%
+
+% file rpta_fixpoint_table.m.
+% main author: Quan Phan
+%
+% Defines the fixpoint table used in the region points-to analysis.
+
+:- module rpta_fixpoint_table.
+
+:- interface.
+
+:- import_module hlds.
+:- import_module hlds.hlds_pred.
+:- import_module rpta_info.
+
+:- import_module list.
+
+:- type rpta_info_fixpoint_table.
+
+	% Initialise the fixpoint table for the given set of pred_proc_id's. 
+    %
+:- pred rpta_info_fixpoint_table_init(list(pred_proc_id)::in, 
+    rpta_info_table::in, rpta_info_fixpoint_table::out) is det.
+
+	% Add the results of a new analysis pass to the already existing
+	% fixpoint table. 
+    %
+:- pred rpta_info_fixpoint_table_new_run(rpta_info_fixpoint_table::in, 
+    rpta_info_fixpoint_table::out) is det.
+
+	% The fixpoint table keeps track of the number of analysis passes. This
+	% predicate returns this number.
+    %
+:- pred rpta_info_fixpoint_table_which_run(rpta_info_fixpoint_table::in, 
+    int::out) is det.
+
+	% A fixpoint is reached if all entries in the table are stable,
+	% i.e. haven't been modified by the last analysis pass. 
+    %
+:- pred rpta_info_fixpoint_table_all_stable(rpta_info_fixpoint_table::in) 
+    is semidet.
+
+	% Enter the newly computed region points-to information for a given 
+    % procedure.
+	% If the description is different from the one that was already stored
+	% for that procedure, the stability of the fixpoint table is set to
+	% "unstable". 
+	% Aborts if the procedure is not already in the fixpoint table. 
+    %
+:- pred rpta_info_fixpoint_table_new_rpta_info(pred_proc_id::in, 
+    rpta_info::in, rpta_info_fixpoint_table::in, 
+    rpta_info_fixpoint_table::out) is det.
+
+	% Retreive the rpta_info of a given pred_proc_id. If this information 
+    % is not available, this means that
+	% the set of pred_proc_id's to which the fixpoint table relates are
+	% mutually recursive, hence the table is characterised as recursive. 
+	% Fails if the procedure is not in the table. 
+    %
+:- pred rpta_info_fixpoint_table_get_rpta_info(pred_proc_id::in, 
+    rpta_info::out, rpta_info_fixpoint_table::in, 
+    rpta_info_fixpoint_table::out) is semidet.
+
+	% Retreive rpta_info, without changing the table. To be used after 
+    % fixpoint has been reached. Aborts if the procedure is not in the
table.
+    %
+:- pred rpta_info_fixpoint_table_get_final_rpta_info(pred_proc_id::in, 
+    rpta_info::out, rpta_info_fixpoint_table::in) is det.
+
+	% Same as rpta_info_fixpoint_table_get_final_as, yet fails instead of 
+    % aborting if the procedure is not in the table.
+    %
+:- pred
rpta_info_fixpoint_table_get_final_rpta_info_semidet(pred_proc_id::in, 
+    rpta_info::out, rpta_info_fixpoint_table::in) is semidet.
+
+:- implementation.
+
+:- import_module fixpoint_table.
+:- import_module libs.
+:- import_module libs.compiler_util.
+
+:- import_module string.
+
+:- type rpta_info_fixpoint_table == 
+		fixpoint_table(pred_proc_id, rpta_info). 
+
+:- pred wrapped_init(rpta_info_table::in, pred_proc_id::in, rpta_info::out)
+    is det.
+    
+wrapped_init(InfoTable, PredProcId, E) :- 
+	( if 
+		rpta_info_table_search_rpta_info(PredProcId, InfoTable) = Entry
+	  then
+		E = Entry
+	  else
+        % The information we are looking for should be there after the
+        % intraprocedural analysis.
+		unexpected(thisfile, "wrapper_init: rpta_info should exist.")
+	).
+
+rpta_info_fixpoint_table_init(Keys, InfoTable, Table):-
+	fp_init(wrapped_init(InfoTable), Keys, Table).
+
+rpta_info_fixpoint_table_new_run(Tin, Tout) :-
+	fp_new_run(Tin,Tout).
+
+rpta_info_fixpoint_table_which_run(Tin, Run) :-
+	Run = fp_which_run(Tin).
+
+rpta_info_fixpoint_table_all_stable(Table) :-
+	fp_stable(Table).
+
+rpta_info_fixpoint_table_new_rpta_info(PredProcId, RptaInfo, Tin, Tout) :-
+	fp_add(
+		pred(TabledElem::in, Elem::in) is semidet :-
+		(
+			rpta_info_equal(Elem, TabledElem)
+		), 
+		PredProcId, RptaInfo, Tin, Tout).
+
+rpta_info_fixpoint_table_get_rpta_info(PredProcId, RptaInfo, Tin, Tout) :-
+	fp_get(PredProcId, RptaInfo, Tin, Tout).
+
+rpta_info_fixpoint_table_get_final_rpta_info(PredProcId, RptaInfo, T):-
+	fp_get_final(PredProcId, RptaInfo, T).
+
+rpta_info_fixpoint_table_get_final_rpta_info_semidet(PredProcId,
RptaInfo, T):-
+	fp_get_final_semidet(PredProcId, RptaInfo, T).
+
+:- func thisfile = string.
+thisfile = "rpta_fixpoint_table.m".
Index: rpta_info.m
===================================================================
RCS file: rpta_info.m
diff -N rpta_info.m
--- /dev/null	1 Jan 1970 00:00:00 -0000
+++ rpta_info.m	15 May 2007 05:31:06 -0000
@@ -0,0 +1,102 @@
+%-----------------------------------------------------------------------------%
+% vim: ft=mercury ts=4 sw=4 et
+%-----------------------------------------------------------------------------%
+% Copyright (C) 2005-2007 The University of Melbourne.
+% This file may only be copied under the terms of the GNU General
+% Public License - see the file COPYING in the Mercury distribution.
+%-----------------------------------------------------------------------------%
+%
+% File rpta_info.m.
+% Main author: Quan Phan.
+% 
+% This module defines the "rpta_info" and "rpta_info_table" types. 
+% rpta_info_table maps a procedure to its corresponding rpt information 
+% (i.e., the rpt graph and the alpha mappings (at the call sites in it)).
+
+:- module rpta_info.
+
+:- interface.
+
+:- import_module hlds.
+:- import_module hlds.hlds_pred.
+:- import_module rpt_graph.
+:- import_module rpt_alpha_mapping.
+
+:- import_module map.
+
+:- type rpta_info_table == map(pred_proc_id, rpta_info).
+
+:- func rpta_info_table_init = rpta_info_table.
+
+:- func rpta_info_table_search_rpta_info(pred_proc_id, rpta_info_table) 
+    = rpta_info is semidet.
+:- pred rpta_info_table_set_rpta_info(pred_proc_id::in, rpta_info::in, 
+    rpta_info_table::in, rpta_info_table::out) is det.
+
+    % type rpta_info and operations
+    %
+:- type rpta_info 
+        ---> rpta_info(rpt_graph, rpt_alpha_mapping).	
+
+    % The rpta_info will be for a specific procedure, at the beginning 
+    % the alpha mapping is empty and the rpt graph contains all the nodes 
+    % corresponding to all the variables appear in the procedure.
+    %
+:- pred rpta_info_init(proc_info::in, rpta_info::out) is det.
+:- func rpta_info_init(proc_info) = rpta_info.
+
+:- pred rpta_info_equal(rpta_info::in, rpta_info::in) is semidet.
+
+:- implementation.
+
+:- import_module hlds.hlds_module. 
+:- import_module parse_tree.
+:- import_module parse_tree.prog_data.
+:- import_module check_hlds.
+:- import_module check_hlds.type_util.
+:- import_module smm_data.
+
+:- import_module std_util.
+:- import_module set. 
+:- import_module list.
+:- import_module varset.
+:- import_module assoc_list.
+:- import_module string.
+:- import_module int.
+
+rpta_info_table_init = map.init. 
+rpta_info_table_search_rpta_info(PredProcId, Table) = RptaInfo :- 
+    Table^elem(PredProcId) = RptaInfo.
+rpta_info_table_set_rpta_info(PredProcId, RptaInfo, Table0, Table) :- 
+    Table = Table0^elem(PredProcId) := RptaInfo.
+
+    % The rpta_info will be for a specific procedure, so at the beginning 
+    % the alpha mapping is empty and the rpt graph contains all the nodes 
+    % corresponding to all the variables appear in the procedure.
+    %
+rpta_info_init(ProcInfo, RptaInfo) :-
+    proc_info_get_vartypes(ProcInfo, VarTypes),
+    map.keys(VarTypes, Vars),
+    list.foldl2(add_node_from_var(VarTypes), Vars, 1, _Reg, 
+	rpt_graph_init, Graph),
+    map.init(AlphaMapping),
+    RptaInfo = rpta_info(Graph, AlphaMapping).
+rpta_info_init(ProcInfo) = RptaInfo :- 
+    rpta_info_init(ProcInfo, RptaInfo).
+
+:- pred add_node_from_var(map(prog_var, mer_type)::in, prog_var::in,
int::in,
+    int::out, rpt_graph::in, rpt_graph::out) is det.
+add_node_from_var(VarTypes, Var, Reg0, Reg, !Graph) :-
+    map.lookup(VarTypes, Var, NodeType), 
+    set.init(Varset0),
+    set.insert(Varset0, Var, Varset),
+    Reg = Reg0 + 1,
+    string.append("R", string.int_to_string(Reg0), RegName),
+    NodeInfo = rptg_node_content(Varset, RegName, set.init, NodeType),
+    rptg_set_node(NodeInfo, _Node, !Graph).
+
+rpta_info_equal(RptaInfo1, RptaInfo2):-
+    RptaInfo1 = rpta_info(Graph1, Alpha1),
+    RptaInfo2 = rpta_info(Graph2, Alpha2), 
+    rptg_equal(Graph1, Graph2),
+    rpt_alpha_mapping_equal(Alpha1, Alpha2).
Index: rpta_run.m
===================================================================
RCS file: rpta_run.m
diff -N rpta_run.m
--- /dev/null	1 Jan 1970 00:00:00 -0000
+++ rpta_run.m	15 May 2007 02:35:56 -0000
@@ -0,0 +1,1199 @@
+%-----------------------------------------------------------------------------%
+% vim: ft=mercury ts=4 sw=4 et
+%-----------------------------------------------------------------------------%
+% Copyright (C) 2005-2007 The University of Melbourne.
+% This file may only be copied under the terms of the GNU General
+% Public License - see the file COPYING in the Mercury distribution.
+%-----------------------------------------------------------------------------%
+%
+% File rpta_run.m.
+% Main author: Quan Phan.
+%
+% This module implements the region points-to analysis (rpta), which
collects 
+% for each procedure a region points-to graph representing the
splitting of 
+% the heap used by the procedure into regions, i.e., which variables are 
+% stored in which regions. Because the region model is polymorphic,
i.e., we 
+% can pass different actual regions for region arguments, the analysis
also 
+% gathers the alpha mapping, which maps formal region parameters to actual 
+% ones at each call site in a procedure. So there are 2 sorts of
information:
+% region points-to graph (rptg) and alpha mapping.
+%
+% The analysis is composed of 2 phases:
+%	1. intraprocedural analysis: only analyses unifications and compute only
+%   rptgs.
+%	2. interprocedural analysis: only analyses (plain) procedure calls, 
+%   compute both rptgs and alpha mappings.
+% 
+%
+% Currently the analysis ONLY collects the information, do NOT record
it into 
+% the HLDS.
+% 
+
+:- module rpta_run.
+
+:- interface.
+
+:- import_module hlds.
+:- import_module hlds.hlds_module.
+:- import_module rpta_info.
+
+:- import_module io.
+
+:- pred region_points_to_analysis(rpta_info_table::out, module_info::in,
+    module_info::out, io::di, io::uo) is det.
+
+:- implementation.
+
+:- import_module check_hlds.
+:- import_module check_hlds.goal_path. 
+:- import_module hlds.hlds_goal.
+:- import_module hlds.hlds_pred.
+:- import_module hlds.hlds_data.
+:- import_module parse_tree.
+:- import_module parse_tree.mercury_to_mercury.
+:- import_module parse_tree.modules.
+:- import_module parse_tree.prog_data. 
+:- import_module parse_tree.prog_out.
+:- import_module parse_tree.prog_util.
+:- import_module transform_hlds.
+:- import_module transform_hlds.dependency_graph.
+:- import_module libs.
+:- import_module libs.compiler_util.
+:- import_module rpt_graph.
+:- import_module rpta_fixpoint_table. 
+:- import_module smm_data. 
+:- import_module smm_utils.
+
+:- import_module assoc_list.
+:- import_module bool.
+:- import_module int.
+:- import_module list.
+:- import_module map.
+:- import_module set.
+:- import_module std_util. 
+:- import_module string.
+:- import_module term.
+:- import_module varset.
+:- import_module pair.
+:- import_module svmap.
+:- import_module maybe.
+:- import_module solutions.
+
+region_points_to_analysis(InfoTable, !ModuleInfo, !IO) :-
+    rpta_info_table_init = InfoTable0,
+    intra_proc_rpta(!.ModuleInfo, InfoTable0, InfoTable1, !IO),	
+    inter_proc_rpta(!.ModuleInfo, InfoTable1, InfoTable, !IO).
+
+%----------------------------------------------------------------------------%
+%
+% The part for intraprocedural analysis
+%
+
+:- pred intra_proc_rpta(module_info::in, rpta_info_table::in, 
+    rpta_info_table::out, io::di, io::uo) is det.
+
+intra_proc_rpta(ModuleInfo, !InfoTable, !IO) :-
+    module_info_predids(PredIds, ModuleInfo, _),
+    list.foldl2(intra_proc_rpta_pred(ModuleInfo), PredIds, !InfoTable,
!IO).
+
+:- pred intra_proc_rpta_pred(module_info::in, pred_id::in, 
+    rpta_info_table::in, rpta_info_table::out, io::di, io::uo) is det.
+
+intra_proc_rpta_pred(ModuleInfo, PredId, !InfoTable, !IO) :-
+    module_info_pred_info(ModuleInfo, PredId, PredInfo),
+    pred_info_non_imported_procids(PredInfo) = ProcIds,
+    list.foldl2(intra_proc_rpta_proc(ModuleInfo, PredId), ProcIds,
!InfoTable,          !IO).
+
+:- pred intra_proc_rpta_proc(module_info::in, pred_id::in, proc_id::in, 
+    rpta_info_table::in, rpta_info_table::out, io::di, io::uo) is det.
+
+intra_proc_rpta_proc(ModuleInfo, PredId, ProcId, !InfoTable, !IO) :-
+    PredProcId = proc(PredId, ProcId),
+    intra_analyse_pred_proc(ModuleInfo, PredProcId, !InfoTable, !IO).
+
+:- pred intra_analyse_pred_proc(module_info::in, pred_proc_id::in, 
+    rpta_info_table::in, rpta_info_table::out, io::di, io::uo) is det.
+
+intra_analyse_pred_proc(ModuleInfo, PredProcId, !InfoTable, !IO) :-
+    ( if 
+        smm_utils.some_are_special_preds([PredProcId], ModuleInfo)
+      then
+        true
+      else
+        module_info_pred_proc_info(ModuleInfo, PredProcId, PredInfo,
ProcInfo),
+        rpta_info_init(ProcInfo, RptaInfo0),
+
+        proc_info_get_goal(ProcInfo, Goal),
+        intra_analyse_goal(ProcInfo, PredInfo, ModuleInfo, Goal,
RptaInfo0, 
+            RptaInfo, !IO),
+        
+        rpta_info_table_set_rpta_info(PredProcId, RptaInfo, !InfoTable)
+    ).
+
+:- pred intra_analyse_goal(proc_info::in, pred_info::in, module_info::in, 
+    hlds_goal::in, rpta_info::in, rpta_info::out, io::di, io::uo) is det.
+
+intra_analyse_goal(ProcInfo, PredInfo, HLDS, Goal, !RptaInfo, !IO) :- 
+    Goal = hlds_goal(GoalExpr, _GoalInfo), 
+    intra_analyse_goal_expr(GoalExpr, ProcInfo, PredInfo, HLDS, 
+        !RptaInfo, !IO).
+	
+:- pred intra_analyse_goal_expr(hlds_goal_expr::in, proc_info::in, 
+    pred_info::in, module_info::in, rpta_info::in, rpta_info::out, io::di, 
+    io::uo) is det.
+
+intra_analyse_goal_expr(conj(_ConjType, Goals), ProcInfo, PredInfo, HLDS , 
+    !RptaInfo, !IO) :- 
+    list.foldl2(intra_analyse_goal(ProcInfo, PredInfo, HLDS), 
+        Goals, !RptaInfo, !IO). 
+
+    % procedure calls are ignored in the intraprocedural analysis
+    %
+intra_analyse_goal_expr(plain_call(_PredId, _ProcId, _ARGS, _, _, _PName), 
+    _ProcInfo, _PredInfo, _HLDS, !RptaInfo, !IO).
+
+intra_analyse_goal_expr(generic_call(_GenCall,_,_,_), _ProcInfo, _P, 
+    _HLDS , !RptaInfo, !IO) :-
+    unexpected(thisfile, "intra_analyse_goal_expr: generic_call not
handled").
+
+intra_analyse_goal_expr(switch(_Var, _CF, Cases), ProcInfo, PredInfo, 
+    HLDS, !RptaInfo, !IO) :- 
+    list.foldl2(intra_analyse_case(ProcInfo, PredInfo, HLDS), Cases, 
+        !RptaInfo, !IO).
+
+:- pred intra_analyse_case(proc_info::in, pred_info::in, module_info::in, 
+    case::in, rpta_info::in, rpta_info::out, io::di, io::uo) is det.
+
+intra_analyse_case(ProcInfo, PredInfo, HLDS, Case, !RptaInfo, !IO) :-
+    Case = case(_, Goal),
+    intra_analyse_goal(ProcInfo, PredInfo, HLDS, Goal, !RptaInfo, !IO).
+
+    % Most of the processing in intraprocedural analysis happens to
+    % unifications.
+    %
+intra_analyse_goal_expr(unify(_, _, _, Unification, _), ProcInfo,
_PredInfo, 
+    _HLDS, !RptaInfo, !IO) :- 
+    process_unification(Unification, ProcInfo, !RptaInfo, !IO).
+
+intra_analyse_goal_expr(disj(Goals), ProcInfo, PredInfo, HLDS, !RptaInfo, 
+    !IO) :-
+    list.foldl2(intra_analyse_goal(ProcInfo, PredInfo, HLDS),  
+        Goals, !RptaInfo, !IO). 
+
+intra_analyse_goal_expr(negation(Goal), ProcInfo, PredInfo, HLDS, 
+    !RptaInfo, !IO) :- 
+    intra_analyse_goal(ProcInfo, PredInfo, HLDS, Goal, !RptaInfo, !IO). 
+
+    % scope 
+    % XXX: only analyse the goal. May need to take into account the
Reason. 
+    %
+intra_analyse_goal_expr(scope(_Reason, Goal), ProcInfo, PredInfo, HLDS, 
+        !RptaInfo, !IO) :- 
+%    (
+%	    ( Reason = exist_quant(_)
+%        ; Reason = promise_solutions(_, _)      % XXX ???
+%        ; Reason = promise_purity(_, _)
+%        ; Reason = commit(_)                    % XXX ???
+%        ; Reason = barrier(_)
+%        ; Reason = trace_goal(_, _, _, _, _)
+%        ; Reason = from_ground_term(_)
+%        ),
+        intra_analyse_goal(ProcInfo, PredInfo, HLDS, Goal, !RptaInfo, !IO).
+%    ;
+%        Msg = "intra_analyse_goal_expr: Scope's reason of
from_ground_term "
+%            ++ "not handled",
+%        unexpected(thisfile, Msg)
+%    ).
+
+intra_analyse_goal_expr(if_then_else(_Vars, If, Then, Else), ProcInfo, 
+        PredInfo, HLDS, !RptaInfo, !IO) :- 
+    intra_analyse_goal(ProcInfo, PredInfo, HLDS, If, !RptaInfo, !IO),
+    intra_analyse_goal(ProcInfo, PredInfo, HLDS, Then, !RptaInfo, !IO),
+    intra_analyse_goal(ProcInfo, PredInfo, HLDS, Else, !RptaInfo, !IO).
+
+intra_analyse_goal_expr(GoalExpr, _ProcInfo, _PredInfo, _HLDS, !RptaInfo, 
+        !IO) :-
+    GoalExpr = call_foreign_proc(_, _, _, _, _, _, _),
+    unexpected(thisfile, "intra_analyse_goal_expr: call_foreign_proc"
+        ++ " not handled").
+
+intra_analyse_goal_expr(shorthand(_), _, _, _, !RptaInfo, !IO) :- 
+    unexpected(thisfile, "intra_analyse_goal_expr: shorthand not handled").
+    
+:- pred process_unification(unification::in, proc_info::in, rpta_info::in,
+    rpta_info::out, io::di, io::uo) is det.
+
+    % for construction and deconstruction, add edges from LVar to each of
+    % RVars
+    process_unification(construct(LVar, ConsId, RVars, _, _, _, _),
ProcInfo, 
+    !RptaInfo, !IO) :-
+    list.foldl3(process_cons_and_decons(LVar, ConsId, ProcInfo), RVars, 
+        1, _NewComponent, !RptaInfo, !IO).
+	
+process_unification(deconstruct(LVar, ConsId, RVars, _, _, _), ProcInfo, 
+    !RptaInfo, !IO) :-
+    list.foldl3(process_cons_and_decons(LVar, ConsId, ProcInfo), RVars, 
+        1, _NewComponent, !RptaInfo, !IO).
+
+:- pred process_cons_and_decons(prog_var::in, cons_id::in, proc_info::in, 
+    prog_var::in, int::in, int::out, rpta_info::in, rpta_info::out, 
+    io::di, io::uo) is det.
+process_cons_and_decons(LVar, ConsId, ProcInfo, RVar, Component,
NewComponent, 
+    RptaInfo0, RptaInfo, !IO) :-
+    RptaInfo0 = rpta_info(Graph0, AlphaMapping),
+    get_node_by_variable(Graph0, LVar, Node1),
+    get_node_by_variable(Graph0, RVar, Node2),
+    Sel = [termsel(ConsId, Component)],
+    ArcContent = rptg_arc_content(Sel),
+
+    % only add the edge if it is not in the graph
+    % It is more suitable to the edge_operator's semantics if we check 
+    % this inside the edge_operator. But we also want to know if the edge
+    % is actually added or not so it is convenient to check the edge's 
+    % existence outside edge_operator. Otherwise we can extend
edge_operator
+    % with one more argument to indicate that.
+    ( if
+        edge_in_graph(Node1, ArcContent, Node2, Graph0)
+      then
+        RptaInfo = RptaInfo0
+      else
+        rpt_graph.edge_operator(Node1, Node2, ArcContent, Graph0, Graph1),
+        RptaInfo1 = rpta_info(Graph1, AlphaMapping), 
+
+        % after an edge is added, rules P2 and P3 are applied to ensure 
+        % the invariants of the graph.
+        apply_rule_2(Node1, Node2, ConsId, Component, RptaInfo1,
RptaInfo2, 
+            !IO),
+        RptaInfo2 = rpta_info(Graph2, _),	
+        rpt_graph.get_node_by_variable(Graph2, RVar, RVarNode),
+        apply_rule_3(RVarNode, ProcInfo, RptaInfo2, RptaInfo, !IO)
+    ),
+    NewComponent = Component + 1.
+
+    % unification is an assigment: merge the corresponding nodes of ToVar 
+    % and FromVar.
+    % 
+process_unification(assign(ToVar, FromVar), _ProcInfo, RptaInfo0,
RptaInfo, 
+        !IO) :-
+    RptaInfo0 = rpta_info(Graph0, AlphaMapping),
+    get_node_by_variable(Graph0, ToVar, ToNode),
+    get_node_by_variable(Graph0, FromVar, FromNode),
+    ( if
+        ToNode = FromNode
+      then
+        RptaInfo = RptaInfo0
+      else
+        rpt_graph.unify_operator(ToNode, FromNode, Graph0, Graph1),
+        RptaInfo1 = rpta_info(Graph1, AlphaMapping),
+        % after the merge of two nodes, apply rule P1 to ensure rptg's 
+        % invariants
+        apply_rule_1(ToNode, RptaInfo1, RptaInfo, !IO)
+    ).
+
+    % do nothing with the simple test.
+    %
+process_unification(simple_test(_, _), _ProcInfo, !RptaInfo, !IO).
+
+    % XXX: do not consider this for the time being.
+    %
+process_unification(complicated_unify(_, _, _), _ProcInfo, !RptaInfo, !IO).
+
+%-----------------------------------------------------------------------------%
+%
+% The part for interprocedural analysis
+%
+
+    % The interprocedural analysis requires fixpoint computation.
+    %
+:- pred inter_proc_rpta(module_info::in, rpta_info_table::in, 
+    rpta_info_table::out, io::di, io::uo) is det.
+
+inter_proc_rpta(HLDS, !InfoTable, !IO) :-
+    % strongly connected components needed
+    module_info_ensure_dependency_info(HLDS, HLDS1),
+    module_info_get_maybe_dependency_info(HLDS1, MaybeDepInfo),
+    (
+        MaybeDepInfo = yes(DepInfo) 
+    ->
+        hlds_dependency_info_get_dependency_ordering(DepInfo, DepOrdering),
+
+        % perform the analysis and update the rpta_info of the procedures
+        run_with_dependencies(DepOrdering, HLDS1, !InfoTable, !IO)
+    ;
+        unexpected(thisfile, "inter_proc_rpta: no dependency information")
+    ).
+
+:- pred run_with_dependencies(dependency_ordering::in, module_info::in, 
+    rpta_info_table::in, rpta_info_table::out, io::di, io::uo) is det.
+
+run_with_dependencies(Deps, HLDS, !InfoTable, !IO) :- 
+    list.foldl2(run_with_dependency(HLDS), Deps, !InfoTable, !IO).
+
+:- pred run_with_dependency(module_info::in, list(pred_proc_id)::in, 
+    rpta_info_table::in, rpta_info_table::out, io::di, io::uo) is det.
+
+run_with_dependency(HLDS, SCC, !InfoTable, !IO) :- 
+    (
+        % analysis ignores special predicates
+        smm_utils.some_are_special_preds(SCC, HLDS)
+    ->
+        true
+    ;
+        % for each list of strongly connected components, 
+        % perform a fixpoint computation.
+        rpta_info_fixpoint_table_init(SCC, !.InfoTable, FPtable0),
+        run_with_dependency_until_fixpoint(SCC, FPtable0, HLDS, 
+            !InfoTable, !IO)
+    ).
+
+:- pred run_with_dependency_until_fixpoint(list(pred_proc_id)::in, 
+    rpta_info_fixpoint_table::in, module_info::in, rpta_info_table::in, 
+    rpta_info_table::out, io::di, io::uo) is det.
+
+run_with_dependency_until_fixpoint(SCC, FPtable0, HLDS, !InfoTable, !IO) :-
+    list.foldl2(inter_analyse_pred_proc(HLDS, !.InfoTable), SCC, 
+        FPtable0, FPtable1, !IO),
+    (
+        rpta_info_fixpoint_table_all_stable(FPtable1)
+    ->
+        % if all rpta_info's in the FPTable are intact
+        % update the main InfoTable.
+        list.foldl(update_rpta_info_in_rpta_info_table(FPtable1), SCC, 
+            !InfoTable)
+    ;
+        % some is not fixed, start all over again 
+        rpta_info_fixpoint_table_new_run(FPtable1, FPtable2),
+        run_with_dependency_until_fixpoint(SCC, FPtable2, HLDS,
!InfoTable, 
+            !IO)
+    ).
+
+:- pred inter_analyse_pred_proc(module_info::in, rpta_info_table::in, 
+    pred_proc_id::in, rpta_info_fixpoint_table::in, 
+    rpta_info_fixpoint_table::out, io::di, io::uo) is det.
+
+inter_analyse_pred_proc(HLDS, InfoTable, PredProcId, !FPTable, !IO) :- 
+    % look up the caller's rpta_info
+    % it should be there already after the intraprocedural analysis.
+    % We start the interprocedural analysis of a procedure with this 
+    % rpta_info.
+    lookup_rpta_info(PredProcId, InfoTable, !FPTable, RptaInfo0, _),
+	
+    % start the analysis of the procedure's body.
+    %
+    % We will need the information about program point for storing alpha
+    % mapping.
+    module_info_pred_proc_info(HLDS, PredProcId, PredInfo, ProcInfo),
+    fill_goal_path_slots(HLDS, ProcInfo, ProcInfo1),
+
+    % Goal now will contain information of program points
+    proc_info_get_goal(ProcInfo1, Goal),
+
+    % ProcInfo1 and PredInfo here are currently not use inside
+    % inter_analyse_goal. They were used for printing purpose.
+    inter_analyse_goal(ProcInfo1, PredInfo, HLDS, InfoTable, Goal,
!FPTable, 
+        RptaInfo0, RptaInfo, !IO),
+   
+    % put into the fixpoint table	
+    rpta_info_fixpoint_table_new_rpta_info(PredProcId, RptaInfo, !FPTable).
+
+	% analyse a given goal, with module_info and fixpoint table
+	% to lookup extra information, starting from an initial abstract
+	% substitution, and creating a new one. During this process,
+	% the fixpoint table might change (when recursive predicates are
+	% encountered).
+	%
+:- pred inter_analyse_goal(proc_info::in, pred_info::in, module_info::in, 
+    rpta_info_table::in, hlds_goal::in, rpta_info_fixpoint_table::in, 
+    rpta_info_fixpoint_table::out, rpta_info::in, rpta_info::out, io::di, 
+    io::uo) is det.
+
+inter_analyse_goal(ProcInfo, PredInfo, HLDS, InfoTable, Goal, !FPtable, 
+    !RptaInfo, !IO) :- 
+    Goal = hlds_goal(GoalExpr, GoalInfo), 
+    inter_analyse_goal_expr(GoalExpr, GoalInfo, ProcInfo, PredInfo, HLDS, 
+        InfoTable, !FPtable, !RptaInfo, !IO).
+	
+:- pred inter_analyse_goal_expr(hlds_goal_expr::in, hlds_goal_info::in, 
+    proc_info::in, pred_info::in, module_info::in, rpta_info_table::in, 
+    rpta_info_fixpoint_table::in, rpta_info_fixpoint_table::out,
+    rpta_info::in, rpta_info::out, io__state::di, io__state::uo) is det.
+
+inter_analyse_goal_expr(conj(_ConjType, Goals), _Info, ProcInfo, PredInfo, 
+    HLDS, InfoTable, !FPTable, !RptaInfo, !IO) :- 
+    list.foldl3(inter_analyse_goal(ProcInfo, PredInfo, HLDS, InfoTable), 
+        Goals, !FPTable, !RptaInfo, !IO). 
+
+    % There are two rpta_info's: 
+    % one is of the currently-analysed procedure (caller) which we are
going 
+    % to update, the other is of the called procedure (callee).
+    %
+    % The input RptaInfo0 is caller's, if the procedure calls itself then 
+    % this is also that of the callee but we will retrieve it again
from the 
+    % InfoTable.
+    %
+inter_analyse_goal_expr(plain_call(PredId, ProcId, ActualParams, _,_,
_PName), 
+    GoalInfo, ProcInfo, _PredInfo, HLDS, InfoTable, FPTable0, FPTable, 
+    RptaInfo0, RptaInfo, !IO) :- 
+  
+    CalleeId = proc(PredId, ProcId),
+
+    % callee's rpta_info
+    % As what I assume now, after the intraprocedural analysis we have all
+    % the rpta_info's of all the procedures in the InfoTable, therefore
+    % this lookup cannot fail. But it sometimes fails because the callee
+    % can be imported procedures, built-ins and so forth which are not 
+    % analysed by the intraprocedural analysis. In such cases, I assume
that
+    % the rpta_info of the caller is not updated, because no information is
+    % available from the callee.
+    % When IsInit = no, the CalleeRptaInfo is dummy.
+    lookup_rpta_info(CalleeId, InfoTable, FPTable0, FPTable,
CalleeRptaInfo, 
+        IsInit),
+	
+    (
+        IsInit = bool.yes,
+        RptaInfo = RptaInfo0
+    ;
+        IsInit = bool.no,
+        program_point_init(GoalInfo, PP),
+        CalleeRptaInfo = rpta_info(Ge, _Ae),
+
+        % collect alpha mapping at this call site
+        module_info_pred_proc_info(HLDS, CalleeId, _CalleePredInfo, 
+            CalleeProcInfo),
+        proc_info_get_headvars(CalleeProcInfo, FormalParams),
+        RptaInfo0 = rpta_info(Gr0, AlphaMappings0),
+        alpha_mapping_at_call_site(FormalParams, ActualParams, Ge, 
+            Gr0, Gr1, map.init, AlphaMappingAtPP, !IO),
+        map.set(AlphaMappings0, PP, AlphaMappingAtPP, AlphaMappings1),
+        RptaInfo1 = rpta_info(Gr1, AlphaMappings1),
+    
+        % follow the edges from the nodes rooted at the formal parameters
+        % (in the callee's graph) and apply the interprocedural rules to
+        % complete the alpha mapping and update the caller's graph with
+        % the information from the callee's graph
+        map.keys(AlphaMappingAtPP, FormalNodes),
+        apply_rules(FormalNodes, PP, [], CalleeRptaInfo, ProcInfo,
RptaInfo1, 
+            RptaInfo, !IO)
+    ).
+
+inter_analyse_goal_expr(generic_call(_GenCall,_,_,_), _Info, _ProcInfo,
_P, 
+        _HLDS , _InfoTable, !FPTable, !RptaInfo, !IO) :-
+    unexpected(thisfile, "inter_analyse_goal_expr: generic_call not
handled").
+
+inter_analyse_goal_expr(switch(_Var, _CF, Cases), _Info, ProcInfo,
PredInfo, 
+    HLDS, InfoTable, !FPTable, !RptaInfo, !IO) :- 
+    list.foldl3(inter_analyse_case(ProcInfo, PredInfo, HLDS, InfoTable), 
+        Cases, !FPTable, !RptaInfo, !IO).
+
+:- pred inter_analyse_case(proc_info::in, pred_info::in, module_info::in, 
+    rpta_info_table::in, case::in, rpta_info_fixpoint_table::in, 
+    rpta_info_fixpoint_table::out, rpta_info::in, rpta_info::out, io::di, 
+    io::uo) is det.
+inter_analyse_case(ProcInfo, PredInfo, HLDS, InfoTable, Case, !FPtable, 
+    !RptaInfo, !IO) :-
+    Case = case(_, Goal),
+    inter_analyse_goal(ProcInfo, PredInfo, HLDS, InfoTable, Goal,
!FPtable, 
+        !RptaInfo, !IO).
+
+    % unifications are ignored in interprocedural analysis
+    %
+inter_analyse_goal_expr(unify(_, _, _, _Unification, _), _Info, _ProcInfo, 
+    _PredInfo, _HLDS, _InfoTable, !FPTable, !RptaInfo, !IO). 
+
+inter_analyse_goal_expr(disj(Goals), _Info, ProcInfo, PredInfo, HLDS, 
+    InfoTable, !FPTable, !RptaInfo, !IO) :- 
+    list.foldl3(inter_analyse_goal(ProcInfo, PredInfo, HLDS, InfoTable), 
+    	Goals, !FPTable, !RptaInfo, !IO). 
+       
+inter_analyse_goal_expr(negation(Goal), _Info, ProcInfo, PredInfo, HLDS, 
+    InfoTable, !FPTable, !RptaInfo, !IO) :- 
+    inter_analyse_goal(ProcInfo, PredInfo, HLDS, InfoTable, Goal, !FPTable,
+	!RptaInfo, !IO). 
+
+    % XXX: may need to take into account the Reason.
+    % for now just analyse the goal.
+    %
+inter_analyse_goal_expr(scope(_Reason, Goal), _Info, ProcInfo,
PredInfo, HLDS, 
+    InfoTable, !FPTable, !RptaInfo, !IO) :-
+%    (
+%        ( Reason = exist_quant(_)
+%        ; Reason = promise_solutions(_, _)      % XXX ???
+%        ; Reason = promise_purity(_, _)
+%        ; Reason = commit(_)                    % XXX ???
+%        ; Reason = barrier(_)
+%        ; Reason = trace_goal(_, _, _, _, _)
+%        ; Reason = from_ground_term(_)
+%        ),
+        inter_analyse_goal(ProcInfo, PredInfo, HLDS, InfoTable, Goal, 
+            !FPTable, !RptaInfo, !IO).
+%    ;
+%        Msg = "inter_analyse_goal_expr: Scope's reason of
from_ground_term "
+%            ++ "not handled",
+%        unexpected(thisfile, Msg)
+%    ).
+  
+inter_analyse_goal_expr(if_then_else(_Vars, If, Then, Else), _Info, 
+    ProcInfo, PredInfo, HLDS, InfoTable, !FPTable, !RptaInfo, !IO) :- 
+    inter_analyse_goal(ProcInfo, PredInfo, HLDS, InfoTable, If, !FPTable, 
+        !RptaInfo, !IO),
+    inter_analyse_goal(ProcInfo, PredInfo, HLDS, InfoTable, Then,
!FPTable, 
+        !RptaInfo, !IO),
+    inter_analyse_goal(ProcInfo, PredInfo, HLDS, InfoTable, Else,
!FPTable, 
+        !RptaInfo, !IO).
+
+inter_analyse_goal_expr(GoalExpr, _Info, _ProcInfo, _PredInfo, _HLDS, _IT, 
+        !FPTable, !RptaInfo, !IO) :- 
+    GoalExpr = call_foreign_proc(_, _, _, _, _, _, _),
+    unexpected(thisfile, "inter_analyse_goal_expr: foreign code not
handled").
+
+inter_analyse_goal_expr(shorthand(_), _, _,  _, _, _IT, !FPTable,
!RptaInfo, 
+        !IO) :- 
+    unexpected(thisfile, 
+        "inter_analyse_goal_expr: shorthand goal not handled").
+
+    % As said above, the rpta_info of a procedure when it is looked 
+    % up in interprocedural analysis is either in the InfoTable or in the 
+    % fixpoint table. If the procedure happens to be imported ones,
built-ins,
+    % and so on, we returns no and initialize the lookup value to a dummy 
+    % value. 
+    %
+:- pred lookup_rpta_info(pred_proc_id::in, rpta_info_table::in, 
+    rpta_info_fixpoint_table::in, rpta_info_fixpoint_table::out,
+    rpta_info::out, bool::out) is det.
+
+lookup_rpta_info(PredProcId, InfoTable, FPtable0, FPtable, RptaInfo,
Init) :- 
+    ( if
+        % 1 - look up in the current fixpoint table
+        rpta_info_fixpoint_table_get_rpta_info(PredProcId, RptaInfo1, 
+            FPtable0, FPtable1)
+      then
+        RptaInfo  = RptaInfo1,
+        FPtable = FPtable1,
+        Init = bool.no 
+      else
+	    % 2 - look up among already recorded rpta_info.
+        ( if 
+            RptaInfo1 = rpta_info_table_search_rpta_info(PredProcId, 
+                InfoTable)
+          then
+            RptaInfo = RptaInfo1,
+            FPtable = FPtable0,
+            Init = bool.no
+          else
+            % initialize a dummy
+            RptaInfo = rpta_info(rpt_graph_init, map.init),
+            Init = bool.yes,
+            FPtable = FPtable0
+        )
+    ).
+
+:- pred update_rpta_info_in_rpta_info_table(rpta_info_fixpoint_table::in, 
+    pred_proc_id::in, rpta_info_table::in, rpta_info_table::out) is det.
+
+update_rpta_info_in_rpta_info_table(FPTable, PredProcId, InfoTable0, 
+    InfoTable) :-
+    rpta_info_fixpoint_table_get_final_rpta_info(PredProcId, RptaInfo, 
+        FPTable), 
+    rpta_info_table_set_rpta_info(PredProcId, RptaInfo, InfoTable0,
InfoTable). 
+:- pred apply_rule_1(rptg_node::in, rpta_info::in, rpta_info::out, io::di, 
+    io::uo) is det.
+
+apply_rule_1(Node, !RptaInfo, !IO) :-
+    !.RptaInfo = rpta_info(Graph0, Alpha),
+    rptg_node_contents(Graph0, Node, Content),
+    rule_1(Content^varset, Graph0, Graph1, !IO),
+    !:RptaInfo = rpta_info(Graph1, Alpha).
+
+    %------------ Rule 1
----------------------------------------------------%
+    % When a node is unified and rule 1 is applied recursively the node 
+    % itself may be removed from the graph so we need to trace "it" by the 
+    % variables assigned to it.	Node is one of the nodes that have just
been 
+    % unified. The algorithm is as follows.  
+    % 1. If the node has no or one out-arc we have to do nothing and the 
+    % predicate quits. 
+    % 2. The node has > 1 out-arc, take one of them, find in the rest 
+    % another arc that has a same label, unify the end nodes of the two
arcs. 
+    % Because of this unification of the end nodes, more unifications are 
+    % probably triggered.
+    % 3. Start all over again with the same node and the *updated* graph. 
+    %
+:- pred rule_1(set(prog_var)::in, rpt_graph::in, rpt_graph::out, 
+    io::di, io::uo) is det.
+
+rule_1(VarSet, !Graph, !IO) :-
+    get_node_by_varset(!.Graph, VarSet, Node),
+    rptg_get_edgemap(!.Graph, EdgeMap),
+    map.lookup(EdgeMap, Node, OutEdgesMap),
+    map.keys(OutEdgesMap, OutArcs),
+    ( 
+        OutArcs = [A|As],
+        merge_nodes_reached_by_same_labelled_arcs(A, As, As, !Graph,
Happened, 
+            !IO),
+        (
+            Happened = bool.no
+        ;
+            % some nodes have been merged, so size of Graph1 is strictly 
+            % smaller than that of Graph0 and at some point this predicate 
+            % will end up in the then-branch
+            Happened = bool.yes,
+            rule_1(VarSet, !Graph, !IO)
+        )
+    ;
+        OutArcs = []
+    ).
+	
+    % This predicate finds in the list the very first arc, which has
the same 
+    % label as the input arc, and unify the end nodes of them. When one
such 
+    % an arc found, the predicate will not look further in the list.
+    % The unification of nodes, if happends, will be propagated by calling 
+    % rule_1 predicate mutually recursively. 
+    %
+:- pred merge_nodes_reached_by_same_labelled_arcs(rptg_arc::in,
+    list(rptg_arc)::in, list(rptg_arc)::in, rpt_graph::in, rpt_graph::out, 
+    bool::out, io::di, io::uo) is det.
+
+    % The loop in this predicate is similar to
+    % for i = ... to N - 1
+    %    for j = i+1 to N ...
+    % ...	
+    % this clause is reached at the end of the inner loop. No unification 
+    % has happened so far therefore the list of arcs (Rest = [A | As]) 
+    % are still safe to use.
+    %
+    % reach this clause means that no unification of nodes happened and 
+    % all the out-arcs have been processed (Rest = []).
+    %
+merge_nodes_reached_by_same_labelled_arcs(_Arc, [], [], !Graph,
bool.no, !IO). 
+
+    % some out-arcs still need to be processed
+    %
+merge_nodes_reached_by_same_labelled_arcs(_Arc, [], [A|As], !Graph,
Happened, 
+    !IO) :-
+    merge_nodes_reached_by_same_labelled_arcs(A, As, As, !Graph, Happened, 
+        !IO).
+
+merge_nodes_reached_by_same_labelled_arcs(Arc, [A|As], Rest, Graph0,
Graph, 
+    Happened, !IO) :-
+    % for a node, we do not allow two arcs with the same label to another 
+    % node. So End and E below must be definitely different nodes and we 
+    % only need to compare labels.   
+    rptg_arc_contents(Graph0, Arc, _Start, End, ArcContent),
+    rptg_arc_contents(Graph0, A, _S, E, AC),
+    ( if
+         ArcContent = AC
+      then
+         % unify the two end nodes
+         rpt_graph.unify_operator(End, E, Graph0, Graph1),
+
+         % apply rule 1 after the above unification 
+         rptg_node_contents(Graph1, End, Content),
+         rule_1(Content^varset, Graph1, Graph, !IO),
+         Happened = bool.yes 
+      else
+         % still not found an arc with the same label, continue the inner 
+         % loop
+         merge_nodes_reached_by_same_labelled_arcs(Arc, As, Rest, Graph0, 
+            Graph, Happened, !IO)
+    ).
+
+%------------------------- end of rule 1
-------------------------------------%
+
+:- pred apply_rule_2(rptg_node::in, rptg_node::in, cons_id::in, int::in,
+    rpta_info::in, rpta_info::out, io::di, io::uo) is det.
+
+apply_rule_2(Start, End, ConsId, Component, !RptaInfo, !IO) :-
+    !.RptaInfo = rpta_info(Graph0, AlphaMapping),
+    rptg_node_contents(Graph0, Start, SContent),
+    rptg_node_contents(Graph0, End, EContent),
+    rule_2(SContent^varset, EContent^varset, ConsId, Component, 
+	Graph0, Graph, !IO),
+    !:RptaInfo = rpta_info(Graph, AlphaMapping).
+
+    %---------------------------- Rule 2
-------------------------------------%
+    % Node is the Start node of the newly-added edge.
+    % Because this predicate will be called any time a new edge is
added so 
+    % there is at most one existing edge with the same label to a
different 
+    % node. Because of that the rule_2 predicate need not be recursive.
+    %
+:- pred rule_2(set(prog_var)::in, set(prog_var)::in, cons_id::in, int::in, 
+    rpt_graph::in, rpt_graph::out, io::di, io::uo) is det.
+
+rule_2(SVarSet, EVarSet, ConsId, Component, !Graph, !IO) :-
+    get_node_by_varset(!.Graph, SVarSet, N),
+    get_node_by_varset(!.Graph, EVarSet, M),
+    Sel = [termsel(ConsId, Component)], 
+    rptg_get_edgemap(!.Graph, EdgeMap),
+    map.lookup(EdgeMap, N, OutEdgesMap),
+    map.keys(OutEdgesMap, OutArcs),
+    merge_nodes_reached_by_same_labelled_arc(Sel, M, OutArcs, !Graph, !IO).
+
+    % If an A(rc) in OutArcs (from N) has the same label Sel then merge M
+    % with the node (MPrime) that the A(rc) points to.
+    %
+:- pred merge_nodes_reached_by_same_labelled_arc(selector::in,
rptg_node::in, 
+    list(rptg_arc)::in, rpt_graph::in, rpt_graph::out, io::di, io::uo)
is det.
+
+merge_nodes_reached_by_same_labelled_arc(_, _, [], !Graph, !IO).
+merge_nodes_reached_by_same_labelled_arc(Sel, M, [A|As], Graph0, Graph,
!IO) :-
+    rptg_arc_contents(Graph0, A, _, MPrime, ArcContent),
+    ( if
+        ArcContent = rptg_arc_content(Selector),
+        Selector = Sel,
+        MPrime \= M
+      then
+        rpt_graph.unify_operator(M, MPrime, Graph0, Graph1),
+        rptg_node_contents(Graph1, M, Content),
+        rule_1(Content^varset, Graph1, Graph, !IO)
+      else
+        % still not found an arc with the same label, continue the loop
+        merge_nodes_reached_by_same_labelled_arc(Sel, M, As, Graph0,
Graph, 
+            !IO)
+    ).
+%---------------------------- end of rule 2
----------------------------------%
+
+    % This predicate is just to wrap the call to rule_3 so that the
changed 
+    % graph is put into rpta_info structure.
+    %
+:- pred apply_rule_3(rptg_node::in, proc_info::in, rpta_info::in, 
+    rpta_info::out, io::di, io::uo) is det.
+
+apply_rule_3(Node, _ProcInfo, !RptaInfo, !IO) :-
+	!.RptaInfo = rpta_info(Graph0, AlphaMapping),
+	rule_3(Node, Graph0, Graph1, !IO),
+	!:RptaInfo = rpta_info(Graph1, AlphaMapping).
+
+    %------------------------------- Rule 3
----------------------------------%
+    % This rule is applied after an edge is added TO the Node to
enforce the 
+    % invariant that a subterm of the same type as the compounding term is 
+    % stored in the same region as the compounding term.
+    % 
+    % This algorithm may not be an efficient one because it checks all the 
+    % nodes in the graph one by one to see if a node can reach the node
or not.
+    %
+    % We enforce the invariant (in the sense that whenever the
invariant is 
+    % made invalid this rule will correct it) therefore whenever we find a 
+    % satisfied node and unify it with Node we can stop. This is
indicated by  
+    % Happened.
+    % 
+:- pred rule_3(rptg_node::in, rpt_graph::in, rpt_graph::out, io::di,
io::uo) 
+    is det.
+
+rule_3(Node, !Graph, !IO) :-
+    rptg_get_nodemap(!.Graph, NodeMap),
+    map.keys(NodeMap, Nodes),
+    (  
+        Nodes = [_N|_NS],
+        % the graph has some node(s), so check each node to see if it 
+        % satisfies the condition of rule 3 or not, if yes unify it with NY
+        get_node_by_node(!.Graph, Node, NY),
+        rule_3_2(Nodes, NY, !Graph, Happened, !IO),
+
+        % This predicate will quit when Happened = no, i.e. no more nodes 
+        % need to be unified.
+        ( 
+            Happened = bool.yes, 
+            % a node in Nodes has been unified with NY, so we start all 
+            % over again. Note that the node that has been unified has 
+            % been removed, so it will not be in the Graph1 in the below 
+            % call. So this predicate can terminate at some point (due
to the
+            % fact that the "size" of Graph1 is smaller than that of
Graph0).
+            rule_3(Node, !Graph, !IO)
+          ;
+            % no node in Nodes has been unified with NY, which means that 
+            % no more nodes need to be unified, so just quit.
+            Happened = bool.no
+	)
+    ; 
+        Nodes = [],
+        % no node in the graph, impossible
+        unexpected(thisfile, "rule_3: impossible having no node in graph")
+    ).
+
+    % check each node in the list to see if it satisfies the condition of 
+    % rule 3 or not. 
+    %	1. If the predicate finds out such a node, it unifies it with NY
(also 
+    %	apply rule 1 here) and quit with Happend = 1.
+    %	2. if no such a node found, it processes the rest of the list. The 
+    %	process continues like that until either 1. happens (the case above) 
+    %	or the list becomes empty and the predicate quits with Happened = 0.
+    %
+:- pred rule_3_2(list(rptg_node)::in, rptg_node::in, rpt_graph::in, 
+    rpt_graph::out, bool::out, io::di, io::uo) is det.
+
+rule_3_2([], _NY, !Graph, bool.no, !IO).
+rule_3_2([NZ|NZs], NY, Graph0, Graph, Happened, !IO) :-
+    ( if
+        rule_3_condition(NZ, NY, Graph0, NZ1)
+      then
+        rpt_graph.unify_operator(NZ, NZ1, Graph0, Graph1),
+        
+        % apply rule 1
+        rptg_node_contents(Graph1, NZ, Content),
+        rule_1(Content^varset, Graph1, Graph, !IO),
+        Happened = bool.yes
+      else
+        % try with the rest, namely NS
+        rule_3_2(NZs, NY, Graph0, Graph, Happened, !IO)
+    ).
+
+:- pred rule_3_condition(rptg_node::in, rptg_node::in, rpt_graph::in, 
+    rptg_node::out) is semidet.
+
+rule_3_condition(NZ, NY, Graph, NZ1) :-
+    rptg_path(Graph, NZ, NY, _),
+    rptg_lookup_node_type(Graph, NZ) = NZType,
+    % a node reachable from NY, with the same type as NZ
+    % The node can be exactly NY
+    reachable_and_having_type(Graph, NY, NZType, NZ1),
+    NZ \= NZ1.
+	
+%----------------------- end of rule 3
---------------------------------------%
+
+%-----------------------------------------------------------------------------%
+%
+% code for collect alpha mapping and application of rule P4
+%
+
+	% build up the alpha mapping (node -> node) and apply rule P4
+    % to ensure that it is actually a function.
+	%
+:- pred alpha_mapping_at_call_site(list(prog_var)::in, list(prog_var)::in, 
+    rpt_graph::in, rpt_graph::in, rpt_graph::out, 
+    map(rptg_node, rptg_node)::in, map(rptg_node, rptg_node)::out, 
+    io::di, io::uo) is det.
+
+alpha_mapping_at_call_site([], [], _Ge, !Gr, !AlphaMap, !IO). 
+alpha_mapping_at_call_site([], [_Yi|_Ys], _Ge, !Gr, !AlphaMap, 
+    !IO) :-
+    unexpected(thisfile, 
+        "alpha_mapping_at_call_site: actuals and formals not match").
+alpha_mapping_at_call_site([_Xi|_Xs], [], _Ge, !Gr, !AlphaMap, 
+    !IO) :-
+    unexpected(thisfile, 
+        "alpha_mapping_at_call_site: actuals and formals not match").
+alpha_mapping_at_call_site([Xi|Xs], [Yi|Ys], Ge, Gr0, Gr, !AlphaMap,
!IO) :-
+    get_node_by_variable(Ge, Xi, N_Xi),
+    get_node_by_variable(Gr0, Yi, N_Yi),
+    ( if
+        map.search(!.AlphaMap, N_Xi, N_Y)
+      then
+            % alpha(N_Xi) = N_Y, alpha(N_Xi) = N_Yi, N_Y != N_Yi
+            %
+        ( if
+            N_Y \= N_Yi
+          then
+            % apply rule P4
+            unify_operator(N_Y, N_Yi, Gr0, Gr1),
+
+            % apply rule P1 after some nodes are unified
+            rptg_node_contents(Gr1, N_Y, Content),
+            rule_1(Content^varset, Gr1, Gr2, !IO)
+          else
+            Gr2 = Gr0
+        )
+      else
+        svmap.set(N_Xi, N_Yi, !AlphaMap),
+        Gr2 = Gr0
+    ),
+    alpha_mapping_at_call_site(Xs, Ys, Ge, Gr2, Gr, !AlphaMap, !IO).
+
+%-----------------------------------------------------------------------------%
+
+%-----------------------------------------------------------------------------%
+%
+% The part for rules P5-P8 and their application
+%
+% The application happens at a call site, so related to a caller and a
callee.
+% The idea of those rules is to complete the alpha mapping and update the 
+% rptg of the caller based on the rptg of the callee. The information
of the 
+% caller will end with "R", of the callee with "E".
+% 
+% We will follow each out-coming edge in the rptg_E exactly once and
try to 
+% apply the rules.
+%
+
+:- pred apply_rules(list(rptg_node)::in, program_point::in, 
+    list(rptg_node)::in, rpta_info::in, proc_info::in, rpta_info::in, 
+    rpta_info::out, io::di, io::uo) is det.  
+
+apply_rules([], _PP, _Processed, _RptaInfoE, _ProcInfo, !RptaInfoR, !IO).
+apply_rules([Ne|Others], PP, Processed, RptaInfoE, ProcInfo, !RptaInfoR, 
+    !IO) :-
+    RptaInfoE = rpta_info(Ge, _Ae),
+    
+    % find alpha mappings relate to this node at this call site.
+    !.RptaInfoR = rpta_info(_Gr0, Ar0),
+    map.lookup(Ar0, PP, AlphaAtPP),
+    map.lookup(AlphaAtPP, Ne, Nr),
+    
+    % prepare a to-be-processed list of nodes
+    rptg_successors(Ge, Ne, Ss),
+    set.to_sorted_list(Ss, SsList),
+    get_unprocessed_nodes(SsList, Ge, Processed, [], ToBeProcessed),
+    list.append(ToBeProcessed, Others, NewOthers),
+
+    % follow Ne and apply rules for its edges
+    apply_rules_node(PP, Ne, RptaInfoE, ProcInfo, Nr, !RptaInfoR, !IO),
+
+    % process the others
+    apply_rules(NewOthers, PP, [Ne|Processed], RptaInfoE, ProcInfo, 
+        !RptaInfoR, !IO).
+
+    % return a list of non-processed prog_var's of the input nodes. 
+    %
+:- pred get_unprocessed_nodes(list(rptg_node)::in, rpt_graph::in, 
+    list(rptg_node)::in, list(rptg_node)::in, list(rptg_node)::out) is det.
+get_unprocessed_nodes([], _Graph, _Processed, Unprocessed, Unprocessed).
+get_unprocessed_nodes([N|Ns], Graph, Processed, Unprocessed0,
Unprocessed) :-
+    ( if
+        in_list(Processed, N, Graph)
+      then
+        get_unprocessed_nodes(Ns, Graph, Processed, Unprocessed0,
Unprocessed)
+      else
+        get_unprocessed_nodes(Ns, Graph, Processed, [N|Unprocessed0], 
+            Unprocessed)
+    ).
+
+:- pred in_list(list(rptg_node)::in, rptg_node::in, rpt_graph::in) is
semidet.
+in_list([N|Ns], Node, Graph) :- 
+    get_node_by_node(Graph, N, Real1),
+    get_node_by_node(Graph, Node, Real2),
+    ( if
+        Real1 = Real2
+      then
+        true
+      else
+        in_list(Ns, Node, Graph)
+    ).
+
+:- pred apply_rules_node(program_point::in, rptg_node::in, rpta_info::in, 
+    proc_info::in, rptg_node::in, rpta_info::in, rpta_info::out, 
+    io::di, io::uo) is det.
+
+apply_rules_node(PP, Ne, RptaInfoE, ProcInfo, Nr, !RptaInfoR, !IO) :-
+    RptaInfoE = rpta_info(Ge, _Ae),
+    get_node_by_node(Ge, Ne, RealNe),
+
+    % apply rules P5-P8 for each out-edge of Ne 
+    rptg_get_edgemap(Ge, EdgeMap),
+    map.lookup(EdgeMap, RealNe, NeOutEdges),
+    map.keys(NeOutEdges, NeOutArcs),
+    apply_rules_arcs(NeOutArcs, Nr, PP, RptaInfoE, ProcInfo,
!RptaInfoR, !IO).
+
+:- pred apply_rules_arcs(list(rptg_arc)::in, rptg_node::in, 
+    program_point::in, rpta_info::in, proc_info::in, rpta_info::in, 
+    rpta_info::out, io::di, io::uo) is det.
+
+apply_rules_arcs([], _Nr, _PP, _RptaInfoE, _ProcInfo, !RptaInfoR, !IO).
+apply_rules_arcs([Arc | Arcs], Nr, PP, RptaInfoE, ProcInfo, !RptaInfoR, 
+    !IO) :-
+	rule_5(Arc, PP, RptaInfoE, ProcInfo, Nr, !RptaInfoR, !IO),
+	rule_6(Arc, PP, RptaInfoE, ProcInfo, Nr, !RptaInfoR, !IO),
+	rule_7(Arc, PP, RptaInfoE, ProcInfo, Nr, !RptaInfoR, !IO),
+	rule_8(Arc, PP, RptaInfoE, ProcInfo, Nr, !RptaInfoR, !IO),
+	apply_rules_arcs(Arcs, Nr, PP, RptaInfoE, ProcInfo, !RptaInfoR, !IO).
+
+    % Arc is an out-arc of Ne
+    % this predicate applies rule P5 to the Arc and an alpha-mapping
relating 
+    % to Nr
+    %
+:- pred rule_5(rptg_arc::in, program_point::in, rpta_info::in,
proc_info::in, 
+    rptg_node::in, rpta_info::in, rpta_info::out, io::di, io::uo) is det.
+rule_5(Arc, PP, RptaInfoE, _ProcInfo, Nr, !RptaInfoR, !IO) :-
+    % find an out-arc in the caller's graph that has a same label 
+    % the label of the out-arc in callee's graph
+    RptaInfoE = rpta_info(Ge, _Ae),
+    rptg_arc_contents(Ge, Arc, _Ne, Me, ContentE),
+
+    % out-arcs of Nr in the caller's graph
+    !.RptaInfoR = rpta_info(Gr0, Ar0),
+    rptg_get_edgemap(Gr0, EdgeMap),
+    get_node_by_node(Gr0, Nr, RealNr),
+    map.lookup(EdgeMap, RealNr, OutEdges),
+    map.keys(OutEdges, OutArcs),
+
+    % when the premises of rule P5 are satisfied, nodes are unified and
+    % rule P1 applied to ensure invariants
+    ( if
+        % (Nr, sel, Mr') exists.
+        same_label(ContentE, OutArcs, Gr0, MrPrime), 
+        map.search(Ar0, PP, AlphaAtPP),
+        map.search(AlphaAtPP, Me, Mr),
+        get_node_by_node(Gr0, Mr, RealMr),
+        MrPrime \= RealMr
+      then
+        rpt_graph.unify_operator(RealMr, MrPrime, Gr0, Gr1),
+        !:RptaInfoR = rpta_info(Gr1, Ar0),
+        apply_rule_1(RealMr, !RptaInfoR, !IO)
+      else
+        true
+    ).
+
+	% Arc is an out-arc of Ne
+	% this predicate applies rule P6 to the Arc and an alpha-mapping 
+	% relating to Ve and Nr
+	%
+:- pred rule_6(rptg_arc::in, program_point::in, rpta_info::in,
proc_info::in,
+    rptg_node::in, rpta_info::in, rpta_info::out, io::di, io::uo) is det.
+rule_6(Arc, PP, RptaInfoE, _ProcInfo, Nr, !RptaInfoR, !IO) :-
+
+    % find an out-arc in the caller's graph that has a same label 
+    % the label of the out-arc in callee's graph
+    RptaInfoE = rpta_info(Ge, _Ae),
+    rptg_arc_contents(Ge, Arc, _Ne, Me, ContentE),
+
+    % out-arcs of Nr in the caller's graph
+    !.RptaInfoR = rpta_info(Gr0, Ar0),
+    rptg_get_edgemap(Gr0, EdgeMap),
+    get_node_by_node(Gr0, Nr, RealNr),
+    map.lookup(EdgeMap, RealNr, OutEdges),
+    map.keys(OutEdges, OutArcs),
+
+    % apply rule P6 when its premises are satisfied
+    map.lookup(Ar0, PP, AlphaAtPP0),
+    ( if
+        same_label(ContentE, OutArcs, Gr0, Mr)
+      then
+        % (Nr, sel, Mr) in the graph
+        ( if
+            map.search(AlphaAtPP0, Me, _Mr)
+          then
+            % alpha(Me) = Mr so ignore
+            true
+          else
+            % record alpha(Me) = Mr
+            map.set(AlphaAtPP0, Me, Mr, AlphaAtPP1),
+            map.set(Ar0, PP, AlphaAtPP1, Ar1),
+            !:RptaInfoR = rpta_info(Gr0, Ar1)
+        )
+      else
+        true
+    ).
+
+    % Arc is an out-arc of Ne
+    % this predicate applies rule P7 to the Arc and an alpha-mapping
relating 
+    % to Ve and Nr
+    %
+:- pred rule_7(rptg_arc::in, program_point::in, rpta_info::in,
proc_info::in,
+    rptg_node::in, rpta_info::in, rpta_info::out, io::di, io::uo) is det.
+rule_7(Arc, PP, RptaInfoE, _ProcInfo, Nr, !RptaInfoR, !IO) :-
+    % find an out-arc in the caller's graph that has a same label 
+    % the label of the out-arc in callee's graph
+    RptaInfoE = rpta_info(Ge, _Ae),
+    rptg_arc_contents(Ge, Arc, _Ne, Me, ContentE),
+	
+    % out-arcs of Nr in the caller's graph
+    !.RptaInfoR = rpta_info(Gr0, Ar0),
+    rptg_get_edgemap(Gr0, EdgeMap),
+    get_node_by_node(Gr0, Nr, RealNr),
+    map.lookup(EdgeMap, RealNr, OutEdges),
+    map.keys(OutEdges, OutArcs),
+    
+    ( if
+        same_label(ContentE, OutArcs, Gr0, _Mr)
+      then
+        true
+      else
+        ( if
+            map.lookup(Ar0, PP, AlphaAtPP),
+            map.search(AlphaAtPP, Me, Mr)
+          then
+            % reach here means all the premises of rule P7 are satisfied, 
+            % add (Nr, sel, Mr)
+            get_node_by_node(Gr0, Mr, RealMr),
+            % This is just a check point
+            ( if 
+                edge_in_graph(RealNr, ContentE, RealMr, Gr0)
+              then
+                % should be an error if reach here because if this edge
+                % is in the graph, same_label should find it.
+                unexpected(thisfile, 
+                    "rule_7: same_label does not work correctly")
+              else
+                rpt_graph.edge_operator(RealNr, RealMr, ContentE, Gr0,
Gr1),
+            
+                % need to apply rule 3 here
+                rule_3(RealMr, Gr1, Gr2, !IO),
+                !:RptaInfoR = rpta_info(Gr2, Ar0)
+            )	
+          else
+            true
+        )
+    ).
+
+    % Arc is an out-arc of Ne
+    % this predicate applies rule P8 to the Arc and an alpha-mapping
relating 
+    % to Ve and Nr
+    %
+:- pred rule_8(rptg_arc::in, program_point::in, rpta_info::in,
proc_info::in,
+    rptg_node::in, rpta_info::in, rpta_info::out, io::di, io::uo) is det.
+rule_8(Arc, PP, RptaInfoE, _ProcInfo, Nr, !RptaInfoR, !IO) :-
+        % find an out-arc in the caller's graph that has a same label 
+        % the label of the out-arc in callee's graph
+        %
+    RptaInfoE = rpta_info(Ge, _Ae),
+    rptg_arc_contents(Ge, Arc, _Ne, Me, ContentE),
+
+        % out-arcs of Nr in the caller's graph
+        %
+    !.RptaInfoR = rpta_info(Gr0, Ar0),
+    rptg_get_edgemap(Gr0, EdgeMap),
+    get_node_by_node(Gr0, Nr, RealNr),
+    map.lookup(EdgeMap, RealNr, OutEdges),
+    map.keys(OutEdges, OutArcs),
+    ( if
+        same_label(ContentE, OutArcs, Gr0, _Mr)
+      then
+        true
+      else
+        ( if 
+            map.lookup(Ar0, PP, AlphaAtPP),
+            map.search(AlphaAtPP, Me, _Mr)
+          then
+            true
+          else
+                % rule 7: add node Mr, alpha(Me) = Mr, edge(Nr, sel, Mr)
+                %
+            rptg_get_node_supply(Gr0, NS0),
+            string.append("R", string.int_to_string(NS0 + 1), RegName),
+            MrContent = rptg_node_content(set.init, RegName, set.init, 
+                rptg_lookup_node_type(Ge, Me)),
+            rptg_set_node(MrContent, Mr, Gr0, Gr1),
+            edge_operator(RealNr, Mr, ContentE, Gr1, Gr2),
+            
+            map.lookup(Ar0, PP, AlphaAtPP0),
+            map.set(AlphaAtPP0, Me, Mr, AlphaAtPP1),
+            map.set(Ar0, PP, AlphaAtPP1, Ar1),
+                
+            rule_3(Mr, Gr2, Gr3, !IO),
+            !:RptaInfoR = rpta_info(Gr3, Ar1)
+        )
+    ).
+
+    % This predicate checks if an arc in the list has the label
ContentE, if 
+    % yes, it also returns the node to which the arc points.
+    %
+:- pred same_label(rptg_arc_content::in, list(rptg_arc)::in,
rpt_graph::in, 
+    rptg_node::out) is semidet.
+same_label(ArcContent, [Arc|Arcs], G, M) :-
+    rptg_arc_contents(G, Arc, _N, M0, ArcContent0),
+    ( if
+        ArcContent0 = ArcContent
+      then
+        M = M0
+      else
+        same_label(ArcContent, Arcs, G, M)
+    ).
+	    
+    % Check if an edge (Start, Label, End) is already in the Graph or not
+    %
+:- pred edge_in_graph(rptg_node::in, rptg_arc_content::in, rptg_node::in, 
+    rpt_graph::in) is semidet.
+edge_in_graph(Start, Label, End, Graph) :-
+    rptg_get_edgemap(Graph, EdgeMap),
+    map.lookup(EdgeMap, Start, OutEdges),
+    solutions(map.inverse_search(OutEdges, End), ArcList),
+    same_label(Label, ArcList, Graph, _).
+
+:- func thisfile = string.
+thisfile = "rpta_run.m".
Index: smm_data.m
===================================================================
RCS file: smm_data.m
diff -N smm_data.m
--- /dev/null	1 Jan 1970 00:00:00 -0000
+++ smm_data.m	14 May 2007 02:37:01 -0000
@@ -0,0 +1,60 @@
+%-----------------------------------------------------------------------------%
+% vim: ft=mercury ts=4 sw=4 et
+%-----------------------------------------------------------------------------%
+% Copyright (C) 2000-2002,2006 The University of Melbourne.
+% This file may only be copied under the terms of the GNU General
+% Public License - see the file COPYING in the Mercury distribution.
+%-----------------------------------------------------------------------------%
+%
+% file smm_data.m.
+% Main author: Quan Phan.
+% 
+% The program_point type defined here is exactly the same as in 
+% structure_reuse.direct.m, which is used for CTGC.
+% That is why I name this module smm_data, I tentatively assume it will 
+% defines several data structures that are used in Static Memory
Management 
+% analyses, e.g., CTGC, RBMM. 
+
+:- module smm_data.
+
+:- interface.
+
+:- import_module hlds.
+:- import_module hlds.hlds_goal.
+
+:- import_module mdbcomp.
+:- import_module mdbcomp.program_representation.
+
+:- import_module term, io.
+
+%-----------------------------------------------------------------------------%
+:- type program_point 
+    ---> 	pp( 
+                pp_context	:: term.context, 
+                pp_path		:: goal_path
+            ).
+
+:- pred program_point_init(hlds_goal_info, program_point).
+:- mode program_point_init(in, out) is det.
+
+:- pred print_program_point(program_point, io__state, io__state).
+:- mode print_program_point(in, di, uo) is det.
+
+%-----------------------------------------------------------------------------%
+
+:- implementation.
+
+:- import_module list, string.
+	
+	% Note: for a meaningful use of this predicate the goal needs to be
filled 
+	% with path information, i.e. call to fill_goal_path_slots(...).
+    %
+program_point_init(GoalInfo, PP) :-
+	goal_info_get_context(GoalInfo, Context),
+	goal_info_get_goal_path(GoalInfo, GoalPath),
+	PP = pp(Context, GoalPath).
+	
+print_program_point(PP, !IO) :-
+	PP^pp_context = Context,
+	Context = term.context(FileName, LineNum),
+	io.format("In file %s at line %d: ", [s(FileName), i(LineNum)], !IO).
Index: smm_utils.m
===================================================================
RCS file: smm_utils.m
diff -N smm_utils.m
--- /dev/null	1 Jan 1970 00:00:00 -0000
+++ smm_utils.m	15 May 2007 05:35:08 -0000
@@ -0,0 +1,140 @@
+%-----------------------------------------------------------------------------%%
vim: ft=mercury ts=4 sw=4 et
+%-----------------------------------------------------------------------------%
+% Copyright (C) 2000-2002,2006 The University of Melbourne.
+% This file may only be copied under the terms of the GNU General
+% Public License - see the file COPYING in the Mercury distribution.
+%-----------------------------------------------------------------------------%
+
+% file smm_utils.m.
+% main author: Quan Phan.
+%
+% This module implements some common utilities for static memory
management 
+% analyses, e.g. CTGC, RBMM
+
+:- module smm_utils.
+
+:- interface.
+
+:- import_module hlds.
+:- import_module hlds.hlds_pred.
+:- import_module hlds.hlds_module.
+:- import_module parse_tree.
+:- import_module parse_tree.prog_data. 
+
+:- import_module list, io.
+
+
+:- pred check_type_of_node(module_info::in, mer_type::in, selector::in) 
+        is semidet.
+
+:- pred some_are_special_preds(list(pred_proc_id)::in, module_info::in) 
+        is semidet.
+
+	% to annotate each goal with pre-post birth-death sets.
+	% this proc may no longer need for live variable analysis because
+	% we do not rely on those sets in that analysis.
+	%
+:- pred annotate_all_liveness_in_module(module_info::in, module_info::out, 
+        io::di, io::uo) is det.
+
+:- implementation.
+
+:- import_module ll_backend.
+:- import_module ll_backend.liveness.
+:- import_module check_hlds.
+:- import_module check_hlds.type_util. 
+
+:- import_module map.
+:- import_module bool.
+:- import_module int.
+
+	% check if the selector is valid w.r.t the type.
+    %
+check_type_of_node(HLDS, StartType, Selector) :-
+	(
+		Selector = [US|RestSelector],
+		(
+			US = termsel(CONS_ID, CHOICE),
+			select_subtype(HLDS, StartType, CONS_ID, 
+				CHOICE, SubType) 
+		; 
+			US = typesel(SubType)
+		),
+		check_type_of_node(HLDS, SubType, RestSelector)
+	;
+		Selector = []
+	).
+
+	% select the subtype of a type Type, selecting ConsId's position
+	% Position. Position counts starting from 1 (and not 0). 
+	% Predicate aborts if subtype cannot be determined. 
+    %
+:- pred select_subtype(module_info::in, mer_type::in, cons_id::in,
int::in, 
+    mer_type::out) is semidet.
+select_subtype(HLDS, Type, ConsID, Choice, SubType) :-
+	type_util.get_cons_id_non_existential_arg_types(HLDS, Type, ConsID, 
+        ArgTypes),
+	list.index1(ArgTypes, Choice, SubType).
+
+some_are_special_preds(PredProcIds, ModuleInfo):- 
+	module_info_get_special_pred_map(ModuleInfo, SpecialPredMap), 
+	map.values(SpecialPredMap, SpecialPreds), 
+	(
+            % either some of the predicates are special 
+            % preds, such as __Unify__ and others
+            %
+		list.filter(pred_id_in(SpecialPreds), PredProcIds, SpecialPredProcs),
+		SpecialPredProcs = [_|_]
+
+	; 
+            % or some of the predicates are not defined in this module. 
+            %
+		list.filter(not_defined_in_this_module(ModuleInfo),	PredProcIds,
+            FilteredPredProcIds), 
+		FilteredPredProcIds = [_|_]
+	).
+
+:- pred pred_id_in(list(pred_id)::in, pred_proc_id::in) is semidet.
+pred_id_in(PredIds, PredProcId):-
+	PredProcId = proc(PredId, _),
+	list.member(PredId, PredIds). 
+
+:- pred not_defined_in_this_module(module_info::in, pred_proc_id::in) 
+    is semidet.
+not_defined_in_this_module(ModuleInfo, proc(Predid, _)):-
+	pred_not_defined_in_this_module(ModuleInfo, Predid).
+
+:- pred pred_not_defined_in_this_module(module_info::in, pred_id::in) 
+    is semidet.
+pred_not_defined_in_this_module(HLDS, PredId) :-
+	module_info_pred_info(HLDS, PredId, PredInfo),
+	pred_info_get_import_status(PredInfo, Status),
+	status_defined_in_this_module(Status) = no.
+
+annotate_all_liveness_in_module(HLDSin, HLDSout, !IO) :-
+	module_info_predids(PredIds0, HLDSin, _), 
+	module_info_get_special_pred_map(HLDSin, SpecialPredMap),
+	map.values(SpecialPredMap, SpecialPredIds),
+	list.delete_elems(PredIds0, SpecialPredIds, PredIds),
+	list.foldl2(annotate_all_liveness_in_module_2, PredIds, HLDSin, HLDSout, 
+        !IO).
+
+:- pred annotate_all_liveness_in_module_2(pred_id::in, module_info::in, 
+		module_info::out, io::di, io::uo) is det.
+annotate_all_liveness_in_module_2(PredId, HLDSin, HLDSout, !IO) :-
+	module_info_pred_info(HLDSin, PredId, PredInfo0),
+	pred_info_non_imported_procids(PredInfo0) = ProcIds, 
+	list.foldl2(annotate_all_liveness_in_pred(PredId, HLDSin),
+			ProcIds, PredInfo0, PredInfo, !IO),
+	module_info_set_pred_info(PredId, PredInfo, HLDSin, HLDSout).
+
+:- pred annotate_all_liveness_in_pred(pred_id::in, module_info::in, 
+    proc_id::in, pred_info::in, pred_info::out, io::di, io::uo) is det.
+annotate_all_liveness_in_pred(PredId, HLDS, ProcId, PredInfo0, PredInfo, 
+    !IO) :-
+	pred_info_get_procedures(PredInfo0, Procedures0),
+	map.lookup(Procedures0, ProcId, ProcInfo0),
+	liveness.detect_liveness_proc(PredId, ProcId, HLDS, ProcInfo0, ProcInfo, 
+        !IO),
+	map.det_update(Procedures0, ProcId, ProcInfo, Procedures),
+	pred_info_set_procedures(Procedures, PredInfo0, PredInfo).
cvs diff: Diffing notes


--------------------------------------------------------------------------
mercury-reviews mailing list
Post messages to:       mercury-reviews at csse.unimelb.edu.au
Administrative Queries: owner-mercury-reviews at csse.unimelb.edu.au
Subscriptions:          mercury-reviews-request at csse.unimelb.edu.au
--------------------------------------------------------------------------



More information about the reviews mailing list